From 1dc4f0d0e7ba87e1e4d2497de69de08327c736e8 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 23 Jan 2024 23:49:09 +0530 Subject: [PATCH 01/61] Add soundness for propositional connectives --- src/Realizability/Tripos/Logic/Semantics.agda | 83 +++++++++++++++---- 1 file changed, 66 insertions(+), 17 deletions(-) diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index 123d1f0..c0232d4 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -123,32 +123,81 @@ module Interpretation substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i → substitutionTermSound subs t i x .snd substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i → f (substitutionTermSound subs t i x) - substitutionFormulaSound : ∀ {Γ Δ} → (subs : Substitution Γ Δ) → (f : Formula Δ) → ⟦ substitutionFormula subs f ⟧ᶠ ≡ ⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ Δ ⟧ᶜ) ⟦ subs ⟧ᴮ ⟦ f ⟧ᶠ + semanticSubstitution : ∀ {Γ Δ} → (subs : Substitution Γ Δ) → Predicate ⟨ ⟦ Δ ⟧ᶜ ⟩ → Predicate ⟨ ⟦ Γ ⟧ᶜ ⟩ + semanticSubstitution {Γ} {Δ} subs = ⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ Δ ⟧ᶜ) ⟦ subs ⟧ᴮ + + -- Due to a shortcut in the soundness of negation termination checking fails + -- TODO : Fix + {-# TERMINATING #-} + substitutionFormulaSound : ∀ {Γ Δ} → (subs : Substitution Γ Δ) → (f : Formula Δ) → ⟦ substitutionFormula subs f ⟧ᶠ ≡ semanticSubstitution subs ⟦ f ⟧ᶠ substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ = Predicate≡ ⟨ ⟦ Γ ⟧ᶜ ⟩ (pre1 ⟨ ⟦ Γ ⟧ᶜ ⟩ (str ⟦ Γ ⟧ᶜ) isNonTrivial) - (⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ Δ ⟧ᶜ) ⟦ subs ⟧ᴮ (pre1 ⟨ ⟦ Δ ⟧ᶜ ⟩ (str ⟦ Δ ⟧ᶜ) isNonTrivial)) - (λ ⟦Γ⟧ a tt* → tt*) - λ ⟦Γ⟧ a tt* → tt* - substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ = - Predicate≡ - ⟨ ⟦ Γ ⟧ᶜ ⟩ - (pre0 ⟨ ⟦ Γ ⟧ᶜ ⟩ (str ⟦ Γ ⟧ᶜ) isNonTrivial) - (⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ Δ ⟧ᶜ) ⟦ subs ⟧ᴮ (pre0 ⟨ ⟦ Δ ⟧ᶜ ⟩ (str ⟦ Δ ⟧ᶜ) isNonTrivial)) - (λ ⟦Γ⟧ a bot → ⊥rec* bot) - λ ⟦Γ⟧ a bot → bot + (semanticSubstitution subs (pre1 ⟨ ⟦ Δ ⟧ᶜ ⟩ (str ⟦ Δ ⟧ᶜ) isNonTrivial)) + (λ γ a a⊩1γ → tt*) + λ γ a a⊩1subsγ → tt* + substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ = {!!} substitutionFormulaSound {Γ} {Δ} subs (f `∨ f₁) = Predicate≡ ⟨ ⟦ Γ ⟧ᶜ ⟩ (_⊔_ ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦ substitutionFormula subs f ⟧ᶠ ⟦ substitutionFormula subs f₁ ⟧ᶠ) - (⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ Δ ⟧ᶜ) ⟦ subs ⟧ᴮ (_⊔_ ⟨ ⟦ Δ ⟧ᶜ ⟩ ⟦ f ⟧ᶠ ⟦ f₁ ⟧ᶠ)) - (λ ⟦Γ⟧ a a⊩f'⊔f₁' → {!!}) + (semanticSubstitution subs (_⊔_ ⟨ ⟦ Δ ⟧ᶜ ⟩ ⟦ f ⟧ᶠ ⟦ f₁ ⟧ᶠ)) + (λ γ a a⊩substFormFs → + a⊩substFormFs >>= + λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) → ∣ inl (pr₁a≡k , subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁ + ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) → ∣ inr (pr₁a≡k' , subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ }) + λ γ a a⊩semanticSubsFs → + a⊩semanticSubsFs >>= + λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) → ∣ inl (pr₁a≡k , (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁ + ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) → ∣ inr (pr₁a≡k' , (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ } + substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) = + Predicate≡ + ⟨ ⟦ Γ ⟧ᶜ ⟩ + (_⊓_ ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦ substitutionFormula subs f ⟧ᶠ ⟦ substitutionFormula subs f₁ ⟧ᶠ) + (semanticSubstitution subs (_⊓_ ⟨ ⟦ Δ ⟧ᶜ ⟩ ⟦ f ⟧ᶠ ⟦ f₁ ⟧ᶠ)) + (λ γ a a⊩substFormulaFs → + (subst (λ form → (pr₁ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f) (a⊩substFormulaFs .fst)) , + (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f₁) (a⊩substFormulaFs .snd))) + λ γ a a⊩semanticSubstFs → + (subst (λ form → (pr₁ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f)) (a⊩semanticSubstFs .fst)) , + (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f₁)) (a⊩semanticSubstFs .snd)) + substitutionFormulaSound {Γ} {Δ} subs (f `→ f₁) = + Predicate≡ + ⟨ ⟦ Γ ⟧ᶜ ⟩ + (_⇒_ ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦ substitutionFormula subs f ⟧ᶠ ⟦ substitutionFormula subs f₁ ⟧ᶠ) + (semanticSubstitution subs (_⇒_ ⟨ ⟦ Δ ⟧ᶜ ⟩ ⟦ f ⟧ᶠ ⟦ f₁ ⟧ᶠ)) + (λ γ a a⊩substFormulaFs → + λ b b⊩semanticSubstFs → + subst + (λ form → (a ⨾ b) ⊩ ∣ form ∣ γ) + (substitutionFormulaSound subs f₁) + (a⊩substFormulaFs + b + (subst + (λ form → b ⊩ ∣ form ∣ γ) + (sym (substitutionFormulaSound subs f)) + b⊩semanticSubstFs))) + λ γ a a⊩semanticSubstFs → + λ b b⊩substFormulaFs → + subst + (λ form → (a ⨾ b) ⊩ ∣ form ∣ γ) + (sym (substitutionFormulaSound subs f₁)) + (a⊩semanticSubstFs + b + (subst + (λ form → b ⊩ ∣ form ∣ γ) + (substitutionFormulaSound subs f) + b⊩substFormulaFs)) + substitutionFormulaSound {Γ} {Δ} subs (`¬ f) = + substitutionFormulaSound subs (f `→ ⊥ᵗ) + substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) = + Predicate≡ + ⟨ ⟦ Γ ⟧ᶜ ⟩ + (`∃[ isSet× (str ⟦ Γ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Γ ⟧ᶜ) (λ { (f , s) → f }) ⟦ substitutionFormula (var here , drop subs) f ⟧ᶠ) + (semanticSubstitution subs (`∃[ isSet× (str ⟦ Δ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Δ ⟧ᶜ) (λ { (γ , b) → γ }) ⟦ f ⟧ᶠ)) + (λ γ a a⊩πSubstFormulaF → {!!}) {!!} - substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) = {!!} - substitutionFormulaSound {Γ} {Δ} subs (f `→ f₁) = {!!} - substitutionFormulaSound {Γ} {Δ} subs (`¬ f) = {!!} - substitutionFormulaSound {Γ} {Δ} subs (`∃ f) = {!!} substitutionFormulaSound {Γ} {Δ} subs (`∀ f) = {!!} substitutionFormulaSound {Γ} {Δ} subs (rel k₁ x) = {!!} From 51c7bfcb55cdbb1f73331d826342973acb3d95bf Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 25 Jan 2024 00:37:54 +0530 Subject: [PATCH 02/61] Update case for existentials --- src/Realizability/Tripos/Logic/Semantics.agda | 26 ++++++++++--------- 1 file changed, 14 insertions(+), 12 deletions(-) diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index c0232d4..633b01a 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -196,7 +196,8 @@ module Interpretation ⟨ ⟦ Γ ⟧ᶜ ⟩ (`∃[ isSet× (str ⟦ Γ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Γ ⟧ᶜ) (λ { (f , s) → f }) ⟦ substitutionFormula (var here , drop subs) f ⟧ᶠ) (semanticSubstitution subs (`∃[ isSet× (str ⟦ Δ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Δ ⟧ᶜ) (λ { (γ , b) → γ }) ⟦ f ⟧ᶠ)) - (λ γ a a⊩πSubstFormulaF → {!!}) + (λ γ a a⊩πSubstFormulaF → + a⊩πSubstFormulaF >>= λ { (γ' , γ'≡γ , a⊩substFormFγ') → {!!} }) {!!} substitutionFormulaSound {Γ} {Δ} subs (`∀ f) = {!!} substitutionFormulaSound {Γ} {Δ} subs (rel k₁ x) = {!!} @@ -210,19 +211,20 @@ module Soundness open Interpretation relSym ⟦_⟧ʳ isNonTrivial infix 24 _⊨_ + + module PredProps = PredicateProperties + + _⊨_ : ∀ {Γ} → Formula Γ → Formula Γ → Type (ℓ-max (ℓ-max ℓ ℓ'') ℓ') + _⊨_ {Γ} ϕ ψ = ⟦ ϕ ⟧ᶠ ≤ ⟦ ψ ⟧ᶠ where open PredProps ⟨ ⟦ Γ ⟧ᶜ ⟩ - module _ {Γ : Context} where + private + variable + Γ Δ : Context + ϕ ψ : Formula Γ + θ χ : Formula Δ - open PredicateProperties {ℓ'' = ℓ''} ⟨ ⟦ Γ ⟧ᶜ ⟩ + cut : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ ψ → ψ ⊨ θ → ϕ ⊨ θ + cut {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤ ⟦ ϕ ⟧ᶠ ⟦ ψ ⟧ᶠ ⟦ θ ⟧ᶠ ϕ⊨ψ ψ⊨θ where open PredProps ⟨ ⟦ Γ ⟧ᶜ ⟩ - _⊨_ : Formula Γ → Formula Γ → Type _ - ϕ ⊨ ψ = ⟦ ϕ ⟧ᶠ ≤ ⟦ ψ ⟧ᶠ - - private - variable - ϕ ψ θ χ : Formula Γ - - cut : ∀ {ϕ ψ θ} → ϕ ⊨ ψ → ψ ⊨ θ → ϕ ⊨ θ - cut {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤ ⟦ ϕ ⟧ᶠ ⟦ ψ ⟧ᶠ ⟦ θ ⟧ᶠ ϕ⊨ψ ψ⊨θ From 98df49c2b80328910d13bbbfa00aa8fff41ad9ae Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 25 Jan 2024 19:08:00 +0530 Subject: [PATCH 03/61] Prove soundness of substitution for existential --- src/Realizability/Tripos/Logic/Semantics.agda | 24 +++++++++++++++---- 1 file changed, 19 insertions(+), 5 deletions(-) diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index 633b01a..12b2340 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -128,7 +128,6 @@ module Interpretation -- Due to a shortcut in the soundness of negation termination checking fails -- TODO : Fix - {-# TERMINATING #-} substitutionFormulaSound : ∀ {Γ Δ} → (subs : Substitution Γ Δ) → (f : Formula Δ) → ⟦ substitutionFormula subs f ⟧ᶠ ≡ semanticSubstitution subs ⟦ f ⟧ᶠ substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ = Predicate≡ @@ -190,17 +189,32 @@ module Interpretation (substitutionFormulaSound subs f) b⊩substFormulaFs)) substitutionFormulaSound {Γ} {Δ} subs (`¬ f) = - substitutionFormulaSound subs (f `→ ⊥ᵗ) + {!!} substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) = Predicate≡ ⟨ ⟦ Γ ⟧ᶜ ⟩ (`∃[ isSet× (str ⟦ Γ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Γ ⟧ᶜ) (λ { (f , s) → f }) ⟦ substitutionFormula (var here , drop subs) f ⟧ᶠ) (semanticSubstitution subs (`∃[ isSet× (str ⟦ Δ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Δ ⟧ᶜ) (λ { (γ , b) → γ }) ⟦ f ⟧ᶠ)) (λ γ a a⊩πSubstFormulaF → - a⊩πSubstFormulaF >>= λ { (γ' , γ'≡γ , a⊩substFormFγ') → {!!} }) - {!!} + a⊩πSubstFormulaF >>= + λ { ((γ' , b) , γ'≡γ , a⊩substFormFγ') → + ∣ ((⟦ subs ⟧ᴮ γ') , b) , + ((cong ⟦ subs ⟧ᴮ γ'≡γ) , + (subst + (λ form → a ⊩ ∣ form ∣ (γ' , b)) + (substitutionFormulaSound (var here , drop subs) f) + a⊩substFormFγ' )) ∣₁ }) + λ γ a a⊩semanticSubstF → + a⊩semanticSubstF >>= + λ (x@(δ , b) , δ≡subsγ , a⊩fx) → + ∣ (γ , b) , + (refl , + (subst + (λ form → a ⊩ ∣ form ∣ (γ , b)) + (sym (substitutionFormulaSound (var here , drop subs) f)) + (subst (λ x → a ⊩ ∣ ⟦ f ⟧ᶠ ∣ (x , b)) δ≡subsγ a⊩fx))) ∣₁ substitutionFormulaSound {Γ} {Δ} subs (`∀ f) = {!!} - substitutionFormulaSound {Γ} {Δ} subs (rel k₁ x) = {!!} + substitutionFormulaSound {Γ} {Δ} subs (rel R t) = {!!} module Soundness {n} From c6a18cafa60bd4faca93631ce9f3e771ce6c8e14 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Fri, 26 Jan 2024 00:39:38 +0530 Subject: [PATCH 04/61] Soundness of substitution and some combinators --- src/Realizability/Tripos/Logic/Semantics.agda | 126 +++++++++++++++--- 1 file changed, 110 insertions(+), 16 deletions(-) diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index 12b2340..b326f95 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -1,6 +1,6 @@ {-# OPTIONS --allow-unsolved-metas #-} open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv @@ -23,12 +23,15 @@ open import Cubical.Relation.Binary.Order.Preorder module Realizability.Tripos.Logic.Semantics {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open CombinatoryAlgebra ca +private λ*ComputationRule = `λ*ComputationRule as fefermanStructure +private λ* = `λ* as fefermanStructure + open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') open import Realizability.Tripos.Prealgebra.Predicate.Properties ca open import Realizability.Tripos.Prealgebra.Meets.Identity ca open import Realizability.Tripos.Prealgebra.Joins.Identity ca open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} -open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) open Predicate' open PredicateProperties hiding (_≤_ ; isTrans≤) @@ -107,8 +110,8 @@ module Interpretation -- TODO : Fix unsolved constraints funExt λ { x@(⟦Γ⟧ , ⟦s⟧) → - ⟦ substitutionVar (drop subs) t ⟧ᵗ (⟦Γ⟧ , ⟦s⟧) - ≡[ i ]⟨ renamingTermSound (drop id) (substitutionVar subs t) i (⟦Γ⟧ , ⟦s⟧) ⟩ + ⟦ substitutionVar (drop subs) t ⟧ᵗ x + ≡[ i ]⟨ renamingTermSound (drop id) (substitutionVar subs t) i x ⟩ ⟦ substitutionVar subs t ⟧ᵗ (⟦ drop id ⟧ᴿ x) ≡⟨ refl ⟩ ⟦ substitutionVar subs t ⟧ᵗ ⟦Γ⟧ @@ -128,6 +131,7 @@ module Interpretation -- Due to a shortcut in the soundness of negation termination checking fails -- TODO : Fix + {-# TERMINATING #-} substitutionFormulaSound : ∀ {Γ Δ} → (subs : Substitution Γ Δ) → (f : Formula Δ) → ⟦ substitutionFormula subs f ⟧ᶠ ≡ semanticSubstitution subs ⟦ f ⟧ᶠ substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ = Predicate≡ @@ -136,7 +140,13 @@ module Interpretation (semanticSubstitution subs (pre1 ⟨ ⟦ Δ ⟧ᶜ ⟩ (str ⟦ Δ ⟧ᶜ) isNonTrivial)) (λ γ a a⊩1γ → tt*) λ γ a a⊩1subsγ → tt* - substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ = {!!} + substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ = + Predicate≡ + ⟨ ⟦ Γ ⟧ᶜ ⟩ + (pre0 ⟨ ⟦ Γ ⟧ᶜ ⟩ (str ⟦ Γ ⟧ᶜ) isNonTrivial) + (semanticSubstitution subs (pre0 ⟨ ⟦ Δ ⟧ᶜ ⟩ (str ⟦ Δ ⟧ᶜ) isNonTrivial)) + (λ _ _ bot → ⊥rec* bot) + λ _ _ bot → bot substitutionFormulaSound {Γ} {Δ} subs (f `∨ f₁) = Predicate≡ ⟨ ⟦ Γ ⟧ᶜ ⟩ @@ -144,12 +154,16 @@ module Interpretation (semanticSubstitution subs (_⊔_ ⟨ ⟦ Δ ⟧ᶜ ⟩ ⟦ f ⟧ᶠ ⟦ f₁ ⟧ᶠ)) (λ γ a a⊩substFormFs → a⊩substFormFs >>= - λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) → ∣ inl (pr₁a≡k , subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁ - ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) → ∣ inr (pr₁a≡k' , subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ }) + λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) → + ∣ inl (pr₁a≡k , subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁ + ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) → + ∣ inr (pr₁a≡k' , subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ }) λ γ a a⊩semanticSubsFs → a⊩semanticSubsFs >>= - λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) → ∣ inl (pr₁a≡k , (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁ - ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) → ∣ inr (pr₁a≡k' , (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ } + λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) → + ∣ inl (pr₁a≡k , (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁ + ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) → + ∣ inr (pr₁a≡k' , (subst (λ form → (pr₂ ⨾ a) ⊩ ∣ form ∣ γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ } substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) = Predicate≡ ⟨ ⟦ Γ ⟧ᶜ ⟩ @@ -189,7 +203,7 @@ module Interpretation (substitutionFormulaSound subs f) b⊩substFormulaFs)) substitutionFormulaSound {Γ} {Δ} subs (`¬ f) = - {!!} + substitutionFormulaSound subs (f `→ ⊥ᵗ) substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) = Predicate≡ ⟨ ⟦ Γ ⟧ᶜ ⟩ @@ -213,8 +227,38 @@ module Interpretation (λ form → a ⊩ ∣ form ∣ (γ , b)) (sym (substitutionFormulaSound (var here , drop subs) f)) (subst (λ x → a ⊩ ∣ ⟦ f ⟧ᶠ ∣ (x , b)) δ≡subsγ a⊩fx))) ∣₁ - substitutionFormulaSound {Γ} {Δ} subs (`∀ f) = {!!} - substitutionFormulaSound {Γ} {Δ} subs (rel R t) = {!!} + substitutionFormulaSound {Γ} {Δ} subs (`∀ {B = B} f) = + Predicate≡ + ⟨ ⟦ Γ ⟧ᶜ ⟩ + (`∀[ isSet× (str ⟦ Γ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Γ ⟧ᶜ) (λ { (f , s) → f }) ⟦ substitutionFormula (var here , drop subs) f ⟧ᶠ) + (semanticSubstitution subs (`∀[ isSet× (str ⟦ Δ ⟧ᶜ) (str ⟦ B ⟧ˢ) ] (str ⟦ Δ ⟧ᶜ) (λ { (f , s) → f }) ⟦ f ⟧ᶠ)) + (λ γ a a⊩substFormF → + λ { r x@(δ , b) δ≡subsγ → + subst + (λ g → (a ⨾ r) ⊩ ∣ ⟦ f ⟧ᶠ ∣ (g , b)) + (sym δ≡subsγ) + (subst + (λ form → (a ⨾ r) ⊩ ∣ form ∣ (γ , b)) + (substitutionFormulaSound (var here , drop subs) f) + (a⊩substFormF r (γ , b) refl)) }) + λ γ a a⊩semanticSubsF → + λ { r x@(γ' , b) γ'≡γ → + subst + (λ form → (a ⨾ r) ⊩ ∣ form ∣ (γ' , b)) + (sym (substitutionFormulaSound (var here , drop subs) f)) + (subst + (λ g → (a ⨾ r) ⊩ ∣ ⟦ f ⟧ᶠ ∣ (g , b)) + (cong ⟦ subs ⟧ᴮ (sym γ'≡γ)) + (a⊩semanticSubsF r (⟦ subs ⟧ᴮ γ , b) refl)) } + substitutionFormulaSound {Γ} {Δ} subs (rel R t) = + Predicate≡ + ⟨ ⟦ Γ ⟧ᶜ ⟩ + (⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ lookup R relSym ⟧ˢ) ⟦ substitutionTerm subs t ⟧ᵗ ⟦ R ⟧ʳ) + (semanticSubstitution subs (⋆_ (str ⟦ Δ ⟧ᶜ) (str ⟦ lookup R relSym ⟧ˢ) ⟦ t ⟧ᵗ ⟦ R ⟧ʳ)) + (λ γ a a⊩substTR → + subst (λ transform → a ⊩ ∣ ⟦ R ⟧ʳ ∣ (transform γ)) (substitutionTermSound subs t) a⊩substTR) + λ γ a a⊩semSubst → + subst (λ transform → a ⊩ ∣ ⟦ R ⟧ʳ ∣ (transform γ)) (sym (substitutionTermSound subs t)) a⊩semSubst module Soundness {n} @@ -223,22 +267,72 @@ module Soundness (⟦_⟧ʳ : RelationInterpretation relSym) where open Relational relSym open Interpretation relSym ⟦_⟧ʳ isNonTrivial - - infix 24 _⊨_ + -- Acknowledgements : 1lab's "the internal logic of a regular hyperdoctrine" + infix 35 _⊨_ module PredProps = PredicateProperties _⊨_ : ∀ {Γ} → Formula Γ → Formula Γ → Type (ℓ-max (ℓ-max ℓ ℓ'') ℓ') _⊨_ {Γ} ϕ ψ = ⟦ ϕ ⟧ᶠ ≤ ⟦ ψ ⟧ᶠ where open PredProps ⟨ ⟦ Γ ⟧ᶜ ⟩ + entails = _⊨_ + private variable Γ Δ : Context - ϕ ψ : Formula Γ - θ χ : Formula Δ + ϕ ψ θ : Formula Γ + χ μ ν : Formula Δ cut : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ ψ → ψ ⊨ θ → ϕ ⊨ θ cut {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤ ⟦ ϕ ⟧ᶠ ⟦ ψ ⟧ᶠ ⟦ θ ⟧ᶠ ϕ⊨ψ ψ⊨θ where open PredProps ⟨ ⟦ Γ ⟧ᶜ ⟩ + substitutionEntailment : ∀ {Γ Δ} (subs : Substitution Γ Δ) → {ϕ ψ : Formula Δ} → ϕ ⊨ ψ → substitutionFormula subs ϕ ⊨ substitutionFormula subs ψ + substitutionEntailment {Γ} {Δ} subs {ϕ} {ψ} ϕ⊨ψ = + subst2 + (λ ϕ' ψ' → ϕ' ≤Γ ψ') + (sym (substitutionFormulaSound subs ϕ)) + (sym (substitutionFormulaSound subs ψ)) + (ϕ⊨ψ >>= + λ { (a , a⊩ϕ≤ψ) → + ∣ a , (λ γ b b⊩ϕsubsγ → a⊩ϕ≤ψ (⟦ subs ⟧ᴮ γ) b b⊩ϕsubsγ) ∣₁ }) where + open PredProps {ℓ'' = ℓ''} ⟨ ⟦ Γ ⟧ᶜ ⟩ renaming (_≤_ to _≤Γ_) + open PredProps {ℓ'' = ℓ''} ⟨ ⟦ Δ ⟧ᶜ ⟩ renaming (_≤_ to _≤Δ_) + + `∧intro : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ ψ → entails ϕ θ → entails ϕ (ψ `∧ θ) + `∧intro {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ϕ⊨θ = + do + (a , a⊩ϕ⊨ψ) ← ϕ⊨ψ + (b , b⊩ϕ⊨θ) ← ϕ⊨θ + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` a ̇ # fzero) ̇ (` b ̇ # fzero) + return + (λ* prover , + λ γ r r⊩ϕγ → + let + proofEq : λ* prover ⨾ r ≡ pair ⨾ (a ⨾ r) ⨾ (b ⨾ r) + proofEq = λ*ComputationRule prover (r ∷ []) + + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ a ⨾ r + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) proofEq ⟩ + pr₁ ⨾ (pair ⨾ (a ⨾ r) ⨾ (b ⨾ r)) + ≡⟨ pr₁pxy≡x _ _ ⟩ + a ⨾ r + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ b ⨾ r + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₂ ⨾ x) proofEq ⟩ + pr₂ ⨾ (pair ⨾ (a ⨾ r) ⨾ (b ⨾ r)) + ≡⟨ pr₂pxy≡y _ _ ⟩ + b ⨾ r + ∎ + in + subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) (sym pr₁proofEq) (a⊩ϕ⊨ψ γ r r⊩ϕγ) , + subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym pr₂proofEq) (b⊩ϕ⊨θ γ r r⊩ϕγ)) + From eb9a962e181d5408924f99eed76173325db17fb1 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 29 Jan 2024 16:41:16 +0530 Subject: [PATCH 05/61] Interpret quantifiers --- src/Realizability/Tripos/Logic/Semantics.agda | 86 ++++++++++++++++++- src/Realizability/Tripos/Logic/Syntax.agda | 2 + 2 files changed, 87 insertions(+), 1 deletion(-) diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index b326f95..a194d13 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -260,6 +260,13 @@ module Interpretation λ γ a a⊩semSubst → subst (λ transform → a ⊩ ∣ ⟦ R ⟧ʳ ∣ (transform γ)) (sym (substitutionTermSound subs t)) a⊩semSubst + weakenFormulaMonotonic : ∀ {Γ B} → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) → (ϕ : Formula Γ) → (a : A) → (b : ⟨ ⟦ B ⟧ˢ ⟩) → a ⊩ ∣ ⟦ ϕ ⟧ᶠ ∣ γ ≡ a ⊩ ∣ ⟦ weakenFormula {S = B} ϕ ⟧ᶠ ∣ (γ , b) + weakenFormulaMonotonic {Γ} {B} γ ϕ a b = + hPropExt + (⟦ ϕ ⟧ᶠ .isPropValued γ a) + (⟦ weakenFormula ϕ ⟧ᶠ .isPropValued (γ , b) a) + (λ a⊩ϕγ → subst (λ form → a ⊩ ∣ form ∣ (γ , b)) (sym (substitutionFormulaSound (drop id) ϕ)) a⊩ϕγ) + λ a⊩weakenϕγb → subst (λ form → a ⊩ ∣ form ∣ (γ , b)) (substitutionFormulaSound (drop id) ϕ) a⊩weakenϕγb module Soundness {n} {relSym : Vec Sort n} @@ -334,5 +341,82 @@ module Soundness subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) (sym pr₁proofEq) (a⊩ϕ⊨ψ γ r r⊩ϕγ) , subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym pr₂proofEq) (b⊩ϕ⊨θ γ r r⊩ϕγ)) + `∧elim1 : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ (ψ `∧ θ) → ϕ ⊨ ψ + `∧elim1 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ = + do + (a , a⊩ϕ⊨ψ∧θ) ← ϕ⊨ψ∧θ + let + prover : ApplStrTerm as 1 + prover = ` pr₁ ̇ (` a ̇ # fzero) + return + (λ* prover , + λ γ b b⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) (sym (λ*ComputationRule prover (b ∷ []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst)) + + `∧elim2 : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ (ψ `∧ θ) → ϕ ⊨ θ + `∧elim2 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ = + do + (a , a⊩ϕ⊨ψ∧θ) ← ϕ⊨ψ∧θ + let + prover : ApplStrTerm as 1 + prover = ` pr₂ ̇ (` a ̇ # fzero) + return + (λ* prover , + λ γ b b⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym (λ*ComputationRule prover (b ∷ []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd)) + + `∃intro : ∀ {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ ′ B)} {t : Term Γ B} → ϕ ⊨ substitutionFormula (t , id) ψ → ϕ ⊨ `∃ ψ + `∃intro {Γ} {ϕ} {B} {ψ} {t} ϕ⊨ψ[t/x] = + do + (a , a⊩ϕ⊨ψ[t/x]) ← ϕ⊨ψ[t/x] + return + (a , (λ γ b b⊩ϕγ → ∣ (γ , (⟦ t ⟧ᵗ γ)) , + (refl , (subst (λ form → (a ⨾ b) ⊩ ∣ form ∣ γ) (substitutionFormulaSound (t , id) ψ) (a⊩ϕ⊨ψ[t/x] γ b b⊩ϕγ))) ∣₁)) + + `∃elim : ∀ {Γ} {ϕ θ : Formula Γ} {B} {ψ : Formula (Γ ′ B)} → ϕ ⊨ `∃ ψ → (weakenFormula ϕ `∧ ψ) ⊨ weakenFormula θ → ϕ ⊨ θ + `∃elim {Γ} {ϕ} {θ} {B} {ψ} ϕ⊨∃ψ ϕ∧ψ⊨θ = + do + (a , a⊩ϕ⊨∃ψ) ← ϕ⊨∃ψ + (b , b⊩ϕ∧ψ⊨θ) ← ϕ∧ψ⊨θ + let + prover : ApplStrTerm as 1 + prover = ` b ̇ (` pair ̇ # fzero ̇ (` a ̇ # fzero)) + return + (λ* prover , + (λ γ c c⊩ϕγ → + subst + (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) + (sym (λ*ComputationRule prover (c ∷ []))) + (transport + (propTruncIdempotent (⟦ θ ⟧ᶠ .isPropValued γ (b ⨾ (pair ⨾ c ⨾ (a ⨾ c))))) + (a⊩ϕ⊨∃ψ γ c c⊩ϕγ >>= + λ { (x@(γ' , b') , (γ'≡γ , a⨾c⊩ψx)) → + ∣ transport + (sym + (weakenFormulaMonotonic γ θ (b ⨾ (pair ⨾ c ⨾ (a ⨾ c))) b')) + (b⊩ϕ∧ψ⊨θ + (γ , b') + (pair ⨾ c ⨾ (a ⨾ c)) + (subst + (λ r → r ⊩ ∣ ⟦ weakenFormula ϕ ⟧ᶠ ∣ (γ , b')) + (sym (pr₁pxy≡x _ _)) + (transport + (weakenFormulaMonotonic γ ϕ c b') c⊩ϕγ) , + subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ (γ , b')) (sym (pr₂pxy≡y _ _)) (subst (λ g → (a ⨾ c) ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ (g , b')) γ'≡γ a⨾c⊩ψx)) ) ∣₁ })))) + + `∀intro : ∀ {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ ′ B)} → weakenFormula ϕ ⊨ ψ → ϕ ⊨ `∀ ψ + `∀intro {Γ} {ϕ} {B} {ψ} ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ) ← ϕ⊨ψ + let + prover : ApplStrTerm as 2 + prover = ` a ̇ # fzero + return + (λ* prover , + (λ γ b b⊩ϕ → λ { c x@(γ' , b') γ'≡γ → + subst + (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ (γ' , b')) + (sym (λ*ComputationRule prover (b ∷ c ∷ []))) + (a⊩ϕ⊨ψ + (γ' , b') + b + (transport (weakenFormulaMonotonic γ' ϕ b b') (subst (λ g → b ⊩ ∣ ⟦ ϕ ⟧ᶠ ∣ g) (sym γ'≡γ) b⊩ϕ))) })) - diff --git a/src/Realizability/Tripos/Logic/Syntax.agda b/src/Realizability/Tripos/Logic/Syntax.agda index 61f0973..e04c2b0 100644 --- a/src/Realizability/Tripos/Logic/Syntax.agda +++ b/src/Realizability/Tripos/Logic/Syntax.agda @@ -109,3 +109,5 @@ module Relational {n} (relSym : Vec Sort n) where substitutionFormula {Γ} {Δ} subs (`∀ form) = `∀ (substitutionFormula (var here , drop subs) form) substitutionFormula {Γ} {Δ} subs (rel k x) = rel k (substitutionTerm subs x) + weakenFormula : ∀ {Γ} {S} → Formula Γ → Formula (Γ ′ S) + weakenFormula {Γ} {S} form = substitutionFormula (drop id) form From 7223be450526de4780d19f7a6d46de54660f888d Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 30 Jan 2024 17:55:43 +0530 Subject: [PATCH 06/61] Define topos objects --- src/Realizability/Topos/Object.agda | 117 +++++++++++ src/Realizability/Tripos/Logic/Semantics.agda | 197 +++++++++++++++++- .../Tripos/Prealgebra/Implication.agda | 5 +- 3 files changed, 307 insertions(+), 12 deletions(-) create mode 100644 src/Realizability/Topos/Object.agda diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda new file mode 100644 index 0000000..72dcff7 --- /dev/null +++ b/src/Realizability/Topos/Object.agda @@ -0,0 +1,117 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Data.Vec +open import Cubical.Data.Nat +open import Cubical.Data.FinData renaming (zero to fzero) +open import Cubical.Data.Sigma +open import Cubical.Data.Empty + +module Realizability.Topos.Object + {ℓ ℓ' ℓ''} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + where + +open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} +open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_) +open import Realizability.Tripos.Prealgebra.Predicate.Properties ca +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate' renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +private + _⊩_ : ∀ a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) → Type _ + a ⊩ P = a pre⊩ ∣ P ∣ tt* + +record ToposObject (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + field + isSetX : isSet X + `X : Sort + `X = ↑ (X , isSetX) + + `X×X : Sort + `X×X = `X `× `X + + ~relSym : Vec Sort 1 + ~relSym = `X×X ∷ [] + + module X×XRelational = Relational ~relSym + open X×XRelational + + field + _~_ : Predicate ⟨ ⟦ lookup fzero ~relSym ⟧ˢ ⟩ + + ~relSymInterpretation : RelationInterpretation ~relSym + ~relSymInterpretation = λ { fzero → _~_ } + + module ~Interpretation = Interpretation ~relSym ~relSymInterpretation isNonTrivial + open ~Interpretation + + module ~Soundness = Soundness isNonTrivial ~relSymInterpretation + open ~Soundness + + -- Partial equivalence relations + symContext : Context + symContext = ([] ′ `X) ′ `X + + x∈symContext : `X ∈ symContext + x∈symContext = there here + + y∈symContext : `X ∈ symContext + y∈symContext = here + + xˢ : Term symContext `X + xˢ = var x∈symContext + + yˢ : Term symContext `X + yˢ = var y∈symContext + + symPrequantFormula : Formula symContext + symPrequantFormula = rel fzero (xˢ `, yˢ) `→ rel fzero (yˢ `, xˢ) + + -- ~ ⊢ ∀ x. ∀ y. x ~ y → y ~ x + symFormula : Formula [] + symFormula = `∀ (`∀ symPrequantFormula) + + field + symWitness : A + symIsWitnessed : symWitness ⊩ ⟦ symFormula ⟧ᶠ + + transContext : Context + transContext = (([] ′ `X) ′ `X) ′ `X + + z∈transContext : `X ∈ transContext + z∈transContext = here + + y∈transContext : `X ∈ transContext + y∈transContext = there here + + x∈transContext : `X ∈ transContext + x∈transContext = there (there here) + + zᵗ : Term transContext `X + zᵗ = var z∈transContext + + yᵗ : Term transContext `X + yᵗ = var y∈transContext + + xᵗ : Term transContext `X + xᵗ = var x∈transContext + + -- ~ : Rel X X, x : X, y : X, z : X ⊢ x ~ y ∧ y ~ z → x ~ z + transPrequantFormula : Formula transContext + transPrequantFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) + + transFormula : Formula [] + transFormula = `∀ (`∀ (`∀ transPrequantFormula)) + + field + transWitness : A + transIsWitnessed : transWitness ⊩ ⟦ transFormula ⟧ᶠ + diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index a194d13..1ffe530 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -31,6 +31,7 @@ open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicat open import Realizability.Tripos.Prealgebra.Predicate.Properties ca open import Realizability.Tripos.Prealgebra.Meets.Identity ca open import Realizability.Tripos.Prealgebra.Joins.Identity ca +open import Realizability.Tripos.Prealgebra.Implication ca open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) open Predicate' @@ -91,16 +92,26 @@ module Interpretation renamingTermSound : ∀ {Γ Δ s} → (ren : Renaming Γ Δ) → (t : Term Δ s) → ⟦ renamingTerm ren t ⟧ᵗ ≡ ⟦ t ⟧ᵗ ∘ ⟦ ren ⟧ᴿ renamingTermSound {Γ} {.Γ} {s} id t = refl - renamingTermSound {.(_ ′ _)} {Δ} {s} (drop ren) (var x) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren x i ⟦Γ⟧ } - renamingTermSound {.(_ ′ _)} {Δ} {.(_ `× _)} r@(drop ren) (t `, t₁) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) } - renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₁ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } - renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₂ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } - renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (fun f t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (var v) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingVarSound r v i x } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {.(_ `× _)} r@(keep ren) (t `, t₁) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₁ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₂ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (fun f t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } + renamingTermSound {.(_ ′ _)} {Δ} {s} (drop ren) (var x) = + funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren x i ⟦Γ⟧ } + renamingTermSound {.(_ ′ _)} {Δ} {.(_ `× _)} r@(drop ren) (t `, t₁) = + funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) } + renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₁ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } + renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₂ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } + renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (fun f t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } + renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (var v) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingVarSound r v i x } + renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {.(_ `× _)} r@(keep ren) (t `, t₁) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) } + renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₁ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } + renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₂ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } + renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (fun f t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } substitutionVarSound : ∀ {Γ Δ s} → (subs : Substitution Γ Δ) → (v : s ∈ Δ) → ⟦ substitutionVar subs v ⟧ᵗ ≡ ⟦ v ⟧ⁿ ∘ ⟦ subs ⟧ᴮ substitutionVarSound {Γ} {.Γ} {s} id t = refl @@ -420,3 +431,169 @@ module Soundness b (transport (weakenFormulaMonotonic γ' ϕ b b') (subst (λ g → b ⊩ ∣ ⟦ ϕ ⟧ᶠ ∣ g) (sym γ'≡γ) b⊩ϕ))) })) + `∀elim : ∀ {Γ} {B} {ϕ : Formula Γ} {ψ : Formula (Γ ′ B)} → ϕ ⊨ `∀ ψ → (t : Term Γ B) → ϕ ⊨ substitutionFormula (t , id) ψ + `∀elim {Γ} {B} {ϕ} {ψ} ϕ⊨∀ψ t = + do + (a , a⊩ϕ⊨∀ψ) ← ϕ⊨∀ψ + let + prover : ApplStrTerm as 1 + prover = ` a ̇ # fzero ̇ ` k + return + (λ* prover , + (λ γ b b⊩ϕγ → + subst + (λ form → (λ* prover ⨾ b) ⊩ ∣ form ∣ γ) + (sym (substitutionFormulaSound (t , id) ψ)) + (subst + (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ (γ , ⟦ t ⟧ᵗ γ)) + (sym (λ*ComputationRule prover (b ∷ []))) + (a⊩ϕ⊨∀ψ γ b b⊩ϕγ k (γ , ⟦ t ⟧ᵗ γ) refl)))) + + `→intro : ∀ {Γ} {ϕ ψ θ : Formula Γ} → (ϕ `∧ ψ) ⊨ θ → ϕ ⊨ (ψ `→ θ) + `→intro {Γ} {ϕ} {ψ} {θ} ϕ∧ψ⊨θ = a⊓b≤c→a≤b⇒c ⟨ ⟦ Γ ⟧ᶜ ⟩ (str ⟦ Γ ⟧ᶜ) ⟦ ϕ ⟧ᶠ ⟦ ψ ⟧ᶠ ⟦ θ ⟧ᶠ ϕ∧ψ⊨θ + + `→elim : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ (ψ `→ θ) → ϕ ⊨ ψ → ϕ ⊨ θ + `→elim {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ→θ ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ→θ) ← ϕ⊨ψ→θ + (b , b⊩ϕ⊨ψ) ← ϕ⊨ψ + let + prover : ApplStrTerm as 1 + prover = ` a ̇ (# fzero) ̇ (` b ̇ # fzero) + return + (λ* prover , + (λ γ c c⊩ϕγ → + subst + (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) + (sym (λ*ComputationRule prover (c ∷ []))) + (a⊩ϕ⊨ψ→θ γ c c⊩ϕγ (b ⨾ c) (b⊩ϕ⊨ψ γ c c⊩ϕγ)))) + + `∨introR : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ ψ → ϕ ⊨ (ψ `∨ θ) + `∨introR {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ) ← ϕ⊨ψ + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ ` true ̇ (` a ̇ # fzero) + return + ((λ* prover) , + (λ γ b b⊩ϕγ → + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ b) ≡ true + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ b) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ true ⨾ (a ⨾ b)) + ≡⟨ pr₁pxy≡x _ _ ⟩ + true + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ b) ≡ a ⨾ b + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ b) + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + pr₂ ⨾ (pair ⨾ true ⨾ (a ⨾ b)) + ≡⟨ pr₂pxy≡y _ _ ⟩ + a ⨾ b + ∎ + in ∣ inl (pr₁proofEq , subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁)) + + `∨introL : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ ψ → ϕ ⊨ (θ `∨ ψ) + `∨introL {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ) ← ϕ⊨ψ + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ ` false ̇ (` a ̇ # fzero) + return + ((λ* prover) , + (λ γ b b⊩ϕγ → + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ b) ≡ false + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ b) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ false ⨾ (a ⨾ b)) + ≡⟨ pr₁pxy≡x _ _ ⟩ + false + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ b) ≡ a ⨾ b + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ b) + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + pr₂ ⨾ (pair ⨾ false ⨾ (a ⨾ b)) + ≡⟨ pr₂pxy≡y _ _ ⟩ + a ⨾ b + ∎ + in ∣ inr (pr₁proofEq , subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁)) + + -- Pretty sure this is code duplication + `if_then_else_ : ∀ {as n} → ApplStrTerm as n → ApplStrTerm as n → ApplStrTerm as n → ApplStrTerm as n + `if a then b else c = ` Id ̇ a ̇ b ̇ c + + `∨elim : ∀ {Γ} {ϕ ψ θ χ : Formula Γ} → (ϕ `∧ ψ) ⊨ χ → (ϕ `∧ θ) ⊨ χ → (ϕ `∧ (ψ `∨ θ)) ⊨ χ + `∨elim {Γ} {ϕ} {ψ} {θ} {χ} ϕ∧ψ⊨χ ϕ∧θ⊨χ = + do + (a , a⊩ϕ∧ψ⊨χ) ← ϕ∧ψ⊨χ + (b , b⊩ϕ∧θ⊨χ) ← ϕ∧θ⊨χ + let + prover : ApplStrTerm as 1 + prover = + (`if ` pr₁ ̇ (` pr₂ ̇ # fzero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) + return + (λ* prover , + (λ + { γ c (pr₁⨾c⊩ϕγ , pr₂⨾c⊩ψ∨θ) → + transport + (propTruncIdempotent (⟦ χ ⟧ᶠ .isPropValued γ (λ* prover ⨾ c))) + (pr₂⨾c⊩ψ∨θ >>= + λ { (inl (pr₁⨾pr₂⨾c≡true , pr₂⨾pr₂⨾c⊩ψ)) → + let + proofEq : λ* prover ⨾ c ≡ a ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + proofEq = + λ* prover ⨾ c + ≡⟨ λ*ComputationRule prover (c ∷ []) ⟩ + (if (pr₁ ⨾ (pr₂ ⨾ c)) then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + ≡⟨ cong (λ x → (if x then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c)))) pr₁⨾pr₂⨾c≡true ⟩ + (if true then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + ≡⟨ cong (λ x → x ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c)))) (ifTrueThen a b) ⟩ + a ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + ∎ + in + ∣ subst + (λ r → r ⊩ ∣ ⟦ χ ⟧ᶠ ∣ γ) + (sym proofEq) + (a⊩ϕ∧ψ⊨χ + γ + (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + (( + subst + (λ r → r ⊩ ∣ ⟦ ϕ ⟧ᶠ ∣ γ) + (sym (pr₁pxy≡x _ _)) + pr₁⨾c⊩ϕγ) , + subst + (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) + (sym (pr₂pxy≡y _ _)) + pr₂⨾pr₂⨾c⊩ψ)) ∣₁ + ; (inr (pr₁pr₂⨾c≡false , pr₂⨾pr₂⨾c⊩θ)) → + let + proofEq : λ* prover ⨾ c ≡ b ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + proofEq = + λ* prover ⨾ c + ≡⟨ λ*ComputationRule prover (c ∷ []) ⟩ + (if (pr₁ ⨾ (pr₂ ⨾ c)) then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + ≡⟨ cong (λ x → (if x then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c)))) pr₁pr₂⨾c≡false ⟩ + (if false then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + ≡⟨ cong (λ x → x ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c)))) (ifFalseElse a b) ⟩ + b ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + ∎ + in + ∣ subst + (λ r → r ⊩ ∣ ⟦ χ ⟧ᶠ ∣ γ) + (sym proofEq) + (b⊩ϕ∧θ⊨χ + γ + (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) + ((subst (λ r → r ⊩ ∣ ⟦ ϕ ⟧ᶠ ∣ γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) , + (subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) })) diff --git a/src/Realizability/Tripos/Prealgebra/Implication.agda b/src/Realizability/Tripos/Prealgebra/Implication.agda index 374c42f..edeb3f2 100644 --- a/src/Realizability/Tripos/Prealgebra/Implication.agda +++ b/src/Realizability/Tripos/Prealgebra/Implication.agda @@ -14,8 +14,9 @@ open import Realizability.Tripos.Prealgebra.Predicate ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -λ*ComputationRule = `λ*ComputationRule as fefermanStructure -λ* = `λ* as fefermanStructure +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where PredicateX = Predicate {ℓ'' = ℓ''} X From 940dcfe90069c29b2ac38638b3c1df71a3e01807 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 30 Jan 2024 20:12:35 +0530 Subject: [PATCH 07/61] Render HTML --- docs/Realizability.Topos.Everything.html | 6 + ...ealizability.Topos.FunctionalRelation.html | 92 +++ docs/Realizability.Topos.Object.html | 124 +++ .../Realizability.Tripos.Logic.Semantics.html | 782 +++++++++++++----- docs/Realizability.Tripos.Logic.Syntax.html | 15 +- ...ability.Tripos.Prealgebra.Implication.html | 135 +-- docs/index.html | 20 +- src/Realizability/Topos/Everything.agda | 4 + .../Topos/FunctionalRelation.agda | 90 ++ src/Realizability/Topos/Object.agda | 83 +- src/index.agda | 4 +- 11 files changed, 1049 insertions(+), 306 deletions(-) create mode 100644 docs/Realizability.Topos.Everything.html create mode 100644 docs/Realizability.Topos.FunctionalRelation.html create mode 100644 docs/Realizability.Topos.Object.html create mode 100644 src/Realizability/Topos/Everything.agda create mode 100644 src/Realizability/Topos/FunctionalRelation.agda diff --git a/docs/Realizability.Topos.Everything.html b/docs/Realizability.Topos.Everything.html new file mode 100644 index 0000000..8a2ca6a --- /dev/null +++ b/docs/Realizability.Topos.Everything.html @@ -0,0 +1,6 @@ + +Realizability.Topos.Everything
module Realizability.Topos.Everything where
+
+open import Realizability.Topos.Object
+open import Realizability.Topos.FunctionalRelation
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.FunctionalRelation.html b/docs/Realizability.Topos.FunctionalRelation.html new file mode 100644 index 0000000..3e4946c --- /dev/null +++ b/docs/Realizability.Topos.FunctionalRelation.html @@ -0,0 +1,92 @@ + +Realizability.Topos.FunctionalRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData renaming (zero to fzero)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+
+module Realizability.Topos.FunctionalRelation
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
+
+open import Realizability.Tripos.Logic.Syntax { = ℓ'}
+open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_)
+open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate' renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+private
+  _⊩_ :  a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*)  Type _
+  a  P = a pre⊩  P  tt*
+
+record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  open PartialEquivalenceRelation
+  field
+    perX : PartialEquivalenceRelation X
+    perY : PartialEquivalenceRelation Y
+  _~X_ = perX ._~_
+  _~Y_ = perY ._~_
+
+  field
+    relation : Predicate (X × Y)
+
+  private
+    `X : Sort
+    `X =  (X , perX .isSetX)
+
+    `Y : Sort
+    `Y =  (Y , perY .isSetX)
+
+    relationSymbol : Vec Sort 3
+    relationSymbol = (`X  `Y)  `X  `X  `Y  `Y  []
+
+    relationSymbolInterpretation : RelationInterpretation relationSymbol
+    relationSymbolInterpretation fzero = relation
+    relationSymbolInterpretation one = _~X_
+    relationSymbolInterpretation two = _~Y_
+
+    `relation : Fin 3
+    `relation = fzero
+    `~X : Fin 3
+    `~X = one
+    `~Y : Fin 3
+    `~Y = two
+
+  module RelationInterpretation = Interpretation relationSymbol relationSymbolInterpretation isNonTrivial
+  open RelationInterpretation
+
+  module RelationSoundness = Soundness isNonTrivial relationSymbolInterpretation
+  open RelationSoundness
+
+  -- Strictness
+  -- If the functional relation holds for x and y then x and y "exist"
+  private
+    strictnessContext : Context
+    strictnessContext = ([]  `X)  `Y
+
+    x∈strictnessContext : `X  strictnessContext
+    x∈strictnessContext = there here
+
+    y∈strictnessContext : `Y  strictnessContext
+    y∈strictnessContext = here
+
+    xˢᵗ : Term strictnessContext `X
+    xˢᵗ = var x∈strictnessContext
+
+    yˢᵗ : Term strictnessContext `Y
+    yˢᵗ = var y∈strictnessContext
+  
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.Object.html b/docs/Realizability.Topos.Object.html new file mode 100644 index 0000000..ac46ab0 --- /dev/null +++ b/docs/Realizability.Topos.Object.html @@ -0,0 +1,124 @@ + +Realizability.Topos.Object
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData renaming (zero to fzero)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+
+module Realizability.Topos.Object
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
+
+open import Realizability.Tripos.Logic.Syntax { = ℓ'}
+open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_)
+open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate' renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+private
+  _⊩_ :  a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*)  Type _
+  a  P = a pre⊩  P  tt*
+
+record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  field
+    isSetX : isSet X
+  private
+    `X : Sort
+    `X =  (X , isSetX)
+
+    `X×X : Sort
+    `X×X = `X  `X
+
+    ~relSym : Vec Sort 1
+    ~relSym = `X×X  []
+
+  module X×XRelational = Relational ~relSym
+  open X×XRelational
+
+  field
+    _~_ : Predicate   lookup fzero ~relSym ⟧ˢ 
+
+  private
+    ~relSymInterpretation : RelationInterpretation ~relSym
+    ~relSymInterpretation = λ { fzero  _~_ }
+  
+  module ~Interpretation = Interpretation ~relSym ~relSymInterpretation isNonTrivial
+  open ~Interpretation
+
+  module ~Soundness = Soundness isNonTrivial ~relSymInterpretation
+  open ~Soundness
+
+  -- Partial equivalence relations
+
+  private
+    symContext : Context
+    symContext = ([]  `X)  `X
+
+    x∈symContext : `X  symContext
+    x∈symContext = there here
+
+    y∈symContext : `X  symContext
+    y∈symContext = here
+
+     : Term symContext `X
+     = var x∈symContext
+
+     : Term symContext `X
+     = var y∈symContext
+
+    symPrequantFormula : Formula symContext
+    symPrequantFormula = rel fzero ( `, ) `→ rel fzero ( `, )
+
+  -- ~ ⊢ ∀ x. ∀ y. x ~ y → y ~ x
+  symFormula : Formula []
+  symFormula = `∀ (`∀ symPrequantFormula)
+
+  field
+    symWitness : A
+    symIsWitnessed : symWitness   symFormula ⟧ᶠ
+
+  private
+    transContext : Context
+    transContext = (([]  `X)  `X)  `X
+
+    z∈transContext : `X  transContext
+    z∈transContext = here
+
+    y∈transContext : `X  transContext
+    y∈transContext = there here
+
+    x∈transContext : `X  transContext
+    x∈transContext = there (there here)
+
+    zᵗ : Term transContext `X
+    zᵗ = var z∈transContext
+
+    yᵗ : Term transContext `X
+    yᵗ = var y∈transContext
+
+    xᵗ : Term transContext `X
+    xᵗ = var x∈transContext
+
+    -- ~ : Rel X X, x : X, y : X, z : X ⊢ x ~ y ∧ y ~ z → x ~ z
+    transPrequantFormula : Formula transContext
+    transPrequantFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ)
+
+  transFormula : Formula []
+  transFormula = `∀ (`∀ (`∀ transPrequantFormula))
+
+  field
+    transWitness : A
+    transIsWitnessed : transWitness   transFormula ⟧ᶠ
+   
+
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Logic.Semantics.html b/docs/Realizability.Tripos.Logic.Semantics.html index 59facd9..b3dfc9c 100644 --- a/docs/Realizability.Tripos.Logic.Semantics.html +++ b/docs/Realizability.Tripos.Logic.Semantics.html @@ -1,189 +1,601 @@ Realizability.Tripos.Logic.Semantics
{-# OPTIONS --allow-unsolved-metas #-}
 open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Univalence
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Function
-open import Cubical.Foundations.Structure
-open import Cubical.Functions.FunExtEquiv
-open import Cubical.Data.Sigma
-open import Cubical.Data.Empty
-open import Cubical.Data.Unit
-open import Cubical.Data.Sum
-open import Cubical.Data.Vec
-open import Cubical.Data.Nat
-open import Cubical.Data.Fin
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-
-module
-  Realizability.Tripos.Logic.Semantics
-  { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A)  where
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate')
-open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
-open import Realizability.Tripos.Prealgebra.Meets.Identity ca
-open import Realizability.Tripos.Prealgebra.Joins.Identity ca
-open import Realizability.Tripos.Logic.Syntax { = ℓ'}
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Predicate'
-open PredicateProperties hiding (_≤_ ; isTrans≤)
-open Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''}
-Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''}
-RelationInterpretation :  {n : }  (Vec Sort n)  Type _
-RelationInterpretation {n} relSym = (∀ i   Predicate   lookup i relSym ⟧ˢ )
-module Interpretation
-  {n : }
-  (relSym : Vec Sort n)
-  (⟦_⟧ʳ : RelationInterpretation relSym) (isNonTrivial : s  k  ) where
-  open Relational relSym
-
-  ⟦_⟧ᶜ : Context  hSet ℓ'
-   [] ⟧ᶜ = Unit* , isSetUnit* 
-   c  x ⟧ᶜ = ( c ⟧ᶜ .fst ×  x ⟧ˢ .fst) , isSet× ( c ⟧ᶜ .snd) ( x ⟧ˢ .snd)
-
-  ⟦_⟧ⁿ :  {Γ s}  s  Γ    Γ ⟧ᶜ     s ⟧ˢ 
-  ⟦_⟧ⁿ {.(_  s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧
-  ⟦_⟧ⁿ {.(_  _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) =  s∈Γ ⟧ⁿ ⟦Γ⟧
-
-  ⟦_⟧ᵗ :  {Γ s}  Term Γ s    Γ ⟧ᶜ     s ⟧ˢ 
-  ⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ =  x ⟧ⁿ ⟦Γ⟧
-  ⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = ( t ⟧ᵗ ⟦Γ⟧) , ( t₁ ⟧ᵗ ⟦Γ⟧)
-  ⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst ( t ⟧ᵗ ⟦Γ⟧)
-  ⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd ( t ⟧ᵗ ⟦Γ⟧)
-  ⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x ( t ⟧ᵗ ⟦Γ⟧)
-
-  ⟦_⟧ᶠ :  {Γ}  Formula Γ  Predicate   Γ ⟧ᶜ 
-  ⟦_⟧ᶠ {[]} ⊤ᵗ = pre1 Unit* isSetUnit* isNonTrivial
-  ⟦_⟧ᶠ {[]} ⊥ᵗ = pre0 Unit* isSetUnit* isNonTrivial
-  ⟦_⟧ᶠ {[]} (f `∨ f₁) = _⊔_ Unit*  f ⟧ᶠ  f₁ ⟧ᶠ
-  ⟦_⟧ᶠ {[]} (f `∧ f₁) = _⊓_ Unit*  f ⟧ᶠ  f₁ ⟧ᶠ
-  ⟦_⟧ᶠ {[]} (f `→ f₁) = _⇒_ Unit*  f ⟧ᶠ  f₁ ⟧ᶠ
-  ⟦_⟧ᶠ {[]} ( f) = _⇒_ Unit*  f ⟧ᶠ  ⊥ᵗ {Γ = []} ⟧ᶠ
-  ⟦_⟧ᶠ {[]} (`∃ {B = B} f) =
-    `∃[_]
-    {X = Unit* ×  B ⟧ˢ .fst}
-    {Y = Unit*}
-    (isSet× isSetUnit* ( B ⟧ˢ .snd))
-    isSetUnit*
-    fst
-     f ⟧ᶠ
-  ⟦_⟧ᶠ {[]} (`∀ {B = B} f) =
-    `∀[_]
-    {X = Unit* ×  B ⟧ˢ .fst}
-    {Y = Unit*}
-    (isSet× isSetUnit* ( B ⟧ˢ .snd))
-    isSetUnit*
-    fst
-     f ⟧ᶠ
-  ⟦_⟧ᶠ {[]} (rel R t) = ⋆_ isSetUnit* (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ
-  ⟦_⟧ᶠ {Γ  x} ⊤ᵗ = pre1 ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst) (isSet× ( Γ ⟧ᶜ .snd) ( x ⟧ˢ .snd)) isNonTrivial
-  ⟦_⟧ᶠ {Γ  x} ⊥ᵗ = pre0 ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst) (isSet× ( Γ ⟧ᶜ .snd) ( x ⟧ˢ .snd)) isNonTrivial
-  ⟦_⟧ᶠ {Γ  x} (f `∨ f₁) = _⊔_ ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst)  f ⟧ᶠ  f₁ ⟧ᶠ
-  ⟦_⟧ᶠ {Γ  x} (f `∧ f₁) = _⊓_ ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst)  f ⟧ᶠ  f₁ ⟧ᶠ
-  ⟦_⟧ᶠ {Γ  x} (f `→ f₁) = _⇒_ ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst)  f ⟧ᶠ  f₁ ⟧ᶠ
-  ⟦_⟧ᶠ {Γ  x} ( f) = _⇒_ ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst)  f ⟧ᶠ  ⊥ᵗ {Γ = Γ  x} ⟧ᶠ
-  ⟦_⟧ᶠ {Γ  x} (`∃ {B = B} f) =
-    `∃[_]
-    {X = ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst) ×  B ⟧ˢ .fst}
-    {Y =  Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst}
-    (isSet× (isSet× ( Γ ⟧ᶜ .snd) ( x ⟧ˢ .snd)) ( B ⟧ˢ .snd))
-    (isSet× ( Γ ⟧ᶜ .snd) ( x ⟧ˢ .snd))
-    fst
-    ( f ⟧ᶠ)
-  ⟦_⟧ᶠ {Γ  x} (`∀ {B = B} f) =
-    `∀[_]
-    {X = ( Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst) ×  B ⟧ˢ .fst}
-    {Y =  Γ ⟧ᶜ .fst ×  x ⟧ˢ .fst}
-    (isSet× (isSet× ( Γ ⟧ᶜ .snd) ( x ⟧ˢ .snd)) ( B ⟧ˢ .snd))
-    (isSet× ( Γ ⟧ᶜ .snd) ( x ⟧ˢ .snd))
-    fst
-    ( f ⟧ᶠ)
-  ⟦_⟧ᶠ {Γ  x} (rel R t) = ⋆_ (str  Γ  x ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ
-
-  -- R for renamings and r for relations
-  ⟦_⟧ᴿ :  {Γ Δ}  Renaming Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
-   id ⟧ᴿ ctx = ctx
-   drop ren ⟧ᴿ (ctx , _) =  ren ⟧ᴿ ctx
-   keep ren ⟧ᴿ (ctx , s) = ( ren ⟧ᴿ ctx) , s
-
-  -- B for suBstitution and s for sorts
-  ⟦_⟧ᴮ :  {Γ Δ}  Substitution Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
-   id ⟧ᴮ ctx = ctx
-   t , sub ⟧ᴮ ctx = ( sub ⟧ᴮ ctx) , ( t ⟧ᵗ ctx)
-   drop sub ⟧ᴮ (ctx , s) =  sub ⟧ᴮ ctx
-
-  renamingVarSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (v : s  Δ)   renamingVar ren v ⟧ⁿ   v ⟧ⁿ   ren ⟧ᴿ
-  renamingVarSound {Γ} {.Γ} {s} id v = refl
-  renamingVarSound {.(_  _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
-  renamingVarSound {.(_  s)} {.(_  s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  ⟦s⟧ }
-  renamingVarSound {.(_  _)} {.(_  _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
-
-  renamingTermSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (t : Term Δ s)   renamingTerm ren t ⟧ᵗ   t ⟧ᵗ   ren ⟧ᴿ
-  renamingTermSound {Γ} {.Γ} {s} id t = refl
-  renamingTermSound {.(_  _)} {Δ} {s} (drop ren) (var x) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren x i ⟦Γ⟧ }
-  renamingTermSound {.(_  _)} {Δ} {.(_  _)} r@(drop ren) (t `, t₁) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) }
-  renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₁ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
-  renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₂ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
-  renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (fun f t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
-  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (var v) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingVarSound r v i x }
-  renamingTermSound {.(_  _)} {.(_  _)} {.(_  _)} r@(keep ren) (t `, t₁) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) }
-  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₁ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
-  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₂ t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
-  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (fun f t) = funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
-
-  substitutionVarSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (v : s  Δ)   substitutionVar subs v ⟧ᵗ   v ⟧ⁿ   subs ⟧ᴮ
-  substitutionVarSound {Γ} {.Γ} {s} id t = refl
-  substitutionVarSound {Γ} {.(_  s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧  refl
-  substitutionVarSound {Γ} {.(_  _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i  substitutionVarSound subs t i ⟦Γ⟧
-  substitutionVarSound {.(_  _)} {Δ} {s} r@(drop subs) t =
-    -- TODO : Fix unsolved constraints
-    funExt
-      λ { x@(⟦Γ⟧ , ⟦s⟧) 
-         substitutionVar (drop subs) t ⟧ᵗ (⟦Γ⟧ , ⟦s⟧)
-          ≡[ i ]⟨  renamingTermSound (drop id) (substitutionVar subs t) i (⟦Γ⟧ , ⟦s⟧)  
-         substitutionVar subs t ⟧ᵗ ( drop id ⟧ᴿ x)
-          ≡⟨ refl 
-         substitutionVar subs t ⟧ᵗ ⟦Γ⟧
-          ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ 
-         t ⟧ⁿ ( subs ⟧ᴮ ⟦Γ⟧)
-          }
-
-  substitutionTermSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (t : Term Δ s)   substitutionTerm subs t ⟧ᵗ   t ⟧ᵗ   subs ⟧ᴮ
-  substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x
-  substitutionTermSound {Γ} {Δ} {.(_  _)} subs (t `, t₁) = funExt λ x i  (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x)
-  substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i  substitutionTermSound subs t i x .fst
-  substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i  substitutionTermSound subs t i x .snd
-  substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i  f (substitutionTermSound subs t i x)
-
-module Soundness
-  {n}
-  {relSym : Vec Sort n}
-  (isNonTrivial : s  k  )
-  (⟦_⟧ʳ : RelationInterpretation relSym) where
-  open Relational relSym
-  open Interpretation relSym ⟦_⟧ʳ isNonTrivial
-
-  infix 24 _⊨_
-
-  module _ {Γ : Context} where
-
-    open PredicateProperties {ℓ'' = ℓ''}   Γ ⟧ᶜ 
-
-    _⊨_ : Formula Γ  Formula Γ  Type _
-    ϕ  ψ =  ϕ ⟧ᶠ   ψ ⟧ᶠ
-
-    private
-      variable
-        ϕ ψ θ χ : Formula Γ
-
-    cut :  {ϕ ψ θ}  ϕ  ψ  ψ  θ  ϕ  θ
-    cut {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤  ϕ ⟧ᶠ  ψ ⟧ᶠ  θ ⟧ᶠ ϕ⊨ψ ψ⊨θ
-
-    
+open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Structure
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty renaming (elim to ⊥elim ; rec* to ⊥rec*)
+open import Cubical.Data.Unit
+open import Cubical.Data.Sum
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.Fin
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+
+module
+  Realizability.Tripos.Logic.Semantics
+  { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A)  where
+open CombinatoryAlgebra ca
+private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
+private λ* = `λ* as fefermanStructure
+
+open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate')
+open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity ca
+open import Realizability.Tripos.Prealgebra.Joins.Identity ca
+open import Realizability.Tripos.Prealgebra.Implication ca
+open import Realizability.Tripos.Logic.Syntax { = ℓ'}
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate'
+open PredicateProperties hiding (_≤_ ; isTrans≤)
+open Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''}
+Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''}
+RelationInterpretation :  {n : }  (Vec Sort n)  Type _
+RelationInterpretation {n} relSym = (∀ i   Predicate   lookup i relSym ⟧ˢ )
+module Interpretation
+  {n : }
+  (relSym : Vec Sort n)
+  (⟦_⟧ʳ : RelationInterpretation relSym) (isNonTrivial : s  k  ) where
+  open Relational relSym
+
+  ⟦_⟧ᶜ : Context  hSet ℓ'
+   [] ⟧ᶜ = Unit* , isSetUnit* 
+   c  x ⟧ᶜ = ( c ⟧ᶜ .fst ×  x ⟧ˢ .fst) , isSet× ( c ⟧ᶜ .snd) ( x ⟧ˢ .snd)
+
+  ⟦_⟧ⁿ :  {Γ s}  s  Γ    Γ ⟧ᶜ     s ⟧ˢ 
+  ⟦_⟧ⁿ {.(_  s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧
+  ⟦_⟧ⁿ {.(_  _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) =  s∈Γ ⟧ⁿ ⟦Γ⟧
+
+  ⟦_⟧ᵗ :  {Γ s}  Term Γ s    Γ ⟧ᶜ     s ⟧ˢ 
+  ⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ =  x ⟧ⁿ ⟦Γ⟧
+  ⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = ( t ⟧ᵗ ⟦Γ⟧) , ( t₁ ⟧ᵗ ⟦Γ⟧)
+  ⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst ( t ⟧ᵗ ⟦Γ⟧)
+  ⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd ( t ⟧ᵗ ⟦Γ⟧)
+  ⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x ( t ⟧ᵗ ⟦Γ⟧)
+
+  ⟦_⟧ᶠ :  {Γ}  Formula Γ  Predicate   Γ ⟧ᶜ 
+  ⟦_⟧ᶠ {Γ} ⊤ᵗ = pre1   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial
+  ⟦_⟧ᶠ {Γ} ⊥ᵗ = pre0   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial
+  ⟦_⟧ᶠ {Γ} (form `∨ form₁) = _⊔_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (form `∧ form₁) = _⊓_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (form `→ form₁) = _⇒_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} ( form) = _⇒_   Γ ⟧ᶜ   form ⟧ᶠ  ⊥ᵗ {Γ = Γ} ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (`∃ {B = B} form) = `∃[_] (isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ)) (str  Γ ⟧ᶜ)  { (⟦Γ⟧ , ⟦B⟧)  ⟦Γ⟧ })  form ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (`∀ {B = B} form) = `∀[_] (isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ)) (str  Γ ⟧ᶜ)  { (⟦Γ⟧ , ⟦B⟧)  ⟦Γ⟧ })  form ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (rel R t) = ⋆_ (str  Γ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ
+
+  -- R for renamings and r for relations
+  ⟦_⟧ᴿ :  {Γ Δ}  Renaming Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
+   id ⟧ᴿ ctx = ctx
+   drop ren ⟧ᴿ (ctx , _) =  ren ⟧ᴿ ctx
+   keep ren ⟧ᴿ (ctx , s) = ( ren ⟧ᴿ ctx) , s
+
+  -- B for suBstitution and s for sorts
+  ⟦_⟧ᴮ :  {Γ Δ}  Substitution Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
+   id ⟧ᴮ ctx = ctx
+   t , sub ⟧ᴮ ctx = ( sub ⟧ᴮ ctx) , ( t ⟧ᵗ ctx)
+   drop sub ⟧ᴮ (ctx , s) =  sub ⟧ᴮ ctx
+
+  renamingVarSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (v : s  Δ)   renamingVar ren v ⟧ⁿ   v ⟧ⁿ   ren ⟧ᴿ
+  renamingVarSound {Γ} {.Γ} {s} id v = refl
+  renamingVarSound {.(_  _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
+  renamingVarSound {.(_  s)} {.(_  s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  ⟦s⟧ }
+  renamingVarSound {.(_  _)} {.(_  _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
+
+  renamingTermSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (t : Term Δ s)   renamingTerm ren t ⟧ᵗ   t ⟧ᵗ   ren ⟧ᴿ
+  renamingTermSound {Γ} {.Γ} {s} id t = refl
+  renamingTermSound {.(_  _)} {Δ} {s} (drop ren) (var x) =
+    funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren x i ⟦Γ⟧ }
+  renamingTermSound {.(_  _)} {Δ} {.(_  _)} r@(drop ren) (t `, t₁) =
+    funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) }
+  renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₁ t) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
+  renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₂ t) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
+  renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (fun f t) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
+  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (var v) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingVarSound r v i x }
+  renamingTermSound {.(_  _)} {.(_  _)} {.(_  _)} r@(keep ren) (t `, t₁) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) }
+  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₁ t) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
+  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₂ t) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
+  renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (fun f t) =
+    funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
+
+  substitutionVarSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (v : s  Δ)   substitutionVar subs v ⟧ᵗ   v ⟧ⁿ   subs ⟧ᴮ
+  substitutionVarSound {Γ} {.Γ} {s} id t = refl
+  substitutionVarSound {Γ} {.(_  s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧  refl
+  substitutionVarSound {Γ} {.(_  _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i  substitutionVarSound subs t i ⟦Γ⟧
+  substitutionVarSound {.(_  _)} {Δ} {s} r@(drop subs) t =
+    -- TODO : Fix unsolved constraints
+    funExt
+      λ { x@(⟦Γ⟧ , ⟦s⟧) 
+         substitutionVar (drop subs) t ⟧ᵗ x
+          ≡[ i ]⟨  renamingTermSound (drop id) (substitutionVar subs t) i x  
+         substitutionVar subs t ⟧ᵗ ( drop id ⟧ᴿ x)
+          ≡⟨ refl 
+         substitutionVar subs t ⟧ᵗ ⟦Γ⟧
+          ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ 
+         t ⟧ⁿ ( subs ⟧ᴮ ⟦Γ⟧)
+          }
+
+  substitutionTermSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (t : Term Δ s)   substitutionTerm subs t ⟧ᵗ   t ⟧ᵗ   subs ⟧ᴮ
+  substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x
+  substitutionTermSound {Γ} {Δ} {.(_  _)} subs (t `, t₁) = funExt λ x i  (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x)
+  substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i  substitutionTermSound subs t i x .fst
+  substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i  substitutionTermSound subs t i x .snd
+  substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i  f (substitutionTermSound subs t i x)
+
+  semanticSubstitution :  {Γ Δ}  (subs : Substitution Γ Δ)  Predicate   Δ ⟧ᶜ   Predicate   Γ ⟧ᶜ 
+  semanticSubstitution {Γ} {Δ} subs = ⋆_ (str  Γ ⟧ᶜ) (str  Δ ⟧ᶜ)  subs ⟧ᴮ
+
+  -- Due to a shortcut in the soundness of negation termination checking fails
+  -- TODO : Fix
+  {-# TERMINATING #-}
+  substitutionFormulaSound :  {Γ Δ}  (subs : Substitution Γ Δ)  (f : Formula Δ)   substitutionFormula subs f ⟧ᶠ  semanticSubstitution subs  f ⟧ᶠ
+  substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (pre1   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial)
+      (semanticSubstitution subs (pre1   Δ ⟧ᶜ  (str  Δ ⟧ᶜ) isNonTrivial))
+       γ a a⊩1γ  tt*)
+      λ γ a a⊩1subsγ  tt*
+  substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (pre0   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial)
+      (semanticSubstitution subs (pre0   Δ ⟧ᶜ  (str  Δ ⟧ᶜ) isNonTrivial))
+       _ _ bot  ⊥rec* bot)
+      λ _ _ bot  bot
+  substitutionFormulaSound {Γ} {Δ} subs (f `∨ f₁) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (_⊔_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
+      (semanticSubstitution subs (_⊔_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
+       γ a a⊩substFormFs 
+        a⊩substFormFs >>=
+          λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) 
+                    inl (pr₁a≡k , subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁
+            ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) 
+                    inr (pr₁a≡k' , subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ })
+      λ γ a a⊩semanticSubsFs 
+        a⊩semanticSubsFs >>=
+          λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) 
+                    inl (pr₁a≡k , (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁
+            ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) 
+                    inr (pr₁a≡k' , (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ }
+  substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (_⊓_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
+      (semanticSubstitution subs (_⊓_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
+       γ a a⊩substFormulaFs 
+        (subst  form  (pr₁  a)   form  γ) (substitutionFormulaSound subs f) (a⊩substFormulaFs .fst)) ,
+        (subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f₁) (a⊩substFormulaFs .snd)))
+      λ γ a a⊩semanticSubstFs 
+        (subst  form  (pr₁  a)   form  γ) (sym (substitutionFormulaSound subs f)) (a⊩semanticSubstFs .fst)) ,
+        (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f₁)) (a⊩semanticSubstFs .snd))
+  substitutionFormulaSound {Γ} {Δ} subs (f `→ f₁) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (_⇒_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
+      (semanticSubstitution subs (_⇒_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
+       γ a a⊩substFormulaFs 
+        λ b b⊩semanticSubstFs 
+          subst
+             form  (a  b)   form  γ)
+            (substitutionFormulaSound subs f₁)
+            (a⊩substFormulaFs
+              b
+              (subst
+                 form  b   form  γ)
+                (sym (substitutionFormulaSound subs f))
+                b⊩semanticSubstFs)))
+      λ γ a a⊩semanticSubstFs 
+        λ b b⊩substFormulaFs 
+          subst
+             form  (a  b)   form  γ)
+            (sym (substitutionFormulaSound subs f₁))
+            (a⊩semanticSubstFs
+              b
+              (subst
+                 form  b   form  γ)
+                (substitutionFormulaSound subs f)
+                b⊩substFormulaFs))
+  substitutionFormulaSound {Γ} {Δ} subs ( f) =
+    substitutionFormulaSound subs (f `→ ⊥ᵗ)
+  substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (`∃[ isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Γ ⟧ᶜ)  { (f , s)  f })  substitutionFormula (var here , drop subs) f ⟧ᶠ)
+      (semanticSubstitution subs (`∃[ isSet× (str  Δ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Δ ⟧ᶜ)  { (γ , b)  γ })  f ⟧ᶠ))
+       γ a a⊩πSubstFormulaF 
+        a⊩πSubstFormulaF >>=
+          λ { ((γ' , b) , γ'≡γ , a⊩substFormFγ') 
+             (( subs ⟧ᴮ γ') , b) ,
+              ((cong  subs ⟧ᴮ γ'≡γ) ,
+                (subst
+                   form  a   form  (γ' , b))
+                  (substitutionFormulaSound (var here , drop subs) f)
+                  a⊩substFormFγ' )) ∣₁ })
+      λ γ a a⊩semanticSubstF 
+        a⊩semanticSubstF >>=
+          λ (x@(δ , b) , δ≡subsγ , a⊩fx) 
+             (γ , b) ,
+              (refl ,
+                (subst
+                   form  a   form  (γ , b))
+                  (sym (substitutionFormulaSound (var here , drop subs) f))
+                  (subst  x  a    f ⟧ᶠ  (x , b)) δ≡subsγ a⊩fx))) ∣₁
+  substitutionFormulaSound {Γ} {Δ} subs (`∀ {B = B} f) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (`∀[ isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Γ ⟧ᶜ)  { (f , s)  f })  substitutionFormula (var here , drop subs) f ⟧ᶠ)
+      (semanticSubstitution subs (`∀[ isSet× (str  Δ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Δ ⟧ᶜ)  { (f , s)  f })  f ⟧ᶠ))
+       γ a a⊩substFormF 
+        λ { r x@(δ , b) δ≡subsγ 
+          subst
+             g  (a  r)    f ⟧ᶠ  (g , b))
+            (sym δ≡subsγ)
+            (subst
+               form  (a  r)   form  (γ , b))
+              (substitutionFormulaSound (var here , drop subs) f)
+              (a⊩substFormF r (γ , b) refl)) })
+      λ γ a a⊩semanticSubsF 
+        λ { r x@(γ' , b) γ'≡γ 
+          subst
+             form  (a  r)   form  (γ' , b))
+            (sym (substitutionFormulaSound (var here , drop subs) f))
+            (subst
+               g  (a  r)    f ⟧ᶠ  (g , b))
+              (cong  subs ⟧ᴮ (sym γ'≡γ))
+              (a⊩semanticSubsF r ( subs ⟧ᴮ γ , b) refl)) }
+  substitutionFormulaSound {Γ} {Δ} subs (rel R t) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (⋆_ (str  Γ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  substitutionTerm subs t ⟧ᵗ  R ⟧ʳ)
+      (semanticSubstitution subs (⋆_ (str  Δ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ))
+       γ a a⊩substTR 
+        subst  transform  a    R ⟧ʳ  (transform γ)) (substitutionTermSound subs t) a⊩substTR)
+      λ γ a a⊩semSubst 
+        subst  transform  a    R ⟧ʳ  (transform γ)) (sym (substitutionTermSound subs t)) a⊩semSubst
+
+  weakenFormulaMonotonic :  {Γ B}  (γ :   Γ ⟧ᶜ )  (ϕ : Formula Γ)  (a : A)  (b :   B ⟧ˢ )  a    ϕ ⟧ᶠ  γ  a    weakenFormula {S = B} ϕ ⟧ᶠ  (γ , b)
+  weakenFormulaMonotonic {Γ} {B} γ ϕ a b =
+    hPropExt
+      ( ϕ ⟧ᶠ .isPropValued γ a)
+      ( weakenFormula ϕ ⟧ᶠ .isPropValued (γ , b) a)
+       a⊩ϕγ  subst  form  a   form  (γ , b)) (sym (substitutionFormulaSound (drop id) ϕ)) a⊩ϕγ)
+      λ a⊩weakenϕγb  subst  form  a   form  (γ , b)) (substitutionFormulaSound (drop id) ϕ) a⊩weakenϕγb
+module Soundness
+  {n}
+  {relSym : Vec Sort n}
+  (isNonTrivial : s  k  )
+  (⟦_⟧ʳ : RelationInterpretation relSym) where
+  open Relational relSym
+  open Interpretation relSym ⟦_⟧ʳ isNonTrivial
+  -- Acknowledgements : 1lab's "the internal logic of a regular hyperdoctrine"
+  infix 35 _⊨_
+  
+  module PredProps = PredicateProperties
+  
+  _⊨_ :  {Γ}  Formula Γ  Formula Γ  Type (ℓ-max (ℓ-max  ℓ'') ℓ')
+  _⊨_ {Γ} ϕ ψ =  ϕ ⟧ᶠ   ψ ⟧ᶠ where open PredProps   Γ ⟧ᶜ 
+
+  entails = _⊨_
+
+  private
+    variable
+      Γ Δ : Context
+      ϕ ψ θ : Formula Γ
+      χ μ ν : Formula Δ
+
+  cut :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ψ  θ  ϕ  θ
+  cut {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤  ϕ ⟧ᶠ  ψ ⟧ᶠ  θ ⟧ᶠ ϕ⊨ψ ψ⊨θ where open PredProps   Γ ⟧ᶜ 
+
+  substitutionEntailment :  {Γ Δ} (subs : Substitution Γ Δ)  {ϕ ψ : Formula Δ}  ϕ  ψ  substitutionFormula subs ϕ  substitutionFormula subs ψ
+  substitutionEntailment {Γ} {Δ} subs {ϕ} {ψ} ϕ⊨ψ =
+    subst2
+       ϕ' ψ'  ϕ' ≤Γ ψ')
+      (sym (substitutionFormulaSound subs ϕ))
+      (sym (substitutionFormulaSound subs ψ))
+      (ϕ⊨ψ >>=
+        λ { (a , a⊩ϕ≤ψ) 
+           a ,  γ b b⊩ϕsubsγ  a⊩ϕ≤ψ ( subs ⟧ᴮ γ) b b⊩ϕsubsγ) ∣₁ }) where
+      open PredProps {ℓ'' = ℓ''}   Γ ⟧ᶜ  renaming (_≤_ to _≤Γ_)
+      open PredProps {ℓ'' = ℓ''}   Δ ⟧ᶜ  renaming (_≤_ to _≤Δ_)
+
+  `∧intro :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  entails ϕ θ  entails ϕ (ψ `∧ θ)
+  `∧intro {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ϕ⊨θ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      (b , b⊩ϕ⊨θ)  ϕ⊨θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pair ̇ (` a ̇ # fzero) ̇ (` b ̇ # fzero)
+      return
+        (λ* prover ,
+          λ γ r r⊩ϕγ 
+            let
+              proofEq : λ* prover  r  pair  (a  r)  (b  r)
+              proofEq = λ*ComputationRule prover (r  [])
+
+              pr₁proofEq : pr₁  (λ* prover  r)  a  r
+              pr₁proofEq =
+                pr₁  (λ* prover  r)
+                  ≡⟨ cong  x  pr₁  x) proofEq 
+                pr₁  (pair  (a  r)  (b  r))
+                  ≡⟨ pr₁pxy≡x _ _ 
+                a  r
+                  
+
+              pr₂proofEq : pr₂  (λ* prover  r)  b  r
+              pr₂proofEq =
+                pr₂  (λ* prover  r)
+                  ≡⟨ cong  x  pr₂  x) proofEq 
+                pr₂  (pair  (a  r)  (b  r))
+                  ≡⟨ pr₂pxy≡y _ _ 
+                b  r
+                  
+            in
+            subst  r  r    ψ ⟧ᶠ  γ) (sym pr₁proofEq) (a⊩ϕ⊨ψ γ r r⊩ϕγ) ,
+            subst  r  r    θ ⟧ᶠ  γ) (sym pr₂proofEq) (b⊩ϕ⊨θ γ r r⊩ϕγ))
+
+  `∧elim1 :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `∧ θ)  ϕ  ψ
+  `∧elim1 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ =
+    do
+      (a , a⊩ϕ⊨ψ∧θ)  ϕ⊨ψ∧θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pr₁ ̇ (` a ̇ # fzero)
+      return
+        (λ* prover ,
+          λ γ b b⊩ϕγ  subst  r  r    ψ ⟧ᶠ  γ) (sym (λ*ComputationRule prover (b  []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst))
+          
+  `∧elim2 :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `∧ θ)  ϕ  θ
+  `∧elim2 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ =
+    do
+      (a , a⊩ϕ⊨ψ∧θ)  ϕ⊨ψ∧θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pr₂ ̇ (` a ̇ # fzero)
+      return
+        (λ* prover ,
+          λ γ b b⊩ϕγ  subst  r  r    θ ⟧ᶠ  γ) (sym (λ*ComputationRule prover (b  []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd))
+
+  `∃intro :  {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ  B)} {t : Term Γ B}  ϕ  substitutionFormula (t , id) ψ  ϕ  `∃ ψ
+  `∃intro {Γ} {ϕ} {B} {ψ} {t} ϕ⊨ψ[t/x] =
+    do
+      (a , a⊩ϕ⊨ψ[t/x])  ϕ⊨ψ[t/x]
+      return
+        (a ,  γ b b⊩ϕγ   (γ , ( t ⟧ᵗ γ)) ,
+        (refl , (subst  form  (a  b)   form  γ) (substitutionFormulaSound (t , id) ψ) (a⊩ϕ⊨ψ[t/x] γ b b⊩ϕγ))) ∣₁))
+
+  `∃elim :  {Γ} {ϕ θ : Formula Γ} {B} {ψ : Formula (Γ  B)}  ϕ  `∃ ψ  (weakenFormula ϕ `∧ ψ)  weakenFormula θ  ϕ  θ
+  `∃elim {Γ} {ϕ} {θ} {B} {ψ} ϕ⊨∃ψ ϕ∧ψ⊨θ =
+    do
+      (a , a⊩ϕ⊨∃ψ)  ϕ⊨∃ψ
+      (b , b⊩ϕ∧ψ⊨θ)  ϕ∧ψ⊨θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` b ̇ (` pair ̇ # fzero ̇ (` a ̇ # fzero))
+      return
+        (λ* prover ,
+         γ c c⊩ϕγ 
+          subst
+             r  r    θ ⟧ᶠ  γ)
+            (sym (λ*ComputationRule prover (c  [])))
+            (transport
+              (propTruncIdempotent ( θ ⟧ᶠ .isPropValued γ (b  (pair  c  (a  c)))))
+              (a⊩ϕ⊨∃ψ γ c c⊩ϕγ >>=
+                λ { (x@(γ' , b') , (γ'≡γ , a⨾c⊩ψx)) 
+                   transport
+                    (sym
+                      (weakenFormulaMonotonic γ θ (b  (pair  c  (a  c))) b'))
+                    (b⊩ϕ∧ψ⊨θ
+                      (γ , b')
+                      (pair  c  (a  c))
+                      (subst
+                         r  r    weakenFormula ϕ ⟧ᶠ  (γ , b'))
+                        (sym (pr₁pxy≡x _ _))
+                        (transport
+                          (weakenFormulaMonotonic γ ϕ c b') c⊩ϕγ) ,
+                      subst  r  r    ψ ⟧ᶠ  (γ , b')) (sym (pr₂pxy≡y _ _)) (subst  g  (a  c)    ψ ⟧ᶠ  (g , b')) γ'≡γ a⨾c⊩ψx)) ) ∣₁ }))))
+
+  `∀intro :  {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ  B)}  weakenFormula ϕ  ψ  ϕ  `∀ ψ
+  `∀intro {Γ} {ϕ} {B} {ψ} ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 2
+        prover = ` a ̇ # fzero
+      return
+        (λ* prover ,
+         γ b b⊩ϕ  λ { c x@(γ' , b') γ'≡γ 
+          subst
+             r  r    ψ ⟧ᶠ  (γ' , b'))
+            (sym (λ*ComputationRule prover (b  c  [])))
+            (a⊩ϕ⊨ψ
+              (γ' , b')
+              b
+              (transport (weakenFormulaMonotonic γ' ϕ b b') (subst  g  b    ϕ ⟧ᶠ  g) (sym γ'≡γ) b⊩ϕ))) }))
+
+  `∀elim :  {Γ} {B} {ϕ : Formula Γ} {ψ : Formula (Γ  B)}  ϕ  `∀ ψ  (t : Term Γ B)  ϕ  substitutionFormula (t , id) ψ
+  `∀elim {Γ} {B} {ϕ} {ψ} ϕ⊨∀ψ t =
+    do
+      (a , a⊩ϕ⊨∀ψ)  ϕ⊨∀ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` a ̇ # fzero ̇ ` k
+      return
+        (λ* prover ,
+         γ b b⊩ϕγ 
+          subst
+           form  (λ* prover  b)   form  γ)
+          (sym (substitutionFormulaSound (t , id) ψ))
+          (subst
+             r  r    ψ ⟧ᶠ  (γ ,  t ⟧ᵗ γ))
+            (sym (λ*ComputationRule prover (b  [])))
+            (a⊩ϕ⊨∀ψ γ b b⊩ϕγ k (γ ,  t ⟧ᵗ γ) refl))))
+
+  `→intro :  {Γ} {ϕ ψ θ : Formula Γ}  (ϕ `∧ ψ)  θ  ϕ  (ψ `→ θ)
+  `→intro {Γ} {ϕ} {ψ} {θ} ϕ∧ψ⊨θ = a⊓b≤c→a≤b⇒c   Γ ⟧ᶜ  (str  Γ ⟧ᶜ)  ϕ ⟧ᶠ  ψ ⟧ᶠ  θ ⟧ᶠ ϕ∧ψ⊨θ
+
+  `→elim :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `→ θ)  ϕ  ψ  ϕ  θ
+  `→elim {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ→θ ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ→θ)  ϕ⊨ψ→θ
+      (b , b⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` a ̇ (# fzero) ̇ (` b ̇ # fzero)
+      return
+        (λ* prover ,
+         γ c c⊩ϕγ 
+          subst
+             r  r    θ ⟧ᶠ  γ)
+            (sym (λ*ComputationRule prover (c  [])))
+            (a⊩ϕ⊨ψ→θ γ c c⊩ϕγ (b  c) (b⊩ϕ⊨ψ γ c c⊩ϕγ))))
+
+  `∨introR :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ϕ  (ψ `∨ θ)
+  `∨introR {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pair ̇ ` true ̇ (` a ̇ # fzero)
+      return
+        ((λ* prover) ,
+         γ b b⊩ϕγ 
+          let
+            pr₁proofEq : pr₁  (λ* prover  b)  true
+            pr₁proofEq =
+              pr₁  (λ* prover  b)
+                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (b  [])) 
+              pr₁  (pair  true  (a  b))
+                ≡⟨ pr₁pxy≡x _ _ 
+              true
+                
+
+            pr₂proofEq : pr₂  (λ* prover  b)  a  b
+            pr₂proofEq =
+              pr₂  (λ* prover  b)
+                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (b  [])) 
+              pr₂  (pair  true  (a  b))
+                ≡⟨ pr₂pxy≡y _ _ 
+              a  b
+                
+          in  inl (pr₁proofEq , subst  r  r    ψ ⟧ᶠ  γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁))
+
+  `∨introL :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ϕ  (θ `∨ ψ)
+  `∨introL {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pair ̇ ` false ̇ (` a ̇ # fzero)
+      return
+        ((λ* prover) ,
+         γ b b⊩ϕγ 
+          let
+            pr₁proofEq : pr₁  (λ* prover  b)  false
+            pr₁proofEq =
+              pr₁  (λ* prover  b)
+                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (b  [])) 
+              pr₁  (pair  false  (a  b))
+                ≡⟨ pr₁pxy≡x _ _ 
+              false
+                
+
+            pr₂proofEq : pr₂  (λ* prover  b)  a  b
+            pr₂proofEq =
+              pr₂  (λ* prover  b)
+                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (b  [])) 
+              pr₂  (pair  false  (a  b))
+                ≡⟨ pr₂pxy≡y _ _ 
+              a  b
+                
+          in  inr (pr₁proofEq , subst  r  r    ψ ⟧ᶠ  γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁))
+
+  -- Pretty sure this is code duplication
+  `if_then_else_ :  {as n}  ApplStrTerm as n  ApplStrTerm as n  ApplStrTerm as n  ApplStrTerm as n
+  `if a then b else c = ` Id ̇ a ̇ b ̇ c
+
+  `∨elim :  {Γ} {ϕ ψ θ χ : Formula Γ}  (ϕ `∧ ψ)  χ  (ϕ `∧ θ)  χ  (ϕ `∧ (ψ `∨ θ))  χ
+  `∨elim {Γ} {ϕ} {ψ} {θ} {χ} ϕ∧ψ⊨χ ϕ∧θ⊨χ =
+    do
+      (a , a⊩ϕ∧ψ⊨χ)  ϕ∧ψ⊨χ
+      (b , b⊩ϕ∧θ⊨χ)  ϕ∧θ⊨χ
+      let
+        prover : ApplStrTerm as 1
+        prover =
+          (`if ` pr₁ ̇ (` pr₂ ̇ # fzero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))
+      return
+        (λ* prover ,
+        
+          { γ c (pr₁⨾c⊩ϕγ , pr₂⨾c⊩ψ∨θ) 
+            transport
+            (propTruncIdempotent ( χ ⟧ᶠ .isPropValued γ (λ* prover  c)))
+            (pr₂⨾c⊩ψ∨θ >>=
+              λ { (inl (pr₁⨾pr₂⨾c≡true , pr₂⨾pr₂⨾c⊩ψ)) 
+                  let
+                    proofEq : λ* prover  c  a  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                    proofEq =
+                      λ* prover  c
+                        ≡⟨ λ*ComputationRule prover (c  []) 
+                      (if (pr₁  (pr₂  c)) then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  (if x then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) pr₁⨾pr₂⨾c≡true 
+                      (if true then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  x  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) (ifTrueThen a b) 
+                      a  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        
+                  in
+                   subst
+                     r  r    χ ⟧ᶠ  γ)
+                    (sym proofEq)
+                    (a⊩ϕ∧ψ⊨χ
+                      γ
+                      (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                      ((
+                      subst
+                         r  r    ϕ ⟧ᶠ  γ)
+                        (sym (pr₁pxy≡x _ _))
+                        pr₁⨾c⊩ϕγ) ,
+                      subst
+                         r  r    ψ ⟧ᶠ  γ)
+                        (sym (pr₂pxy≡y _ _))
+                        pr₂⨾pr₂⨾c⊩ψ)) ∣₁
+                ; (inr (pr₁pr₂⨾c≡false , pr₂⨾pr₂⨾c⊩θ)) 
+                  let
+                    proofEq : λ* prover  c  b  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                    proofEq =
+                      λ* prover  c
+                        ≡⟨ λ*ComputationRule prover (c  []) 
+                      (if (pr₁  (pr₂  c)) then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  (if x then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) pr₁pr₂⨾c≡false 
+                      (if false then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  x  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) (ifFalseElse a b) 
+                      b  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        
+                  in
+                   subst
+                     r  r    χ ⟧ᶠ  γ)
+                    (sym proofEq)
+                    (b⊩ϕ∧θ⊨χ
+                      γ
+                      (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                      ((subst  r  r    ϕ ⟧ᶠ  γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) ,
+                       (subst  r  r    θ ⟧ᶠ  γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) }))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Logic.Syntax.html b/docs/Realizability.Tripos.Logic.Syntax.html index 09411b6..cbef8a1 100644 --- a/docs/Realizability.Tripos.Logic.Syntax.html +++ b/docs/Realizability.Tripos.Logic.Syntax.html @@ -98,5 +98,18 @@ `∃ : {Γ B} Formula (Γ B) Formula Γ `∀ : {Γ B} Formula (Γ B) Formula Γ rel : {Γ} (k : Fin n) Term Γ (lookup k relSym) Formula Γ - + + substitutionFormula : {Γ Δ} Substitution Γ Δ Formula Δ Formula Γ + substitutionFormula {Γ} {Δ} subs ⊤ᵗ = ⊤ᵗ + substitutionFormula {Γ} {Δ} subs ⊥ᵗ = ⊥ᵗ + substitutionFormula {Γ} {Δ} subs (form `∨ form₁) = substitutionFormula subs form `∨ substitutionFormula subs form₁ + substitutionFormula {Γ} {Δ} subs (form `∧ form₁) = substitutionFormula subs form `∧ substitutionFormula subs form₁ + substitutionFormula {Γ} {Δ} subs (form `→ form₁) = substitutionFormula subs form `→ substitutionFormula subs form₁ + substitutionFormula {Γ} {Δ} subs ( form) = substitutionFormula subs form + substitutionFormula {Γ} {Δ} subs (`∃ form) = `∃ (substitutionFormula (var here , drop subs) form ) + substitutionFormula {Γ} {Δ} subs (`∀ form) = `∀ (substitutionFormula (var here , drop subs) form) + substitutionFormula {Γ} {Δ} subs (rel k x) = rel k (substitutionTerm subs x) + + weakenFormula : {Γ} {S} Formula Γ Formula (Γ S) + weakenFormula {Γ} {S} form = substitutionFormula (drop id) form \ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Implication.html b/docs/Realizability.Tripos.Prealgebra.Implication.html index 846cb4c..968a8e7 100644 --- a/docs/Realizability.Tripos.Prealgebra.Implication.html +++ b/docs/Realizability.Tripos.Prealgebra.Implication.html @@ -15,76 +15,77 @@ open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -λ*ComputationRule = `λ*ComputationRule as fefermanStructure -λ* = `λ* as fefermanStructure +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where - PredicateX = Predicate {ℓ'' = ℓ''} X - open Predicate - open PredicateProperties {ℓ'' = ℓ''} X - -- ⇒ is Heyting implication - a⊓b≤c→a≤b⇒c : a b c (a b c) a (b c) - a⊓b≤c→a≤b⇒c a b c a⊓b≤c = - do - (a~ , a~proves) a⊓b≤c - let prover = (` a~ ̇ (` pair ̇ (# fzero) ̇ (# fone))) - return - (λ* prover , - λ x aₓ aₓ⊩ax bₓ bₓ⊩bx - subst - r r c x) - (sym (λ*ComputationRule prover (aₓ bₓ []))) - (a~proves - x - (pair aₓ bₓ) - ((subst r r a x) (sym (pr₁pxy≡x _ _)) aₓ⊩ax) , - (subst r r b x) (sym (pr₂pxy≡y _ _)) bₓ⊩bx)))) +module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where + PredicateX = Predicate {ℓ'' = ℓ''} X + open Predicate + open PredicateProperties {ℓ'' = ℓ''} X + -- ⇒ is Heyting implication + a⊓b≤c→a≤b⇒c : a b c (a b c) a (b c) + a⊓b≤c→a≤b⇒c a b c a⊓b≤c = + do + (a~ , a~proves) a⊓b≤c + let prover = (` a~ ̇ (` pair ̇ (# fzero) ̇ (# fone))) + return + (λ* prover , + λ x aₓ aₓ⊩ax bₓ bₓ⊩bx + subst + r r c x) + (sym (λ*ComputationRule prover (aₓ bₓ []))) + (a~proves + x + (pair aₓ bₓ) + ((subst r r a x) (sym (pr₁pxy≡x _ _)) aₓ⊩ax) , + (subst r r b x) (sym (pr₂pxy≡y _ _)) bₓ⊩bx)))) - a≤b⇒c→a⊓b≤c : a b c a (b c) (a b c) - a≤b⇒c→a⊓b≤c a b c a≤b⇒c = - do - (a~ , a~proves) a≤b⇒c - let prover = ` a~ ̇ (` pr₁ ̇ (# fzero)) ̇ (` pr₂ ̇ (# fzero)) - return - (λ* prover , - λ { x abₓ (pr₁abₓ⊩ax , pr₂abₓ⊩bx) - subst - r r c x) - (sym (λ*ComputationRule prover (abₓ []))) - (a~proves x (pr₁ abₓ) pr₁abₓ⊩ax (pr₂ abₓ) pr₂abₓ⊩bx) }) + a≤b⇒c→a⊓b≤c : a b c a (b c) (a b c) + a≤b⇒c→a⊓b≤c a b c a≤b⇒c = + do + (a~ , a~proves) a≤b⇒c + let prover = ` a~ ̇ (` pr₁ ̇ (# fzero)) ̇ (` pr₂ ̇ (# fzero)) + return + (λ* prover , + λ { x abₓ (pr₁abₓ⊩ax , pr₂abₓ⊩bx) + subst + r r c x) + (sym (λ*ComputationRule prover (abₓ []))) + (a~proves x (pr₁ abₓ) pr₁abₓ⊩ax (pr₂ abₓ) pr₂abₓ⊩bx) }) - ⇒isRightAdjointOf⊓ : a b c (a b c) (a b c) - ⇒isRightAdjointOf⊓ a b c = hPropExt (isProp≤ (a b) c) (isProp≤ a (b c)) (a⊓b≤c→a≤b⇒c a b c) (a≤b⇒c→a⊓b≤c a b c) + ⇒isRightAdjointOf⊓ : a b c (a b c) (a b c) + ⇒isRightAdjointOf⊓ a b c = hPropExt (isProp≤ (a b) c) (isProp≤ a (b c)) (a⊓b≤c→a≤b⇒c a b c) (a≤b⇒c→a⊓b≤c a b c) - antiSym→a⇒c≤b⇒c : a b c a b b a (a c) (b c) - antiSym→a⇒c≤b⇒c a b c a≤b b≤a = - do - (α , αProves) a≤b - (β , βProves) b≤a - let - prover : Term as 2 - prover = (# fzero) ̇ (` β ̇ # fone) - return - (λ* prover , - x r r⊩a⇒c r' r'⊩b - subst - witness witness c x) - (sym (λ*ComputationRule prover (r r' []))) - (r⊩a⇒c (β r') (βProves x r' r'⊩b)))) + antiSym→a⇒c≤b⇒c : a b c a b b a (a c) (b c) + antiSym→a⇒c≤b⇒c a b c a≤b b≤a = + do + (α , αProves) a≤b + (β , βProves) b≤a + let + prover : Term as 2 + prover = (# fzero) ̇ (` β ̇ # fone) + return + (λ* prover , + x r r⊩a⇒c r' r'⊩b + subst + witness witness c x) + (sym (λ*ComputationRule prover (r r' []))) + (r⊩a⇒c (β r') (βProves x r' r'⊩b)))) - antiSym→a⇒b≤a⇒c : a b c b c c b (a b) (a c) - antiSym→a⇒b≤a⇒c a b c b≤c c≤b = - do - (β , βProves) b≤c - (γ , γProves) c≤b - let - prover : Term as 2 - prover = ` β ̇ ((# fzero) ̇ (# fone)) - return - (λ* prover , - x α α⊩a⇒b a' a'⊩a - subst - r r c x) - (sym (λ*ComputationRule prover (α a' []))) - (βProves x (α a') (α⊩a⇒b a' a'⊩a)))) + antiSym→a⇒b≤a⇒c : a b c b c c b (a b) (a c) + antiSym→a⇒b≤a⇒c a b c b≤c c≤b = + do + (β , βProves) b≤c + (γ , γProves) c≤b + let + prover : Term as 2 + prover = ` β ̇ ((# fzero) ̇ (# fone)) + return + (λ* prover , + x α α⊩a⇒b a' a'⊩a + subst + r r c x) + (sym (λ*ComputationRule prover (α a' []))) + (βProves x (α a') (α⊩a⇒b a' a'⊩a)))) \ No newline at end of file diff --git a/docs/index.html b/docs/index.html index b4c0290..414742e 100644 --- a/docs/index.html +++ b/docs/index.html @@ -2,15 +2,13 @@ index
{-# OPTIONS --cubical #-}
 module index where
 
---open import Realizability.Partiality
---open import Realizability.PartialApplicativeStructure
---open import Realizability.PartialCombinatoryAlgebra
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure
-open import Realizability.Assembly.Everything
-open import Realizability.Tripos.Everything
-open import Realizability.Choice
-open import Tripoi.Tripos
-open import Tripoi.HeytingAlgebra
-open import Tripoi.PosetReflection
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.Assembly.Everything
+open import Realizability.Tripos.Everything
+open import Realizability.Topos.Everything
+open import Realizability.Choice
+open import Tripoi.Tripos
+open import Tripoi.HeytingAlgebra
+open import Tripoi.PosetReflection
 
\ No newline at end of file diff --git a/src/Realizability/Topos/Everything.agda b/src/Realizability/Topos/Everything.agda new file mode 100644 index 0000000..7cad4c4 --- /dev/null +++ b/src/Realizability/Topos/Everything.agda @@ -0,0 +1,4 @@ +module Realizability.Topos.Everything where + +open import Realizability.Topos.Object +open import Realizability.Topos.FunctionalRelation diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda new file mode 100644 index 0000000..3e537ac --- /dev/null +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -0,0 +1,90 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Data.Vec +open import Cubical.Data.Nat +open import Cubical.Data.FinData renaming (zero to fzero) +open import Cubical.Data.Sigma +open import Cubical.Data.Empty + +module Realizability.Topos.FunctionalRelation + {ℓ ℓ' ℓ''} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + where + +open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} +open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_) +open import Realizability.Tripos.Prealgebra.Predicate.Properties ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate' renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +private + _⊩_ : ∀ a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) → Type _ + a ⊩ P = a pre⊩ ∣ P ∣ tt* + +record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + open PartialEquivalenceRelation + field + perX : PartialEquivalenceRelation X + perY : PartialEquivalenceRelation Y + _~X_ = perX ._~_ + _~Y_ = perY ._~_ + + field + relation : Predicate (X × Y) + + private + `X : Sort + `X = ↑ (X , perX .isSetX) + + `Y : Sort + `Y = ↑ (Y , perY .isSetX) + + relationSymbol : Vec Sort 3 + relationSymbol = (`X `× `Y) ∷ `X `× `X ∷ `Y `× `Y ∷ [] + + relationSymbolInterpretation : RelationInterpretation relationSymbol + relationSymbolInterpretation fzero = relation + relationSymbolInterpretation one = _~X_ + relationSymbolInterpretation two = _~Y_ + + `relation : Fin 3 + `relation = fzero + `~X : Fin 3 + `~X = one + `~Y : Fin 3 + `~Y = two + + module RelationInterpretation = Interpretation relationSymbol relationSymbolInterpretation isNonTrivial + open RelationInterpretation + + module RelationSoundness = Soundness isNonTrivial relationSymbolInterpretation + open RelationSoundness + + -- Strictness + -- If the functional relation holds for x and y then x and y "exist" + private + strictnessContext : Context + strictnessContext = ([] ′ `X) ′ `Y + + x∈strictnessContext : `X ∈ strictnessContext + x∈strictnessContext = there here + + y∈strictnessContext : `Y ∈ strictnessContext + y∈strictnessContext = here + + xˢᵗ : Term strictnessContext `X + xˢᵗ = var x∈strictnessContext + + yˢᵗ : Term strictnessContext `Y + yˢᵗ = var y∈strictnessContext + diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 72dcff7..2c80f27 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -29,17 +29,18 @@ private _⊩_ : ∀ a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) → Type _ a ⊩ P = a pre⊩ ∣ P ∣ tt* -record ToposObject (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where +record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field isSetX : isSet X - `X : Sort - `X = ↑ (X , isSetX) + private + `X : Sort + `X = ↑ (X , isSetX) - `X×X : Sort - `X×X = `X `× `X + `X×X : Sort + `X×X = `X `× `X - ~relSym : Vec Sort 1 - ~relSym = `X×X ∷ [] + ~relSym : Vec Sort 1 + ~relSym = `X×X ∷ [] module X×XRelational = Relational ~relSym open X×XRelational @@ -47,8 +48,9 @@ record ToposObject (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-s field _~_ : Predicate ⟨ ⟦ lookup fzero ~relSym ⟧ˢ ⟩ - ~relSymInterpretation : RelationInterpretation ~relSym - ~relSymInterpretation = λ { fzero → _~_ } + private + ~relSymInterpretation : RelationInterpretation ~relSym + ~relSymInterpretation = λ { fzero → _~_ } module ~Interpretation = Interpretation ~relSym ~relSymInterpretation isNonTrivial open ~Interpretation @@ -57,24 +59,26 @@ record ToposObject (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-s open ~Soundness -- Partial equivalence relations - symContext : Context - symContext = ([] ′ `X) ′ `X - x∈symContext : `X ∈ symContext - x∈symContext = there here + private + symContext : Context + symContext = ([] ′ `X) ′ `X - y∈symContext : `X ∈ symContext - y∈symContext = here + x∈symContext : `X ∈ symContext + x∈symContext = there here - xˢ : Term symContext `X - xˢ = var x∈symContext + y∈symContext : `X ∈ symContext + y∈symContext = here - yˢ : Term symContext `X - yˢ = var y∈symContext + xˢ : Term symContext `X + xˢ = var x∈symContext + + yˢ : Term symContext `X + yˢ = var y∈symContext + + symPrequantFormula : Formula symContext + symPrequantFormula = rel fzero (xˢ `, yˢ) `→ rel fzero (yˢ `, xˢ) - symPrequantFormula : Formula symContext - symPrequantFormula = rel fzero (xˢ `, yˢ) `→ rel fzero (yˢ `, xˢ) - -- ~ ⊢ ∀ x. ∀ y. x ~ y → y ~ x symFormula : Formula [] symFormula = `∀ (`∀ symPrequantFormula) @@ -83,30 +87,31 @@ record ToposObject (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-s symWitness : A symIsWitnessed : symWitness ⊩ ⟦ symFormula ⟧ᶠ - transContext : Context - transContext = (([] ′ `X) ′ `X) ′ `X + private + transContext : Context + transContext = (([] ′ `X) ′ `X) ′ `X - z∈transContext : `X ∈ transContext - z∈transContext = here + z∈transContext : `X ∈ transContext + z∈transContext = here - y∈transContext : `X ∈ transContext - y∈transContext = there here + y∈transContext : `X ∈ transContext + y∈transContext = there here - x∈transContext : `X ∈ transContext - x∈transContext = there (there here) + x∈transContext : `X ∈ transContext + x∈transContext = there (there here) - zᵗ : Term transContext `X - zᵗ = var z∈transContext + zᵗ : Term transContext `X + zᵗ = var z∈transContext - yᵗ : Term transContext `X - yᵗ = var y∈transContext + yᵗ : Term transContext `X + yᵗ = var y∈transContext - xᵗ : Term transContext `X - xᵗ = var x∈transContext + xᵗ : Term transContext `X + xᵗ = var x∈transContext - -- ~ : Rel X X, x : X, y : X, z : X ⊢ x ~ y ∧ y ~ z → x ~ z - transPrequantFormula : Formula transContext - transPrequantFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) + -- ~ : Rel X X, x : X, y : X, z : X ⊢ x ~ y ∧ y ~ z → x ~ z + transPrequantFormula : Formula transContext + transPrequantFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) transFormula : Formula [] transFormula = `∀ (`∀ (`∀ transPrequantFormula)) diff --git a/src/index.agda b/src/index.agda index bd6dec1..83d4cfd 100644 --- a/src/index.agda +++ b/src/index.agda @@ -1,13 +1,11 @@ {-# OPTIONS --cubical #-} module index where ---open import Realizability.Partiality ---open import Realizability.PartialApplicativeStructure ---open import Realizability.PartialCombinatoryAlgebra open import Realizability.CombinatoryAlgebra open import Realizability.ApplicativeStructure open import Realizability.Assembly.Everything open import Realizability.Tripos.Everything +open import Realizability.Topos.Everything open import Realizability.Choice open import Tripoi.Tripos open import Tripoi.HeytingAlgebra From f6af9d5d10c6118481fee3082b10331d2c916595 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 30 Jan 2024 23:10:26 +0530 Subject: [PATCH 08/61] Define functional relations --- .../Topos/FunctionalRelation.agda | 120 ++++++++++++++++-- src/Realizability/Tripos/Logic/Syntax.agda | 1 + 2 files changed, 109 insertions(+), 12 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 3e537ac..78635da 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -52,25 +52,22 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc relationSymbol : Vec Sort 3 relationSymbol = (`X `× `Y) ∷ `X `× `X ∷ `Y `× `Y ∷ [] - relationSymbolInterpretation : RelationInterpretation relationSymbol - relationSymbolInterpretation fzero = relation - relationSymbolInterpretation one = _~X_ - relationSymbolInterpretation two = _~Y_ - - `relation : Fin 3 - `relation = fzero + `F : Fin 3 + `F = fzero `~X : Fin 3 `~X = one `~Y : Fin 3 `~Y = two - module RelationInterpretation = Interpretation relationSymbol relationSymbolInterpretation isNonTrivial - open RelationInterpretation + open Relational relationSymbol + + module RelationInterpretation' = Interpretation relationSymbol (λ { fzero → relation ; one → _~X_ ; two → _~Y_ }) isNonTrivial + open RelationInterpretation' - module RelationSoundness = Soundness isNonTrivial relationSymbolInterpretation + module RelationSoundness = Soundness {relSym = relationSymbol} isNonTrivial (λ { fzero → relation ; one → _~X_ ; two → _~Y_ }) open RelationSoundness - -- Strictness + -- # Strictness -- If the functional relation holds for x and y then x and y "exist" private strictnessContext : Context @@ -87,4 +84,103 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc yˢᵗ : Term strictnessContext `Y yˢᵗ = var y∈strictnessContext - + + -- F : Rel X Y, _~X_ : Rel X X, _~Y_ : Rel Y Y ⊢ F(x ,y) → (x ~X x) ∧ (y ~Y y) + strictnessPrequantFormula : Formula strictnessContext + strictnessPrequantFormula = rel `F (xˢᵗ `, yˢᵗ) `→ (rel `~X (xˢᵗ `, xˢᵗ) `∧ rel `~Y (yˢᵗ `, yˢᵗ)) + + strictnessFormula : Formula [] + strictnessFormula = `∀ (`∀ strictnessPrequantFormula) + + field + strictnessWitness : A + isStrict : strictnessWitness ⊩ ⟦ strictnessFormula ⟧ᶠ + + -- # Relational + -- The functional relation preserves equality + -- "Substitutive" might be a better term for this property + private + relationalContext : Context + relationalContext = + [] ′ `Y ′ `Y ′ `X ′ `X + + x₁∈relationalContext : `X ∈ relationalContext + x₁∈relationalContext = there here + + x₂∈relationalContext : `X ∈ relationalContext + x₂∈relationalContext = here + + y₁∈relationalContext : `Y ∈ relationalContext + y₁∈relationalContext = there (there here) + + y₂∈relationalContext : `Y ∈ relationalContext + y₂∈relationalContext = there (there (there here)) + + x₁ = var x₁∈relationalContext + x₂ = var x₂∈relationalContext + y₁ = var y₁∈relationalContext + y₂ = var y₂∈relationalContext + + relationalPrequantFormula : Formula relationalContext + relationalPrequantFormula = (rel `F (x₁ `, y₁) `∧ (rel `~X (x₁ `, x₂) `∧ rel `~Y (y₁ `, y₂))) `→ rel `F (x₂ `, y₂) + + relationalFormula : Formula [] + relationalFormula = `∀ (`∀ (`∀ (`∀ relationalPrequantFormula))) + + field + relationalWitness : A + isRelational : relationalWitness ⊩ ⟦ relationalFormula ⟧ᶠ + + -- # Single-valued + -- Self explanatory + private + singleValuedContext : Context + singleValuedContext = + [] ′ `Y ′ `Y ′ `X + + x∈singleValuedContext : `X ∈ singleValuedContext + x∈singleValuedContext = here + + y₁∈singleValuedContext : `Y ∈ singleValuedContext + y₁∈singleValuedContext = there here + + y₂∈singleValuedContext : `Y ∈ singleValuedContext + y₂∈singleValuedContext = there (there here) + + xˢᵛ = var x∈singleValuedContext + y₁ˢᵛ = var y₁∈singleValuedContext + y₂ˢᵛ = var y₂∈singleValuedContext + + singleValuedPrequantFormula : Formula singleValuedContext + singleValuedPrequantFormula = + (rel `F (xˢᵛ `, y₁ˢᵛ) `∧ rel `F (xˢᵛ `, y₂ˢᵛ)) `→ rel `~Y (y₁ˢᵛ `, y₂ˢᵛ) + + singleValuedFormula : Formula [] + singleValuedFormula = `∀ (`∀ (`∀ singleValuedPrequantFormula)) + + field + singleValuedWitness : A + isSingleValued : singleValuedWitness ⊩ ⟦ singleValuedFormula ⟧ᶠ + + -- # Total + -- For all existent elements in the domain x there is an element in the codomain y + -- such that F(x, y) + private + totalContext : Context + totalContext = + [] ′ `X + + x∈totalContext : `X ∈ totalContext + x∈totalContext = here + + xᵗˡ = var x∈totalContext + + totalPrequantFormula : Formula totalContext + totalPrequantFormula = rel `~X (xᵗˡ `, xᵗˡ) `→ `∃ (rel `F (renamingTerm (drop id) xᵗˡ `, var here)) + + totalFormula : Formula [] + totalFormula = `∀ totalPrequantFormula + + field + totalWitness : A + isTotal : totalWitness ⊩ ⟦ totalFormula ⟧ᶠ diff --git a/src/Realizability/Tripos/Logic/Syntax.agda b/src/Realizability/Tripos/Logic/Syntax.agda index e04c2b0..8a3cf87 100644 --- a/src/Realizability/Tripos/Logic/Syntax.agda +++ b/src/Realizability/Tripos/Logic/Syntax.agda @@ -17,6 +17,7 @@ data Sort : Type (ℓ-suc ℓ) where ⟦ ↑ a ⟧ˢ = a ⟦ a `× b ⟧ˢ = ⟦ a ⟧ˢ .fst × ⟦ b ⟧ˢ .fst , isSet× (⟦ a ⟧ˢ .snd) (⟦ b ⟧ˢ .snd) +infixl 30 _′_ data Context : Type (ℓ-suc ℓ) where [] : Context _′_ : Context → Sort → Context From 83130e959991224329987a6577651f16621366be Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 31 Jan 2024 02:07:58 +0530 Subject: [PATCH 09/61] Add introduction rules for natural deduction --- src/Realizability/Tripos/Logic/Semantics.agda | 39 +++++++++++++++++++ 1 file changed, 39 insertions(+) diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index 1ffe530..e0557a0 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -597,3 +597,42 @@ module Soundness (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) ((subst (λ r → r ⊩ ∣ ⟦ ϕ ⟧ᶠ ∣ γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) , (subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) })) + +module NaturalDeduction + {n} + {relSym : Vec Sort n} + (isNonTrivial : s ≡ k → ⊥) + (⟦_⟧ʳ : RelationInterpretation relSym) where + open Relational relSym + open Interpretation relSym ⟦_⟧ʳ isNonTrivial + open Soundness isNonTrivial ⟦_⟧ʳ + + -- ahh yes, the reduction of natural deduction to sequent calculus + -- cause I do not like sequent calculus + -- and it is much easier for formalisation to use natural deduction + + data PropContext (Γ : Context) : Type (ℓ-suc ℓ') where + [] : PropContext Γ + _`&_ : PropContext Γ → Formula Γ → PropContext Γ + + data _∈ᶠ_ : ∀ {Γ} → Formula Γ → PropContext Γ → Type (ℓ-suc ℓ') where + here : ∀ {Γ} {p : Formula Γ} → p ∈ᶠ ([] `& p) + there : ∀ {Γ Ξ ξ} {p : Formula Γ} → p ∈ᶠ Ξ → p ∈ᶠ (Ξ `& ξ) + + weakenPropContext : ∀ {Γ S} → PropContext Γ → PropContext (Γ ′ S) + weakenPropContext {Γ} {S} [] = [] + weakenPropContext {Γ} {S} (ctx `& x) = weakenPropContext ctx `& weakenFormula x + + -- Judgements of natural deduction + data _∣_⊢_ : (Γ : Context) → PropContext Γ → Formula Γ → Type (ℓ-suc ℓ') where + -- Introduction rules + intro⊤ : ∀ {Γ} {Ξ} → Γ ∣ Ξ ⊢ ⊤ᵗ + intro∧ : ∀ {Γ Ξ A B} → Γ ∣ Ξ ⊢ A → Γ ∣ Ξ ⊢ B → Γ ∣ Ξ ⊢ (A `∧ B) + intro∨L : ∀ {Γ Ξ A B} → Γ ∣ Ξ ⊢ A → Γ ∣ Ξ ⊢ (A `∨ B) + intro∨R : ∀ {Γ Ξ A B} → Γ ∣ Ξ ⊢ B → Γ ∣ Ξ ⊢ (A `∨ B) + intro∀ : ∀ {Γ} {S A} → {Ξ : PropContext Γ} → (Γ ′ S) ∣ weakenPropContext Ξ ⊢ A → Γ ∣ Ξ ⊢ `∀ A + intro∃ : ∀ {Γ S A Ξ} → (t : Term Γ S) → Γ ∣ Ξ ⊢ substitutionFormula (t , id) A → Γ ∣ Ξ ⊢ `∃ A + intro→ : ∀ {Γ Ξ A B} → Γ ∣ Ξ `& A ⊢ B → Γ ∣ Ξ ⊢ (A `→ B) + intro∈ : ∀ {Γ Ξ p} → p ∈ᶠ Ξ → Γ ∣ Ξ ⊢ p + + From 458d257b09c3fa08414e9174816e6834e8698be4 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 31 Jan 2024 02:08:25 +0530 Subject: [PATCH 10/61] Add functional relation stuff --- .../Topos/FunctionalRelation.agda | 87 +++++++++++++++++-- 1 file changed, 81 insertions(+), 6 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 78635da..0295ed8 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -31,8 +31,9 @@ private _⊩_ : ∀ a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) → Type _ a ⊩ P = a pre⊩ ∣ P ∣ tt* +open PartialEquivalenceRelation + record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - open PartialEquivalenceRelation field perX : PartialEquivalenceRelation X perY : PartialEquivalenceRelation Y @@ -42,13 +43,13 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc field relation : Predicate (X × Y) - private - `X : Sort - `X = ↑ (X , perX .isSetX) + `X : Sort + `X = ↑ (X , perX .isSetX) - `Y : Sort - `Y = ↑ (Y , perY .isSetX) + `Y : Sort + `Y = ↑ (Y , perY .isSetX) + private relationSymbol : Vec Sort 3 relationSymbol = (`X `× `Y) ∷ `X `× `X ∷ `Y `× `Y ∷ [] @@ -184,3 +185,77 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc field totalWitness : A isTotal : totalWitness ⊩ ⟦ totalFormula ⟧ᶠ + +open FunctionalRelation hiding (`X; `Y) + +pointwiseEntailment : ∀ {X Y : Type ℓ'} → FunctionalRelation X Y → FunctionalRelation X Y → Type _ +pointwiseEntailment {X} {Y} F G = Σ[ a ∈ A ] (a ⊩ ⟦ entailmentFormula ⟧ᶠ) where + + `X : Sort + `Y : Sort + + `X = ↑ (X , F .perX .isSetX) + `Y = ↑ (Y , F .perY .isSetX) + + relationSymbols : Vec Sort 2 + relationSymbols = (`X `× `Y) ∷ (`X `× `Y) ∷ [] + + `F : Fin 2 + `F = fzero + + `G : Fin 2 + `G = one + + open Relational relationSymbols + + module RelationalInterpretation = Interpretation relationSymbols (λ { fzero → F .relation ; one → G .relation }) isNonTrivial + open RelationalInterpretation + + entailmentContext : Context + entailmentContext = [] ′ `X ′ `Y + + x∈entailmentContext : `X ∈ entailmentContext + x∈entailmentContext = there here + + y∈entailmentContext : `Y ∈ entailmentContext + y∈entailmentContext = here + + x = var x∈entailmentContext + y = var y∈entailmentContext + + entailmentPrequantFormula : Formula entailmentContext + entailmentPrequantFormula = rel `F (x `, y) `→ rel `G (x `, y) + + entailmentFormula : Formula [] + entailmentFormula = `∀ (`∀ entailmentPrequantFormula) + + +functionalRelationIsomorphism : ∀ {X Y : Type ℓ'} → FunctionalRelation X Y → FunctionalRelation X Y → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') +functionalRelationIsomorphism {X} {Y} F G = + pointwiseEntailment F G × pointwiseEntailment G F + +pointwiseEntailment→functionalRelationIsomorphism : ∀ {X Y : Type ℓ'} → (F G : FunctionalRelation X Y) → pointwiseEntailment F G → functionalRelationIsomorphism F G +pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G (a , a⊩peFG) = {!!} where + + `X : Sort + `Y : Sort + + `X = ↑ (X , F .perX .isSetX) + `Y = ↑ (Y , F .perY .isSetX) + + relationSymbols : Vec Sort 4 + relationSymbols = (`X `× `Y) ∷ (`X `× `Y) ∷ (`X `× `X) ∷ (`Y `× `Y) ∷ [] + + `F : Fin 4 + `F = fzero + + `G : Fin 4 + `G = one + + `~X : Fin 4 + `~X = two + + `~Y : Fin 4 + `~Y = three + + From 3375b73ccc935cf65754f68ce6e2ccc3fc2c89b9 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Fri, 2 Feb 2024 22:34:16 +0530 Subject: [PATCH 11/61] Commit before refactor --- .../Topos/FunctionalRelation.agda | 60 ++++++++++++++++++- 1 file changed, 58 insertions(+), 2 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 0295ed8..5d1beb6 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -7,6 +7,9 @@ open import Cubical.Data.Nat open import Cubical.Data.FinData renaming (zero to fzero) open import Cubical.Data.Sigma open import Cubical.Data.Empty +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad module Realizability.Topos.FunctionalRelation {ℓ ℓ' ℓ''} @@ -19,6 +22,7 @@ open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_) open import Realizability.Tripos.Prealgebra.Predicate.Properties ca +open import Realizability.Tripos.Prealgebra.Meets.Identity ca open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial open CombinatoryAlgebra ca @@ -189,7 +193,7 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc open FunctionalRelation hiding (`X; `Y) pointwiseEntailment : ∀ {X Y : Type ℓ'} → FunctionalRelation X Y → FunctionalRelation X Y → Type _ -pointwiseEntailment {X} {Y} F G = Σ[ a ∈ A ] (a ⊩ ⟦ entailmentFormula ⟧ᶠ) where +pointwiseEntailment {X} {Y} F G = ∃[ a ∈ A ] (a ⊩ ⟦ entailmentFormula ⟧ᶠ) where `X : Sort `Y : Sort @@ -235,7 +239,7 @@ functionalRelationIsomorphism {X} {Y} F G = pointwiseEntailment F G × pointwiseEntailment G F pointwiseEntailment→functionalRelationIsomorphism : ∀ {X Y : Type ℓ'} → (F G : FunctionalRelation X Y) → pointwiseEntailment F G → functionalRelationIsomorphism F G -pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G (a , a⊩peFG) = {!!} where +pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , {!!} where `X : Sort `Y : Sort @@ -258,4 +262,56 @@ pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G (a , a⊩peFG) = `~Y : Fin 4 `~Y = three + open Interpretation relationSymbols (λ { fzero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_}) isNonTrivial + open Soundness {relSym = relationSymbols} isNonTrivial ((λ { fzero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_})) + open Relational relationSymbols + -- What we need to prove is that + -- F ≤ G ⊨ G ≤ F + -- We will use the semantic combinators we borrowed from the 1lab + -- We know that a ⊩ F ≤ G + -- or in other words + -- a ⊩ ∀ x. ∀ y. F(x , y) → G(x ,y) + + -- We will firstly use the introduction rule + -- for ∀ to get an x : X and a y : Y in context + proof : pointwiseEntailment G F + proof = + do + (a , a⊩F≤G) ← F≤G + let + context : Context + context = [] ′ `X ′ `Y + + x : Term context `X + x = var (there here) + + y : Term context `Y + y = var here + (b , b⊩) ← + `∀intro + {Γ = []} + {ϕ = ⊤ᵗ} + {B = `X} + {ψ = + `∀ {B = `Y} + (rel `G (x `, y) `→ rel `F (x `, y))} + (`∀intro + {Γ = [] ′ `X} + {ϕ = ⊤ᵗ} + {B = `Y} + {ψ = rel `G (x `, y) `→ rel `F (x `, y)} + (`→intro + {Γ = context} + {ϕ = ⊤ᵗ} + {ψ = rel `G (x `, y)} + {θ = rel `F (x `, y)} + (cut + {Γ = context} + {ϕ = ⊤ᵗ `∧ rel `G (x `, y)} + {ψ = rel `~X (x `, x)} + {θ = rel `F (x `, y)} + {!!} + {!!}))) + return (b , {!b⊩!}) + From 0a0ef50d189244317e4bb6c0f44706da2e52be02 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 6 Feb 2024 22:07:24 +0530 Subject: [PATCH 12/61] Start working on relational context structural operations --- .../Tripos/Logic/RelContextStructural.agda | 113 +++++++++ src/Realizability/Tripos/Logic/Semantics.agda | 214 ++++++++---------- 2 files changed, 204 insertions(+), 123 deletions(-) create mode 100644 src/Realizability/Tripos/Logic/RelContextStructural.agda diff --git a/src/Realizability/Tripos/Logic/RelContextStructural.agda b/src/Realizability/Tripos/Logic/RelContextStructural.agda new file mode 100644 index 0000000..d581dab --- /dev/null +++ b/src/Realizability/Tripos/Logic/RelContextStructural.agda @@ -0,0 +1,113 @@ +{-# OPTIONS --lossy-unification #-} +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Data.Vec +open import Cubical.Data.Nat +open import Cubical.Data.FinData +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Sum +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.Tripos.Logic.Syntax +import Realizability.Tripos.Logic.Semantics as Semantics + +module Realizability.Tripos.Logic.RelContextStructural where + +module WeakenSyntax + {n} + {ℓ} + (Ξ : Vec (Sort {ℓ = ℓ}) n) + (R : Sort) where + + private module SyntaxΞ = Relational Ξ + open SyntaxΞ renaming (Formula to ΞFormula) + + private module SyntaxΞR = Relational (R ∷ Ξ) + open SyntaxΞR renaming (Formula to ΞRFormula) + + transportAlongWeakening : ∀ {Γ} → ΞFormula Γ → ΞRFormula Γ + transportAlongWeakening {Γ} Relational.⊤ᵗ = SyntaxΞR.⊤ᵗ + transportAlongWeakening {Γ} Relational.⊥ᵗ = SyntaxΞR.⊥ᵗ + transportAlongWeakening {Γ} (form Relational.`∨ form₁) = transportAlongWeakening form SyntaxΞR.`∨ transportAlongWeakening form₁ + transportAlongWeakening {Γ} (form Relational.`∧ form₁) = transportAlongWeakening form SyntaxΞR.`∧ transportAlongWeakening form₁ + transportAlongWeakening {Γ} (form Relational.`→ form₁) = transportAlongWeakening form SyntaxΞR.`→ transportAlongWeakening form₁ + transportAlongWeakening {Γ} (Relational.`¬ form) = SyntaxΞR.`¬ transportAlongWeakening form + transportAlongWeakening {Γ} (Relational.`∃ form) = SyntaxΞR.`∃ (transportAlongWeakening form) + transportAlongWeakening {Γ} (Relational.`∀ form) = SyntaxΞR.`∀ (transportAlongWeakening form) + transportAlongWeakening {Γ} (Relational.rel k x) = SyntaxΞR.rel (suc k) x + +module WeakenSemantics + {n} + {ℓ ℓ' ℓ''} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + (Ξ : Vec (Sort {ℓ = ℓ'}) n) + (R : Sort {ℓ = ℓ'}) + (Ξsem : Semantics.RelationInterpretation {ℓ'' = ℓ''} ca Ξ) where + open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') + open import Realizability.Tripos.Prealgebra.Predicate.Properties ca + open CombinatoryAlgebra ca + open Combinators ca + open WeakenSyntax Ξ R + open Predicate' + open Semantics {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca + Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} + + module SyntaxΞ = Relational Ξ + open SyntaxΞ renaming (Formula to ΞFormula) + + module SyntaxΞR = Relational (R ∷ Ξ) + open SyntaxΞR renaming (Formula to ΞRFormula) + + module _ (Rsem : Predicate ⟨ ⟦ R ⟧ˢ ⟩) where + + RΞsem : Semantics.RelationInterpretation {ℓ'' = ℓ''} ca (R ∷ Ξ) + RΞsem zero = Rsem + RΞsem (suc f) = Ξsem f + + module InterpretationΞ = Interpretation Ξ Ξsem isNonTrivial + module InterpretationΞR = Interpretation (R ∷ Ξ) RΞsem isNonTrivial + module SoundnessΞ = Soundness {relSym = Ξ} isNonTrivial Ξsem + module SoundnessΞR = Soundness {relSym = R ∷ Ξ} isNonTrivial RΞsem + + syntacticTransportPreservesRealizers⁺ : ∀ {Γ} → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) → (r : A) → (f : ΞFormula Γ) → ∣ InterpretationΞ.⟦ f ⟧ᶠ ∣ γ r → r ⊩ ∣ InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ ∣ γ + syntacticTransportPreservesRealizers⁻ : ∀ {Γ} → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) → (r : A) → (f : ΞFormula Γ) → r ⊩ ∣ InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ ∣ γ → r ⊩ ∣ InterpretationΞ.⟦ f ⟧ᶠ ∣ γ + + syntacticTransportPreservesRealizers⁺ {Γ} γ r Relational.⊤ᵗ r⊩⟦f⟧Ξ = r⊩⟦f⟧Ξ + syntacticTransportPreservesRealizers⁺ {Γ} γ r (f Relational.`∨ f₁) r⊩⟦f⟧Ξ = + r⊩⟦f⟧Ξ >>= + λ { (inl (pr₁r≡k , pr₂r⊩⟦f⟧)) → ∣ inl (pr₁r≡k , syntacticTransportPreservesRealizers⁺ γ (pr₂ ⨾ r) f pr₂r⊩⟦f⟧) ∣₁ + ; (inr (pr₁r≡k' , pr₂r⊩⟦f₁⟧)) → ∣ inr (pr₁r≡k' , (syntacticTransportPreservesRealizers⁺ γ (pr₂ ⨾ r) f₁ pr₂r⊩⟦f₁⟧)) ∣₁ } + syntacticTransportPreservesRealizers⁺ {Γ} γ r (f Relational.`∧ f₁) (pr₁r⊩⟦f⟧ , pr₂r⊩⟦f₁⟧) = + syntacticTransportPreservesRealizers⁺ γ (pr₁ ⨾ r) f pr₁r⊩⟦f⟧ , + syntacticTransportPreservesRealizers⁺ γ (pr₂ ⨾ r) f₁ pr₂r⊩⟦f₁⟧ + syntacticTransportPreservesRealizers⁺ {Γ} γ r (f Relational.`→ f₁) r⊩⟦f⟧Ξ = + λ b b⊩⟦f⟧ΞR → syntacticTransportPreservesRealizers⁺ γ (r ⨾ b) f₁ (r⊩⟦f⟧Ξ b (syntacticTransportPreservesRealizers⁻ γ b f b⊩⟦f⟧ΞR)) + syntacticTransportPreservesRealizers⁺ {Γ} γ r (Relational.`¬ f) r⊩⟦f⟧Ξ = {!!} + syntacticTransportPreservesRealizers⁺ {Γ} γ r (Relational.`∃ f) r⊩⟦f⟧Ξ = + r⊩⟦f⟧Ξ >>= + λ { ((γ' , b) , γ'≡γ , r⊩⟦f⟧Ξγ'b) → ∣ (γ' , b) , (γ'≡γ , (syntacticTransportPreservesRealizers⁺ (γ' , b) r f r⊩⟦f⟧Ξγ'b)) ∣₁ } + syntacticTransportPreservesRealizers⁺ {Γ} γ r (Relational.`∀ f) r⊩⟦f⟧Ξ = + λ { b (γ' , b') γ'≡γ → syntacticTransportPreservesRealizers⁺ (γ' , b') (r ⨾ b) f (r⊩⟦f⟧Ξ b (γ' , b') γ'≡γ) } + syntacticTransportPreservesRealizers⁺ {Γ} γ r (Relational.rel Rsym t) r⊩⟦f⟧Ξ = + subst + (λ R' → r ⊩ ∣ R' ∣ (⟦ t ⟧ᵗ γ)) + (sym (RΞsem (suc Rsym) ≡⟨ refl ⟩ (Ξsem Rsym ∎))) + r⊩⟦f⟧Ξ + + syntacticTransportPreservesRealizers⁻ {Γ} γ r Relational.⊤ᵗ r⊩⟦f⟧ΞR = r⊩⟦f⟧ΞR + syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`∨ f₁) r⊩⟦f⟧ΞR = {!!} + syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`∧ f₁) r⊩⟦f⟧ΞR = {!!} + syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`→ f₁) r⊩⟦f⟧ΞR = {!!} + syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`¬ f) r⊩⟦f⟧ΞR = {!!} + syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`∃ f) r⊩⟦f⟧ΞR = {!!} + syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`∀ f) r⊩⟦f⟧ΞR = {!!} + syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.rel k₁ x) r⊩⟦f⟧ΞR = {!!} + + transportPreservesHoldsInTripos : ∀ {Γ} → (f : ΞFormula Γ) → SoundnessΞ.holdsInTripos f → SoundnessΞR.holdsInTripos (transportAlongWeakening f) + transportPreservesHoldsInTripos {Γ} f holds = {!!} + diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index e0557a0..f23506c 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -37,30 +37,100 @@ open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a t open Predicate' open PredicateProperties hiding (_≤_ ; isTrans≤) open Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''} -Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} +private + Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} RelationInterpretation : ∀ {n : ℕ} → (Vec Sort n) → Type _ RelationInterpretation {n} relSym = (∀ i → Predicate ⟨ ⟦ lookup i relSym ⟧ˢ ⟩) + +⟦_⟧ᶜ : Context → hSet ℓ' +⟦ [] ⟧ᶜ = Unit* , isSetUnit* +⟦ c ′ x ⟧ᶜ = (⟦ c ⟧ᶜ .fst × ⟦ x ⟧ˢ .fst) , isSet× (⟦ c ⟧ᶜ .snd) (⟦ x ⟧ˢ .snd) + +⟦_⟧ⁿ : ∀ {Γ s} → s ∈ Γ → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ s ⟧ˢ ⟩ +⟦_⟧ⁿ {.(_ ′ s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧ +⟦_⟧ⁿ {.(_ ′ _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) = ⟦ s∈Γ ⟧ⁿ ⟦Γ⟧ + +⟦_⟧ᵗ : ∀ {Γ s} → Term Γ s → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ s ⟧ˢ ⟩ +⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ = ⟦ x ⟧ⁿ ⟦Γ⟧ +⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = (⟦ t ⟧ᵗ ⟦Γ⟧) , (⟦ t₁ ⟧ᵗ ⟦Γ⟧) +⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst (⟦ t ⟧ᵗ ⟦Γ⟧) +⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd (⟦ t ⟧ᵗ ⟦Γ⟧) +⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x (⟦ t ⟧ᵗ ⟦Γ⟧) + +-- R for renamings and r for relations +⟦_⟧ᴿ : ∀ {Γ Δ} → Renaming Γ Δ → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ Δ ⟧ᶜ ⟩ +⟦ id ⟧ᴿ ctx = ctx +⟦ drop ren ⟧ᴿ (ctx , _) = ⟦ ren ⟧ᴿ ctx +⟦ keep ren ⟧ᴿ (ctx , s) = (⟦ ren ⟧ᴿ ctx) , s + +-- B for suBstitution and s for sorts +⟦_⟧ᴮ : ∀ {Γ Δ} → Substitution Γ Δ → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ Δ ⟧ᶜ ⟩ +⟦ id ⟧ᴮ ctx = ctx +⟦ t , sub ⟧ᴮ ctx = (⟦ sub ⟧ᴮ ctx) , (⟦ t ⟧ᵗ ctx) +⟦ drop sub ⟧ᴮ (ctx , s) = ⟦ sub ⟧ᴮ ctx + +renamingVarSound : ∀ {Γ Δ s} → (ren : Renaming Γ Δ) → (v : s ∈ Δ) → ⟦ renamingVar ren v ⟧ⁿ ≡ ⟦ v ⟧ⁿ ∘ ⟦ ren ⟧ᴿ +renamingVarSound {Γ} {.Γ} {s} id v = refl +renamingVarSound {.(_ ′ _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren v i ⟦Γ⟧ } +renamingVarSound {.(_ ′ s)} {.(_ ′ s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → ⟦s⟧ } +renamingVarSound {.(_ ′ _)} {.(_ ′ _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren v i ⟦Γ⟧ } + +renamingTermSound : ∀ {Γ Δ s} → (ren : Renaming Γ Δ) → (t : Term Δ s) → ⟦ renamingTerm ren t ⟧ᵗ ≡ ⟦ t ⟧ᵗ ∘ ⟦ ren ⟧ᴿ +renamingTermSound {Γ} {.Γ} {s} id t = refl +renamingTermSound {.(_ ′ _)} {Δ} {s} (drop ren) (var x) = + funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren x i ⟦Γ⟧ } +renamingTermSound {.(_ ′ _)} {Δ} {.(_ `× _)} r@(drop ren) (t `, t₁) = + funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) } +renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₁ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } +renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₂ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } +renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (fun f t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } +renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (var v) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingVarSound r v i x } +renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {.(_ `× _)} r@(keep ren) (t `, t₁) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) } +renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₁ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } +renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₂ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } +renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (fun f t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } + + +substitutionVarSound : ∀ {Γ Δ s} → (subs : Substitution Γ Δ) → (v : s ∈ Δ) → ⟦ substitutionVar subs v ⟧ᵗ ≡ ⟦ v ⟧ⁿ ∘ ⟦ subs ⟧ᴮ +substitutionVarSound {Γ} {.Γ} {s} id t = refl +substitutionVarSound {Γ} {.(_ ′ s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧ → refl +substitutionVarSound {Γ} {.(_ ′ _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i → substitutionVarSound subs t i ⟦Γ⟧ +substitutionVarSound {.(_ ′ _)} {Δ} {s} r@(drop subs) t = + -- TODO : Fix unsolved constraints + funExt + λ { x@(⟦Γ⟧ , ⟦s⟧) → + ⟦ substitutionVar (drop subs) t ⟧ᵗ x + ≡[ i ]⟨ renamingTermSound (drop id) (substitutionVar subs t) i x ⟩ + ⟦ substitutionVar subs t ⟧ᵗ (⟦ drop id ⟧ᴿ x) + ≡⟨ refl ⟩ + ⟦ substitutionVar subs t ⟧ᵗ ⟦Γ⟧ + ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ ⟩ + ⟦ t ⟧ⁿ (⟦ subs ⟧ᴮ ⟦Γ⟧) + ∎} + +substitutionTermSound : ∀ {Γ Δ s} → (subs : Substitution Γ Δ) → (t : Term Δ s) → ⟦ substitutionTerm subs t ⟧ᵗ ≡ ⟦ t ⟧ᵗ ∘ ⟦ subs ⟧ᴮ +substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x +substitutionTermSound {Γ} {Δ} {.(_ `× _)} subs (t `, t₁) = funExt λ x i → (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x) +substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i → substitutionTermSound subs t i x .fst +substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i → substitutionTermSound subs t i x .snd +substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i → f (substitutionTermSound subs t i x) + +semanticSubstitution : ∀ {Γ Δ} → (subs : Substitution Γ Δ) → Predicate ⟨ ⟦ Δ ⟧ᶜ ⟩ → Predicate ⟨ ⟦ Γ ⟧ᶜ ⟩ +semanticSubstitution {Γ} {Δ} subs = ⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ Δ ⟧ᶜ) ⟦ subs ⟧ᴮ + module Interpretation {n : ℕ} (relSym : Vec Sort n) (⟦_⟧ʳ : RelationInterpretation relSym) (isNonTrivial : s ≡ k → ⊥) where open Relational relSym - - ⟦_⟧ᶜ : Context → hSet ℓ' - ⟦ [] ⟧ᶜ = Unit* , isSetUnit* - ⟦ c ′ x ⟧ᶜ = (⟦ c ⟧ᶜ .fst × ⟦ x ⟧ˢ .fst) , isSet× (⟦ c ⟧ᶜ .snd) (⟦ x ⟧ˢ .snd) - - ⟦_⟧ⁿ : ∀ {Γ s} → s ∈ Γ → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ s ⟧ˢ ⟩ - ⟦_⟧ⁿ {.(_ ′ s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧ - ⟦_⟧ⁿ {.(_ ′ _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) = ⟦ s∈Γ ⟧ⁿ ⟦Γ⟧ - - ⟦_⟧ᵗ : ∀ {Γ s} → Term Γ s → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ s ⟧ˢ ⟩ - ⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ = ⟦ x ⟧ⁿ ⟦Γ⟧ - ⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = (⟦ t ⟧ᵗ ⟦Γ⟧) , (⟦ t₁ ⟧ᵗ ⟦Γ⟧) - ⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst (⟦ t ⟧ᵗ ⟦Γ⟧) - ⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd (⟦ t ⟧ᵗ ⟦Γ⟧) - ⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x (⟦ t ⟧ᵗ ⟦Γ⟧) - ⟦_⟧ᶠ : ∀ {Γ} → Formula Γ → Predicate ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦_⟧ᶠ {Γ} ⊤ᵗ = pre1 ⟨ ⟦ Γ ⟧ᶜ ⟩ (str ⟦ Γ ⟧ᶜ) isNonTrivial ⟦_⟧ᶠ {Γ} ⊥ᵗ = pre0 ⟨ ⟦ Γ ⟧ᶜ ⟩ (str ⟦ Γ ⟧ᶜ) isNonTrivial @@ -72,74 +142,6 @@ module Interpretation ⟦_⟧ᶠ {Γ} (`∀ {B = B} form) = `∀[_] (isSet× (str ⟦ Γ ⟧ᶜ) (str ⟦ B ⟧ˢ)) (str ⟦ Γ ⟧ᶜ) (λ { (⟦Γ⟧ , ⟦B⟧) → ⟦Γ⟧ }) ⟦ form ⟧ᶠ ⟦_⟧ᶠ {Γ} (rel R t) = ⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ lookup R relSym ⟧ˢ) ⟦ t ⟧ᵗ ⟦ R ⟧ʳ - -- R for renamings and r for relations - ⟦_⟧ᴿ : ∀ {Γ Δ} → Renaming Γ Δ → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ Δ ⟧ᶜ ⟩ - ⟦ id ⟧ᴿ ctx = ctx - ⟦ drop ren ⟧ᴿ (ctx , _) = ⟦ ren ⟧ᴿ ctx - ⟦ keep ren ⟧ᴿ (ctx , s) = (⟦ ren ⟧ᴿ ctx) , s - - -- B for suBstitution and s for sorts - ⟦_⟧ᴮ : ∀ {Γ Δ} → Substitution Γ Δ → ⟨ ⟦ Γ ⟧ᶜ ⟩ → ⟨ ⟦ Δ ⟧ᶜ ⟩ - ⟦ id ⟧ᴮ ctx = ctx - ⟦ t , sub ⟧ᴮ ctx = (⟦ sub ⟧ᴮ ctx) , (⟦ t ⟧ᵗ ctx) - ⟦ drop sub ⟧ᴮ (ctx , s) = ⟦ sub ⟧ᴮ ctx - - renamingVarSound : ∀ {Γ Δ s} → (ren : Renaming Γ Δ) → (v : s ∈ Δ) → ⟦ renamingVar ren v ⟧ⁿ ≡ ⟦ v ⟧ⁿ ∘ ⟦ ren ⟧ᴿ - renamingVarSound {Γ} {.Γ} {s} id v = refl - renamingVarSound {.(_ ′ _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren v i ⟦Γ⟧ } - renamingVarSound {.(_ ′ s)} {.(_ ′ s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → ⟦s⟧ } - renamingVarSound {.(_ ′ _)} {.(_ ′ _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren v i ⟦Γ⟧ } - - renamingTermSound : ∀ {Γ Δ s} → (ren : Renaming Γ Δ) → (t : Term Δ s) → ⟦ renamingTerm ren t ⟧ᵗ ≡ ⟦ t ⟧ᵗ ∘ ⟦ ren ⟧ᴿ - renamingTermSound {Γ} {.Γ} {s} id t = refl - renamingTermSound {.(_ ′ _)} {Δ} {s} (drop ren) (var x) = - funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingVarSound ren x i ⟦Γ⟧ } - renamingTermSound {.(_ ′ _)} {Δ} {.(_ `× _)} r@(drop ren) (t `, t₁) = - funExt λ { (⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) } - renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₁ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } - renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (π₂ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } - renamingTermSound {.(_ ′ _)} {Δ} {s} r@(drop ren) (fun f t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (var v) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingVarSound r v i x } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {.(_ `× _)} r@(keep ren) (t `, t₁) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₁ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .fst } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (π₂ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → renamingTermSound r t i x .snd } - renamingTermSound {.(_ ′ _)} {.(_ ′ _)} {s} r@(keep ren) (fun f t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i → f (renamingTermSound r t i x) } - - substitutionVarSound : ∀ {Γ Δ s} → (subs : Substitution Γ Δ) → (v : s ∈ Δ) → ⟦ substitutionVar subs v ⟧ᵗ ≡ ⟦ v ⟧ⁿ ∘ ⟦ subs ⟧ᴮ - substitutionVarSound {Γ} {.Γ} {s} id t = refl - substitutionVarSound {Γ} {.(_ ′ s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧ → refl - substitutionVarSound {Γ} {.(_ ′ _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i → substitutionVarSound subs t i ⟦Γ⟧ - substitutionVarSound {.(_ ′ _)} {Δ} {s} r@(drop subs) t = - -- TODO : Fix unsolved constraints - funExt - λ { x@(⟦Γ⟧ , ⟦s⟧) → - ⟦ substitutionVar (drop subs) t ⟧ᵗ x - ≡[ i ]⟨ renamingTermSound (drop id) (substitutionVar subs t) i x ⟩ - ⟦ substitutionVar subs t ⟧ᵗ (⟦ drop id ⟧ᴿ x) - ≡⟨ refl ⟩ - ⟦ substitutionVar subs t ⟧ᵗ ⟦Γ⟧ - ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ ⟩ - ⟦ t ⟧ⁿ (⟦ subs ⟧ᴮ ⟦Γ⟧) - ∎} - - substitutionTermSound : ∀ {Γ Δ s} → (subs : Substitution Γ Δ) → (t : Term Δ s) → ⟦ substitutionTerm subs t ⟧ᵗ ≡ ⟦ t ⟧ᵗ ∘ ⟦ subs ⟧ᴮ - substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x - substitutionTermSound {Γ} {Δ} {.(_ `× _)} subs (t `, t₁) = funExt λ x i → (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x) - substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i → substitutionTermSound subs t i x .fst - substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i → substitutionTermSound subs t i x .snd - substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i → f (substitutionTermSound subs t i x) - - semanticSubstitution : ∀ {Γ Δ} → (subs : Substitution Γ Δ) → Predicate ⟨ ⟦ Δ ⟧ᶜ ⟩ → Predicate ⟨ ⟦ Γ ⟧ᶜ ⟩ - semanticSubstitution {Γ} {Δ} subs = ⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ Δ ⟧ᶜ) ⟦ subs ⟧ᴮ - -- Due to a shortcut in the soundness of negation termination checking fails -- TODO : Fix {-# TERMINATING #-} @@ -295,6 +297,9 @@ module Soundness entails = _⊨_ + holdsInTripos : ∀ {Γ} → Formula Γ → Type (ℓ-max (ℓ-max ℓ ℓ'') ℓ') + holdsInTripos {Γ} form = ⊤ᵗ ⊨ form + private variable Γ Δ : Context @@ -598,41 +603,4 @@ module Soundness ((subst (λ r → r ⊩ ∣ ⟦ ϕ ⟧ᶠ ∣ γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) , (subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) })) -module NaturalDeduction - {n} - {relSym : Vec Sort n} - (isNonTrivial : s ≡ k → ⊥) - (⟦_⟧ʳ : RelationInterpretation relSym) where - open Relational relSym - open Interpretation relSym ⟦_⟧ʳ isNonTrivial - open Soundness isNonTrivial ⟦_⟧ʳ - - -- ahh yes, the reduction of natural deduction to sequent calculus - -- cause I do not like sequent calculus - -- and it is much easier for formalisation to use natural deduction - - data PropContext (Γ : Context) : Type (ℓ-suc ℓ') where - [] : PropContext Γ - _`&_ : PropContext Γ → Formula Γ → PropContext Γ - - data _∈ᶠ_ : ∀ {Γ} → Formula Γ → PropContext Γ → Type (ℓ-suc ℓ') where - here : ∀ {Γ} {p : Formula Γ} → p ∈ᶠ ([] `& p) - there : ∀ {Γ Ξ ξ} {p : Formula Γ} → p ∈ᶠ Ξ → p ∈ᶠ (Ξ `& ξ) - - weakenPropContext : ∀ {Γ S} → PropContext Γ → PropContext (Γ ′ S) - weakenPropContext {Γ} {S} [] = [] - weakenPropContext {Γ} {S} (ctx `& x) = weakenPropContext ctx `& weakenFormula x - - -- Judgements of natural deduction - data _∣_⊢_ : (Γ : Context) → PropContext Γ → Formula Γ → Type (ℓ-suc ℓ') where - -- Introduction rules - intro⊤ : ∀ {Γ} {Ξ} → Γ ∣ Ξ ⊢ ⊤ᵗ - intro∧ : ∀ {Γ Ξ A B} → Γ ∣ Ξ ⊢ A → Γ ∣ Ξ ⊢ B → Γ ∣ Ξ ⊢ (A `∧ B) - intro∨L : ∀ {Γ Ξ A B} → Γ ∣ Ξ ⊢ A → Γ ∣ Ξ ⊢ (A `∨ B) - intro∨R : ∀ {Γ Ξ A B} → Γ ∣ Ξ ⊢ B → Γ ∣ Ξ ⊢ (A `∨ B) - intro∀ : ∀ {Γ} {S A} → {Ξ : PropContext Γ} → (Γ ′ S) ∣ weakenPropContext Ξ ⊢ A → Γ ∣ Ξ ⊢ `∀ A - intro∃ : ∀ {Γ S A Ξ} → (t : Term Γ S) → Γ ∣ Ξ ⊢ substitutionFormula (t , id) A → Γ ∣ Ξ ⊢ `∃ A - intro→ : ∀ {Γ Ξ A B} → Γ ∣ Ξ `& A ⊢ B → Γ ∣ Ξ ⊢ (A `→ B) - intro∈ : ∀ {Γ Ξ p} → p ∈ᶠ Ξ → Γ ∣ Ξ ⊢ p - - + From 1019d7993aedb775c347f34b6632a86e26b1bfdd Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 8 Feb 2024 19:49:11 +0530 Subject: [PATCH 13/61] Prove that weakening of relational context preserves realizers --- .../Tripos/Logic/RelContextStructural.agda | 67 +++++++++++++++---- src/Realizability/Tripos/Logic/Semantics.agda | 3 - src/Realizability/Tripos/Logic/Syntax.agda | 6 +- 3 files changed, 58 insertions(+), 18 deletions(-) diff --git a/src/Realizability/Tripos/Logic/RelContextStructural.agda b/src/Realizability/Tripos/Logic/RelContextStructural.agda index d581dab..1d4afca 100644 --- a/src/Realizability/Tripos/Logic/RelContextStructural.agda +++ b/src/Realizability/Tripos/Logic/RelContextStructural.agda @@ -1,6 +1,7 @@ {-# OPTIONS --lossy-unification #-} open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure +open import Cubical.Foundations.Univalence open import Cubical.Data.Vec open import Cubical.Data.Nat open import Cubical.Data.FinData @@ -34,7 +35,6 @@ module WeakenSyntax transportAlongWeakening {Γ} (form Relational.`∨ form₁) = transportAlongWeakening form SyntaxΞR.`∨ transportAlongWeakening form₁ transportAlongWeakening {Γ} (form Relational.`∧ form₁) = transportAlongWeakening form SyntaxΞR.`∧ transportAlongWeakening form₁ transportAlongWeakening {Γ} (form Relational.`→ form₁) = transportAlongWeakening form SyntaxΞR.`→ transportAlongWeakening form₁ - transportAlongWeakening {Γ} (Relational.`¬ form) = SyntaxΞR.`¬ transportAlongWeakening form transportAlongWeakening {Γ} (Relational.`∃ form) = SyntaxΞR.`∃ (transportAlongWeakening form) transportAlongWeakening {Γ} (Relational.`∀ form) = SyntaxΞR.`∀ (transportAlongWeakening form) transportAlongWeakening {Γ} (Relational.rel k x) = SyntaxΞR.rel (suc k) x @@ -74,8 +74,22 @@ module WeakenSemantics module SoundnessΞ = Soundness {relSym = Ξ} isNonTrivial Ξsem module SoundnessΞR = Soundness {relSym = R ∷ Ξ} isNonTrivial RΞsem - syntacticTransportPreservesRealizers⁺ : ∀ {Γ} → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) → (r : A) → (f : ΞFormula Γ) → ∣ InterpretationΞ.⟦ f ⟧ᶠ ∣ γ r → r ⊩ ∣ InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ ∣ γ - syntacticTransportPreservesRealizers⁻ : ∀ {Γ} → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) → (r : A) → (f : ΞFormula Γ) → r ⊩ ∣ InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ ∣ γ → r ⊩ ∣ InterpretationΞ.⟦ f ⟧ᶠ ∣ γ + syntacticTransportPreservesRealizers⁺ : + ∀ {Γ} + → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) + → (r : A) + → (f : ΞFormula Γ) + → ∣ InterpretationΞ.⟦ f ⟧ᶠ ∣ γ r + ------------------------------------------------------------- + → r ⊩ ∣ InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ ∣ γ + syntacticTransportPreservesRealizers⁻ : + ∀ {Γ} + → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) + → (r : A) + → (f : ΞFormula Γ) + → r ⊩ ∣ InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ ∣ γ + ------------------------------------------------------------- + → r ⊩ ∣ InterpretationΞ.⟦ f ⟧ᶠ ∣ γ syntacticTransportPreservesRealizers⁺ {Γ} γ r Relational.⊤ᵗ r⊩⟦f⟧Ξ = r⊩⟦f⟧Ξ syntacticTransportPreservesRealizers⁺ {Γ} γ r (f Relational.`∨ f₁) r⊩⟦f⟧Ξ = @@ -87,7 +101,6 @@ module WeakenSemantics syntacticTransportPreservesRealizers⁺ γ (pr₂ ⨾ r) f₁ pr₂r⊩⟦f₁⟧ syntacticTransportPreservesRealizers⁺ {Γ} γ r (f Relational.`→ f₁) r⊩⟦f⟧Ξ = λ b b⊩⟦f⟧ΞR → syntacticTransportPreservesRealizers⁺ γ (r ⨾ b) f₁ (r⊩⟦f⟧Ξ b (syntacticTransportPreservesRealizers⁻ γ b f b⊩⟦f⟧ΞR)) - syntacticTransportPreservesRealizers⁺ {Γ} γ r (Relational.`¬ f) r⊩⟦f⟧Ξ = {!!} syntacticTransportPreservesRealizers⁺ {Γ} γ r (Relational.`∃ f) r⊩⟦f⟧Ξ = r⊩⟦f⟧Ξ >>= λ { ((γ' , b) , γ'≡γ , r⊩⟦f⟧Ξγ'b) → ∣ (γ' , b) , (γ'≡γ , (syntacticTransportPreservesRealizers⁺ (γ' , b) r f r⊩⟦f⟧Ξγ'b)) ∣₁ } @@ -100,14 +113,44 @@ module WeakenSemantics r⊩⟦f⟧Ξ syntacticTransportPreservesRealizers⁻ {Γ} γ r Relational.⊤ᵗ r⊩⟦f⟧ΞR = r⊩⟦f⟧ΞR - syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`∨ f₁) r⊩⟦f⟧ΞR = {!!} - syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`∧ f₁) r⊩⟦f⟧ΞR = {!!} - syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`→ f₁) r⊩⟦f⟧ΞR = {!!} - syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`¬ f) r⊩⟦f⟧ΞR = {!!} - syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`∃ f) r⊩⟦f⟧ΞR = {!!} - syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`∀ f) r⊩⟦f⟧ΞR = {!!} - syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.rel k₁ x) r⊩⟦f⟧ΞR = {!!} + syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`∨ f₁) r⊩⟦f⟧ΞR = + r⊩⟦f⟧ΞR >>= + λ { (inl (pr₁r≡k , pr₂r⊩⟦f⟧)) → + ∣ inl (pr₁r≡k , (syntacticTransportPreservesRealizers⁻ γ (pr₂ ⨾ r) f pr₂r⊩⟦f⟧)) ∣₁ + ; (inr (pr₁r≡k' , pr₂r⊩⟦f₁⟧)) → + ∣ inr (pr₁r≡k' , (syntacticTransportPreservesRealizers⁻ γ (pr₂ ⨾ r) f₁ pr₂r⊩⟦f₁⟧)) ∣₁ } + syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`∧ f₁) r⊩⟦f⟧ΞR = + (syntacticTransportPreservesRealizers⁻ γ (pr₁ ⨾ r) f (r⊩⟦f⟧ΞR .fst)) , + (syntacticTransportPreservesRealizers⁻ γ (pr₂ ⨾ r) f₁ (r⊩⟦f⟧ΞR .snd)) + syntacticTransportPreservesRealizers⁻ {Γ} γ r (f Relational.`→ f₁) r⊩⟦f→f₁⟧ΞR = + λ b b⊩⟦f⟧Ξ → syntacticTransportPreservesRealizers⁻ γ (r ⨾ b) f₁ (r⊩⟦f→f₁⟧ΞR b (syntacticTransportPreservesRealizers⁺ γ b f b⊩⟦f⟧Ξ)) + syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`∃ f) r⊩⟦f⟧ΞR = + r⊩⟦f⟧ΞR >>= + λ { ((γ' , b) , γ'≡γ , r⊩⟦f⟧γ'b) → ∣ (γ' , b) , γ'≡γ , (syntacticTransportPreservesRealizers⁻ (γ' , b) r f r⊩⟦f⟧γ'b) ∣₁ } + syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.`∀ f) r⊩⟦f⟧ΞR = + λ { b (γ' , b') γ'≡γ → syntacticTransportPreservesRealizers⁻ (γ' , b') (r ⨾ b) f (r⊩⟦f⟧ΞR b (γ' , b') γ'≡γ) } + syntacticTransportPreservesRealizers⁻ {Γ} γ r (Relational.rel Rsym t) r⊩⟦f⟧ΞR = + subst + (λ R' → r ⊩ ∣ R' ∣ (⟦ t ⟧ᵗ γ)) + (sym (RΞsem (suc Rsym) ≡⟨ refl ⟩ (Ξsem Rsym ∎))) + r⊩⟦f⟧ΞR + + syntacticTransportPreservesRealizers : + ∀ {Γ} + → (γ : ⟨ ⟦ Γ ⟧ᶜ ⟩) + → (r : A) + → (f : ΞFormula Γ) + ---------------------------------------------------------- + → r ⊩ ∣ InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ ∣ γ + ≡ r ⊩ ∣ InterpretationΞ.⟦ f ⟧ᶠ ∣ γ + syntacticTransportPreservesRealizers {Γ} γ r f = + hPropExt + (InterpretationΞR.⟦ transportAlongWeakening f ⟧ᶠ .isPropValued γ r) + (InterpretationΞ.⟦ f ⟧ᶠ .isPropValued γ r) + (λ r⊩⟦f⟧ΞR → syntacticTransportPreservesRealizers⁻ γ r f r⊩⟦f⟧ΞR) + λ r⊩⟦f⟧Ξ → syntacticTransportPreservesRealizers⁺ γ r f r⊩⟦f⟧Ξ transportPreservesHoldsInTripos : ∀ {Γ} → (f : ΞFormula Γ) → SoundnessΞ.holdsInTripos f → SoundnessΞR.holdsInTripos (transportAlongWeakening f) - transportPreservesHoldsInTripos {Γ} f holds = {!!} + transportPreservesHoldsInTripos {Γ} f holds = + {!!} diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index f23506c..231925d 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -137,7 +137,6 @@ module Interpretation ⟦_⟧ᶠ {Γ} (form `∨ form₁) = _⊔_ ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦ form ⟧ᶠ ⟦ form₁ ⟧ᶠ ⟦_⟧ᶠ {Γ} (form `∧ form₁) = _⊓_ ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦ form ⟧ᶠ ⟦ form₁ ⟧ᶠ ⟦_⟧ᶠ {Γ} (form `→ form₁) = _⇒_ ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦ form ⟧ᶠ ⟦ form₁ ⟧ᶠ - ⟦_⟧ᶠ {Γ} (`¬ form) = _⇒_ ⟨ ⟦ Γ ⟧ᶜ ⟩ ⟦ form ⟧ᶠ ⟦ ⊥ᵗ {Γ = Γ} ⟧ᶠ ⟦_⟧ᶠ {Γ} (`∃ {B = B} form) = `∃[_] (isSet× (str ⟦ Γ ⟧ᶜ) (str ⟦ B ⟧ˢ)) (str ⟦ Γ ⟧ᶜ) (λ { (⟦Γ⟧ , ⟦B⟧) → ⟦Γ⟧ }) ⟦ form ⟧ᶠ ⟦_⟧ᶠ {Γ} (`∀ {B = B} form) = `∀[_] (isSet× (str ⟦ Γ ⟧ᶜ) (str ⟦ B ⟧ˢ)) (str ⟦ Γ ⟧ᶜ) (λ { (⟦Γ⟧ , ⟦B⟧) → ⟦Γ⟧ }) ⟦ form ⟧ᶠ ⟦_⟧ᶠ {Γ} (rel R t) = ⋆_ (str ⟦ Γ ⟧ᶜ) (str ⟦ lookup R relSym ⟧ˢ) ⟦ t ⟧ᵗ ⟦ R ⟧ʳ @@ -215,8 +214,6 @@ module Interpretation (λ form → b ⊩ ∣ form ∣ γ) (substitutionFormulaSound subs f) b⊩substFormulaFs)) - substitutionFormulaSound {Γ} {Δ} subs (`¬ f) = - substitutionFormulaSound subs (f `→ ⊥ᵗ) substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) = Predicate≡ ⟨ ⟦ Γ ⟧ᶜ ⟩ diff --git a/src/Realizability/Tripos/Logic/Syntax.agda b/src/Realizability/Tripos/Logic/Syntax.agda index 8a3cf87..ee80ffc 100644 --- a/src/Realizability/Tripos/Logic/Syntax.agda +++ b/src/Realizability/Tripos/Logic/Syntax.agda @@ -93,19 +93,19 @@ module Relational {n} (relSym : Vec Sort n) where ⊥ᵗ : ∀ {Γ} → Formula Γ _`∨_ : ∀ {Γ} → Formula Γ → Formula Γ → Formula Γ _`∧_ : ∀ {Γ} → Formula Γ → Formula Γ → Formula Γ - _`→_ : ∀ {Γ} → Formula Γ → Formula Γ → Formula Γ - `¬_ : ∀ {Γ} → Formula Γ → Formula Γ + _`→_ : ∀ {Γ} → Formula Γ → Formula Γ → Formula Γ `∃ : ∀ {Γ B} → Formula (Γ ′ B) → Formula Γ `∀ : ∀ {Γ B} → Formula (Γ ′ B) → Formula Γ rel : ∀ {Γ} (k : Fin n) → Term Γ (lookup k relSym) → Formula Γ + pattern `¬ f = f `→ ⊥ᵗ + substitutionFormula : ∀ {Γ Δ} → Substitution Γ Δ → Formula Δ → Formula Γ substitutionFormula {Γ} {Δ} subs ⊤ᵗ = ⊤ᵗ substitutionFormula {Γ} {Δ} subs ⊥ᵗ = ⊥ᵗ substitutionFormula {Γ} {Δ} subs (form `∨ form₁) = substitutionFormula subs form `∨ substitutionFormula subs form₁ substitutionFormula {Γ} {Δ} subs (form `∧ form₁) = substitutionFormula subs form `∧ substitutionFormula subs form₁ substitutionFormula {Γ} {Δ} subs (form `→ form₁) = substitutionFormula subs form `→ substitutionFormula subs form₁ - substitutionFormula {Γ} {Δ} subs (`¬ form) = `¬ substitutionFormula subs form substitutionFormula {Γ} {Δ} subs (`∃ form) = `∃ (substitutionFormula (var here , drop subs) form ) substitutionFormula {Γ} {Δ} subs (`∀ form) = `∀ (substitutionFormula (var here , drop subs) form) substitutionFormula {Γ} {Δ} subs (rel k x) = rel k (substitutionTerm subs x) From bce456a87a48465354d4839bc16b896338c4621d Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 8 Feb 2024 19:49:31 +0530 Subject: [PATCH 14/61] Update Topos Modules --- .../Topos/FunctionalRelation.agda | 198 ++++++++++-------- src/Realizability/Topos/Object.agda | 28 +-- 2 files changed, 122 insertions(+), 104 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 5d1beb6..82270ee 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -10,6 +10,8 @@ open import Cubical.Data.Empty open import Cubical.Data.Unit open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients +open import Cubical.Categories.Category module Realizability.Topos.FunctionalRelation {ℓ ℓ' ℓ''} @@ -90,16 +92,12 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc yˢᵗ : Term strictnessContext `Y yˢᵗ = var y∈strictnessContext - -- F : Rel X Y, _~X_ : Rel X X, _~Y_ : Rel Y Y ⊢ F(x ,y) → (x ~X x) ∧ (y ~Y y) - strictnessPrequantFormula : Formula strictnessContext - strictnessPrequantFormula = rel `F (xˢᵗ `, yˢᵗ) `→ (rel `~X (xˢᵗ `, xˢᵗ) `∧ rel `~Y (yˢᵗ `, yˢᵗ)) - - strictnessFormula : Formula [] - strictnessFormula = `∀ (`∀ strictnessPrequantFormula) + -- F : Rel X Y, _~X_ : Rel X X, _~Y_ : Rel Y Y ⊢ F(x ,y) → (x ~X x) ∧ (y ~Y y) + strictnessFormula : Formula strictnessContext + strictnessFormula = rel `F (xˢᵗ `, yˢᵗ) `→ (rel `~X (xˢᵗ `, xˢᵗ) `∧ rel `~Y (yˢᵗ `, yˢᵗ)) field - strictnessWitness : A - isStrict : strictnessWitness ⊩ ⟦ strictnessFormula ⟧ᶠ + isStrict : holdsInTripos strictnessFormula -- # Relational -- The functional relation preserves equality @@ -126,15 +124,11 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc y₁ = var y₁∈relationalContext y₂ = var y₂∈relationalContext - relationalPrequantFormula : Formula relationalContext - relationalPrequantFormula = (rel `F (x₁ `, y₁) `∧ (rel `~X (x₁ `, x₂) `∧ rel `~Y (y₁ `, y₂))) `→ rel `F (x₂ `, y₂) - - relationalFormula : Formula [] - relationalFormula = `∀ (`∀ (`∀ (`∀ relationalPrequantFormula))) + relationalFormula : Formula relationalContext + relationalFormula = (rel `F (x₁ `, y₁) `∧ (rel `~X (x₁ `, x₂) `∧ rel `~Y (y₁ `, y₂))) `→ rel `F (x₂ `, y₂) field - relationalWitness : A - isRelational : relationalWitness ⊩ ⟦ relationalFormula ⟧ᶠ + isRelational : holdsInTripos relationalFormula -- # Single-valued -- Self explanatory @@ -156,16 +150,12 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc y₁ˢᵛ = var y₁∈singleValuedContext y₂ˢᵛ = var y₂∈singleValuedContext - singleValuedPrequantFormula : Formula singleValuedContext - singleValuedPrequantFormula = - (rel `F (xˢᵛ `, y₁ˢᵛ) `∧ rel `F (xˢᵛ `, y₂ˢᵛ)) `→ rel `~Y (y₁ˢᵛ `, y₂ˢᵛ) - - singleValuedFormula : Formula [] - singleValuedFormula = `∀ (`∀ (`∀ singleValuedPrequantFormula)) + singleValuedFormula : Formula singleValuedContext + singleValuedFormula = + (rel `F (xˢᵛ `, y₁ˢᵛ) `∧ rel `F (xˢᵛ `, y₂ˢᵛ)) `→ rel `~Y (y₁ˢᵛ `, y₂ˢᵛ) field - singleValuedWitness : A - isSingleValued : singleValuedWitness ⊩ ⟦ singleValuedFormula ⟧ᶠ + isSingleValued : holdsInTripos singleValuedFormula -- # Total -- For all existent elements in the domain x there is an element in the codomain y @@ -180,20 +170,16 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc xᵗˡ = var x∈totalContext - totalPrequantFormula : Formula totalContext - totalPrequantFormula = rel `~X (xᵗˡ `, xᵗˡ) `→ `∃ (rel `F (renamingTerm (drop id) xᵗˡ `, var here)) - - totalFormula : Formula [] - totalFormula = `∀ totalPrequantFormula + totalFormula : Formula totalContext + totalFormula = rel `~X (xᵗˡ `, xᵗˡ) `→ `∃ (rel `F (renamingTerm (drop id) xᵗˡ `, var here)) field - totalWitness : A - isTotal : totalWitness ⊩ ⟦ totalFormula ⟧ᶠ + isTotal : holdsInTripos totalFormula open FunctionalRelation hiding (`X; `Y) pointwiseEntailment : ∀ {X Y : Type ℓ'} → FunctionalRelation X Y → FunctionalRelation X Y → Type _ -pointwiseEntailment {X} {Y} F G = ∃[ a ∈ A ] (a ⊩ ⟦ entailmentFormula ⟧ᶠ) where +pointwiseEntailment {X} {Y} F G = holdsInTripos entailmentFormula where `X : Sort `Y : Sort @@ -215,6 +201,9 @@ pointwiseEntailment {X} {Y} F G = ∃[ a ∈ A ] (a ⊩ ⟦ entailmentFormula module RelationalInterpretation = Interpretation relationSymbols (λ { fzero → F .relation ; one → G .relation }) isNonTrivial open RelationalInterpretation + module RelationalSoundness = Soundness {relSym = relationSymbols} isNonTrivial (λ { fzero → F .relation ; one → G .relation }) + open RelationalSoundness + entailmentContext : Context entailmentContext = [] ′ `X ′ `Y @@ -227,19 +216,16 @@ pointwiseEntailment {X} {Y} F G = ∃[ a ∈ A ] (a ⊩ ⟦ entailmentFormula x = var x∈entailmentContext y = var y∈entailmentContext - entailmentPrequantFormula : Formula entailmentContext - entailmentPrequantFormula = rel `F (x `, y) `→ rel `G (x `, y) - - entailmentFormula : Formula [] - entailmentFormula = `∀ (`∀ entailmentPrequantFormula) - + entailmentFormula : Formula entailmentContext + entailmentFormula = rel `F (x `, y) `→ rel `G (x `, y) functionalRelationIsomorphism : ∀ {X Y : Type ℓ'} → FunctionalRelation X Y → FunctionalRelation X Y → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') functionalRelationIsomorphism {X} {Y} F G = pointwiseEntailment F G × pointwiseEntailment G F + pointwiseEntailment→functionalRelationIsomorphism : ∀ {X Y : Type ℓ'} → (F G : FunctionalRelation X Y) → pointwiseEntailment F G → functionalRelationIsomorphism F G -pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , {!!} where +pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , G≤F where `X : Sort `Y : Sort @@ -265,53 +251,97 @@ pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , open Interpretation relationSymbols (λ { fzero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_}) isNonTrivial open Soundness {relSym = relationSymbols} isNonTrivial ((λ { fzero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_})) open Relational relationSymbols + -- What we need to prove is that -- F ≤ G ⊨ G ≤ F -- We will use the semantic combinators we borrowed from the 1lab - -- We know that a ⊩ F ≤ G - -- or in other words - -- a ⊩ ∀ x. ∀ y. F(x , y) → G(x ,y) - - -- We will firstly use the introduction rule - -- for ∀ to get an x : X and a y : Y in context - proof : pointwiseEntailment G F - proof = - do - (a , a⊩F≤G) ← F≤G - let - context : Context - context = [] ′ `X ′ `Y - - x : Term context `X - x = var (there here) - - y : Term context `Y - y = var here - (b , b⊩) ← - `∀intro - {Γ = []} - {ϕ = ⊤ᵗ} - {B = `X} - {ψ = - `∀ {B = `Y} - (rel `G (x `, y) `→ rel `F (x `, y))} - (`∀intro - {Γ = [] ′ `X} - {ϕ = ⊤ᵗ} - {B = `Y} - {ψ = rel `G (x `, y) `→ rel `F (x `, y)} - (`→intro - {Γ = context} - {ϕ = ⊤ᵗ} - {ψ = rel `G (x `, y)} - {θ = rel `F (x `, y)} - (cut - {Γ = context} - {ϕ = ⊤ᵗ `∧ rel `G (x `, y)} - {ψ = rel `~X (x `, x)} - {θ = rel `F (x `, y)} - {!!} - {!!}))) - return (b , {!b⊩!}) - + entailmentContext : Context + entailmentContext = [] ′ `X ′ `Y + + x : Term entailmentContext `X + x = var (there here) + + y : Term entailmentContext `Y + y = var here + + G⊨x~x : (⊤ᵗ `∧ rel `G (x `, y)) ⊨ rel `~X (x `, x) + G⊨x~x = + `→elim + {Γ = entailmentContext} + {ϕ = ⊤ᵗ `∧ (rel `G (x `, y))} + {ψ = rel `G (x `, y)} + {θ = rel `~X (x `, x)} + {!G .isStrict!} + {!!} + + ⊤∧G⊨F : (⊤ᵗ `∧ rel `G (x `, y)) ⊨ rel `F (x `, y) + ⊤∧G⊨F = cut {Γ = entailmentContext} {ϕ = ⊤ᵗ `∧ rel `G (x `, y)} {ψ = rel `~X (x `, x)} + {θ = rel `F (x `, y)} + G⊨x~x + {!!} + + G≤F : pointwiseEntailment G F + G≤F = `→intro {Γ = entailmentContext} {ϕ = ⊤ᵗ} {ψ = rel `G (x `, y)} {θ = rel `F (x `, y)} ⊤∧G⊨F + +RTMorphism : (X Y : Type ℓ') → Type _ +RTMorphism X Y = FunctionalRelation X Y / λ F G → functionalRelationIsomorphism F G + +RTObject : Type _ +RTObject = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X + +idMorphism : (ob : RTObject) → RTMorphism (ob .fst) (ob .fst) +idMorphism ob = + [ record + { perX = ob .snd + ; perY = ob .snd + ; relation = ob .snd ._~_ + ; isStrict = {!!} + ; isRelational = {!!} + ; isSingleValued = {!!} + ; isTotal = {!!} + } ] where + + `X : Sort + `X = ↑ (ob .fst , ob .snd .isSetX) + + relationSymbols : Vec Sort 3 + relationSymbols = (`X `× `X) ∷ (`X `× `X) ∷ (`X `× `X) ∷ [] + + open Interpretation relationSymbols (λ { fzero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) isNonTrivial + open Soundness {relSym = relationSymbols} isNonTrivial (λ { fzero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) + open Relational relationSymbols + + isStrictContext : Context + isStrictContext = [] ′ `X ′ `X + + isStrictId : holdsInTripos (rel fzero (var (there here) `, var here) `→ (rel one (var (there here) `, var here) `∧ rel two (var (there here) `, var here))) + isStrictId = + `→intro + {Γ = isStrictContext} + {ϕ = ⊤ᵗ} + {ψ = rel fzero (var (there here) `, var here)} + {θ = rel one (var (there here) `, var here) `∧ rel two (var (there here) `, var here)} + (`∧intro + {Γ = isStrictContext} + {ϕ = ⊤ᵗ `∧ rel fzero (var (there here) `, var here)} + {ψ = rel one (var (there here) `, var here)} + {θ = rel two (var (there here) `, var here)} + {!`∧elim2 + {Γ = isStrictContext} + {ϕ = ⊤ᵗ} + {ψ = rel fzero (var (there here) `, var here)} + {θ = rel one (var (there here) `, var here)} + ?!} + {!!}) + + +RT : Category (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) ((ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ''))) +Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X +Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism X Y +Category.id RT {X , perX} = {!!} +Category._⋆_ RT = {!!} +Category.⋆IdL RT = {!!} +Category.⋆IdR RT = {!!} +Category.⋆Assoc RT = {!!} +Category.isSetHom RT = {!!} diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 2c80f27..ada3cfc 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -25,9 +25,7 @@ open Predicate' renaming (isSetX to isSetPredicateBase) open PredicateProperties open Morphism -private - _⊩_ : ∀ a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) → Type _ - a ⊩ P = a pre⊩ ∣ P ∣ tt* +Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field @@ -42,7 +40,7 @@ record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ- ~relSym : Vec Sort 1 ~relSym = `X×X ∷ [] - module X×XRelational = Relational ~relSym + module X×XRelational = Relational {n = 1} ~relSym open X×XRelational field @@ -76,16 +74,11 @@ record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ- yˢ : Term symContext `X yˢ = var y∈symContext - symPrequantFormula : Formula symContext - symPrequantFormula = rel fzero (xˢ `, yˢ) `→ rel fzero (yˢ `, xˢ) - - -- ~ ⊢ ∀ x. ∀ y. x ~ y → y ~ x - symFormula : Formula [] - symFormula = `∀ (`∀ symPrequantFormula) + symmetryFormula : Formula symContext + symmetryFormula = rel fzero (xˢ `, yˢ) `→ rel fzero (yˢ `, xˢ) field - symWitness : A - symIsWitnessed : symWitness ⊩ ⟦ symFormula ⟧ᶠ + symmetry : holdsInTripos symmetryFormula private transContext : Context @@ -109,14 +102,9 @@ record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ- xᵗ : Term transContext `X xᵗ = var x∈transContext - -- ~ : Rel X X, x : X, y : X, z : X ⊢ x ~ y ∧ y ~ z → x ~ z - transPrequantFormula : Formula transContext - transPrequantFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) - - transFormula : Formula [] - transFormula = `∀ (`∀ (`∀ transPrequantFormula)) + transitivityFormula : Formula transContext + transitivityFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) field - transWitness : A - transIsWitnessed : transWitness ⊩ ⟦ transFormula ⟧ᶠ + transitivity : holdsInTripos transitivityFormula From b268c8e723fbc4e000f04569239d7337963fb00c Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 8 Feb 2024 23:32:01 +0530 Subject: [PATCH 15/61] Finish weakening boilerplate --- src/Realizability/Tripos/Logic/RelContextStructural.agda | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/Realizability/Tripos/Logic/RelContextStructural.agda b/src/Realizability/Tripos/Logic/RelContextStructural.agda index 1d4afca..b520904 100644 --- a/src/Realizability/Tripos/Logic/RelContextStructural.agda +++ b/src/Realizability/Tripos/Logic/RelContextStructural.agda @@ -152,5 +152,7 @@ module WeakenSemantics transportPreservesHoldsInTripos : ∀ {Γ} → (f : ΞFormula Γ) → SoundnessΞ.holdsInTripos f → SoundnessΞR.holdsInTripos (transportAlongWeakening f) transportPreservesHoldsInTripos {Γ} f holds = - {!!} + do + (a , a⊩holds) ← holds + return (a , λ { γ b tt* → syntacticTransportPreservesRealizers⁺ γ (a ⨾ b) f (a⊩holds γ b tt*) }) From 452726a53e1f5cbd1fe8a2644a80e332c719b114 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Fri, 9 Feb 2024 00:26:46 +0530 Subject: [PATCH 16/61] Prove strictness of identity --- .../Topos/FunctionalRelation.agda | 98 +++++++++++++------ 1 file changed, 66 insertions(+), 32 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 82270ee..a4203c5 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -1,10 +1,13 @@ -open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) + open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels open import Cubical.Data.Vec open import Cubical.Data.Nat -open import Cubical.Data.FinData renaming (zero to fzero) +open import Cubical.Data.FinData +open import Cubical.Data.Fin hiding (Fin; _/_) open import Cubical.Data.Sigma open import Cubical.Data.Empty open import Cubical.Data.Unit @@ -37,6 +40,11 @@ private _⊩_ : ∀ a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) → Type _ a ⊩ P = a pre⊩ ∣ P ∣ tt* + +private λ*ComputationRule = `λ*ComputationRule as fefermanStructure +private λ* = `λ* as fefermanStructure + + open PartialEquivalenceRelation record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where @@ -60,7 +68,7 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc relationSymbol = (`X `× `Y) ∷ `X `× `X ∷ `Y `× `Y ∷ [] `F : Fin 3 - `F = fzero + `F = zero `~X : Fin 3 `~X = one `~Y : Fin 3 @@ -68,10 +76,10 @@ record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc open Relational relationSymbol - module RelationInterpretation' = Interpretation relationSymbol (λ { fzero → relation ; one → _~X_ ; two → _~Y_ }) isNonTrivial + module RelationInterpretation' = Interpretation relationSymbol (λ { zero → relation ; one → _~X_ ; two → _~Y_ }) isNonTrivial open RelationInterpretation' - module RelationSoundness = Soundness {relSym = relationSymbol} isNonTrivial (λ { fzero → relation ; one → _~X_ ; two → _~Y_ }) + module RelationSoundness = Soundness {relSym = relationSymbol} isNonTrivial (λ { zero → relation ; one → _~X_ ; two → _~Y_ }) open RelationSoundness -- # Strictness @@ -191,17 +199,17 @@ pointwiseEntailment {X} {Y} F G = holdsInTripos entailmentFormula where relationSymbols = (`X `× `Y) ∷ (`X `× `Y) ∷ [] `F : Fin 2 - `F = fzero + `F = zero `G : Fin 2 `G = one open Relational relationSymbols - module RelationalInterpretation = Interpretation relationSymbols (λ { fzero → F .relation ; one → G .relation }) isNonTrivial + module RelationalInterpretation = Interpretation relationSymbols (λ { zero → F .relation ; one → G .relation }) isNonTrivial open RelationalInterpretation - module RelationalSoundness = Soundness {relSym = relationSymbols} isNonTrivial (λ { fzero → F .relation ; one → G .relation }) + module RelationalSoundness = Soundness {relSym = relationSymbols} isNonTrivial (λ { zero → F .relation ; one → G .relation }) open RelationalSoundness entailmentContext : Context @@ -237,7 +245,7 @@ pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , relationSymbols = (`X `× `Y) ∷ (`X `× `Y) ∷ (`X `× `X) ∷ (`Y `× `Y) ∷ [] `F : Fin 4 - `F = fzero + `F = zero `G : Fin 4 `G = one @@ -248,8 +256,8 @@ pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , `~Y : Fin 4 `~Y = three - open Interpretation relationSymbols (λ { fzero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_}) isNonTrivial - open Soundness {relSym = relationSymbols} isNonTrivial ((λ { fzero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_})) + open Interpretation relationSymbols (λ { zero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_}) isNonTrivial + open Soundness {relSym = relationSymbols} isNonTrivial ((λ { zero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_})) open Relational relationSymbols -- What we need to prove is that @@ -308,33 +316,59 @@ idMorphism ob = relationSymbols : Vec Sort 3 relationSymbols = (`X `× `X) ∷ (`X `× `X) ∷ (`X `× `X) ∷ [] - open Interpretation relationSymbols (λ { fzero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) isNonTrivial - open Soundness {relSym = relationSymbols} isNonTrivial (λ { fzero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) + open Interpretation relationSymbols (λ { zero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) isNonTrivial + open Soundness {relSym = relationSymbols} isNonTrivial (λ { zero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) open Relational relationSymbols isStrictContext : Context isStrictContext = [] ′ `X ′ `X - isStrictId : holdsInTripos (rel fzero (var (there here) `, var here) `→ (rel one (var (there here) `, var here) `∧ rel two (var (there here) `, var here))) - isStrictId = - `→intro - {Γ = isStrictContext} - {ϕ = ⊤ᵗ} - {ψ = rel fzero (var (there here) `, var here)} - {θ = rel one (var (there here) `, var here) `∧ rel two (var (there here) `, var here)} - (`∧intro - {Γ = isStrictContext} - {ϕ = ⊤ᵗ `∧ rel fzero (var (there here) `, var here)} - {ψ = rel one (var (there here) `, var here)} - {θ = rel two (var (there here) `, var here)} - {!`∧elim2 - {Γ = isStrictContext} - {ϕ = ⊤ᵗ} - {ψ = rel fzero (var (there here) `, var here)} - {θ = rel one (var (there here) `, var here)} - ?!} - {!!}) + x : Term isStrictContext `X + x = var (there here) + + y : Term isStrictContext `X + y = var here + holdsInTriposIsStrict : holdsInTripos (rel zero (x `, y) `→ (rel one (x `, y) `∧ rel two (x `, y))) + holdsInTriposIsStrict = + do + let + prover : ApplStrTerm as 2 + prover = + ` pair ̇ # fone ̇ # fone + return + (λ* prover , + (λ { ((tt* , x') , y') a tt* b b⊩x'~y' + → + let + proofEq : λ* prover ⨾ a ⨾ b ≡ pair ⨾ b ⨾ b + proofEq = + λ*ComputationRule prover (a ∷ b ∷ []) + + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ a ⨾ b) ≡ b + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ a ⨾ b) + ≡⟨ cong (λ x → pr₁ ⨾ x) proofEq ⟩ + pr₁ ⨾ (pair ⨾ b ⨾ b) + ≡⟨ pr₁pxy≡x b b ⟩ + b + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ a ⨾ b) ≡ b + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ a ⨾ b) + ≡⟨ cong (λ x → pr₂ ⨾ x) proofEq ⟩ + pr₂ ⨾ (pair ⨾ b ⨾ b) + ≡⟨ pr₂pxy≡y b b ⟩ + b + ∎ + in + (subst + (λ r → r pre⊩ ∣ (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX)) (λ γ → snd (fst γ) , snd γ) (ob .snd ._~_)) ∣ ((tt* , x') , y')) + (sym pr₁proofEq) b⊩x'~y') , + subst + (λ r → r pre⊩ ∣ (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX)) (λ γ → snd (fst γ) , snd γ) (ob .snd ._~_)) ∣ ((tt* , x') , y')) (sym pr₂proofEq) b⊩x'~y' })) + RT : Category (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) ((ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ''))) Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X From 1ceb218c14a942afc3ce040b12dfc1bfa2d4473b Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Fri, 9 Feb 2024 00:30:14 +0530 Subject: [PATCH 17/61] Add documentation --- docs/Cubical.Data.Bool.Properties.html | 20 + docs/Cubical.Data.Fin.Properties.html | 26 +- docs/Cubical.Data.FinData.Properties.html | 10 +- docs/Cubical.Data.Nat.Base.html | 141 ++- docs/Cubical.Data.Nat.Order.html | 498 ++++---- docs/Cubical.Data.Nat.Properties.html | 18 +- ...bical.Foundations.Pointed.Homogeneous.html | 306 ++--- ...Ts.PropositionalTruncation.MagicTrick.html | 2 +- ...ealizability.Topos.FunctionalRelation.html | 470 +++++-- docs/Realizability.Topos.Object.html | 126 +- .../Realizability.Tripos.Logic.Semantics.html | 1118 +++++++++-------- docs/Realizability.Tripos.Logic.Syntax.html | 189 +-- .../Topos/FunctionalRelation.agda | 1 + 13 files changed, 1628 insertions(+), 1297 deletions(-) diff --git a/docs/Cubical.Data.Bool.Properties.html b/docs/Cubical.Data.Bool.Properties.html index 41d4b08..85e36f9 100644 --- a/docs/Cubical.Data.Bool.Properties.html +++ b/docs/Cubical.Data.Bool.Properties.html @@ -410,4 +410,24 @@ separatedBool : Separated Bool separatedBool = Discrete→Separated _≟_ + + +Bool→Bool→∙Bool : Bool (Bool , true) →∙ (Bool , true) +Bool→Bool→∙Bool false = idfun∙ _ +Bool→Bool→∙Bool true = const∙ _ _ + +Iso-Bool→∙Bool-Bool : Iso ((Bool , true) →∙ (Bool , true)) Bool +Iso.fun Iso-Bool→∙Bool-Bool f = fst f false +Iso.inv Iso-Bool→∙Bool-Bool = Bool→Bool→∙Bool +Iso.rightInv Iso-Bool→∙Bool-Bool false = refl +Iso.rightInv Iso-Bool→∙Bool-Bool true = refl +Iso.leftInv Iso-Bool→∙Bool-Bool f = Σ≡Prop _ isSetBool _ _) (help _ refl) + where + help : (x : Bool) fst f false x + Bool→Bool→∙Bool (fst f false) .fst f .fst + help false p = funExt + λ { false j Bool→Bool→∙Bool (p j) .fst false) sym p + ; true j Bool→Bool→∙Bool (p j) .fst true) sym (snd f)} + help true p = j Bool→Bool→∙Bool (p j) .fst) + funExt λ { false sym p ; true sym (snd f)} \ No newline at end of file diff --git a/docs/Cubical.Data.Fin.Properties.html b/docs/Cubical.Data.Fin.Properties.html index a53ed38..def4b12 100644 --- a/docs/Cubical.Data.Fin.Properties.html +++ b/docs/Cubical.Data.Fin.Properties.html @@ -70,7 +70,7 @@ ... | no prf = no λ h prf (cong fst h) inject<-ne : {n} (i : Fin n) ¬ inject< ≤-refl i (n , ≤-refl) -inject<-ne {n} (k , k<n) p = <→≢ k<n (cong fst p) +inject<-ne {n} (k , k<n) p = <→≢ k<n (cong fst p) Fin-fst-≡ : {n} {i j : Fin n} fst i fst j i j Fin-fst-≡ = Σ≡Prop _ isProp≤) @@ -203,12 +203,12 @@ step n = transport (Residue≡ k n) reduce : n Residue k n - reduce = +induction k₀ (Residue k) base step + reduce = +induction k₀ (Residue k) base step reduce≡ : n transport (Residue≡ k n) (reduce n) reduce (k + n) reduce≡ n - = sym (+inductionStep k₀ _ base step n) + = sym (+inductionStep k₀ _ base step n) reduceP : n PathP i Residue≡ k n i) (reduce n) (reduce (k + n)) @@ -467,7 +467,7 @@ ≤-helper : m n ≤-helper = ≤-·sk-cancel (pred-≤-pred (<≤-trans p (≤-suc ≤-refl))) goal : m < n - goal = case <-split (suc-≤-suc ≤-helper) of λ + goal = case <-split (suc-≤-suc ≤-helper) of λ { (inl g) g ; (inr e) Empty.rec (¬m<m (subst m m · suc k < n · suc k) e p)) } @@ -481,11 +481,11 @@ nm = expand× (nn , toℕ mm) nm<n·m : toℕ mm · suc n + toℕ nn < suc n · m nm<n·m = - toℕ mm · suc n + toℕ nn <≤⟨ <-k+ (snd nn) - toℕ mm · suc n + suc n ≡≤⟨ +-comm _ (suc n) - suc (toℕ mm) · suc n ≤≡⟨ ≤-·k (snd mm) + toℕ mm · suc n + toℕ nn <≤⟨ <-k+ (snd nn) + toℕ mm · suc n + suc n ≡≤⟨ +-comm _ (suc n) + suc (toℕ mm) · suc n ≤≡⟨ ≤-·k (snd mm) m · suc n ≡⟨ ·-comm _ (suc n) - suc n · m where open <-Reasoning + suc n · m where open <-Reasoning intro-injective : {o} {p} intro o intro p o p intro-injective {o} {p} io≡ip = λ i io′≡ip′ i .fst , toℕ-injective {fj = snd o} {fk = snd p} (cong snd io′≡ip′) i where @@ -508,10 +508,10 @@ nmsnd = subst l l < suc n · m) (sym nmmoddiv) (snd nm) mm·sn<m·sn : mm · suc n < m · suc n mm·sn<m·sn = - mm · suc n ≤<⟨ nn , +-comm nn (mm · suc n) - mm · suc n + nn <≡⟨ nmsnd + mm · suc n ≤<⟨ nn , +-comm nn (mm · suc n) + mm · suc n + nn <≡⟨ nmsnd suc n · m ≡⟨ ·-comm (suc n) m - m · suc n where open <-Reasoning + m · suc n where open <-Reasoning mm<m : mm < m mm<m = <-·sk-cancel mm·sn<m·sn @@ -678,7 +678,7 @@ Fin≤1→isProp (suc (suc n)) p = Empty.rec (¬-<-zero (pred-≤-pred p)) isProp→Fin≤1 : (n : ) isProp (Fin n) n 1 -isProp→Fin≤1 0 _ = ≤-solver 0 1 -isProp→Fin≤1 1 _ = ≤-solver 1 1 +isProp→Fin≤1 0 _ = ≤-solver 0 1 +isProp→Fin≤1 1 _ = ≤-solver 1 1 isProp→Fin≤1 (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) \ No newline at end of file diff --git a/docs/Cubical.Data.FinData.Properties.html b/docs/Cubical.Data.FinData.Properties.html index b92439d..aad12b8 100644 --- a/docs/Cubical.Data.FinData.Properties.html +++ b/docs/Cubical.Data.FinData.Properties.html @@ -157,13 +157,13 @@ toFin : {n : } (m : ) m < n Fin n toFin {n = ℕzero} _ m<0 = ⊥.rec (¬-<-zero m<0) toFin {n = ℕsuc n} _ (ℕzero , _) = fromℕ n --in this case we have m≡n -toFin {n = ℕsuc n} m (ℕsuc k , p) = weakenFin (toFin m (k , cong predℕ p)) +toFin {n = ℕsuc n} m (ℕsuc k , p) = weakenFin (toFin m (k , cong predℕ p)) toFin0≡0 : {n : } (p : 0 < ℕsuc n) toFin 0 p zero -toFin0≡0 (ℕzero , p) = subst x fromℕ x zero) (cong predℕ p) refl -toFin0≡0 {ℕzero} (ℕsuc k , p) = ⊥.rec (ℕsnotz (+-comm 1 k (cong predℕ p))) +toFin0≡0 (ℕzero , p) = subst x fromℕ x zero) (cong predℕ p) refl +toFin0≡0 {ℕzero} (ℕsuc k , p) = ⊥.rec (ℕsnotz (+-comm 1 k (cong predℕ p))) toFin0≡0 {ℕsuc n} (ℕsuc k , p) = - subst x weakenFin x zero) (sym (toFin0≡0 (k , cong predℕ p))) refl + subst x weakenFin x zero) (sym (toFin0≡0 (k , cong predℕ p))) refl genδ-FinVec : (n k : ) (a b : A) FinVec A n genδ-FinVec (ℕsuc n) ℕzero a b zero = a @@ -208,7 +208,7 @@ P (enum (ℕsuc k) p) ((m : )(q : m < n)(q' : m ℕsuc k) P (enum m q)) enumIndStep P k p f x m q q' = - case (≤-split q') return _ P (enum m q)) of + case (≤-split q') return _ P (enum m q)) of λ { (inl r') f m q (pred-≤-pred r') ; (inr r') subst P (enumExt p q (sym r')) x } diff --git a/docs/Cubical.Data.Nat.Base.html b/docs/Cubical.Data.Nat.Base.html index b06be8d..b2ad75f 100644 --- a/docs/Cubical.Data.Nat.Base.html +++ b/docs/Cubical.Data.Nat.Base.html @@ -13,69 +13,80 @@ open import Cubical.Data.Sum.Base hiding (elim) open import Cubical.Data.Empty.Base hiding (elim) open import Cubical.Data.Unit.Base - -predℕ : -predℕ zero = zero -predℕ (suc n) = n - -caseNat : {} {A : Type } (a0 aS : A) A -caseNat a0 aS zero = a0 -caseNat a0 aS (suc n) = aS - -doubleℕ : -doubleℕ zero = zero -doubleℕ (suc x) = suc (suc (doubleℕ x)) - --- doublesℕ n m = 2^n · m -doublesℕ : -doublesℕ zero m = m -doublesℕ (suc n) m = doublesℕ n (doubleℕ m) - --- iterate -iter : {} {A : Type } (A A) A A -iter zero f z = z -iter (suc n) f z = f (iter n f z) - -elim : {} {A : Type } - A zero - ((n : ) A n A (suc n)) - (n : ) A n -elim a₀ _ zero = a₀ -elim a₀ f (suc n) = f n (elim a₀ f n) - -elim+2 : {} {A : Type } A 0 A 1 - ((n : ) (A (suc n) A (suc (suc n)))) - (n : ) A n -elim+2 a0 a1 ind zero = a0 -elim+2 a0 a1 ind (suc zero) = a1 -elim+2 {A = A} a0 a1 ind (suc (suc n)) = - ind n (elim+2 {A = A} a0 a1 ind (suc n)) - -isEven isOdd : Bool -isEven zero = true -isEven (suc n) = isOdd n -isOdd zero = false -isOdd (suc n) = isEven n - ---Typed version -private - toType : Bool Type - toType false = - toType true = Unit - -isEvenT : Type -isEvenT n = toType (isEven n) - -isOddT : Type -isOddT n = isEvenT (suc n) - -isZero : Bool -isZero zero = true -isZero (suc n) = false - --- exponential - -_^_ : -m ^ 0 = 1 -m ^ (suc n) = m · m ^ n +open import Cubical.Data.Sigma.Base + +predℕ : +predℕ zero = zero +predℕ (suc n) = n + +caseNat : {} {A : Type } (a0 aS : A) A +caseNat a0 aS zero = a0 +caseNat a0 aS (suc n) = aS + +doubleℕ : +doubleℕ zero = zero +doubleℕ (suc x) = suc (suc (doubleℕ x)) + +-- doublesℕ n m = 2^n · m +doublesℕ : +doublesℕ zero m = m +doublesℕ (suc n) m = doublesℕ n (doubleℕ m) + +-- iterate +iter : {} {A : Type } (A A) A A +iter zero f z = z +iter (suc n) f z = f (iter n f z) + +elim : {} {A : Type } + A zero + ((n : ) A n A (suc n)) + (n : ) A n +elim a₀ _ zero = a₀ +elim a₀ f (suc n) = f n (elim a₀ f n) + +elim+2 : {} {A : Type } A 0 A 1 + ((n : ) (A (suc n) A (suc (suc n)))) + (n : ) A n +elim+2 a0 a1 ind zero = a0 +elim+2 a0 a1 ind (suc zero) = a1 +elim+2 {A = A} a0 a1 ind (suc (suc n)) = + ind n (elim+2 {A = A} a0 a1 ind (suc n)) + +isEven isOdd : Bool +isEven zero = true +isEven (suc n) = isOdd n +isOdd zero = false +isOdd (suc n) = isEven n + +--Typed version +private + toType : Bool Type + toType false = + toType true = Unit + +isEvenT : Type +isEvenT n = toType (isEven n) + +isOddT : Type +isOddT n = isEvenT (suc n) + +isZero : Bool +isZero zero = true +isZero (suc n) = false + +-- exponential + +_^_ : +m ^ 0 = 1 +m ^ (suc n) = m · m ^ n + + +-- Iterated product +_ˣ_ : {} (A : Type ) (n : ) Type +A ˣ zero = A zero +A ˣ suc n = (A ˣ n) × A (suc n) + + : {} (A : Type ) (0A : (n : ) A n) (n : ) A ˣ n + A 0A zero = 0A zero + A 0A (suc n) = ( A 0A n) , (0A (suc n)) \ No newline at end of file diff --git a/docs/Cubical.Data.Nat.Order.html b/docs/Cubical.Data.Nat.Order.html index ea393bd..6055c47 100644 --- a/docs/Cubical.Data.Nat.Order.html +++ b/docs/Cubical.Data.Nat.Order.html @@ -93,7 +93,7 @@ ≤-sucℕ : n suc n ≤-sucℕ = ≤-suc ≤-refl -≤-predℕ : predℕ n n +≤-predℕ : predℕ n n ≤-predℕ {zero} = ≤-refl ≤-predℕ {suc n} = ≤-suc ≤-refl @@ -156,7 +156,7 @@ ≤0→≡0 {zero} ineq = refl ≤0→≡0 {suc n} ineq = ⊥.rec (¬-<-zero ineq) -predℕ-≤-predℕ : m n (predℕ m) (predℕ n) +predℕ-≤-predℕ : m n (predℕ m) (predℕ n) predℕ-≤-predℕ {zero} {zero} ineq = ≤-refl predℕ-≤-predℕ {zero} {suc n} ineq = zero-≤ predℕ-≤-predℕ {suc m} {zero} ineq = ⊥.rec (¬-<-zero ineq) @@ -273,253 +273,259 @@ suc m zero = gt (m , +-comm m 1) suc m suc n = Trichotomy-suc (m n) -splitℕ-≤ : (m n : ) (m n) (n < m) -splitℕ-≤ m n with m n -... | lt x = inl (<-weaken x) -... | eq x = inl (0 , x) -... | gt x = inr x - -splitℕ-< : (m n : ) (m < n) (n m) -splitℕ-< m n with m n -... | lt x = inl x -... | eq x = inr (0 , (sym x)) -... | gt x = inr (<-weaken x) - -≤CaseInduction : {P : Type } {n m : } - (n m P n m) (m n P n m) - P n m -≤CaseInduction {n = n} {m = m} p q with n m -... | lt x = p (<-weaken x) -... | eq x = p (subst (n ≤_) x ≤-refl) -... | gt x = q (<-weaken x) - -<-split : m < suc n (m < n) (m n) -<-split {n = zero} = inr snd m+n≡0→m≡0×n≡0 snd pred-≤-pred -<-split {zero} {suc n} = λ _ inl (suc-≤-suc zero-≤) -<-split {suc m} {suc n} = ⊎.map suc-≤-suc (cong suc) <-split pred-≤-pred - -≤-split : m n (m < n) (m n) -≤-split p = <-split (suc-≤-suc p) - -≤→< : m n ¬ m n m < n -≤→< p q = - case (≤-split p) of - λ { (inl r) r - ; (inr r) ⊥.rec (q r) } - -≤-suc-≢ : m suc n (m suc n ) m n -≤-suc-≢ p ¬q = pred-≤-pred (≤→< p ¬q) - -≤-+-split : n m k k n + m (n k) (m (n + m) k) -≤-+-split n m k k≤n+m with n k -... | eq p = inl (0 , p) -... | lt n<k = inl (<-weaken n<k) -... | gt k<n with m ((n + m) k) -... | eq p = inr (0 , p) -... | lt m<n+m∸k = inr (<-weaken m<n+m∸k) -... | gt n+m∸k<m = - ⊥.rec (¬m<m (transport i ≤-∸-+-cancel k≤n+m i < +-comm m n i) (<-+-< n+m∸k<m k<n))) - -<-asym'-case : Trichotomy m n ¬ m < n n m -<-asym'-case (lt p) q = ⊥.rec (q p) -<-asym'-case (eq p) _ = _ , sym p -<-asym'-case (gt p) _ = <-weaken p - -<-asym' : ¬ m < n n m -<-asym' = <-asym'-case (_≟_ _ _) - -private - acc-suc : Acc _<_ n Acc _<_ (suc n) - acc-suc a - = acc λ y y<sn - case <-split y<sn of λ - { (inl y<n) access a y y<n - ; (inr y≡n) subst _ (sym y≡n) a - } - -<-wellfounded : WellFounded _<_ -<-wellfounded zero = acc λ _ ⊥.rec ¬-<-zero -<-wellfounded (suc n) = acc-suc (<-wellfounded n) - -<→≢ : n < m ¬ n m -<→≢ {n} {m} p q = ¬m<m (subst (_< m) q p) - -module _ - (b₀ : ) - (P : Type₀) - (base : n n < suc b₀ P n) - (step : n P n P (suc b₀ + n)) - where - open WFI (<-wellfounded) - - private - dichotomy : b n (n < b) (Σ[ m ] n b + m) - dichotomy b n - = case n b return _ (n < b) (Σ[ m ] n b + m)) of λ - { (lt o) inl o - ; (eq p) inr (0 , p sym (+-zero b)) - ; (gt (m , p)) inr (suc m , sym p +-suc m b +-comm (suc m) b) - } - - dichotomy<≡ : b n (n<b : n < b) dichotomy b n inl n<b - dichotomy<≡ b n n<b - = case dichotomy b n return d d inl n<b) of λ - { (inl x) cong inl (isProp≤ x n<b) - ; (inr (m , p)) ⊥.rec (<-asym n<b (m , sym (p +-comm b m))) - } - - dichotomy+≡ : b m n (p : n b + m) dichotomy b n inr (m , p) - dichotomy+≡ b m n p - = case dichotomy b n return d d inr (m , p)) of λ - { (inl n<b) ⊥.rec (<-asym n<b (m , +-comm m b sym p)) - ; (inr (m' , q)) - cong inr (Σ≡Prop x isSetℕ n (b + x)) (inj-m+ {m = b} (sym q p))) - } - - b = suc b₀ - - lemma₁ : ∀{x y z} x suc z + y y < x - lemma₁ {y = y} {z} p = z , +-suc z y sym p - - subStep : (n : ) (∀ m m < n P m) (n < b) (Σ[ m ] n b + m) P n - subStep n _ (inl l) = base n l - subStep n rec (inr (m , p)) - = transport (cong P (sym p)) (step m (rec m (lemma₁ p))) - - wfStep : (n : ) (∀ m m < n P m) P n - wfStep n rec = subStep n rec (dichotomy b n) - - wfStepLemma₀ : n (n<b : n < b) rec wfStep n rec base n n<b - wfStepLemma₀ n n<b rec = cong (subStep n rec) (dichotomy<≡ b n n<b) - - wfStepLemma₁ : n rec wfStep (b + n) rec step n (rec n (lemma₁ refl)) - wfStepLemma₁ n rec - = cong (subStep (b + n) rec) (dichotomy+≡ b n (b + n) refl) - transportRefl _ - - +induction : n P n - +induction = induction wfStep - - +inductionBase : n (l : n < b) +induction n base n l - +inductionBase n l = induction-compute wfStep n wfStepLemma₀ n l _ - - +inductionStep : n +induction (b + n) step n (+induction n) - +inductionStep n = induction-compute wfStep (b + n) wfStepLemma₁ n _ - -module <-Reasoning where - -- TODO: would it be better to mirror the way it is done in the agda-stdlib? - infixr 2 _<⟨_⟩_ _≤<⟨_⟩_ _≤⟨_⟩_ _<≤⟨_⟩_ _≡<⟨_⟩_ _≡≤⟨_⟩_ _<≡⟨_⟩_ _≤≡⟨_⟩_ - _<⟨_⟩_ : k k < n n < m k < m - _ <⟨ p q = <-trans p q - - _≤<⟨_⟩_ : k k n n < m k < m - _ ≤<⟨ p q = ≤<-trans p q - - _≤⟨_⟩_ : k k n n m k m - _ ≤⟨ p q = ≤-trans p q - - _<≤⟨_⟩_ : k k < n n m k < m - _ <≤⟨ p q = <≤-trans p q - - _≡≤⟨_⟩_ : k k l l m k m - _ ≡≤⟨ p q = subst k k _) (sym p) q - - _≡<⟨_⟩_ : k k l l < m k < m - _ ≡<⟨ p q = _ ≡≤⟨ cong suc p q - - _≤≡⟨_⟩_ : k k l l m k m - _ ≤≡⟨ p q = subst l _ l) q p - - _<≡⟨_⟩_ : k k < l l m k < m - _ <≡⟨ p q = _ ≤≡⟨ p q - - --- Some lemmas about ∸ -suc∸-fst : (n m : ) m < n suc (n m) (suc n) m -suc∸-fst zero zero p = refl -suc∸-fst zero (suc m) p = ⊥.rec (¬-<-zero p) -suc∸-fst (suc n) zero p = refl -suc∸-fst (suc n) (suc m) p = (suc∸-fst n m (pred-≤-pred p)) - -n∸m≡0 : (n m : ) n m (n m) 0 -n∸m≡0 zero zero p = refl -n∸m≡0 (suc n) zero p = ⊥.rec (¬-<-zero p) -n∸m≡0 zero (suc m) p = refl -n∸m≡0 (suc n) (suc m) p = n∸m≡0 n m (pred-≤-pred p) - -n∸n≡0 : (n : ) n n 0 -n∸n≡0 zero = refl -n∸n≡0 (suc n) = n∸n≡0 n - -n∸l>0 : (n l : ) (l < n) 0 < (n l) -n∸l>0 zero zero r = ⊥.rec (¬-<-zero r) -n∸l>0 zero (suc l) r = ⊥.rec (¬-<-zero r) -n∸l>0 (suc n) zero r = suc-≤-suc zero-≤ -n∸l>0 (suc n) (suc l) r = n∸l>0 n l (pred-≤-pred r) - --- automation - -≤-solver-type : (m n : ) Trichotomy m n Type -≤-solver-type m n (lt p) = m n -≤-solver-type m n (eq p) = m n -≤-solver-type m n (gt p) = n < m - -≤-solver-case : (m n : ) (p : Trichotomy m n) ≤-solver-type m n p -≤-solver-case m n (lt p) = <-weaken p -≤-solver-case m n (eq p) = _ , p -≤-solver-case m n (gt p) = p - -≤-solver : (m n : ) ≤-solver-type m n (m n) -≤-solver m n = ≤-solver-case m n (m n) +Dichotomyℕ : (n m : ) (n m) (n > m) +Dichotomyℕ n m with (n m) +... | lt x = inl (<-weaken x) +... | Trichotomy.eq x = inl (0 , x) +... | gt x = inr x + +splitℕ-≤ : (m n : ) (m n) (n < m) +splitℕ-≤ m n with m n +... | lt x = inl (<-weaken x) +... | eq x = inl (0 , x) +... | gt x = inr x + +splitℕ-< : (m n : ) (m < n) (n m) +splitℕ-< m n with m n +... | lt x = inl x +... | eq x = inr (0 , (sym x)) +... | gt x = inr (<-weaken x) + +≤CaseInduction : {P : Type } {n m : } + (n m P n m) (m n P n m) + P n m +≤CaseInduction {n = n} {m = m} p q with n m +... | lt x = p (<-weaken x) +... | eq x = p (subst (n ≤_) x ≤-refl) +... | gt x = q (<-weaken x) + +<-split : m < suc n (m < n) (m n) +<-split {n = zero} = inr snd m+n≡0→m≡0×n≡0 snd pred-≤-pred +<-split {zero} {suc n} = λ _ inl (suc-≤-suc zero-≤) +<-split {suc m} {suc n} = ⊎.map suc-≤-suc (cong suc) <-split pred-≤-pred + +≤-split : m n (m < n) (m n) +≤-split p = <-split (suc-≤-suc p) + +≤→< : m n ¬ m n m < n +≤→< p q = + case (≤-split p) of + λ { (inl r) r + ; (inr r) ⊥.rec (q r) } + +≤-suc-≢ : m suc n (m suc n ) m n +≤-suc-≢ p ¬q = pred-≤-pred (≤→< p ¬q) + +≤-+-split : n m k k n + m (n k) (m (n + m) k) +≤-+-split n m k k≤n+m with n k +... | eq p = inl (0 , p) +... | lt n<k = inl (<-weaken n<k) +... | gt k<n with m ((n + m) k) +... | eq p = inr (0 , p) +... | lt m<n+m∸k = inr (<-weaken m<n+m∸k) +... | gt n+m∸k<m = + ⊥.rec (¬m<m (transport i ≤-∸-+-cancel k≤n+m i < +-comm m n i) (<-+-< n+m∸k<m k<n))) + +<-asym'-case : Trichotomy m n ¬ m < n n m +<-asym'-case (lt p) q = ⊥.rec (q p) +<-asym'-case (eq p) _ = _ , sym p +<-asym'-case (gt p) _ = <-weaken p + +<-asym' : ¬ m < n n m +<-asym' = <-asym'-case (_≟_ _ _) + +private + acc-suc : Acc _<_ n Acc _<_ (suc n) + acc-suc a + = acc λ y y<sn + case <-split y<sn of λ + { (inl y<n) access a y y<n + ; (inr y≡n) subst _ (sym y≡n) a + } + +<-wellfounded : WellFounded _<_ +<-wellfounded zero = acc λ _ ⊥.rec ¬-<-zero +<-wellfounded (suc n) = acc-suc (<-wellfounded n) + +<→≢ : n < m ¬ n m +<→≢ {n} {m} p q = ¬m<m (subst (_< m) q p) + +module _ + (b₀ : ) + (P : Type₀) + (base : n n < suc b₀ P n) + (step : n P n P (suc b₀ + n)) + where + open WFI (<-wellfounded) + + private + dichotomy : b n (n < b) (Σ[ m ] n b + m) + dichotomy b n + = case n b return _ (n < b) (Σ[ m ] n b + m)) of λ + { (lt o) inl o + ; (eq p) inr (0 , p sym (+-zero b)) + ; (gt (m , p)) inr (suc m , sym p +-suc m b +-comm (suc m) b) + } + + dichotomy<≡ : b n (n<b : n < b) dichotomy b n inl n<b + dichotomy<≡ b n n<b + = case dichotomy b n return d d inl n<b) of λ + { (inl x) cong inl (isProp≤ x n<b) + ; (inr (m , p)) ⊥.rec (<-asym n<b (m , sym (p +-comm b m))) + } + + dichotomy+≡ : b m n (p : n b + m) dichotomy b n inr (m , p) + dichotomy+≡ b m n p + = case dichotomy b n return d d inr (m , p)) of λ + { (inl n<b) ⊥.rec (<-asym n<b (m , +-comm m b sym p)) + ; (inr (m' , q)) + cong inr (Σ≡Prop x isSetℕ n (b + x)) (inj-m+ {m = b} (sym q p))) + } + + b = suc b₀ + + lemma₁ : ∀{x y z} x suc z + y y < x + lemma₁ {y = y} {z} p = z , +-suc z y sym p + + subStep : (n : ) (∀ m m < n P m) (n < b) (Σ[ m ] n b + m) P n + subStep n _ (inl l) = base n l + subStep n rec (inr (m , p)) + = transport (cong P (sym p)) (step m (rec m (lemma₁ p))) + + wfStep : (n : ) (∀ m m < n P m) P n + wfStep n rec = subStep n rec (dichotomy b n) + + wfStepLemma₀ : n (n<b : n < b) rec wfStep n rec base n n<b + wfStepLemma₀ n n<b rec = cong (subStep n rec) (dichotomy<≡ b n n<b) + + wfStepLemma₁ : n rec wfStep (b + n) rec step n (rec n (lemma₁ refl)) + wfStepLemma₁ n rec + = cong (subStep (b + n) rec) (dichotomy+≡ b n (b + n) refl) + transportRefl _ + + +induction : n P n + +induction = induction wfStep + + +inductionBase : n (l : n < b) +induction n base n l + +inductionBase n l = induction-compute wfStep n wfStepLemma₀ n l _ + + +inductionStep : n +induction (b + n) step n (+induction n) + +inductionStep n = induction-compute wfStep (b + n) wfStepLemma₁ n _ + +module <-Reasoning where + -- TODO: would it be better to mirror the way it is done in the agda-stdlib? + infixr 2 _<⟨_⟩_ _≤<⟨_⟩_ _≤⟨_⟩_ _<≤⟨_⟩_ _≡<⟨_⟩_ _≡≤⟨_⟩_ _<≡⟨_⟩_ _≤≡⟨_⟩_ + _<⟨_⟩_ : k k < n n < m k < m + _ <⟨ p q = <-trans p q + + _≤<⟨_⟩_ : k k n n < m k < m + _ ≤<⟨ p q = ≤<-trans p q + + _≤⟨_⟩_ : k k n n m k m + _ ≤⟨ p q = ≤-trans p q + + _<≤⟨_⟩_ : k k < n n m k < m + _ <≤⟨ p q = <≤-trans p q + + _≡≤⟨_⟩_ : k k l l m k m + _ ≡≤⟨ p q = subst k k _) (sym p) q + + _≡<⟨_⟩_ : k k l l < m k < m + _ ≡<⟨ p q = _ ≡≤⟨ cong suc p q + + _≤≡⟨_⟩_ : k k l l m k m + _ ≤≡⟨ p q = subst l _ l) q p + + _<≡⟨_⟩_ : k k < l l m k < m + _ <≡⟨ p q = _ ≤≡⟨ p q + + +-- Some lemmas about ∸ +suc∸-fst : (n m : ) m < n suc (n m) (suc n) m +suc∸-fst zero zero p = refl +suc∸-fst zero (suc m) p = ⊥.rec (¬-<-zero p) +suc∸-fst (suc n) zero p = refl +suc∸-fst (suc n) (suc m) p = (suc∸-fst n m (pred-≤-pred p)) + +n∸m≡0 : (n m : ) n m (n m) 0 +n∸m≡0 zero zero p = refl +n∸m≡0 (suc n) zero p = ⊥.rec (¬-<-zero p) +n∸m≡0 zero (suc m) p = refl +n∸m≡0 (suc n) (suc m) p = n∸m≡0 n m (pred-≤-pred p) + +n∸n≡0 : (n : ) n n 0 +n∸n≡0 zero = refl +n∸n≡0 (suc n) = n∸n≡0 n + +n∸l>0 : (n l : ) (l < n) 0 < (n l) +n∸l>0 zero zero r = ⊥.rec (¬-<-zero r) +n∸l>0 zero (suc l) r = ⊥.rec (¬-<-zero r) +n∸l>0 (suc n) zero r = suc-≤-suc zero-≤ +n∸l>0 (suc n) (suc l) r = n∸l>0 n l (pred-≤-pred r) + +-- automation + +≤-solver-type : (m n : ) Trichotomy m n Type +≤-solver-type m n (lt p) = m n +≤-solver-type m n (eq p) = m n +≤-solver-type m n (gt p) = n < m + +≤-solver-case : (m n : ) (p : Trichotomy m n) ≤-solver-type m n p +≤-solver-case m n (lt p) = <-weaken p +≤-solver-case m n (eq p) = _ , p +≤-solver-case m n (gt p) = p + +≤-solver : (m n : ) ≤-solver-type m n (m n) +≤-solver m n = ≤-solver-case m n (m n) --- inductive order relation taken from agda-stdlib -data _≤'_ : Type where - z≤ : {n} zero ≤' n - s≤s : {m n} m ≤' n suc m ≤' suc n - -_<'_ : Type -m <' n = suc m ≤' n - --- Smart constructors of _<_ -pattern z<s {n} = s≤s (z≤ {n}) -pattern s<s {m} {n} m<n = s≤s {m} {n} m<n +-- inductive order relation taken from agda-stdlib +data _≤'_ : Type where + z≤ : {n} zero ≤' n + s≤s : {m n} m ≤' n suc m ≤' suc n + +_<'_ : Type +m <' n = suc m ≤' n + +-- Smart constructors of _<_ +pattern z<s {n} = s≤s (z≤ {n}) +pattern s<s {m} {n} m<n = s≤s {m} {n} m<n -¬-<'-zero : ¬ m <' 0 -¬-<'-zero {zero} () -¬-<'-zero {suc m} () - -≤'Dec : m n Dec (m ≤' n) -≤'Dec zero n = yes z≤ -≤'Dec (suc m) zero = no ¬-<'-zero -≤'Dec (suc m) (suc n) with ≤'Dec m n -... | yes m≤'n = yes (s≤s m≤'n) -... | no m≰'n = no λ { (s≤s m≤'n) m≰'n m≤'n } +¬-<'-zero : ¬ m <' 0 +¬-<'-zero {zero} () +¬-<'-zero {suc m} () + +≤'Dec : m n Dec (m ≤' n) +≤'Dec zero n = yes z≤ +≤'Dec (suc m) zero = no ¬-<'-zero +≤'Dec (suc m) (suc n) with ≤'Dec m n +... | yes m≤'n = yes (s≤s m≤'n) +... | no m≰'n = no λ { (s≤s m≤'n) m≰'n m≤'n } -≤'IsPropValued : m n isProp (m ≤' n) -≤'IsPropValued zero n z≤ z≤ = refl -≤'IsPropValued (suc m) zero () -≤'IsPropValued (suc m) (suc n) (s≤s x) (s≤s y) = cong s≤s (≤'IsPropValued m n x y) - -≤-∸-≤ : m n l m n m l n l -≤-∸-≤ m n zero r = r -≤-∸-≤ zero zero (suc l) r = ≤-refl -≤-∸-≤ zero (suc n) (suc l) r = (n l) , (+-zero _) -≤-∸-≤ (suc m) zero (suc l) r = ⊥.rec (¬-<-zero r) -≤-∸-≤ (suc m) (suc n) (suc l) r = ≤-∸-≤ m n l (pred-≤-pred r) - -<-∸-< : m n l m < n l < n m l < n l -<-∸-< m n zero r q = r -<-∸-< zero zero (suc l) r q = ⊥.rec (¬-<-zero r) -<-∸-< zero (suc n) (suc l) r q = n∸l>0 (suc n) (suc l) q -<-∸-< (suc m) zero (suc l) r q = ⊥.rec (¬-<-zero r) -<-∸-< (suc m) (suc n) (suc l) r q = <-∸-< m n l (pred-≤-pred r) (pred-≤-pred q) - -≤-∸-≥ : n l k l k n k n l -≤-∸-≥ n zero zero r = ≤-refl -≤-∸-≥ n zero (suc k) r = ∸-≤ n (suc k) -≤-∸-≥ n (suc l) zero r = ⊥.rec (¬-<-zero r) -≤-∸-≥ zero (suc l) (suc k) r = ≤-refl -≤-∸-≥ (suc n) (suc l) (suc k) r = ≤-∸-≥ n l k (pred-≤-pred r) +≤'IsPropValued : m n isProp (m ≤' n) +≤'IsPropValued zero n z≤ z≤ = refl +≤'IsPropValued (suc m) zero () +≤'IsPropValued (suc m) (suc n) (s≤s x) (s≤s y) = cong s≤s (≤'IsPropValued m n x y) + +≤-∸-≤ : m n l m n m l n l +≤-∸-≤ m n zero r = r +≤-∸-≤ zero zero (suc l) r = ≤-refl +≤-∸-≤ zero (suc n) (suc l) r = (n l) , (+-zero _) +≤-∸-≤ (suc m) zero (suc l) r = ⊥.rec (¬-<-zero r) +≤-∸-≤ (suc m) (suc n) (suc l) r = ≤-∸-≤ m n l (pred-≤-pred r) + +<-∸-< : m n l m < n l < n m l < n l +<-∸-< m n zero r q = r +<-∸-< zero zero (suc l) r q = ⊥.rec (¬-<-zero r) +<-∸-< zero (suc n) (suc l) r q = n∸l>0 (suc n) (suc l) q +<-∸-< (suc m) zero (suc l) r q = ⊥.rec (¬-<-zero r) +<-∸-< (suc m) (suc n) (suc l) r q = <-∸-< m n l (pred-≤-pred r) (pred-≤-pred q) + +≤-∸-≥ : n l k l k n k n l +≤-∸-≥ n zero zero r = ≤-refl +≤-∸-≥ n zero (suc k) r = ∸-≤ n (suc k) +≤-∸-≥ n (suc l) zero r = ⊥.rec (¬-<-zero r) +≤-∸-≥ zero (suc l) (suc k) r = ≤-refl +≤-∸-≥ (suc n) (suc l) (suc k) r = ≤-∸-≥ n l k (pred-≤-pred r) \ No newline at end of file diff --git a/docs/Cubical.Data.Nat.Properties.html b/docs/Cubical.Data.Nat.Properties.html index 6db4d26..85a2d95 100644 --- a/docs/Cubical.Data.Nat.Properties.html +++ b/docs/Cubical.Data.Nat.Properties.html @@ -42,13 +42,13 @@ maxComm (suc n) (suc m) = cong suc (maxComm n m) znots : ¬ (0 suc n) -znots eq = subst (caseNat ) eq 0 +znots eq = subst (caseNat ) eq 0 snotz : ¬ (suc n 0) -snotz eq = subst (caseNat ) eq 0 +snotz eq = subst (caseNat ) eq 0 injSuc : suc m suc n m n -injSuc p = cong predℕ p +injSuc p = cong predℕ p -- encode decode caracterisation of equality codeℕ : (n m : ) Type ℓ-zero @@ -136,7 +136,7 @@ -- Arithmetic facts about predℕ -suc-predℕ : n ¬ n 0 n suc (predℕ n) +suc-predℕ : n ¬ n 0 n suc (predℕ n) suc-predℕ zero p = ⊥.rec (p refl) suc-predℕ (suc n) p = refl @@ -274,25 +274,25 @@ zero choose suc k = 0 suc n choose suc k = n choose (suc k) + n choose k -evenOrOdd : (n : ) isEvenT n isOddT n +evenOrOdd : (n : ) isEvenT n isOddT n evenOrOdd zero = inl tt evenOrOdd (suc zero) = inr tt evenOrOdd (suc (suc n)) = evenOrOdd n -¬evenAndOdd : (n : ) ¬ isEvenT n × isOddT n +¬evenAndOdd : (n : ) ¬ isEvenT n × isOddT n ¬evenAndOdd zero (p , ()) ¬evenAndOdd (suc zero) () ¬evenAndOdd (suc (suc n)) = ¬evenAndOdd n -isPropIsEvenT : (n : ) isProp (isEvenT n) +isPropIsEvenT : (n : ) isProp (isEvenT n) isPropIsEvenT zero x y = refl isPropIsEvenT (suc zero) = isProp⊥ isPropIsEvenT (suc (suc n)) = isPropIsEvenT n -isPropIsOddT : (n : ) isProp (isOddT n) +isPropIsOddT : (n : ) isProp (isOddT n) isPropIsOddT n = isPropIsEvenT (suc n) -isPropEvenOrOdd : (n : ) isProp (isEvenT n isOddT n) +isPropEvenOrOdd : (n : ) isProp (isEvenT n isOddT n) isPropEvenOrOdd n (inl x) (inl x₁) = cong inl (isPropIsEvenT n x x₁) isPropEvenOrOdd n (inl x) (inr x₁) = ⊥.rec (¬evenAndOdd n (x , x₁)) isPropEvenOrOdd n (inr x) (inl x₁) = ⊥.rec (¬evenAndOdd (suc n) (x , x₁)) diff --git a/docs/Cubical.Foundations.Pointed.Homogeneous.html b/docs/Cubical.Foundations.Pointed.Homogeneous.html index 96d4fef..8c97d49 100644 --- a/docs/Cubical.Foundations.Pointed.Homogeneous.html +++ b/docs/Cubical.Foundations.Pointed.Homogeneous.html @@ -58,153 +58,161 @@ }) (sym (h (pt B∙)) h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) -→∙Homogeneous≡Path : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} - (h : isHomogeneous B∙) (p q : f∙ g∙) cong fst p cong fst q p q -→∙Homogeneous≡Path {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(f , f₀)} {g∙@(g , g₀)} h p q r = - transport k - PathP i - PathP j (A , a₀) →∙ newPath-refl p q r i j (~ k)) - (f , f₀) (g , g₀)) p q) - (badPath p q r) - where - newPath : (p q : f∙ g∙) (r : cong fst p cong fst q) - Square (refl {x = b}) refl refl refl - newPath p q r i j = - hcomp k λ {(i = i0) cong snd p j k - ; (i = i1) cong snd q j k - ; (j = i0) f₀ k - ; (j = i1) g₀ k}) - (r i j a₀) - - newPath-refl : (p q : f∙ g∙) (r : cong fst p cong fst q) - PathP i (PathP j B∙ (B , newPath p q r i j))) refl refl) refl refl - newPath-refl p q r i j k = - hcomp w λ { (i = i0) lCancel (h b) w k - ; (i = i1) lCancel (h b) w k - ; (j = i0) lCancel (h b) w k - ; (j = i1) lCancel (h b) w k - ; (k = i0) lCancel (h b) w k - ; (k = i1) B , newPath p q r i j}) - ((sym (h b) h (newPath p q r i j)) k) - - badPath : (p q : f∙ g∙) (r : cong fst p cong fst q) - PathP i - PathP j A∙ →∙ (B , newPath p q r i j)) - (f , f₀) (g , g₀)) - p q - fst (badPath p q r i j) = r i j - snd (badPath p q s i j) k = - hcomp r λ { (i = i0) snd (p j) (r k) - ; (i = i1) snd (q j) (r k) - ; (j = i0) f₀ (k r) - ; (j = i1) g₀ (k r) - ; (k = i0) s i j a₀}) - (s i j a₀) - -→∙HomogeneousSquare : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ h∙ l∙ : A∙ →∙ B∙} - (h : isHomogeneous B∙) (s : f∙ h∙) (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) - Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) - Square p q s t -→∙HomogeneousSquare {f∙ = f∙} {g∙ = g∙} {h∙ = h∙} {l∙ = l∙} h = - J h∙ s (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) - Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) - Square p q s t) - (J l∙ t (p : f∙ g∙) (q : f∙ l∙) - Square (cong fst p) (cong fst q) refl (cong fst t) - Square p q refl t) - (→∙Homogeneous≡Path {f∙ = f∙} {g∙ = g∙} h)) - -isHomogeneousPi : { ℓ'} {A : Type } {B∙ : A Pointed ℓ'} - (∀ a isHomogeneous (B∙ a)) isHomogeneous (Πᵘ∙ A B∙) -isHomogeneousPi h f i .fst = a typ (h a (f a) i) -isHomogeneousPi h f i .snd a = pt (h a (f a) i) - -isHomogeneousΠ∙ : { ℓ'} (A : Pointed ) (B : typ A Type ℓ') - (b₀ : B (pt A)) - ((a : typ A) (x : B a) isHomogeneous (B a , x)) - (f : Π∙ A B b₀) - isHomogeneous (Π∙ A B b₀ , f) -fst (isHomogeneousΠ∙ A B b₀ h f g i) = - Σ[ r ((a : typ A) fst ((h a (fst f a) (fst g a)) i)) ] - r (pt A) hcomp k λ {(i = i0) snd f k - ; (i = i1) snd g k}) - (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) -snd (isHomogeneousΠ∙ A B b₀ h f g i) = - a snd (h a (fst f a) (fst g a) i)) - , λ j hcomp k λ { (i = i0) snd f (k j) - ; (i = i1) snd g (k j) - ; (j = i0) snd (h (pt A) (fst f (pt A)) - (fst g (pt A)) i)}) - (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) - -isHomogeneous→∙ : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} - isHomogeneous B∙ isHomogeneous (A∙ →∙ B∙ ) -isHomogeneous→∙ {A∙ = A∙} {B∙} h f∙ = - ΣPathP - ( i Π∙ A∙ a T a i) (t₀ i)) - , PathPIsoPath _ _ _ .Iso.inv - (→∙Homogeneous≡ h - (PathPIsoPath i (a : typ A∙) T a i) _ pt B∙) _ .Iso.fun - i a pt (h (f∙ .fst a) i)))) - ) - where - T : a typ B∙ typ B∙ - T a i = typ (h (f∙ .fst a) i) - - t₀ : PathP i T (pt A∙) i) (pt B∙) (pt B∙) - t₀ = cong pt (h (f∙ .fst (pt A∙))) f∙ .snd - -isHomogeneousProd : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} - isHomogeneous A∙ isHomogeneous B∙ isHomogeneous (A∙ ×∙ B∙) -isHomogeneousProd hA hB (a , b) i .fst = typ (hA a i) × typ (hB b i) -isHomogeneousProd hA hB (a , b) i .snd .fst = pt (hA a i) -isHomogeneousProd hA hB (a , b) i .snd .snd = pt (hB b i) - -isHomogeneousPath : {} (A : Type ) {x y : A} (p : x y) isHomogeneous ((x y) , p) -isHomogeneousPath A {x} {y} p q - = pointed-sip ((x y) , p) ((x y) , q) (eqv , compPathr-cancel p q) - where eqv : (x y) (x y) - eqv = compPathlEquiv (q sym p) - -module HomogeneousDiscrete {} {A∙ : Pointed } (dA : Discrete (typ A∙)) (y : typ A∙) where - - -- switches pt A∙ with y - switch : typ A∙ typ A∙ - switch x with dA x (pt A∙) - ... | yes _ = y - ... | no _ with dA x y - ... | yes _ = pt A∙ - ... | no _ = x - - switch-ptA∙ : switch (pt A∙) y - switch-ptA∙ with dA (pt A∙) (pt A∙) - ... | yes _ = refl - ... | no ¬p = ⊥.rec (¬p refl) - - switch-idp : x switch (switch x) x - switch-idp x with dA x (pt A∙) - switch-idp x | yes p with dA y (pt A∙) - switch-idp x | yes p | yes q = q sym p - switch-idp x | yes p | no _ with dA y y - switch-idp x | yes p | no _ | yes _ = sym p - switch-idp x | yes p | no _ | no ¬p = ⊥.rec (¬p refl) - switch-idp x | no ¬p with dA x y - switch-idp x | no ¬p | yes p with dA y (pt A∙) - switch-idp x | no ¬p | yes p | yes q = ⊥.rec (¬p (p q)) - switch-idp x | no ¬p | yes p | no _ with dA (pt A∙) (pt A∙) - switch-idp x | no ¬p | yes p | no _ | yes _ = sym p - switch-idp x | no ¬p | yes p | no _ | no ¬q = ⊥.rec (¬q refl) - switch-idp x | no ¬p | no ¬q with dA x (pt A∙) - switch-idp x | no ¬p | no ¬q | yes p = ⊥.rec (¬p p) - switch-idp x | no ¬p | no ¬q | no _ with dA x y - switch-idp x | no ¬p | no ¬q | no _ | yes q = ⊥.rec (¬q q) - switch-idp x | no ¬p | no ¬q | no _ | no _ = refl - - switch-eqv : typ A∙ typ A∙ - switch-eqv = isoToEquiv (iso switch switch switch-idp switch-idp) - -isHomogeneousDiscrete : {} {A∙ : Pointed } (dA : Discrete (typ A∙)) isHomogeneous A∙ -isHomogeneousDiscrete {} {A∙} dA y - = pointed-sip (typ A∙ , pt A∙) (typ A∙ , y) (switch-eqv , switch-ptA∙) - where open HomogeneousDiscrete {} {A∙} dA y +→∙HomogeneousPathP : + { ℓ'} {A∙ A∙' : Pointed } {B∙ B∙' : Pointed ℓ'} + {f∙ : A∙ →∙ B∙} {g∙ : A∙' →∙ B∙'} (p : A∙ A∙') (q : B∙ B∙') + (h : isHomogeneous B∙') + PathP i fst (p i) fst (q i)) (f∙ .fst) (g∙ .fst) + PathP i p i →∙ q i) f∙ g∙ +→∙HomogeneousPathP p q h r = toPathP (→∙Homogeneous≡ h (fromPathP r)) + +→∙Homogeneous≡Path : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) (p q : f∙ g∙) cong fst p cong fst q p q +→∙Homogeneous≡Path {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(f , f₀)} {g∙@(g , g₀)} h p q r = + transport k + PathP i + PathP j (A , a₀) →∙ newPath-refl p q r i j (~ k)) + (f , f₀) (g , g₀)) p q) + (badPath p q r) + where + newPath : (p q : f∙ g∙) (r : cong fst p cong fst q) + Square (refl {x = b}) refl refl refl + newPath p q r i j = + hcomp k λ {(i = i0) cong snd p j k + ; (i = i1) cong snd q j k + ; (j = i0) f₀ k + ; (j = i1) g₀ k}) + (r i j a₀) + + newPath-refl : (p q : f∙ g∙) (r : cong fst p cong fst q) + PathP i (PathP j B∙ (B , newPath p q r i j))) refl refl) refl refl + newPath-refl p q r i j k = + hcomp w λ { (i = i0) lCancel (h b) w k + ; (i = i1) lCancel (h b) w k + ; (j = i0) lCancel (h b) w k + ; (j = i1) lCancel (h b) w k + ; (k = i0) lCancel (h b) w k + ; (k = i1) B , newPath p q r i j}) + ((sym (h b) h (newPath p q r i j)) k) + + badPath : (p q : f∙ g∙) (r : cong fst p cong fst q) + PathP i + PathP j A∙ →∙ (B , newPath p q r i j)) + (f , f₀) (g , g₀)) + p q + fst (badPath p q r i j) = r i j + snd (badPath p q s i j) k = + hcomp r λ { (i = i0) snd (p j) (r k) + ; (i = i1) snd (q j) (r k) + ; (j = i0) f₀ (k r) + ; (j = i1) g₀ (k r) + ; (k = i0) s i j a₀}) + (s i j a₀) + +→∙HomogeneousSquare : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ h∙ l∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) (s : f∙ h∙) (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) + Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) + Square p q s t +→∙HomogeneousSquare {f∙ = f∙} {g∙ = g∙} {h∙ = h∙} {l∙ = l∙} h = + J h∙ s (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) + Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) + Square p q s t) + (J l∙ t (p : f∙ g∙) (q : f∙ l∙) + Square (cong fst p) (cong fst q) refl (cong fst t) + Square p q refl t) + (→∙Homogeneous≡Path {f∙ = f∙} {g∙ = g∙} h)) + +isHomogeneousPi : { ℓ'} {A : Type } {B∙ : A Pointed ℓ'} + (∀ a isHomogeneous (B∙ a)) isHomogeneous (Πᵘ∙ A B∙) +isHomogeneousPi h f i .fst = a typ (h a (f a) i) +isHomogeneousPi h f i .snd a = pt (h a (f a) i) + +isHomogeneousΠ∙ : { ℓ'} (A : Pointed ) (B : typ A Type ℓ') + (b₀ : B (pt A)) + ((a : typ A) (x : B a) isHomogeneous (B a , x)) + (f : Π∙ A B b₀) + isHomogeneous (Π∙ A B b₀ , f) +fst (isHomogeneousΠ∙ A B b₀ h f g i) = + Σ[ r ((a : typ A) fst ((h a (fst f a) (fst g a)) i)) ] + r (pt A) hcomp k λ {(i = i0) snd f k + ; (i = i1) snd g k}) + (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) +snd (isHomogeneousΠ∙ A B b₀ h f g i) = + a snd (h a (fst f a) (fst g a) i)) + , λ j hcomp k λ { (i = i0) snd f (k j) + ; (i = i1) snd g (k j) + ; (j = i0) snd (h (pt A) (fst f (pt A)) + (fst g (pt A)) i)}) + (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) + +isHomogeneous→∙ : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} + isHomogeneous B∙ isHomogeneous (A∙ →∙ B∙ ) +isHomogeneous→∙ {A∙ = A∙} {B∙} h f∙ = + ΣPathP + ( i Π∙ A∙ a T a i) (t₀ i)) + , PathPIsoPath _ _ _ .Iso.inv + (→∙Homogeneous≡ h + (PathPIsoPath i (a : typ A∙) T a i) _ pt B∙) _ .Iso.fun + i a pt (h (f∙ .fst a) i)))) + ) + where + T : a typ B∙ typ B∙ + T a i = typ (h (f∙ .fst a) i) + + t₀ : PathP i T (pt A∙) i) (pt B∙) (pt B∙) + t₀ = cong pt (h (f∙ .fst (pt A∙))) f∙ .snd + +isHomogeneousProd : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} + isHomogeneous A∙ isHomogeneous B∙ isHomogeneous (A∙ ×∙ B∙) +isHomogeneousProd hA hB (a , b) i .fst = typ (hA a i) × typ (hB b i) +isHomogeneousProd hA hB (a , b) i .snd .fst = pt (hA a i) +isHomogeneousProd hA hB (a , b) i .snd .snd = pt (hB b i) + +isHomogeneousPath : {} (A : Type ) {x y : A} (p : x y) isHomogeneous ((x y) , p) +isHomogeneousPath A {x} {y} p q + = pointed-sip ((x y) , p) ((x y) , q) (eqv , compPathr-cancel p q) + where eqv : (x y) (x y) + eqv = compPathlEquiv (q sym p) + +module HomogeneousDiscrete {} {A∙ : Pointed } (dA : Discrete (typ A∙)) (y : typ A∙) where + + -- switches pt A∙ with y + switch : typ A∙ typ A∙ + switch x with dA x (pt A∙) + ... | yes _ = y + ... | no _ with dA x y + ... | yes _ = pt A∙ + ... | no _ = x + + switch-ptA∙ : switch (pt A∙) y + switch-ptA∙ with dA (pt A∙) (pt A∙) + ... | yes _ = refl + ... | no ¬p = ⊥.rec (¬p refl) + + switch-idp : x switch (switch x) x + switch-idp x with dA x (pt A∙) + switch-idp x | yes p with dA y (pt A∙) + switch-idp x | yes p | yes q = q sym p + switch-idp x | yes p | no _ with dA y y + switch-idp x | yes p | no _ | yes _ = sym p + switch-idp x | yes p | no _ | no ¬p = ⊥.rec (¬p refl) + switch-idp x | no ¬p with dA x y + switch-idp x | no ¬p | yes p with dA y (pt A∙) + switch-idp x | no ¬p | yes p | yes q = ⊥.rec (¬p (p q)) + switch-idp x | no ¬p | yes p | no _ with dA (pt A∙) (pt A∙) + switch-idp x | no ¬p | yes p | no _ | yes _ = sym p + switch-idp x | no ¬p | yes p | no _ | no ¬q = ⊥.rec (¬q refl) + switch-idp x | no ¬p | no ¬q with dA x (pt A∙) + switch-idp x | no ¬p | no ¬q | yes p = ⊥.rec (¬p p) + switch-idp x | no ¬p | no ¬q | no _ with dA x y + switch-idp x | no ¬p | no ¬q | no _ | yes q = ⊥.rec (¬q q) + switch-idp x | no ¬p | no ¬q | no _ | no _ = refl + + switch-eqv : typ A∙ typ A∙ + switch-eqv = isoToEquiv (iso switch switch switch-idp switch-idp) + +isHomogeneousDiscrete : {} {A∙ : Pointed } (dA : Discrete (typ A∙)) isHomogeneous A∙ +isHomogeneousDiscrete {} {A∙} dA y + = pointed-sip (typ A∙ , pt A∙) (typ A∙ , y) (switch-eqv , switch-ptA∙) + where open HomogeneousDiscrete {} {A∙} dA y \ No newline at end of file diff --git a/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html b/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html index 62289a3..945ae72 100644 --- a/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html +++ b/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html @@ -68,7 +68,7 @@ private open import Cubical.Data.Nat - open Recover ( , zero) (isHomogeneousDiscrete discreteℕ) + open Recover ( , zero) (isHomogeneousDiscrete discreteℕ) -- only `∣hidden∣` is exported, `hidden` is no longer in scope module _ where diff --git a/docs/Realizability.Topos.FunctionalRelation.html b/docs/Realizability.Topos.FunctionalRelation.html index 3e4946c..9896e56 100644 --- a/docs/Realizability.Topos.FunctionalRelation.html +++ b/docs/Realizability.Topos.FunctionalRelation.html @@ -1,92 +1,384 @@ -Realizability.Topos.FunctionalRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
-open import Realizability.CombinatoryAlgebra
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Structure
-open import Cubical.Data.Vec
-open import Cubical.Data.Nat
-open import Cubical.Data.FinData renaming (zero to fzero)
-open import Cubical.Data.Sigma
-open import Cubical.Data.Empty
-
-module Realizability.Topos.FunctionalRelation
-  { ℓ' ℓ''}
-  {A : Type }
-  (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
-  where
-
-open import Realizability.Tripos.Logic.Syntax { = ℓ'}
-open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_)
-open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
-open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
-
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Predicate' renaming (isSetX to isSetPredicateBase)
-open PredicateProperties
-open Morphism
-
-private
-  _⊩_ :  a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*)  Type _
-  a  P = a pre⊩  P  tt*
-
-record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
-  open PartialEquivalenceRelation
-  field
-    perX : PartialEquivalenceRelation X
-    perY : PartialEquivalenceRelation Y
-  _~X_ = perX ._~_
-  _~Y_ = perY ._~_
-
-  field
-    relation : Predicate (X × Y)
-
-  private
-    `X : Sort
-    `X =  (X , perX .isSetX)
-
-    `Y : Sort
-    `Y =  (Y , perY .isSetX)
-
-    relationSymbol : Vec Sort 3
-    relationSymbol = (`X  `Y)  `X  `X  `Y  `Y  []
-
-    relationSymbolInterpretation : RelationInterpretation relationSymbol
-    relationSymbolInterpretation fzero = relation
-    relationSymbolInterpretation one = _~X_
-    relationSymbolInterpretation two = _~Y_
-
-    `relation : Fin 3
-    `relation = fzero
-    `~X : Fin 3
-    `~X = one
-    `~Y : Fin 3
-    `~Y = two
-
-  module RelationInterpretation = Interpretation relationSymbol relationSymbolInterpretation isNonTrivial
-  open RelationInterpretation
-
-  module RelationSoundness = Soundness isNonTrivial relationSymbolInterpretation
-  open RelationSoundness
-
-  -- Strictness
-  -- If the functional relation holds for x and y then x and y "exist"
-  private
-    strictnessContext : Context
-    strictnessContext = ([]  `X)  `Y
-
-    x∈strictnessContext : `X  strictnessContext
-    x∈strictnessContext = there here
-
-    y∈strictnessContext : `Y  strictnessContext
-    y∈strictnessContext = here
-
-    xˢᵗ : Term strictnessContext `X
-    xˢᵗ = var x∈strictnessContext
-
-    yˢᵗ : Term strictnessContext `Y
-    yˢᵗ = var y∈strictnessContext
+Realizability.Topos.FunctionalRelation
{-# OPTIONS --allow-unsolved-metas #-}
+open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData
+open import Cubical.Data.Fin hiding (Fin; _/_)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients
+open import Cubical.Categories.Category
+
+module Realizability.Topos.FunctionalRelation
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
+
+open import Realizability.Tripos.Logic.Syntax { = ℓ'}
+open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_)
+open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate' renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+private
+  _⊩_ :  a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*)  Type _
+  a  P = a pre⊩  P  tt*
+
+  
+private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
+private λ* = `λ* as fefermanStructure
+
+
+open PartialEquivalenceRelation
+
+record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  field
+    perX : PartialEquivalenceRelation X
+    perY : PartialEquivalenceRelation Y
+  _~X_ = perX ._~_
+  _~Y_ = perY ._~_
+
+  field
+    relation : Predicate (X × Y)
+
+  `X : Sort
+  `X =  (X , perX .isSetX)
+
+  `Y : Sort
+  `Y =  (Y , perY .isSetX)
+
+  private
+    relationSymbol : Vec Sort 3
+    relationSymbol = (`X  `Y)  `X  `X  `Y  `Y  []
+
+    `F : Fin 3
+    `F = zero
+    `~X : Fin 3
+    `~X = one
+    `~Y : Fin 3
+    `~Y = two
+
+  open Relational relationSymbol
+
+  module RelationInterpretation' = Interpretation relationSymbol  { zero  relation ; one  _~X_ ; two  _~Y_ }) isNonTrivial
+  open RelationInterpretation'
+
+  module RelationSoundness = Soundness {relSym = relationSymbol} isNonTrivial  { zero  relation ; one  _~X_ ; two  _~Y_ })
+  open RelationSoundness
+
+  -- # Strictness
+  -- If the functional relation holds for x and y then x and y "exist"
+  private
+    strictnessContext : Context
+    strictnessContext = ([]  `X)  `Y
+
+    x∈strictnessContext : `X  strictnessContext
+    x∈strictnessContext = there here
+
+    y∈strictnessContext : `Y  strictnessContext
+    y∈strictnessContext = here
+
+    xˢᵗ : Term strictnessContext `X
+    xˢᵗ = var x∈strictnessContext
+
+    yˢᵗ : Term strictnessContext `Y
+    yˢᵗ = var y∈strictnessContext
+
+  -- F : Rel X Y, _~X_ : Rel X X, _~Y_ : Rel Y Y ⊢ F(x ,y) → (x ~X x) ∧ (y ~Y y)
+  strictnessFormula : Formula strictnessContext
+  strictnessFormula = rel `F (xˢᵗ `, yˢᵗ) `→ (rel `~X (xˢᵗ `, xˢᵗ) `∧ rel `~Y (yˢᵗ `, yˢᵗ))
+
+  field
+    isStrict : holdsInTripos strictnessFormula
+
+  -- # Relational
+  -- The functional relation preserves equality
+  -- "Substitutive" might be a better term for this property
+  private
+    relationalContext : Context
+    relationalContext =
+      []  `Y  `Y  `X  `X
+
+    x₁∈relationalContext : `X  relationalContext
+    x₁∈relationalContext = there here
+
+    x₂∈relationalContext : `X  relationalContext
+    x₂∈relationalContext = here
+
+    y₁∈relationalContext : `Y  relationalContext
+    y₁∈relationalContext = there (there here)
+
+    y₂∈relationalContext : `Y  relationalContext
+    y₂∈relationalContext = there (there (there here))
+
+    x₁ = var x₁∈relationalContext
+    x₂ = var x₂∈relationalContext
+    y₁ = var y₁∈relationalContext
+    y₂ = var y₂∈relationalContext
+
+  relationalFormula : Formula relationalContext
+  relationalFormula = (rel `F (x₁ `, y₁) `∧ (rel `~X (x₁ `, x₂) `∧ rel `~Y (y₁ `, y₂))) `→ rel `F (x₂ `, y₂)
+
+  field
+    isRelational : holdsInTripos relationalFormula
+
+  -- # Single-valued
+  -- Self explanatory
+  private
+    singleValuedContext : Context
+    singleValuedContext =
+      []  `Y  `Y  `X
+
+    x∈singleValuedContext : `X  singleValuedContext
+    x∈singleValuedContext = here
+
+    y₁∈singleValuedContext : `Y  singleValuedContext
+    y₁∈singleValuedContext = there here
+
+    y₂∈singleValuedContext : `Y  singleValuedContext
+    y₂∈singleValuedContext = there (there here)
+
+    xˢᵛ = var x∈singleValuedContext
+    y₁ˢᵛ = var y₁∈singleValuedContext
+    y₂ˢᵛ = var y₂∈singleValuedContext
+
+  singleValuedFormula : Formula singleValuedContext
+  singleValuedFormula =
+    (rel `F (xˢᵛ `, y₁ˢᵛ) `∧ rel `F (xˢᵛ `, y₂ˢᵛ)) `→ rel `~Y (y₁ˢᵛ `, y₂ˢᵛ)
+
+  field
+    isSingleValued : holdsInTripos singleValuedFormula
+
+  -- # Total
+  -- For all existent elements in the domain x there is an element in the codomain y
+  -- such that F(x, y)
+  private
+    totalContext : Context
+    totalContext =
+      []  `X
+
+    x∈totalContext : `X  totalContext
+    x∈totalContext = here
+
+    xᵗˡ = var x∈totalContext
+
+  totalFormula : Formula totalContext
+  totalFormula = rel `~X (xᵗˡ `, xᵗˡ)  `→ `∃ (rel `F (renamingTerm (drop id) xᵗˡ `, var here))
+
+  field
+    isTotal : holdsInTripos totalFormula
+
+open FunctionalRelation hiding (`X; `Y)
+
+pointwiseEntailment :  {X Y : Type ℓ'}  FunctionalRelation X Y  FunctionalRelation X Y  Type _
+pointwiseEntailment {X} {Y} F G = holdsInTripos entailmentFormula where
   
+  `X : Sort
+  `Y : Sort
+
+  `X =  (X , F .perX .isSetX)
+  `Y =  (Y , F .perY .isSetX)
+
+  relationSymbols : Vec Sort 2
+  relationSymbols = (`X  `Y)  (`X  `Y)  []
+
+  `F : Fin 2
+  `F = zero
+
+  `G : Fin 2
+  `G = one
+
+  open Relational relationSymbols
+
+  module RelationalInterpretation = Interpretation relationSymbols  { zero  F .relation ; one  G .relation }) isNonTrivial
+  open RelationalInterpretation
+
+  module RelationalSoundness = Soundness {relSym = relationSymbols} isNonTrivial  { zero  F .relation ; one  G .relation })
+  open RelationalSoundness
+
+  entailmentContext : Context
+  entailmentContext = []  `X  `Y
+
+  x∈entailmentContext : `X  entailmentContext
+  x∈entailmentContext = there here
+
+  y∈entailmentContext : `Y  entailmentContext
+  y∈entailmentContext = here
+
+  x = var x∈entailmentContext
+  y = var y∈entailmentContext
+
+  entailmentFormula : Formula entailmentContext
+  entailmentFormula = rel `F (x `, y) `→ rel `G (x `, y)
+
+functionalRelationIsomorphism :  {X Y : Type ℓ'}  FunctionalRelation X Y  FunctionalRelation X Y  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
+functionalRelationIsomorphism {X} {Y} F G =
+  pointwiseEntailment F G × pointwiseEntailment G F
+
+
+pointwiseEntailment→functionalRelationIsomorphism :  {X Y : Type ℓ'}  (F G : FunctionalRelation X Y)  pointwiseEntailment F G  functionalRelationIsomorphism F G
+pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , G≤F where
+    
+  `X : Sort
+  `Y : Sort
+
+  `X =  (X , F .perX .isSetX)
+  `Y =  (Y , F .perY .isSetX)
+
+  relationSymbols : Vec Sort 4
+  relationSymbols = (`X  `Y)  (`X  `Y)  (`X  `X)  (`Y  `Y)  []
+
+  `F : Fin 4
+  `F = zero
+
+  `G : Fin 4
+  `G = one
+
+  `~X : Fin 4
+  `~X = two
+
+  `~Y : Fin 4
+  `~Y = three
+
+  open Interpretation relationSymbols  { zero  F .relation ; one  G .relation ; two  F .perX ._~_ ; three  G .perY ._~_}) isNonTrivial
+  open Soundness {relSym = relationSymbols} isNonTrivial ((λ { zero  F .relation ; one  G .relation ; two  F .perX ._~_ ; three  G .perY ._~_}))
+  open Relational relationSymbols
+
+  -- What we need to prove is that
+  -- F ≤ G ⊨ G ≤ F
+  -- We will use the semantic combinators we borrowed from the 1lab
+
+  entailmentContext : Context
+  entailmentContext = []  `X  `Y
+
+  x : Term entailmentContext `X
+  x = var (there here)
+
+  y : Term entailmentContext `Y
+  y = var here
+
+  G⊨x~x : (⊤ᵗ `∧ rel `G (x `, y))  rel `~X (x `, x)
+  G⊨x~x =
+    `→elim
+      {Γ = entailmentContext}
+      {ϕ = ⊤ᵗ `∧ (rel `G (x `, y))}
+      {ψ = rel `G (x `, y)}
+      {θ = rel `~X (x `, x)}
+      {!G .isStrict!}
+      {!!}
+
+  ⊤∧G⊨F : (⊤ᵗ `∧ rel `G (x `, y))  rel `F (x `, y)
+  ⊤∧G⊨F = cut {Γ = entailmentContext} {ϕ = ⊤ᵗ `∧ rel `G (x `, y)} {ψ = rel `~X (x `, x)}
+           {θ = rel `F (x `, y)}
+           G⊨x~x
+           {!!}
+
+  G≤F : pointwiseEntailment G F
+  G≤F = `→intro {Γ = entailmentContext} {ϕ = ⊤ᵗ} {ψ = rel `G (x `, y)} {θ = rel `F (x `, y)} ⊤∧G⊨F
+
+RTMorphism : (X Y : Type ℓ')   Type _
+RTMorphism X Y = FunctionalRelation X Y / λ F G  functionalRelationIsomorphism F G
+
+RTObject : Type _
+RTObject = Σ[ X  Type ℓ' ] PartialEquivalenceRelation X
+
+idMorphism : (ob : RTObject)  RTMorphism (ob .fst) (ob .fst)
+idMorphism ob =
+  [ record
+     { perX = ob .snd
+     ; perY = ob .snd
+     ; relation = ob .snd ._~_
+     ; isStrict = {!!}
+     ; isRelational = {!!}
+     ; isSingleValued = {!!}
+     ; isTotal = {!!}
+     } ] where
+
+  `X : Sort
+  `X =  (ob .fst , ob .snd .isSetX)
+
+  relationSymbols : Vec Sort 3
+  relationSymbols = (`X  `X)  (`X  `X)  (`X  `X)  []
+
+  open Interpretation relationSymbols  { zero  ob .snd ._~_ ; one  ob .snd ._~_ ; two  ob .snd ._~_ }) isNonTrivial
+  open Soundness {relSym = relationSymbols} isNonTrivial  { zero  ob .snd ._~_ ; one  ob .snd ._~_ ; two  ob .snd ._~_ })
+  open Relational relationSymbols
+
+  isStrictContext : Context
+  isStrictContext = []  `X  `X
+
+  x : Term isStrictContext `X
+  x = var (there here)
+
+  y : Term isStrictContext `X
+  y = var here
+
+  holdsInTriposIsStrict : holdsInTripos (rel zero (x `, y) `→ (rel one (x `, y) `∧ rel two (x `, y)))
+  holdsInTriposIsStrict =
+    do
+    let
+      prover : ApplStrTerm as 2
+      prover =
+        ` pair ̇ # fone ̇ # fone
+    return
+      (λ* prover ,
+       { ((tt* , x') , y') a tt* b b⊩x'~y'
+        
+          let
+            proofEq : λ* prover  a  b  pair  b  b
+            proofEq =
+              λ*ComputationRule prover (a  b  [])
+
+            pr₁proofEq : pr₁  (λ* prover  a  b)  b
+            pr₁proofEq =
+              pr₁  (λ* prover  a  b)
+                ≡⟨ cong  x  pr₁  x) proofEq 
+              pr₁  (pair  b  b)
+                ≡⟨ pr₁pxy≡x b b 
+              b
+                
+
+            pr₂proofEq : pr₂  (λ* prover  a  b)  b
+            pr₂proofEq =
+              pr₂  (λ* prover  a  b)
+                ≡⟨ cong  x  pr₂  x) proofEq 
+              pr₂  (pair  b  b)
+                ≡⟨ pr₂pxy≡y b b 
+              b
+                
+          in
+          (subst
+             r  r pre⊩  (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX))  γ  snd (fst γ) , snd γ) (ob .snd ._~_))  ((tt* , x') , y'))
+          (sym pr₁proofEq) b⊩x'~y') ,
+          subst
+             r  r pre⊩  (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX))  γ  snd (fst γ) , snd γ) (ob .snd ._~_))  ((tt* , x') , y')) (sym pr₂proofEq) b⊩x'~y' }))
+  
+
+RT : Category (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) ((ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')))
+Category.ob RT = Σ[ X  Type ℓ' ] PartialEquivalenceRelation X
+Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism X Y
+Category.id RT {X , perX} = {!!}
+Category._⋆_ RT = {!!}
+Category.⋆IdL RT = {!!}
+Category.⋆IdR RT = {!!}
+Category.⋆Assoc RT = {!!}
+Category.isSetHom RT = {!!}
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.Object.html b/docs/Realizability.Topos.Object.html index ac46ab0..70a72f3 100644 --- a/docs/Realizability.Topos.Object.html +++ b/docs/Realizability.Topos.Object.html @@ -26,99 +26,87 @@ open PredicateProperties open Morphism -private - _⊩_ : a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) Type _ - a P = a pre⊩ P tt* +Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} -record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - field - isSetX : isSet X - private - `X : Sort - `X = (X , isSetX) +record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + field + isSetX : isSet X + private + `X : Sort + `X = (X , isSetX) - `X×X : Sort - `X×X = `X `X + `X×X : Sort + `X×X = `X `X - ~relSym : Vec Sort 1 - ~relSym = `X×X [] + ~relSym : Vec Sort 1 + ~relSym = `X×X [] - module X×XRelational = Relational ~relSym - open X×XRelational + module X×XRelational = Relational {n = 1} ~relSym + open X×XRelational - field - _~_ : Predicate lookup fzero ~relSym ⟧ˢ + field + _~_ : Predicate lookup fzero ~relSym ⟧ˢ - private - ~relSymInterpretation : RelationInterpretation ~relSym - ~relSymInterpretation = λ { fzero _~_ } + private + ~relSymInterpretation : RelationInterpretation ~relSym + ~relSymInterpretation = λ { fzero _~_ } - module ~Interpretation = Interpretation ~relSym ~relSymInterpretation isNonTrivial - open ~Interpretation + module ~Interpretation = Interpretation ~relSym ~relSymInterpretation isNonTrivial + open ~Interpretation - module ~Soundness = Soundness isNonTrivial ~relSymInterpretation - open ~Soundness + module ~Soundness = Soundness isNonTrivial ~relSymInterpretation + open ~Soundness - -- Partial equivalence relations + -- Partial equivalence relations - private - symContext : Context - symContext = ([] `X) `X + private + symContext : Context + symContext = ([] `X) `X - x∈symContext : `X symContext - x∈symContext = there here + x∈symContext : `X symContext + x∈symContext = there here - y∈symContext : `X symContext - y∈symContext = here + y∈symContext : `X symContext + y∈symContext = here - : Term symContext `X - = var x∈symContext + : Term symContext `X + = var x∈symContext - : Term symContext `X - = var y∈symContext + : Term symContext `X + = var y∈symContext - symPrequantFormula : Formula symContext - symPrequantFormula = rel fzero ( `, ) `→ rel fzero ( `, ) + symmetryFormula : Formula symContext + symmetryFormula = rel fzero ( `, ) `→ rel fzero ( `, ) - -- ~ ⊢ ∀ x. ∀ y. x ~ y → y ~ x - symFormula : Formula [] - symFormula = `∀ (`∀ symPrequantFormula) + field + symmetry : holdsInTripos symmetryFormula - field - symWitness : A - symIsWitnessed : symWitness symFormula ⟧ᶠ + private + transContext : Context + transContext = (([] `X) `X) `X - private - transContext : Context - transContext = (([] `X) `X) `X + z∈transContext : `X transContext + z∈transContext = here - z∈transContext : `X transContext - z∈transContext = here + y∈transContext : `X transContext + y∈transContext = there here - y∈transContext : `X transContext - y∈transContext = there here + x∈transContext : `X transContext + x∈transContext = there (there here) - x∈transContext : `X transContext - x∈transContext = there (there here) + zᵗ : Term transContext `X + zᵗ = var z∈transContext - zᵗ : Term transContext `X - zᵗ = var z∈transContext + yᵗ : Term transContext `X + yᵗ = var y∈transContext - yᵗ : Term transContext `X - yᵗ = var y∈transContext + xᵗ : Term transContext `X + xᵗ = var x∈transContext - xᵗ : Term transContext `X - xᵗ = var x∈transContext + transitivityFormula : Formula transContext + transitivityFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) - -- ~ : Rel X X, x : X, y : X, z : X ⊢ x ~ y ∧ y ~ z → x ~ z - transPrequantFormula : Formula transContext - transPrequantFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) - - transFormula : Formula [] - transFormula = `∀ (`∀ (`∀ transPrequantFormula)) - - field - transWitness : A - transIsWitnessed : transWitness transFormula ⟧ᶠ + field + transitivity : holdsInTripos transitivityFormula
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Logic.Semantics.html b/docs/Realizability.Tripos.Logic.Semantics.html index b3dfc9c..af8380b 100644 --- a/docs/Realizability.Tripos.Logic.Semantics.html +++ b/docs/Realizability.Tripos.Logic.Semantics.html @@ -38,564 +38,568 @@ open Predicate' open PredicateProperties hiding (_≤_ ; isTrans≤) open Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''} -Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} -RelationInterpretation : {n : } (Vec Sort n) Type _ -RelationInterpretation {n} relSym = (∀ i Predicate lookup i relSym ⟧ˢ ) -module Interpretation - {n : } - (relSym : Vec Sort n) - (⟦_⟧ʳ : RelationInterpretation relSym) (isNonTrivial : s k ) where - open Relational relSym - - ⟦_⟧ᶜ : Context hSet ℓ' - [] ⟧ᶜ = Unit* , isSetUnit* - c x ⟧ᶜ = ( c ⟧ᶜ .fst × x ⟧ˢ .fst) , isSet× ( c ⟧ᶜ .snd) ( x ⟧ˢ .snd) - - ⟦_⟧ⁿ : {Γ s} s Γ Γ ⟧ᶜ s ⟧ˢ - ⟦_⟧ⁿ {.(_ s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧ - ⟦_⟧ⁿ {.(_ _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) = s∈Γ ⟧ⁿ ⟦Γ⟧ - - ⟦_⟧ᵗ : {Γ s} Term Γ s Γ ⟧ᶜ s ⟧ˢ - ⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ = x ⟧ⁿ ⟦Γ⟧ - ⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = ( t ⟧ᵗ ⟦Γ⟧) , ( t₁ ⟧ᵗ ⟦Γ⟧) - ⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst ( t ⟧ᵗ ⟦Γ⟧) - ⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd ( t ⟧ᵗ ⟦Γ⟧) - ⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x ( t ⟧ᵗ ⟦Γ⟧) - - ⟦_⟧ᶠ : {Γ} Formula Γ Predicate Γ ⟧ᶜ - ⟦_⟧ᶠ {Γ} ⊤ᵗ = pre1 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial - ⟦_⟧ᶠ {Γ} ⊥ᵗ = pre0 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial - ⟦_⟧ᶠ {Γ} (form `∨ form₁) = _⊔_ Γ ⟧ᶜ form ⟧ᶠ form₁ ⟧ᶠ - ⟦_⟧ᶠ {Γ} (form `∧ form₁) = _⊓_ Γ ⟧ᶜ form ⟧ᶠ form₁ ⟧ᶠ - ⟦_⟧ᶠ {Γ} (form `→ form₁) = _⇒_ Γ ⟧ᶜ form ⟧ᶠ form₁ ⟧ᶠ - ⟦_⟧ᶠ {Γ} ( form) = _⇒_ Γ ⟧ᶜ form ⟧ᶠ ⊥ᵗ {Γ = Γ} ⟧ᶠ - ⟦_⟧ᶠ {Γ} (`∃ {B = B} form) = `∃[_] (isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ)) (str Γ ⟧ᶜ) { (⟦Γ⟧ , ⟦B⟧) ⟦Γ⟧ }) form ⟧ᶠ - ⟦_⟧ᶠ {Γ} (`∀ {B = B} form) = `∀[_] (isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ)) (str Γ ⟧ᶜ) { (⟦Γ⟧ , ⟦B⟧) ⟦Γ⟧ }) form ⟧ᶠ - ⟦_⟧ᶠ {Γ} (rel R t) = ⋆_ (str Γ ⟧ᶜ) (str lookup R relSym ⟧ˢ) t ⟧ᵗ R ⟧ʳ - - -- R for renamings and r for relations - ⟦_⟧ᴿ : {Γ Δ} Renaming Γ Δ Γ ⟧ᶜ Δ ⟧ᶜ - id ⟧ᴿ ctx = ctx - drop ren ⟧ᴿ (ctx , _) = ren ⟧ᴿ ctx - keep ren ⟧ᴿ (ctx , s) = ( ren ⟧ᴿ ctx) , s - - -- B for suBstitution and s for sorts - ⟦_⟧ᴮ : {Γ Δ} Substitution Γ Δ Γ ⟧ᶜ Δ ⟧ᶜ - id ⟧ᴮ ctx = ctx - t , sub ⟧ᴮ ctx = ( sub ⟧ᴮ ctx) , ( t ⟧ᵗ ctx) - drop sub ⟧ᴮ (ctx , s) = sub ⟧ᴮ ctx - - renamingVarSound : {Γ Δ s} (ren : Renaming Γ Δ) (v : s Δ) renamingVar ren v ⟧ⁿ v ⟧ⁿ ren ⟧ᴿ - renamingVarSound {Γ} {.Γ} {s} id v = refl - renamingVarSound {.(_ _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingVarSound ren v i ⟦Γ⟧ } - renamingVarSound {.(_ s)} {.(_ s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i ⟦s⟧ } - renamingVarSound {.(_ _)} {.(_ _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingVarSound ren v i ⟦Γ⟧ } - - renamingTermSound : {Γ Δ s} (ren : Renaming Γ Δ) (t : Term Δ s) renamingTerm ren t ⟧ᵗ t ⟧ᵗ ren ⟧ᴿ - renamingTermSound {Γ} {.Γ} {s} id t = refl - renamingTermSound {.(_ _)} {Δ} {s} (drop ren) (var x) = - funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingVarSound ren x i ⟦Γ⟧ } - renamingTermSound {.(_ _)} {Δ} {.(_ _)} r@(drop ren) (t `, t₁) = - funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) } - renamingTermSound {.(_ _)} {Δ} {s} r@(drop ren) (π₁ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .fst } - renamingTermSound {.(_ _)} {Δ} {s} r@(drop ren) (π₂ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .snd } - renamingTermSound {.(_ _)} {Δ} {s} r@(drop ren) (fun f t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i f (renamingTermSound r t i x) } - renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (var v) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingVarSound r v i x } - renamingTermSound {.(_ _)} {.(_ _)} {.(_ _)} r@(keep ren) (t `, t₁) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) } - renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (π₁ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .fst } - renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (π₂ t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .snd } - renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (fun f t) = - funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i f (renamingTermSound r t i x) } - - substitutionVarSound : {Γ Δ s} (subs : Substitution Γ Δ) (v : s Δ) substitutionVar subs v ⟧ᵗ v ⟧ⁿ subs ⟧ᴮ - substitutionVarSound {Γ} {.Γ} {s} id t = refl - substitutionVarSound {Γ} {.(_ s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧ refl - substitutionVarSound {Γ} {.(_ _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i substitutionVarSound subs t i ⟦Γ⟧ - substitutionVarSound {.(_ _)} {Δ} {s} r@(drop subs) t = - -- TODO : Fix unsolved constraints - funExt - λ { x@(⟦Γ⟧ , ⟦s⟧) - substitutionVar (drop subs) t ⟧ᵗ x - ≡[ i ]⟨ renamingTermSound (drop id) (substitutionVar subs t) i x - substitutionVar subs t ⟧ᵗ ( drop id ⟧ᴿ x) - ≡⟨ refl - substitutionVar subs t ⟧ᵗ ⟦Γ⟧ - ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ - t ⟧ⁿ ( subs ⟧ᴮ ⟦Γ⟧) - } - - substitutionTermSound : {Γ Δ s} (subs : Substitution Γ Δ) (t : Term Δ s) substitutionTerm subs t ⟧ᵗ t ⟧ᵗ subs ⟧ᴮ - substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x - substitutionTermSound {Γ} {Δ} {.(_ _)} subs (t `, t₁) = funExt λ x i (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x) - substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i substitutionTermSound subs t i x .fst - substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i substitutionTermSound subs t i x .snd - substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i f (substitutionTermSound subs t i x) - - semanticSubstitution : {Γ Δ} (subs : Substitution Γ Δ) Predicate Δ ⟧ᶜ Predicate Γ ⟧ᶜ - semanticSubstitution {Γ} {Δ} subs = ⋆_ (str Γ ⟧ᶜ) (str Δ ⟧ᶜ) subs ⟧ᴮ - - -- Due to a shortcut in the soundness of negation termination checking fails - -- TODO : Fix - {-# TERMINATING #-} - substitutionFormulaSound : {Γ Δ} (subs : Substitution Γ Δ) (f : Formula Δ) substitutionFormula subs f ⟧ᶠ semanticSubstitution subs f ⟧ᶠ - substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ = - Predicate≡ - Γ ⟧ᶜ - (pre1 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial) - (semanticSubstitution subs (pre1 Δ ⟧ᶜ (str Δ ⟧ᶜ) isNonTrivial)) - γ a a⊩1γ tt*) - λ γ a a⊩1subsγ tt* - substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ = - Predicate≡ - Γ ⟧ᶜ - (pre0 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial) - (semanticSubstitution subs (pre0 Δ ⟧ᶜ (str Δ ⟧ᶜ) isNonTrivial)) - _ _ bot ⊥rec* bot) - λ _ _ bot bot - substitutionFormulaSound {Γ} {Δ} subs (f `∨ f₁) = - Predicate≡ - Γ ⟧ᶜ - (_⊔_ Γ ⟧ᶜ substitutionFormula subs f ⟧ᶠ substitutionFormula subs f₁ ⟧ᶠ) - (semanticSubstitution subs (_⊔_ Δ ⟧ᶜ f ⟧ᶠ f₁ ⟧ᶠ)) - γ a a⊩substFormFs - a⊩substFormFs >>= - λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) - inl (pr₁a≡k , subst form (pr₂ a) form γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁ - ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) - inr (pr₁a≡k' , subst form (pr₂ a) form γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ }) - λ γ a a⊩semanticSubsFs - a⊩semanticSubsFs >>= - λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) - inl (pr₁a≡k , (subst form (pr₂ a) form γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁ - ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) - inr (pr₁a≡k' , (subst form (pr₂ a) form γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ } - substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) = - Predicate≡ - Γ ⟧ᶜ - (_⊓_ Γ ⟧ᶜ substitutionFormula subs f ⟧ᶠ substitutionFormula subs f₁ ⟧ᶠ) - (semanticSubstitution subs (_⊓_ Δ ⟧ᶜ f ⟧ᶠ f₁ ⟧ᶠ)) - γ a a⊩substFormulaFs - (subst form (pr₁ a) form γ) (substitutionFormulaSound subs f) (a⊩substFormulaFs .fst)) , - (subst form (pr₂ a) form γ) (substitutionFormulaSound subs f₁) (a⊩substFormulaFs .snd))) - λ γ a a⊩semanticSubstFs - (subst form (pr₁ a) form γ) (sym (substitutionFormulaSound subs f)) (a⊩semanticSubstFs .fst)) , - (subst form (pr₂ a) form γ) (sym (substitutionFormulaSound subs f₁)) (a⊩semanticSubstFs .snd)) - substitutionFormulaSound {Γ} {Δ} subs (f `→ f₁) = - Predicate≡ - Γ ⟧ᶜ - (_⇒_ Γ ⟧ᶜ substitutionFormula subs f ⟧ᶠ substitutionFormula subs f₁ ⟧ᶠ) - (semanticSubstitution subs (_⇒_ Δ ⟧ᶜ f ⟧ᶠ f₁ ⟧ᶠ)) - γ a a⊩substFormulaFs - λ b b⊩semanticSubstFs - subst - form (a b) form γ) - (substitutionFormulaSound subs f₁) - (a⊩substFormulaFs - b - (subst - form b form γ) - (sym (substitutionFormulaSound subs f)) - b⊩semanticSubstFs))) - λ γ a a⊩semanticSubstFs - λ b b⊩substFormulaFs - subst - form (a b) form γ) - (sym (substitutionFormulaSound subs f₁)) - (a⊩semanticSubstFs - b - (subst - form b form γ) - (substitutionFormulaSound subs f) - b⊩substFormulaFs)) - substitutionFormulaSound {Γ} {Δ} subs ( f) = - substitutionFormulaSound subs (f `→ ⊥ᵗ) - substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) = - Predicate≡ - Γ ⟧ᶜ - (`∃[ isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ) ] (str Γ ⟧ᶜ) { (f , s) f }) substitutionFormula (var here , drop subs) f ⟧ᶠ) - (semanticSubstitution subs (`∃[ isSet× (str Δ ⟧ᶜ) (str B ⟧ˢ) ] (str Δ ⟧ᶜ) { (γ , b) γ }) f ⟧ᶠ)) - γ a a⊩πSubstFormulaF - a⊩πSubstFormulaF >>= - λ { ((γ' , b) , γ'≡γ , a⊩substFormFγ') - (( subs ⟧ᴮ γ') , b) , - ((cong subs ⟧ᴮ γ'≡γ) , - (subst - form a form (γ' , b)) - (substitutionFormulaSound (var here , drop subs) f) - a⊩substFormFγ' )) ∣₁ }) - λ γ a a⊩semanticSubstF - a⊩semanticSubstF >>= - λ (x@(δ , b) , δ≡subsγ , a⊩fx) - (γ , b) , - (refl , - (subst - form a form (γ , b)) - (sym (substitutionFormulaSound (var here , drop subs) f)) - (subst x a f ⟧ᶠ (x , b)) δ≡subsγ a⊩fx))) ∣₁ - substitutionFormulaSound {Γ} {Δ} subs (`∀ {B = B} f) = - Predicate≡ - Γ ⟧ᶜ - (`∀[ isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ) ] (str Γ ⟧ᶜ) { (f , s) f }) substitutionFormula (var here , drop subs) f ⟧ᶠ) - (semanticSubstitution subs (`∀[ isSet× (str Δ ⟧ᶜ) (str B ⟧ˢ) ] (str Δ ⟧ᶜ) { (f , s) f }) f ⟧ᶠ)) - γ a a⊩substFormF - λ { r x@(δ , b) δ≡subsγ - subst - g (a r) f ⟧ᶠ (g , b)) - (sym δ≡subsγ) - (subst - form (a r) form (γ , b)) - (substitutionFormulaSound (var here , drop subs) f) - (a⊩substFormF r (γ , b) refl)) }) - λ γ a a⊩semanticSubsF - λ { r x@(γ' , b) γ'≡γ - subst - form (a r) form (γ' , b)) - (sym (substitutionFormulaSound (var here , drop subs) f)) - (subst - g (a r) f ⟧ᶠ (g , b)) - (cong subs ⟧ᴮ (sym γ'≡γ)) - (a⊩semanticSubsF r ( subs ⟧ᴮ γ , b) refl)) } - substitutionFormulaSound {Γ} {Δ} subs (rel R t) = - Predicate≡ - Γ ⟧ᶜ - (⋆_ (str Γ ⟧ᶜ) (str lookup R relSym ⟧ˢ) substitutionTerm subs t ⟧ᵗ R ⟧ʳ) - (semanticSubstitution subs (⋆_ (str Δ ⟧ᶜ) (str lookup R relSym ⟧ˢ) t ⟧ᵗ R ⟧ʳ)) - γ a a⊩substTR - subst transform a R ⟧ʳ (transform γ)) (substitutionTermSound subs t) a⊩substTR) - λ γ a a⊩semSubst - subst transform a R ⟧ʳ (transform γ)) (sym (substitutionTermSound subs t)) a⊩semSubst - - weakenFormulaMonotonic : {Γ B} (γ : Γ ⟧ᶜ ) (ϕ : Formula Γ) (a : A) (b : B ⟧ˢ ) a ϕ ⟧ᶠ γ a weakenFormula {S = B} ϕ ⟧ᶠ (γ , b) - weakenFormulaMonotonic {Γ} {B} γ ϕ a b = - hPropExt - ( ϕ ⟧ᶠ .isPropValued γ a) - ( weakenFormula ϕ ⟧ᶠ .isPropValued (γ , b) a) - a⊩ϕγ subst form a form (γ , b)) (sym (substitutionFormulaSound (drop id) ϕ)) a⊩ϕγ) - λ a⊩weakenϕγb subst form a form (γ , b)) (substitutionFormulaSound (drop id) ϕ) a⊩weakenϕγb -module Soundness - {n} - {relSym : Vec Sort n} - (isNonTrivial : s k ) - (⟦_⟧ʳ : RelationInterpretation relSym) where - open Relational relSym - open Interpretation relSym ⟦_⟧ʳ isNonTrivial - -- Acknowledgements : 1lab's "the internal logic of a regular hyperdoctrine" - infix 35 _⊨_ +private + Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} +RelationInterpretation : {n : } (Vec Sort n) Type _ +RelationInterpretation {n} relSym = (∀ i Predicate lookup i relSym ⟧ˢ ) + +⟦_⟧ᶜ : Context hSet ℓ' + [] ⟧ᶜ = Unit* , isSetUnit* + c x ⟧ᶜ = ( c ⟧ᶜ .fst × x ⟧ˢ .fst) , isSet× ( c ⟧ᶜ .snd) ( x ⟧ˢ .snd) + +⟦_⟧ⁿ : {Γ s} s Γ Γ ⟧ᶜ s ⟧ˢ +⟦_⟧ⁿ {.(_ s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧ +⟦_⟧ⁿ {.(_ _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) = s∈Γ ⟧ⁿ ⟦Γ⟧ + +⟦_⟧ᵗ : {Γ s} Term Γ s Γ ⟧ᶜ s ⟧ˢ +⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ = x ⟧ⁿ ⟦Γ⟧ +⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = ( t ⟧ᵗ ⟦Γ⟧) , ( t₁ ⟧ᵗ ⟦Γ⟧) +⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst ( t ⟧ᵗ ⟦Γ⟧) +⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd ( t ⟧ᵗ ⟦Γ⟧) +⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x ( t ⟧ᵗ ⟦Γ⟧) + +-- R for renamings and r for relations +⟦_⟧ᴿ : {Γ Δ} Renaming Γ Δ Γ ⟧ᶜ Δ ⟧ᶜ + id ⟧ᴿ ctx = ctx + drop ren ⟧ᴿ (ctx , _) = ren ⟧ᴿ ctx + keep ren ⟧ᴿ (ctx , s) = ( ren ⟧ᴿ ctx) , s + +-- B for suBstitution and s for sorts +⟦_⟧ᴮ : {Γ Δ} Substitution Γ Δ Γ ⟧ᶜ Δ ⟧ᶜ + id ⟧ᴮ ctx = ctx + t , sub ⟧ᴮ ctx = ( sub ⟧ᴮ ctx) , ( t ⟧ᵗ ctx) + drop sub ⟧ᴮ (ctx , s) = sub ⟧ᴮ ctx + +renamingVarSound : {Γ Δ s} (ren : Renaming Γ Δ) (v : s Δ) renamingVar ren v ⟧ⁿ v ⟧ⁿ ren ⟧ᴿ +renamingVarSound {Γ} {.Γ} {s} id v = refl +renamingVarSound {.(_ _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingVarSound ren v i ⟦Γ⟧ } +renamingVarSound {.(_ s)} {.(_ s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i ⟦s⟧ } +renamingVarSound {.(_ _)} {.(_ _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingVarSound ren v i ⟦Γ⟧ } + +renamingTermSound : {Γ Δ s} (ren : Renaming Γ Δ) (t : Term Δ s) renamingTerm ren t ⟧ᵗ t ⟧ᵗ ren ⟧ᴿ +renamingTermSound {Γ} {.Γ} {s} id t = refl +renamingTermSound {.(_ _)} {Δ} {s} (drop ren) (var x) = + funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingVarSound ren x i ⟦Γ⟧ } +renamingTermSound {.(_ _)} {Δ} {.(_ _)} r@(drop ren) (t `, t₁) = + funExt λ { (⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) } +renamingTermSound {.(_ _)} {Δ} {s} r@(drop ren) (π₁ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .fst } +renamingTermSound {.(_ _)} {Δ} {s} r@(drop ren) (π₂ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .snd } +renamingTermSound {.(_ _)} {Δ} {s} r@(drop ren) (fun f t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i f (renamingTermSound r t i x) } +renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (var v) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingVarSound r v i x } +renamingTermSound {.(_ _)} {.(_ _)} {.(_ _)} r@(keep ren) (t `, t₁) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) } +renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (π₁ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .fst } +renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (π₂ t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i renamingTermSound r t i x .snd } +renamingTermSound {.(_ _)} {.(_ _)} {s} r@(keep ren) (fun f t) = + funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i f (renamingTermSound r t i x) } + + +substitutionVarSound : {Γ Δ s} (subs : Substitution Γ Δ) (v : s Δ) substitutionVar subs v ⟧ᵗ v ⟧ⁿ subs ⟧ᴮ +substitutionVarSound {Γ} {.Γ} {s} id t = refl +substitutionVarSound {Γ} {.(_ s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧ refl +substitutionVarSound {Γ} {.(_ _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i substitutionVarSound subs t i ⟦Γ⟧ +substitutionVarSound {.(_ _)} {Δ} {s} r@(drop subs) t = + -- TODO : Fix unsolved constraints + funExt + λ { x@(⟦Γ⟧ , ⟦s⟧) + substitutionVar (drop subs) t ⟧ᵗ x + ≡[ i ]⟨ renamingTermSound (drop id) (substitutionVar subs t) i x + substitutionVar subs t ⟧ᵗ ( drop id ⟧ᴿ x) + ≡⟨ refl + substitutionVar subs t ⟧ᵗ ⟦Γ⟧ + ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ + t ⟧ⁿ ( subs ⟧ᴮ ⟦Γ⟧) + } + +substitutionTermSound : {Γ Δ s} (subs : Substitution Γ Δ) (t : Term Δ s) substitutionTerm subs t ⟧ᵗ t ⟧ᵗ subs ⟧ᴮ +substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x +substitutionTermSound {Γ} {Δ} {.(_ _)} subs (t `, t₁) = funExt λ x i (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x) +substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i substitutionTermSound subs t i x .fst +substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i substitutionTermSound subs t i x .snd +substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i f (substitutionTermSound subs t i x) + +semanticSubstitution : {Γ Δ} (subs : Substitution Γ Δ) Predicate Δ ⟧ᶜ Predicate Γ ⟧ᶜ +semanticSubstitution {Γ} {Δ} subs = ⋆_ (str Γ ⟧ᶜ) (str Δ ⟧ᶜ) subs ⟧ᴮ + +module Interpretation + {n : } + (relSym : Vec Sort n) + (⟦_⟧ʳ : RelationInterpretation relSym) (isNonTrivial : s k ) where + open Relational relSym + ⟦_⟧ᶠ : {Γ} Formula Γ Predicate Γ ⟧ᶜ + ⟦_⟧ᶠ {Γ} ⊤ᵗ = pre1 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial + ⟦_⟧ᶠ {Γ} ⊥ᵗ = pre0 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial + ⟦_⟧ᶠ {Γ} (form `∨ form₁) = _⊔_ Γ ⟧ᶜ form ⟧ᶠ form₁ ⟧ᶠ + ⟦_⟧ᶠ {Γ} (form `∧ form₁) = _⊓_ Γ ⟧ᶜ form ⟧ᶠ form₁ ⟧ᶠ + ⟦_⟧ᶠ {Γ} (form `→ form₁) = _⇒_ Γ ⟧ᶜ form ⟧ᶠ form₁ ⟧ᶠ + ⟦_⟧ᶠ {Γ} (`∃ {B = B} form) = `∃[_] (isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ)) (str Γ ⟧ᶜ) { (⟦Γ⟧ , ⟦B⟧) ⟦Γ⟧ }) form ⟧ᶠ + ⟦_⟧ᶠ {Γ} (`∀ {B = B} form) = `∀[_] (isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ)) (str Γ ⟧ᶜ) { (⟦Γ⟧ , ⟦B⟧) ⟦Γ⟧ }) form ⟧ᶠ + ⟦_⟧ᶠ {Γ} (rel R t) = ⋆_ (str Γ ⟧ᶜ) (str lookup R relSym ⟧ˢ) t ⟧ᵗ R ⟧ʳ + + -- Due to a shortcut in the soundness of negation termination checking fails + -- TODO : Fix + {-# TERMINATING #-} + substitutionFormulaSound : {Γ Δ} (subs : Substitution Γ Δ) (f : Formula Δ) substitutionFormula subs f ⟧ᶠ semanticSubstitution subs f ⟧ᶠ + substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ = + Predicate≡ + Γ ⟧ᶜ + (pre1 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial) + (semanticSubstitution subs (pre1 Δ ⟧ᶜ (str Δ ⟧ᶜ) isNonTrivial)) + γ a a⊩1γ tt*) + λ γ a a⊩1subsγ tt* + substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ = + Predicate≡ + Γ ⟧ᶜ + (pre0 Γ ⟧ᶜ (str Γ ⟧ᶜ) isNonTrivial) + (semanticSubstitution subs (pre0 Δ ⟧ᶜ (str Δ ⟧ᶜ) isNonTrivial)) + _ _ bot ⊥rec* bot) + λ _ _ bot bot + substitutionFormulaSound {Γ} {Δ} subs (f `∨ f₁) = + Predicate≡ + Γ ⟧ᶜ + (_⊔_ Γ ⟧ᶜ substitutionFormula subs f ⟧ᶠ substitutionFormula subs f₁ ⟧ᶠ) + (semanticSubstitution subs (_⊔_ Δ ⟧ᶜ f ⟧ᶠ f₁ ⟧ᶠ)) + γ a a⊩substFormFs + a⊩substFormFs >>= + λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) + inl (pr₁a≡k , subst form (pr₂ a) form γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁ + ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) + inr (pr₁a≡k' , subst form (pr₂ a) form γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ }) + λ γ a a⊩semanticSubsFs + a⊩semanticSubsFs >>= + λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) + inl (pr₁a≡k , (subst form (pr₂ a) form γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁ + ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) + inr (pr₁a≡k' , (subst form (pr₂ a) form γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ } + substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) = + Predicate≡ + Γ ⟧ᶜ + (_⊓_ Γ ⟧ᶜ substitutionFormula subs f ⟧ᶠ substitutionFormula subs f₁ ⟧ᶠ) + (semanticSubstitution subs (_⊓_ Δ ⟧ᶜ f ⟧ᶠ f₁ ⟧ᶠ)) + γ a a⊩substFormulaFs + (subst form (pr₁ a) form γ) (substitutionFormulaSound subs f) (a⊩substFormulaFs .fst)) , + (subst form (pr₂ a) form γ) (substitutionFormulaSound subs f₁) (a⊩substFormulaFs .snd))) + λ γ a a⊩semanticSubstFs + (subst form (pr₁ a) form γ) (sym (substitutionFormulaSound subs f)) (a⊩semanticSubstFs .fst)) , + (subst form (pr₂ a) form γ) (sym (substitutionFormulaSound subs f₁)) (a⊩semanticSubstFs .snd)) + substitutionFormulaSound {Γ} {Δ} subs (f `→ f₁) = + Predicate≡ + Γ ⟧ᶜ + (_⇒_ Γ ⟧ᶜ substitutionFormula subs f ⟧ᶠ substitutionFormula subs f₁ ⟧ᶠ) + (semanticSubstitution subs (_⇒_ Δ ⟧ᶜ f ⟧ᶠ f₁ ⟧ᶠ)) + γ a a⊩substFormulaFs + λ b b⊩semanticSubstFs + subst + form (a b) form γ) + (substitutionFormulaSound subs f₁) + (a⊩substFormulaFs + b + (subst + form b form γ) + (sym (substitutionFormulaSound subs f)) + b⊩semanticSubstFs))) + λ γ a a⊩semanticSubstFs + λ b b⊩substFormulaFs + subst + form (a b) form γ) + (sym (substitutionFormulaSound subs f₁)) + (a⊩semanticSubstFs + b + (subst + form b form γ) + (substitutionFormulaSound subs f) + b⊩substFormulaFs)) + substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) = + Predicate≡ + Γ ⟧ᶜ + (`∃[ isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ) ] (str Γ ⟧ᶜ) { (f , s) f }) substitutionFormula (var here , drop subs) f ⟧ᶠ) + (semanticSubstitution subs (`∃[ isSet× (str Δ ⟧ᶜ) (str B ⟧ˢ) ] (str Δ ⟧ᶜ) { (γ , b) γ }) f ⟧ᶠ)) + γ a a⊩πSubstFormulaF + a⊩πSubstFormulaF >>= + λ { ((γ' , b) , γ'≡γ , a⊩substFormFγ') + (( subs ⟧ᴮ γ') , b) , + ((cong subs ⟧ᴮ γ'≡γ) , + (subst + form a form (γ' , b)) + (substitutionFormulaSound (var here , drop subs) f) + a⊩substFormFγ' )) ∣₁ }) + λ γ a a⊩semanticSubstF + a⊩semanticSubstF >>= + λ (x@(δ , b) , δ≡subsγ , a⊩fx) + (γ , b) , + (refl , + (subst + form a form (γ , b)) + (sym (substitutionFormulaSound (var here , drop subs) f)) + (subst x a f ⟧ᶠ (x , b)) δ≡subsγ a⊩fx))) ∣₁ + substitutionFormulaSound {Γ} {Δ} subs (`∀ {B = B} f) = + Predicate≡ + Γ ⟧ᶜ + (`∀[ isSet× (str Γ ⟧ᶜ) (str B ⟧ˢ) ] (str Γ ⟧ᶜ) { (f , s) f }) substitutionFormula (var here , drop subs) f ⟧ᶠ) + (semanticSubstitution subs (`∀[ isSet× (str Δ ⟧ᶜ) (str B ⟧ˢ) ] (str Δ ⟧ᶜ) { (f , s) f }) f ⟧ᶠ)) + γ a a⊩substFormF + λ { r x@(δ , b) δ≡subsγ + subst + g (a r) f ⟧ᶠ (g , b)) + (sym δ≡subsγ) + (subst + form (a r) form (γ , b)) + (substitutionFormulaSound (var here , drop subs) f) + (a⊩substFormF r (γ , b) refl)) }) + λ γ a a⊩semanticSubsF + λ { r x@(γ' , b) γ'≡γ + subst + form (a r) form (γ' , b)) + (sym (substitutionFormulaSound (var here , drop subs) f)) + (subst + g (a r) f ⟧ᶠ (g , b)) + (cong subs ⟧ᴮ (sym γ'≡γ)) + (a⊩semanticSubsF r ( subs ⟧ᴮ γ , b) refl)) } + substitutionFormulaSound {Γ} {Δ} subs (rel R t) = + Predicate≡ + Γ ⟧ᶜ + (⋆_ (str Γ ⟧ᶜ) (str lookup R relSym ⟧ˢ) substitutionTerm subs t ⟧ᵗ R ⟧ʳ) + (semanticSubstitution subs (⋆_ (str Δ ⟧ᶜ) (str lookup R relSym ⟧ˢ) t ⟧ᵗ R ⟧ʳ)) + γ a a⊩substTR + subst transform a R ⟧ʳ (transform γ)) (substitutionTermSound subs t) a⊩substTR) + λ γ a a⊩semSubst + subst transform a R ⟧ʳ (transform γ)) (sym (substitutionTermSound subs t)) a⊩semSubst + + weakenFormulaMonotonic : {Γ B} (γ : Γ ⟧ᶜ ) (ϕ : Formula Γ) (a : A) (b : B ⟧ˢ ) a ϕ ⟧ᶠ γ a weakenFormula {S = B} ϕ ⟧ᶠ (γ , b) + weakenFormulaMonotonic {Γ} {B} γ ϕ a b = + hPropExt + ( ϕ ⟧ᶠ .isPropValued γ a) + ( weakenFormula ϕ ⟧ᶠ .isPropValued (γ , b) a) + a⊩ϕγ subst form a form (γ , b)) (sym (substitutionFormulaSound (drop id) ϕ)) a⊩ϕγ) + λ a⊩weakenϕγb subst form a form (γ , b)) (substitutionFormulaSound (drop id) ϕ) a⊩weakenϕγb +module Soundness + {n} + {relSym : Vec Sort n} + (isNonTrivial : s k ) + (⟦_⟧ʳ : RelationInterpretation relSym) where + open Relational relSym + open Interpretation relSym ⟦_⟧ʳ isNonTrivial + -- Acknowledgements : 1lab's "the internal logic of a regular hyperdoctrine" + infix 35 _⊨_ - module PredProps = PredicateProperties + module PredProps = PredicateProperties - _⊨_ : {Γ} Formula Γ Formula Γ Type (ℓ-max (ℓ-max ℓ'') ℓ') - _⊨_ {Γ} ϕ ψ = ϕ ⟧ᶠ ψ ⟧ᶠ where open PredProps Γ ⟧ᶜ - - entails = _⊨_ - - private - variable - Γ Δ : Context - ϕ ψ θ : Formula Γ - χ μ ν : Formula Δ - - cut : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ ψ θ ϕ θ - cut {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤ ϕ ⟧ᶠ ψ ⟧ᶠ θ ⟧ᶠ ϕ⊨ψ ψ⊨θ where open PredProps Γ ⟧ᶜ - - substitutionEntailment : {Γ Δ} (subs : Substitution Γ Δ) {ϕ ψ : Formula Δ} ϕ ψ substitutionFormula subs ϕ substitutionFormula subs ψ - substitutionEntailment {Γ} {Δ} subs {ϕ} {ψ} ϕ⊨ψ = - subst2 - ϕ' ψ' ϕ' ≤Γ ψ') - (sym (substitutionFormulaSound subs ϕ)) - (sym (substitutionFormulaSound subs ψ)) - (ϕ⊨ψ >>= - λ { (a , a⊩ϕ≤ψ) - a , γ b b⊩ϕsubsγ a⊩ϕ≤ψ ( subs ⟧ᴮ γ) b b⊩ϕsubsγ) ∣₁ }) where - open PredProps {ℓ'' = ℓ''} Γ ⟧ᶜ renaming (_≤_ to _≤Γ_) - open PredProps {ℓ'' = ℓ''} Δ ⟧ᶜ renaming (_≤_ to _≤Δ_) - - `∧intro : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ entails ϕ θ entails ϕ (ψ `∧ θ) - `∧intro {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ϕ⊨θ = - do - (a , a⊩ϕ⊨ψ) ϕ⊨ψ - (b , b⊩ϕ⊨θ) ϕ⊨θ - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` a ̇ # fzero) ̇ (` b ̇ # fzero) - return - (λ* prover , - λ γ r r⊩ϕγ - let - proofEq : λ* prover r pair (a r) (b r) - proofEq = λ*ComputationRule prover (r []) - - pr₁proofEq : pr₁ (λ* prover r) a r - pr₁proofEq = - pr₁ (λ* prover r) - ≡⟨ cong x pr₁ x) proofEq - pr₁ (pair (a r) (b r)) - ≡⟨ pr₁pxy≡x _ _ - a r - - - pr₂proofEq : pr₂ (λ* prover r) b r - pr₂proofEq = - pr₂ (λ* prover r) - ≡⟨ cong x pr₂ x) proofEq - pr₂ (pair (a r) (b r)) - ≡⟨ pr₂pxy≡y _ _ - b r - - in - subst r r ψ ⟧ᶠ γ) (sym pr₁proofEq) (a⊩ϕ⊨ψ γ r r⊩ϕγ) , - subst r r θ ⟧ᶠ γ) (sym pr₂proofEq) (b⊩ϕ⊨θ γ r r⊩ϕγ)) - - `∧elim1 : {Γ} {ϕ ψ θ : Formula Γ} ϕ (ψ `∧ θ) ϕ ψ - `∧elim1 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ = - do - (a , a⊩ϕ⊨ψ∧θ) ϕ⊨ψ∧θ - let - prover : ApplStrTerm as 1 - prover = ` pr₁ ̇ (` a ̇ # fzero) - return - (λ* prover , - λ γ b b⊩ϕγ subst r r ψ ⟧ᶠ γ) (sym (λ*ComputationRule prover (b []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst)) + _⊨_ : {Γ} Formula Γ Formula Γ Type (ℓ-max (ℓ-max ℓ'') ℓ') + _⊨_ {Γ} ϕ ψ = ϕ ⟧ᶠ ψ ⟧ᶠ where open PredProps Γ ⟧ᶜ + + entails = _⊨_ + + holdsInTripos : {Γ} Formula Γ Type (ℓ-max (ℓ-max ℓ'') ℓ') + holdsInTripos {Γ} form = ⊤ᵗ form + + private + variable + Γ Δ : Context + ϕ ψ θ : Formula Γ + χ μ ν : Formula Δ + + cut : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ ψ θ ϕ θ + cut {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤ ϕ ⟧ᶠ ψ ⟧ᶠ θ ⟧ᶠ ϕ⊨ψ ψ⊨θ where open PredProps Γ ⟧ᶜ + + substitutionEntailment : {Γ Δ} (subs : Substitution Γ Δ) {ϕ ψ : Formula Δ} ϕ ψ substitutionFormula subs ϕ substitutionFormula subs ψ + substitutionEntailment {Γ} {Δ} subs {ϕ} {ψ} ϕ⊨ψ = + subst2 + ϕ' ψ' ϕ' ≤Γ ψ') + (sym (substitutionFormulaSound subs ϕ)) + (sym (substitutionFormulaSound subs ψ)) + (ϕ⊨ψ >>= + λ { (a , a⊩ϕ≤ψ) + a , γ b b⊩ϕsubsγ a⊩ϕ≤ψ ( subs ⟧ᴮ γ) b b⊩ϕsubsγ) ∣₁ }) where + open PredProps {ℓ'' = ℓ''} Γ ⟧ᶜ renaming (_≤_ to _≤Γ_) + open PredProps {ℓ'' = ℓ''} Δ ⟧ᶜ renaming (_≤_ to _≤Δ_) + + `∧intro : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ entails ϕ θ entails ϕ (ψ `∧ θ) + `∧intro {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ϕ⊨θ = + do + (a , a⊩ϕ⊨ψ) ϕ⊨ψ + (b , b⊩ϕ⊨θ) ϕ⊨θ + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` a ̇ # fzero) ̇ (` b ̇ # fzero) + return + (λ* prover , + λ γ r r⊩ϕγ + let + proofEq : λ* prover r pair (a r) (b r) + proofEq = λ*ComputationRule prover (r []) + + pr₁proofEq : pr₁ (λ* prover r) a r + pr₁proofEq = + pr₁ (λ* prover r) + ≡⟨ cong x pr₁ x) proofEq + pr₁ (pair (a r) (b r)) + ≡⟨ pr₁pxy≡x _ _ + a r + + + pr₂proofEq : pr₂ (λ* prover r) b r + pr₂proofEq = + pr₂ (λ* prover r) + ≡⟨ cong x pr₂ x) proofEq + pr₂ (pair (a r) (b r)) + ≡⟨ pr₂pxy≡y _ _ + b r + + in + subst r r ψ ⟧ᶠ γ) (sym pr₁proofEq) (a⊩ϕ⊨ψ γ r r⊩ϕγ) , + subst r r θ ⟧ᶠ γ) (sym pr₂proofEq) (b⊩ϕ⊨θ γ r r⊩ϕγ)) + + `∧elim1 : {Γ} {ϕ ψ θ : Formula Γ} ϕ (ψ `∧ θ) ϕ ψ + `∧elim1 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ = + do + (a , a⊩ϕ⊨ψ∧θ) ϕ⊨ψ∧θ + let + prover : ApplStrTerm as 1 + prover = ` pr₁ ̇ (` a ̇ # fzero) + return + (λ* prover , + λ γ b b⊩ϕγ subst r r ψ ⟧ᶠ γ) (sym (λ*ComputationRule prover (b []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst)) - `∧elim2 : {Γ} {ϕ ψ θ : Formula Γ} ϕ (ψ `∧ θ) ϕ θ - `∧elim2 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ = - do - (a , a⊩ϕ⊨ψ∧θ) ϕ⊨ψ∧θ - let - prover : ApplStrTerm as 1 - prover = ` pr₂ ̇ (` a ̇ # fzero) - return - (λ* prover , - λ γ b b⊩ϕγ subst r r θ ⟧ᶠ γ) (sym (λ*ComputationRule prover (b []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd)) - - `∃intro : {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ B)} {t : Term Γ B} ϕ substitutionFormula (t , id) ψ ϕ `∃ ψ - `∃intro {Γ} {ϕ} {B} {ψ} {t} ϕ⊨ψ[t/x] = - do - (a , a⊩ϕ⊨ψ[t/x]) ϕ⊨ψ[t/x] - return - (a , γ b b⊩ϕγ (γ , ( t ⟧ᵗ γ)) , - (refl , (subst form (a b) form γ) (substitutionFormulaSound (t , id) ψ) (a⊩ϕ⊨ψ[t/x] γ b b⊩ϕγ))) ∣₁)) - - `∃elim : {Γ} {ϕ θ : Formula Γ} {B} {ψ : Formula (Γ B)} ϕ `∃ ψ (weakenFormula ϕ `∧ ψ) weakenFormula θ ϕ θ - `∃elim {Γ} {ϕ} {θ} {B} {ψ} ϕ⊨∃ψ ϕ∧ψ⊨θ = - do - (a , a⊩ϕ⊨∃ψ) ϕ⊨∃ψ - (b , b⊩ϕ∧ψ⊨θ) ϕ∧ψ⊨θ - let - prover : ApplStrTerm as 1 - prover = ` b ̇ (` pair ̇ # fzero ̇ (` a ̇ # fzero)) - return - (λ* prover , - γ c c⊩ϕγ - subst - r r θ ⟧ᶠ γ) - (sym (λ*ComputationRule prover (c []))) - (transport - (propTruncIdempotent ( θ ⟧ᶠ .isPropValued γ (b (pair c (a c))))) - (a⊩ϕ⊨∃ψ γ c c⊩ϕγ >>= - λ { (x@(γ' , b') , (γ'≡γ , a⨾c⊩ψx)) - transport - (sym - (weakenFormulaMonotonic γ θ (b (pair c (a c))) b')) - (b⊩ϕ∧ψ⊨θ - (γ , b') - (pair c (a c)) - (subst - r r weakenFormula ϕ ⟧ᶠ (γ , b')) - (sym (pr₁pxy≡x _ _)) - (transport - (weakenFormulaMonotonic γ ϕ c b') c⊩ϕγ) , - subst r r ψ ⟧ᶠ (γ , b')) (sym (pr₂pxy≡y _ _)) (subst g (a c) ψ ⟧ᶠ (g , b')) γ'≡γ a⨾c⊩ψx)) ) ∣₁ })))) - - `∀intro : {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ B)} weakenFormula ϕ ψ ϕ `∀ ψ - `∀intro {Γ} {ϕ} {B} {ψ} ϕ⊨ψ = - do - (a , a⊩ϕ⊨ψ) ϕ⊨ψ - let - prover : ApplStrTerm as 2 - prover = ` a ̇ # fzero - return - (λ* prover , - γ b b⊩ϕ λ { c x@(γ' , b') γ'≡γ - subst - r r ψ ⟧ᶠ (γ' , b')) - (sym (λ*ComputationRule prover (b c []))) - (a⊩ϕ⊨ψ - (γ' , b') - b - (transport (weakenFormulaMonotonic γ' ϕ b b') (subst g b ϕ ⟧ᶠ g) (sym γ'≡γ) b⊩ϕ))) })) - - `∀elim : {Γ} {B} {ϕ : Formula Γ} {ψ : Formula (Γ B)} ϕ `∀ ψ (t : Term Γ B) ϕ substitutionFormula (t , id) ψ - `∀elim {Γ} {B} {ϕ} {ψ} ϕ⊨∀ψ t = - do - (a , a⊩ϕ⊨∀ψ) ϕ⊨∀ψ - let - prover : ApplStrTerm as 1 - prover = ` a ̇ # fzero ̇ ` k - return - (λ* prover , - γ b b⊩ϕγ - subst - form (λ* prover b) form γ) - (sym (substitutionFormulaSound (t , id) ψ)) - (subst - r r ψ ⟧ᶠ (γ , t ⟧ᵗ γ)) - (sym (λ*ComputationRule prover (b []))) - (a⊩ϕ⊨∀ψ γ b b⊩ϕγ k (γ , t ⟧ᵗ γ) refl)))) - - `→intro : {Γ} {ϕ ψ θ : Formula Γ} (ϕ `∧ ψ) θ ϕ (ψ `→ θ) - `→intro {Γ} {ϕ} {ψ} {θ} ϕ∧ψ⊨θ = a⊓b≤c→a≤b⇒c Γ ⟧ᶜ (str Γ ⟧ᶜ) ϕ ⟧ᶠ ψ ⟧ᶠ θ ⟧ᶠ ϕ∧ψ⊨θ - - `→elim : {Γ} {ϕ ψ θ : Formula Γ} ϕ (ψ `→ θ) ϕ ψ ϕ θ - `→elim {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ→θ ϕ⊨ψ = - do - (a , a⊩ϕ⊨ψ→θ) ϕ⊨ψ→θ - (b , b⊩ϕ⊨ψ) ϕ⊨ψ - let - prover : ApplStrTerm as 1 - prover = ` a ̇ (# fzero) ̇ (` b ̇ # fzero) - return - (λ* prover , - γ c c⊩ϕγ - subst - r r θ ⟧ᶠ γ) - (sym (λ*ComputationRule prover (c []))) - (a⊩ϕ⊨ψ→θ γ c c⊩ϕγ (b c) (b⊩ϕ⊨ψ γ c c⊩ϕγ)))) - - `∨introR : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ ϕ (ψ `∨ θ) - `∨introR {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ = - do - (a , a⊩ϕ⊨ψ) ϕ⊨ψ - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ ` true ̇ (` a ̇ # fzero) - return - ((λ* prover) , - γ b b⊩ϕγ - let - pr₁proofEq : pr₁ (λ* prover b) true - pr₁proofEq = - pr₁ (λ* prover b) - ≡⟨ cong x pr₁ x) (λ*ComputationRule prover (b [])) - pr₁ (pair true (a b)) - ≡⟨ pr₁pxy≡x _ _ - true - - - pr₂proofEq : pr₂ (λ* prover b) a b - pr₂proofEq = - pr₂ (λ* prover b) - ≡⟨ cong x pr₂ x) (λ*ComputationRule prover (b [])) - pr₂ (pair true (a b)) - ≡⟨ pr₂pxy≡y _ _ - a b - - in inl (pr₁proofEq , subst r r ψ ⟧ᶠ γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁)) - - `∨introL : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ ϕ (θ `∨ ψ) - `∨introL {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ = - do - (a , a⊩ϕ⊨ψ) ϕ⊨ψ - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ ` false ̇ (` a ̇ # fzero) - return - ((λ* prover) , - γ b b⊩ϕγ - let - pr₁proofEq : pr₁ (λ* prover b) false - pr₁proofEq = - pr₁ (λ* prover b) - ≡⟨ cong x pr₁ x) (λ*ComputationRule prover (b [])) - pr₁ (pair false (a b)) - ≡⟨ pr₁pxy≡x _ _ - false - - - pr₂proofEq : pr₂ (λ* prover b) a b - pr₂proofEq = - pr₂ (λ* prover b) - ≡⟨ cong x pr₂ x) (λ*ComputationRule prover (b [])) - pr₂ (pair false (a b)) - ≡⟨ pr₂pxy≡y _ _ - a b - - in inr (pr₁proofEq , subst r r ψ ⟧ᶠ γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁)) - - -- Pretty sure this is code duplication - `if_then_else_ : {as n} ApplStrTerm as n ApplStrTerm as n ApplStrTerm as n ApplStrTerm as n - `if a then b else c = ` Id ̇ a ̇ b ̇ c - - `∨elim : {Γ} {ϕ ψ θ χ : Formula Γ} (ϕ `∧ ψ) χ (ϕ `∧ θ) χ (ϕ `∧ (ψ `∨ θ)) χ - `∨elim {Γ} {ϕ} {ψ} {θ} {χ} ϕ∧ψ⊨χ ϕ∧θ⊨χ = - do - (a , a⊩ϕ∧ψ⊨χ) ϕ∧ψ⊨χ - (b , b⊩ϕ∧θ⊨χ) ϕ∧θ⊨χ - let - prover : ApplStrTerm as 1 - prover = - (`if ` pr₁ ̇ (` pr₂ ̇ # fzero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) - return - (λ* prover , - - { γ c (pr₁⨾c⊩ϕγ , pr₂⨾c⊩ψ∨θ) - transport - (propTruncIdempotent ( χ ⟧ᶠ .isPropValued γ (λ* prover c))) - (pr₂⨾c⊩ψ∨θ >>= - λ { (inl (pr₁⨾pr₂⨾c≡true , pr₂⨾pr₂⨾c⊩ψ)) - let - proofEq : λ* prover c a (pair (pr₁ c) (pr₂ (pr₂ c))) - proofEq = - λ* prover c - ≡⟨ λ*ComputationRule prover (c []) - (if (pr₁ (pr₂ c)) then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) - ≡⟨ cong x (if x then a else b) (pair (pr₁ c) (pr₂ (pr₂ c)))) pr₁⨾pr₂⨾c≡true - (if true then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) - ≡⟨ cong x x (pair (pr₁ c) (pr₂ (pr₂ c)))) (ifTrueThen a b) - a (pair (pr₁ c) (pr₂ (pr₂ c))) - - in - subst - r r χ ⟧ᶠ γ) - (sym proofEq) - (a⊩ϕ∧ψ⊨χ - γ - (pair (pr₁ c) (pr₂ (pr₂ c))) - (( - subst - r r ϕ ⟧ᶠ γ) - (sym (pr₁pxy≡x _ _)) - pr₁⨾c⊩ϕγ) , - subst - r r ψ ⟧ᶠ γ) - (sym (pr₂pxy≡y _ _)) - pr₂⨾pr₂⨾c⊩ψ)) ∣₁ - ; (inr (pr₁pr₂⨾c≡false , pr₂⨾pr₂⨾c⊩θ)) - let - proofEq : λ* prover c b (pair (pr₁ c) (pr₂ (pr₂ c))) - proofEq = - λ* prover c - ≡⟨ λ*ComputationRule prover (c []) - (if (pr₁ (pr₂ c)) then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) - ≡⟨ cong x (if x then a else b) (pair (pr₁ c) (pr₂ (pr₂ c)))) pr₁pr₂⨾c≡false - (if false then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) - ≡⟨ cong x x (pair (pr₁ c) (pr₂ (pr₂ c)))) (ifFalseElse a b) - b (pair (pr₁ c) (pr₂ (pr₂ c))) - - in - subst - r r χ ⟧ᶠ γ) - (sym proofEq) - (b⊩ϕ∧θ⊨χ - γ - (pair (pr₁ c) (pr₂ (pr₂ c))) - ((subst r r ϕ ⟧ᶠ γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) , - (subst r r θ ⟧ᶠ γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) })) + `∧elim2 : {Γ} {ϕ ψ θ : Formula Γ} ϕ (ψ `∧ θ) ϕ θ + `∧elim2 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ = + do + (a , a⊩ϕ⊨ψ∧θ) ϕ⊨ψ∧θ + let + prover : ApplStrTerm as 1 + prover = ` pr₂ ̇ (` a ̇ # fzero) + return + (λ* prover , + λ γ b b⊩ϕγ subst r r θ ⟧ᶠ γ) (sym (λ*ComputationRule prover (b []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd)) + + `∃intro : {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ B)} {t : Term Γ B} ϕ substitutionFormula (t , id) ψ ϕ `∃ ψ + `∃intro {Γ} {ϕ} {B} {ψ} {t} ϕ⊨ψ[t/x] = + do + (a , a⊩ϕ⊨ψ[t/x]) ϕ⊨ψ[t/x] + return + (a , γ b b⊩ϕγ (γ , ( t ⟧ᵗ γ)) , + (refl , (subst form (a b) form γ) (substitutionFormulaSound (t , id) ψ) (a⊩ϕ⊨ψ[t/x] γ b b⊩ϕγ))) ∣₁)) + + `∃elim : {Γ} {ϕ θ : Formula Γ} {B} {ψ : Formula (Γ B)} ϕ `∃ ψ (weakenFormula ϕ `∧ ψ) weakenFormula θ ϕ θ + `∃elim {Γ} {ϕ} {θ} {B} {ψ} ϕ⊨∃ψ ϕ∧ψ⊨θ = + do + (a , a⊩ϕ⊨∃ψ) ϕ⊨∃ψ + (b , b⊩ϕ∧ψ⊨θ) ϕ∧ψ⊨θ + let + prover : ApplStrTerm as 1 + prover = ` b ̇ (` pair ̇ # fzero ̇ (` a ̇ # fzero)) + return + (λ* prover , + γ c c⊩ϕγ + subst + r r θ ⟧ᶠ γ) + (sym (λ*ComputationRule prover (c []))) + (transport + (propTruncIdempotent ( θ ⟧ᶠ .isPropValued γ (b (pair c (a c))))) + (a⊩ϕ⊨∃ψ γ c c⊩ϕγ >>= + λ { (x@(γ' , b') , (γ'≡γ , a⨾c⊩ψx)) + transport + (sym + (weakenFormulaMonotonic γ θ (b (pair c (a c))) b')) + (b⊩ϕ∧ψ⊨θ + (γ , b') + (pair c (a c)) + (subst + r r weakenFormula ϕ ⟧ᶠ (γ , b')) + (sym (pr₁pxy≡x _ _)) + (transport + (weakenFormulaMonotonic γ ϕ c b') c⊩ϕγ) , + subst r r ψ ⟧ᶠ (γ , b')) (sym (pr₂pxy≡y _ _)) (subst g (a c) ψ ⟧ᶠ (g , b')) γ'≡γ a⨾c⊩ψx)) ) ∣₁ })))) + + `∀intro : {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ B)} weakenFormula ϕ ψ ϕ `∀ ψ + `∀intro {Γ} {ϕ} {B} {ψ} ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ) ϕ⊨ψ + let + prover : ApplStrTerm as 2 + prover = ` a ̇ # fzero + return + (λ* prover , + γ b b⊩ϕ λ { c x@(γ' , b') γ'≡γ + subst + r r ψ ⟧ᶠ (γ' , b')) + (sym (λ*ComputationRule prover (b c []))) + (a⊩ϕ⊨ψ + (γ' , b') + b + (transport (weakenFormulaMonotonic γ' ϕ b b') (subst g b ϕ ⟧ᶠ g) (sym γ'≡γ) b⊩ϕ))) })) + + `∀elim : {Γ} {B} {ϕ : Formula Γ} {ψ : Formula (Γ B)} ϕ `∀ ψ (t : Term Γ B) ϕ substitutionFormula (t , id) ψ + `∀elim {Γ} {B} {ϕ} {ψ} ϕ⊨∀ψ t = + do + (a , a⊩ϕ⊨∀ψ) ϕ⊨∀ψ + let + prover : ApplStrTerm as 1 + prover = ` a ̇ # fzero ̇ ` k + return + (λ* prover , + γ b b⊩ϕγ + subst + form (λ* prover b) form γ) + (sym (substitutionFormulaSound (t , id) ψ)) + (subst + r r ψ ⟧ᶠ (γ , t ⟧ᵗ γ)) + (sym (λ*ComputationRule prover (b []))) + (a⊩ϕ⊨∀ψ γ b b⊩ϕγ k (γ , t ⟧ᵗ γ) refl)))) + + `→intro : {Γ} {ϕ ψ θ : Formula Γ} (ϕ `∧ ψ) θ ϕ (ψ `→ θ) + `→intro {Γ} {ϕ} {ψ} {θ} ϕ∧ψ⊨θ = a⊓b≤c→a≤b⇒c Γ ⟧ᶜ (str Γ ⟧ᶜ) ϕ ⟧ᶠ ψ ⟧ᶠ θ ⟧ᶠ ϕ∧ψ⊨θ + + `→elim : {Γ} {ϕ ψ θ : Formula Γ} ϕ (ψ `→ θ) ϕ ψ ϕ θ + `→elim {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ→θ ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ→θ) ϕ⊨ψ→θ + (b , b⊩ϕ⊨ψ) ϕ⊨ψ + let + prover : ApplStrTerm as 1 + prover = ` a ̇ (# fzero) ̇ (` b ̇ # fzero) + return + (λ* prover , + γ c c⊩ϕγ + subst + r r θ ⟧ᶠ γ) + (sym (λ*ComputationRule prover (c []))) + (a⊩ϕ⊨ψ→θ γ c c⊩ϕγ (b c) (b⊩ϕ⊨ψ γ c c⊩ϕγ)))) + + `∨introR : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ ϕ (ψ `∨ θ) + `∨introR {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ) ϕ⊨ψ + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ ` true ̇ (` a ̇ # fzero) + return + ((λ* prover) , + γ b b⊩ϕγ + let + pr₁proofEq : pr₁ (λ* prover b) true + pr₁proofEq = + pr₁ (λ* prover b) + ≡⟨ cong x pr₁ x) (λ*ComputationRule prover (b [])) + pr₁ (pair true (a b)) + ≡⟨ pr₁pxy≡x _ _ + true + + + pr₂proofEq : pr₂ (λ* prover b) a b + pr₂proofEq = + pr₂ (λ* prover b) + ≡⟨ cong x pr₂ x) (λ*ComputationRule prover (b [])) + pr₂ (pair true (a b)) + ≡⟨ pr₂pxy≡y _ _ + a b + + in inl (pr₁proofEq , subst r r ψ ⟧ᶠ γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁)) + + `∨introL : {Γ} {ϕ ψ θ : Formula Γ} ϕ ψ ϕ (θ `∨ ψ) + `∨introL {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ = + do + (a , a⊩ϕ⊨ψ) ϕ⊨ψ + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ ` false ̇ (` a ̇ # fzero) + return + ((λ* prover) , + γ b b⊩ϕγ + let + pr₁proofEq : pr₁ (λ* prover b) false + pr₁proofEq = + pr₁ (λ* prover b) + ≡⟨ cong x pr₁ x) (λ*ComputationRule prover (b [])) + pr₁ (pair false (a b)) + ≡⟨ pr₁pxy≡x _ _ + false + + + pr₂proofEq : pr₂ (λ* prover b) a b + pr₂proofEq = + pr₂ (λ* prover b) + ≡⟨ cong x pr₂ x) (λ*ComputationRule prover (b [])) + pr₂ (pair false (a b)) + ≡⟨ pr₂pxy≡y _ _ + a b + + in inr (pr₁proofEq , subst r r ψ ⟧ᶠ γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁)) + + -- Pretty sure this is code duplication + `if_then_else_ : {as n} ApplStrTerm as n ApplStrTerm as n ApplStrTerm as n ApplStrTerm as n + `if a then b else c = ` Id ̇ a ̇ b ̇ c + + `∨elim : {Γ} {ϕ ψ θ χ : Formula Γ} (ϕ `∧ ψ) χ (ϕ `∧ θ) χ (ϕ `∧ (ψ `∨ θ)) χ + `∨elim {Γ} {ϕ} {ψ} {θ} {χ} ϕ∧ψ⊨χ ϕ∧θ⊨χ = + do + (a , a⊩ϕ∧ψ⊨χ) ϕ∧ψ⊨χ + (b , b⊩ϕ∧θ⊨χ) ϕ∧θ⊨χ + let + prover : ApplStrTerm as 1 + prover = + (`if ` pr₁ ̇ (` pr₂ ̇ # fzero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) + return + (λ* prover , + + { γ c (pr₁⨾c⊩ϕγ , pr₂⨾c⊩ψ∨θ) + transport + (propTruncIdempotent ( χ ⟧ᶠ .isPropValued γ (λ* prover c))) + (pr₂⨾c⊩ψ∨θ >>= + λ { (inl (pr₁⨾pr₂⨾c≡true , pr₂⨾pr₂⨾c⊩ψ)) + let + proofEq : λ* prover c a (pair (pr₁ c) (pr₂ (pr₂ c))) + proofEq = + λ* prover c + ≡⟨ λ*ComputationRule prover (c []) + (if (pr₁ (pr₂ c)) then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) + ≡⟨ cong x (if x then a else b) (pair (pr₁ c) (pr₂ (pr₂ c)))) pr₁⨾pr₂⨾c≡true + (if true then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) + ≡⟨ cong x x (pair (pr₁ c) (pr₂ (pr₂ c)))) (ifTrueThen a b) + a (pair (pr₁ c) (pr₂ (pr₂ c))) + + in + subst + r r χ ⟧ᶠ γ) + (sym proofEq) + (a⊩ϕ∧ψ⊨χ + γ + (pair (pr₁ c) (pr₂ (pr₂ c))) + (( + subst + r r ϕ ⟧ᶠ γ) + (sym (pr₁pxy≡x _ _)) + pr₁⨾c⊩ϕγ) , + subst + r r ψ ⟧ᶠ γ) + (sym (pr₂pxy≡y _ _)) + pr₂⨾pr₂⨾c⊩ψ)) ∣₁ + ; (inr (pr₁pr₂⨾c≡false , pr₂⨾pr₂⨾c⊩θ)) + let + proofEq : λ* prover c b (pair (pr₁ c) (pr₂ (pr₂ c))) + proofEq = + λ* prover c + ≡⟨ λ*ComputationRule prover (c []) + (if (pr₁ (pr₂ c)) then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) + ≡⟨ cong x (if x then a else b) (pair (pr₁ c) (pr₂ (pr₂ c)))) pr₁pr₂⨾c≡false + (if false then a else b) (pair (pr₁ c) (pr₂ (pr₂ c))) + ≡⟨ cong x x (pair (pr₁ c) (pr₂ (pr₂ c)))) (ifFalseElse a b) + b (pair (pr₁ c) (pr₂ (pr₂ c))) + + in + subst + r r χ ⟧ᶠ γ) + (sym proofEq) + (b⊩ϕ∧θ⊨χ + γ + (pair (pr₁ c) (pr₂ (pr₂ c))) + ((subst r r ϕ ⟧ᶠ γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) , + (subst r r θ ⟧ᶠ γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) })) + + \ No newline at end of file diff --git a/docs/Realizability.Tripos.Logic.Syntax.html b/docs/Realizability.Tripos.Logic.Syntax.html index cbef8a1..218efb7 100644 --- a/docs/Realizability.Tripos.Logic.Syntax.html +++ b/docs/Realizability.Tripos.Logic.Syntax.html @@ -18,98 +18,99 @@ a ⟧ˢ = a a b ⟧ˢ = a ⟧ˢ .fst × b ⟧ˢ .fst , isSet× ( a ⟧ˢ .snd) ( b ⟧ˢ .snd) -data Context : Type (ℓ-suc ) where - [] : Context - _′_ : Context Sort Context - -data _∈_ : Sort Context Type (ℓ-suc ) where - here : {Γ A} A (Γ A) - there : {Γ A B} A Γ A (Γ B) - -data Term : Context Sort Type (ℓ-suc ) where - var : {Γ A} A Γ Term Γ A - _`,_ : {Γ A B} Term Γ A Term Γ B Term Γ (A B) - π₁ : {Γ A B} Term Γ (A B) Term Γ A - π₂ : {Γ A B} Term Γ (A B) Term Γ B - fun : {Γ A B} ( A ⟧ˢ .fst B ⟧ˢ .fst) Term Γ A Term Γ B - -data Renaming : Context Context Type (ℓ-suc ) where - id : {Γ} Renaming Γ Γ - drop : {Γ Δ s} Renaming Γ Δ Renaming (Γ s) Δ - keep : {Γ Δ s} Renaming Γ Δ Renaming (Γ s) (Δ s) - -data Substitution : Context Context Type (ℓ-suc ) where - id : {Γ} Substitution Γ Γ - _,_ : {Γ Δ s} (t : Term Γ s) Substitution Γ Δ Substitution Γ (Δ s) - drop : {Γ Δ s} Substitution Γ Δ Substitution (Γ s) Δ - -terminatingSubstitution : {Γ} Substitution Γ [] -terminatingSubstitution {[]} = id -terminatingSubstitution {Γ x} = drop terminatingSubstitution - -renamingCompose : {Γ Δ Θ} Renaming Γ Δ Renaming Δ Θ Renaming Γ Θ -renamingCompose {Γ} {.Γ} {Θ} id Δ→Θ = Δ→Θ -renamingCompose {.(_ _)} {Δ} {Θ} (drop Γ→Δ) Δ→Θ = drop (renamingCompose Γ→Δ Δ→Θ) -renamingCompose {.(_ _)} {.(_ _)} {.(_ _)} (keep Γ→Δ) id = keep Γ→Δ -renamingCompose {.(_ _)} {.(_ _)} {Θ} (keep Γ→Δ) (drop Δ→Θ) = drop (renamingCompose Γ→Δ Δ→Θ) -renamingCompose {.(_ _)} {.(_ _)} {.(_ _)} (keep Γ→Δ) (keep Δ→Θ) = keep (renamingCompose Γ→Δ Δ→Θ) - -renamingVar : {Γ Δ s} Renaming Γ Δ s Δ s Γ -renamingVar {Γ} {.Γ} {s} id s∈Δ = s∈Δ -renamingVar {.(_ _)} {Δ} {s} (drop ren) s∈Δ = there (renamingVar ren s∈Δ) -renamingVar {.(_ s)} {.(_ s)} {s} (keep ren) here = here -renamingVar {.(_ _)} {.(_ _)} {s} (keep ren) (there s∈Δ) = there (renamingVar ren s∈Δ) - -renamingTerm : {Γ Δ s} Renaming Γ Δ Term Δ s Term Γ s -renamingTerm {Γ} {.Γ} {s} id term = term -renamingTerm {.(_ _)} {Δ} {s} (drop ren) (var x) = var (renamingVar (drop ren) x) -renamingTerm {.(_ _)} {Δ} {.(_ _)} (drop ren) (term `, term₁) = renamingTerm (drop ren) term `, renamingTerm (drop ren) term₁ -renamingTerm {.(_ _)} {Δ} {s} (drop ren) (π₁ term) = π₁ (renamingTerm (drop ren) term) -renamingTerm {.(_ _)} {Δ} {s} (drop ren) (π₂ term) = π₂ (renamingTerm (drop ren) term) -renamingTerm {.(_ _)} {Δ} {s} (drop ren) (fun f term) = fun f (renamingTerm (drop ren) term) -renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (var x) = var (renamingVar (keep ren) x) -renamingTerm {.(_ _)} {.(_ _)} {.(_ _)} (keep ren) (term `, term₁) = renamingTerm (keep ren) term `, renamingTerm (keep ren) term₁ -renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (π₁ term) = π₁ (renamingTerm (keep ren) term) -renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (π₂ term) = π₂ (renamingTerm (keep ren) term) -renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (fun f term) = fun f (renamingTerm (keep ren) term) - -substitutionVar : {Γ Δ s} Substitution Γ Δ s Δ Term Γ s -substitutionVar {Γ} {.Γ} {s} id s∈Δ = var s∈Δ -substitutionVar {Γ} {.(_ s)} {s} (t , subs) here = t -substitutionVar {Γ} {.(_ _)} {s} (t , subs) (there s∈Δ) = substitutionVar subs s∈Δ -substitutionVar {.(_ _)} {Δ} {s} (drop subs) s∈Δ = renamingTerm (drop id) (substitutionVar subs s∈Δ) - -substitutionTerm : {Γ Δ s} Substitution Γ Δ Term Δ s Term Γ s -substitutionTerm {Γ} {Δ} {s} subs (var x) = substitutionVar subs x -substitutionTerm {Γ} {Δ} {.(_ _)} subs (t `, t₁) = substitutionTerm subs t `, substitutionTerm subs t₁ -substitutionTerm {Γ} {Δ} {s} subs (π₁ t) = π₁ (substitutionTerm subs t) -substitutionTerm {Γ} {Δ} {s} subs (π₂ t) = π₂ (substitutionTerm subs t) -substitutionTerm {Γ} {Δ} {s} subs (fun x t) = fun x (substitutionTerm subs t) - -module Relational {n} (relSym : Vec Sort n) where - - data Formula : Context Type (ℓ-suc ) where - ⊤ᵗ : {Γ} Formula Γ - ⊥ᵗ : {Γ} Formula Γ - _`∨_ : {Γ} Formula Γ Formula Γ Formula Γ - _`∧_ : {Γ} Formula Γ Formula Γ Formula Γ - _`→_ : {Γ} Formula Γ Formula Γ Formula Γ - `¬_ : {Γ} Formula Γ Formula Γ - `∃ : {Γ B} Formula (Γ B) Formula Γ - `∀ : {Γ B} Formula (Γ B) Formula Γ - rel : {Γ} (k : Fin n) Term Γ (lookup k relSym) Formula Γ - - substitutionFormula : {Γ Δ} Substitution Γ Δ Formula Δ Formula Γ - substitutionFormula {Γ} {Δ} subs ⊤ᵗ = ⊤ᵗ - substitutionFormula {Γ} {Δ} subs ⊥ᵗ = ⊥ᵗ - substitutionFormula {Γ} {Δ} subs (form `∨ form₁) = substitutionFormula subs form `∨ substitutionFormula subs form₁ - substitutionFormula {Γ} {Δ} subs (form `∧ form₁) = substitutionFormula subs form `∧ substitutionFormula subs form₁ - substitutionFormula {Γ} {Δ} subs (form `→ form₁) = substitutionFormula subs form `→ substitutionFormula subs form₁ - substitutionFormula {Γ} {Δ} subs ( form) = substitutionFormula subs form - substitutionFormula {Γ} {Δ} subs (`∃ form) = `∃ (substitutionFormula (var here , drop subs) form ) - substitutionFormula {Γ} {Δ} subs (`∀ form) = `∀ (substitutionFormula (var here , drop subs) form) - substitutionFormula {Γ} {Δ} subs (rel k x) = rel k (substitutionTerm subs x) - - weakenFormula : {Γ} {S} Formula Γ Formula (Γ S) - weakenFormula {Γ} {S} form = substitutionFormula (drop id) form +infixl 30 _′_ +data Context : Type (ℓ-suc ) where + [] : Context + _′_ : Context Sort Context + +data _∈_ : Sort Context Type (ℓ-suc ) where + here : {Γ A} A (Γ A) + there : {Γ A B} A Γ A (Γ B) + +data Term : Context Sort Type (ℓ-suc ) where + var : {Γ A} A Γ Term Γ A + _`,_ : {Γ A B} Term Γ A Term Γ B Term Γ (A B) + π₁ : {Γ A B} Term Γ (A B) Term Γ A + π₂ : {Γ A B} Term Γ (A B) Term Γ B + fun : {Γ A B} ( A ⟧ˢ .fst B ⟧ˢ .fst) Term Γ A Term Γ B + +data Renaming : Context Context Type (ℓ-suc ) where + id : {Γ} Renaming Γ Γ + drop : {Γ Δ s} Renaming Γ Δ Renaming (Γ s) Δ + keep : {Γ Δ s} Renaming Γ Δ Renaming (Γ s) (Δ s) + +data Substitution : Context Context Type (ℓ-suc ) where + id : {Γ} Substitution Γ Γ + _,_ : {Γ Δ s} (t : Term Γ s) Substitution Γ Δ Substitution Γ (Δ s) + drop : {Γ Δ s} Substitution Γ Δ Substitution (Γ s) Δ + +terminatingSubstitution : {Γ} Substitution Γ [] +terminatingSubstitution {[]} = id +terminatingSubstitution {Γ x} = drop terminatingSubstitution + +renamingCompose : {Γ Δ Θ} Renaming Γ Δ Renaming Δ Θ Renaming Γ Θ +renamingCompose {Γ} {.Γ} {Θ} id Δ→Θ = Δ→Θ +renamingCompose {.(_ _)} {Δ} {Θ} (drop Γ→Δ) Δ→Θ = drop (renamingCompose Γ→Δ Δ→Θ) +renamingCompose {.(_ _)} {.(_ _)} {.(_ _)} (keep Γ→Δ) id = keep Γ→Δ +renamingCompose {.(_ _)} {.(_ _)} {Θ} (keep Γ→Δ) (drop Δ→Θ) = drop (renamingCompose Γ→Δ Δ→Θ) +renamingCompose {.(_ _)} {.(_ _)} {.(_ _)} (keep Γ→Δ) (keep Δ→Θ) = keep (renamingCompose Γ→Δ Δ→Θ) + +renamingVar : {Γ Δ s} Renaming Γ Δ s Δ s Γ +renamingVar {Γ} {.Γ} {s} id s∈Δ = s∈Δ +renamingVar {.(_ _)} {Δ} {s} (drop ren) s∈Δ = there (renamingVar ren s∈Δ) +renamingVar {.(_ s)} {.(_ s)} {s} (keep ren) here = here +renamingVar {.(_ _)} {.(_ _)} {s} (keep ren) (there s∈Δ) = there (renamingVar ren s∈Δ) + +renamingTerm : {Γ Δ s} Renaming Γ Δ Term Δ s Term Γ s +renamingTerm {Γ} {.Γ} {s} id term = term +renamingTerm {.(_ _)} {Δ} {s} (drop ren) (var x) = var (renamingVar (drop ren) x) +renamingTerm {.(_ _)} {Δ} {.(_ _)} (drop ren) (term `, term₁) = renamingTerm (drop ren) term `, renamingTerm (drop ren) term₁ +renamingTerm {.(_ _)} {Δ} {s} (drop ren) (π₁ term) = π₁ (renamingTerm (drop ren) term) +renamingTerm {.(_ _)} {Δ} {s} (drop ren) (π₂ term) = π₂ (renamingTerm (drop ren) term) +renamingTerm {.(_ _)} {Δ} {s} (drop ren) (fun f term) = fun f (renamingTerm (drop ren) term) +renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (var x) = var (renamingVar (keep ren) x) +renamingTerm {.(_ _)} {.(_ _)} {.(_ _)} (keep ren) (term `, term₁) = renamingTerm (keep ren) term `, renamingTerm (keep ren) term₁ +renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (π₁ term) = π₁ (renamingTerm (keep ren) term) +renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (π₂ term) = π₂ (renamingTerm (keep ren) term) +renamingTerm {.(_ _)} {.(_ _)} {s} (keep ren) (fun f term) = fun f (renamingTerm (keep ren) term) + +substitutionVar : {Γ Δ s} Substitution Γ Δ s Δ Term Γ s +substitutionVar {Γ} {.Γ} {s} id s∈Δ = var s∈Δ +substitutionVar {Γ} {.(_ s)} {s} (t , subs) here = t +substitutionVar {Γ} {.(_ _)} {s} (t , subs) (there s∈Δ) = substitutionVar subs s∈Δ +substitutionVar {.(_ _)} {Δ} {s} (drop subs) s∈Δ = renamingTerm (drop id) (substitutionVar subs s∈Δ) + +substitutionTerm : {Γ Δ s} Substitution Γ Δ Term Δ s Term Γ s +substitutionTerm {Γ} {Δ} {s} subs (var x) = substitutionVar subs x +substitutionTerm {Γ} {Δ} {.(_ _)} subs (t `, t₁) = substitutionTerm subs t `, substitutionTerm subs t₁ +substitutionTerm {Γ} {Δ} {s} subs (π₁ t) = π₁ (substitutionTerm subs t) +substitutionTerm {Γ} {Δ} {s} subs (π₂ t) = π₂ (substitutionTerm subs t) +substitutionTerm {Γ} {Δ} {s} subs (fun x t) = fun x (substitutionTerm subs t) + +module Relational {n} (relSym : Vec Sort n) where + + data Formula : Context Type (ℓ-suc ) where + ⊤ᵗ : {Γ} Formula Γ + ⊥ᵗ : {Γ} Formula Γ + _`∨_ : {Γ} Formula Γ Formula Γ Formula Γ + _`∧_ : {Γ} Formula Γ Formula Γ Formula Γ + _`→_ : {Γ} Formula Γ Formula Γ Formula Γ + `∃ : {Γ B} Formula (Γ B) Formula Γ + `∀ : {Γ B} Formula (Γ B) Formula Γ + rel : {Γ} (k : Fin n) Term Γ (lookup k relSym) Formula Γ + + pattern f = f `→ ⊥ᵗ + + substitutionFormula : {Γ Δ} Substitution Γ Δ Formula Δ Formula Γ + substitutionFormula {Γ} {Δ} subs ⊤ᵗ = ⊤ᵗ + substitutionFormula {Γ} {Δ} subs ⊥ᵗ = ⊥ᵗ + substitutionFormula {Γ} {Δ} subs (form `∨ form₁) = substitutionFormula subs form `∨ substitutionFormula subs form₁ + substitutionFormula {Γ} {Δ} subs (form `∧ form₁) = substitutionFormula subs form `∧ substitutionFormula subs form₁ + substitutionFormula {Γ} {Δ} subs (form `→ form₁) = substitutionFormula subs form `→ substitutionFormula subs form₁ + substitutionFormula {Γ} {Δ} subs (`∃ form) = `∃ (substitutionFormula (var here , drop subs) form ) + substitutionFormula {Γ} {Δ} subs (`∀ form) = `∀ (substitutionFormula (var here , drop subs) form) + substitutionFormula {Γ} {Δ} subs (rel k x) = rel k (substitutionTerm subs x) + + weakenFormula : {Γ} {S} Formula Γ Formula (Γ S) + weakenFormula {Γ} {S} form = substitutionFormula (drop id) form \ No newline at end of file diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index a4203c5..9516a10 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -1,3 +1,4 @@ +{-# OPTIONS --allow-unsolved-metas #-} open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) open import Realizability.CombinatoryAlgebra From 60e4c402b19a497b1a81f33598b0e9fb85a214ea Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 12 Feb 2024 23:46:27 +0530 Subject: [PATCH 18/61] Define topos object and morphisms --- ...ealizability.Topos.FunctionalRelation.html | 753 +++++++++--------- docs/Realizability.Topos.Object.html | 109 +-- .../Topos/FunctionalRelation.agda | 649 ++++++++------- src/Realizability/Topos/Object.agda | 85 +- 4 files changed, 720 insertions(+), 876 deletions(-) diff --git a/docs/Realizability.Topos.FunctionalRelation.html b/docs/Realizability.Topos.FunctionalRelation.html index 9896e56..0fca2f8 100644 --- a/docs/Realizability.Topos.FunctionalRelation.html +++ b/docs/Realizability.Topos.FunctionalRelation.html @@ -1,384 +1,377 @@ Realizability.Topos.FunctionalRelation
{-# OPTIONS --allow-unsolved-metas #-}
 open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
-
-open import Realizability.CombinatoryAlgebra
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Structure
-open import Cubical.Foundations.HLevels
-open import Cubical.Data.Vec
-open import Cubical.Data.Nat
-open import Cubical.Data.FinData
-open import Cubical.Data.Fin hiding (Fin; _/_)
-open import Cubical.Data.Sigma
-open import Cubical.Data.Empty
-open import Cubical.Data.Unit
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.HITs.SetQuotients
-open import Cubical.Categories.Category
-
-module Realizability.Topos.FunctionalRelation
-  { ℓ' ℓ''}
-  {A : Type }
-  (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
-  where
-
-open import Realizability.Tripos.Logic.Syntax { = ℓ'}
-open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_)
-open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
-open import Realizability.Tripos.Prealgebra.Meets.Identity ca
-open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
-
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Predicate' renaming (isSetX to isSetPredicateBase)
-open PredicateProperties
-open Morphism
-
-private
-  _⊩_ :  a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*)  Type _
-  a  P = a pre⊩  P  tt*
-
-  
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
-
-
-open PartialEquivalenceRelation
-
-record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
-  field
-    perX : PartialEquivalenceRelation X
-    perY : PartialEquivalenceRelation Y
-  _~X_ = perX ._~_
-  _~Y_ = perY ._~_
-
-  field
-    relation : Predicate (X × Y)
-
-  `X : Sort
-  `X =  (X , perX .isSetX)
-
-  `Y : Sort
-  `Y =  (Y , perY .isSetX)
-
-  private
-    relationSymbol : Vec Sort 3
-    relationSymbol = (`X  `Y)  `X  `X  `Y  `Y  []
-
-    `F : Fin 3
-    `F = zero
-    `~X : Fin 3
-    `~X = one
-    `~Y : Fin 3
-    `~Y = two
-
-  open Relational relationSymbol
-
-  module RelationInterpretation' = Interpretation relationSymbol  { zero  relation ; one  _~X_ ; two  _~Y_ }) isNonTrivial
-  open RelationInterpretation'
-
-  module RelationSoundness = Soundness {relSym = relationSymbol} isNonTrivial  { zero  relation ; one  _~X_ ; two  _~Y_ })
-  open RelationSoundness
-
-  -- # Strictness
-  -- If the functional relation holds for x and y then x and y "exist"
-  private
-    strictnessContext : Context
-    strictnessContext = ([]  `X)  `Y
-
-    x∈strictnessContext : `X  strictnessContext
-    x∈strictnessContext = there here
-
-    y∈strictnessContext : `Y  strictnessContext
-    y∈strictnessContext = here
-
-    xˢᵗ : Term strictnessContext `X
-    xˢᵗ = var x∈strictnessContext
-
-    yˢᵗ : Term strictnessContext `Y
-    yˢᵗ = var y∈strictnessContext
-
-  -- F : Rel X Y, _~X_ : Rel X X, _~Y_ : Rel Y Y ⊢ F(x ,y) → (x ~X x) ∧ (y ~Y y)
-  strictnessFormula : Formula strictnessContext
-  strictnessFormula = rel `F (xˢᵗ `, yˢᵗ) `→ (rel `~X (xˢᵗ `, xˢᵗ) `∧ rel `~Y (yˢᵗ `, yˢᵗ))
-
-  field
-    isStrict : holdsInTripos strictnessFormula
-
-  -- # Relational
-  -- The functional relation preserves equality
-  -- "Substitutive" might be a better term for this property
-  private
-    relationalContext : Context
-    relationalContext =
-      []  `Y  `Y  `X  `X
-
-    x₁∈relationalContext : `X  relationalContext
-    x₁∈relationalContext = there here
-
-    x₂∈relationalContext : `X  relationalContext
-    x₂∈relationalContext = here
-
-    y₁∈relationalContext : `Y  relationalContext
-    y₁∈relationalContext = there (there here)
-
-    y₂∈relationalContext : `Y  relationalContext
-    y₂∈relationalContext = there (there (there here))
-
-    x₁ = var x₁∈relationalContext
-    x₂ = var x₂∈relationalContext
-    y₁ = var y₁∈relationalContext
-    y₂ = var y₂∈relationalContext
-
-  relationalFormula : Formula relationalContext
-  relationalFormula = (rel `F (x₁ `, y₁) `∧ (rel `~X (x₁ `, x₂) `∧ rel `~Y (y₁ `, y₂))) `→ rel `F (x₂ `, y₂)
-
-  field
-    isRelational : holdsInTripos relationalFormula
-
-  -- # Single-valued
-  -- Self explanatory
-  private
-    singleValuedContext : Context
-    singleValuedContext =
-      []  `Y  `Y  `X
-
-    x∈singleValuedContext : `X  singleValuedContext
-    x∈singleValuedContext = here
-
-    y₁∈singleValuedContext : `Y  singleValuedContext
-    y₁∈singleValuedContext = there here
-
-    y₂∈singleValuedContext : `Y  singleValuedContext
-    y₂∈singleValuedContext = there (there here)
-
-    xˢᵛ = var x∈singleValuedContext
-    y₁ˢᵛ = var y₁∈singleValuedContext
-    y₂ˢᵛ = var y₂∈singleValuedContext
-
-  singleValuedFormula : Formula singleValuedContext
-  singleValuedFormula =
-    (rel `F (xˢᵛ `, y₁ˢᵛ) `∧ rel `F (xˢᵛ `, y₂ˢᵛ)) `→ rel `~Y (y₁ˢᵛ `, y₂ˢᵛ)
-
-  field
-    isSingleValued : holdsInTripos singleValuedFormula
-
-  -- # Total
-  -- For all existent elements in the domain x there is an element in the codomain y
-  -- such that F(x, y)
-  private
-    totalContext : Context
-    totalContext =
-      []  `X
-
-    x∈totalContext : `X  totalContext
-    x∈totalContext = here
-
-    xᵗˡ = var x∈totalContext
-
-  totalFormula : Formula totalContext
-  totalFormula = rel `~X (xᵗˡ `, xᵗˡ)  `→ `∃ (rel `F (renamingTerm (drop id) xᵗˡ `, var here))
-
-  field
-    isTotal : holdsInTripos totalFormula
-
-open FunctionalRelation hiding (`X; `Y)
-
-pointwiseEntailment :  {X Y : Type ℓ'}  FunctionalRelation X Y  FunctionalRelation X Y  Type _
-pointwiseEntailment {X} {Y} F G = holdsInTripos entailmentFormula where
-  
-  `X : Sort
-  `Y : Sort
-
-  `X =  (X , F .perX .isSetX)
-  `Y =  (Y , F .perY .isSetX)
-
-  relationSymbols : Vec Sort 2
-  relationSymbols = (`X  `Y)  (`X  `Y)  []
-
-  `F : Fin 2
-  `F = zero
-
-  `G : Fin 2
-  `G = one
-
-  open Relational relationSymbols
-
-  module RelationalInterpretation = Interpretation relationSymbols  { zero  F .relation ; one  G .relation }) isNonTrivial
-  open RelationalInterpretation
-
-  module RelationalSoundness = Soundness {relSym = relationSymbols} isNonTrivial  { zero  F .relation ; one  G .relation })
-  open RelationalSoundness
-
-  entailmentContext : Context
-  entailmentContext = []  `X  `Y
-
-  x∈entailmentContext : `X  entailmentContext
-  x∈entailmentContext = there here
-
-  y∈entailmentContext : `Y  entailmentContext
-  y∈entailmentContext = here
-
-  x = var x∈entailmentContext
-  y = var y∈entailmentContext
-
-  entailmentFormula : Formula entailmentContext
-  entailmentFormula = rel `F (x `, y) `→ rel `G (x `, y)
-
-functionalRelationIsomorphism :  {X Y : Type ℓ'}  FunctionalRelation X Y  FunctionalRelation X Y  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
-functionalRelationIsomorphism {X} {Y} F G =
-  pointwiseEntailment F G × pointwiseEntailment G F
-
-
-pointwiseEntailment→functionalRelationIsomorphism :  {X Y : Type ℓ'}  (F G : FunctionalRelation X Y)  pointwiseEntailment F G  functionalRelationIsomorphism F G
-pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , G≤F where
-    
-  `X : Sort
-  `Y : Sort
-
-  `X =  (X , F .perX .isSetX)
-  `Y =  (Y , F .perY .isSetX)
-
-  relationSymbols : Vec Sort 4
-  relationSymbols = (`X  `Y)  (`X  `Y)  (`X  `X)  (`Y  `Y)  []
-
-  `F : Fin 4
-  `F = zero
-
-  `G : Fin 4
-  `G = one
-
-  `~X : Fin 4
-  `~X = two
-
-  `~Y : Fin 4
-  `~Y = three
-
-  open Interpretation relationSymbols  { zero  F .relation ; one  G .relation ; two  F .perX ._~_ ; three  G .perY ._~_}) isNonTrivial
-  open Soundness {relSym = relationSymbols} isNonTrivial ((λ { zero  F .relation ; one  G .relation ; two  F .perX ._~_ ; three  G .perY ._~_}))
-  open Relational relationSymbols
-
-  -- What we need to prove is that
-  -- F ≤ G ⊨ G ≤ F
-  -- We will use the semantic combinators we borrowed from the 1lab
-
-  entailmentContext : Context
-  entailmentContext = []  `X  `Y
-
-  x : Term entailmentContext `X
-  x = var (there here)
-
-  y : Term entailmentContext `Y
-  y = var here
-
-  G⊨x~x : (⊤ᵗ `∧ rel `G (x `, y))  rel `~X (x `, x)
-  G⊨x~x =
-    `→elim
-      {Γ = entailmentContext}
-      {ϕ = ⊤ᵗ `∧ (rel `G (x `, y))}
-      {ψ = rel `G (x `, y)}
-      {θ = rel `~X (x `, x)}
-      {!G .isStrict!}
-      {!!}
-
-  ⊤∧G⊨F : (⊤ᵗ `∧ rel `G (x `, y))  rel `F (x `, y)
-  ⊤∧G⊨F = cut {Γ = entailmentContext} {ϕ = ⊤ᵗ `∧ rel `G (x `, y)} {ψ = rel `~X (x `, x)}
-           {θ = rel `F (x `, y)}
-           G⊨x~x
-           {!!}
-
-  G≤F : pointwiseEntailment G F
-  G≤F = `→intro {Γ = entailmentContext} {ϕ = ⊤ᵗ} {ψ = rel `G (x `, y)} {θ = rel `F (x `, y)} ⊤∧G⊨F
-
-RTMorphism : (X Y : Type ℓ')   Type _
-RTMorphism X Y = FunctionalRelation X Y / λ F G  functionalRelationIsomorphism F G
-
-RTObject : Type _
-RTObject = Σ[ X  Type ℓ' ] PartialEquivalenceRelation X
-
-idMorphism : (ob : RTObject)  RTMorphism (ob .fst) (ob .fst)
-idMorphism ob =
-  [ record
-     { perX = ob .snd
-     ; perY = ob .snd
-     ; relation = ob .snd ._~_
-     ; isStrict = {!!}
-     ; isRelational = {!!}
-     ; isSingleValued = {!!}
-     ; isTotal = {!!}
-     } ] where
-
-  `X : Sort
-  `X =  (ob .fst , ob .snd .isSetX)
-
-  relationSymbols : Vec Sort 3
-  relationSymbols = (`X  `X)  (`X  `X)  (`X  `X)  []
-
-  open Interpretation relationSymbols  { zero  ob .snd ._~_ ; one  ob .snd ._~_ ; two  ob .snd ._~_ }) isNonTrivial
-  open Soundness {relSym = relationSymbols} isNonTrivial  { zero  ob .snd ._~_ ; one  ob .snd ._~_ ; two  ob .snd ._~_ })
-  open Relational relationSymbols
-
-  isStrictContext : Context
-  isStrictContext = []  `X  `X
-
-  x : Term isStrictContext `X
-  x = var (there here)
-
-  y : Term isStrictContext `X
-  y = var here
-
-  holdsInTriposIsStrict : holdsInTripos (rel zero (x `, y) `→ (rel one (x `, y) `∧ rel two (x `, y)))
-  holdsInTriposIsStrict =
-    do
-    let
-      prover : ApplStrTerm as 2
-      prover =
-        ` pair ̇ # fone ̇ # fone
-    return
-      (λ* prover ,
-       { ((tt* , x') , y') a tt* b b⊩x'~y'
-        
-          let
-            proofEq : λ* prover  a  b  pair  b  b
-            proofEq =
-              λ*ComputationRule prover (a  b  [])
-
-            pr₁proofEq : pr₁  (λ* prover  a  b)  b
-            pr₁proofEq =
-              pr₁  (λ* prover  a  b)
-                ≡⟨ cong  x  pr₁  x) proofEq 
-              pr₁  (pair  b  b)
-                ≡⟨ pr₁pxy≡x b b 
-              b
-                
-
-            pr₂proofEq : pr₂  (λ* prover  a  b)  b
-            pr₂proofEq =
-              pr₂  (λ* prover  a  b)
-                ≡⟨ cong  x  pr₂  x) proofEq 
-              pr₂  (pair  b  b)
-                ≡⟨ pr₂pxy≡y b b 
-              b
-                
-          in
-          (subst
-             r  r pre⊩  (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX))  γ  snd (fst γ) , snd γ) (ob .snd ._~_))  ((tt* , x') , y'))
-          (sym pr₁proofEq) b⊩x'~y') ,
-          subst
-             r  r pre⊩  (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX))  γ  snd (fst γ) , snd γ) (ob .snd ._~_))  ((tt* , x') , y')) (sym pr₂proofEq) b⊩x'~y' }))
-  
-
-RT : Category (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) ((ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')))
-Category.ob RT = Σ[ X  Type ℓ' ] PartialEquivalenceRelation X
-Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism X Y
-Category.id RT {X , perX} = {!!}
-Category._⋆_ RT = {!!}
-Category.⋆IdL RT = {!!}
-Category.⋆IdR RT = {!!}
-Category.⋆Assoc RT = {!!}
-Category.isSetHom RT = {!!}
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData
+open import Cubical.Data.Fin hiding (Fin; _/_)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients renaming (rec to setQuotRec; rec2 to setQuotRec2; elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2; elimProp3 to setQuotElimProp3)
+open import Cubical.Categories.Category
+
+module Realizability.Topos.FunctionalRelation
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
+
+open import Realizability.Tripos.Logic.Syntax { = ℓ'}
+open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate')
+open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate' renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
+private λ* = `λ* as fefermanStructure
+
+private
+  Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''}
+
+open PartialEquivalenceRelation
+
+record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  equalityX = perX .equality
+  equalityY = perY .equality
+
+  field
+    relation : Predicate (X × Y)
+    isStrict :
+      ∃[ st  A ]
+      (∀ x y r
+       r   relation  (x , y)
+      ------------------------------------------------------------------------------------
+       (pr₁  (st  r))   equalityX  (x , x) × ((pr₂  (st  r))   equalityY  (y , y)))
+    isRelational :
+      ∃[ rl  A ]
+      (∀ x x' y y' a b c
+       a   equalityX  (x , x')
+       b   relation  (x , y)
+       c   equalityY  (y , y')
+      ------------------------------------------
+       (rl  (pair  a  (pair  b  c)))   relation  (x' , y'))
+    isSingleValued :
+      ∃[ sv  A ]
+      (∀ x y y' r₁ r₂
+       r₁   relation  (x , y)
+       r₂   relation  (x , y')
+      -----------------------------------
+       (sv  (pair  r₁  r₂))   equalityY  (y , y'))
+    isTotal :
+      ∃[ tl  A ]
+      (∀ x r  r   equalityX  (x , x)  ∃[ y  Y ] (tl  r)   relation  (x , y))
+
+open FunctionalRelation
+
+pointwiseEntailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  (F G : FunctionalRelation perX perY)  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
+pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe  A ] (∀ x y r  r   F .relation  (x , y)  (pe  r)   G .relation  (x , y))
+
+-- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong
+-- Lemma 4.3.5
+F≤G→G≤F :
+   {X Y : Type ℓ'}
+   (perX : PartialEquivalenceRelation X)
+   (perY : PartialEquivalenceRelation Y)
+   (F G : FunctionalRelation perX perY)
+   pointwiseEntailment perX perY F G
+   pointwiseEntailment perX perY G F
+F≤G→G≤F {X} {Y} perX perY F G F≤G =
+  do
+    (r , r⊩F≤G)  F≤G
+    (stG , stG⊩isStrictG)  G .isStrict
+    (tlF , tlF⊩isTotalF)  F .isTotal
+    (svG , svG⊩isSingleValuedG)  G .isSingleValued
+    (rlF , rlF⊩isRelational)  F .isRelational
+    let
+      prover : ApplStrTerm as 1
+      prover = ` rlF ̇ (` pair ̇ (` pr₁ ̇ (` stG ̇ # fzero)) ̇ (` pair ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero))) ̇ (` svG ̇ (` pair ̇ (` r ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero)))) ̇ # fzero))))
+    return
+      (λ* prover ,
+       x y r' r'⊩Gxy 
+        subst
+           r''  r''   F .relation  (x , y))
+          (sym (λ*ComputationRule prover (r'  [])))
+          (transport
+            (propTruncIdempotent (F .relation .isPropValued (x , y) _))
+            (do
+              (y' , Fxy')  tlF⊩isTotalF x (pr₁  (stG  r')) (stG⊩isStrictG x y r' r'⊩Gxy .fst)
+              return
+                (rlF⊩isRelational
+                  x x y' y
+                  (pr₁  (stG  r'))
+                  (tlF  (pr₁  (stG  r')))
+                  (svG  (pair  (r  (tlF  (pr₁  (stG  r'))))  r'))
+                  (stG⊩isStrictG x y r' r'⊩Gxy .fst)
+                  Fxy'
+                  (svG⊩isSingleValuedG x y' y (r  (tlF  (pr₁  (stG  r')))) r' (r⊩F≤G x y' (tlF  (pr₁  (stG  r'))) Fxy') r'⊩Gxy))))))
+
+equalityFuncRel :  {X : Type ℓ'}  (perX : PartialEquivalenceRelation X)  FunctionalRelation perX perX
+relation (equalityFuncRel {X} perX) = perX .equality
+isStrict (equalityFuncRel {X} perX) =
+  do
+    (s , s⊩isSymmetric)  perX .isSymmetric
+    (t , t⊩isTransitive)  perX .isTransitive
+    let
+      prover : ApplStrTerm as 1
+      prover = ` pair ̇ (` t ̇ (` pair ̇ # fzero ̇ (` s ̇ # fzero))) ̇ (` t ̇ (` pair ̇ (` s ̇ # fzero) ̇ # fzero))
+    return
+      (λ* prover ,
+      λ x x' r r⊩x~x' 
+        let
+          pr₁proofEq : pr₁  (λ* prover  r)  t  (pair  r  (s  r))
+          pr₁proofEq =
+            pr₁  (λ* prover  r)
+              ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (r  [])) 
+            pr₁  (pair  (t  (pair  r  (s  r)))  (t  (pair  (s  r)  r)))
+              ≡⟨ pr₁pxy≡x _ _ 
+            t  (pair  r  (s  r))
+              
+
+          pr₂proofEq : pr₂  (λ* prover  r)  t  (pair  (s  r)  r)
+          pr₂proofEq =
+            pr₂  (λ* prover  r)
+              ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (r  [])) 
+            pr₂  (pair  (t  (pair  r  (s  r)))  (t  (pair  (s  r)  r)))
+              ≡⟨ pr₂pxy≡y _ _ 
+            t  (pair  (s  r)  r)
+              
+        in
+        subst  r'  r'   perX .equality  (x , x)) (sym pr₁proofEq) (t⊩isTransitive x x' x r (s  r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')) ,
+        subst  r'  r'   perX .equality  (x' , x')) (sym pr₂proofEq) (t⊩isTransitive x' x x' (s  r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x'))
+isRelational (equalityFuncRel {X} perX) =
+  do
+    (s , s⊩isSymmetric)  perX .isSymmetric
+    (t , t⊩isTransitive)  perX .isTransitive
+    let
+      prover : ApplStrTerm as 1
+      prover = ` t ̇ (` pair ̇ (` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))
+
+    return
+      (λ* prover ,
+       x x' y y' a b c a⊩x~x' b⊩x~y c⊩y~y' 
+        let
+          proofNormalForm : (λ* prover  (pair  a  (pair  b  c)))  (t  (pair  (t  (pair  (s  a)  b))  c))
+          proofNormalForm =
+            λ* prover  (pair  a  (pair  b  c))
+              ≡⟨ λ*ComputationRule prover ((pair  a  (pair  b  c))  []) 
+            (t 
+            (pair 
+             (t 
+              (pair  (s  (pr₁  (pair  a  (pair  b  c)))) 
+               (pr₁  (pr₂  (pair  a  (pair  b  c))))))
+              (pr₂  (pr₂  (pair  a  (pair  b  c))))))
+             ≡⟨
+               cong₂
+                  x y  (t  (pair  (t  (pair  (s  x)  y))  (pr₂  (pr₂  (pair  a  (pair  b  c)))))))
+                 (pr₁pxy≡x _ _)
+                 (pr₁  (pr₂  (pair  a  (pair  b  c)))
+                   ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                 pr₁  (pair  b  c)
+                   ≡⟨ pr₁pxy≡x _ _ 
+                 b
+                   )
+              
+            t  (pair  (t  (pair  (s  a)  b))  (pr₂  (pr₂  (pair  a  (pair  b  c)))))
+              ≡⟨
+                cong
+                   x  t  (pair  (t  (pair  (s  a)  b))  x))
+                  (pr₂  (pr₂  (pair  a  (pair  b  c)))
+                    ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                  pr₂  (pair  b  c)
+                    ≡⟨ pr₂pxy≡y _ _ 
+                  c
+                    )
+               
+            t  (pair  (t  (pair  (s  a)  b))  c)
+              
+        in
+        subst
+           r  r   perX .equality  (x' , y'))
+          (sym proofNormalForm)
+          (t⊩isTransitive x' y y' (t  (pair  (s  a)  b)) c (t⊩isTransitive x' x y (s  a) b (s⊩isSymmetric x x' a a⊩x~x') b⊩x~y) c⊩y~y')))
+isSingleValued (equalityFuncRel {X} perX) =
+  do
+    (s , s⊩isSymmetric)  perX .isSymmetric
+    (t , t⊩isTransitive)  perX .isTransitive
+    let
+      prover : ApplStrTerm as 1
+      prover = ` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero))
+    return
+      (λ* prover ,
+       x y y' r₁ r₂ r₁⊩x~y r₂⊩x~y' 
+        let
+          proofEq : λ* prover  (pair  r₁  r₂)  t  (pair  (s  r₁)  r₂)
+          proofEq = {!!}
+        in
+        subst
+           r  r   perX .equality  (y , y'))
+          (sym proofEq)
+          (t⊩isTransitive y x y' (s  r₁) r₂ (s⊩isSymmetric x y r₁ r₁⊩x~y) r₂⊩x~y')))
+isTotal (equalityFuncRel {X} perX) =
+  do
+    (s , s⊩isSymmetric)  perX .isSymmetric
+    (t , t⊩isTransitive)  perX .isTransitive
+    return (Id ,  x r r⊩x~x   x , (subst  r  r   perX .equality  (x , x)) (sym (Ida≡a r)) r⊩x~x) ∣₁))
+
+composeFuncRel :
+   {X Y Z : Type ℓ'}
+   (perX : PartialEquivalenceRelation X)
+   (perY : PartialEquivalenceRelation Y)
+   (perZ : PartialEquivalenceRelation Z)
+   (F : FunctionalRelation perX perY)
+   (G : FunctionalRelation perY perZ)
+   FunctionalRelation perX perZ
+
+(relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
+  record
+    { isSetX = isSet× (perX .isSetX) (perZ .isSetX)
+    ; ∣_∣ = λ { (x , z) r  ∃[ y  Y ] (pr₁  r)   F .relation  (x , y) × (pr₂  r)   G .relation  (y , z) }
+    ; isPropValued = λ x a  isPropPropTrunc }
+isStrict (composeFuncRel {X} {Y} {Z} perX perY perZ F G) =
+  do
+    (stF , stF⊩isStrictF)  F .isStrict
+    (stG , stG⊩isStrictG)  G .isStrict
+    let
+      prover : ApplStrTerm as 1
+      prover = ` pair ̇ (` pr₁ ̇ (` stF ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ # fzero)))
+    return
+      (λ* prover ,
+       x z r r⊩∃y 
+        transport (propTruncIdempotent {!!})
+        (do
+          (y , pr₁r⊩Fxy , pr₂r⊩Gxy)  r⊩∃y
+          let
+            pr₁proofEq : pr₁  (λ* prover  r)  pr₁  (stF  (pr₁  r))
+            pr₁proofEq =
+              pr₁  (λ* prover  r)
+                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (r  [])) 
+              pr₁  (pair  (pr₁  (stF  (pr₁  r)))  (pr₂  (stG  (pr₂  r))))
+                ≡⟨ pr₁pxy≡x _ _ 
+              pr₁  (stF  (pr₁  r))
+                
+
+            pr₂proofEq : pr₂  (λ* prover  r)  pr₂  (stG  (pr₂  r))
+            pr₂proofEq = {!!}
+          return
+            ((subst  r'  r'   perX .equality  (x , x)) (sym pr₁proofEq) (stF⊩isStrictF x y (pr₁  r) pr₁r⊩Fxy .fst)) ,
+              subst  r'  r'   perZ .equality  (z , z)) (sym pr₂proofEq) (stG⊩isStrictG y z (pr₂  r) pr₂r⊩Gxy .snd)))))
+isRelational (composeFuncRel {X} {Y} {Z} perX perY perZ F G) =
+  do
+    (rlF , rlF⊩isRelationalF)  F .isRelational
+    (stF , stF⊩isStrictF)  F .isStrict
+    (rlG , rlG⊩isRelationalG)  G .isRelational
+    let
+      prover : ApplStrTerm as 1
+      prover =
+        ` pair ̇
+          (` rlF ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))))))) ̇
+          (` rlG ̇ (` pair ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))))) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))))
+    return
+      (λ* prover ,
+       x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' 
+        do
+          (y , pr₁b⊩Fxy , pr₂b⊩Gyz)  b⊩∃y
+          let
+            proofEq : λ* prover  (pair  a  (pair  b  c))  pair  (rlF  (pair  a  (pair  (pr₁  b)  (pr₂  (stF  (pr₁  b))))))  (rlG  (pair  (pr₂  (stF  (pr₁  b)))  (pair  (pr₂  b)  c)))
+            proofEq =
+              λ* prover  (pair  a  (pair  b  c))
+                ≡⟨ λ*ComputationRule prover ((pair  a  (pair  b  c))  []) 
+              {!!}
+          return
+            (y ,
+            (subst  r  r   F .relation  (x' , y)) (sym {!!}) (rlF⊩isRelationalF x x' y y a (pr₁  b) (pr₂  (stF  (pr₁  b))) a⊩x~x' pr₁b⊩Fxy (stF⊩isStrictF x y (pr₁  b) pr₁b⊩Fxy .snd)) ,
+            (subst  r  r   G .relation  (y , z')) (sym {!!}) (rlG⊩isRelationalG y y z z' (pr₂  (stF  (pr₁  b))) (pr₂  b) c (stF⊩isStrictF x y (pr₁  b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z'))))))
+isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!}
+isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!}
+
+RTMorphism :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  Type _
+RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G  pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F
+
+idRTMorphism :  {X : Type ℓ'}  (perX : PartialEquivalenceRelation X)  RTMorphism perX perX
+idRTMorphism {X} perX = [ equalityFuncRel perX ]
+
+composeRTMorphism :
+   {X Y Z : Type ℓ'}
+   (perX : PartialEquivalenceRelation X)
+   (perY : PartialEquivalenceRelation Y)
+   (perZ : PartialEquivalenceRelation Z)
+   (f : RTMorphism perX perY)
+   (g : RTMorphism perY perZ)
+  ----------------------------------------
+   RTMorphism perX perZ
+composeRTMorphism {X} {Y} {Z} perX perY perZ f g =
+  setQuotRec2
+    squash/
+     F G  [ composeFuncRel perX perY perZ F G ])
+     { F F' G (F≤F' , F'≤F)  eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) ({!!} , {!!}) })
+     { F G G' (G≤G' , G'≤G)  eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') ({!!} , {!!}) })
+    f g
+
+idLRTMorphism :
+   {X Y : Type ℓ'}
+   (perX : PartialEquivalenceRelation X)
+   (perY : PartialEquivalenceRelation Y)
+   (f : RTMorphism perX perY)
+   composeRTMorphism perX perX perY (idRTMorphism perX) f  f
+idLRTMorphism {X} {Y} perX perY f =
+  setQuotElimProp
+     f  squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f)
+     f  eq/ _ _ ({!!} , {!!}))
+    f
+
+idRRTMorphism :
+   {X Y : Type ℓ'}
+   (perX : PartialEquivalenceRelation X)
+   (perY : PartialEquivalenceRelation Y)
+   (f : RTMorphism perX perY)
+   composeRTMorphism perX perY perY f (idRTMorphism perY)  f
+idRRTMorphism {X} {Y} perX perY f =
+  setQuotElimProp
+     f  squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f)
+     f  eq/ _ _ ({!!} , {!!}))
+    f
+
+assocRTMorphism :
+   {X Y Z W : Type ℓ'}
+   (perX : PartialEquivalenceRelation X)
+   (perY : PartialEquivalenceRelation Y)
+   (perZ : PartialEquivalenceRelation Z)
+   (perW : PartialEquivalenceRelation W)
+   (f : RTMorphism perX perY)
+   (g : RTMorphism perY perZ)
+   (h : RTMorphism perZ perW)
+   composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h  composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h)
+assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h =
+  setQuotElimProp3
+     f g h 
+      squash/
+        (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h)
+        (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h)))
+     f g h  eq/ _ _ ({!!} , {!!}))
+    f g h
+
+RT : Category (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ'')))
+Category.ob RT = Σ[ X  Type ℓ' ] PartialEquivalenceRelation X
+Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism perX perY
+Category.id RT {X , perX} = idRTMorphism perX
+Category._⋆_ RT {X , perX} {y , perY} {Z , perZ} f g = composeRTMorphism perX perY perZ f g
+Category.⋆IdL RT {X , perX} {Y , perY} f = idLRTMorphism perX perY f
+Category.⋆IdR RT {X , perX} {Y , perY} f = idRRTMorphism perX perY f
+Category.⋆Assoc RT {X , perX} {Y , perY} {Z , perZ} {W , perW} f g h = assocRTMorphism perX perY perZ perW f g h
+Category.isSetHom RT = squash/
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.Object.html b/docs/Realizability.Topos.Object.html index 70a72f3..b87ceb6 100644 --- a/docs/Realizability.Topos.Object.html +++ b/docs/Realizability.Topos.Object.html @@ -18,95 +18,24 @@ open import Realizability.Tripos.Logic.Syntax { = ℓ'} open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_) -open import Realizability.Tripos.Prealgebra.Predicate.Properties ca -open CombinatoryAlgebra ca -open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open Predicate' renaming (isSetX to isSetPredicateBase) -open PredicateProperties -open Morphism - -Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} - -record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - field - isSetX : isSet X - private - `X : Sort - `X = (X , isSetX) - - `X×X : Sort - `X×X = `X `X - - ~relSym : Vec Sort 1 - ~relSym = `X×X [] - - module X×XRelational = Relational {n = 1} ~relSym - open X×XRelational - - field - _~_ : Predicate lookup fzero ~relSym ⟧ˢ - - private - ~relSymInterpretation : RelationInterpretation ~relSym - ~relSymInterpretation = λ { fzero _~_ } +open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') +open import Realizability.Tripos.Prealgebra.Predicate.Properties ca +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate' renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +private + Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} + +record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + field + isSetX : isSet X + equality : Predicate (X × X) + isSymmetric : ∃[ s A ] (∀ x y r r equality (x , y) (s r) equality (y , x)) + open PredicateProperties {ℓ'' = ℓ''} (X × X) + field + isTransitive : ∃[ t A ] (∀ x y z a b a equality (x , y) b equality (y , z) (t (pair a b)) equality (x , z)) - module ~Interpretation = Interpretation ~relSym ~relSymInterpretation isNonTrivial - open ~Interpretation - - module ~Soundness = Soundness isNonTrivial ~relSymInterpretation - open ~Soundness - - -- Partial equivalence relations - - private - symContext : Context - symContext = ([] `X) `X - - x∈symContext : `X symContext - x∈symContext = there here - - y∈symContext : `X symContext - y∈symContext = here - - : Term symContext `X - = var x∈symContext - - : Term symContext `X - = var y∈symContext - - symmetryFormula : Formula symContext - symmetryFormula = rel fzero ( `, ) `→ rel fzero ( `, ) - - field - symmetry : holdsInTripos symmetryFormula - - private - transContext : Context - transContext = (([] `X) `X) `X - - z∈transContext : `X transContext - z∈transContext = here - - y∈transContext : `X transContext - y∈transContext = there here - - x∈transContext : `X transContext - x∈transContext = there (there here) - - zᵗ : Term transContext `X - zᵗ = var z∈transContext - - yᵗ : Term transContext `X - yᵗ = var y∈transContext - - xᵗ : Term transContext `X - xᵗ = var x∈transContext - - transitivityFormula : Formula transContext - transitivityFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) - - field - transitivity : holdsInTripos transitivityFormula - \ No newline at end of file diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 9516a10..cb3ee63 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -1,6 +1,5 @@ {-# OPTIONS --allow-unsolved-metas #-} open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) - open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure @@ -14,7 +13,7 @@ open import Cubical.Data.Empty open import Cubical.Data.Unit open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients +open import Cubical.HITs.SetQuotients renaming (rec to setQuotRec; rec2 to setQuotRec2; elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2; elimProp3 to setQuotElimProp3) open import Cubical.Categories.Category module Realizability.Topos.FunctionalRelation @@ -26,7 +25,7 @@ module Realizability.Topos.FunctionalRelation open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_) +open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') open import Realizability.Tripos.Prealgebra.Predicate.Properties ca open import Realizability.Tripos.Prealgebra.Meets.Identity ca open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial @@ -37,346 +36,340 @@ open Predicate' renaming (isSetX to isSetPredicateBase) open PredicateProperties open Morphism -private - _⊩_ : ∀ a (P : Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} Unit*) → Type _ - a ⊩ P = a pre⊩ ∣ P ∣ tt* - - private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure +private + Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} open PartialEquivalenceRelation -record FunctionalRelation (X Y : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - field - perX : PartialEquivalenceRelation X - perY : PartialEquivalenceRelation Y - _~X_ = perX ._~_ - _~Y_ = perY ._~_ +record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + equalityX = perX .equality + equalityY = perY .equality field relation : Predicate (X × Y) + isStrict : + ∃[ st ∈ A ] + (∀ x y r + → r ⊩ ∣ relation ∣ (x , y) + ------------------------------------------------------------------------------------ + → (pr₁ ⨾ (st ⨾ r)) ⊩ ∣ equalityX ∣ (x , x) × ((pr₂ ⨾ (st ⨾ r)) ⊩ ∣ equalityY ∣ (y , y))) + isRelational : + ∃[ rl ∈ A ] + (∀ x x' y y' a b c + → a ⊩ ∣ equalityX ∣ (x , x') + → b ⊩ ∣ relation ∣ (x , y) + → c ⊩ ∣ equalityY ∣ (y , y') + ------------------------------------------ + → (rl ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ⊩ ∣ relation ∣ (x' , y')) + isSingleValued : + ∃[ sv ∈ A ] + (∀ x y y' r₁ r₂ + → r₁ ⊩ ∣ relation ∣ (x , y) + → r₂ ⊩ ∣ relation ∣ (x , y') + ----------------------------------- + → (sv ⨾ (pair ⨾ r₁ ⨾ r₂)) ⊩ ∣ equalityY ∣ (y , y')) + isTotal : + ∃[ tl ∈ A ] + (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) + +open FunctionalRelation + +pointwiseEntailment : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → (F G : FunctionalRelation perX perY) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') +pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe ∈ A ] (∀ x y r → r ⊩ ∣ F .relation ∣ (x , y) → (pe ⨾ r) ⊩ ∣ G .relation ∣ (x , y)) + +-- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong +-- Lemma 4.3.5 +F≤G→G≤F : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (F G : FunctionalRelation perX perY) + → pointwiseEntailment perX perY F G + → pointwiseEntailment perX perY G F +F≤G→G≤F {X} {Y} perX perY F G F≤G = + do + (r , r⊩F≤G) ← F≤G + (stG , stG⊩isStrictG) ← G .isStrict + (tlF , tlF⊩isTotalF) ← F .isTotal + (svG , svG⊩isSingleValuedG) ← G .isSingleValued + (rlF , rlF⊩isRelational) ← F .isRelational + let + prover : ApplStrTerm as 1 + prover = ` rlF ̇ (` pair ̇ (` pr₁ ̇ (` stG ̇ # fzero)) ̇ (` pair ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero))) ̇ (` svG ̇ (` pair ̇ (` r ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero)))) ̇ # fzero)))) + return + (λ* prover , + (λ x y r' r'⊩Gxy → + subst + (λ r'' → r'' ⊩ ∣ F .relation ∣ (x , y)) + (sym (λ*ComputationRule prover (r' ∷ []))) + (transport + (propTruncIdempotent (F .relation .isPropValued (x , y) _)) + (do + (y' , Fxy') ← tlF⊩isTotalF x (pr₁ ⨾ (stG ⨾ r')) (stG⊩isStrictG x y r' r'⊩Gxy .fst) + return + (rlF⊩isRelational + x x y' y + (pr₁ ⨾ (stG ⨾ r')) + (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) + (svG ⨾ (pair ⨾ (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) ⨾ r')) + (stG⊩isStrictG x y r' r'⊩Gxy .fst) + Fxy' + (svG⊩isSingleValuedG x y' y (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) r' (r⊩F≤G x y' (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) Fxy') r'⊩Gxy)))))) + +equalityFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX +relation (equalityFuncRel {X} perX) = perX .equality +isStrict (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` t ̇ (` pair ̇ # fzero ̇ (` s ̇ # fzero))) ̇ (` t ̇ (` pair ̇ (` s ̇ # fzero) ̇ # fzero)) + return + (λ* prover , + λ x x' r r⊩x~x' → + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) + ≡⟨ pr₁pxy≡x _ _ ⟩ + t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₂ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) + ≡⟨ pr₂pxy≡y _ _ ⟩ + t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) + ∎ + in + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (t⊩isTransitive x x' x r (s ⨾ r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')) , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) (sym pr₂proofEq) (t⊩isTransitive x' x x' (s ⨾ r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x')) +isRelational (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` t ̇ (` pair ̇ (` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) - `X : Sort - `X = ↑ (X , perX .isSetX) - - `Y : Sort - `Y = ↑ (Y , perY .isSetX) - - private - relationSymbol : Vec Sort 3 - relationSymbol = (`X `× `Y) ∷ `X `× `X ∷ `Y `× `Y ∷ [] - - `F : Fin 3 - `F = zero - `~X : Fin 3 - `~X = one - `~Y : Fin 3 - `~Y = two - - open Relational relationSymbol - - module RelationInterpretation' = Interpretation relationSymbol (λ { zero → relation ; one → _~X_ ; two → _~Y_ }) isNonTrivial - open RelationInterpretation' - - module RelationSoundness = Soundness {relSym = relationSymbol} isNonTrivial (λ { zero → relation ; one → _~X_ ; two → _~Y_ }) - open RelationSoundness - - -- # Strictness - -- If the functional relation holds for x and y then x and y "exist" - private - strictnessContext : Context - strictnessContext = ([] ′ `X) ′ `Y - - x∈strictnessContext : `X ∈ strictnessContext - x∈strictnessContext = there here - - y∈strictnessContext : `Y ∈ strictnessContext - y∈strictnessContext = here - - xˢᵗ : Term strictnessContext `X - xˢᵗ = var x∈strictnessContext - - yˢᵗ : Term strictnessContext `Y - yˢᵗ = var y∈strictnessContext - - -- F : Rel X Y, _~X_ : Rel X X, _~Y_ : Rel Y Y ⊢ F(x ,y) → (x ~X x) ∧ (y ~Y y) - strictnessFormula : Formula strictnessContext - strictnessFormula = rel `F (xˢᵗ `, yˢᵗ) `→ (rel `~X (xˢᵗ `, xˢᵗ) `∧ rel `~Y (yˢᵗ `, yˢᵗ)) - - field - isStrict : holdsInTripos strictnessFormula - - -- # Relational - -- The functional relation preserves equality - -- "Substitutive" might be a better term for this property - private - relationalContext : Context - relationalContext = - [] ′ `Y ′ `Y ′ `X ′ `X - - x₁∈relationalContext : `X ∈ relationalContext - x₁∈relationalContext = there here - - x₂∈relationalContext : `X ∈ relationalContext - x₂∈relationalContext = here - - y₁∈relationalContext : `Y ∈ relationalContext - y₁∈relationalContext = there (there here) - - y₂∈relationalContext : `Y ∈ relationalContext - y₂∈relationalContext = there (there (there here)) - - x₁ = var x₁∈relationalContext - x₂ = var x₂∈relationalContext - y₁ = var y₁∈relationalContext - y₂ = var y₂∈relationalContext - - relationalFormula : Formula relationalContext - relationalFormula = (rel `F (x₁ `, y₁) `∧ (rel `~X (x₁ `, x₂) `∧ rel `~Y (y₁ `, y₂))) `→ rel `F (x₂ `, y₂) - - field - isRelational : holdsInTripos relationalFormula - - -- # Single-valued - -- Self explanatory - private - singleValuedContext : Context - singleValuedContext = - [] ′ `Y ′ `Y ′ `X - - x∈singleValuedContext : `X ∈ singleValuedContext - x∈singleValuedContext = here - - y₁∈singleValuedContext : `Y ∈ singleValuedContext - y₁∈singleValuedContext = there here - - y₂∈singleValuedContext : `Y ∈ singleValuedContext - y₂∈singleValuedContext = there (there here) - - xˢᵛ = var x∈singleValuedContext - y₁ˢᵛ = var y₁∈singleValuedContext - y₂ˢᵛ = var y₂∈singleValuedContext - - singleValuedFormula : Formula singleValuedContext - singleValuedFormula = - (rel `F (xˢᵛ `, y₁ˢᵛ) `∧ rel `F (xˢᵛ `, y₂ˢᵛ)) `→ rel `~Y (y₁ˢᵛ `, y₂ˢᵛ) - - field - isSingleValued : holdsInTripos singleValuedFormula - - -- # Total - -- For all existent elements in the domain x there is an element in the codomain y - -- such that F(x, y) - private - totalContext : Context - totalContext = - [] ′ `X - - x∈totalContext : `X ∈ totalContext - x∈totalContext = here - - xᵗˡ = var x∈totalContext - - totalFormula : Formula totalContext - totalFormula = rel `~X (xᵗˡ `, xᵗˡ) `→ `∃ (rel `F (renamingTerm (drop id) xᵗˡ `, var here)) - - field - isTotal : holdsInTripos totalFormula - -open FunctionalRelation hiding (`X; `Y) - -pointwiseEntailment : ∀ {X Y : Type ℓ'} → FunctionalRelation X Y → FunctionalRelation X Y → Type _ -pointwiseEntailment {X} {Y} F G = holdsInTripos entailmentFormula where - - `X : Sort - `Y : Sort - - `X = ↑ (X , F .perX .isSetX) - `Y = ↑ (Y , F .perY .isSetX) - - relationSymbols : Vec Sort 2 - relationSymbols = (`X `× `Y) ∷ (`X `× `Y) ∷ [] - - `F : Fin 2 - `F = zero - - `G : Fin 2 - `G = one - - open Relational relationSymbols - - module RelationalInterpretation = Interpretation relationSymbols (λ { zero → F .relation ; one → G .relation }) isNonTrivial - open RelationalInterpretation - - module RelationalSoundness = Soundness {relSym = relationSymbols} isNonTrivial (λ { zero → F .relation ; one → G .relation }) - open RelationalSoundness - - entailmentContext : Context - entailmentContext = [] ′ `X ′ `Y - - x∈entailmentContext : `X ∈ entailmentContext - x∈entailmentContext = there here - - y∈entailmentContext : `Y ∈ entailmentContext - y∈entailmentContext = here - - x = var x∈entailmentContext - y = var y∈entailmentContext - - entailmentFormula : Formula entailmentContext - entailmentFormula = rel `F (x `, y) `→ rel `G (x `, y) - -functionalRelationIsomorphism : ∀ {X Y : Type ℓ'} → FunctionalRelation X Y → FunctionalRelation X Y → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') -functionalRelationIsomorphism {X} {Y} F G = - pointwiseEntailment F G × pointwiseEntailment G F - - -pointwiseEntailment→functionalRelationIsomorphism : ∀ {X Y : Type ℓ'} → (F G : FunctionalRelation X Y) → pointwiseEntailment F G → functionalRelationIsomorphism F G -pointwiseEntailment→functionalRelationIsomorphism {X} {Y} F G F≤G = F≤G , G≤F where - - `X : Sort - `Y : Sort - - `X = ↑ (X , F .perX .isSetX) - `Y = ↑ (Y , F .perY .isSetX) - - relationSymbols : Vec Sort 4 - relationSymbols = (`X `× `Y) ∷ (`X `× `Y) ∷ (`X `× `X) ∷ (`Y `× `Y) ∷ [] - - `F : Fin 4 - `F = zero - - `G : Fin 4 - `G = one - - `~X : Fin 4 - `~X = two - - `~Y : Fin 4 - `~Y = three - - open Interpretation relationSymbols (λ { zero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_}) isNonTrivial - open Soundness {relSym = relationSymbols} isNonTrivial ((λ { zero → F .relation ; one → G .relation ; two → F .perX ._~_ ; three → G .perY ._~_})) - open Relational relationSymbols - - -- What we need to prove is that - -- F ≤ G ⊨ G ≤ F - -- We will use the semantic combinators we borrowed from the 1lab - - entailmentContext : Context - entailmentContext = [] ′ `X ′ `Y - - x : Term entailmentContext `X - x = var (there here) - - y : Term entailmentContext `Y - y = var here - - G⊨x~x : (⊤ᵗ `∧ rel `G (x `, y)) ⊨ rel `~X (x `, x) - G⊨x~x = - `→elim - {Γ = entailmentContext} - {ϕ = ⊤ᵗ `∧ (rel `G (x `, y))} - {ψ = rel `G (x `, y)} - {θ = rel `~X (x `, x)} - {!G .isStrict!} - {!!} - - ⊤∧G⊨F : (⊤ᵗ `∧ rel `G (x `, y)) ⊨ rel `F (x `, y) - ⊤∧G⊨F = cut {Γ = entailmentContext} {ϕ = ⊤ᵗ `∧ rel `G (x `, y)} {ψ = rel `~X (x `, x)} - {θ = rel `F (x `, y)} - G⊨x~x - {!!} - - G≤F : pointwiseEntailment G F - G≤F = `→intro {Γ = entailmentContext} {ϕ = ⊤ᵗ} {ψ = rel `G (x `, y)} {θ = rel `F (x `, y)} ⊤∧G⊨F - -RTMorphism : (X Y : Type ℓ') → Type _ -RTMorphism X Y = FunctionalRelation X Y / λ F G → functionalRelationIsomorphism F G - -RTObject : Type _ -RTObject = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X - -idMorphism : (ob : RTObject) → RTMorphism (ob .fst) (ob .fst) -idMorphism ob = - [ record - { perX = ob .snd - ; perY = ob .snd - ; relation = ob .snd ._~_ - ; isStrict = {!!} - ; isRelational = {!!} - ; isSingleValued = {!!} - ; isTotal = {!!} - } ] where - - `X : Sort - `X = ↑ (ob .fst , ob .snd .isSetX) - - relationSymbols : Vec Sort 3 - relationSymbols = (`X `× `X) ∷ (`X `× `X) ∷ (`X `× `X) ∷ [] - - open Interpretation relationSymbols (λ { zero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) isNonTrivial - open Soundness {relSym = relationSymbols} isNonTrivial (λ { zero → ob .snd ._~_ ; one → ob .snd ._~_ ; two → ob .snd ._~_ }) - open Relational relationSymbols - - isStrictContext : Context - isStrictContext = [] ′ `X ′ `X - - x : Term isStrictContext `X - x = var (there here) - - y : Term isStrictContext `X - y = var here - - holdsInTriposIsStrict : holdsInTripos (rel zero (x `, y) `→ (rel one (x `, y) `∧ rel two (x `, y))) - holdsInTriposIsStrict = - do + return + (λ* prover , + (λ x x' y y' a b c a⊩x~x' b⊩x~y c⊩y~y' → + let + proofNormalForm : (λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ≡ (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c)) + proofNormalForm = + λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) + ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ + (t ⨾ + (pair ⨾ + (t ⨾ + (pair ⨾ (s ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))) ⨾ + (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) + ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) + ≡⟨ + cong₂ + (λ x y → (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ x) ⨾ y)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))))) + (pr₁pxy≡x _ _) + (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) + ≡⟨ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ⟩ + pr₁ ⨾ (pair ⨾ b ⨾ c) + ≡⟨ pr₁pxy≡x _ _ ⟩ + b + ∎) + ⟩ + t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))) + ≡⟨ + cong + (λ x → t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ x)) + (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) + ≡⟨ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ⟩ + pr₂ ⨾ (pair ⨾ b ⨾ c) + ≡⟨ pr₂pxy≡y _ _ ⟩ + c + ∎) + ⟩ + t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c) + ∎ + in + subst + (λ r → r ⊩ ∣ perX .equality ∣ (x' , y')) + (sym proofNormalForm) + (t⊩isTransitive x' y y' (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) c (t⊩isTransitive x' x y (s ⨾ a) b (s⊩isSymmetric x x' a a⊩x~x') b⊩x~y) c⊩y~y'))) +isSingleValued (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive let - prover : ApplStrTerm as 2 - prover = - ` pair ̇ # fone ̇ # fone + prover : ApplStrTerm as 1 + prover = ` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero)) return (λ* prover , - (λ { ((tt* , x') , y') a tt* b b⊩x'~y' - → + (λ x y y' r₁ r₂ r₁⊩x~y r₂⊩x~y' → + let + proofEq : λ* prover ⨾ (pair ⨾ r₁ ⨾ r₂) ≡ t ⨾ (pair ⨾ (s ⨾ r₁) ⨾ r₂) + proofEq = {!!} + in + subst + (λ r → r ⊩ ∣ perX .equality ∣ (y , y')) + (sym proofEq) + (t⊩isTransitive y x y' (s ⨾ r₁) r₂ (s⊩isSymmetric x y r₁ r₁⊩x~y) r₂⊩x~y'))) +isTotal (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + return (Id , (λ x r r⊩x~x → ∣ x , (subst (λ r → r ⊩ ∣ perX .equality ∣ (x , x)) (sym (Ida≡a r)) r⊩x~x) ∣₁)) + +composeFuncRel : + ∀ {X Y Z : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (F : FunctionalRelation perX perY) + → (G : FunctionalRelation perY perZ) + → FunctionalRelation perX perZ + +(relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = + record + { isSetX = isSet× (perX .isSetX) (perZ .isSetX) + ; ∣_∣ = λ { (x , z) r → ∃[ y ∈ Y ] (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (x , y) × (pr₂ ⨾ r) ⊩ ∣ G .relation ∣ (y , z) } + ; isPropValued = λ x a → isPropPropTrunc } +isStrict (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (stF , stF⊩isStrictF) ← F .isStrict + (stG , stG⊩isStrictG) ← G .isStrict + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` pr₁ ̇ (` stF ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ # fzero))) + return + (λ* prover , + (λ x z r r⊩∃y → + transport (propTruncIdempotent {!!}) + (do + (y , pr₁r⊩Fxy , pr₂r⊩Gxy) ← r⊩∃y let - proofEq : λ* prover ⨾ a ⨾ b ≡ pair ⨾ b ⨾ b - proofEq = - λ*ComputationRule prover (a ∷ b ∷ []) - - pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ a ⨾ b) ≡ b + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) pr₁proofEq = - pr₁ ⨾ (λ* prover ⨾ a ⨾ b) - ≡⟨ cong (λ x → pr₁ ⨾ x) proofEq ⟩ - pr₁ ⨾ (pair ⨾ b ⨾ b) - ≡⟨ pr₁pxy≡x b b ⟩ - b - ∎ - - pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ a ⨾ b) ≡ b - pr₂proofEq = - pr₂ ⨾ (λ* prover ⨾ a ⨾ b) - ≡⟨ cong (λ x → pr₂ ⨾ x) proofEq ⟩ - pr₂ ⨾ (pair ⨾ b ⨾ b) - ≡⟨ pr₂pxy≡y b b ⟩ - b + pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ (pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r))) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)))) + ≡⟨ pr₁pxy≡x _ _ ⟩ + pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) ∎ - in - (subst - (λ r → r pre⊩ ∣ (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX)) (λ γ → snd (fst γ) , snd γ) (ob .snd ._~_)) ∣ ((tt* , x') , y')) - (sym pr₁proofEq) b⊩x'~y') , - subst - (λ r → r pre⊩ ∣ (⋆_ (isSet× (isSet× isSetUnit* (ob .snd .isSetX)) (ob .snd .isSetX)) (isSet× (ob .snd .isSetX) (ob .snd .isSetX)) (λ γ → snd (fst γ) , snd γ) (ob .snd ._~_)) ∣ ((tt* , x') , y')) (sym pr₂proofEq) b⊩x'~y' })) - -RT : Category (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) ((ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ''))) + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)) + pr₂proofEq = {!!} + return + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (stF⊩isStrictF x y (pr₁ ⨾ r) pr₁r⊩Fxy .fst)) , + subst (λ r' → r' ⊩ ∣ perZ .equality ∣ (z , z)) (sym pr₂proofEq) (stG⊩isStrictG y z (pr₂ ⨾ r) pr₂r⊩Gxy .snd))))) +isRelational (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (rlF , rlF⊩isRelationalF) ← F .isRelational + (stF , stF⊩isStrictF) ← F .isStrict + (rlG , rlG⊩isRelationalG) ← G .isRelational + let + prover : ApplStrTerm as 1 + prover = + ` pair ̇ + (` rlF ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))))))) ̇ + (` rlG ̇ (` pair ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))))) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))))) + return + (λ* prover , + (λ x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' → + do + (y , pr₁b⊩Fxy , pr₂b⊩Gyz) ← b⊩∃y + let + proofEq : λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ≡ pair ⨾ (rlF ⨾ (pair ⨾ a ⨾ (pair ⨾ (pr₁ ⨾ b) ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b)))))) ⨾ (rlG ⨾ (pair ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) ⨾ (pair ⨾ (pr₂ ⨾ b) ⨾ c))) + proofEq = + λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) + ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ + {!!} + return + (y , + (subst (λ r → r ⊩ ∣ F .relation ∣ (x' , y)) (sym {!!}) (rlF⊩isRelationalF x x' y y a (pr₁ ⨾ b) (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) a⊩x~x' pr₁b⊩Fxy (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd)) , + (subst (λ r → r ⊩ ∣ G .relation ∣ (y , z')) (sym {!!}) (rlG⊩isRelationalG y y z z' (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) (pr₂ ⨾ b) c (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z')))))) +isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} +isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} + +RTMorphism : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → Type _ +RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G → pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F + +idRTMorphism : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → RTMorphism perX perX +idRTMorphism {X} perX = [ equalityFuncRel perX ] + +composeRTMorphism : + ∀ {X Y Z : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (f : RTMorphism perX perY) + → (g : RTMorphism perY perZ) + ---------------------------------------- + → RTMorphism perX perZ +composeRTMorphism {X} {Y} {Z} perX perY perZ f g = + setQuotRec2 + squash/ + (λ F G → [ composeFuncRel perX perY perZ F G ]) + (λ { F F' G (F≤F' , F'≤F) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) ({!!} , {!!}) }) + (λ { F G G' (G≤G' , G'≤G) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') ({!!} , {!!}) }) + f g + +idLRTMorphism : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (f : RTMorphism perX perY) + → composeRTMorphism perX perX perY (idRTMorphism perX) f ≡ f +idLRTMorphism {X} {Y} perX perY f = + setQuotElimProp + (λ f → squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f) + (λ f → eq/ _ _ ({!!} , {!!})) + f + +idRRTMorphism : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (f : RTMorphism perX perY) + → composeRTMorphism perX perY perY f (idRTMorphism perY) ≡ f +idRRTMorphism {X} {Y} perX perY f = + setQuotElimProp + (λ f → squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f) + (λ f → eq/ _ _ ({!!} , {!!})) + f + +assocRTMorphism : + ∀ {X Y Z W : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (perW : PartialEquivalenceRelation W) + → (f : RTMorphism perX perY) + → (g : RTMorphism perY perZ) + → (h : RTMorphism perZ perW) + → composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h ≡ composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h) +assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h = + setQuotElimProp3 + (λ f g h → + squash/ + (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h) + (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h))) + (λ f g h → eq/ _ _ ({!!} , {!!})) + f g h + +RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X -Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism X Y -Category.id RT {X , perX} = {!!} -Category._⋆_ RT = {!!} -Category.⋆IdL RT = {!!} -Category.⋆IdR RT = {!!} -Category.⋆Assoc RT = {!!} -Category.isSetHom RT = {!!} +Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism perX perY +Category.id RT {X , perX} = idRTMorphism perX +Category._⋆_ RT {X , perX} {y , perY} {Z , perZ} f g = composeRTMorphism perX perY perZ f g +Category.⋆IdL RT {X , perX} {Y , perY} f = idLRTMorphism perX perY f +Category.⋆IdR RT {X , perX} {Y , perY} f = idRRTMorphism perX perY f +Category.⋆Assoc RT {X , perX} {Y , perY} {Z , perZ} {W , perW} f g h = assocRTMorphism perX perY perZ perW f g h +Category.isSetHom RT = squash/ diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index ada3cfc..8054bf9 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -17,7 +17,7 @@ module Realizability.Topos.Object open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate'; _⊩_ to _pre⊩_) +open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') open import Realizability.Tripos.Prealgebra.Predicate.Properties ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -25,86 +25,15 @@ open Predicate' renaming (isSetX to isSetPredicateBase) open PredicateProperties open Morphism -Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} +private + Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field isSetX : isSet X - private - `X : Sort - `X = ↑ (X , isSetX) - - `X×X : Sort - `X×X = `X `× `X - - ~relSym : Vec Sort 1 - ~relSym = `X×X ∷ [] - - module X×XRelational = Relational {n = 1} ~relSym - open X×XRelational - + equality : Predicate (X × X) + isSymmetric : ∃[ s ∈ A ] (∀ x y r → r ⊩ ∣ equality ∣ (x , y) → (s ⨾ r) ⊩ ∣ equality ∣ (y , x)) + open PredicateProperties {ℓ'' = ℓ''} (X × X) field - _~_ : Predicate ⟨ ⟦ lookup fzero ~relSym ⟧ˢ ⟩ - - private - ~relSymInterpretation : RelationInterpretation ~relSym - ~relSymInterpretation = λ { fzero → _~_ } + isTransitive : ∃[ t ∈ A ] (∀ x y z a b → a ⊩ ∣ equality ∣ (x , y) → b ⊩ ∣ equality ∣ (y , z) → (t ⨾ (pair ⨾ a ⨾ b)) ⊩ ∣ equality ∣ (x , z)) - module ~Interpretation = Interpretation ~relSym ~relSymInterpretation isNonTrivial - open ~Interpretation - - module ~Soundness = Soundness isNonTrivial ~relSymInterpretation - open ~Soundness - - -- Partial equivalence relations - - private - symContext : Context - symContext = ([] ′ `X) ′ `X - - x∈symContext : `X ∈ symContext - x∈symContext = there here - - y∈symContext : `X ∈ symContext - y∈symContext = here - - xˢ : Term symContext `X - xˢ = var x∈symContext - - yˢ : Term symContext `X - yˢ = var y∈symContext - - symmetryFormula : Formula symContext - symmetryFormula = rel fzero (xˢ `, yˢ) `→ rel fzero (yˢ `, xˢ) - - field - symmetry : holdsInTripos symmetryFormula - - private - transContext : Context - transContext = (([] ′ `X) ′ `X) ′ `X - - z∈transContext : `X ∈ transContext - z∈transContext = here - - y∈transContext : `X ∈ transContext - y∈transContext = there here - - x∈transContext : `X ∈ transContext - x∈transContext = there (there here) - - zᵗ : Term transContext `X - zᵗ = var z∈transContext - - yᵗ : Term transContext `X - yᵗ = var y∈transContext - - xᵗ : Term transContext `X - xᵗ = var x∈transContext - - transitivityFormula : Formula transContext - transitivityFormula = (rel fzero (xᵗ `, yᵗ) `∧ rel fzero (yᵗ `, zᵗ)) `→ rel fzero (xᵗ `, zᵗ) - - field - transitivity : holdsInTripos transitivityFormula - From 8f196fc64c81b1f3f5b827f575851500188297b3 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 13 Feb 2024 20:06:50 +0530 Subject: [PATCH 19/61] Finish goals --- .../Topos/FunctionalRelation.agda | 110 +++++++++++++++++- 1 file changed, 104 insertions(+), 6 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index cb3ee63..8300ca7 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -1,4 +1,3 @@ -{-# OPTIONS --allow-unsolved-metas #-} open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude @@ -251,7 +250,7 @@ isStrict (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = return (λ* prover , (λ x z r r⊩∃y → - transport (propTruncIdempotent {!!}) + transport (propTruncIdempotent (isProp× (perX .equality .isPropValued (x , x) (pr₁ ⨾ (λ* prover ⨾ r))) ( perZ .equality .isPropValued (z , z) (pr₂ ⨾ (λ* prover ⨾ r))))) (do (y , pr₁r⊩Fxy , pr₂r⊩Gxy) ← r⊩∃y let @@ -293,10 +292,109 @@ isRelational (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} return (y , - (subst (λ r → r ⊩ ∣ F .relation ∣ (x' , y)) (sym {!!}) (rlF⊩isRelationalF x x' y y a (pr₁ ⨾ b) (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) a⊩x~x' pr₁b⊩Fxy (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd)) , - (subst (λ r → r ⊩ ∣ G .relation ∣ (y , z')) (sym {!!}) (rlG⊩isRelationalG y y z z' (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) (pr₂ ⨾ b) c (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z')))))) -isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} -isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} + (subst + (λ r → r ⊩ ∣ F .relation ∣ (x' , y)) + (sym {!!}) + (rlF⊩isRelationalF + x x' y y a + (pr₁ ⨾ b) + (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) + a⊩x~x' pr₁b⊩Fxy + (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd)) , + (subst + (λ r → r ⊩ ∣ G .relation ∣ (y , z')) + (sym {!!}) + (rlG⊩isRelationalG + y y z z' + (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) + (pr₂ ⨾ b) + c + (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z')))))) +isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + (svG , svG⊩isSingleValuedG) ← G .isSingleValued + (rlG , rlG⊩isRelationalG) ← G .isRelational + (stG , stG⊩isStrictG) ← G .isStrict + let + prover : ApplStrTerm as 1 + prover = + ` svG ̇ + (` pair ̇ + (` rlG ̇ + (` pair ̇ + (` svF ̇ + (` pair ̇ + (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ + (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ + (` pair ̇ + (` pr₂ ̇ (` pr₁ ̇ # fzero)) ̇ + (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))))))) ̇ + (` pr₂ ̇ (` pr₂ ̇ # fzero))) + return + (λ* prover , + (λ x z z' r₁ r₂ r₁⊩∃y r₂⊩∃y' → + transport + (propTruncIdempotent (perZ .equality .isPropValued (z , z') _)) + (do + (y , pr₁r₁⊩Fxy , pr₂r₁⊩Gyz) ← r₁⊩∃y + (y' , pr₁r₂⊩Fxy' , pr₂r₂⊩Gy'z) ← r₂⊩∃y' + return + (subst + (λ r → r ⊩ ∣ perZ .equality ∣ (z , z')) + (sym {!!}) + (svG⊩isSingleValuedG + y' z z' + (rlG ⨾ (pair ⨾ (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) ⨾ (pair ⨾ (pr₂ ⨾ r₁) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁)))))) (pr₂ ⨾ r₂) + (rlG⊩isRelationalG + y y' z z + (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) + (pr₂ ⨾ r₁) + (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁))) + (svF⊩isSingleValuedF + x y y' + (pr₁ ⨾ r₁) (pr₁ ⨾ r₂) + pr₁r₁⊩Fxy pr₁r₂⊩Fxy') + pr₂r₁⊩Gyz + (stG⊩isStrictG y z (pr₂ ⨾ r₁) pr₂r₁⊩Gyz .snd)) pr₂r₂⊩Gy'z))))) +isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (tlF , tlF⊩isTotalF) ← F .isTotal + (tlG , tlG⊩isTotalG) ← G .isTotal + (stF , stF⊩isStrictF) ← F .isStrict + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` tlF ̇ # fzero) ̇ (` tlG ̇ (` pr₂ ̇ (` stF ̇ (` tlF ̇ # fzero)))) + return + (λ* prover , + (λ x r r⊩x~x → + do + (y , tlF⨾r⊩Fxy) ← tlF⊩isTotalF x r r⊩x~x + (z , ⊩Gyz) ← tlG⊩isTotalG y (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) (stF⊩isStrictF x y (tlF ⨾ r) tlF⨾r⊩Fxy .snd) + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ tlF ⨾ r + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) + ≡⟨ pr₁pxy≡x _ _ ⟩ + tlF ⨾ r + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₂ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) + ≡⟨ pr₂pxy≡y _ _ ⟩ + tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) + ∎ + return + (z , + return + (y , + (subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym pr₁proofEq) tlF⨾r⊩Fxy , + (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , z)) (sym pr₂proofEq) ⊩Gyz)))))) RTMorphism : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → Type _ RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G → pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F From 2e1ee9710ddf833fc84c38639770b74b3e64eece Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sat, 17 Feb 2024 03:09:49 +0530 Subject: [PATCH 20/61] Add terminal object of RT --- .../Topos/FunctionalRelation.agda | 1 + src/Realizability/Topos/TerminalObject.agda | 128 ++++++++++++++++++ 2 files changed, 129 insertions(+) create mode 100644 src/Realizability/Topos/TerminalObject.agda diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 8300ca7..f455036 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -1,3 +1,4 @@ +{-# OPTIONS --allow-unsolved-metas #-} open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude diff --git a/src/Realizability/Topos/TerminalObject.agda b/src/Realizability/Topos/TerminalObject.agda new file mode 100644 index 0000000..eaeca3d --- /dev/null +++ b/src/Realizability/Topos/TerminalObject.agda @@ -0,0 +1,128 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Fin +open import Cubical.Data.Vec +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients renaming (elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2) +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.Terminal + +module Realizability.Topos.TerminalObject + {ℓ ℓ' ℓ''} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where + +open CombinatoryAlgebra ca +open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') +open import Realizability.Tripos.Prealgebra.Predicate.Properties ca +open import Realizability.Topos.Object {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial + +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open PartialEquivalenceRelation +open Predicate' renaming (isSetX to isSetPredicateBase) +private + Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + +terminalPartialEquivalenceRelation : PartialEquivalenceRelation Unit* +isSetX terminalPartialEquivalenceRelation = isSetUnit* +isSetPredicateBase (equality terminalPartialEquivalenceRelation) = isSet× isSetUnit* isSetUnit* +∣ equality terminalPartialEquivalenceRelation ∣ (tt* , tt*) r = Unit* +isPropValued (equality terminalPartialEquivalenceRelation) (tt* , tt*) r = isPropUnit* +isSymmetric terminalPartialEquivalenceRelation = return (k , (λ { tt* tt* _ tt* → tt* })) +isTransitive terminalPartialEquivalenceRelation = return (k , (λ { tt* tt* tt* _ _ tt* tt* → tt* })) + +open FunctionalRelation +-- I have officially taken the inlining too far +-- TODO : Refactor +isTerminalTerminalPartialEquivalenceRelation : ∀ {Y : Type ℓ'} → (perY : PartialEquivalenceRelation Y) → isContr (RTMorphism perY terminalPartialEquivalenceRelation) +isTerminalTerminalPartialEquivalenceRelation {Y} perY = + inhProp→isContr + [ record + { relation = + record + { isSetX = isSet× (perY .isSetX) isSetUnit* + ; ∣_∣ = λ { (y , tt*) r → r ⊩ ∣ perY .equality ∣ (y , y) } + ; isPropValued = λ { (y , tt*) r → perY .equality .isPropValued _ _ } } + ; isStrict = + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ # fzero ̇ # fzero + in + return + ((λ* prover) , + (λ { y tt* r r⊩y~y → + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) + (sym + (pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ r ⨾ r) + ≡⟨ pr₁pxy≡x _ _ ⟩ + r + ∎)) + r⊩y~y , + tt* })) + ; isRelational = + do + (trY , trY⊩isTransitiveY) ← perY .isTransitive + (smY , smY⊩isSymmetricY) ← perY .isSymmetric + let + prover : ApplStrTerm as 1 + prover = ` trY ̇ (` pair ̇ (` smY ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero)) + return + (λ* prover , + (λ { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt* → + let + proofEq : λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ≡ trY ⨾ (pair ⨾ (smY ⨾ a) ⨾ a) + proofEq = + λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) + ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ + (trY ⨾ (pair ⨾ (smY ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))) ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))) + ≡⟨ cong₂ (λ x y → trY ⨾ (pair ⨾ (smY ⨾ x) ⨾ y)) (pr₁pxy≡x _ _) (pr₁pxy≡x _ _) ⟩ + trY ⨾ (pair ⨾ (smY ⨾ a) ⨾ a) + ∎ + in + subst + (λ r → r ⊩ ∣ perY .equality ∣ (y' , y')) + (sym proofEq) + (trY⊩isTransitiveY y' y y' (smY ⨾ a) a (smY⊩isSymmetricY y y' a a⊩y~y') a⊩y~y') })) + ; isSingleValued = return (k , (λ { _ tt* tt* _ _ _ _ → tt* })) -- nice + ; isTotal = return (Id , (λ y r r⊩y~y → return (tt* , subst (λ r → r ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y))) + } ] + λ f g → + setQuotElimProp2 + (λ f g → squash/ f g) + (λ F G → + eq/ + F G + let + F≤G : pointwiseEntailment perY terminalPartialEquivalenceRelation F G + F≤G = + (do + (tlG , tlG⊩isTotalG) ← G .isTotal + (stF , stF⊩isStrictF) ← F .isStrict + let + prover : ApplStrTerm as 1 + prover = ` tlG ̇ (` pr₁ ̇ (` stF ̇ # fzero)) + return + (λ* prover , + (λ { y tt* r r⊩Fx → + transport + (propTruncIdempotent (G .relation .isPropValued _ _)) + (tlG⊩isTotalG y (pr₁ ⨾ (stF ⨾ r)) (stF⊩isStrictF y tt* r r⊩Fx .fst) + >>= λ { (tt* , ⊩Gy) → return (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , tt*)) (sym (λ*ComputationRule prover (r ∷ []))) ⊩Gy) }) }))) + in F≤G , (F≤G→G≤F perY terminalPartialEquivalenceRelation F G F≤G)) + f g + +TerminalRT : Terminal RT +TerminalRT = + (Unit* , terminalPartialEquivalenceRelation) , (λ { (Y , perY) → isTerminalTerminalPartialEquivalenceRelation perY}) From 4d7bc8ef8d9849d70d73ce2f2961294f5702719d Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sat, 24 Feb 2024 14:00:31 +0530 Subject: [PATCH 21/61] Add algebraic realisability relatiobn --- src/Realizability/Tripos/Algebra/Base.agda | 51 +++++++++++++++++++--- 1 file changed, 44 insertions(+), 7 deletions(-) diff --git a/src/Realizability/Tripos/Algebra/Base.agda b/src/Realizability/Tripos/Algebra/Base.agda index 4287b4a..55615b3 100644 --- a/src/Realizability/Tripos/Algebra/Base.agda +++ b/src/Realizability/Tripos/Algebra/Base.agda @@ -8,6 +8,7 @@ open import Cubical.Foundations.Univalence open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels open import Cubical.Functions.FunExtEquiv open import Cubical.Data.Fin open import Cubical.Data.Sum renaming (rec to sumRec) @@ -17,7 +18,7 @@ open import Cubical.Data.Empty renaming (elim to ⊥elim) open import Cubical.Data.Unit open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients renaming (rec to quotRec; rec2 to quotRec2) +open import Cubical.HITs.SetQuotients as SQ renaming (rec to quotRec; rec2 to quotRec2) open import Cubical.Relation.Binary.Order.Preorder open import Cubical.Relation.Binary.Order.Poset open import Cubical.Algebra.Lattice @@ -26,8 +27,8 @@ open import Cubical.Algebra.CommMonoid open import Cubical.Algebra.Monoid open import Cubical.Algebra.Semigroup -module Realizability.Tripos.Algebra.Base {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +module Realizability.Tripos.Algebra.Base {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +import Realizability.Tripos.Prealgebra.Predicate ca as Pred open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca open import Realizability.Tripos.Prealgebra.Joins.Identity ca open import Realizability.Tripos.Prealgebra.Joins.Idempotency ca @@ -41,10 +42,46 @@ open import Realizability.Tripos.Prealgebra.Absorbtion ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -λ*ComputationRule = `λ*ComputationRule as fefermanStructure -λ* = `λ* as fefermanStructure - -module AlgebraicProperties {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + +AlgebraicPredicate : Type ℓ' → Type _ +AlgebraicPredicate X = PosetReflection (Pred.PredicateProperties._≤_ {ℓ'' = ℓ''} X) + +infixl 50 _⊩[_]_ +opaque + _⊩[_]_ : ∀ {X : Type ℓ'} → A → AlgebraicPredicate X → X → Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) + r ⊩[ ϕ ] x = + ⟨ + SQ.rec + isSetHProp + (λ Ψ → (∃[ s ∈ A ] Pred.Predicate.∣ Ψ ∣ x (s ⨾ r)) , isPropPropTrunc) + (λ { Ψ Ξ (Ψ≤Ξ , Ξ≤Ψ) → + Σ≡Prop + (λ _ → isPropIsProp) + (hPropExt isPropPropTrunc isPropPropTrunc + (λ Ψholds → + do + (s , s⊩Ψ≤Ξ) ← Ψ≤Ξ + (p , p⊩Ψ) ← Ψholds + let + prover : Term as 1 + prover = ` s ̇ (` p ̇ # fzero) + return (λ* prover , subst (λ r' → Pred.Predicate.∣ Ξ ∣ x r') (sym (λ*ComputationRule prover (r ∷ []))) (s⊩Ψ≤Ξ x (p ⨾ r) p⊩Ψ))) + (λ Ξholds → + do + (s , s⊩Ξ≤Ψ) ← Ξ≤Ψ + (p , p⊩Ξ) ← Ξholds + let + prover : Term as 1 + prover = ` s ̇ (` p ̇ # fzero) + return (λ* prover , subst (λ r' → Pred.Predicate.∣ Ψ ∣ x r') (sym (λ*ComputationRule prover (r ∷ []))) (s⊩Ξ≤Ψ x (p ⨾ r) p⊩Ξ)))) }) + ϕ + ⟩ + +module AlgebraicProperties (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where + open Pred private PredicateX = Predicate {ℓ'' = ℓ''} X open Predicate open PredicateProperties {ℓ'' = ℓ''} X From b35fcdb61b9c42b70f702b130eee812870a8aab8 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sat, 24 Feb 2024 14:22:02 +0530 Subject: [PATCH 22/61] Start using algebraic predicates --- .../Topos/FunctionalRelation.agda | 463 ++---------------- src/Realizability/Topos/Object.agda | 15 +- 2 files changed, 35 insertions(+), 443 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index f455036..77c2bda 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -13,7 +13,7 @@ open import Cubical.Data.Empty open import Cubical.Data.Unit open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients renaming (rec to setQuotRec; rec2 to setQuotRec2; elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2; elimProp3 to setQuotElimProp3) +open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category module Realizability.Topos.FunctionalRelation @@ -23,25 +23,15 @@ module Realizability.Topos.FunctionalRelation (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where -open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} -open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') -open import Realizability.Tripos.Prealgebra.Predicate.Properties ca -open import Realizability.Tripos.Prealgebra.Meets.Identity ca +open import Realizability.Tripos.Algebra.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca renaming (AlgebraicPredicate to Predicate) open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open Predicate' renaming (isSetX to isSetPredicateBase) -open PredicateProperties -open Morphism private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure -private - Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} - open PartialEquivalenceRelation record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where @@ -50,425 +40,36 @@ record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X field relation : Predicate (X × Y) - isStrict : - ∃[ st ∈ A ] - (∀ x y r - → r ⊩ ∣ relation ∣ (x , y) - ------------------------------------------------------------------------------------ - → (pr₁ ⨾ (st ⨾ r)) ⊩ ∣ equalityX ∣ (x , x) × ((pr₂ ⨾ (st ⨾ r)) ⊩ ∣ equalityY ∣ (y , y))) - isRelational : - ∃[ rl ∈ A ] - (∀ x x' y y' a b c - → a ⊩ ∣ equalityX ∣ (x , x') - → b ⊩ ∣ relation ∣ (x , y) - → c ⊩ ∣ equalityY ∣ (y , y') - ------------------------------------------ - → (rl ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ⊩ ∣ relation ∣ (x' , y')) - isSingleValued : - ∃[ sv ∈ A ] - (∀ x y y' r₁ r₂ - → r₁ ⊩ ∣ relation ∣ (x , y) - → r₂ ⊩ ∣ relation ∣ (x , y') - ----------------------------------- - → (sv ⨾ (pair ⨾ r₁ ⨾ r₂)) ⊩ ∣ equalityY ∣ (y , y')) - isTotal : - ∃[ tl ∈ A ] - (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) - -open FunctionalRelation - -pointwiseEntailment : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → (F G : FunctionalRelation perX perY) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') -pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe ∈ A ] (∀ x y r → r ⊩ ∣ F .relation ∣ (x , y) → (pe ⨾ r) ⊩ ∣ G .relation ∣ (x , y)) - --- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong --- Lemma 4.3.5 -F≤G→G≤F : - ∀ {X Y : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (F G : FunctionalRelation perX perY) - → pointwiseEntailment perX perY F G - → pointwiseEntailment perX perY G F -F≤G→G≤F {X} {Y} perX perY F G F≤G = - do - (r , r⊩F≤G) ← F≤G - (stG , stG⊩isStrictG) ← G .isStrict - (tlF , tlF⊩isTotalF) ← F .isTotal - (svG , svG⊩isSingleValuedG) ← G .isSingleValued - (rlF , rlF⊩isRelational) ← F .isRelational - let - prover : ApplStrTerm as 1 - prover = ` rlF ̇ (` pair ̇ (` pr₁ ̇ (` stG ̇ # fzero)) ̇ (` pair ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero))) ̇ (` svG ̇ (` pair ̇ (` r ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero)))) ̇ # fzero)))) - return - (λ* prover , - (λ x y r' r'⊩Gxy → - subst - (λ r'' → r'' ⊩ ∣ F .relation ∣ (x , y)) - (sym (λ*ComputationRule prover (r' ∷ []))) - (transport - (propTruncIdempotent (F .relation .isPropValued (x , y) _)) - (do - (y' , Fxy') ← tlF⊩isTotalF x (pr₁ ⨾ (stG ⨾ r')) (stG⊩isStrictG x y r' r'⊩Gxy .fst) - return - (rlF⊩isRelational - x x y' y - (pr₁ ⨾ (stG ⨾ r')) - (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) - (svG ⨾ (pair ⨾ (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) ⨾ r')) - (stG⊩isStrictG x y r' r'⊩Gxy .fst) - Fxy' - (svG⊩isSingleValuedG x y' y (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) r' (r⊩F≤G x y' (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) Fxy') r'⊩Gxy)))))) - -equalityFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX -relation (equalityFuncRel {X} perX) = perX .equality -isStrict (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` t ̇ (` pair ̇ # fzero ̇ (` s ̇ # fzero))) ̇ (` t ̇ (` pair ̇ (` s ̇ # fzero) ̇ # fzero)) - return - (λ* prover , - λ x x' r r⊩x~x' → - let - pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) - pr₁proofEq = - pr₁ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₁ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) - ≡⟨ pr₁pxy≡x _ _ ⟩ - t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) - ∎ - - pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) - pr₂proofEq = - pr₂ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₂ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) - ≡⟨ pr₂pxy≡y _ _ ⟩ - t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) - ∎ - in - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (t⊩isTransitive x x' x r (s ⨾ r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')) , - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) (sym pr₂proofEq) (t⊩isTransitive x' x x' (s ⨾ r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x')) -isRelational (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` t ̇ (` pair ̇ (` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) - - return - (λ* prover , - (λ x x' y y' a b c a⊩x~x' b⊩x~y c⊩y~y' → - let - proofNormalForm : (λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ≡ (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c)) - proofNormalForm = - λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) - ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ - (t ⨾ - (pair ⨾ - (t ⨾ - (pair ⨾ (s ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))) ⨾ - (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) - ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) - ≡⟨ - cong₂ - (λ x y → (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ x) ⨾ y)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))))) - (pr₁pxy≡x _ _) - (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) - ≡⟨ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ⟩ - pr₁ ⨾ (pair ⨾ b ⨾ c) - ≡⟨ pr₁pxy≡x _ _ ⟩ - b - ∎) - ⟩ - t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))) - ≡⟨ - cong - (λ x → t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ x)) - (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) - ≡⟨ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ⟩ - pr₂ ⨾ (pair ⨾ b ⨾ c) - ≡⟨ pr₂pxy≡y _ _ ⟩ - c - ∎) - ⟩ - t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c) - ∎ - in - subst - (λ r → r ⊩ ∣ perX .equality ∣ (x' , y')) - (sym proofNormalForm) - (t⊩isTransitive x' y y' (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) c (t⊩isTransitive x' x y (s ⨾ a) b (s⊩isSymmetric x x' a a⊩x~x') b⊩x~y) c⊩y~y'))) -isSingleValued (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero)) - return - (λ* prover , - (λ x y y' r₁ r₂ r₁⊩x~y r₂⊩x~y' → - let - proofEq : λ* prover ⨾ (pair ⨾ r₁ ⨾ r₂) ≡ t ⨾ (pair ⨾ (s ⨾ r₁) ⨾ r₂) - proofEq = {!!} - in - subst - (λ r → r ⊩ ∣ perX .equality ∣ (y , y')) - (sym proofEq) - (t⊩isTransitive y x y' (s ⨾ r₁) r₂ (s⊩isSymmetric x y r₁ r₁⊩x~y) r₂⊩x~y'))) -isTotal (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - return (Id , (λ x r r⊩x~x → ∣ x , (subst (λ r → r ⊩ ∣ perX .equality ∣ (x , x)) (sym (Ida≡a r)) r⊩x~x) ∣₁)) - -composeFuncRel : - ∀ {X Y Z : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (perZ : PartialEquivalenceRelation Z) - → (F : FunctionalRelation perX perY) - → (G : FunctionalRelation perY perZ) - → FunctionalRelation perX perZ + isStrictDomain : ∃[ stD ∈ A ] (∀ x y r → r ⊩[ relation ] (x , y) → (stD ⨾ r) ⊩[ equalityX ] (x , x)) + isStrictCodomain : ∃[ stC ∈ A ] (∀ x y r → r ⊩[ relation ] (x , y) → (stC ⨾ r) ⊩[ equalityY ] (y , y)) + isExtensional : ∃[ ext ∈ A ] (∀ x x' y y' r s p → r ⊩[ equalityX ] (x , x') → s ⊩[ equalityY ] (y , y') → p ⊩[ relation ] (x , y) → (ext ⨾ r ⨾ s ⨾ p) ⊩[ relation ] (x' , y')) + isSingleValued : ∃[ sv ∈ A ] (∀ x y y' r s → r ⊩[ relation ] (x , y) → s ⊩[ relation ] (x , y') → (sv ⨾ r ⨾ s) ⊩[ equalityY ] (y , y')) + isTotal : ∃[ tl ∈ A ] (∀ x x' r → r ⊩[ equalityX ] (x , x') → ∃[ y ∈ Y ] (tl ⨾ r) ⊩[ relation ] (x , y)) -(relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = - record - { isSetX = isSet× (perX .isSetX) (perZ .isSetX) - ; ∣_∣ = λ { (x , z) r → ∃[ y ∈ Y ] (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (x , y) × (pr₂ ⨾ r) ⊩ ∣ G .relation ∣ (y , z) } - ; isPropValued = λ x a → isPropPropTrunc } -isStrict (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (stF , stF⊩isStrictF) ← F .isStrict - (stG , stG⊩isStrictG) ← G .isStrict - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` pr₁ ̇ (` stF ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ # fzero))) - return - (λ* prover , - (λ x z r r⊩∃y → - transport (propTruncIdempotent (isProp× (perX .equality .isPropValued (x , x) (pr₁ ⨾ (λ* prover ⨾ r))) ( perZ .equality .isPropValued (z , z) (pr₂ ⨾ (λ* prover ⨾ r))))) - (do - (y , pr₁r⊩Fxy , pr₂r⊩Gxy) ← r⊩∃y - let - pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) - pr₁proofEq = - pr₁ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₁ ⨾ (pair ⨾ (pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r))) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)))) - ≡⟨ pr₁pxy≡x _ _ ⟩ - pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) - ∎ - pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)) - pr₂proofEq = {!!} - return - ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (stF⊩isStrictF x y (pr₁ ⨾ r) pr₁r⊩Fxy .fst)) , - subst (λ r' → r' ⊩ ∣ perZ .equality ∣ (z , z)) (sym pr₂proofEq) (stG⊩isStrictG y z (pr₂ ⨾ r) pr₂r⊩Gxy .snd))))) -isRelational (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (rlF , rlF⊩isRelationalF) ← F .isRelational - (stF , stF⊩isStrictF) ← F .isStrict - (rlG , rlG⊩isRelationalG) ← G .isRelational - let - prover : ApplStrTerm as 1 - prover = - ` pair ̇ - (` rlF ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))))))) ̇ - (` rlG ̇ (` pair ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))))) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))))) - return - (λ* prover , - (λ x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' → - do - (y , pr₁b⊩Fxy , pr₂b⊩Gyz) ← b⊩∃y - let - proofEq : λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ≡ pair ⨾ (rlF ⨾ (pair ⨾ a ⨾ (pair ⨾ (pr₁ ⨾ b) ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b)))))) ⨾ (rlG ⨾ (pair ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) ⨾ (pair ⨾ (pr₂ ⨾ b) ⨾ c))) - proofEq = - λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) - ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ - {!!} - return - (y , - (subst - (λ r → r ⊩ ∣ F .relation ∣ (x' , y)) - (sym {!!}) - (rlF⊩isRelationalF - x x' y y a - (pr₁ ⨾ b) - (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) - a⊩x~x' pr₁b⊩Fxy - (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd)) , - (subst - (λ r → r ⊩ ∣ G .relation ∣ (y , z')) - (sym {!!}) - (rlG⊩isRelationalG - y y z z' - (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) - (pr₂ ⨾ b) - c - (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z')))))) -isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (svF , svF⊩isSingleValuedF) ← F .isSingleValued - (svG , svG⊩isSingleValuedG) ← G .isSingleValued - (rlG , rlG⊩isRelationalG) ← G .isRelational - (stG , stG⊩isStrictG) ← G .isStrict - let - prover : ApplStrTerm as 1 - prover = - ` svG ̇ - (` pair ̇ - (` rlG ̇ - (` pair ̇ - (` svF ̇ - (` pair ̇ - (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ - (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ - (` pair ̇ - (` pr₂ ̇ (` pr₁ ̇ # fzero)) ̇ - (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))))))) ̇ - (` pr₂ ̇ (` pr₂ ̇ # fzero))) - return - (λ* prover , - (λ x z z' r₁ r₂ r₁⊩∃y r₂⊩∃y' → - transport - (propTruncIdempotent (perZ .equality .isPropValued (z , z') _)) - (do - (y , pr₁r₁⊩Fxy , pr₂r₁⊩Gyz) ← r₁⊩∃y - (y' , pr₁r₂⊩Fxy' , pr₂r₂⊩Gy'z) ← r₂⊩∃y' - return - (subst - (λ r → r ⊩ ∣ perZ .equality ∣ (z , z')) - (sym {!!}) - (svG⊩isSingleValuedG - y' z z' - (rlG ⨾ (pair ⨾ (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) ⨾ (pair ⨾ (pr₂ ⨾ r₁) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁)))))) (pr₂ ⨾ r₂) - (rlG⊩isRelationalG - y y' z z - (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) - (pr₂ ⨾ r₁) - (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁))) - (svF⊩isSingleValuedF - x y y' - (pr₁ ⨾ r₁) (pr₁ ⨾ r₂) - pr₁r₁⊩Fxy pr₁r₂⊩Fxy') - pr₂r₁⊩Gyz - (stG⊩isStrictG y z (pr₂ ⨾ r₁) pr₂r₁⊩Gyz .snd)) pr₂r₂⊩Gy'z))))) -isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (tlF , tlF⊩isTotalF) ← F .isTotal - (tlG , tlG⊩isTotalG) ← G .isTotal - (stF , stF⊩isStrictF) ← F .isStrict - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` tlF ̇ # fzero) ̇ (` tlG ̇ (` pr₂ ̇ (` stF ̇ (` tlF ̇ # fzero)))) - return - (λ* prover , - (λ x r r⊩x~x → - do - (y , tlF⨾r⊩Fxy) ← tlF⊩isTotalF x r r⊩x~x - (z , ⊩Gyz) ← tlG⊩isTotalG y (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) (stF⊩isStrictF x y (tlF ⨾ r) tlF⨾r⊩Fxy .snd) - let - pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ tlF ⨾ r - pr₁proofEq = - pr₁ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₁ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) - ≡⟨ pr₁pxy≡x _ _ ⟩ - tlF ⨾ r - ∎ - - pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) - pr₂proofEq = - pr₂ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₂ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) - ≡⟨ pr₂pxy≡y _ _ ⟩ - tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) - ∎ - return - (z , - return - (y , - (subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym pr₁proofEq) tlF⨾r⊩Fxy , - (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , z)) (sym pr₂proofEq) ⊩Gyz)))))) - -RTMorphism : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → Type _ -RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G → pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F - -idRTMorphism : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → RTMorphism perX perX -idRTMorphism {X} perX = [ equalityFuncRel perX ] - -composeRTMorphism : - ∀ {X Y Z : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (perZ : PartialEquivalenceRelation Z) - → (f : RTMorphism perX perY) - → (g : RTMorphism perY perZ) - ---------------------------------------- - → RTMorphism perX perZ -composeRTMorphism {X} {Y} {Z} perX perY perZ f g = - setQuotRec2 - squash/ - (λ F G → [ composeFuncRel perX perY perZ F G ]) - (λ { F F' G (F≤F' , F'≤F) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) ({!!} , {!!}) }) - (λ { F G G' (G≤G' , G'≤G) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') ({!!} , {!!}) }) - f g - -idLRTMorphism : - ∀ {X Y : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (f : RTMorphism perX perY) - → composeRTMorphism perX perX perY (idRTMorphism perX) f ≡ f -idLRTMorphism {X} {Y} perX perY f = - setQuotElimProp - (λ f → squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f) - (λ f → eq/ _ _ ({!!} , {!!})) - f - -idRRTMorphism : - ∀ {X Y : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (f : RTMorphism perX perY) - → composeRTMorphism perX perY perY f (idRTMorphism perY) ≡ f -idRRTMorphism {X} {Y} perX perY f = - setQuotElimProp - (λ f → squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f) - (λ f → eq/ _ _ ({!!} , {!!})) - f - -assocRTMorphism : - ∀ {X Y Z W : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (perZ : PartialEquivalenceRelation Z) - → (perW : PartialEquivalenceRelation W) - → (f : RTMorphism perX perY) - → (g : RTMorphism perY perZ) - → (h : RTMorphism perZ perW) - → composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h ≡ composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h) -assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h = - setQuotElimProp3 - (λ f g h → - squash/ - (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h) - (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h))) - (λ f g h → eq/ _ _ ({!!} , {!!})) - f g h +open FunctionalRelation -RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) -Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X -Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism perX perY -Category.id RT {X , perX} = idRTMorphism perX -Category._⋆_ RT {X , perX} {y , perY} {Z , perZ} f g = composeRTMorphism perX perY perZ f g -Category.⋆IdL RT {X , perX} {Y , perY} f = idLRTMorphism perX perY f -Category.⋆IdR RT {X , perX} {Y , perY} f = idRRTMorphism perX perY f -Category.⋆Assoc RT {X , perX} {Y , perY} {Z , perZ} {W , perW} f g h = assocRTMorphism perX perY perZ perW f g h -Category.isSetHom RT = squash/ +opaque + idFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX + relation (idFuncRel {X} perX) = perX .equality + isStrictDomain (idFuncRel {X} perX) = + do + (sm , sm⊩isSymmetric) ← perX .isSymmetric + (ts , ts⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` ts ̇ # fzero ̇ (` sm ̇ # fzero) + return + (λ* prover , + λ x x' r r⊩x~x' → + subst + (λ r' → r' ⊩[ perX .equality ] (x , x)) + (sym (λ*ComputationRule prover (r ∷ []))) + (ts⊩isTransitive x x' x r (sm ⨾ r) r⊩x~x' (sm⊩isSymmetric x x' r r⊩x~x'))) + isStrictCodomain (idFuncRel {X} perX) = {!!} + isExtensional (idFuncRel {X} perX) = {!!} + isSingleValued (idFuncRel {X} perX) = {!!} + isTotal (idFuncRel {X} perX) = {!!} + + RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) + RT = {!!} diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 8054bf9..0a46eac 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -17,23 +17,14 @@ module Realizability.Topos.Object open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') -open import Realizability.Tripos.Prealgebra.Predicate.Properties ca +open import Realizability.Tripos.Algebra.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca renaming (AlgebraicPredicate to Predicate) open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open Predicate' renaming (isSetX to isSetPredicateBase) -open PredicateProperties -open Morphism - -private - Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field isSetX : isSet X equality : Predicate (X × X) - isSymmetric : ∃[ s ∈ A ] (∀ x y r → r ⊩ ∣ equality ∣ (x , y) → (s ⨾ r) ⊩ ∣ equality ∣ (y , x)) - open PredicateProperties {ℓ'' = ℓ''} (X × X) - field - isTransitive : ∃[ t ∈ A ] (∀ x y z a b → a ⊩ ∣ equality ∣ (x , y) → b ⊩ ∣ equality ∣ (y , z) → (t ⨾ (pair ⨾ a ⨾ b)) ⊩ ∣ equality ∣ (x , z)) + isSymmetric : ∃[ sm ∈ A ] (∀ x x' r → r ⊩[ equality ] (x , x') → (sm ⨾ r) ⊩[ equality ] (x' , x)) + isTransitive : ∃[ ts ∈ A ] (∀ x x' x'' r s → r ⊩[ equality ] (x , x') → s ⊩[ equality ] (x' , x'') → (ts ⨾ r ⨾ s) ⊩[ equality ] (x , x'')) From 8d69eab3e88503913200820d6d8d17c05896e070 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 27 Feb 2024 00:33:46 +0530 Subject: [PATCH 23/61] Archive commit --- src/Experiments/Counterexample.agda | 10 ++ .../Topos/FunctionalRelation.agda | 101 ++++++++++++++++-- src/Realizability/Topos/Object.agda | 57 +++++++++- src/Realizability/Tripos/Algebra/Base.agda | 24 ++++- .../Prealgebra/Predicate/Properties.agda | 61 +++++++++-- 5 files changed, 228 insertions(+), 25 deletions(-) create mode 100644 src/Experiments/Counterexample.agda diff --git a/src/Experiments/Counterexample.agda b/src/Experiments/Counterexample.agda new file mode 100644 index 0000000..e02a410 --- /dev/null +++ b/src/Experiments/Counterexample.agda @@ -0,0 +1,10 @@ +module Experiments.Counterexample where + +open import Cubical.Foundations.Prelude +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Data.Nat +open import Cubical.Data.Nat.IsEven +open import Cubical.Data.Sigma +open import Cubical.Data.Bool +open import Cubical.Data.Sum as Sum + diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 77c2bda..f615895 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -4,6 +4,7 @@ open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence open import Cubical.Data.Vec open import Cubical.Data.Nat open import Cubical.Data.FinData @@ -11,7 +12,7 @@ open import Cubical.Data.Fin hiding (Fin; _/_) open import Cubical.Data.Sigma open import Cubical.Data.Empty open import Cubical.Data.Unit -open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation as PT open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category @@ -23,8 +24,9 @@ module Realizability.Topos.FunctionalRelation (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where +open import Realizability.Tripos.Prealgebra.Predicate ca as Pred hiding (Predicate) open import Realizability.Tripos.Algebra.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca renaming (AlgebraicPredicate to Predicate) -open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -66,10 +68,95 @@ opaque (λ r' → r' ⊩[ perX .equality ] (x , x)) (sym (λ*ComputationRule prover (r ∷ []))) (ts⊩isTransitive x x' x r (sm ⨾ r) r⊩x~x' (sm⊩isSymmetric x x' r r⊩x~x'))) - isStrictCodomain (idFuncRel {X} perX) = {!!} - isExtensional (idFuncRel {X} perX) = {!!} - isSingleValued (idFuncRel {X} perX) = {!!} - isTotal (idFuncRel {X} perX) = {!!} + isStrictCodomain (idFuncRel {X} perX) = + do + (sm , sm⊩isSymmetric) ← perX .isSymmetric + (ts , ts⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` ts ̇ (` sm ̇ # fzero) ̇ # fzero + return + ((λ* prover) , + (λ x x' r r⊩x~x' → subst (λ r' → r' ⊩[ perX .equality ] (x' , x')) (sym (λ*ComputationRule prover (r ∷ []))) (ts⊩isTransitive x' x x' (sm ⨾ r) r (sm⊩isSymmetric x x' r r⊩x~x') r⊩x~x'))) + isExtensional (idFuncRel {X} perX) = + do + (sm , sm⊩isSymmetric) ← perX .isSymmetric + (ts , ts⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 3 + prover = ` ts ̇ (` ts ̇ (` sm ̇ # 0) ̇ # 2) ̇ # 1 + return + (λ* prover , + (λ x₁ x₂ x₃ x₄ r s p r⊩x₁~x₂ s⊩x₃~x₄ p⊩x₁~x₃ → + subst + (λ r' → r' ⊩[ perX .equality ] (x₂ , x₄)) + (sym (λ*ComputationRule prover (r ∷ s ∷ p ∷ []))) + (ts⊩isTransitive + x₂ x₃ x₄ + (ts ⨾ (sm ⨾ r) ⨾ p) + s + (ts⊩isTransitive + x₂ x₁ x₃ + (sm ⨾ r) p + (sm⊩isSymmetric x₁ x₂ r r⊩x₁~x₂) + p⊩x₁~x₃) s⊩x₃~x₄))) + isSingleValued (idFuncRel {X} perX) = + do + (sm , sm⊩isSymmetric) ← perX .isSymmetric + (ts , ts⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 2 + prover = ` ts ̇ (` sm ̇ # 0) ̇ # 1 + return (λ* prover , (λ x x' x'' r s r⊩x~x' s⊩x~x'' → subst (λ r' → r' ⊩[ perX .equality ] (x' , x'')) (sym (λ*ComputationRule prover (r ∷ s ∷ []))) (ts⊩isTransitive x' x x'' (sm ⨾ r) s (sm⊩isSymmetric x x' r r⊩x~x') s⊩x~x''))) + isTotal (idFuncRel {X} perX) = + do + (sm , sm⊩isSymmetric) ← perX .isSymmetric + (ts , ts⊩isTransitive) ← perX .isTransitive + return (Id , (λ x x' r r⊩x~x' → ∣ x' , (subst (λ r' → r' ⊩[ perX .equality ] (x , x')) (sym (Ida≡a _)) r⊩x~x') ∣₁)) + +opaque + unfolding _⊩[_]_ + composeFuncRel : + ∀ {X Y Z : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → FunctionalRelation perX perY + → FunctionalRelation perY perZ + → FunctionalRelation perX perZ + relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + [ record + { isSetX = isSet× (perX .isSetX) (perZ .isSetX) + ; ∣_∣ = λ { (x , z) r → ∃[ y ∈ Y ] (pr₁ ⨾ r) ⊩[ F .relation ] (x , y) × (pr₂ ⨾ r) ⊩[ G .relation ] (y , z) } + ; isPropValued = λ { (x , z) r → isPropPropTrunc } } ] + isStrictDomain (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (stF , stF⊩isStrictF) ← F .isStrictDomain + return + (Id , + (λ x z r r⊩∃y → + transport + (propTruncIdempotent (isProp⊩ _ (perX .equality) (x , x))) + (do + (s , sr⊩) ← r⊩∃y + (y , pr₁sr⊩Fxy , pr₂sr⊩Gyz) ← sr⊩ + (eqX , [eqX]≡perX) ← []surjective (perX .equality) + return + (subst + (λ per → (Id ⨾ r) ⊩[ per ] (x , x)) + [eqX]≡perX + (transformRealizes (Id ⨾ r) eqX (x , x) {!!}))))) + isStrictCodomain (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} + isExtensional (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} + isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} + isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) - RT = {!!} + Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X + Category.Hom[_,_] RT (X , perX) (Y , perY) = FunctionalRelation perX perY + Category.id RT {X , perX} = idFuncRel perX + Category._⋆_ RT {X , perX} {Y , perY} {Z , perZ} F G = {!!} + Category.⋆IdL RT = {!!} + Category.⋆IdR RT = {!!} + Category.⋆Assoc RT = {!!} + Category.isSetHom RT = {!!} diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 0a46eac..7092fc5 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -2,11 +2,16 @@ open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.HITs.PropositionalTruncation open import Cubical.Data.Vec open import Cubical.Data.Nat -open import Cubical.Data.FinData renaming (zero to fzero) +open import Cubical.Data.FinData open import Cubical.Data.Sigma open import Cubical.Data.Empty +open import Cubical.Reflection.RecordEquiv module Realizability.Topos.Object {ℓ ℓ' ℓ''} @@ -21,10 +26,54 @@ open import Realizability.Tripos.Algebra.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca r open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where +record isPartialEquivalenceRelation (X : Type ℓ') (equality : Predicate (X × X)) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field isSetX : isSet X - equality : Predicate (X × X) isSymmetric : ∃[ sm ∈ A ] (∀ x x' r → r ⊩[ equality ] (x , x') → (sm ⨾ r) ⊩[ equality ] (x' , x)) isTransitive : ∃[ ts ∈ A ] (∀ x x' x'' r s → r ⊩[ equality ] (x , x') → s ⊩[ equality ] (x' , x'') → (ts ⨾ r ⨾ s) ⊩[ equality ] (x , x'')) - + +opaque + isPropIsPartialEquivalenceRelation : ∀ X equality → isProp (isPartialEquivalenceRelation X equality) + isPropIsPartialEquivalenceRelation X equality x y i = + record { isSetX = isPropIsSet {A = X} (x .isSetX) (y .isSetX) i ; isSymmetric = squash₁ (x .isSymmetric) (y .isSymmetric) i ; isTransitive = squash₁ (x .isTransitive) (y .isTransitive) i } where + open isPartialEquivalenceRelation + +record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + field + equality : Predicate (X × X) + isPerEquality : isPartialEquivalenceRelation X equality + open isPartialEquivalenceRelation isPerEquality public + +open PartialEquivalenceRelation + +unquoteDecl PartialEquivalenceRelationIsoΣ = declareRecordIsoΣ PartialEquivalenceRelationIsoΣ (quote PartialEquivalenceRelation) + +PartialEquivalenceRelationΣ : (X : Type ℓ') → Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) +PartialEquivalenceRelationΣ X = Σ[ equality ∈ Predicate (X × X) ] isPartialEquivalenceRelation X equality + + +module _ (X : Type ℓ') where opaque + open Iso + PartialEquivalenceRelationΣ≡ : (perA perB : PartialEquivalenceRelationΣ X) → perA .fst ≡ perB .fst → perA ≡ perB + PartialEquivalenceRelationΣ≡ perA perB predicateEq = Σ≡Prop (λ x → isPropIsPartialEquivalenceRelation X x) predicateEq + + PartialEquivalenceRelationΣ≃ : (perA perB : PartialEquivalenceRelationΣ X) → (perA .fst ≡ perB .fst) ≃ (perA ≡ perB) + PartialEquivalenceRelationΣ≃ perA perB = Σ≡PropEquiv λ x → isPropIsPartialEquivalenceRelation X x + + PartialEquivalenceRelationIso : (perA perB : PartialEquivalenceRelation X) → Iso (Iso.fun PartialEquivalenceRelationIsoΣ perA ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB) (perA ≡ perB) + Iso.fun (PartialEquivalenceRelationIso perA perB) p i = Iso.inv PartialEquivalenceRelationIsoΣ (p i) + inv (PartialEquivalenceRelationIso perA perB) = cong (λ x → Iso.fun PartialEquivalenceRelationIsoΣ x) + rightInv (PartialEquivalenceRelationIso perA perB) b = refl + leftInv (PartialEquivalenceRelationIso perA perB) a = refl + + -- Main SIP + PartialEquivalenceRelation≃ : (perA perB : PartialEquivalenceRelation X) → (perA .equality ≡ perB .equality) ≃ (perA ≡ perB) + PartialEquivalenceRelation≃ perA perB = + perA .equality ≡ perB .equality + ≃⟨ idEquiv (perA .equality ≡ perB .equality) ⟩ + Iso.fun PartialEquivalenceRelationIsoΣ perA .fst ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB .fst + ≃⟨ PartialEquivalenceRelationΣ≃ (Iso.fun PartialEquivalenceRelationIsoΣ perA) (Iso.fun PartialEquivalenceRelationIsoΣ perB) ⟩ + Iso.fun PartialEquivalenceRelationIsoΣ perA ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB + ≃⟨ isoToEquiv (PartialEquivalenceRelationIso perA perB) ⟩ + perA ≡ perB + ■ diff --git a/src/Realizability/Tripos/Algebra/Base.agda b/src/Realizability/Tripos/Algebra/Base.agda index 55615b3..a1d63ba 100644 --- a/src/Realizability/Tripos/Algebra/Base.agda +++ b/src/Realizability/Tripos/Algebra/Base.agda @@ -51,10 +51,9 @@ AlgebraicPredicate X = PosetReflection (Pred.PredicateProperties._≤_ {ℓ'' = infixl 50 _⊩[_]_ opaque - _⊩[_]_ : ∀ {X : Type ℓ'} → A → AlgebraicPredicate X → X → Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) - r ⊩[ ϕ ] x = - ⟨ - SQ.rec + realizes : ∀ {X : Type ℓ'} → A → AlgebraicPredicate X → X → hProp (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) + realizes {X} r ϕ x = + SQ.rec isSetHProp (λ Ψ → (∃[ s ∈ A ] Pred.Predicate.∣ Ψ ∣ x (s ⨾ r)) , isPropPropTrunc) (λ { Ψ Ξ (Ψ≤Ξ , Ξ≤Ψ) → @@ -78,7 +77,22 @@ opaque prover = ` s ̇ (` p ̇ # fzero) return (λ* prover , subst (λ r' → Pred.Predicate.∣ Ψ ∣ x r') (sym (λ*ComputationRule prover (r ∷ []))) (s⊩Ξ≤Ψ x (p ⨾ r) p⊩Ξ)))) }) ϕ - ⟩ + + _⊩[_]_ : ∀ {X : Type ℓ'} → A → AlgebraicPredicate X → X → Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) + r ⊩[ ϕ ] x = ⟨ realizes r ϕ x ⟩ + + isProp⊩ : ∀ {X : Type ℓ'} → (a : A) → (ϕ : AlgebraicPredicate X) → (x : X) → isProp (a ⊩[ ϕ ] x) + isProp⊩ {X} a ϕ x = realizes a ϕ x .snd + + transformRealizes : ∀ {X : Type ℓ'} → (r : A) → (ϕ : Pred.Predicate X) → (x : X) → (∃[ s ∈ A ] (s ⨾ r) ⊩[ [ ϕ ] ] x) → r ⊩[ [ ϕ ] ] x + transformRealizes {X} r ϕ x ∃ = + do + (s , s⊩ϕx) ← ∃ + (p , ps⊩ϕx) ← s⊩ϕx + let + prover : Term as 1 + prover = ` p ̇ (` s ̇ # fzero) + return (λ* prover , subst (λ r' → Pred.Predicate.∣ ϕ ∣ x r') (sym (λ*ComputationRule prover (r ∷ []))) ps⊩ϕx) module AlgebraicProperties (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where open Pred diff --git a/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda b/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda index db3ff90..4a01d04 100644 --- a/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda +++ b/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda @@ -1,6 +1,6 @@ open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure -open import Cubical.Foundations.Prelude +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Cubical.Foundations.Prelude as P open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence @@ -11,20 +11,24 @@ open import Cubical.Data.Sigma open import Cubical.Data.Empty open import Cubical.Data.Unit open import Cubical.Data.Sum +open import Cubical.Data.Vec open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Relation.Binary.Order.Preorder module Realizability.Tripos.Prealgebra.Predicate.Properties - {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where open import Realizability.Tripos.Prealgebra.Predicate.Base ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) open Predicate -module PredicateProperties {ℓ' ℓ''} (X : Type ℓ') where +private λ*ComputationRule = `λ*ComputationRule as fefermanStructure +private λ* = `λ* as fefermanStructure + +module PredicateProperties (X : Type ℓ') where private PredicateX = Predicate {ℓ'' = ℓ''} X open Predicate _≤_ : Predicate {ℓ'' = ℓ''} X → Predicate {ℓ'' = ℓ''} X → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') @@ -83,12 +87,52 @@ module PredicateProperties {ℓ' ℓ''} (X : Type ℓ') where ∣ ϕ ⇒ ψ ∣ x a = ∀ b → (b ⊩ ∣ ϕ ∣ x) → (a ⨾ b) ⊩ ∣ ψ ∣ x (ϕ ⇒ ψ) .isPropValued x a = isPropΠ λ a → isPropΠ λ a⊩ϕx → ψ .isPropValued _ _ +module _ where + open PredicateProperties Unit* + private + Predicate' = Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} + module NotAntiSym (antiSym : ∀ (a b : Predicate Unit*) → (a≤b : a ≤ b) → (b≤a : b ≤ a) → a ≡ b) where + Lift' = Lift {i = ℓ} {j = (ℓ-max ℓ' ℓ'')} + + kRealized : Predicate' Unit* + kRealized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a → Lift' (a ≡ k) ; isPropValued = λ x a → isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k) } + + k'Realized : Predicate' Unit* + k'Realized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a → Lift' (a ≡ k') ; isPropValued = λ x a → isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k') } + + kRealized≤k'Realized : kRealized ≤ k'Realized + kRealized≤k'Realized = + do + let + prover : ApplStrTerm as 1 + prover = ` k' + return (λ* prover , λ { x a (lift a≡k) → lift (λ*ComputationRule prover (a ∷ [])) }) + + k'Realized≤kRealized : k'Realized ≤ kRealized + k'Realized≤kRealized = + do + let + prover : ApplStrTerm as 1 + prover = ` k + return (λ* prover , λ { x a (lift a≡k') → lift (λ*ComputationRule prover (a ∷ [])) }) + + kRealized≡k'Realized : kRealized ≡ k'Realized + kRealized≡k'Realized = antiSym kRealized k'Realized kRealized≤k'Realized k'Realized≤kRealized + + Lift≡ : Lift' (k ≡ k) ≡ Lift' (k ≡ k') + Lift≡ i = ∣ kRealized≡k'Realized i ∣ tt* k + + Liftk≡k' : Lift' (k ≡ k') + Liftk≡k' = transport Lift≡ (lift refl) + + k≡k' : k ≡ k' + k≡k' = Liftk≡k' .lower -module Morphism {ℓ' ℓ''} {X Y : Type ℓ'} (isSetX : isSet X) (isSetY : isSet Y) where +module Morphism {X Y : Type ℓ'} (isSetX : isSet X) (isSetY : isSet Y) where PredicateX = Predicate {ℓ'' = ℓ''} X PredicateY = Predicate {ℓ'' = ℓ''} Y - module PredicatePropertiesX = PredicateProperties {ℓ'' = ℓ''} X - module PredicatePropertiesY = PredicateProperties {ℓ'' = ℓ''} Y + module PredicatePropertiesX = PredicateProperties X + module PredicatePropertiesY = PredicateProperties Y open PredicatePropertiesX renaming (_≤_ to _≤X_ ; isProp≤ to isProp≤X) open PredicatePropertiesY renaming (_≤_ to _≤Y_ ; isProp≤ to isProp≤Y) open Predicate hiding (isSetX) @@ -232,7 +276,6 @@ module Morphism {ℓ' ℓ''} {X Y : Type ℓ'} (isSetX : isSet X) (isSetY : isSe -- The proof is trivial but I am the reader it was left to as an exercise module BeckChevalley - {ℓ' ℓ'' : Level} (I J K : Type ℓ') (isSetI : isSet I) (isSetJ : isSet J) @@ -240,7 +283,7 @@ module BeckChevalley (f : J → I) (g : K → I) where - module Morphism' = Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''} + module Morphism' = Morphism open Morphism' L = Σ[ k ∈ K ] Σ[ j ∈ J ] (g k ≡ f j) From 026a925ccd9d2f0f2f60e1d5d18c594630cf3c35 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 27 Feb 2024 20:08:43 +0530 Subject: [PATCH 24/61] Revert to preorder predicates --- .../Topos/FunctionalRelation.agda | 546 ++++++++++++++---- src/Realizability/Topos/Object.agda | 65 +-- src/Realizability/Tripos/Logic/Semantics.agda | 20 +- .../Tripos/Prealgebra/Implication.agda | 10 +- .../Prealgebra/Joins/Commutativity.agda | 8 +- .../Tripos/Prealgebra/Joins/Identity.agda | 10 +- .../Prealgebra/Meets/Commutativity.agda | 10 +- .../Tripos/Prealgebra/Meets/Identity.agda | 12 +- .../Tripos/Prealgebra/Predicate.agda | 6 +- .../Tripos/Prealgebra/Predicate/Base.agda | 31 +- .../Prealgebra/Predicate/Properties.agda | 18 +- 11 files changed, 497 insertions(+), 239 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index f615895..14f9883 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -4,7 +4,6 @@ open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Univalence open import Cubical.Data.Vec open import Cubical.Data.Nat open import Cubical.Data.FinData @@ -12,9 +11,9 @@ open import Cubical.Data.Fin hiding (Fin; _/_) open import Cubical.Data.Sigma open import Cubical.Data.Empty open import Cubical.Data.Unit -open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients as SQ +open import Cubical.HITs.SetQuotients renaming (rec to setQuotRec; rec2 to setQuotRec2; elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2; elimProp3 to setQuotElimProp3) open import Cubical.Categories.Category module Realizability.Topos.FunctionalRelation @@ -24,12 +23,17 @@ module Realizability.Topos.FunctionalRelation (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where -open import Realizability.Tripos.Prealgebra.Predicate ca as Pred hiding (Predicate) -open import Realizability.Tripos.Algebra.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca renaming (AlgebraicPredicate to Predicate) -open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} +open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure @@ -42,121 +46,425 @@ record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X field relation : Predicate (X × Y) - isStrictDomain : ∃[ stD ∈ A ] (∀ x y r → r ⊩[ relation ] (x , y) → (stD ⨾ r) ⊩[ equalityX ] (x , x)) - isStrictCodomain : ∃[ stC ∈ A ] (∀ x y r → r ⊩[ relation ] (x , y) → (stC ⨾ r) ⊩[ equalityY ] (y , y)) - isExtensional : ∃[ ext ∈ A ] (∀ x x' y y' r s p → r ⊩[ equalityX ] (x , x') → s ⊩[ equalityY ] (y , y') → p ⊩[ relation ] (x , y) → (ext ⨾ r ⨾ s ⨾ p) ⊩[ relation ] (x' , y')) - isSingleValued : ∃[ sv ∈ A ] (∀ x y y' r s → r ⊩[ relation ] (x , y) → s ⊩[ relation ] (x , y') → (sv ⨾ r ⨾ s) ⊩[ equalityY ] (y , y')) - isTotal : ∃[ tl ∈ A ] (∀ x x' r → r ⊩[ equalityX ] (x , x') → ∃[ y ∈ Y ] (tl ⨾ r) ⊩[ relation ] (x , y)) - + isStrict : + ∃[ st ∈ A ] + (∀ x y r + → r ⊩ ∣ relation ∣ (x , y) + ------------------------------------------------------------------------------------ + → (pr₁ ⨾ (st ⨾ r)) ⊩ ∣ equalityX ∣ (x , x) × ((pr₂ ⨾ (st ⨾ r)) ⊩ ∣ equalityY ∣ (y , y))) + isRelational : + ∃[ rl ∈ A ] + (∀ x x' y y' a b c + → a ⊩ ∣ equalityX ∣ (x , x') + → b ⊩ ∣ relation ∣ (x , y) + → c ⊩ ∣ equalityY ∣ (y , y') + ------------------------------------------ + → (rl ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ⊩ ∣ relation ∣ (x' , y')) + isSingleValued : + ∃[ sv ∈ A ] + (∀ x y y' r₁ r₂ + → r₁ ⊩ ∣ relation ∣ (x , y) + → r₂ ⊩ ∣ relation ∣ (x , y') + ----------------------------------- + → (sv ⨾ (pair ⨾ r₁ ⨾ r₂)) ⊩ ∣ equalityY ∣ (y , y')) + isTotal : + ∃[ tl ∈ A ] + (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) open FunctionalRelation -opaque - idFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX - relation (idFuncRel {X} perX) = perX .equality - isStrictDomain (idFuncRel {X} perX) = - do - (sm , sm⊩isSymmetric) ← perX .isSymmetric - (ts , ts⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` ts ̇ # fzero ̇ (` sm ̇ # fzero) - return - (λ* prover , - λ x x' r r⊩x~x' → - subst - (λ r' → r' ⊩[ perX .equality ] (x , x)) - (sym (λ*ComputationRule prover (r ∷ []))) - (ts⊩isTransitive x x' x r (sm ⨾ r) r⊩x~x' (sm⊩isSymmetric x x' r r⊩x~x'))) - isStrictCodomain (idFuncRel {X} perX) = - do - (sm , sm⊩isSymmetric) ← perX .isSymmetric - (ts , ts⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` ts ̇ (` sm ̇ # fzero) ̇ # fzero - return - ((λ* prover) , - (λ x x' r r⊩x~x' → subst (λ r' → r' ⊩[ perX .equality ] (x' , x')) (sym (λ*ComputationRule prover (r ∷ []))) (ts⊩isTransitive x' x x' (sm ⨾ r) r (sm⊩isSymmetric x x' r r⊩x~x') r⊩x~x'))) - isExtensional (idFuncRel {X} perX) = - do - (sm , sm⊩isSymmetric) ← perX .isSymmetric - (ts , ts⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 3 - prover = ` ts ̇ (` ts ̇ (` sm ̇ # 0) ̇ # 2) ̇ # 1 - return - (λ* prover , - (λ x₁ x₂ x₃ x₄ r s p r⊩x₁~x₂ s⊩x₃~x₄ p⊩x₁~x₃ → - subst - (λ r' → r' ⊩[ perX .equality ] (x₂ , x₄)) - (sym (λ*ComputationRule prover (r ∷ s ∷ p ∷ []))) - (ts⊩isTransitive - x₂ x₃ x₄ - (ts ⨾ (sm ⨾ r) ⨾ p) - s - (ts⊩isTransitive - x₂ x₁ x₃ - (sm ⨾ r) p - (sm⊩isSymmetric x₁ x₂ r r⊩x₁~x₂) - p⊩x₁~x₃) s⊩x₃~x₄))) - isSingleValued (idFuncRel {X} perX) = - do - (sm , sm⊩isSymmetric) ← perX .isSymmetric - (ts , ts⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 2 - prover = ` ts ̇ (` sm ̇ # 0) ̇ # 1 - return (λ* prover , (λ x x' x'' r s r⊩x~x' s⊩x~x'' → subst (λ r' → r' ⊩[ perX .equality ] (x' , x'')) (sym (λ*ComputationRule prover (r ∷ s ∷ []))) (ts⊩isTransitive x' x x'' (sm ⨾ r) s (sm⊩isSymmetric x x' r r⊩x~x') s⊩x~x''))) - isTotal (idFuncRel {X} perX) = - do - (sm , sm⊩isSymmetric) ← perX .isSymmetric - (ts , ts⊩isTransitive) ← perX .isTransitive - return (Id , (λ x x' r r⊩x~x' → ∣ x' , (subst (λ r' → r' ⊩[ perX .equality ] (x , x')) (sym (Ida≡a _)) r⊩x~x') ∣₁)) - -opaque - unfolding _⊩[_]_ - composeFuncRel : - ∀ {X Y Z : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (perZ : PartialEquivalenceRelation Z) - → FunctionalRelation perX perY - → FunctionalRelation perY perZ - → FunctionalRelation perX perZ - relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - [ record - { isSetX = isSet× (perX .isSetX) (perZ .isSetX) - ; ∣_∣ = λ { (x , z) r → ∃[ y ∈ Y ] (pr₁ ⨾ r) ⊩[ F .relation ] (x , y) × (pr₂ ⨾ r) ⊩[ G .relation ] (y , z) } - ; isPropValued = λ { (x , z) r → isPropPropTrunc } } ] - isStrictDomain (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (stF , stF⊩isStrictF) ← F .isStrictDomain - return - (Id , - (λ x z r r⊩∃y → - transport - (propTruncIdempotent (isProp⊩ _ (perX .equality) (x , x))) +pointwiseEntailment : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → (F G : FunctionalRelation perX perY) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') +pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe ∈ A ] (∀ x y r → r ⊩ ∣ F .relation ∣ (x , y) → (pe ⨾ r) ⊩ ∣ G .relation ∣ (x , y)) + +-- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong +-- Lemma 4.3.5 +F≤G→G≤F : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (F G : FunctionalRelation perX perY) + → pointwiseEntailment perX perY F G + → pointwiseEntailment perX perY G F +F≤G→G≤F {X} {Y} perX perY F G F≤G = + do + (r , r⊩F≤G) ← F≤G + (stG , stG⊩isStrictG) ← G .isStrict + (tlF , tlF⊩isTotalF) ← F .isTotal + (svG , svG⊩isSingleValuedG) ← G .isSingleValued + (rlF , rlF⊩isRelational) ← F .isRelational + let + prover : ApplStrTerm as 1 + prover = ` rlF ̇ (` pair ̇ (` pr₁ ̇ (` stG ̇ # fzero)) ̇ (` pair ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero))) ̇ (` svG ̇ (` pair ̇ (` r ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero)))) ̇ # fzero)))) + return + (λ* prover , + (λ x y r' r'⊩Gxy → + subst + (λ r'' → r'' ⊩ ∣ F .relation ∣ (x , y)) + (sym (λ*ComputationRule prover (r' ∷ []))) + (transport + (propTruncIdempotent (F .relation .isPropValued (x , y) _)) (do - (s , sr⊩) ← r⊩∃y - (y , pr₁sr⊩Fxy , pr₂sr⊩Gyz) ← sr⊩ - (eqX , [eqX]≡perX) ← []surjective (perX .equality) + (y' , Fxy') ← tlF⊩isTotalF x (pr₁ ⨾ (stG ⨾ r')) (stG⊩isStrictG x y r' r'⊩Gxy .fst) return - (subst - (λ per → (Id ⨾ r) ⊩[ per ] (x , x)) - [eqX]≡perX - (transformRealizes (Id ⨾ r) eqX (x , x) {!!}))))) - isStrictCodomain (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} - isExtensional (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} - isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} - isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!} - - RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) - Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X - Category.Hom[_,_] RT (X , perX) (Y , perY) = FunctionalRelation perX perY - Category.id RT {X , perX} = idFuncRel perX - Category._⋆_ RT {X , perX} {Y , perY} {Z , perZ} F G = {!!} - Category.⋆IdL RT = {!!} - Category.⋆IdR RT = {!!} - Category.⋆Assoc RT = {!!} - Category.isSetHom RT = {!!} + (rlF⊩isRelational + x x y' y + (pr₁ ⨾ (stG ⨾ r')) + (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) + (svG ⨾ (pair ⨾ (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) ⨾ r')) + (stG⊩isStrictG x y r' r'⊩Gxy .fst) + Fxy' + (svG⊩isSingleValuedG x y' y (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) r' (r⊩F≤G x y' (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) Fxy') r'⊩Gxy)))))) + +equalityFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX +relation (equalityFuncRel {X} perX) = perX .equality +isStrict (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` t ̇ (` pair ̇ # fzero ̇ (` s ̇ # fzero))) ̇ (` t ̇ (` pair ̇ (` s ̇ # fzero) ̇ # fzero)) + return + (λ* prover , + λ x x' r r⊩x~x' → + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) + ≡⟨ pr₁pxy≡x _ _ ⟩ + t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₂ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) + ≡⟨ pr₂pxy≡y _ _ ⟩ + t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) + ∎ + in + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (t⊩isTransitive x x' x r (s ⨾ r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')) , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) (sym pr₂proofEq) (t⊩isTransitive x' x x' (s ⨾ r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x')) +isRelational (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` t ̇ (` pair ̇ (` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) + + return + (λ* prover , + (λ x x' y y' a b c a⊩x~x' b⊩x~y c⊩y~y' → + let + proofNormalForm : (λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ≡ (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c)) + proofNormalForm = + λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) + ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ + (t ⨾ + (pair ⨾ + (t ⨾ + (pair ⨾ (s ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))) ⨾ + (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) + ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) + ≡⟨ + cong₂ + (λ x y → (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ x) ⨾ y)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))))) + (pr₁pxy≡x _ _) + (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) + ≡⟨ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ⟩ + pr₁ ⨾ (pair ⨾ b ⨾ c) + ≡⟨ pr₁pxy≡x _ _ ⟩ + b + ∎) + ⟩ + t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))) + ≡⟨ + cong + (λ x → t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ x)) + (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) + ≡⟨ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ⟩ + pr₂ ⨾ (pair ⨾ b ⨾ c) + ≡⟨ pr₂pxy≡y _ _ ⟩ + c + ∎) + ⟩ + t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c) + ∎ + in + subst + (λ r → r ⊩ ∣ perX .equality ∣ (x' , y')) + (sym proofNormalForm) + (t⊩isTransitive x' y y' (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) c (t⊩isTransitive x' x y (s ⨾ a) b (s⊩isSymmetric x x' a a⊩x~x') b⊩x~y) c⊩y~y'))) +isSingleValued (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + (λ x y y' r₁ r₂ r₁⊩x~y r₂⊩x~y' → + let + proofEq : λ* prover ⨾ (pair ⨾ r₁ ⨾ r₂) ≡ t ⨾ (pair ⨾ (s ⨾ r₁) ⨾ r₂) + proofEq = {!!} + in + subst + (λ r → r ⊩ ∣ perX .equality ∣ (y , y')) + (sym proofEq) + (t⊩isTransitive y x y' (s ⨾ r₁) r₂ (s⊩isSymmetric x y r₁ r₁⊩x~y) r₂⊩x~y'))) +isTotal (equalityFuncRel {X} perX) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + return (Id , (λ x r r⊩x~x → ∣ x , (subst (λ r → r ⊩ ∣ perX .equality ∣ (x , x)) (sym (Ida≡a r)) r⊩x~x) ∣₁)) + +composeFuncRel : + ∀ {X Y Z : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (F : FunctionalRelation perX perY) + → (G : FunctionalRelation perY perZ) + → FunctionalRelation perX perZ + +(relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = + record + { isSetX = isSet× (perX .isSetX) (perZ .isSetX) + ; ∣_∣ = λ { (x , z) r → ∃[ y ∈ Y ] (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (x , y) × (pr₂ ⨾ r) ⊩ ∣ G .relation ∣ (y , z) } + ; isPropValued = λ x a → isPropPropTrunc } +isStrict (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (stF , stF⊩isStrictF) ← F .isStrict + (stG , stG⊩isStrictG) ← G .isStrict + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` pr₁ ̇ (` stF ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ # fzero))) + return + (λ* prover , + (λ x z r r⊩∃y → + transport (propTruncIdempotent (isProp× (perX .equality .isPropValued (x , x) (pr₁ ⨾ (λ* prover ⨾ r))) ( perZ .equality .isPropValued (z , z) (pr₂ ⨾ (λ* prover ⨾ r))))) + (do + (y , pr₁r⊩Fxy , pr₂r⊩Gxy) ← r⊩∃y + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ (pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r))) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)))) + ≡⟨ pr₁pxy≡x _ _ ⟩ + pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)) + pr₂proofEq = {!!} + return + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (stF⊩isStrictF x y (pr₁ ⨾ r) pr₁r⊩Fxy .fst)) , + subst (λ r' → r' ⊩ ∣ perZ .equality ∣ (z , z)) (sym pr₂proofEq) (stG⊩isStrictG y z (pr₂ ⨾ r) pr₂r⊩Gxy .snd))))) +isRelational (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (rlF , rlF⊩isRelationalF) ← F .isRelational + (stF , stF⊩isStrictF) ← F .isStrict + (rlG , rlG⊩isRelationalG) ← G .isRelational + let + prover : ApplStrTerm as 1 + prover = + ` pair ̇ + (` rlF ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))))))) ̇ + (` rlG ̇ (` pair ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))))) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))))) + return + (λ* prover , + (λ x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' → + do + (y , pr₁b⊩Fxy , pr₂b⊩Gyz) ← b⊩∃y + let + proofEq : λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ≡ pair ⨾ (rlF ⨾ (pair ⨾ a ⨾ (pair ⨾ (pr₁ ⨾ b) ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b)))))) ⨾ (rlG ⨾ (pair ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) ⨾ (pair ⨾ (pr₂ ⨾ b) ⨾ c))) + proofEq = + λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) + ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ + {!!} + return + (y , + (subst + (λ r → r ⊩ ∣ F .relation ∣ (x' , y)) + (sym {!!}) + (rlF⊩isRelationalF + x x' y y a + (pr₁ ⨾ b) + (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) + a⊩x~x' pr₁b⊩Fxy + (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd)) , + (subst + (λ r → r ⊩ ∣ G .relation ∣ (y , z')) + (sym {!!}) + (rlG⊩isRelationalG + y y z z' + (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) + (pr₂ ⨾ b) + c + (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z')))))) +isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + (svG , svG⊩isSingleValuedG) ← G .isSingleValued + (rlG , rlG⊩isRelationalG) ← G .isRelational + (stG , stG⊩isStrictG) ← G .isStrict + let + prover : ApplStrTerm as 1 + prover = + ` svG ̇ + (` pair ̇ + (` rlG ̇ + (` pair ̇ + (` svF ̇ + (` pair ̇ + (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ + (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ + (` pair ̇ + (` pr₂ ̇ (` pr₁ ̇ # fzero)) ̇ + (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))))))) ̇ + (` pr₂ ̇ (` pr₂ ̇ # fzero))) + return + (λ* prover , + (λ x z z' r₁ r₂ r₁⊩∃y r₂⊩∃y' → + transport + (propTruncIdempotent (perZ .equality .isPropValued (z , z') _)) + (do + (y , pr₁r₁⊩Fxy , pr₂r₁⊩Gyz) ← r₁⊩∃y + (y' , pr₁r₂⊩Fxy' , pr₂r₂⊩Gy'z) ← r₂⊩∃y' + return + (subst + (λ r → r ⊩ ∣ perZ .equality ∣ (z , z')) + (sym {!!}) + (svG⊩isSingleValuedG + y' z z' + (rlG ⨾ (pair ⨾ (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) ⨾ (pair ⨾ (pr₂ ⨾ r₁) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁)))))) (pr₂ ⨾ r₂) + (rlG⊩isRelationalG + y y' z z + (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) + (pr₂ ⨾ r₁) + (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁))) + (svF⊩isSingleValuedF + x y y' + (pr₁ ⨾ r₁) (pr₁ ⨾ r₂) + pr₁r₁⊩Fxy pr₁r₂⊩Fxy') + pr₂r₁⊩Gyz + (stG⊩isStrictG y z (pr₂ ⨾ r₁) pr₂r₁⊩Gyz .snd)) pr₂r₂⊩Gy'z))))) +isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = + do + (tlF , tlF⊩isTotalF) ← F .isTotal + (tlG , tlG⊩isTotalG) ← G .isTotal + (stF , stF⊩isStrictF) ← F .isStrict + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` tlF ̇ # fzero) ̇ (` tlG ̇ (` pr₂ ̇ (` stF ̇ (` tlF ̇ # fzero)))) + return + (λ* prover , + (λ x r r⊩x~x → + do + (y , tlF⨾r⊩Fxy) ← tlF⊩isTotalF x r r⊩x~x + (z , ⊩Gyz) ← tlG⊩isTotalG y (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) (stF⊩isStrictF x y (tlF ⨾ r) tlF⨾r⊩Fxy .snd) + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ tlF ⨾ r + pr₁proofEq = + pr₁ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) + ≡⟨ pr₁pxy≡x _ _ ⟩ + tlF ⨾ r + ∎ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) + pr₂proofEq = + pr₂ ⨾ (λ* prover ⨾ r) + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ + pr₂ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) + ≡⟨ pr₂pxy≡y _ _ ⟩ + tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) + ∎ + return + (z , + return + (y , + (subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym pr₁proofEq) tlF⨾r⊩Fxy , + (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , z)) (sym pr₂proofEq) ⊩Gyz)))))) + +RTMorphism : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → Type _ +RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G → pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F + +idRTMorphism : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → RTMorphism perX perX +idRTMorphism {X} perX = [ equalityFuncRel perX ] + +composeRTMorphism : + ∀ {X Y Z : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (f : RTMorphism perX perY) + → (g : RTMorphism perY perZ) + ---------------------------------------- + → RTMorphism perX perZ +composeRTMorphism {X} {Y} {Z} perX perY perZ f g = + setQuotRec2 + squash/ + (λ F G → [ composeFuncRel perX perY perZ F G ]) + (λ { F F' G (F≤F' , F'≤F) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) ({!!} , {!!}) }) + (λ { F G G' (G≤G' , G'≤G) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') ({!!} , {!!}) }) + f g + +idLRTMorphism : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (f : RTMorphism perX perY) + → composeRTMorphism perX perX perY (idRTMorphism perX) f ≡ f +idLRTMorphism {X} {Y} perX perY f = + setQuotElimProp + (λ f → squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f) + (λ f → eq/ _ _ ({!!} , {!!})) + f + +idRRTMorphism : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (f : RTMorphism perX perY) + → composeRTMorphism perX perY perY f (idRTMorphism perY) ≡ f +idRRTMorphism {X} {Y} perX perY f = + setQuotElimProp + (λ f → squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f) + (λ f → eq/ _ _ ({!!} , {!!})) + f + +assocRTMorphism : + ∀ {X Y Z W : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (perW : PartialEquivalenceRelation W) + → (f : RTMorphism perX perY) + → (g : RTMorphism perY perZ) + → (h : RTMorphism perZ perW) + → composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h ≡ composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h) +assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h = + setQuotElimProp3 + (λ f g h → + squash/ + (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h) + (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h))) + (λ f g h → eq/ _ _ ({!!} , {!!})) + f g h + +RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) +Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X +Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism perX perY +Category.id RT {X , perX} = idRTMorphism perX +Category._⋆_ RT {X , perX} {y , perY} {Z , perZ} f g = composeRTMorphism perX perY perZ f g +Category.⋆IdL RT {X , perX} {Y , perY} f = idLRTMorphism perX perY f +Category.⋆IdR RT {X , perX} {Y , perY} f = idRRTMorphism perX perY f +Category.⋆Assoc RT {X , perX} {Y , perY} {Z , perZ} {W , perW} f g h = assocRTMorphism perX perY perZ perW f g h +Category.isSetHom RT = squash/ diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 7092fc5..60340a8 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -2,16 +2,11 @@ open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.HITs.PropositionalTruncation open import Cubical.Data.Vec open import Cubical.Data.Nat -open import Cubical.Data.FinData +open import Cubical.Data.FinData renaming (zero to fzero) open import Cubical.Data.Sigma open import Cubical.Data.Empty -open import Cubical.Reflection.RecordEquiv module Realizability.Topos.Object {ℓ ℓ' ℓ''} @@ -22,58 +17,18 @@ module Realizability.Topos.Object open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Tripos.Algebra.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca renaming (AlgebraicPredicate to Predicate) +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) - -record isPartialEquivalenceRelation (X : Type ℓ') (equality : Predicate (X × X)) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - field - isSetX : isSet X - isSymmetric : ∃[ sm ∈ A ] (∀ x x' r → r ⊩[ equality ] (x , x') → (sm ⨾ r) ⊩[ equality ] (x' , x)) - isTransitive : ∃[ ts ∈ A ] (∀ x x' x'' r s → r ⊩[ equality ] (x , x') → s ⊩[ equality ] (x' , x'') → (ts ⨾ r ⨾ s) ⊩[ equality ] (x , x'')) - -opaque - isPropIsPartialEquivalenceRelation : ∀ X equality → isProp (isPartialEquivalenceRelation X equality) - isPropIsPartialEquivalenceRelation X equality x y i = - record { isSetX = isPropIsSet {A = X} (x .isSetX) (y .isSetX) i ; isSymmetric = squash₁ (x .isSymmetric) (y .isSymmetric) i ; isTransitive = squash₁ (x .isTransitive) (y .isTransitive) i } where - open isPartialEquivalenceRelation +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field + isSetX : isSet X equality : Predicate (X × X) - isPerEquality : isPartialEquivalenceRelation X equality - open isPartialEquivalenceRelation isPerEquality public - -open PartialEquivalenceRelation - -unquoteDecl PartialEquivalenceRelationIsoΣ = declareRecordIsoΣ PartialEquivalenceRelationIsoΣ (quote PartialEquivalenceRelation) - -PartialEquivalenceRelationΣ : (X : Type ℓ') → Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) -PartialEquivalenceRelationΣ X = Σ[ equality ∈ Predicate (X × X) ] isPartialEquivalenceRelation X equality - - -module _ (X : Type ℓ') where opaque - open Iso - PartialEquivalenceRelationΣ≡ : (perA perB : PartialEquivalenceRelationΣ X) → perA .fst ≡ perB .fst → perA ≡ perB - PartialEquivalenceRelationΣ≡ perA perB predicateEq = Σ≡Prop (λ x → isPropIsPartialEquivalenceRelation X x) predicateEq - - PartialEquivalenceRelationΣ≃ : (perA perB : PartialEquivalenceRelationΣ X) → (perA .fst ≡ perB .fst) ≃ (perA ≡ perB) - PartialEquivalenceRelationΣ≃ perA perB = Σ≡PropEquiv λ x → isPropIsPartialEquivalenceRelation X x - - PartialEquivalenceRelationIso : (perA perB : PartialEquivalenceRelation X) → Iso (Iso.fun PartialEquivalenceRelationIsoΣ perA ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB) (perA ≡ perB) - Iso.fun (PartialEquivalenceRelationIso perA perB) p i = Iso.inv PartialEquivalenceRelationIsoΣ (p i) - inv (PartialEquivalenceRelationIso perA perB) = cong (λ x → Iso.fun PartialEquivalenceRelationIsoΣ x) - rightInv (PartialEquivalenceRelationIso perA perB) b = refl - leftInv (PartialEquivalenceRelationIso perA perB) a = refl - - -- Main SIP - PartialEquivalenceRelation≃ : (perA perB : PartialEquivalenceRelation X) → (perA .equality ≡ perB .equality) ≃ (perA ≡ perB) - PartialEquivalenceRelation≃ perA perB = - perA .equality ≡ perB .equality - ≃⟨ idEquiv (perA .equality ≡ perB .equality) ⟩ - Iso.fun PartialEquivalenceRelationIsoΣ perA .fst ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB .fst - ≃⟨ PartialEquivalenceRelationΣ≃ (Iso.fun PartialEquivalenceRelationIsoΣ perA) (Iso.fun PartialEquivalenceRelationIsoΣ perB) ⟩ - Iso.fun PartialEquivalenceRelationIsoΣ perA ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB - ≃⟨ isoToEquiv (PartialEquivalenceRelationIso perA perB) ⟩ - perA ≡ perB - ■ + isSymmetric : ∃[ s ∈ A ] (∀ x y r → r ⊩ ∣ equality ∣ (x , y) → (s ⨾ r) ⊩ ∣ equality ∣ (y , x)) + field + isTransitive : ∃[ t ∈ A ] (∀ x y z a b → a ⊩ ∣ equality ∣ (x , y) → b ⊩ ∣ equality ∣ (y , z) → (t ⨾ (pair ⨾ a ⨾ b)) ⊩ ∣ equality ∣ (x , z)) + diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index 231925d..d19a65a 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -27,18 +27,16 @@ open CombinatoryAlgebra ca private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure -open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') -open import Realizability.Tripos.Prealgebra.Predicate.Properties ca -open import Realizability.Tripos.Prealgebra.Meets.Identity ca -open import Realizability.Tripos.Prealgebra.Joins.Identity ca -open import Realizability.Tripos.Prealgebra.Implication ca +open import Realizability.Tripos.Prealgebra.Predicate.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Predicate.Properties {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Joins.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Implication {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open Predicate' +open Predicate open PredicateProperties hiding (_≤_ ; isTrans≤) -open Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''} -private - Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} +open Morphism RelationInterpretation : ∀ {n : ℕ} → (Vec Sort n) → Type _ RelationInterpretation {n} relSym = (∀ i → Predicate ⟨ ⟦ lookup i relSym ⟧ˢ ⟩) @@ -315,8 +313,8 @@ module Soundness (ϕ⊨ψ >>= λ { (a , a⊩ϕ≤ψ) → ∣ a , (λ γ b b⊩ϕsubsγ → a⊩ϕ≤ψ (⟦ subs ⟧ᴮ γ) b b⊩ϕsubsγ) ∣₁ }) where - open PredProps {ℓ'' = ℓ''} ⟨ ⟦ Γ ⟧ᶜ ⟩ renaming (_≤_ to _≤Γ_) - open PredProps {ℓ'' = ℓ''} ⟨ ⟦ Δ ⟧ᶜ ⟩ renaming (_≤_ to _≤Δ_) + open PredProps ⟨ ⟦ Γ ⟧ᶜ ⟩ renaming (_≤_ to _≤Γ_) + open PredProps ⟨ ⟦ Δ ⟧ᶜ ⟩ renaming (_≤_ to _≤Δ_) `∧intro : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ ψ → entails ϕ θ → entails ϕ (ψ `∧ θ) `∧intro {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ϕ⊨θ = diff --git a/src/Realizability/Tripos/Prealgebra/Implication.agda b/src/Realizability/Tripos/Prealgebra/Implication.agda index edeb3f2..f3d53d7 100644 --- a/src/Realizability/Tripos/Prealgebra/Implication.agda +++ b/src/Realizability/Tripos/Prealgebra/Implication.agda @@ -7,9 +7,9 @@ open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Data.Fin open import Cubical.Data.Vec -module Realizability.Tripos.Prealgebra.Implication {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +module Realizability.Tripos.Prealgebra.Implication {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -18,10 +18,10 @@ private λ*ComputationRule = `λ*ComputationRule as fefermanStructure λ* = `λ* as fefermanStructure -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where - PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) where + PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X -- ⇒ is Heyting implication a⊓b≤c→a≤b⇒c : ∀ a b c → (a ⊓ b ≤ c) → a ≤ (b ⇒ c) a⊓b≤c→a≤b⇒c a b c a⊓b≤c = diff --git a/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda b/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda index dac032a..5affaf6 100644 --- a/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda +++ b/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda @@ -11,8 +11,6 @@ open import Cubical.HITs.PropositionalTruncation.Monad module Realizability.Tripos.Prealgebra.Joins.Commutativity {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca - open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -20,10 +18,10 @@ private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where - - private PredicateX = Predicate {ℓ'' = ℓ''} X + open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca + private PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X open PreorderReasoning preorder≤ -- ⊔ is commutative upto anti-symmetry diff --git a/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda b/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda index 99675ac..f8cb0f1 100644 --- a/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda +++ b/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda @@ -12,8 +12,8 @@ open import Cubical.Relation.Binary.Order.Preorder open import Realizability.CombinatoryAlgebra open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) -module Realizability.Tripos.Prealgebra.Joins.Identity {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +module Realizability.Tripos.Prealgebra.Joins.Identity {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -21,10 +21,10 @@ open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a t private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - private PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where + private PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X open PreorderReasoning preorder≤ diff --git a/src/Realizability/Tripos/Prealgebra/Meets/Commutativity.agda b/src/Realizability/Tripos/Prealgebra/Meets/Commutativity.agda index 4e4b22b..ead32c9 100644 --- a/src/Realizability/Tripos/Prealgebra/Meets/Commutativity.agda +++ b/src/Realizability/Tripos/Prealgebra/Meets/Commutativity.agda @@ -9,9 +9,9 @@ open import Cubical.Relation.Binary.Order.Preorder open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -module Realizability.Tripos.Prealgebra.Meets.Commutativity {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +module Realizability.Tripos.Prealgebra.Meets.Commutativity {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -19,11 +19,11 @@ open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a t private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where +module _ (X : Type ℓ') (isSetX' : isSet X) where - private PredicateX = Predicate {ℓ'' = ℓ''} X + private PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X open PreorderReasoning preorder≤ x⊓y≤y⊓x : ∀ x y → x ⊓ y ≤ y ⊓ x diff --git a/src/Realizability/Tripos/Prealgebra/Meets/Identity.agda b/src/Realizability/Tripos/Prealgebra/Meets/Identity.agda index 6d6946a..ff41138 100644 --- a/src/Realizability/Tripos/Prealgebra/Meets/Identity.agda +++ b/src/Realizability/Tripos/Prealgebra/Meets/Identity.agda @@ -11,19 +11,19 @@ open import Cubical.Relation.Binary.Order.Preorder open import Realizability.CombinatoryAlgebra open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) -module Realizability.Tripos.Prealgebra.Meets.Identity {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca -open import Realizability.Tripos.Prealgebra.Meets.Commutativity ca +module Realizability.Tripos.Prealgebra.Meets.Identity {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Meets.Commutativity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - private PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where + private PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X open PreorderReasoning preorder≤ pre1 : PredicateX diff --git a/src/Realizability/Tripos/Prealgebra/Predicate.agda b/src/Realizability/Tripos/Prealgebra/Predicate.agda index a12e0b6..c99a67d 100644 --- a/src/Realizability/Tripos/Prealgebra/Predicate.agda +++ b/src/Realizability/Tripos/Prealgebra/Predicate.agda @@ -1,7 +1,7 @@ open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude -module Realizability.Tripos.Prealgebra.Predicate {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +module Realizability.Tripos.Prealgebra.Predicate {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate.Base ca public -open import Realizability.Tripos.Prealgebra.Predicate.Properties ca public +open import Realizability.Tripos.Prealgebra.Predicate.Base {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca public +open import Realizability.Tripos.Prealgebra.Predicate.Properties {ℓ' = ℓ'} {ℓ'' = ℓ''} ca public diff --git a/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda b/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda index 5850d7c..75790d8 100644 --- a/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda +++ b/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda @@ -7,12 +7,12 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Data.Sigma open import Cubical.Functions.FunExtEquiv -module Realizability.Tripos.Prealgebra.Predicate.Base {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +module Realizability.Tripos.Prealgebra.Predicate.Base {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -record Predicate {ℓ' ℓ''} (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where +record Predicate (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field isSetX : isSet X ∣_∣ : X → A → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') @@ -24,40 +24,39 @@ infix 26 _⊩_ _⊩_ : ∀ {ℓ'} → A → (A → Type ℓ') → Type ℓ' a ⊩ ϕ = ϕ a -PredicateΣ : ∀ {ℓ' ℓ''} → (X : Type ℓ') → Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) -PredicateΣ {ℓ'} {ℓ''} X = Σ[ rel ∈ (X → A → hProp (ℓ-max (ℓ-max ℓ ℓ') ℓ'')) ] (isSet X) +PredicateΣ : ∀ (X : Type ℓ') → Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) +PredicateΣ X = Σ[ rel ∈ (X → A → hProp (ℓ-max (ℓ-max ℓ ℓ') ℓ'')) ] (isSet X) -isSetPredicateΣ : ∀ {ℓ' ℓ''} (X : Type ℓ') → isSet (PredicateΣ {ℓ'' = ℓ''} X) +isSetPredicateΣ : ∀ (X : Type ℓ') → isSet (PredicateΣ X) isSetPredicateΣ X = isSetΣ (isSetΠ (λ x → isSetΠ λ a → isSetHProp)) λ _ → isProp→isSet isPropIsSet -PredicateIsoΣ : ∀ {ℓ' ℓ''} (X : Type ℓ') → Iso (Predicate {ℓ'' = ℓ''} X) (PredicateΣ {ℓ'' = ℓ''} X) -PredicateIsoΣ {ℓ'} {ℓ''} X = +PredicateIsoΣ : ∀ (X : Type ℓ') → Iso (Predicate X) (PredicateΣ X) +PredicateIsoΣ X = iso (λ p → (λ x a → (a ⊩ ∣ p ∣ x) , p .isPropValued x a) , p .isSetX) (λ p → record { isSetX = p .snd ; ∣_∣ = λ x a → p .fst x a .fst ; isPropValued = λ x a → p .fst x a .snd }) (λ b → refl) λ a → refl -Predicate≡PredicateΣ : ∀ {ℓ' ℓ''} (X : Type ℓ') → Predicate {ℓ'' = ℓ''} X ≡ PredicateΣ {ℓ'' = ℓ''} X -Predicate≡PredicateΣ {ℓ'} {ℓ''} X = isoToPath (PredicateIsoΣ X) +Predicate≡PredicateΣ : ∀ (X : Type ℓ') → Predicate X ≡ PredicateΣ X +Predicate≡PredicateΣ X = isoToPath (PredicateIsoΣ X) -isSetPredicate : ∀ {ℓ' ℓ''} (X : Type ℓ') → isSet (Predicate {ℓ'' = ℓ''} X) -isSetPredicate {ℓ'} {ℓ''} X = subst (λ predicateType → isSet predicateType) (sym (Predicate≡PredicateΣ X)) (isSetPredicateΣ {ℓ'' = ℓ''} X) +isSetPredicate : ∀ (X : Type ℓ') → isSet (Predicate X) +isSetPredicate X = subst (λ predicateType → isSet predicateType) (sym (Predicate≡PredicateΣ X)) (isSetPredicateΣ X) -PredicateΣ≡ : ∀ {ℓ' ℓ''} (X : Type ℓ') → (P Q : PredicateΣ {ℓ'' = ℓ''} X) → (P .fst ≡ Q .fst) → P ≡ Q +PredicateΣ≡ : ∀ (X : Type ℓ') → (P Q : PredicateΣ X) → (P .fst ≡ Q .fst) → P ≡ Q PredicateΣ≡ X P Q ∣P∣≡∣Q∣ = Σ≡Prop (λ _ → isPropIsSet) ∣P∣≡∣Q∣ Predicate≡ : - ∀ {ℓ' ℓ''} (X : Type ℓ') - → (P Q : Predicate {ℓ'' = ℓ''} X) + ∀ (X : Type ℓ') + → (P Q : Predicate X) → (∀ x a → a ⊩ ∣ P ∣ x → a ⊩ ∣ Q ∣ x) → (∀ x a → a ⊩ ∣ Q ∣ x → a ⊩ ∣ P ∣ x) ----------------------------------- → P ≡ Q -Predicate≡ {ℓ'} {ℓ''} X P Q P→Q Q→P i = +Predicate≡ X P Q P→Q Q→P i = PredicateIsoΣ X .inv (PredicateΣ≡ - {ℓ'' = ℓ''} X (PredicateIsoΣ X .fun P) (PredicateIsoΣ X .fun Q) diff --git a/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda b/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda index 4a01d04..bcb28e7 100644 --- a/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda +++ b/src/Realizability/Tripos/Prealgebra/Predicate/Properties.agda @@ -20,7 +20,7 @@ module Realizability.Tripos.Prealgebra.Predicate.Properties {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate.Base ca +open import Realizability.Tripos.Prealgebra.Predicate.Base {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -29,9 +29,9 @@ private λ*ComputationRule = `λ*ComputationRule as fefermanStructure private λ* = `λ* as fefermanStructure module PredicateProperties (X : Type ℓ') where - private PredicateX = Predicate {ℓ'' = ℓ''} X + private PredicateX = Predicate X open Predicate - _≤_ : Predicate {ℓ'' = ℓ''} X → Predicate {ℓ'' = ℓ''} X → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') + _≤_ : Predicate X → Predicate X → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') ϕ ≤ ψ = ∃[ b ∈ A ] (∀ (x : X) (a : A) → a ⊩ (∣ ϕ ∣ x) → (b ⨾ a) ⊩ ∣ ψ ∣ x) isProp≤ : ∀ ϕ ψ → isProp (ϕ ≤ ψ) @@ -90,8 +90,8 @@ module PredicateProperties (X : Type ℓ') where module _ where open PredicateProperties Unit* private - Predicate' = Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} - module NotAntiSym (antiSym : ∀ (a b : Predicate Unit*) → (a≤b : a ≤ b) → (b≤a : b ≤ a) → a ≡ b) where + Predicate' = Predicate + module NotAntiSym (antiSym : ∀ (a b : Predicate' Unit*) → (a≤b : a ≤ b) → (b≤a : b ≤ a) → a ≡ b) where Lift' = Lift {i = ℓ} {j = (ℓ-max ℓ' ℓ'')} kRealized : Predicate' Unit* @@ -129,8 +129,8 @@ module _ where k≡k' = Liftk≡k' .lower module Morphism {X Y : Type ℓ'} (isSetX : isSet X) (isSetY : isSet Y) where - PredicateX = Predicate {ℓ'' = ℓ''} X - PredicateY = Predicate {ℓ'' = ℓ''} Y + PredicateX = Predicate X + PredicateY = Predicate Y module PredicatePropertiesX = PredicateProperties X module PredicatePropertiesY = PredicateProperties Y open PredicatePropertiesX renaming (_≤_ to _≤X_ ; isProp≤ to isProp≤X) @@ -323,7 +323,7 @@ module BeckChevalley `∃BeckChevalley = funExt λ ϕ i → PredicateIsoΣ K .inv - (PredicateΣ≡ {ℓ'' = ℓ''} K + (PredicateΣ≡ K ((λ k a → (∣ (g* ∘ `∃[J→I][ f ]) ϕ ∣ k a) , ((g* ∘ `∃[J→I][ f ]) ϕ .isPropValued k a)) , isSetK) ((λ k a → (∣ (`∃[L→K][ p ] ∘ q*) ϕ ∣ k a) , ((`∃[L→K][ p ] ∘ q*) ϕ .isPropValued k a)) , isSetK) (funExt₂ @@ -347,7 +347,7 @@ module BeckChevalley `∀BeckChevalley = funExt λ ϕ i → PredicateIsoΣ K .inv - (PredicateΣ≡ {ℓ'' = ℓ''} K + (PredicateΣ≡ K ((λ k a → (a ⊩ ∣ g* (`∀[J→I][ f ] ϕ) ∣ k) , (g* (`∀[J→I][ f ] ϕ) .isPropValued k a)) , isSetK) ((λ k a → (a ⊩ ∣ `∀[L→K][ p ] (q* ϕ) ∣ k) , (`∀[L→K][ p ] (q* ϕ) .isPropValued k a)) , isSetK) (funExt₂ From 8029847a456e1525fe1e6fe143861b037a8da81c Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 28 Feb 2024 04:26:43 +0530 Subject: [PATCH 25/61] Archive state --- .../Topos/FunctionalRelation.agda | 354 ++---------------- src/Realizability/Topos/Object.agda | 55 ++- 2 files changed, 71 insertions(+), 338 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 14f9883..73536c3 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -13,7 +13,7 @@ open import Cubical.Data.Empty open import Cubical.Data.Unit open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients renaming (rec to setQuotRec; rec2 to setQuotRec2; elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2; elimProp3 to setQuotElimProp3) +open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category module Realizability.Topos.FunctionalRelation @@ -26,7 +26,7 @@ module Realizability.Topos.FunctionalRelation open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Tripos.Prealgebra.Meets.Identity ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial open CombinatoryAlgebra ca @@ -46,12 +46,18 @@ record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X field relation : Predicate (X × Y) - isStrict : - ∃[ st ∈ A ] + isStrictDomain : + ∃[ stD ∈ A ] (∀ x y r → r ⊩ ∣ relation ∣ (x , y) - ------------------------------------------------------------------------------------ - → (pr₁ ⨾ (st ⨾ r)) ⊩ ∣ equalityX ∣ (x , x) × ((pr₂ ⨾ (st ⨾ r)) ⊩ ∣ equalityY ∣ (y , y))) + ---------------------------------- + → (stD ⨾ r) ⊩ ∣ equalityX ∣ (x , x)) + isStrictCodomain : + ∃[ stC ∈ A ] + (∀ x y r + → r ⊩ ∣ relation ∣ (x , y) + ---------------------------------- + → (stC ⨾ r) ⊩ ∣ equalityY ∣ (y , y)) isRelational : ∃[ rl ∈ A ] (∀ x x' y y' a b c @@ -59,14 +65,14 @@ record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X → b ⊩ ∣ relation ∣ (x , y) → c ⊩ ∣ equalityY ∣ (y , y') ------------------------------------------ - → (rl ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ⊩ ∣ relation ∣ (x' , y')) + → (rl ⨾ a ⨾ b ⨾ c) ⊩ ∣ relation ∣ (x' , y')) isSingleValued : ∃[ sv ∈ A ] (∀ x y y' r₁ r₂ → r₁ ⊩ ∣ relation ∣ (x , y) → r₂ ⊩ ∣ relation ∣ (x , y') ----------------------------------- - → (sv ⨾ (pair ⨾ r₁ ⨾ r₂)) ⊩ ∣ equalityY ∣ (y , y')) + → (sv ⨾ r₁ ⨾ r₂) ⊩ ∣ equalityY ∣ (y , y')) isTotal : ∃[ tl ∈ A ] (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) @@ -88,316 +94,19 @@ F≤G→G≤F : F≤G→G≤F {X} {Y} perX perY F G F≤G = do (r , r⊩F≤G) ← F≤G - (stG , stG⊩isStrictG) ← G .isStrict (tlF , tlF⊩isTotalF) ← F .isTotal (svG , svG⊩isSingleValuedG) ← G .isSingleValued (rlF , rlF⊩isRelational) ← F .isRelational let prover : ApplStrTerm as 1 - prover = ` rlF ̇ (` pair ̇ (` pr₁ ̇ (` stG ̇ # fzero)) ̇ (` pair ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero))) ̇ (` svG ̇ (` pair ̇ (` r ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero)))) ̇ # fzero)))) - return - (λ* prover , - (λ x y r' r'⊩Gxy → - subst - (λ r'' → r'' ⊩ ∣ F .relation ∣ (x , y)) - (sym (λ*ComputationRule prover (r' ∷ []))) - (transport - (propTruncIdempotent (F .relation .isPropValued (x , y) _)) - (do - (y' , Fxy') ← tlF⊩isTotalF x (pr₁ ⨾ (stG ⨾ r')) (stG⊩isStrictG x y r' r'⊩Gxy .fst) - return - (rlF⊩isRelational - x x y' y - (pr₁ ⨾ (stG ⨾ r')) - (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) - (svG ⨾ (pair ⨾ (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) ⨾ r')) - (stG⊩isStrictG x y r' r'⊩Gxy .fst) - Fxy' - (svG⊩isSingleValuedG x y' y (r ⨾ (tlF ⨾ (pr₁ ⨾ (stG ⨾ r')))) r' (r⊩F≤G x y' (tlF ⨾ (pr₁ ⨾ (stG ⨾ r'))) Fxy') r'⊩Gxy)))))) - -equalityFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX -relation (equalityFuncRel {X} perX) = perX .equality -isStrict (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` t ̇ (` pair ̇ # fzero ̇ (` s ̇ # fzero))) ̇ (` t ̇ (` pair ̇ (` s ̇ # fzero) ̇ # fzero)) - return - (λ* prover , - λ x x' r r⊩x~x' → - let - pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) - pr₁proofEq = - pr₁ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₁ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) - ≡⟨ pr₁pxy≡x _ _ ⟩ - t ⨾ (pair ⨾ r ⨾ (s ⨾ r)) - ∎ - - pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) - pr₂proofEq = - pr₂ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₂ ⨾ (pair ⨾ (t ⨾ (pair ⨾ r ⨾ (s ⨾ r))) ⨾ (t ⨾ (pair ⨾ (s ⨾ r) ⨾ r))) - ≡⟨ pr₂pxy≡y _ _ ⟩ - t ⨾ (pair ⨾ (s ⨾ r) ⨾ r) - ∎ - in - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (t⊩isTransitive x x' x r (s ⨾ r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')) , - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) (sym pr₂proofEq) (t⊩isTransitive x' x x' (s ⨾ r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x')) -isRelational (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` t ̇ (` pair ̇ (` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) - - return - (λ* prover , - (λ x x' y y' a b c a⊩x~x' b⊩x~y c⊩y~y' → - let - proofNormalForm : (λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) ≡ (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c)) - proofNormalForm = - λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) - ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ - (t ⨾ - (pair ⨾ - (t ⨾ - (pair ⨾ (s ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))) ⨾ - (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) - ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))) - ≡⟨ - cong₂ - (λ x y → (t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ x) ⨾ y)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))))) - (pr₁pxy≡x _ _) - (pr₁ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) - ≡⟨ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ⟩ - pr₁ ⨾ (pair ⨾ b ⨾ c) - ≡⟨ pr₁pxy≡x _ _ ⟩ - b - ∎) - ⟩ - t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))) - ≡⟨ - cong - (λ x → t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ x)) - (pr₂ ⨾ (pr₂ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))) - ≡⟨ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ⟩ - pr₂ ⨾ (pair ⨾ b ⨾ c) - ≡⟨ pr₂pxy≡y _ _ ⟩ - c - ∎) - ⟩ - t ⨾ (pair ⨾ (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) ⨾ c) - ∎ - in - subst - (λ r → r ⊩ ∣ perX .equality ∣ (x' , y')) - (sym proofNormalForm) - (t⊩isTransitive x' y y' (t ⨾ (pair ⨾ (s ⨾ a) ⨾ b)) c (t⊩isTransitive x' x y (s ⨾ a) b (s⊩isSymmetric x x' a a⊩x~x') b⊩x~y) c⊩y~y'))) -isSingleValued (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - let - prover : ApplStrTerm as 1 - prover = ` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero)) - return - (λ* prover , - (λ x y y' r₁ r₂ r₁⊩x~y r₂⊩x~y' → - let - proofEq : λ* prover ⨾ (pair ⨾ r₁ ⨾ r₂) ≡ t ⨾ (pair ⨾ (s ⨾ r₁) ⨾ r₂) - proofEq = {!!} - in - subst - (λ r → r ⊩ ∣ perX .equality ∣ (y , y')) - (sym proofEq) - (t⊩isTransitive y x y' (s ⨾ r₁) r₂ (s⊩isSymmetric x y r₁ r₁⊩x~y) r₂⊩x~y'))) -isTotal (equalityFuncRel {X} perX) = - do - (s , s⊩isSymmetric) ← perX .isSymmetric - (t , t⊩isTransitive) ← perX .isTransitive - return (Id , (λ x r r⊩x~x → ∣ x , (subst (λ r → r ⊩ ∣ perX .equality ∣ (x , x)) (sym (Ida≡a r)) r⊩x~x) ∣₁)) - -composeFuncRel : - ∀ {X Y Z : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (perZ : PartialEquivalenceRelation Z) - → (F : FunctionalRelation perX perY) - → (G : FunctionalRelation perY perZ) - → FunctionalRelation perX perZ - -(relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = - record - { isSetX = isSet× (perX .isSetX) (perZ .isSetX) - ; ∣_∣ = λ { (x , z) r → ∃[ y ∈ Y ] (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (x , y) × (pr₂ ⨾ r) ⊩ ∣ G .relation ∣ (y , z) } - ; isPropValued = λ x a → isPropPropTrunc } -isStrict (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (stF , stF⊩isStrictF) ← F .isStrict - (stG , stG⊩isStrictG) ← G .isStrict - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` pr₁ ̇ (` stF ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ # fzero))) - return - (λ* prover , - (λ x z r r⊩∃y → - transport (propTruncIdempotent (isProp× (perX .equality .isPropValued (x , x) (pr₁ ⨾ (λ* prover ⨾ r))) ( perZ .equality .isPropValued (z , z) (pr₂ ⨾ (λ* prover ⨾ r))))) - (do - (y , pr₁r⊩Fxy , pr₂r⊩Gxy) ← r⊩∃y - let - pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) - pr₁proofEq = - pr₁ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₁ ⨾ (pair ⨾ (pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r))) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)))) - ≡⟨ pr₁pxy≡x _ _ ⟩ - pr₁ ⨾ (stF ⨾ (pr₁ ⨾ r)) - ∎ - - pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r)) - pr₂proofEq = {!!} - return - ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym pr₁proofEq) (stF⊩isStrictF x y (pr₁ ⨾ r) pr₁r⊩Fxy .fst)) , - subst (λ r' → r' ⊩ ∣ perZ .equality ∣ (z , z)) (sym pr₂proofEq) (stG⊩isStrictG y z (pr₂ ⨾ r) pr₂r⊩Gxy .snd))))) -isRelational (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (rlF , rlF⊩isRelationalF) ← F .isRelational - (stF , stF⊩isStrictF) ← F .isStrict - (rlG , rlG⊩isRelationalG) ← G .isRelational - let - prover : ApplStrTerm as 1 - prover = - ` pair ̇ - (` rlF ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))))))) ̇ - (` rlG ̇ (` pair ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))))) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))))) - return - (λ* prover , - (λ x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' → - do - (y , pr₁b⊩Fxy , pr₂b⊩Gyz) ← b⊩∃y - let - proofEq : λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ≡ pair ⨾ (rlF ⨾ (pair ⨾ a ⨾ (pair ⨾ (pr₁ ⨾ b) ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b)))))) ⨾ (rlG ⨾ (pair ⨾ (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) ⨾ (pair ⨾ (pr₂ ⨾ b) ⨾ c))) - proofEq = - λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) - ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ - {!!} - return - (y , - (subst - (λ r → r ⊩ ∣ F .relation ∣ (x' , y)) - (sym {!!}) - (rlF⊩isRelationalF - x x' y y a - (pr₁ ⨾ b) - (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) - a⊩x~x' pr₁b⊩Fxy - (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd)) , - (subst - (λ r → r ⊩ ∣ G .relation ∣ (y , z')) - (sym {!!}) - (rlG⊩isRelationalG - y y z z' - (pr₂ ⨾ (stF ⨾ (pr₁ ⨾ b))) - (pr₂ ⨾ b) - c - (stF⊩isStrictF x y (pr₁ ⨾ b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z')))))) -isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (svF , svF⊩isSingleValuedF) ← F .isSingleValued - (svG , svG⊩isSingleValuedG) ← G .isSingleValued - (rlG , rlG⊩isRelationalG) ← G .isRelational - (stG , stG⊩isStrictG) ← G .isStrict - let - prover : ApplStrTerm as 1 - prover = - ` svG ̇ - (` pair ̇ - (` rlG ̇ - (` pair ̇ - (` svF ̇ - (` pair ̇ - (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ - (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ - (` pair ̇ - (` pr₂ ̇ (` pr₁ ̇ # fzero)) ̇ - (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))))))) ̇ - (` pr₂ ̇ (` pr₂ ̇ # fzero))) - return - (λ* prover , - (λ x z z' r₁ r₂ r₁⊩∃y r₂⊩∃y' → - transport - (propTruncIdempotent (perZ .equality .isPropValued (z , z') _)) - (do - (y , pr₁r₁⊩Fxy , pr₂r₁⊩Gyz) ← r₁⊩∃y - (y' , pr₁r₂⊩Fxy' , pr₂r₂⊩Gy'z) ← r₂⊩∃y' - return - (subst - (λ r → r ⊩ ∣ perZ .equality ∣ (z , z')) - (sym {!!}) - (svG⊩isSingleValuedG - y' z z' - (rlG ⨾ (pair ⨾ (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) ⨾ (pair ⨾ (pr₂ ⨾ r₁) ⨾ (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁)))))) (pr₂ ⨾ r₂) - (rlG⊩isRelationalG - y y' z z - (svF ⨾ (pair ⨾ (pr₁ ⨾ r₁) ⨾ (pr₁ ⨾ r₂))) - (pr₂ ⨾ r₁) - (pr₂ ⨾ (stG ⨾ (pr₂ ⨾ r₁))) - (svF⊩isSingleValuedF - x y y' - (pr₁ ⨾ r₁) (pr₁ ⨾ r₂) - pr₁r₁⊩Fxy pr₁r₂⊩Fxy') - pr₂r₁⊩Gyz - (stG⊩isStrictG y z (pr₂ ⨾ r₁) pr₂r₁⊩Gyz .snd)) pr₂r₂⊩Gy'z))))) -isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = - do - (tlF , tlF⊩isTotalF) ← F .isTotal - (tlG , tlG⊩isTotalG) ← G .isTotal - (stF , stF⊩isStrictF) ← F .isStrict - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` tlF ̇ # fzero) ̇ (` tlG ̇ (` pr₂ ̇ (` stF ̇ (` tlF ̇ # fzero)))) - return - (λ* prover , - (λ x r r⊩x~x → - do - (y , tlF⨾r⊩Fxy) ← tlF⊩isTotalF x r r⊩x~x - (z , ⊩Gyz) ← tlG⊩isTotalG y (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) (stF⊩isStrictF x y (tlF ⨾ r) tlF⨾r⊩Fxy .snd) - let - pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ tlF ⨾ r - pr₁proofEq = - pr₁ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₁ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) - ≡⟨ pr₁pxy≡x _ _ ⟩ - tlF ⨾ r - ∎ - - pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ r) ≡ tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) - pr₂proofEq = - pr₂ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₂ ⨾ (pair ⨾ (tlF ⨾ r) ⨾ (tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))))) - ≡⟨ pr₂pxy≡y _ _ ⟩ - tlG ⨾ (pr₂ ⨾ (stF ⨾ (tlF ⨾ r))) - ∎ - return - (z , - return - (y , - (subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym pr₁proofEq) tlF⨾r⊩Fxy , - (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , z)) (sym pr₂proofEq) ⊩Gyz)))))) + prover = {!!} + return {!!} RTMorphism : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → Type _ RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G → pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F idRTMorphism : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → RTMorphism perX perX -idRTMorphism {X} perX = [ equalityFuncRel perX ] +idRTMorphism {X} perX = {!!} composeRTMorphism : ∀ {X Y Z : Type ℓ'} @@ -408,13 +117,7 @@ composeRTMorphism : → (g : RTMorphism perY perZ) ---------------------------------------- → RTMorphism perX perZ -composeRTMorphism {X} {Y} {Z} perX perY perZ f g = - setQuotRec2 - squash/ - (λ F G → [ composeFuncRel perX perY perZ F G ]) - (λ { F F' G (F≤F' , F'≤F) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) ({!!} , {!!}) }) - (λ { F G G' (G≤G' , G'≤G) → eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') ({!!} , {!!}) }) - f g +composeRTMorphism {X} {Y} {Z} perX perY perZ f g = {!!} idLRTMorphism : ∀ {X Y : Type ℓ'} @@ -422,11 +125,7 @@ idLRTMorphism : → (perY : PartialEquivalenceRelation Y) → (f : RTMorphism perX perY) → composeRTMorphism perX perX perY (idRTMorphism perX) f ≡ f -idLRTMorphism {X} {Y} perX perY f = - setQuotElimProp - (λ f → squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f) - (λ f → eq/ _ _ ({!!} , {!!})) - f +idLRTMorphism {X} {Y} perX perY f = {!!} idRRTMorphism : ∀ {X Y : Type ℓ'} @@ -434,11 +133,7 @@ idRRTMorphism : → (perY : PartialEquivalenceRelation Y) → (f : RTMorphism perX perY) → composeRTMorphism perX perY perY f (idRTMorphism perY) ≡ f -idRRTMorphism {X} {Y} perX perY f = - setQuotElimProp - (λ f → squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f) - (λ f → eq/ _ _ ({!!} , {!!})) - f +idRRTMorphism {X} {Y} perX perY f = {!!} assocRTMorphism : ∀ {X Y Z W : Type ℓ'} @@ -450,14 +145,7 @@ assocRTMorphism : → (g : RTMorphism perY perZ) → (h : RTMorphism perZ perW) → composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h ≡ composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h) -assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h = - setQuotElimProp3 - (λ f g h → - squash/ - (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h) - (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h))) - (λ f g h → eq/ _ _ ({!!} , {!!})) - f g h +assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h = {!!} RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 60340a8..8dd134b 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -2,11 +2,15 @@ open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv open import Cubical.Data.Vec open import Cubical.Data.Nat -open import Cubical.Data.FinData renaming (zero to fzero) open import Cubical.Data.Sigma open import Cubical.Data.Empty +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Reflection.RecordEquiv module Realizability.Topos.Object {ℓ ℓ' ℓ''} @@ -24,11 +28,52 @@ open Predicate renaming (isSetX to isSetPredicateBase) open PredicateProperties open Morphism -record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where +record isPartialEquivalenceRelation (X : Type ℓ') (equality : Predicate (X × X)) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field isSetX : isSet X - equality : Predicate (X × X) isSymmetric : ∃[ s ∈ A ] (∀ x y r → r ⊩ ∣ equality ∣ (x , y) → (s ⨾ r) ⊩ ∣ equality ∣ (y , x)) - field isTransitive : ∃[ t ∈ A ] (∀ x y z a b → a ⊩ ∣ equality ∣ (x , y) → b ⊩ ∣ equality ∣ (y , z) → (t ⨾ (pair ⨾ a ⨾ b)) ⊩ ∣ equality ∣ (x , z)) - + +open isPartialEquivalenceRelation +isPropIsPartialEquivalenceRelation : ∀ {X : Type ℓ'} → (equality : Predicate (X × X)) → isProp (isPartialEquivalenceRelation X equality) +isPropIsPartialEquivalenceRelation {X} equality x y i = + record { isSetX = isProp→PathP (λ i → isPropIsSet) (x .isSetX) (y .isSetX) i ; isSymmetric = squash₁ (x .isSymmetric) (y .isSymmetric) i ; isTransitive = squash₁ (x .isTransitive) (y .isTransitive) i } + +record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + field + equality : Predicate (X × X) + isPerEquality : isPartialEquivalenceRelation X equality + open isPartialEquivalenceRelation isPerEquality public + +-- Directly from previous commit +unquoteDecl PartialEquivalenceRelationIsoΣ = declareRecordIsoΣ PartialEquivalenceRelationIsoΣ (quote PartialEquivalenceRelation) + +PartialEquivalenceRelationΣ : (X : Type ℓ') → Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) +PartialEquivalenceRelationΣ X = Σ[ equality ∈ Predicate (X × X) ] isPartialEquivalenceRelation X equality + +open PartialEquivalenceRelation +module _ (X : Type ℓ') where opaque + open Iso + PartialEquivalenceRelationΣ≡ : (perA perB : PartialEquivalenceRelationΣ X) → perA .fst ≡ perB .fst → perA ≡ perB + PartialEquivalenceRelationΣ≡ perA perB predicateEq = Σ≡Prop (λ x → isPropIsPartialEquivalenceRelation x) predicateEq + + PartialEquivalenceRelationΣ≃ : (perA perB : PartialEquivalenceRelationΣ X) → (perA .fst ≡ perB .fst) ≃ (perA ≡ perB) + PartialEquivalenceRelationΣ≃ perA perB = Σ≡PropEquiv λ x → isPropIsPartialEquivalenceRelation x + + PartialEquivalenceRelationIso : (perA perB : PartialEquivalenceRelation X) → Iso (Iso.fun PartialEquivalenceRelationIsoΣ perA ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB) (perA ≡ perB) + Iso.fun (PartialEquivalenceRelationIso perA perB) p i = Iso.inv PartialEquivalenceRelationIsoΣ (p i) + inv (PartialEquivalenceRelationIso perA perB) = cong (λ x → Iso.fun PartialEquivalenceRelationIsoΣ x) + rightInv (PartialEquivalenceRelationIso perA perB) b = refl + leftInv (PartialEquivalenceRelationIso perA perB) a = refl + + -- Main SIP + PartialEquivalenceRelation≃ : (perA perB : PartialEquivalenceRelation X) → (perA .equality ≡ perB .equality) ≃ (perA ≡ perB) + PartialEquivalenceRelation≃ perA perB = + perA .equality ≡ perB .equality + ≃⟨ idEquiv (perA .equality ≡ perB .equality) ⟩ + Iso.fun PartialEquivalenceRelationIsoΣ perA .fst ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB .fst + ≃⟨ PartialEquivalenceRelationΣ≃ (Iso.fun PartialEquivalenceRelationIsoΣ perA) (Iso.fun PartialEquivalenceRelationIsoΣ perB) ⟩ + Iso.fun PartialEquivalenceRelationIsoΣ perA ≡ Iso.fun PartialEquivalenceRelationIsoΣ perB + ≃⟨ isoToEquiv (PartialEquivalenceRelationIso perA perB) ⟩ + perA ≡ perB + ■ From 48f8ce2c6775d003b2dca435ce7eb4d2639be307 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Fri, 1 Mar 2024 02:05:08 +0530 Subject: [PATCH 26/61] Switch to new definitions --- .../Topos/FunctionalRelation.agda | 578 +++++++++++++++--- src/Realizability/Topos/Object.agda | 2 +- 2 files changed, 486 insertions(+), 94 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 73536c3..41309e9 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -23,8 +23,6 @@ module Realizability.Topos.FunctionalRelation (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where -open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} -open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial @@ -40,43 +38,53 @@ private λ* = `λ* as fefermanStructure open PartialEquivalenceRelation -record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - equalityX = perX .equality - equalityY = perY .equality +record isFunctionalRelation + {X Y : Type ℓ'} + (perX : PartialEquivalenceRelation X) + (perY : PartialEquivalenceRelation Y) + (relation : Predicate (X × Y)) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + equalityX = perX .equality + equalityY = perY .equality + + field + isStrictDomain : + ∃[ stD ∈ A ] + (∀ x y r + → r ⊩ ∣ relation ∣ (x , y) + ---------------------------------- + → (stD ⨾ r) ⊩ ∣ equalityX ∣ (x , x)) + isStrictCodomain : + ∃[ stC ∈ A ] + (∀ x y r + → r ⊩ ∣ relation ∣ (x , y) + ---------------------------------- + → (stC ⨾ r) ⊩ ∣ equalityY ∣ (y , y)) + isRelational : + ∃[ rl ∈ A ] + (∀ x x' y y' a b c + → a ⊩ ∣ equalityX ∣ (x , x') + → b ⊩ ∣ relation ∣ (x , y) + → c ⊩ ∣ equalityY ∣ (y , y') + ------------------------------------------ + → (rl ⨾ a ⨾ b ⨾ c) ⊩ ∣ relation ∣ (x' , y')) + isSingleValued : + ∃[ sv ∈ A ] + (∀ x y y' r₁ r₂ + → r₁ ⊩ ∣ relation ∣ (x , y) + → r₂ ⊩ ∣ relation ∣ (x , y') + ----------------------------------- + → (sv ⨾ r₁ ⨾ r₂) ⊩ ∣ equalityY ∣ (y , y')) + isTotal : + ∃[ tl ∈ A ] + (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) + +record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field relation : Predicate (X × Y) - isStrictDomain : - ∃[ stD ∈ A ] - (∀ x y r - → r ⊩ ∣ relation ∣ (x , y) - ---------------------------------- - → (stD ⨾ r) ⊩ ∣ equalityX ∣ (x , x)) - isStrictCodomain : - ∃[ stC ∈ A ] - (∀ x y r - → r ⊩ ∣ relation ∣ (x , y) - ---------------------------------- - → (stC ⨾ r) ⊩ ∣ equalityY ∣ (y , y)) - isRelational : - ∃[ rl ∈ A ] - (∀ x x' y y' a b c - → a ⊩ ∣ equalityX ∣ (x , x') - → b ⊩ ∣ relation ∣ (x , y) - → c ⊩ ∣ equalityY ∣ (y , y') - ------------------------------------------ - → (rl ⨾ a ⨾ b ⨾ c) ⊩ ∣ relation ∣ (x' , y')) - isSingleValued : - ∃[ sv ∈ A ] - (∀ x y y' r₁ r₂ - → r₁ ⊩ ∣ relation ∣ (x , y) - → r₂ ⊩ ∣ relation ∣ (x , y') - ----------------------------------- - → (sv ⨾ r₁ ⨾ r₂) ⊩ ∣ equalityY ∣ (y , y')) - isTotal : - ∃[ tl ∈ A ] - (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) - + isFuncRel : isFunctionalRelation perX perY relation + open isFunctionalRelation isFuncRel public + open FunctionalRelation pointwiseEntailment : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → (F G : FunctionalRelation perX perY) → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') @@ -84,68 +92,452 @@ pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe ∈ A ] (∀ x y r → r ⊩ -- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong -- Lemma 4.3.5 -F≤G→G≤F : - ∀ {X Y : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (F G : FunctionalRelation perX perY) - → pointwiseEntailment perX perY F G - → pointwiseEntailment perX perY G F -F≤G→G≤F {X} {Y} perX perY F G F≤G = - do - (r , r⊩F≤G) ← F≤G - (tlF , tlF⊩isTotalF) ← F .isTotal - (svG , svG⊩isSingleValuedG) ← G .isSingleValued - (rlF , rlF⊩isRelational) ← F .isRelational - let - prover : ApplStrTerm as 1 - prover = {!!} - return {!!} +opaque + F≤G→G≤F : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (F G : FunctionalRelation perX perY) + → pointwiseEntailment perX perY F G + → pointwiseEntailment perX perY G F + F≤G→G≤F {X} {Y} perX perY F G F≤G = + do + (r , r⊩F≤G) ← F≤G + (tlF , tlF⊩isTotalF) ← F .isTotal + (svG , svG⊩isSingleValuedG) ← G .isSingleValued + (rlF , rlF⊩isRelationalF) ← F .isRelational + (stGD , stGD⊩isStrictDomainG) ← G .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` rlF ̇ (` stGD ̇ # 0) ̇ (` tlF ̇ (` stGD ̇ # 0)) ̇ (` svG ̇ (` r ̇ (` tlF ̇ (` stGD ̇ # 0))) ̇ # 0) + return + (λ* prover , + (λ x y s s⊩Gxy → + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) + (sym (λ*ComputationRule prover (s ∷ []))) + (transport + (propTruncIdempotent (F .relation .isPropValued _ _)) + (do + (y' , tlF⨾stGDs⊩Fxy') ← tlF⊩isTotalF x (stGD ⨾ s) (stGD⊩isStrictDomainG x y s s⊩Gxy) + return + (rlF⊩isRelationalF + x x y' y + (stGD ⨾ s) (tlF ⨾ (stGD ⨾ s)) (svG ⨾ (r ⨾ (tlF ⨾ (stGD ⨾ s))) ⨾ s) + (stGD⊩isStrictDomainG x y s s⊩Gxy) + tlF⨾stGDs⊩Fxy' + (svG⊩isSingleValuedG x y' y (r ⨾ (tlF ⨾ (stGD ⨾ s))) s (r⊩F≤G x y' (tlF ⨾ (stGD ⨾ s)) tlF⨾stGDs⊩Fxy') s⊩Gxy)))))) RTMorphism : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → Type _ RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G → pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F +opaque + idFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX + relation (idFuncRel {X} perX) = perX .equality + isFunctionalRelation.isStrictDomain (isFuncRel (idFuncRel {X} perX)) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` t ̇ # 0 ̇ (` s ̇ # 0) + return + (λ* prover , + λ x x' r r⊩x~x' → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (λ*ComputationRule prover (r ∷ []))) + (t⊩isTransitive x x' x r (s ⨾ r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x'))) + isFunctionalRelation.isStrictCodomain (isFuncRel (idFuncRel {X} perX)) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 1 + prover = ` t ̇ (` s ̇ # 0) ̇ # 0 + return + (λ* prover , + (λ x x' r r⊩x~x' → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) + (sym (λ*ComputationRule prover (r ∷ []))) + (t⊩isTransitive x' x x' (s ⨾ r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x'))) + isFunctionalRelation.isRelational (isFuncRel (idFuncRel {X} perX)) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 3 + prover = ` t ̇ (` t ̇ (` s ̇ # 0) ̇ # 1) ̇ # 2 + return + (λ* prover , + (λ x₁ x₂ x₃ x₄ a b c a⊩x₁~x₂ b⊩x₁~x₃ c⊩x₃~x₄ → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₄)) + (sym (λ*ComputationRule prover (a ∷ b ∷ c ∷ []))) + (t⊩isTransitive x₂ x₃ x₄ (t ⨾ (s ⨾ a) ⨾ b) c (t⊩isTransitive x₂ x₁ x₃ (s ⨾ a) b (s⊩isSymmetric x₁ x₂ a a⊩x₁~x₂) b⊩x₁~x₃) c⊩x₃~x₄))) + isFunctionalRelation.isSingleValued (isFuncRel (idFuncRel {X} perX)) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + let + prover : ApplStrTerm as 2 + prover = ` t ̇ (` s ̇ # 0) ̇ # 1 + return + (λ* prover , + (λ x₁ x₂ x₃ r₁ r₂ r₁⊩x₁~x₂ r₂⊩x₁~x₃ → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₃)) + (sym (λ*ComputationRule prover (r₁ ∷ r₂ ∷ []))) + (t⊩isTransitive x₂ x₁ x₃ (s ⨾ r₁) r₂ (s⊩isSymmetric x₁ x₂ r₁ r₁⊩x₁~x₂) r₂⊩x₁~x₃))) + isFunctionalRelation.isTotal (isFuncRel (idFuncRel {X} perX)) = + do + (s , s⊩isSymmetric) ← perX .isSymmetric + (t , t⊩isTransitive) ← perX .isTransitive + return + (Id , + (λ x r r⊩x~x → ∣ x , subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (Ida≡a _)) r⊩x~x ∣₁)) + idRTMorphism : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → RTMorphism perX perX -idRTMorphism {X} perX = {!!} - -composeRTMorphism : - ∀ {X Y Z : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (perZ : PartialEquivalenceRelation Z) - → (f : RTMorphism perX perY) - → (g : RTMorphism perY perZ) - ---------------------------------------- - → RTMorphism perX perZ -composeRTMorphism {X} {Y} {Z} perX perY perZ f g = {!!} - -idLRTMorphism : - ∀ {X Y : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (f : RTMorphism perX perY) - → composeRTMorphism perX perX perY (idRTMorphism perX) f ≡ f -idLRTMorphism {X} {Y} perX perY f = {!!} - -idRRTMorphism : - ∀ {X Y : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (f : RTMorphism perX perY) - → composeRTMorphism perX perY perY f (idRTMorphism perY) ≡ f -idRRTMorphism {X} {Y} perX perY f = {!!} - -assocRTMorphism : - ∀ {X Y Z W : Type ℓ'} - → (perX : PartialEquivalenceRelation X) - → (perY : PartialEquivalenceRelation Y) - → (perZ : PartialEquivalenceRelation Z) - → (perW : PartialEquivalenceRelation W) - → (f : RTMorphism perX perY) - → (g : RTMorphism perY perZ) - → (h : RTMorphism perZ perW) - → composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h ≡ composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h) -assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h = {!!} +idRTMorphism {X} perX = [ idFuncRel perX ] + +opaque + {-# TERMINATING #-} -- bye bye, type-checking with --safe 😔💔 + composeFuncRel : + ∀ {X Y Z : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → FunctionalRelation perX perY + → FunctionalRelation perY perZ + → FunctionalRelation perX perZ + isSetPredicateBase (relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = isSet× (perX .isSetX) (perZ .isSetX) + ∣ relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G) ∣ (x , z) r = + ∃[ y ∈ Y ] (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (x , y) × (pr₂ ⨾ r) ⊩ ∣ G .relation ∣ (y , z) + isPropValued (relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) (x , z) r = isPropPropTrunc + isFunctionalRelation.isStrictDomain (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = + do + (stFD , stFD⊩isStrictDomainF) ← F .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` stFD ̇ (` pr₁ ̇ # 0) + return + (λ* prover , + (λ x z r r⊩∃y → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (λ*ComputationRule prover (r ∷ []))) + (transport + (propTruncIdempotent (perX .equality .isPropValued _ _)) + (do + (y , pr₁r⊩Fxy , pr₂r⊩Gyz) ← r⊩∃y + return (stFD⊩isStrictDomainF x y (pr₁ ⨾ r) pr₁r⊩Fxy))))) + isFunctionalRelation.isStrictCodomain (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = + do + (stGC , stGC⊩isStrictCodomainG) ← G .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = ` stGC ̇ (` pr₂ ̇ # 0) + return + (λ* prover , + λ x z r r⊩∃y → + subst + (λ r' → r' ⊩ ∣ perZ .equality ∣ (z , z)) + (sym (λ*ComputationRule prover (r ∷ []))) + (transport + (propTruncIdempotent (perZ .equality .isPropValued _ _)) + (do + (y , pr₁r⊩Fxy , pr₂r⊩Gyz) ← r⊩∃y + return (stGC⊩isStrictCodomainG y z (pr₂ ⨾ r) pr₂r⊩Gyz)))) + isFunctionalRelation.isRelational (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = + do + (rlF , rlF⊩isRelationalF) ← F .isRelational + (rlG , rlG⊩isRelationalG) ← G .isRelational + (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain + let + prover : ApplStrTerm as 3 + prover = ` pair ̇ (` rlF ̇ # 0 ̇ (` pr₁ ̇ # 1) ̇ (` stFC ̇ (` pr₁ ̇ # 1))) ̇ (` rlG ̇ (` stFC ̇ (` pr₁ ̇ # 1)) ̇ (` pr₂ ̇ # 1) ̇ # 2) + return + (λ* prover , + (λ x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' → + do + (y , pr₁b⊩Fxy , pr₂b⊩Gyz) ← b⊩∃y + let + pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ a ⨾ b ⨾ c) ≡ rlF ⨾ a ⨾ (pr₁ ⨾ b) ⨾ (stFC ⨾ (pr₁ ⨾ b)) + pr₁proofEq = cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _ + + pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ a ⨾ b ⨾ c) ≡ rlG ⨾ (stFC ⨾ (pr₁ ⨾ b)) ⨾ (pr₂ ⨾ b) ⨾ c + pr₂proofEq = cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₂pxy≡y _ _ + return + (y , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x' , y)) + (sym pr₁proofEq) + (rlF⊩isRelationalF x x' y y a (pr₁ ⨾ b) (stFC ⨾ (pr₁ ⨾ b)) a⊩x~x' pr₁b⊩Fxy (stFC⊩isStrictCodomainF x y (pr₁ ⨾ b) pr₁b⊩Fxy)) , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (y , z')) + (sym pr₂proofEq) + (rlG⊩isRelationalG y y z z' (stFC ⨾ (pr₁ ⨾ b)) (pr₂ ⨾ b) c (stFC⊩isStrictCodomainF x y (pr₁ ⨾ b) pr₁b⊩Fxy) pr₂b⊩Gyz c⊩z~z')))) + isFunctionalRelation.isSingleValued (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = + do + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + (svG , svG⊩isSingleValuedG) ← G .isSingleValued + (relG , relG⊩isRelationalG) ← G .isRelational + (stGC , stGC⊩isStrictCodomainG) ← G .isStrictCodomain + let + prover : ApplStrTerm as 2 + prover = ` svG ̇ (` pr₂ ̇ # 0) ̇ (` relG ̇ (` svF ̇ (` pr₁ ̇ # 1) ̇ (` pr₁ ̇ # 0)) ̇ (` pr₂ ̇ # 1) ̇ (` stGC ̇ (` pr₂ ̇ # 1))) + return + (λ* prover , + (λ x z z' r₁ r₂ r₁⊩∃y r₂⊩∃y → + transport + (propTruncIdempotent (perZ .equality .isPropValued _ _)) + (do + (y , pr₁r₁⊩Fxy , pr₂r₁⊩Gyz) ← r₁⊩∃y + (y' , pr₁r₂⊩Fxy' , pr₂r₂⊩Gy'z') ← r₂⊩∃y + return + (subst + (λ r' → r' ⊩ ∣ perZ .equality ∣ (z , z')) + (sym (λ*ComputationRule prover (r₁ ∷ r₂ ∷ []))) + (svG⊩isSingleValuedG + y z z' + (pr₂ ⨾ r₁) + (relG ⨾ (svF ⨾ (pr₁ ⨾ r₂) ⨾ (pr₁ ⨾ r₁)) ⨾ (pr₂ ⨾ r₂) ⨾ (stGC ⨾ (pr₂ ⨾ r₂))) + pr₂r₁⊩Gyz + (relG⊩isRelationalG + y' y z' z' + (svF ⨾ (pr₁ ⨾ r₂) ⨾ (pr₁ ⨾ r₁)) + (pr₂ ⨾ r₂) + (stGC ⨾ (pr₂ ⨾ r₂)) + (svF⊩isSingleValuedF x y' y (pr₁ ⨾ r₂) (pr₁ ⨾ r₁) pr₁r₂⊩Fxy' pr₁r₁⊩Fxy) + pr₂r₂⊩Gy'z' + (stGC⊩isStrictCodomainG y' z' (pr₂ ⨾ r₂) pr₂r₂⊩Gy'z'))))))) + isFunctionalRelation.isTotal (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = + do + (tlF , tlF⊩isTotalF) ← F .isTotal + (tlG , tlG⊩isTotalG) ← G .isTotal + (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` tlF ̇ # 0) ̇ (` tlG ̇ (` stFC ̇ (` tlF ̇ # 0))) + return + (λ* prover , + (λ x r r⊩x~x → + do + (y , ⊩Fxy) ← tlF⊩isTotalF x r r⊩x~x + (z , ⊩Gyz) ← tlG⊩isTotalG y (stFC ⨾ (tlF ⨾ r)) (stFC⊩isStrictCodomainF x y (tlF ⨾ r) ⊩Fxy) + return + (z , + return + (y , + ((subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) ⊩Fxy) , + (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , z)) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) ⊩Gyz)))))) + +opaque + unfolding composeFuncRel + composeRTMorphism : + ∀ {X Y Z : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (f : RTMorphism perX perY) + → (g : RTMorphism perY perZ) + ---------------------------------------- + → RTMorphism perX perZ + composeRTMorphism {X} {Y} {Z} perX perY perZ f g = + SQ.rec2 + squash/ + (λ F G → [ composeFuncRel perX perY perZ F G ]) + (λ { F F' G (F≤F' , F'≤F) → + eq/ _ _ + let answer = (do + (s , s⊩F≤F') ← F≤F' + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # 0)) ̇ (` pr₂ ̇ # 0) + return + (λ* prover , + (λ x z r r⊩∃y → + do + (y , pr₁r⊩Fxy , pr₂r⊩Gyz) ← r⊩∃y + return + (y , + subst + (λ r' → r' ⊩ ∣ F' .relation ∣ (x , y)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (s⊩F≤F' x y (pr₁ ⨾ r) pr₁r⊩Fxy) , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (y , z)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + pr₂r⊩Gyz)))) + in + (answer , F≤G→G≤F perX perZ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) answer) }) + (λ { F G G' (G≤G' , G'≤G) → + eq/ _ _ + let answer = (do + (s , s⊩G≤G') ← G≤G' + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` pr₁ ̇ # 0) ̇ (` s ̇ (` pr₂ ̇ # 0)) + return + (λ* prover , + (λ x z r r⊩∃y → + do + (y , pr₁r⊩Fxy , pr₂r⊩Gyz) ← r⊩∃y + + return + (y , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) pr₁r⊩Fxy , + subst (λ r' → r' ⊩ ∣ G' .relation ∣ (y , z)) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) (s⊩G≤G' y z (pr₂ ⨾ r) pr₂r⊩Gyz))))) + in + (answer , F≤G→G≤F perX perZ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') answer) }) + f g + +opaque + unfolding composeRTMorphism + unfolding idFuncRel + idLRTMorphism : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (f : RTMorphism perX perY) + → composeRTMorphism perX perX perY (idRTMorphism perX) f ≡ f + idLRTMorphism {X} {Y} perX perY f = + SQ.elimProp + (λ f → squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f) + (λ F → + let + answer : pointwiseEntailment perX perY (composeFuncRel perX perX perY (idFuncRel perX) F) F + answer = + do + (relF , relF⊩isRelationalF) ← F .isRelational + (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain + (sX , sX⊩isSymmetricX) ← perX .isSymmetric + let + prover : ApplStrTerm as 1 + prover = ` relF ̇ (` sX ̇ (` pr₁ ̇ # 0)) ̇ (` pr₂ ̇ # 0) ̇ (` stFC ̇ (` pr₂ ̇ # 0)) + return + (λ* prover , + (λ x y r r⊩∃x' → + transport + (propTruncIdempotent (F .relation .isPropValued _ _)) + (do + (x' , pr₁r⊩x~x' , pr₂r⊩Fx'y) ← r⊩∃x' + return + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) + (sym (λ*ComputationRule prover (r ∷ []))) + (relF⊩isRelationalF + x' x y y + (sX ⨾ (pr₁ ⨾ r)) (pr₂ ⨾ r) (stFC ⨾ (pr₂ ⨾ r)) + (sX⊩isSymmetricX x x' (pr₁ ⨾ r) pr₁r⊩x~x') + pr₂r⊩Fx'y + (stFC⊩isStrictCodomainF x' y (pr₂ ⨾ r) pr₂r⊩Fx'y)))))) + in + eq/ _ _ (answer , F≤G→G≤F perX perY (composeFuncRel perX perX perY (idFuncRel perX) F) F answer)) + f + +opaque + unfolding composeRTMorphism + unfolding idFuncRel + idRRTMorphism : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (f : RTMorphism perX perY) + → composeRTMorphism perX perY perY f (idRTMorphism perY) ≡ f + idRRTMorphism {X} {Y} perX perY f = + SQ.elimProp + (λ f → squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f) + (λ F → + let + answer : pointwiseEntailment perX perY (composeFuncRel perX perY perY F (idFuncRel perY)) F + answer = + do + (relF , relF⊩isRelationalF) ← F .isRelational + (stFD , stFD⊩isStrictDomainF) ← F .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # 0)) ̇ (` pr₁ ̇ # 0) ̇ (` pr₂ ̇ # 0) + return + (λ* prover , + (λ x y r r⊩∃y' → + transport + (propTruncIdempotent (F .relation .isPropValued _ _)) + (do + (y' , pr₁r⊩Fxy' , pr₂r⊩y'~y) ← r⊩∃y' + return + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) + (sym (λ*ComputationRule prover (r ∷ []))) + (relF⊩isRelationalF x x y' y (stFD ⨾ (pr₁ ⨾ r)) (pr₁ ⨾ r) (pr₂ ⨾ r) (stFD⊩isStrictDomainF x y' (pr₁ ⨾ r) pr₁r⊩Fxy') pr₁r⊩Fxy' pr₂r⊩y'~y))))) + in + eq/ _ _ (answer , F≤G→G≤F perX perY (composeFuncRel perX perY perY F (idFuncRel perY)) F answer)) + f + +opaque + unfolding composeRTMorphism + assocRTMorphism : + ∀ {X Y Z W : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (perZ : PartialEquivalenceRelation Z) + → (perW : PartialEquivalenceRelation W) + → (f : RTMorphism perX perY) + → (g : RTMorphism perY perZ) + → (h : RTMorphism perZ perW) + → composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h ≡ composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h) + assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h = + SQ.elimProp3 + (λ f g h → + squash/ + (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h) + (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h))) + (λ F G H → + let + answer = + do + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # 0)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # 0)) ̇ (` pr₂ ̇ # 0)) + return + (λ* prover , + (λ x w r r⊩∃z → + transport + (propTruncIdempotent isPropPropTrunc) + (do + (z , pr₁r⊩∃y , pr₂r⊩Hzw) ← r⊩∃z + (y , pr₁pr₁r⊩Fxy , pr₂pr₁r⊩Gyz) ← pr₁r⊩∃y + return + (return + (y , + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + pr₁pr₁r⊩Fxy , + return + (z , + ((subst + (λ r' → r' ⊩ ∣ G .relation ∣ (y , z)) + (sym + (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r ∷ [])) ∙ + cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) + pr₂pr₁r⊩Gyz) , + (subst + (λ r' → r' ⊩ ∣ H .relation ∣ (z , w)) + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r ∷ [])) ∙ + cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) + pr₂r⊩Hzw))))))))) + in + eq/ _ _ + (answer , + F≤G→G≤F + perX perW + (composeFuncRel perX perZ perW (composeFuncRel perX perY perZ F G) H) + (composeFuncRel perX perY perW F (composeFuncRel perY perZ perW G H)) + answer)) + f g h RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 8dd134b..0f0bac6 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -32,7 +32,7 @@ record isPartialEquivalenceRelation (X : Type ℓ') (equality : Predicate (X × field isSetX : isSet X isSymmetric : ∃[ s ∈ A ] (∀ x y r → r ⊩ ∣ equality ∣ (x , y) → (s ⨾ r) ⊩ ∣ equality ∣ (y , x)) - isTransitive : ∃[ t ∈ A ] (∀ x y z a b → a ⊩ ∣ equality ∣ (x , y) → b ⊩ ∣ equality ∣ (y , z) → (t ⨾ (pair ⨾ a ⨾ b)) ⊩ ∣ equality ∣ (x , z)) + isTransitive : ∃[ t ∈ A ] (∀ x y z a b → a ⊩ ∣ equality ∣ (x , y) → b ⊩ ∣ equality ∣ (y , z) → (t ⨾ a ⨾ b) ⊩ ∣ equality ∣ (x , z)) open isPartialEquivalenceRelation isPropIsPartialEquivalenceRelation : ∀ {X : Type ℓ'} → (equality : Predicate (X × X)) → isProp (isPartialEquivalenceRelation X equality) From 6e6052fdd17904831940d19f07528730a46acf35 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Fri, 1 Mar 2024 04:01:14 +0530 Subject: [PATCH 27/61] Redo terminal object --- .../Topos/FunctionalRelation.agda | 1 - src/Realizability/Topos/TerminalObject.agda | 173 ++++++++---------- 2 files changed, 79 insertions(+), 95 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 41309e9..e8a7bd9 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -1,4 +1,3 @@ -{-# OPTIONS --allow-unsolved-metas #-} open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) open import Realizability.CombinatoryAlgebra open import Cubical.Foundations.Prelude diff --git a/src/Realizability/Topos/TerminalObject.agda b/src/Realizability/Topos/TerminalObject.agda index eaeca3d..be9f6d5 100644 --- a/src/Realizability/Topos/TerminalObject.agda +++ b/src/Realizability/Topos/TerminalObject.agda @@ -5,10 +5,11 @@ open import Cubical.Foundations.HLevels open import Cubical.Data.Unit open import Cubical.Data.Empty open import Cubical.Data.Fin +open import Cubical.Data.Fin.Literals open import Cubical.Data.Vec open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients renaming (elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2) +open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category open import Cubical.Categories.Limits.Terminal @@ -19,110 +20,94 @@ module Realizability.Topos.TerminalObject (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where open CombinatoryAlgebra ca -open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate') -open import Realizability.Tripos.Prealgebra.Predicate.Properties ca +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Topos.Object {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial open Combinators ca renaming (i to Id; ia≡a to Ida≡a) open PartialEquivalenceRelation -open Predicate' renaming (isSetX to isSetPredicateBase) +open Predicate renaming (isSetX to isSetPredicateBase) private - Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''} λ*ComputationRule = `λ*ComputationRule as fefermanStructure λ* = `λ* as fefermanStructure -terminalPartialEquivalenceRelation : PartialEquivalenceRelation Unit* -isSetX terminalPartialEquivalenceRelation = isSetUnit* -isSetPredicateBase (equality terminalPartialEquivalenceRelation) = isSet× isSetUnit* isSetUnit* -∣ equality terminalPartialEquivalenceRelation ∣ (tt* , tt*) r = Unit* -isPropValued (equality terminalPartialEquivalenceRelation) (tt* , tt*) r = isPropUnit* -isSymmetric terminalPartialEquivalenceRelation = return (k , (λ { tt* tt* _ tt* → tt* })) -isTransitive terminalPartialEquivalenceRelation = return (k , (λ { tt* tt* tt* _ _ tt* tt* → tt* })) +opaque + terminalPer : PartialEquivalenceRelation Unit* + isSetPredicateBase (equality terminalPer) = isSet× isSetUnit* isSetUnit* + ∣ equality terminalPer ∣ (tt* , tt*) _ = Unit* + isPropValued (equality terminalPer) _ _ = isPropUnit* + isPartialEquivalenceRelation.isSetX (isPerEquality terminalPer) = isSetUnit* + isPartialEquivalenceRelation.isSymmetric (isPerEquality terminalPer) = + return (k , (λ { tt* tt* r tt* → tt* })) + isPartialEquivalenceRelation.isTransitive (isPerEquality terminalPer) = + return (k , (λ { tt* tt* tt* _ _ tt* tt* → tt* })) open FunctionalRelation --- I have officially taken the inlining too far --- TODO : Refactor -isTerminalTerminalPartialEquivalenceRelation : ∀ {Y : Type ℓ'} → (perY : PartialEquivalenceRelation Y) → isContr (RTMorphism perY terminalPartialEquivalenceRelation) -isTerminalTerminalPartialEquivalenceRelation {Y} perY = - inhProp→isContr - [ record - { relation = - record - { isSetX = isSet× (perY .isSetX) isSetUnit* - ; ∣_∣ = λ { (y , tt*) r → r ⊩ ∣ perY .equality ∣ (y , y) } - ; isPropValued = λ { (y , tt*) r → perY .equality .isPropValued _ _ } } - ; isStrict = - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ # fzero ̇ # fzero - in - return - ((λ* prover) , - (λ { y tt* r r⊩y~y → - subst - (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) - (sym - (pr₁ ⨾ (λ* prover ⨾ r) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩ - pr₁ ⨾ (pair ⨾ r ⨾ r) - ≡⟨ pr₁pxy≡x _ _ ⟩ - r - ∎)) - r⊩y~y , - tt* })) - ; isRelational = - do - (trY , trY⊩isTransitiveY) ← perY .isTransitive - (smY , smY⊩isSymmetricY) ← perY .isSymmetric - let - prover : ApplStrTerm as 1 - prover = ` trY ̇ (` pair ̇ (` smY ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero)) - return - (λ* prover , - (λ { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt* → - let - proofEq : λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ≡ trY ⨾ (pair ⨾ (smY ⨾ a) ⨾ a) - proofEq = - λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) - ≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩ - (trY ⨾ (pair ⨾ (smY ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))) ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))))) - ≡⟨ cong₂ (λ x y → trY ⨾ (pair ⨾ (smY ⨾ x) ⨾ y)) (pr₁pxy≡x _ _) (pr₁pxy≡x _ _) ⟩ - trY ⨾ (pair ⨾ (smY ⨾ a) ⨾ a) - ∎ - in - subst - (λ r → r ⊩ ∣ perY .equality ∣ (y' , y')) - (sym proofEq) - (trY⊩isTransitiveY y' y y' (smY ⨾ a) a (smY⊩isSymmetricY y y' a a⊩y~y') a⊩y~y') })) - ; isSingleValued = return (k , (λ { _ tt* tt* _ _ _ _ → tt* })) -- nice - ; isTotal = return (Id , (λ y r r⊩y~y → return (tt* , subst (λ r → r ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y))) - } ] - λ f g → - setQuotElimProp2 - (λ f g → squash/ f g) - (λ F G → - eq/ - F G - let - F≤G : pointwiseEntailment perY terminalPartialEquivalenceRelation F G - F≤G = - (do - (tlG , tlG⊩isTotalG) ← G .isTotal - (stF , stF⊩isStrictF) ← F .isStrict - let - prover : ApplStrTerm as 1 - prover = ` tlG ̇ (` pr₁ ̇ (` stF ̇ # fzero)) - return - (λ* prover , - (λ { y tt* r r⊩Fx → - transport - (propTruncIdempotent (G .relation .isPropValued _ _)) - (tlG⊩isTotalG y (pr₁ ⨾ (stF ⨾ r)) (stF⊩isStrictF y tt* r r⊩Fx .fst) - >>= λ { (tt* , ⊩Gy) → return (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , tt*)) (sym (λ*ComputationRule prover (r ∷ []))) ⊩Gy) }) }))) - in F≤G , (F≤G→G≤F perY terminalPartialEquivalenceRelation F G F≤G)) - f g + +opaque + unfolding terminalPer + -- TODO : Refactor into (ugly 😠) records + -- Maybe something to do with η equality for records? + {-# TERMINATING #-} + terminalFuncRel : ∀ {Y : Type ℓ'} → (perY : PartialEquivalenceRelation Y) → FunctionalRelation perY terminalPer + Predicate.isSetX (relation (terminalFuncRel {Y} perY)) = + isSet× (perY .isSetX) isSetUnit* + Predicate.∣ relation (terminalFuncRel {Y} perY) ∣ (y , tt*) r = r ⊩ ∣ perY .equality ∣ (y , y) + Predicate.isPropValued (relation (terminalFuncRel {Y} perY)) (y , tt*) r = perY .equality .isPropValued _ _ + isFunctionalRelation.isStrictDomain (isFuncRel (terminalFuncRel {Y} perY)) = + return (Id , (λ { y tt* r r⊩y~y → subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y })) + isFunctionalRelation.isStrictCodomain (isFuncRel (terminalFuncRel {Y} perY)) = + return (k , (λ { y tt* r r⊩y~y → tt* })) + isFunctionalRelation.isRelational (isFuncRel (terminalFuncRel {Y} perY)) = + do + (t , t⊩isTransitive) ← perY .isTransitive + (s , s⊩isSymmetric) ← perY .isSymmetric + let + prover : ApplStrTerm as 3 + prover = ` t ̇ (` s ̇ # fzero) ̇ # fzero + return + (λ* prover , + (λ { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt* → + subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y')) (sym (λ*ComputationRule prover (a ∷ b ∷ c ∷ []))) (t⊩isTransitive y' y y' (s ⨾ a) a (s⊩isSymmetric y y' a a⊩y~y') a⊩y~y') })) + isFunctionalRelation.isSingleValued (isFuncRel (terminalFuncRel {Y} perY)) = + return (k , (λ { y tt* tt* r₁ r₂ r₁⊩y~y r₂⊩y~y → tt* })) + isFunctionalRelation.isTotal (isFuncRel (terminalFuncRel {Y} perY)) = + return + (Id , + (λ y r r⊩y~y → + return (tt* , (subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y)))) + +opaque + unfolding terminalPer + isTerminalTerminalPer : ∀ {Y : Type ℓ'} → (perY : PartialEquivalenceRelation Y) → isContr (RTMorphism perY terminalPer) + isTerminalTerminalPer {Y} perY = + inhProp→isContr + [ terminalFuncRel perY ] + λ f g → + SQ.elimProp2 + (λ f g → squash/ f g) + (λ F G → + let + answer : pointwiseEntailment perY terminalPer F G + answer = + do + (tlG , tlG⊩isTotalG) ← G .isTotal + (stFD , stFD⊩isStrictDomainF) ← F .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` tlG ̇ (` stFD ̇ # fzero) + return + (λ* prover , + (λ { y tt* r r⊩Fy → + transport + (propTruncIdempotent (G .relation .isPropValued _ _)) + (do + (tt* , tlGstFD⊩Gy) ← tlG⊩isTotalG y (stFD ⨾ r) (stFD⊩isStrictDomainF y tt* r r⊩Fy) + return (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , tt*)) (sym (λ*ComputationRule prover (r ∷ []))) tlGstFD⊩Gy)) })) + in + eq/ _ _ (answer , F≤G→G≤F perY terminalPer F G answer)) + f g TerminalRT : Terminal RT TerminalRT = - (Unit* , terminalPartialEquivalenceRelation) , (λ { (Y , perY) → isTerminalTerminalPartialEquivalenceRelation perY}) + (Unit* , terminalPer) , (λ { (Y , perY) → isTerminalTerminalPer perY}) From 5b2a98f1d51f1d4193a7db1180d9884602a8947c Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sun, 3 Mar 2024 17:53:24 +0530 Subject: [PATCH 28/61] Refactor binary products --- src/Realizability/Topos/#BinProducts.agda# | 200 ++++++++++++++++++++ src/Realizability/Topos/BinProducts.agda | 200 ++++++++++++++++++++ src/Realizability/Topos/TerminalObject.agda | 61 +++--- 3 files changed, 431 insertions(+), 30 deletions(-) create mode 100644 src/Realizability/Topos/#BinProducts.agda# create mode 100644 src/Realizability/Topos/BinProducts.agda diff --git a/src/Realizability/Topos/#BinProducts.agda# b/src/Realizability/Topos/#BinProducts.agda# new file mode 100644 index 0000000..fc454bb --- /dev/null +++ b/src/Realizability/Topos/#BinProducts.agda# @@ -0,0 +1,200 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Fin +open import Cubical.Data.Vec +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.BinProduct + +module Realizability.Topos.BinProducts + {ℓ ℓ' ℓ''} {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + +open FunctionalRelation +open PartialEquivalenceRelation hiding (isSetX) +module _ + {X : Type ℓ'} + {Y : Type ℓ'} + (perX : PartialEquivalenceRelation X) + (perY : PartialEquivalenceRelation Y) where + + opaque + isSetX : isSet X + isSetX = PartialEquivalenceRelation.isSetX perX + isSetY : isSet Y + isSetY = PartialEquivalenceRelation.isSetX perY + + opaque + {-# TERMINATING #-} + binProdObRT : PartialEquivalenceRelation (X × Y) + Predicate.isSetX (PartialEquivalenceRelation.equality binProdObRT) = + isSet× (isSet× isSetX isSetY) (isSet× isSetX isSetY) + Predicate.∣ PartialEquivalenceRelation.equality binProdObRT ∣ ((x , y) , x' , y') r = + (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ perY .equality ∣ (y , y') + Predicate.isPropValued (PartialEquivalenceRelation.equality binProdObRT) ((x , y) , x' , y') r = + isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) + isPartialEquivalenceRelation.isSetX (PartialEquivalenceRelation.isPerEquality binProdObRT) = isSet× isSetX isSetY + isPartialEquivalenceRelation.isSymmetric (PartialEquivalenceRelation.isPerEquality binProdObRT) = + do + (sX , sX⊩isSymmetricX) ← perX .isSymmetric + (sY , sY⊩isSymmetricY) ← perY .isSymmetric + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` sX ̇ (` pr₁ ̇ # fzero)) ̇ (` sY ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + (λ { (x , y) (x' , y') r (pr₁r⊩x~x' , pr₂r⊩y~y') → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (sX⊩isSymmetricX x x' (pr₁ ⨾ r) pr₁r⊩x~x') , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (sY⊩isSymmetricY y y' (pr₂ ⨾ r) pr₂r⊩y~y') })) + isPartialEquivalenceRelation.isTransitive (PartialEquivalenceRelation.isPerEquality binProdObRT) = + do + (tX , tX⊩isTransitiveX) ← perX .isTransitive + (tY , tY⊩isTransitiveY) ← perY .isTransitive + let + prover : ApplStrTerm as 2 + prover = ` pair ̇ (` tX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone)) ̇ (` tY ̇ (` pr₂ ̇ # fzero) ̇ (` pr₂ ̇ # fone)) + return + (λ* prover , + (λ { (x , y) (x' , y') (x'' , y'') a b (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x'~x'' , pr₂b⊩y'~y'') → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x'')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ pr₁pxy≡x _ _)) + (tX⊩isTransitiveX x x' x'' (pr₁ ⨾ a) (pr₁ ⨾ b) pr₁a⊩x~x' pr₁b⊩x'~x'') , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y'')) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ pr₂pxy≡y _ _)) + (tY⊩isTransitiveY y y' y'' (pr₂ ⨾ a) (pr₂ ⨾ b) pr₂a⊩y~y' pr₂b⊩y'~y'') })) + + opaque + unfolding binProdObRT + unfolding idFuncRel + binProdPr₁FuncRel : FunctionalRelation binProdObRT perX + FunctionalRelation.relation binProdPr₁FuncRel = + record + { isSetX = isSet× (isSet× isSetX isSetY) isSetX + ; ∣_∣ = λ { ((x , y) , x') r → (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ perY .equality ∣ (y , y) } + ; isPropValued = (λ { ((x , y) , x') r → isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) }) } + FunctionalRelation.isFuncRel binProdPr₁FuncRel = + record + { isStrictDomain = + do + (stD , stD⊩isStrictDomainEqX) ← idFuncRel perX .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` stD ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ (# fzero)) + return + (λ* prover , + (λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (stD⊩isStrictDomainEqX x x' (pr₁ ⨾ r) pr₁r⊩x~x') , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + pr₂r⊩y~y })) + ; isStrictCodomain = + do + (stC , stC⊩isStrictCodomainEqX) ← idFuncRel perX .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = ` stC ̇ (` pr₁ ̇ # fzero) + return + (λ* prover , + λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) + (sym (λ*ComputationRule prover (r ∷ []))) + (stC⊩isStrictCodomainEqX x x' (pr₁ ⨾ r) pr₁r⊩x~x') }) + ; isRelational = + do + (stC , stC⊩isStrictCodomainEqY) ← idFuncRel perY .isStrictCodomain + (t , t⊩isTransitiveX) ← perX .isTransitive + (s , s⊩isSymmetricX) ← perX .isSymmetric + let + prover : ApplStrTerm as 3 + prover = ` pair ̇ (` t ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` t ̇ (` pr₁ ̇ # fone) ̇ # (fsuc fone))) ̇ (` stC ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + ((λ { (x , y) (x' , y') x'' x''' a b c (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x~x'' , pr₂b⊩y~y) c⊩x''~x''' → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x''')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) + (t⊩isTransitiveX + x' x x''' + (s ⨾ (pr₁ ⨾ a)) (t ⨾ (pr₁ ⨾ b) ⨾ c) + (s⊩isSymmetricX x x' (pr₁ ⨾ a) pr₁a⊩x~x') + (t⊩isTransitiveX x x'' x''' (pr₁ ⨾ b) c pr₁b⊩x~x'' c⊩x''~x''')) , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y')) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₂pxy≡y _ _)) + (stC⊩isStrictCodomainEqY y y' (pr₂ ⨾ a) pr₂a⊩y~y') }))) + ; isSingleValued = + do + (t , t⊩isTransitive) ← perX .isTransitive + (s , s⊩isSymmetric) ← perX .isSymmetric + let + prover : ApplStrTerm as 2 + prover = ` t ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fone) + return + (λ* prover , + (λ { (x , y) x' x'' r₁ r₂ (pr₁r₁⊩x~x' , pr₂r₁⊩y~y) (pr₁r₂⊩x~x'' , pr₂r₂⊩y~y) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x'')) + (sym (λ*ComputationRule prover (r₁ ∷ r₂ ∷ []))) + (t⊩isTransitive x' x x'' (s ⨾ (pr₁ ⨾ r₁)) (pr₁ ⨾ r₂) (s⊩isSymmetric x x' (pr₁ ⨾ r₁) pr₁r₁⊩x~x') pr₁r₂⊩x~x'')})) + ; isTotal = + do + return + (Id , + (λ { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) → + return + (x , + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (cong (λ x → pr₁ ⨾ x) (sym (Ida≡a _))) pr₁r⊩x~x) , + (subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (cong (λ x → pr₂ ⨾ x) (sym (Ida≡a _))) pr₂r⊩y~y))) })) + } + + opaque + binProdPr₁RT : RTMorphism binProdObRT perX + binProdPr₁RT = [ binProdPr₁FuncRel ] + + opaque + + + opaque + binProductRT : BinProduct RT (X , perX) (Y , perY) + BinProduct.binProdOb binProductRT = X × Y , binProdObRT + BinProduct.binProdPr₁ binProductRT = binProdPr₁RT + BinProduct.binProdPr₂ binProductRT = {!!} + BinProduct.univProp binProductRT = {!!} + +binProductsRT : BinProducts RT +binProductsRT (X , perX) (Y , perY) = binProductRT perX perY diff --git a/src/Realizability/Topos/BinProducts.agda b/src/Realizability/Topos/BinProducts.agda new file mode 100644 index 0000000..fc454bb --- /dev/null +++ b/src/Realizability/Topos/BinProducts.agda @@ -0,0 +1,200 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Fin +open import Cubical.Data.Vec +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.BinProduct + +module Realizability.Topos.BinProducts + {ℓ ℓ' ℓ''} {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + +open FunctionalRelation +open PartialEquivalenceRelation hiding (isSetX) +module _ + {X : Type ℓ'} + {Y : Type ℓ'} + (perX : PartialEquivalenceRelation X) + (perY : PartialEquivalenceRelation Y) where + + opaque + isSetX : isSet X + isSetX = PartialEquivalenceRelation.isSetX perX + isSetY : isSet Y + isSetY = PartialEquivalenceRelation.isSetX perY + + opaque + {-# TERMINATING #-} + binProdObRT : PartialEquivalenceRelation (X × Y) + Predicate.isSetX (PartialEquivalenceRelation.equality binProdObRT) = + isSet× (isSet× isSetX isSetY) (isSet× isSetX isSetY) + Predicate.∣ PartialEquivalenceRelation.equality binProdObRT ∣ ((x , y) , x' , y') r = + (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ perY .equality ∣ (y , y') + Predicate.isPropValued (PartialEquivalenceRelation.equality binProdObRT) ((x , y) , x' , y') r = + isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) + isPartialEquivalenceRelation.isSetX (PartialEquivalenceRelation.isPerEquality binProdObRT) = isSet× isSetX isSetY + isPartialEquivalenceRelation.isSymmetric (PartialEquivalenceRelation.isPerEquality binProdObRT) = + do + (sX , sX⊩isSymmetricX) ← perX .isSymmetric + (sY , sY⊩isSymmetricY) ← perY .isSymmetric + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` sX ̇ (` pr₁ ̇ # fzero)) ̇ (` sY ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + (λ { (x , y) (x' , y') r (pr₁r⊩x~x' , pr₂r⊩y~y') → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (sX⊩isSymmetricX x x' (pr₁ ⨾ r) pr₁r⊩x~x') , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (sY⊩isSymmetricY y y' (pr₂ ⨾ r) pr₂r⊩y~y') })) + isPartialEquivalenceRelation.isTransitive (PartialEquivalenceRelation.isPerEquality binProdObRT) = + do + (tX , tX⊩isTransitiveX) ← perX .isTransitive + (tY , tY⊩isTransitiveY) ← perY .isTransitive + let + prover : ApplStrTerm as 2 + prover = ` pair ̇ (` tX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone)) ̇ (` tY ̇ (` pr₂ ̇ # fzero) ̇ (` pr₂ ̇ # fone)) + return + (λ* prover , + (λ { (x , y) (x' , y') (x'' , y'') a b (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x'~x'' , pr₂b⊩y'~y'') → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x'')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ pr₁pxy≡x _ _)) + (tX⊩isTransitiveX x x' x'' (pr₁ ⨾ a) (pr₁ ⨾ b) pr₁a⊩x~x' pr₁b⊩x'~x'') , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y'')) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ pr₂pxy≡y _ _)) + (tY⊩isTransitiveY y y' y'' (pr₂ ⨾ a) (pr₂ ⨾ b) pr₂a⊩y~y' pr₂b⊩y'~y'') })) + + opaque + unfolding binProdObRT + unfolding idFuncRel + binProdPr₁FuncRel : FunctionalRelation binProdObRT perX + FunctionalRelation.relation binProdPr₁FuncRel = + record + { isSetX = isSet× (isSet× isSetX isSetY) isSetX + ; ∣_∣ = λ { ((x , y) , x') r → (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ perY .equality ∣ (y , y) } + ; isPropValued = (λ { ((x , y) , x') r → isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) }) } + FunctionalRelation.isFuncRel binProdPr₁FuncRel = + record + { isStrictDomain = + do + (stD , stD⊩isStrictDomainEqX) ← idFuncRel perX .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` stD ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ (# fzero)) + return + (λ* prover , + (λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (stD⊩isStrictDomainEqX x x' (pr₁ ⨾ r) pr₁r⊩x~x') , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + pr₂r⊩y~y })) + ; isStrictCodomain = + do + (stC , stC⊩isStrictCodomainEqX) ← idFuncRel perX .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = ` stC ̇ (` pr₁ ̇ # fzero) + return + (λ* prover , + λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) + (sym (λ*ComputationRule prover (r ∷ []))) + (stC⊩isStrictCodomainEqX x x' (pr₁ ⨾ r) pr₁r⊩x~x') }) + ; isRelational = + do + (stC , stC⊩isStrictCodomainEqY) ← idFuncRel perY .isStrictCodomain + (t , t⊩isTransitiveX) ← perX .isTransitive + (s , s⊩isSymmetricX) ← perX .isSymmetric + let + prover : ApplStrTerm as 3 + prover = ` pair ̇ (` t ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` t ̇ (` pr₁ ̇ # fone) ̇ # (fsuc fone))) ̇ (` stC ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + ((λ { (x , y) (x' , y') x'' x''' a b c (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x~x'' , pr₂b⊩y~y) c⊩x''~x''' → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x''')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) + (t⊩isTransitiveX + x' x x''' + (s ⨾ (pr₁ ⨾ a)) (t ⨾ (pr₁ ⨾ b) ⨾ c) + (s⊩isSymmetricX x x' (pr₁ ⨾ a) pr₁a⊩x~x') + (t⊩isTransitiveX x x'' x''' (pr₁ ⨾ b) c pr₁b⊩x~x'' c⊩x''~x''')) , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y')) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₂pxy≡y _ _)) + (stC⊩isStrictCodomainEqY y y' (pr₂ ⨾ a) pr₂a⊩y~y') }))) + ; isSingleValued = + do + (t , t⊩isTransitive) ← perX .isTransitive + (s , s⊩isSymmetric) ← perX .isSymmetric + let + prover : ApplStrTerm as 2 + prover = ` t ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fone) + return + (λ* prover , + (λ { (x , y) x' x'' r₁ r₂ (pr₁r₁⊩x~x' , pr₂r₁⊩y~y) (pr₁r₂⊩x~x'' , pr₂r₂⊩y~y) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x'')) + (sym (λ*ComputationRule prover (r₁ ∷ r₂ ∷ []))) + (t⊩isTransitive x' x x'' (s ⨾ (pr₁ ⨾ r₁)) (pr₁ ⨾ r₂) (s⊩isSymmetric x x' (pr₁ ⨾ r₁) pr₁r₁⊩x~x') pr₁r₂⊩x~x'')})) + ; isTotal = + do + return + (Id , + (λ { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) → + return + (x , + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (cong (λ x → pr₁ ⨾ x) (sym (Ida≡a _))) pr₁r⊩x~x) , + (subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (cong (λ x → pr₂ ⨾ x) (sym (Ida≡a _))) pr₂r⊩y~y))) })) + } + + opaque + binProdPr₁RT : RTMorphism binProdObRT perX + binProdPr₁RT = [ binProdPr₁FuncRel ] + + opaque + + + opaque + binProductRT : BinProduct RT (X , perX) (Y , perY) + BinProduct.binProdOb binProductRT = X × Y , binProdObRT + BinProduct.binProdPr₁ binProductRT = binProdPr₁RT + BinProduct.binProdPr₂ binProductRT = {!!} + BinProduct.univProp binProductRT = {!!} + +binProductsRT : BinProducts RT +binProductsRT (X , perX) (Y , perY) = binProductRT perX perY diff --git a/src/Realizability/Topos/TerminalObject.agda b/src/Realizability/Topos/TerminalObject.agda index be9f6d5..331224b 100644 --- a/src/Realizability/Topos/TerminalObject.agda +++ b/src/Realizability/Topos/TerminalObject.agda @@ -46,37 +46,38 @@ open FunctionalRelation opaque unfolding terminalPer - -- TODO : Refactor into (ugly 😠) records - -- Maybe something to do with η equality for records? - {-# TERMINATING #-} terminalFuncRel : ∀ {Y : Type ℓ'} → (perY : PartialEquivalenceRelation Y) → FunctionalRelation perY terminalPer - Predicate.isSetX (relation (terminalFuncRel {Y} perY)) = - isSet× (perY .isSetX) isSetUnit* - Predicate.∣ relation (terminalFuncRel {Y} perY) ∣ (y , tt*) r = r ⊩ ∣ perY .equality ∣ (y , y) - Predicate.isPropValued (relation (terminalFuncRel {Y} perY)) (y , tt*) r = perY .equality .isPropValued _ _ - isFunctionalRelation.isStrictDomain (isFuncRel (terminalFuncRel {Y} perY)) = - return (Id , (λ { y tt* r r⊩y~y → subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y })) - isFunctionalRelation.isStrictCodomain (isFuncRel (terminalFuncRel {Y} perY)) = - return (k , (λ { y tt* r r⊩y~y → tt* })) - isFunctionalRelation.isRelational (isFuncRel (terminalFuncRel {Y} perY)) = - do - (t , t⊩isTransitive) ← perY .isTransitive - (s , s⊩isSymmetric) ← perY .isSymmetric - let - prover : ApplStrTerm as 3 - prover = ` t ̇ (` s ̇ # fzero) ̇ # fzero - return - (λ* prover , - (λ { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt* → - subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y')) (sym (λ*ComputationRule prover (a ∷ b ∷ c ∷ []))) (t⊩isTransitive y' y y' (s ⨾ a) a (s⊩isSymmetric y y' a a⊩y~y') a⊩y~y') })) - isFunctionalRelation.isSingleValued (isFuncRel (terminalFuncRel {Y} perY)) = - return (k , (λ { y tt* tt* r₁ r₂ r₁⊩y~y r₂⊩y~y → tt* })) - isFunctionalRelation.isTotal (isFuncRel (terminalFuncRel {Y} perY)) = - return - (Id , - (λ y r r⊩y~y → - return (tt* , (subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y)))) - + terminalFuncRel {Y} perY = + record + { relation = + record + { isSetX = isSet× (perY .isSetX) isSetUnit* + ; ∣_∣ = λ { (y , tt*) r → r ⊩ ∣ perY .equality ∣ (y , y) } + ; isPropValued = λ { (y , tt*) r → perY .equality .isPropValued _ _ } } + ; isFuncRel = + record + { isStrictDomain = return (Id , (λ { y tt* r r⊩y~y → subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y })) + ; isStrictCodomain = return (k , (λ { y tt* r r⊩y~y → tt* })) + ; isRelational = + (do + (t , t⊩isTransitive) ← perY .isTransitive + (s , s⊩isSymmetric) ← perY .isSymmetric + let + prover : ApplStrTerm as 3 + prover = ` t ̇ (` s ̇ # fzero) ̇ # fzero + return + (λ* prover , + (λ { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt* → + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y')) + (sym (λ*ComputationRule prover (a ∷ b ∷ c ∷ []))) + (t⊩isTransitive y' y y' (s ⨾ a) a (s⊩isSymmetric y y' a a⊩y~y') a⊩y~y') }))) + ; isSingleValued = (return (k , (λ { y tt* tt* r₁ r₂ r₁⊩y~y r₂⊩y~y → tt* }))) + ; isTotal = return + (Id , + (λ y r r⊩y~y → + return (tt* , (subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y)))) + } } opaque unfolding terminalPer isTerminalTerminalPer : ∀ {Y : Type ℓ'} → (perY : PartialEquivalenceRelation Y) → isContr (RTMorphism perY terminalPer) From 18b9d949521eb0b3a39a51f5a8f8ab373cb712b4 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 6 Mar 2024 16:37:59 +0530 Subject: [PATCH 29/61] Universal property of equalisers --- src/Realizability/Topos/#BinProducts.agda# | 200 ------- src/Realizability/Topos/BinProducts.agda | 364 +++++++++++- src/Realizability/Topos/Equalizer.agda | 542 ++++++++++++++++++ .../Topos/FunctionalRelation.agda | 30 +- 4 files changed, 931 insertions(+), 205 deletions(-) delete mode 100644 src/Realizability/Topos/#BinProducts.agda# create mode 100644 src/Realizability/Topos/Equalizer.agda diff --git a/src/Realizability/Topos/#BinProducts.agda# b/src/Realizability/Topos/#BinProducts.agda# deleted file mode 100644 index fc454bb..0000000 --- a/src/Realizability/Topos/#BinProducts.agda# +++ /dev/null @@ -1,200 +0,0 @@ -open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) -open import Realizability.CombinatoryAlgebra -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Data.Unit -open import Cubical.Data.Empty -open import Cubical.Data.Fin -open import Cubical.Data.Vec -open import Cubical.Data.Sigma -open import Cubical.HITs.PropositionalTruncation -open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients as SQ -open import Cubical.Categories.Category -open import Cubical.Categories.Limits.BinProduct - -module Realizability.Topos.BinProducts - {ℓ ℓ' ℓ''} {A : Type ℓ} - (ca : CombinatoryAlgebra A) - (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where - -open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca -open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial -open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial - -open CombinatoryAlgebra ca -open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open Predicate renaming (isSetX to isSetPredicateBase) -open PredicateProperties -open Morphism - -private - λ*ComputationRule = `λ*ComputationRule as fefermanStructure - λ* = `λ* as fefermanStructure - -open FunctionalRelation -open PartialEquivalenceRelation hiding (isSetX) -module _ - {X : Type ℓ'} - {Y : Type ℓ'} - (perX : PartialEquivalenceRelation X) - (perY : PartialEquivalenceRelation Y) where - - opaque - isSetX : isSet X - isSetX = PartialEquivalenceRelation.isSetX perX - isSetY : isSet Y - isSetY = PartialEquivalenceRelation.isSetX perY - - opaque - {-# TERMINATING #-} - binProdObRT : PartialEquivalenceRelation (X × Y) - Predicate.isSetX (PartialEquivalenceRelation.equality binProdObRT) = - isSet× (isSet× isSetX isSetY) (isSet× isSetX isSetY) - Predicate.∣ PartialEquivalenceRelation.equality binProdObRT ∣ ((x , y) , x' , y') r = - (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ perY .equality ∣ (y , y') - Predicate.isPropValued (PartialEquivalenceRelation.equality binProdObRT) ((x , y) , x' , y') r = - isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) - isPartialEquivalenceRelation.isSetX (PartialEquivalenceRelation.isPerEquality binProdObRT) = isSet× isSetX isSetY - isPartialEquivalenceRelation.isSymmetric (PartialEquivalenceRelation.isPerEquality binProdObRT) = - do - (sX , sX⊩isSymmetricX) ← perX .isSymmetric - (sY , sY⊩isSymmetricY) ← perY .isSymmetric - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` sX ̇ (` pr₁ ̇ # fzero)) ̇ (` sY ̇ (` pr₂ ̇ # fzero)) - return - (λ* prover , - (λ { (x , y) (x' , y') r (pr₁r⊩x~x' , pr₂r⊩y~y') → - subst - (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x)) - (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) - (sX⊩isSymmetricX x x' (pr₁ ⨾ r) pr₁r⊩x~x') , - subst - (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y)) - (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) - (sY⊩isSymmetricY y y' (pr₂ ⨾ r) pr₂r⊩y~y') })) - isPartialEquivalenceRelation.isTransitive (PartialEquivalenceRelation.isPerEquality binProdObRT) = - do - (tX , tX⊩isTransitiveX) ← perX .isTransitive - (tY , tY⊩isTransitiveY) ← perY .isTransitive - let - prover : ApplStrTerm as 2 - prover = ` pair ̇ (` tX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone)) ̇ (` tY ̇ (` pr₂ ̇ # fzero) ̇ (` pr₂ ̇ # fone)) - return - (λ* prover , - (λ { (x , y) (x' , y') (x'' , y'') a b (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x'~x'' , pr₂b⊩y'~y'') → - subst - (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x'')) - (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ pr₁pxy≡x _ _)) - (tX⊩isTransitiveX x x' x'' (pr₁ ⨾ a) (pr₁ ⨾ b) pr₁a⊩x~x' pr₁b⊩x'~x'') , - subst - (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y'')) - (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ pr₂pxy≡y _ _)) - (tY⊩isTransitiveY y y' y'' (pr₂ ⨾ a) (pr₂ ⨾ b) pr₂a⊩y~y' pr₂b⊩y'~y'') })) - - opaque - unfolding binProdObRT - unfolding idFuncRel - binProdPr₁FuncRel : FunctionalRelation binProdObRT perX - FunctionalRelation.relation binProdPr₁FuncRel = - record - { isSetX = isSet× (isSet× isSetX isSetY) isSetX - ; ∣_∣ = λ { ((x , y) , x') r → (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ perY .equality ∣ (y , y) } - ; isPropValued = (λ { ((x , y) , x') r → isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) }) } - FunctionalRelation.isFuncRel binProdPr₁FuncRel = - record - { isStrictDomain = - do - (stD , stD⊩isStrictDomainEqX) ← idFuncRel perX .isStrictDomain - let - prover : ApplStrTerm as 1 - prover = ` pair ̇ (` stD ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ (# fzero)) - return - (λ* prover , - (λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) → - subst - (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) - (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) - (stD⊩isStrictDomainEqX x x' (pr₁ ⨾ r) pr₁r⊩x~x') , - subst - (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) - (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) - pr₂r⊩y~y })) - ; isStrictCodomain = - do - (stC , stC⊩isStrictCodomainEqX) ← idFuncRel perX .isStrictCodomain - let - prover : ApplStrTerm as 1 - prover = ` stC ̇ (` pr₁ ̇ # fzero) - return - (λ* prover , - λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) → - subst - (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) - (sym (λ*ComputationRule prover (r ∷ []))) - (stC⊩isStrictCodomainEqX x x' (pr₁ ⨾ r) pr₁r⊩x~x') }) - ; isRelational = - do - (stC , stC⊩isStrictCodomainEqY) ← idFuncRel perY .isStrictCodomain - (t , t⊩isTransitiveX) ← perX .isTransitive - (s , s⊩isSymmetricX) ← perX .isSymmetric - let - prover : ApplStrTerm as 3 - prover = ` pair ̇ (` t ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` t ̇ (` pr₁ ̇ # fone) ̇ # (fsuc fone))) ̇ (` stC ̇ (` pr₂ ̇ # fzero)) - return - (λ* prover , - ((λ { (x , y) (x' , y') x'' x''' a b c (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x~x'' , pr₂b⊩y~y) c⊩x''~x''' → - subst - (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x''')) - (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) - (t⊩isTransitiveX - x' x x''' - (s ⨾ (pr₁ ⨾ a)) (t ⨾ (pr₁ ⨾ b) ⨾ c) - (s⊩isSymmetricX x x' (pr₁ ⨾ a) pr₁a⊩x~x') - (t⊩isTransitiveX x x'' x''' (pr₁ ⨾ b) c pr₁b⊩x~x'' c⊩x''~x''')) , - subst - (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y')) - (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₂pxy≡y _ _)) - (stC⊩isStrictCodomainEqY y y' (pr₂ ⨾ a) pr₂a⊩y~y') }))) - ; isSingleValued = - do - (t , t⊩isTransitive) ← perX .isTransitive - (s , s⊩isSymmetric) ← perX .isSymmetric - let - prover : ApplStrTerm as 2 - prover = ` t ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fone) - return - (λ* prover , - (λ { (x , y) x' x'' r₁ r₂ (pr₁r₁⊩x~x' , pr₂r₁⊩y~y) (pr₁r₂⊩x~x'' , pr₂r₂⊩y~y) → - subst - (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x'')) - (sym (λ*ComputationRule prover (r₁ ∷ r₂ ∷ []))) - (t⊩isTransitive x' x x'' (s ⨾ (pr₁ ⨾ r₁)) (pr₁ ⨾ r₂) (s⊩isSymmetric x x' (pr₁ ⨾ r₁) pr₁r₁⊩x~x') pr₁r₂⊩x~x'')})) - ; isTotal = - do - return - (Id , - (λ { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) → - return - (x , - ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (cong (λ x → pr₁ ⨾ x) (sym (Ida≡a _))) pr₁r⊩x~x) , - (subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (cong (λ x → pr₂ ⨾ x) (sym (Ida≡a _))) pr₂r⊩y~y))) })) - } - - opaque - binProdPr₁RT : RTMorphism binProdObRT perX - binProdPr₁RT = [ binProdPr₁FuncRel ] - - opaque - - - opaque - binProductRT : BinProduct RT (X , perX) (Y , perY) - BinProduct.binProdOb binProductRT = X × Y , binProdObRT - BinProduct.binProdPr₁ binProductRT = binProdPr₁RT - BinProduct.binProdPr₂ binProductRT = {!!} - BinProduct.univProp binProductRT = {!!} - -binProductsRT : BinProducts RT -binProductsRT (X , perX) (Y , perY) = binProductRT perX perY diff --git a/src/Realizability/Topos/BinProducts.agda b/src/Realizability/Topos/BinProducts.agda index fc454bb..79ca40d 100644 --- a/src/Realizability/Topos/BinProducts.agda +++ b/src/Realizability/Topos/BinProducts.agda @@ -40,7 +40,7 @@ module _ (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) where - opaque + opaque private isSetX : isSet X isSetX = PartialEquivalenceRelation.isSetX perX isSetY : isSet Y @@ -186,15 +186,371 @@ module _ binProdPr₁RT : RTMorphism binProdObRT perX binProdPr₁RT = [ binProdPr₁FuncRel ] + -- Code repetition to a degree + -- TODO : Refactor into a proper abstraction opaque - + unfolding binProdObRT + unfolding idFuncRel + binProdPr₂FuncRel : FunctionalRelation binProdObRT perY + FunctionalRelation.relation binProdPr₂FuncRel = + record + { isSetX = isSet× (isSet× isSetX isSetY) isSetY + ; ∣_∣ = λ { ((x , y) , y') r → (pr₁ ⨾ r) ⊩ ∣ perY .equality ∣ (y , y') × (pr₂ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x) } + ; isPropValued = λ { ((x , y) , y') r → isProp× (perY .equality .isPropValued _ _) (perX .equality .isPropValued _ _) } } + FunctionalRelation.isFuncRel binProdPr₂FuncRel = + record + { isStrictDomain = + do + (stD , stD⊩isStrictDomainEqY) ← idFuncRel perY .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` pr₂ ̇ (# fzero)) ̇ (` stD ̇ (` pr₁ ̇ # fzero)) + return + (λ* prover , + (λ { (x , y) y' r (pr₁r⊩y~y' , pr₂r⊩x~x) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + pr₂r⊩x~x , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (stD⊩isStrictDomainEqY y y' (pr₁ ⨾ r) pr₁r⊩y~y') })) + ; isStrictCodomain = + do + (stC , stC⊩isStrictCodomainEqY) ← idFuncRel perY .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = ` stC ̇ (` pr₁ ̇ # fzero) + return + (λ* prover , + (λ { (x , y) y' r (pr₁r⊩y~y' , pr₂r⊩x~x) → + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y')) + (sym (λ*ComputationRule prover (r ∷ []))) + (stC⊩isStrictCodomainEqY y y' (pr₁ ⨾ r) pr₁r⊩y~y') })) + ; isRelational = + do + (stC , stC⊩isStrictCodomainEqX) ← idFuncRel perX .isStrictCodomain + (relY , relY⊩isRelationalEqY) ← idFuncRel perY .isRelational + let + prover : ApplStrTerm as 3 + prover = ` pair ̇ (` relY ̇ (` pr₂ ̇ # fzero) ̇ (` pr₁ ̇ # fone) ̇ # (fsuc fone)) ̇ (` stC ̇ (` pr₁ ̇ # fzero)) + return + (λ* prover , + (λ { (x , y₁) (x' , y₂) y₃ y₄ a b c (pr₁a⊩x~x' , pr₂a⊩y₁~y₂) (pr₁b⊩y₁~y₃ , pr₂b⊩x~x) c⊩y₃~y₄ → + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y₂ , y₄)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) + (relY⊩isRelationalEqY y₁ y₂ y₃ y₄ (pr₂ ⨾ a) (pr₁ ⨾ b) c pr₂a⊩y₁~y₂ pr₁b⊩y₁~y₃ c⊩y₃~y₄) , + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₂pxy≡y _ _)) + (stC⊩isStrictCodomainEqX x x' (pr₁ ⨾ a) pr₁a⊩x~x') })) + ; isSingleValued = + do + (svY , svY⊩isSingleValuedY) ← idFuncRel perY .isSingleValued + let + prover : ApplStrTerm as 2 + prover = ` svY ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone) + return + (λ* prover , + (λ { (x , y) y' y'' r₁ r₂ (pr₁r₁⊩y~y' , pr₂r₁⊩x~x) (pr₁r₂⊩y~y'' , pr₂r₂⊩) → + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y' , y'')) + (sym (λ*ComputationRule prover (r₁ ∷ r₂ ∷ []))) + (svY⊩isSingleValuedY y y' y'' (pr₁ ⨾ r₁) (pr₁ ⨾ r₂) pr₁r₁⊩y~y' pr₁r₂⊩y~y'') })) + ; isTotal = + do + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` pr₂ ̇ # fzero) ̇ (` pr₁ ̇ # fzero) + return + (λ* prover , + (λ { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) → + return + (y , + (subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + pr₂r⊩y~y , + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + pr₁r⊩x~x)) })) + } + + binProdPr₂RT : RTMorphism binProdObRT perY + binProdPr₂RT = [ binProdPr₂FuncRel ] + + module UnivProp + {Z : Type ℓ'} + (perZ : PartialEquivalenceRelation Z) + (f : RTMorphism perZ perX) + (g : RTMorphism perZ perY) where + + isSetZ = PartialEquivalenceRelation.isSetX perZ + + opaque + unfolding binProdObRT + theFuncRel : (F : FunctionalRelation perZ perX) → (G : FunctionalRelation perZ perY) → FunctionalRelation perZ binProdObRT + theFuncRel F G = + record + { relation = + record + { isSetX = isSet× isSetZ (isSet× isSetX isSetY) + ; ∣_∣ = λ { (z , x , y) r → (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (z , x) × (pr₂ ⨾ r) ⊩ ∣ G .relation ∣ (z , y) } + ; isPropValued = λ { (z , x , y) r → isProp× (F .relation .isPropValued _ _) (G .relation .isPropValued _ _) } } + ; isFuncRel = + record + { isStrictDomain = + do + (stFD , stFD⊩isStrictDomain) ← F .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` stFD ̇ (` pr₁ ̇ # fzero) + return + (λ* prover , + (λ { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) → + subst + (λ r' → r' ⊩ ∣ perZ .equality ∣ (z , z)) + (sym (λ*ComputationRule prover (r ∷ []))) + (stFD⊩isStrictDomain z x (pr₁ ⨾ r) pr₁r⊩Fzx) })) + ; isStrictCodomain = + do + (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain + (stGC , stGC⊩isStrictCodomainG) ← G .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` stFC ̇ (` pr₁ ̇ # fzero)) ̇ (` stGC ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + (λ { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (stFC⊩isStrictCodomainF z x (pr₁ ⨾ r) pr₁r⊩Fzx) , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (stGC⊩isStrictCodomainG z y (pr₂ ⨾ r) pr₂r⊩Gzy) })) + ; isRelational = + do + (relF , relF⊩isRelationalF) ← F .isRelational + (relG , relG⊩isRelationalG) ← G .isRelational + let + prover : ApplStrTerm as 3 + prover = ` pair ̇ (` relF ̇ # fzero ̇ (` pr₁ ̇ # fone) ̇ (` pr₁ ̇ # (fsuc fone))) ̇ (` relG ̇ # fzero ̇ (` pr₂ ̇ # fone) ̇ (` pr₂ ̇ # (fsuc fone))) + return + (λ* prover , + (λ { z z' (x , y) (x' , y') a b c a⊩z~z' (pr₁b⊩Fzx , pr₂b⊩Gzy) (pr₁c⊩x~x' , pr₂c⊩y~y') → + (subst + (λ r → r ⊩ ∣ F .relation ∣ (z' , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) + (relF⊩isRelationalF z z' x x' a (pr₁ ⨾ b) (pr₁ ⨾ c) a⊩z~z' pr₁b⊩Fzx pr₁c⊩x~x')) , + (subst + (λ r → r ⊩ ∣ G .relation ∣ (z' , y')) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₂pxy≡y _ _)) + (relG⊩isRelationalG z z' y y' a (pr₂ ⨾ b) (pr₂ ⨾ c) a⊩z~z' pr₂b⊩Gzy pr₂c⊩y~y')) })) + ; isSingleValued = + do + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + (svG , svG⊩isSingleValuedG) ← G .isSingleValued + let + prover : ApplStrTerm as 2 + prover = ` pair ̇ (` svF ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone)) ̇ (` svG ̇ (` pr₂ ̇ # fzero) ̇ (` pr₂ ̇ # fone)) + return + (λ* prover , + (λ { z (x , y) (x' , y') r₁ r₂ (pr₁r₁⊩Fzx , pr₂r₁⊩Gzy) (pr₁r₂⊩Fzx' , pr₂r₂⊩Gzy') → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r₁ ∷ r₂ ∷ [])) ∙ pr₁pxy≡x _ _)) + (svF⊩isSingleValuedF z x x' (pr₁ ⨾ r₁) (pr₁ ⨾ r₂) pr₁r₁⊩Fzx pr₁r₂⊩Fzx') , + subst + (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y')) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r₁ ∷ r₂ ∷ [])) ∙ pr₂pxy≡y _ _)) + (svG⊩isSingleValuedG z y y' (pr₂ ⨾ r₁) (pr₂ ⨾ r₂) pr₂r₁⊩Gzy pr₂r₂⊩Gzy') })) + ; isTotal = + do + (tlF , tlF⊩isTotalF) ← F .isTotal + (tlG , tlG⊩isTotalG) ← G .isTotal + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` tlF ̇ # fzero) ̇ (` tlG ̇ # fzero) + return + (λ* prover , + (λ { z r r⊩z~z → + do + (x , tlFr⊩Fzx) ← tlF⊩isTotalF z r r⊩z~z + (y , tlGr⊩Gzy) ← tlG⊩isTotalG z r r⊩z~z + return + ((x , y) , + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (z , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + tlFr⊩Fzx , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (z , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + tlGr⊩Gzy)) })) + }} + opaque + unfolding binProdObRT + unfolding theFuncRel + theMap : RTMorphism perZ binProdObRT + theMap = + SQ.rec2 + squash/ + (λ F G → + [ theFuncRel F G ]) + (λ { F F' G (F≤F' , F'≤F) → + let + answer = + do + (s , s⊩F≤F') ← F≤F' + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero) + return + (λ* prover , + (λ { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) → + subst + (λ r' → r' ⊩ ∣ F' .relation ∣ (z , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (s⊩F≤F' z x (pr₁ ⨾ r) pr₁r⊩Fzx) , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (z , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + pr₂r⊩Gzy })) + in + eq/ _ _ (answer , F≤G→G≤F perZ binProdObRT (theFuncRel F G) (theFuncRel F' G) answer) }) + (λ { F G G' (G≤G' , G'≤G) → + let + answer = + do + (s , s⊩G≤G') ← G≤G' + let + prover : ApplStrTerm as 1 + prover = ` pair ̇ (` pr₁ ̇ # fzero) ̇ (` s ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + (λ { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) → + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (z , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + pr₁r⊩Fzx) , + (subst + (λ r' → r' ⊩ ∣ G' .relation ∣ (z , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (s⊩G≤G' z y (pr₂ ⨾ r) pr₂r⊩Gzy)) })) + in eq/ _ _ (answer , (F≤G→G≤F perZ binProdObRT (theFuncRel F G) (theFuncRel F G') answer)) }) + f g opaque + unfolding UnivProp.theMap + unfolding UnivProp.theFuncRel + unfolding binProdPr₁RT + unfolding binProdPr₁FuncRel + unfolding composeRTMorphism + unfolding binProdPr₂FuncRel binProductRT : BinProduct RT (X , perX) (Y , perY) BinProduct.binProdOb binProductRT = X × Y , binProdObRT BinProduct.binProdPr₁ binProductRT = binProdPr₁RT - BinProduct.binProdPr₂ binProductRT = {!!} - BinProduct.univProp binProductRT = {!!} + BinProduct.binProdPr₂ binProductRT = binProdPr₂RT + BinProduct.univProp binProductRT {Z , perZ} f g = + uniqueExists + (UnivProp.theMap perZ f g) + -- There is probably a better less kluged version of this proof + -- But this is the best I could do + (SQ.elimProp3 + {P = λ f g theMap' → ∀ (foo : theMap' ≡ (UnivProp.theMap perZ f g)) → composeRTMorphism _ _ _ theMap' binProdPr₁RT ≡ f} + (λ f g h → isPropΠ λ h≡ → squash/ _ _) + (λ F G theFuncRel' [theFuncRel']≡theMap → + let + answer = + do + let + (p , q) = (SQ.effective + (λ a b → isProp× isPropPropTrunc isPropPropTrunc) + (isEquivRelBientailment perZ binProdObRT) + theFuncRel' + (UnivProp.theFuncRel perZ [ F ] [ G ] F G) + [theFuncRel']≡theMap) + (p , p⊩theFuncRel'≤theFuncRel) ← p + (q , q⊩theFuncRel≤theFuncRel') ← q + (relF , relF⊩isRelationalF) ← F .isRelational + (stD , stD⊩isStrictDomain) ← theFuncRel' .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` relF ̇ (` stD ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` p ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + λ z x r r⊩∃ → + transport + (propTruncIdempotent (F .relation .isPropValued _ _)) + (do + ((x' , y) , (pr₁r⊩theFuncRel'zx'y , (pr₁pr₂r⊩x~x' , pr₂pr₂r⊩y~y))) ← r⊩∃ + return + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (z , x)) + (sym (λ*ComputationRule prover (r ∷ []))) + (relF⊩isRelationalF + z z x' x + (stD ⨾ (pr₁ ⨾ r)) (pr₁ ⨾ (p ⨾ (pr₁ ⨾ r))) (pr₁ ⨾ (pr₂ ⨾ r)) + (stD⊩isStrictDomain z (x' , y) (pr₁ ⨾ r) pr₁r⊩theFuncRel'zx'y ) + (p⊩theFuncRel'≤theFuncRel z (x' , y) (pr₁ ⨾ r) pr₁r⊩theFuncRel'zx'y .fst) + pr₁pr₂r⊩x~x')))) + in + eq/ _ _ (answer , F≤G→G≤F perZ perX (composeFuncRel _ _ _ theFuncRel' binProdPr₁FuncRel) F answer)) + f + g + (UnivProp.theMap perZ f g) + refl , + SQ.elimProp3 + {P = λ f g theMap' → ∀ (foo : theMap' ≡ (UnivProp.theMap perZ f g)) → composeRTMorphism _ _ _ theMap' binProdPr₂RT ≡ g} + (λ f g h → isPropΠ λ h≡ → squash/ _ _) + (λ F G theFuncRel' [theFuncRel']≡theMap → + let + answer = + do + let + (p , q) = (SQ.effective + (λ a b → isProp× isPropPropTrunc isPropPropTrunc) + (isEquivRelBientailment perZ binProdObRT) + theFuncRel' + (UnivProp.theFuncRel perZ [ F ] [ G ] F G) + [theFuncRel']≡theMap) + (p , p⊩theFuncRel'≤theFuncRel) ← p + (q , q⊩theFuncRel≤theFuncRel') ← q + (relG , relG⊩isRelationalG) ← G .isRelational + (st , st⊩isStrictDomainTheFuncRel') ← theFuncRel' .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` relG ̇ (` st ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ (` p ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) + return + ({!!} , + (λ z y r r⊩∃ → + transport + (propTruncIdempotent (G .relation .isPropValued _ _)) + (do + ((x , y') , (pr₁r⊩theFuncRel'zxy' , pr₁pr₂r⊩y'~y , pr₂pr₂r⊩x~x)) ← r⊩∃ + return + (subst + (λ r' → r' ⊩ ∣ G .relation ∣ (z , y)) + (sym {!!}) + (relG⊩isRelationalG + z z y' y + (st ⨾ (pr₁ ⨾ r)) (pr₂ ⨾ (p ⨾ (pr₁ ⨾ r))) (pr₁ ⨾ (pr₂ ⨾ r)) + (st⊩isStrictDomainTheFuncRel' z (x , y') (pr₁ ⨾ r) pr₁r⊩theFuncRel'zxy') + (p⊩theFuncRel'≤theFuncRel z (x , y') (pr₁ ⨾ r) pr₁r⊩theFuncRel'zxy' .snd) + pr₁pr₂r⊩y'~y))))) + in + eq/ _ _ (answer , F≤G→G≤F perZ perY (composeFuncRel _ _ _ theFuncRel' binProdPr₂FuncRel) G answer)) + f g + (UnivProp.theMap perZ f g) + refl) + (λ ! → isProp× (squash/ _ _) (squash/ _ _)) + {!!} binProductsRT : BinProducts RT binProductsRT (X , perX) (Y , perY) = binProductRT perX perY diff --git a/src/Realizability/Topos/Equalizer.agda b/src/Realizability/Topos/Equalizer.agda new file mode 100644 index 0000000..001b20a --- /dev/null +++ b/src/Realizability/Topos/Equalizer.agda @@ -0,0 +1,542 @@ +{- + +EQUALISERS IN RT(A) +──────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────── +──────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────── + +Consider two parallel morphisms f g from (X, _=X_) to (Y , _=Y_) + +In order to construct their equaliser we need to first construct an auxillary object (X , _=E_) +and construct the equaliser as an injection from (X , _=E_) to (X , _=X_) + +However, we cannot, in general show that RT has equalisers because the object (X , _=E_) that injects into (X , _=X_) depends +on choice of representatives of f and g. + +We can, however, prove a weaker theorem. We can show that equalisers *merely* exist. + +More formally, we can show that ∃[ ob ∈ RTObject ] ∃[ eq ∈ RTMorphism ob (X , _=X_) ] (univPropEqualizer eq) + +To do this, we firstly show the universal property for the case when we have already been given the +representatives. + +Since we are eliminating a set quotient into a proposition, we can choose any representatives. + +Thus we have shown that RT merely has equalisers. + +The idea of showing the mere existence of equalisers was suggested by Jon Sterling. + +See also : Remark 2.7 of "Tripos Theory" by JHP + +────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────── + +An extra note worth adding is that the code is quite difficult to read and very ugly. This is mostly due to the fact that a lot +of the things that are "implicit" in an informal setting need to be justified here. More so than usual. + +There is additional bureacracy because we have to deal with eliminators of set quotients. This makes things a little more complicated. + +-} +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Fin +open import Cubical.Data.Vec +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.BinProduct + +module Realizability.Topos.Equalizer + {ℓ ℓ' ℓ''} {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + +open FunctionalRelation +open PartialEquivalenceRelation + +equalizerUnivProp : + ∀ {X Y : Type ℓ'} + → (perX : PartialEquivalenceRelation X) + → (perY : PartialEquivalenceRelation Y) + → (f g : RTMorphism perX perY) + → (equalizerOb : PartialEquivalenceRelation X) + → (inc : RTMorphism equalizerOb perX) + → Type _ +equalizerUnivProp {X} {Y} perX perY f g equalizerOb inc = + ∀ {Z : Type ℓ'} (perZ : PartialEquivalenceRelation Z) (inc' : RTMorphism perZ perX) + → (composeRTMorphism _ _ _ inc' f) ≡ (composeRTMorphism _ _ _ inc' g) + ----------------------------------------------------------------------------------- + → ∃![ ! ∈ RTMorphism perZ equalizerOb ] (composeRTMorphism _ _ _ ! inc ≡ inc') + +module _ + {X : Type ℓ'} + {Y : Type ℓ'} + (perX : PartialEquivalenceRelation X) + (perY : PartialEquivalenceRelation Y) + (f g : RTMorphism perX perY) where + + opaque + equalizerPer : ∀ (F G : FunctionalRelation perX perY) → PartialEquivalenceRelation X + equalizerPer F G = + record + { equality = + record + { isSetX = isSet× (perX .isSetX) (perX .isSetX) + ; ∣_∣ = λ { (x , x') r → + ((pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x')) × + (∃[ y ∈ Y ] (pr₁ ⨾ (pr₂ ⨾ r)) ⊩ ∣ F .relation ∣ (x , y) × (pr₂ ⨾ (pr₂ ⨾ r)) ⊩ ∣ G .relation ∣ (x , y)) } + ; isPropValued = λ { (x , x') r → isProp× (perX .equality .isPropValued _ _) isPropPropTrunc } } + ; isPerEquality = + record + { isSetX = perX .isSetX + ; isSymmetric = + do + (s , s⊩isSymmetricX) ← perX .isSymmetric + (relF , relF⊩isRelationalF) ← F .isRelational + (relG , relG⊩isRelationalG) ← G .isRelational + (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = + ` pair ̇ + (` s ̇ (` pr₁ ̇ # fzero)) ̇ + (` pair ̇ + (` relF ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ + (` relG ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))))) + return + (λ* prover , + (λ { x x' r (pr₁r⊩x~x' , pr₂r⊩∃) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (s⊩isSymmetricX x x' (pr₁ ⨾ r) pr₁r⊩x~x') , + do + (y , pr₁pr₂r⊩Fxy , pr₂pr₂r⊩Gxy) ← pr₂r⊩∃ + return + (y , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x' , y)) + (sym (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r ∷ [])) ∙ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) + (relF⊩isRelationalF + x x' y y + (pr₁ ⨾ r) (pr₁ ⨾ (pr₂ ⨾ r)) (stFC ⨾ (pr₁ ⨾ (pr₂ ⨾ r))) + pr₁r⊩x~x' + pr₁pr₂r⊩Fxy + (stFC⊩isStrictCodomainF x y (pr₁ ⨾ (pr₂ ⨾ r)) pr₁pr₂r⊩Fxy)) , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (x' , y)) + (sym (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r ∷ [])) ∙ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) + (relG⊩isRelationalG + x x' y y + (pr₁ ⨾ r) (pr₂ ⨾ (pr₂ ⨾ r)) (stFC ⨾ (pr₁ ⨾ (pr₂ ⨾ r))) + pr₁r⊩x~x' + pr₂pr₂r⊩Gxy + (stFC⊩isStrictCodomainF x y (pr₁ ⨾ (pr₂ ⨾ r)) pr₁pr₂r⊩Fxy))) })) + ; isTransitive = + do + (t , t⊩isTransitiveX) ← perX .isTransitive + (relF , relF⊩isRelationalF) ← F .isRelational + (relG , relG⊩isRelationalG) ← G .isRelational + let + prover : ApplStrTerm as 2 + prover = ` pair ̇ (` t ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) + return + (λ* prover , + λ { x₁ x₂ x₃ a b (pr₁a⊩x₁~x₂ , pr₂a⊩∃) (pr₁b⊩x₂~x₃ , pr₂b⊩∃) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₃)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ pr₁pxy≡x _ _)) + (t⊩isTransitiveX x₁ x₂ x₃ (pr₁ ⨾ a) (pr₁ ⨾ b) pr₁a⊩x₁~x₂ pr₁b⊩x₂~x₃) , + do + (y , (pr₁pr₂a⊩Fx₁y , pr₂pr₂a⊩Gx₁y)) ← pr₂a⊩∃ + (y' , (pr₁pr₂a⊩Fx₂y' , pr₂pr₂a⊩Gx₂y')) ← pr₂b⊩∃ + return + (y , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x₁ , y)) + (sym (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) + pr₁pr₂a⊩Fx₁y , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (x₁ , y)) + (sym (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (a ∷ b ∷ [])) ∙ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) + pr₂pr₂a⊩Gx₁y) }) } } + + opaque + unfolding idFuncRel + unfolding equalizerPer + equalizerFuncRel : ∀ (F G : FunctionalRelation perX perY) → FunctionalRelation (equalizerPer F G) perX + equalizerFuncRel F G = + record + { relation = equalizerPer F G .equality + ; isFuncRel = + record + { isStrictDomain = idFuncRel (equalizerPer F G) .isStrictDomain + ; isStrictCodomain = + do + (stC , stC⊩isStrictCodomain) ← idFuncRel perX .isStrictCodomain + let + prover : ApplStrTerm as 1 + prover = ` stC ̇ (` pr₁ ̇ # fzero) + return + (λ* prover , + (λ { x x' r (pr₁r⊩x~x' , pr₂r⊩∃) → + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) (sym (λ*ComputationRule prover (r ∷ []))) (stC⊩isStrictCodomain x x' (pr₁ ⨾ r) pr₁r⊩x~x') })) + ; isRelational = + do + (relEqX , relEqX⊩isRelationalEqX) ← idFuncRel perX .isRelational + (relF , relF⊩isRelationalF) ← F .isRelational + (relG , relG⊩isRelationalG) ← G .isRelational + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + let + prover : ApplStrTerm as 3 + prover = + ` pair ̇ + (` relEqX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone) ̇ # (fsuc fone)) ̇ + (` pair ̇ + (` relF ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fone)))) ̇ + (` relG ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fone))))) + return + (λ* prover , + (λ x₁ x₂ x₃ x₄ a b c (pr₁a⊩x₁~x₂ , pr₂a⊩) (pr₁b⊩x₁~x₃ , pr₂b⊩) c⊩x₃~x₄ → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₄)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) + (relEqX⊩isRelationalEqX x₁ x₂ x₃ x₄ (pr₁ ⨾ a) (pr₁ ⨾ b) c pr₁a⊩x₁~x₂ pr₁b⊩x₁~x₃ c⊩x₃~x₄) , + do + (y , pr₁pr₂a⊩Fx₁y , pr₂pr₂a⊩Gx₁y) ← pr₂a⊩ + (y' , pr₁pr₂b⊩Fx₁y' , pr₂pr₂b⊩Gx₁y') ← pr₂b⊩ + let + y~y' = svF⊩isSingleValuedF x₁ y y' (pr₁ ⨾ (pr₂ ⨾ a)) (pr₁ ⨾ (pr₂ ⨾ b)) pr₁pr₂a⊩Fx₁y pr₁pr₂b⊩Fx₁y' + return + (y' , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x₂ , y')) + (sym (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) + (relF⊩isRelationalF + x₁ x₂ y y' + (pr₁ ⨾ a) (pr₁ ⨾ (pr₂ ⨾ a)) (svF ⨾ (pr₁ ⨾ (pr₂ ⨾ a)) ⨾ (pr₁ ⨾ (pr₂ ⨾ b))) + pr₁a⊩x₁~x₂ pr₁pr₂a⊩Fx₁y y~y') , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (x₂ , y')) + (sym (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (a ∷ b ∷ c ∷ [])) ∙ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) + (relG⊩isRelationalG + x₁ x₂ y y' + (pr₁ ⨾ a) (pr₂ ⨾ (pr₂ ⨾ a)) (svF ⨾ (pr₁ ⨾ (pr₂ ⨾ a)) ⨾ (pr₁ ⨾ (pr₂ ⨾ b))) + pr₁a⊩x₁~x₂ pr₂pr₂a⊩Gx₁y y~y')))) + ; isSingleValued = + do + (svEqX , svEqX⊩isSingleValuedEqX) ← idFuncRel perX .isSingleValued + let + prover : ApplStrTerm as 2 + prover = ` svEqX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone) + return + (λ* prover , + (λ x₁ x₂ x₃ r₁ r₂ (pr₁r₁⊩x₁~x₂ , pr₁r₁⊩) (pr₁r₂⊩x₁~x₃ , pr₂r₂⊩) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₃)) + (sym (λ*ComputationRule prover (r₁ ∷ r₂ ∷ []))) + (svEqX⊩isSingleValuedEqX x₁ x₂ x₃ (pr₁ ⨾ r₁) (pr₁ ⨾ r₂) pr₁r₁⊩x₁~x₂ pr₁r₂⊩x₁~x₃))) + ; isTotal = idFuncRel (equalizerPer F G) .isTotal + } } + + opaque + equalizerMorphism : ∀ (F G : FunctionalRelation perX perY) → RTMorphism (equalizerPer F G) perX + equalizerMorphism F G = [ equalizerFuncRel F G ] + + module UnivProp + (F G : FunctionalRelation perX perY) + {Z : Type ℓ'} + (perZ : PartialEquivalenceRelation Z) + (h : RTMorphism perZ perX) + (h⋆f≡h⋆g : composeRTMorphism _ _ _ h [ F ] ≡ composeRTMorphism _ _ _ h [ G ]) where opaque + unfolding equalizerPer + unfolding composeRTMorphism + unfolding equalizerMorphism + unfolding equalizerFuncRel + + private + !funcRel : ∀ (H : FunctionalRelation perZ perX) (H⋆F≡H⋆G : composeRTMorphism _ _ _ [ H ] [ F ] ≡ composeRTMorphism _ _ _ [ H ] [ G ]) → FunctionalRelation perZ (equalizerPer F G) + !funcRel H H⋆F≡H⋆G = + let + (p , q) = + SQ.effective + (isPropValuedBientailment perZ perY) + (isEquivRelBientailment perZ perY) + (composeFuncRel _ _ _ H F) + (composeFuncRel _ _ _ H G) + H⋆F≡H⋆G + in + record + { relation = H .relation + ; isFuncRel = + record + { isStrictDomain = H .isStrictDomain + ; isStrictCodomain = + do + (p , p⊩H⋆F≤H⋆G) ← p + (q , q⊩H⋆G≤H⋆F) ← q + (tlF , tlF⊩isTotalF) ← F .isTotal + (stCH , stCH⊩isStrictCodomainH) ← H .isStrictCodomain + (relG , relG⊩isRelationalG) ← G .isRelational + (svH , svH⊩isSingleValuedH) ← H .isSingleValued + (stCF , stCF⊩isStrictCodomainF) ← F .isStrictCodomain + let + -- possibly the ugliest realizer out there + prover : ApplStrTerm as 1 + prover = + ` pair ̇ + (` stCH ̇ # fzero) ̇ + (` pair ̇ + (` tlF ̇ (` stCH ̇ # fzero)) ̇ + (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCH ̇ # fzero))))) ̇ # fzero) ̇ + (` pr₂ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCH ̇ # fzero))))) ̇ + (` stCF ̇ (` tlF ̇ (` stCH ̇ # fzero))))) + return + (λ* prover , + (λ z x r r⊩Hzx → + let + x~x = stCH⊩isStrictCodomainH z x r r⊩Hzx + in + subst (λ r → r ⊩ ∣ perX .equality ∣ (x , x)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ∙ pr₁pxy≡x _ _)) x~x , + (do + (y , ⊩Fxy) ← tlF⊩isTotalF x (stCH ⨾ r) x~x + let + hope = + p⊩H⋆F≤H⋆G + z y (pair ⨾ r ⨾ (tlF ⨾ (stCH ⨾ r))) + (return + (x , + subst (λ r' → r' ⊩ ∣ H .relation ∣ (z , x)) (sym (pr₁pxy≡x _ _)) r⊩Hzx , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym (pr₂pxy≡y _ _)) ⊩Fxy)) + return + (y , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) + (sym + (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r ∷ [])) ∙ + cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₁pxy≡x _ _)) + ⊩Fxy , + -- god I wish there was a better way to do this :( + transport + (propTruncIdempotent (G .relation .isPropValued _ _)) + (do + (x' , ⊩Hzx' , ⊩Gx'y) ← hope + return + (subst + (λ r' → r' ⊩ ∣ G .relation ∣ (x , y)) + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r ∷ [])) ∙ + cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₂pxy≡y _ _)) + (relG⊩isRelationalG x' x y y _ _ _ (svH⊩isSingleValuedH z x' x _ _ ⊩Hzx' r⊩Hzx) ⊩Gx'y (stCF⊩isStrictCodomainF x y _ ⊩Fxy)))))))) + ; isRelational = + do + (relH , relH⊩isRelationalH) ← H .isRelational + let + prover : ApplStrTerm as 3 + prover = ` relH ̇ # fzero ̇ # fone ̇ (` pr₁ ̇ # (fsuc fone)) + return + (λ* prover , + (λ z z' x x' a b c a⊩z~z' b⊩Hzx (pr₁c⊩x~x' , pr₂c⊩∃) → + subst + (λ r' → r' ⊩ ∣ H .relation ∣ (z' , x')) + (sym (λ*ComputationRule prover (a ∷ b ∷ c ∷ []))) + (relH⊩isRelationalH z z' x x' a b (pr₁ ⨾ c) a⊩z~z' b⊩Hzx pr₁c⊩x~x'))) + ; isSingleValued = + do + (svH , svH⊩isSingleValuedH) ← H .isSingleValued + (tlF , tlF⊩isTotalF) ← F .isTotal + (tlG , tlG⊩isTotalG) ← G .isTotal + (stCH , stCH⊩isStrictCodomainH) ← H .isStrictCodomain + (stCF , stCF⊩isStrictCodomainF) ← F .isStrictCodomain + (relG , relG⊩isRelationalG) ← G .isRelational + (p , p⊩H⋆F≤H⋆G) ← p + let + prover : ApplStrTerm as 2 + prover = + ` pair ̇ + (` svH ̇ # fzero ̇ # fone) ̇ + (` pair ̇ + (` tlF ̇ (` stCH ̇ # fzero)) ̇ + (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCH ̇ # fzero))))) ̇ # fzero) ̇ + (` pr₂ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCH ̇ # fzero))))) ̇ + (` stCF ̇(` tlF ̇ (` stCH ̇ # fzero))))) + return + (λ* prover , + (λ z x x' r₁ r₂ r₁⊩Hzx r₂⊩Hzx' → + let + x~x' = svH⊩isSingleValuedH z x x' r₁ r₂ r₁⊩Hzx r₂⊩Hzx' + x~x = stCH⊩isStrictCodomainH z x r₁ r₁⊩Hzx + in + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (r₁ ∷ r₂ ∷ [])) ∙ pr₁pxy≡x _ _)) + x~x' , + do + (y , ⊩Fxy) ← tlF⊩isTotalF x _ x~x + let + y~y = stCF⊩isStrictCodomainF x y _ ⊩Fxy + hope = + p⊩H⋆F≤H⋆G z y + (pair ⨾ r₁ ⨾ (tlF ⨾ (stCH ⨾ r₁))) + (do + return + (x , + (subst (λ r' → r' ⊩ ∣ H .relation ∣ (z , x)) (sym (pr₁pxy≡x _ _)) r₁⊩Hzx) , + (subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym (pr₂pxy≡y _ _)) ⊩Fxy))) + (x'' , ⊩Hzx'' , ⊩Gx''y) ← hope + -- Can not use the fact that x ≐ x'' + let + x''~x = svH⊩isSingleValuedH z x'' x _ _ ⊩Hzx'' r₁⊩Hzx + return + (y , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) + (sym + (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r₁ ∷ r₂ ∷ [])) ∙ + cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₁pxy≡x _ _)) ⊩Fxy , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (x , y)) + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule prover (r₁ ∷ r₂ ∷ [])) ∙ + cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₂pxy≡y _ _)) + (relG⊩isRelationalG x'' x y y _ _ _ x''~x ⊩Gx''y y~y)))) + ; isTotal = H .isTotal + } } + + mainMap : + Σ[ ! ∈ RTMorphism perZ (equalizerPer F G) ] + (composeRTMorphism _ _ _ ! (equalizerMorphism F G) ≡ h) × + (∀ (!' : RTMorphism perZ (equalizerPer F G)) (!'⋆inc≡h : composeRTMorphism _ _ _ !' (equalizerMorphism F G) ≡ h) → !' ≡ !) + mainMap = + SQ.elim + {P = + λ h → + ∀ (h⋆f≡h⋆g : composeRTMorphism _ _ _ h [ F ] ≡ composeRTMorphism _ _ _ h [ G ]) + → Σ[ ! ∈ _ ] (composeRTMorphism _ _ _ ! (equalizerMorphism F G) ≡ h) × + (∀ (!' : RTMorphism perZ (equalizerPer F G)) (!'⋆inc≡h : composeRTMorphism _ _ _ !' (equalizerMorphism F G) ≡ h) → !' ≡ !)} + (λ h → isSetΠ λ _ → isSetΣ squash/ λ ! → isSet× (isProp→isSet (squash/ _ _)) (isSetΠ λ !' → isSetΠ λ _ → isProp→isSet (squash/ !' !))) + (λ H H⋆F≡H⋆G → + [ !funcRel H H⋆F≡H⋆G ] , + let + answer = + do + (relH , relH⊩isRelationalH) ← H .isRelational + (stDH , stDH⊩isStrictDomainH) ← H .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` relH ̇ (` stDH ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) + return + (λ* prover , + (λ z x r r⊩∃x' → + transport + (propTruncIdempotent (H .relation .isPropValued _ _)) + (do + (x' , pr₁r⊩Hzx' , (pr₁pr₂r⊩x'~x , _)) ← r⊩∃x' + return + (subst + (λ r' → r' ⊩ ∣ H .relation ∣ (z , x)) + (sym (λ*ComputationRule prover (r ∷ []))) + (relH⊩isRelationalH z z x' x _ _ _ (stDH⊩isStrictDomainH z x' (pr₁ ⨾ r) pr₁r⊩Hzx') pr₁r⊩Hzx' pr₁pr₂r⊩x'~x))))) + !funcRel⋆inc≡H = eq/ _ _ (answer , F≤G→G≤F _ _ (composeFuncRel _ _ _ (!funcRel H H⋆F≡H⋆G) (equalizerFuncRel F G)) H answer) + in !funcRel⋆inc≡H , + λ !' !'⋆inc≡H → + SQ.elimProp + {P = + λ !' → + ∀ (foo : composeRTMorphism _ _ _ !' (equalizerMorphism F G) ≡ [ H ]) + → !' ≡ [ !funcRel H H⋆F≡H⋆G ]} + (λ !' → isPropΠ λ _ → squash/ _ _) + (λ !'funcRel !'funcRel⋆inc≡H → + let + (p , q) = SQ.effective (isPropValuedBientailment perZ perX) (isEquivRelBientailment perZ perX) (composeFuncRel _ _ _ !'funcRel (equalizerFuncRel F G)) H !'funcRel⋆inc≡H + (p' , q') = SQ.effective (isPropValuedBientailment perZ perY) (isEquivRelBientailment perZ perY) (composeFuncRel _ _ _ H F) (composeFuncRel _ _ _ H G) H⋆F≡H⋆G + answer = + do + (q , q⊩inc⋆!'≤H) ← q + (rel!' , rel!'⊩isRelational!'FuncRel) ← !'funcRel .isRelational + (stDH , stDH⊩isStrictDomainH) ← H .isStrictDomain + let + prover : ApplStrTerm as 1 + prover = ` rel!' ̇ (` stDH ̇ # fzero) ̇ (` pr₁ ̇ (` q ̇ # fzero)) ̇ (` pr₂ ̇ (` q ̇ # fzero)) + return + (λ* prover , + (λ z x r r⊩Hzx → + transport + (propTruncIdempotent (!'funcRel .relation .isPropValued _ _)) + (do + (x' , pr₁qr⊩!'zx' , ⊩x~x' , foo) ← q⊩inc⋆!'≤H z x r r⊩Hzx + return + (subst + (λ r' → r' ⊩ ∣ !'funcRel .relation ∣ (z , x)) + (sym (λ*ComputationRule prover (r ∷ []))) + (rel!'⊩isRelational!'FuncRel + z z x' x _ _ _ + (stDH⊩isStrictDomainH z x r r⊩Hzx) + pr₁qr⊩!'zx' + (⊩x~x' , foo)))))) + in + eq/ _ _ (F≤G→G≤F perZ (equalizerPer F G) (!funcRel H H⋆F≡H⋆G) !'funcRel answer , answer)) + !' + !'⋆inc≡H) + (λ { H H' (H≤H' , H'≤H) → + funExtDep + {A = λ i → composeRTMorphism _ _ _ (eq/ H H' (H≤H' , H'≤H) i) [ F ] ≡ composeRTMorphism _ _ _ (eq/ H H' (H≤H' , H'≤H) i) [ G ]} + λ {H⋆F≡H⋆G} {H'⋆F≡H'⋆G} p → + ΣPathPProp + {A = λ i → RTMorphism perZ (equalizerPer F G)} + {B = λ i ! → + ((composeRTMorphism _ _ _ ! (equalizerMorphism F G)) ≡ eq/ H H' (H≤H' , H'≤H) i) × + (∀ (!' : RTMorphism perZ (equalizerPer F G)) → composeRTMorphism _ _ _ !' (equalizerMorphism F G) ≡ eq/ H H' (H≤H' , H'≤H) i → !' ≡ !)} + (λ ! → isProp× (squash/ _ _) (isPropΠ λ !' → isPropΠ λ !'⋆inc≡H' → squash/ _ _)) + let + answer = + (do + (s , s⊩H≤H') ← H≤H' + return + (s , + (λ z x r r⊩Hzx → + s⊩H≤H' z x r r⊩Hzx))) + in eq/ _ _ (answer , F≤G→G≤F perZ (equalizerPer F G) (!funcRel H H⋆F≡H⋆G) (!funcRel H' H'⋆F≡H'⋆G) answer) }) + h + h⋆f≡h⋆g + + -- We have now done the major work and can simply eliminate f and g + opaque + unfolding idFuncRel + unfolding equalizerPer + equalizer : + ∃[ equalizerOb ∈ PartialEquivalenceRelation X ] + ∃[ inc ∈ RTMorphism equalizerOb perX ] + (equalizerUnivProp perX perY f g equalizerOb inc) + equalizer = + SQ.elimProp2 + {P = λ f g → ∃[ equalizerOb ∈ PartialEquivalenceRelation X ] ∃[ inc ∈ RTMorphism equalizerOb perX ] (equalizerUnivProp perX perY f g equalizerOb inc)} + {!!} + (λ F G → + {!!}) + f g + diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index e8a7bd9..50a9738 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -14,6 +14,7 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category +open import Cubical.Relation.Binary module Realizability.Topos.FunctionalRelation {ℓ ℓ' ℓ''} @@ -127,8 +128,35 @@ opaque tlF⨾stGDs⊩Fxy' (svG⊩isSingleValuedG x y' y (r ⨾ (tlF ⨾ (stGD ⨾ s))) s (r⊩F≤G x y' (tlF ⨾ (stGD ⨾ s)) tlF⨾stGDs⊩Fxy') s⊩Gxy)))))) +bientailment : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → FunctionalRelation perX perY → FunctionalRelation perX perY → Type _ +bientailment {X} {Y} perX perY F G = pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F + +isPropValuedBientailment : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → (F G : FunctionalRelation perX perY) → isProp (bientailment perX perY F G) +isPropValuedBientailment {X} {Y} perX perY F G = isProp× isPropPropTrunc isPropPropTrunc + RTMorphism : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → Type _ -RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G → pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F +RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / bientailment perX perY + +isEquivRelBientailment : ∀ {X Y : Type ℓ'} → (perX : PartialEquivalenceRelation X) → (perY : PartialEquivalenceRelation Y) → BinaryRelation.isEquivRel (bientailment perX perY) +BinaryRelation.isEquivRel.reflexive (isEquivRelBientailment {X} {Y} perX perY) = + λ A → + ∣ Id , (λ x y r r⊩Axy → subst (λ r' → r' ⊩ ∣ A .relation ∣ (x , y)) (sym (Ida≡a _)) r⊩Axy) ∣₁ , + ∣ Id , (λ x y r r⊩Axy → subst (λ r' → r' ⊩ ∣ A .relation ∣ (x , y)) (sym (Ida≡a _)) r⊩Axy) ∣₁ +BinaryRelation.isEquivRel.symmetric (isEquivRelBientailment {X} {Y} perX perY) F G (F≤G , G≤F) = G≤F , F≤G +BinaryRelation.isEquivRel.transitive (isEquivRelBientailment {X} {Y} perX perY) F G H (F≤G , G≤F) (G≤H , H≤G) = + let + answer = + do + (s , s⊩F≤G) ← F≤G + (p , p⊩G≤H) ← G≤H + let + prover : ApplStrTerm as 1 + prover = ` p ̇ (` s ̇ # fzero) + return + (λ* prover , + (λ x y r r⊩Fxy → subst (λ r' → r' ⊩ ∣ H .relation ∣ (x , y)) (sym (λ*ComputationRule prover (r ∷ []))) (p⊩G≤H x y (s ⨾ r) (s⊩F≤G x y r r⊩Fxy)))) + in + answer , F≤G→G≤F perX perY F H answer opaque idFuncRel : ∀ {X : Type ℓ'} → (perX : PartialEquivalenceRelation X) → FunctionalRelation perX perX From 6a7058d3c9ed34274c5d10548a7a654cbda78f35 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 6 Mar 2024 18:13:09 +0530 Subject: [PATCH 30/61] Show mere existence of equalisers --- src/Realizability/Topos/Equalizer.agda | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/src/Realizability/Topos/Equalizer.agda b/src/Realizability/Topos/Equalizer.agda index 001b20a..e9d80a0 100644 --- a/src/Realizability/Topos/Equalizer.agda +++ b/src/Realizability/Topos/Equalizer.agda @@ -535,8 +535,18 @@ module _ equalizer = SQ.elimProp2 {P = λ f g → ∃[ equalizerOb ∈ PartialEquivalenceRelation X ] ∃[ inc ∈ RTMorphism equalizerOb perX ] (equalizerUnivProp perX perY f g equalizerOb inc)} - {!!} + (λ f g → isPropPropTrunc) (λ F G → - {!!}) + return + ((equalizerPer F G) , + (return + ((equalizerMorphism F G) , + (λ perZ h h⋆f≡h⋆g → + uniqueExists + (UnivProp.mainMap F G perZ h h⋆f≡h⋆g .fst) + (UnivProp.mainMap F G perZ h h⋆f≡h⋆g .snd .fst) + (λ !' → squash/ _ _) + λ !' !'⋆inc≡h → + sym (UnivProp.mainMap F G perZ h h⋆f≡h⋆g .snd .snd !' !'⋆inc≡h)))))) f g From fde067155ec86af26d6b1fd1d49e596af5e7ba5e Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 12 Mar 2024 21:20:41 +0530 Subject: [PATCH 31/61] Realizer for universal property of binary products --- src/Realizability/Topos/BinProducts.agda | 60 ++++++++++++++++++++++-- 1 file changed, 57 insertions(+), 3 deletions(-) diff --git a/src/Realizability/Topos/BinProducts.agda b/src/Realizability/Topos/BinProducts.agda index 79ca40d..afff07c 100644 --- a/src/Realizability/Topos/BinProducts.agda +++ b/src/Realizability/Topos/BinProducts.agda @@ -528,7 +528,7 @@ module _ prover : ApplStrTerm as 1 prover = ` relG ̇ (` st ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ (` p ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) return - ({!!} , + (λ* prover , (λ z y r r⊩∃ → transport (propTruncIdempotent (G .relation .isPropValued _ _)) @@ -537,7 +537,7 @@ module _ return (subst (λ r' → r' ⊩ ∣ G .relation ∣ (z , y)) - (sym {!!}) + (sym (λ*ComputationRule prover (r ∷ []))) (relG⊩isRelationalG z z y' y (st ⨾ (pr₁ ⨾ r)) (pr₂ ⨾ (p ⨾ (pr₁ ⨾ r))) (pr₁ ⨾ (pr₂ ⨾ r)) @@ -550,7 +550,61 @@ module _ (UnivProp.theMap perZ f g) refl) (λ ! → isProp× (squash/ _ _) (squash/ _ _)) - {!!} + λ { !' (!'⋆π₁≡f , !'⋆π₂≡g) → + SQ.elimProp3 + {P = λ !' f g → ∀ (foo : composeRTMorphism _ _ _ !' binProdPr₁RT ≡ f) (bar : composeRTMorphism _ _ _ !' binProdPr₂RT ≡ g) → UnivProp.theMap perZ f g ≡ !'} + (λ f g !' → isPropΠ λ _ → isPropΠ λ _ → squash/ _ _) + (λ !' F G !'⋆π₁≡F !'⋆π₂≡G → + let + answer = + do + let + (p , q) = SQ.effective (isPropValuedBientailment perZ perX) (isEquivRelBientailment perZ perX) (composeFuncRel _ _ _ !' binProdPr₁FuncRel) F !'⋆π₁≡F + (p' , q') = SQ.effective (isPropValuedBientailment perZ perY) (isEquivRelBientailment perZ perY) (composeFuncRel _ _ _ !' binProdPr₂FuncRel) G !'⋆π₂≡G + (q , q⊩F≤!'⋆π₁) ← q + (q' , q'⊩G≤!'⋆π₂) ← q' + (rel!' , rel!'⊩isRelational!') ← !' .isRelational + (sv!' , sv!'⊩isSingleValued!') ← !' .isSingleValued + (tX , tX⊩isTransitiveX) ← perX .isTransitive + (sX , sX⊩isSymmetricX) ← perX .isSymmetric + (stD!' , stD!'⊩isStrictDomain!') ← !' .isStrictDomain + let + realizer : ApplStrTerm as 1 -- cursed + realizer = + ` rel!' ̇ (` stD!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # fzero)))) ̇ + (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # fzero))) ̇ + (` pair ̇ + (` tX ̇ + (` sX ̇ + (` pr₁ ̇ + (` sv!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # fzero)))))) ̇ + (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pr₁ ̇ # fzero))))) ̇ + (` pr₁ ̇ (` pr₂ ̇ (` q' ̇ (` pr₂ ̇ # fzero))))) + return + (λ* realizer , + (λ { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) → + transport + (propTruncIdempotent (!' .relation .isPropValued _ _)) + (do + ((x' , y') , ⊩!'zx'y' , ⊩x'~x , ⊩y'~y') ← q⊩F≤!'⋆π₁ z x _ pr₁r⊩Fzx + ((x'' , y'') , ⊩!'zx''y'' , ⊩y''~y , ⊩x''~x'') ← q'⊩G≤!'⋆π₂ z y _ pr₂r⊩Gzy + let + (⊩x'~x'' , ⊩y'~y'') = sv!'⊩isSingleValued!' z (x' , y') (x'' , y'') _ _ ⊩!'zx'y' ⊩!'zx''y'' + ⊩x''~x = tX⊩isTransitiveX x'' x' x _ _ (sX⊩isSymmetricX x' x'' _ ⊩x'~x'') ⊩x'~x + return + (subst + (λ r' → r' ⊩ ∣ !' .relation ∣ (z , x , y)) + (sym (λ*ComputationRule realizer (r ∷ []))) + (rel!'⊩isRelational!' + z z + (x'' , y'') + (x , y) + _ _ _ + (stD!'⊩isStrictDomain!' z (x' , y') _ ⊩!'zx'y') ⊩!'zx''y'' ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x'' , x)) (sym (pr₁pxy≡x _ _)) ⊩x''~x) , (subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y'' , y)) (sym (pr₂pxy≡y _ _)) ⊩y''~y))))) })) + in + eq/ _ _ (answer , F≤G→G≤F perZ binProdObRT (UnivProp.theFuncRel perZ [ F ] [ G ] F G) + !' answer)) + !' f g !'⋆π₁≡f !'⋆π₂≡g } binProductsRT : BinProducts RT binProductsRT (X , perX) (Y , perY) = binProductRT perX perY From 3fe30b94445d3a1478c2d08d935ca3ab9590cd21 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 13 Mar 2024 14:08:48 +0530 Subject: [PATCH 32/61] Add equalisation of equalisers --- src/Realizability/Topos/Equalizer.agda | 91 ++++++++++++++++++++++---- 1 file changed, 80 insertions(+), 11 deletions(-) diff --git a/src/Realizability/Topos/Equalizer.agda b/src/Realizability/Topos/Equalizer.agda index e9d80a0..4e8487c 100644 --- a/src/Realizability/Topos/Equalizer.agda +++ b/src/Realizability/Topos/Equalizer.agda @@ -83,6 +83,7 @@ equalizerUnivProp : → (inc : RTMorphism equalizerOb perX) → Type _ equalizerUnivProp {X} {Y} perX perY f g equalizerOb inc = + ((composeRTMorphism _ _ _ inc f) ≡ (composeRTMorphism _ _ _ inc g)) × ∀ {Z : Type ℓ'} (perZ : PartialEquivalenceRelation Z) (inc' : RTMorphism perZ perX) → (composeRTMorphism _ _ _ inc' f) ≡ (composeRTMorphism _ _ _ inc' g) ----------------------------------------------------------------------------------- @@ -263,6 +264,71 @@ module _ equalizerMorphism : ∀ (F G : FunctionalRelation perX perY) → RTMorphism (equalizerPer F G) perX equalizerMorphism F G = [ equalizerFuncRel F G ] + opaque + unfolding equalizerMorphism + unfolding equalizerFuncRel + unfolding composeRTMorphism + inc⋆f≡inc⋆g : ∀ (F G : FunctionalRelation perX perY) → composeRTMorphism _ _ _ (equalizerMorphism F G) [ F ] ≡ composeRTMorphism _ _ _ (equalizerMorphism F G) [ G ] + inc⋆f≡inc⋆g F G = + let + answer = + do + (relG , relG⊩isRelationalG) ← G .isRelational + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + (relF , relF⊩isRelationalF) ← F .isRelational + (sX , sX⊩isSymmetricX) ← perX .isSymmetric + (stCF , stCF⊩isStrictCodomainF) ← F .isStrictCodomain + let + realizer : ApplStrTerm as 1 + realizer = + ` pair ̇ + (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))))) ̇ + (` relG ̇ (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))) ̇ + (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))) ̇ + (` relF ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ # fzero) ̇ (` stCF ̇ (` pr₂ ̇ # fzero))))) + return + (λ* realizer , + -- unfold everything and bring it back in together + (λ x y r r⊩∃ → + do + (x' , (⊩x~x' , ∃y) , ⊩Fx'y) ← r⊩∃ + (y' , ⊩Fxy' , ⊩Gxy') ← ∃y + let + y'~y = + svF⊩isSingleValuedF x y' y _ _ ⊩Fxy' (relF⊩isRelationalF x' x y y _ _ _ (sX⊩isSymmetricX x x' _ ⊩x~x') ⊩Fx'y (stCF⊩isStrictCodomainF x' y _ ⊩Fx'y)) + return + (x' , + (subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x')) + (sym (cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (λ*ComputationRule realizer (r ∷ [])) ∙ cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₁pxy≡x _ _)) + ⊩x~x' , + do + return + (y' , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (x , y')) + (sym + (cong (λ x → pr₁ ⨾ (pr₂ ⨾ (pr₁ ⨾ x))) (λ*ComputationRule realizer (r ∷ [])) ∙ + cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₁pxy≡x _ _)) + ⊩Fxy' , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (x , y')) + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ (pr₁ ⨾ x))) (λ*ComputationRule realizer (r ∷ [])) ∙ + cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₂pxy≡y _ _)) + ⊩Gxy')) , + subst + (λ r' → r' ⊩ ∣ G .relation ∣ (x' , y)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (relG⊩isRelationalG x x' y' y _ _ _ ⊩x~x' ⊩Gxy' y'~y)))) + in + eq/ _ _ + (answer , F≤G→G≤F (equalizerPer F G) perY (composeFuncRel _ _ _ (equalizerFuncRel F G) F) (composeFuncRel _ _ _ (equalizerFuncRel F G) G) answer) + module UnivProp (F G : FunctionalRelation perX perY) {Z : Type ℓ'} @@ -537,16 +603,19 @@ module _ {P = λ f g → ∃[ equalizerOb ∈ PartialEquivalenceRelation X ] ∃[ inc ∈ RTMorphism equalizerOb perX ] (equalizerUnivProp perX perY f g equalizerOb inc)} (λ f g → isPropPropTrunc) (λ F G → - return - ((equalizerPer F G) , - (return - ((equalizerMorphism F G) , - (λ perZ h h⋆f≡h⋆g → - uniqueExists - (UnivProp.mainMap F G perZ h h⋆f≡h⋆g .fst) - (UnivProp.mainMap F G perZ h h⋆f≡h⋆g .snd .fst) - (λ !' → squash/ _ _) - λ !' !'⋆inc≡h → - sym (UnivProp.mainMap F G perZ h h⋆f≡h⋆g .snd .snd !' !'⋆inc≡h)))))) + return + ((equalizerPer F G) , + (return + ((equalizerMorphism F G) , + ((inc⋆f≡inc⋆g F G) , + (λ {Z} perZ inc' inc'⋆f≡inc'⋆g → + let + (! , !⋆inc≡inc' , unique!) = UnivProp.mainMap F G perZ inc' inc'⋆f≡inc'⋆g + in + uniqueExists + ! + !⋆inc≡inc' + (λ ! → squash/ _ _) + λ !' !'⋆inc≡inc' → sym (unique! !' !'⋆inc≡inc'))))))) f g From adf534a633d008a3f6b822899eb807018bf6f104 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 14 Mar 2024 19:42:16 +0530 Subject: [PATCH 33/61] Functional relations that represent monics --- src/Realizability/Topos/MorphismProps.agda | 259 +++++++++++++++++++++ src/Realizability/Topos/Pullbacks.agda | 39 ++++ 2 files changed, 298 insertions(+) create mode 100644 src/Realizability/Topos/MorphismProps.agda create mode 100644 src/Realizability/Topos/Pullbacks.agda diff --git a/src/Realizability/Topos/MorphismProps.agda b/src/Realizability/Topos/MorphismProps.agda new file mode 100644 index 0000000..3da3e00 --- /dev/null +++ b/src/Realizability/Topos/MorphismProps.agda @@ -0,0 +1,259 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Vec +open import Cubical.Data.Nat +open import Cubical.Data.FinData +open import Cubical.Data.Fin hiding (Fin; _/_) +open import Cubical.Data.Sigma +open import Cubical.Data.Empty +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Morphism +open import Cubical.Relation.Binary + +module Realizability.Topos.MorphismProps + {ℓ ℓ' ℓ''} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.Equalizer {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.BinProducts {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism +open PartialEquivalenceRelation +open FunctionalRelation +open Category RT + +private λ*ComputationRule = `λ*ComputationRule as fefermanStructure +private λ* = `λ* as fefermanStructure + +-- Monics in RT +module _ {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) (F : FunctionalRelation perX perY) where + + opaque + isInjectiveFuncRel : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) + isInjectiveFuncRel = + ∃[ inj ∈ A ] (∀ x x' y r₁ r₂ → r₁ ⊩ ∣ F .relation ∣ (x , y) → r₂ ⊩ ∣ F .relation ∣ (x' , y) → (inj ⨾ r₁ ⨾ r₂) ⊩ ∣ perX .equality ∣ (x , x')) + + opaque + unfolding isInjectiveFuncRel + isPropIsInjectiveFuncRel : isProp isInjectiveFuncRel + isPropIsInjectiveFuncRel = isPropPropTrunc + + -- This is the easier part + -- Essentially just a giant realizer that uses the injectivity + opaque + unfolding composeRTMorphism + unfolding composeFuncRel + unfolding isInjectiveFuncRel + isInjectiveFuncRel→isMonic : isInjectiveFuncRel → isMonic RT [ F ] + isInjectiveFuncRel→isMonic isInjectiveF {Z , perZ} {a} {b} a⋆[F]≡b⋆[F] = + SQ.elimProp2 + {P = λ a b → composeRTMorphism _ _ _ a [ F ] ≡ composeRTMorphism _ _ _ b [ F ] → a ≡ b} + (λ a b → isPropΠ λ _ → squash/ a b) + (λ A B A⋆F≡B⋆F → + let + (p , q) = SQ.effective (isPropValuedBientailment perZ perY) (isEquivRelBientailment perZ perY) _ _ A⋆F≡B⋆F + answer = + do + (p , p⊩A⋆F≤B⋆F) ← p + (stCA , stCA⊩isStrictCodomainA) ← A .isStrictCodomain + (stDA , stDA⊩isStrictDomainA) ← A .isStrictDomain + (tlF , tlF⊩isTotalF) ← F .isTotal + (relB , relB⊩isRelationalB) ← B .isRelational + (injF , injF⊩isInjectiveF) ← isInjectiveF + let + realizer : ApplStrTerm as 1 + realizer = + ` relB ̇ (` stDA ̇ # fzero) ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCA ̇ # fzero))))) ̇ + (` injF ̇ (` pr₂ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCA ̇ # fzero))))) ̇ + (` tlF ̇ (` stCA ̇ # fzero))) + return + (λ* realizer , + (λ z x r r⊩Azx → + transport + (propTruncIdempotent (B .relation .isPropValued _ _)) + (do + let + x~x = stCA⊩isStrictCodomainA z x r r⊩Azx + z~z = stDA⊩isStrictDomainA z x r r⊩Azx + (y , ⊩Fxy) ← tlF⊩isTotalF x _ x~x + (x' , ⊩Bzx' , ⊩Fx'y) ← + p⊩A⋆F≤B⋆F + z y + (pair ⨾ r ⨾ (tlF ⨾ (stCA ⨾ r))) + (return + (x , + subst (λ r' → r' ⊩ ∣ A .relation ∣ (z , x)) (sym (pr₁pxy≡x _ _)) r⊩Azx , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym (pr₂pxy≡y _ _)) ⊩Fxy)) + let + x'~x = injF⊩isInjectiveF x' x y _ _ ⊩Fx'y ⊩Fxy -- this is the only place where we actually need the injectivity + return + (subst + (λ r' → r' ⊩ ∣ B .relation ∣ (z , x)) + (sym (λ*ComputationRule realizer (r ∷ []))) + (relB⊩isRelationalB z z x' x _ _ _ z~z ⊩Bzx' x'~x))))) + in + eq/ A B (answer , F≤G→G≤F perZ perX A B answer)) + a b + a⋆[F]≡b⋆[F] + + opaque + unfolding binProdPr₁RT + unfolding binProdPr₁FuncRel + unfolding binProdPr₂FuncRel + unfolding equalizerMorphism + unfolding composeRTMorphism + + π₁ : FunctionalRelation (binProdObRT perX perX) perX + π₁ = binProdPr₁FuncRel perX perX + + π₂ : FunctionalRelation (binProdObRT perX perX) perX + π₂ = binProdPr₂FuncRel perX perX + + kernelPairEqualizerFuncRel : + FunctionalRelation + (equalizerPer -- hehe + (binProdObRT perX perX) perY + ([ π₁ ] ⋆ [ F ]) + ([ π₂ ] ⋆ [ F ]) + (composeFuncRel _ _ _ π₁ F) + (composeFuncRel _ _ _ π₂ F)) + (binProdObRT perX perX) + kernelPairEqualizerFuncRel = + equalizerFuncRel _ _ + ((binProdPr₁RT perX perX) ⋆ [ F ]) + ((binProdPr₂RT perX perX) ⋆ [ F ]) + (composeFuncRel _ _ _ (binProdPr₁FuncRel perX perX) F) + (composeFuncRel _ _ _ (binProdPr₂FuncRel perX perX) F) + + kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ : + composeRTMorphism _ _ _ [ kernelPairEqualizerFuncRel ] (composeRTMorphism _ _ _ [ π₁ ] [ F ]) ≡ composeRTMorphism _ _ _ [ kernelPairEqualizerFuncRel ] (composeRTMorphism _ _ _ [ π₂ ] [ F ]) + kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ = + inc⋆f≡inc⋆g + (binProdObRT perX perX) perY + (composeRTMorphism _ _ _ [ π₁ ] [ F ]) + (composeRTMorphism _ _ _ [ π₂ ] [ F ]) + (composeFuncRel _ _ _ π₁ F) + (composeFuncRel _ _ _ π₂ F) + + mainKernelPairEquation : ([ kernelPairEqualizerFuncRel ] ⋆ [ π₁ ]) ⋆ [ F ] ≡ ([ kernelPairEqualizerFuncRel ] ⋆ [ π₂ ]) ⋆ [ F ] + mainKernelPairEquation = + ([ kernelPairEqualizerFuncRel ] ⋆ [ π₁ ]) ⋆ [ F ] + ≡⟨ ⋆Assoc [ kernelPairEqualizerFuncRel ] [ π₁ ] [ F ] ⟩ -- Agda cannot solve these as constraints 😔 + [ kernelPairEqualizerFuncRel ] ⋆ ([ π₁ ] ⋆ [ F ]) + ≡⟨ kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ ⟩ + [ kernelPairEqualizerFuncRel ] ⋆ ([ π₂ ] ⋆ [ F ]) + ≡⟨ sym (⋆Assoc [ kernelPairEqualizerFuncRel ] [ π₂ ] [ F ]) ⟩ + ([ kernelPairEqualizerFuncRel ] ⋆ [ π₂ ]) ⋆ [ F ] + ∎ + + opaque + unfolding isInjectiveFuncRel + unfolding composeRTMorphism + unfolding kernelPairEqualizerFuncRel + unfolding equalizerFuncRel + unfolding equalizerPer + unfolding binProdPr₁RT + unfolding binProdPr₂FuncRel + isMonic→isInjectiveFuncRel : isMonic RT [ F ] → isInjectiveFuncRel + isMonic→isInjectiveFuncRel isMonicF = + do + let + equation = isMonicF {a = [ kernelPairEqualizerFuncRel ] ⋆ [ π₁ ]} {a' = [ kernelPairEqualizerFuncRel ] ⋆ [ π₂ ]} mainKernelPairEquation + (p , q) = SQ.effective (isPropValuedBientailment _ _) (isEquivRelBientailment _ _) _ _ equation + (p , p⊩kπ₁≤kπ₂) ← p + (q , q⊩kπ₂≤kπ₁) ← q + (stCF , stCF⊩isStrictCodomainF) ← F .isStrictCodomain + (stDF , stDF⊩isStrictDomainF) ← F .isStrictDomain + (s , s⊩isSymmetricEqX) ← perX .isSymmetric + (t , t⊩isTransitiveEqX) ← perX .isTransitive + return + ({!!} , + (λ x₁ x₂ y r₁ r₂ r₁⊩Fx₁y r₂⊩Fx₂y → + let + x₁~x₁ = stDF⊩isStrictDomainF x₁ y r₁ r₁⊩Fx₁y + x₂~x₂ = stDF⊩isStrictDomainF x₂ y r₂ r₂⊩Fx₂y + foo = + p⊩kπ₁≤kπ₂ (x₁ , x₂) x₁ + (pair ⨾ + (pair ⨾ + (pair ⨾ (stDF ⨾ r₁) ⨾ (stDF ⨾ r₂)) ⨾ + (pair ⨾ (pair ⨾ (pair ⨾ (stDF ⨾ r₁) ⨾ (stDF ⨾ r₂)) ⨾ r₁) ⨾ (pair ⨾ (pair ⨾ (stDF ⨾ r₂) ⨾ (stDF ⨾ r₁)) ⨾ r₂))) ⨾ + (pair ⨾ (stDF ⨾ r₁) ⨾ (stDF ⨾ r₂))) + (return + (((x₁ , x₂)) , + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym (cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₁pxy≡x _ _)) x₁~x₁ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) (sym (cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₂pxy≡y _ _)) x₂~x₂) , + return + (y , + return + (x₁ , + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) + (sym + (cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₁ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₁pxy≡x _ _)) + x₁~x₁ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) + (sym + (cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₁ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₂pxy≡y _ _)) + x₂~x₂) , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x₁ , y)) + (sym + (cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₂ ⨾ x))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₂pxy≡y _ _)) + r₁⊩Fx₁y) , + return + (x₂ , + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) + (sym + (cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₁pxy≡x _ _)) + x₂~x₂ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym ({!!} ∙ {!!} ∙ {!!} ∙ {!!} ∙ {!!})) x₁~x₁) , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x₂ , y)) (sym {!!}) r₂⊩Fx₂y))) , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym {!!}) x₁~x₁ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) (sym {!!}) x₂~x₂)) + in + transport + (propTruncIdempotent {!!}) + (do + ((x₁' , x₂') , ((x₁~x₁' , x₂~x₂') , kp2) , p2) ← foo + (y' , bar1 , bar2) ← kp2 + (x₁'' , (x₁~x₁'' , x₂~'x₂) , ⊩Fx₁''y') ← bar1 + (x₂'' , (x₂~x₂'' , x₁~'x₁) , ⊩Fx₂''y') ← bar2 + let + (x₂'~x₁ , foo') = p2 + return + (subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₂)) + {!!} + (t⊩isTransitiveEqX x₁ x₂' x₂ _ _ (s⊩isSymmetricEqX x₂' x₁ _ x₂'~x₁) (s⊩isSymmetricEqX x₂ x₂' _ x₂~x₂')))))) diff --git a/src/Realizability/Topos/Pullbacks.agda b/src/Realizability/Topos/Pullbacks.agda new file mode 100644 index 0000000..0063c0d --- /dev/null +++ b/src/Realizability/Topos/Pullbacks.agda @@ -0,0 +1,39 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Unit +open import Cubical.Data.Empty +open import Cubical.Data.Fin +open import Cubical.Data.Vec +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation as PT +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.BinProduct + +module Realizability.Topos.Pullbacks + {ℓ ℓ' ℓ''} {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + +open FunctionalRelation +open PartialEquivalenceRelation + From 73a0508994dd07ff8dc50394ce099d3e9f70b0f4 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 26 Mar 2024 18:04:27 +0530 Subject: [PATCH 34/61] Characterise subobjects in terms of strict relations --- .../Topos/FunctionalRelation.agda | 52 +- src/Realizability/Topos/MonicReprFuncRel.agda | 281 ++++++++ src/Realizability/Topos/MorphismProps.agda | 40 +- src/Realizability/Topos/Object.agda | 5 +- src/Realizability/Topos/StrictRelation.agda | 635 ++++++++++++++++++ 5 files changed, 976 insertions(+), 37 deletions(-) create mode 100644 src/Realizability/Topos/MonicReprFuncRel.agda create mode 100644 src/Realizability/Topos/StrictRelation.agda diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 50a9738..6173fb4 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -38,46 +38,48 @@ private λ* = `λ* as fefermanStructure open PartialEquivalenceRelation -record isFunctionalRelation +module _ {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) - (relation : Predicate (X × Y)) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - equalityX = perX .equality - equalityY = perY .equality + (relation : Predicate (X × Y)) where + equalityX = perX .equality + equalityY = perY .equality + + realizesStrictDomain : A → Type _ + realizesStrictDomain stD = (∀ x y r → r ⊩ ∣ relation ∣ (x , y) → (stD ⨾ r) ⊩ ∣ equalityX ∣ (x , x)) - field - isStrictDomain : - ∃[ stD ∈ A ] - (∀ x y r - → r ⊩ ∣ relation ∣ (x , y) - ---------------------------------- - → (stD ⨾ r) ⊩ ∣ equalityX ∣ (x , x)) - isStrictCodomain : - ∃[ stC ∈ A ] - (∀ x y r - → r ⊩ ∣ relation ∣ (x , y) - ---------------------------------- - → (stC ⨾ r) ⊩ ∣ equalityY ∣ (y , y)) - isRelational : - ∃[ rl ∈ A ] + realizesStrictCodomain : A → Type _ + realizesStrictCodomain stC = (∀ x y r → r ⊩ ∣ relation ∣ (x , y) → (stC ⨾ r) ⊩ ∣ equalityY ∣ (y , y)) + + realizesRelational : A → Type _ + realizesRelational rel = (∀ x x' y y' a b c → a ⊩ ∣ equalityX ∣ (x , x') → b ⊩ ∣ relation ∣ (x , y) → c ⊩ ∣ equalityY ∣ (y , y') ------------------------------------------ - → (rl ⨾ a ⨾ b ⨾ c) ⊩ ∣ relation ∣ (x' , y')) - isSingleValued : - ∃[ sv ∈ A ] + → (rel ⨾ a ⨾ b ⨾ c) ⊩ ∣ relation ∣ (x' , y')) + + realizesSingleValued : A → Type _ + realizesSingleValued sv = (∀ x y y' r₁ r₂ → r₁ ⊩ ∣ relation ∣ (x , y) → r₂ ⊩ ∣ relation ∣ (x , y') ----------------------------------- → (sv ⨾ r₁ ⨾ r₂) ⊩ ∣ equalityY ∣ (y , y')) - isTotal : - ∃[ tl ∈ A ] - (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) + realizesTotal : A → Type _ + realizesTotal tl = + (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) + + record isFunctionalRelation : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + field + isStrictDomain : ∃[ stD ∈ A ] (realizesStrictDomain stD) + isStrictCodomain : ∃[ stC ∈ A ] (realizesStrictCodomain stC) + isRelational : ∃[ rl ∈ A ] (realizesRelational rl) + isSingleValued : ∃[ sv ∈ A ] (realizesSingleValued sv) + isTotal : ∃[ tl ∈ A ] (realizesTotal tl) record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where field diff --git a/src/Realizability/Topos/MonicReprFuncRel.agda b/src/Realizability/Topos/MonicReprFuncRel.agda new file mode 100644 index 0000000..195b29f --- /dev/null +++ b/src/Realizability/Topos/MonicReprFuncRel.agda @@ -0,0 +1,281 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Vec +open import Cubical.Data.Nat +open import Cubical.Data.FinData +open import Cubical.Data.Fin hiding (Fin; _/_) +open import Cubical.Data.Sigma +open import Cubical.Data.Empty +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Morphism +open import Cubical.Relation.Binary + +module Realizability.Topos.MonicReprFuncRel + {ℓ ℓ' ℓ''} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.Equalizer {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.BinProducts {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism +open PartialEquivalenceRelation +open FunctionalRelation +open Category RT + +private λ*ComputationRule = `λ*ComputationRule as fefermanStructure +private λ* = `λ* as fefermanStructure + +-- Monics in RT +module _ {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) (F : FunctionalRelation perX perY) where + + opaque + isInjectiveFuncRel : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) + isInjectiveFuncRel = + ∃[ inj ∈ A ] (∀ x x' y r₁ r₂ → r₁ ⊩ ∣ F .relation ∣ (x , y) → r₂ ⊩ ∣ F .relation ∣ (x' , y) → (inj ⨾ r₁ ⨾ r₂) ⊩ ∣ perX .equality ∣ (x , x')) + + opaque + unfolding isInjectiveFuncRel + isPropIsInjectiveFuncRel : isProp isInjectiveFuncRel + isPropIsInjectiveFuncRel = isPropPropTrunc + + -- This is the easier part + -- Essentially just a giant realizer that uses the injectivity + opaque + unfolding composeRTMorphism + unfolding composeFuncRel + unfolding isInjectiveFuncRel + isInjectiveFuncRel→isMonic : isInjectiveFuncRel → isMonic RT [ F ] + isInjectiveFuncRel→isMonic isInjectiveF {Z , perZ} {a} {b} a⋆[F]≡b⋆[F] = + SQ.elimProp2 + {P = λ a b → composeRTMorphism _ _ _ a [ F ] ≡ composeRTMorphism _ _ _ b [ F ] → a ≡ b} + (λ a b → isPropΠ λ _ → squash/ a b) + (λ A B A⋆F≡B⋆F → + let + (p , q) = SQ.effective (isPropValuedBientailment perZ perY) (isEquivRelBientailment perZ perY) _ _ A⋆F≡B⋆F + answer = + do + (p , p⊩A⋆F≤B⋆F) ← p + (stCA , stCA⊩isStrictCodomainA) ← A .isStrictCodomain + (stDA , stDA⊩isStrictDomainA) ← A .isStrictDomain + (tlF , tlF⊩isTotalF) ← F .isTotal + (relB , relB⊩isRelationalB) ← B .isRelational + (injF , injF⊩isInjectiveF) ← isInjectiveF + let + realizer : ApplStrTerm as 1 + realizer = + ` relB ̇ (` stDA ̇ # fzero) ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCA ̇ # fzero))))) ̇ + (` injF ̇ (` pr₂ ̇ (` p ̇ (` pair ̇ # fzero ̇ (` tlF ̇ (` stCA ̇ # fzero))))) ̇ + (` tlF ̇ (` stCA ̇ # fzero))) + return + (λ* realizer , + (λ z x r r⊩Azx → + transport + (propTruncIdempotent (B .relation .isPropValued _ _)) + (do + let + x~x = stCA⊩isStrictCodomainA z x r r⊩Azx + z~z = stDA⊩isStrictDomainA z x r r⊩Azx + (y , ⊩Fxy) ← tlF⊩isTotalF x _ x~x + (x' , ⊩Bzx' , ⊩Fx'y) ← + p⊩A⋆F≤B⋆F + z y + (pair ⨾ r ⨾ (tlF ⨾ (stCA ⨾ r))) + (return + (x , + subst (λ r' → r' ⊩ ∣ A .relation ∣ (z , x)) (sym (pr₁pxy≡x _ _)) r⊩Azx , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x , y)) (sym (pr₂pxy≡y _ _)) ⊩Fxy)) + let + x'~x = injF⊩isInjectiveF x' x y _ _ ⊩Fx'y ⊩Fxy -- this is the only place where we actually need the injectivity + return + (subst + (λ r' → r' ⊩ ∣ B .relation ∣ (z , x)) + (sym (λ*ComputationRule realizer (r ∷ []))) + (relB⊩isRelationalB z z x' x _ _ _ z~z ⊩Bzx' x'~x))))) + in + eq/ A B (answer , F≤G→G≤F perZ perX A B answer)) + a b + a⋆[F]≡b⋆[F] + + opaque + unfolding binProdPr₁RT + unfolding binProdPr₁FuncRel + unfolding binProdPr₂FuncRel + unfolding equalizerMorphism + unfolding composeRTMorphism + + π₁ : FunctionalRelation (binProdObRT perX perX) perX + π₁ = binProdPr₁FuncRel perX perX + + π₂ : FunctionalRelation (binProdObRT perX perX) perX + π₂ = binProdPr₂FuncRel perX perX + + kernelPairEqualizerFuncRel : + FunctionalRelation + (equalizerPer -- hehe + (binProdObRT perX perX) perY + ([ π₁ ] ⋆ [ F ]) + ([ π₂ ] ⋆ [ F ]) + (composeFuncRel _ _ _ π₁ F) + (composeFuncRel _ _ _ π₂ F)) + (binProdObRT perX perX) + kernelPairEqualizerFuncRel = + equalizerFuncRel _ _ + ((binProdPr₁RT perX perX) ⋆ [ F ]) + ((binProdPr₂RT perX perX) ⋆ [ F ]) + (composeFuncRel _ _ _ (binProdPr₁FuncRel perX perX) F) + (composeFuncRel _ _ _ (binProdPr₂FuncRel perX perX) F) + + kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ : + composeRTMorphism _ _ _ [ kernelPairEqualizerFuncRel ] (composeRTMorphism _ _ _ [ π₁ ] [ F ]) ≡ composeRTMorphism _ _ _ [ kernelPairEqualizerFuncRel ] (composeRTMorphism _ _ _ [ π₂ ] [ F ]) + kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ = + inc⋆f≡inc⋆g + (binProdObRT perX perX) perY + (composeRTMorphism _ _ _ [ π₁ ] [ F ]) + (composeRTMorphism _ _ _ [ π₂ ] [ F ]) + (composeFuncRel _ _ _ π₁ F) + (composeFuncRel _ _ _ π₂ F) + + mainKernelPairEquation : ([ kernelPairEqualizerFuncRel ] ⋆ [ π₁ ]) ⋆ [ F ] ≡ ([ kernelPairEqualizerFuncRel ] ⋆ [ π₂ ]) ⋆ [ F ] + mainKernelPairEquation = + ([ kernelPairEqualizerFuncRel ] ⋆ [ π₁ ]) ⋆ [ F ] + ≡⟨ ⋆Assoc [ kernelPairEqualizerFuncRel ] [ π₁ ] [ F ] ⟩ -- Agda cannot solve these as constraints 😔 + [ kernelPairEqualizerFuncRel ] ⋆ ([ π₁ ] ⋆ [ F ]) + ≡⟨ kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ ⟩ + [ kernelPairEqualizerFuncRel ] ⋆ ([ π₂ ] ⋆ [ F ]) + ≡⟨ sym (⋆Assoc [ kernelPairEqualizerFuncRel ] [ π₂ ] [ F ]) ⟩ + ([ kernelPairEqualizerFuncRel ] ⋆ [ π₂ ]) ⋆ [ F ] + ∎ + + opaque + unfolding isInjectiveFuncRel + unfolding composeRTMorphism + unfolding kernelPairEqualizerFuncRel + unfolding equalizerFuncRel + unfolding equalizerPer + unfolding binProdPr₁RT + unfolding binProdPr₂FuncRel + isMonic→isInjectiveFuncRel : isMonic RT [ F ] → isInjectiveFuncRel + isMonic→isInjectiveFuncRel isMonicF = + do + let + equation = isMonicF {a = [ kernelPairEqualizerFuncRel ] ⋆ [ π₁ ]} {a' = [ kernelPairEqualizerFuncRel ] ⋆ [ π₂ ]} mainKernelPairEquation + (p , q) = SQ.effective (isPropValuedBientailment _ _) (isEquivRelBientailment _ _) _ _ equation + (p , p⊩kπ₁≤kπ₂) ← p + (q , q⊩kπ₂≤kπ₁) ← q + (stCF , stCF⊩isStrictCodomainF) ← F .isStrictCodomain + (stDF , stDF⊩isStrictDomainF) ← F .isStrictDomain + (s , s⊩isSymmetricEqX) ← perX .isSymmetric + (t , t⊩isTransitiveEqX) ← perX .isTransitive + let + cursed : ApplStrTerm as 2 + cursed = + (` pair ̇ + (` pair ̇ + (` pair ̇ (` stDF ̇ # fzero) ̇ (` stDF ̇ # fone)) ̇ + (` pair ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # fzero) ̇ (` stDF ̇ # fone)) ̇ # fzero) ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # fone) ̇ (` stDF ̇ # fzero)) ̇ # fone))) ̇ + (` pair ̇ (` stDF ̇ # fzero) ̇ (` stDF ̇ # fone))) + realizer : ApplStrTerm as 2 + realizer = ` t ̇ (` s ̇ (` pr₁ ̇ (` pr₂ ̇ (` p ̇ cursed)))) ̇ (` s ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₁ ̇ (` p ̇ cursed))))) + return + (λ* realizer , + (λ x₁ x₂ y r₁ r₂ r₁⊩Fx₁y r₂⊩Fx₂y → + let + x₁~x₁ = stDF⊩isStrictDomainF x₁ y r₁ r₁⊩Fx₁y + x₂~x₂ = stDF⊩isStrictDomainF x₂ y r₂ r₂⊩Fx₂y + foo = + p⊩kπ₁≤kπ₂ (x₁ , x₂) x₁ + (pair ⨾ + (pair ⨾ + (pair ⨾ (stDF ⨾ r₁) ⨾ (stDF ⨾ r₂)) ⨾ + (pair ⨾ (pair ⨾ (pair ⨾ (stDF ⨾ r₁) ⨾ (stDF ⨾ r₂)) ⨾ r₁) ⨾ (pair ⨾ (pair ⨾ (stDF ⨾ r₂) ⨾ (stDF ⨾ r₁)) ⨾ r₂))) ⨾ + (pair ⨾ (stDF ⨾ r₁) ⨾ (stDF ⨾ r₂))) + (return + (((x₁ , x₂)) , + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym (cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₁pxy≡x _ _)) x₁~x₁ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) (sym (cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₂pxy≡y _ _)) x₂~x₂) , + return + (y , + return + (x₁ , + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) + (sym + (cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₁ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₁pxy≡x _ _)) + x₁~x₁ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) + (sym + (cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₁ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₂pxy≡y _ _)) + x₂~x₂) , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x₁ , y)) + (sym + (cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₂ ⨾ x))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₂pxy≡y _ _)) + r₁⊩Fx₁y) , + return + (x₂ , + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) + (sym + (cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ (pr₂ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₁pxy≡x _ _)) + x₂~x₂ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) + (sym + (cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₂ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₂ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₂pxy≡y _ _)) + x₁~x₁) , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x₂ , y)) + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ (pr₂ ⨾ x))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₂pxy≡y _ _)) r₂⊩Fx₂y))) , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym (cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) x₁~x₁ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) (sym (cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) x₂~x₂)) + in + transport + (propTruncIdempotent (perX .equality .isPropValued _ _)) + (do + ((x₁' , x₂') , ((x₁~x₁' , x₂~x₂') , kp2) , p2) ← foo + (y' , bar1 , bar2) ← kp2 + (x₁'' , (x₁~x₁'' , x₂~'x₂) , ⊩Fx₁''y') ← bar1 + (x₂'' , (x₂~x₂'' , x₁~'x₁) , ⊩Fx₂''y') ← bar2 + let + (x₂'~x₁ , foo') = p2 + return + (subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₂)) + (sym (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ []))) + (t⊩isTransitiveEqX x₁ x₂' x₂ _ _ (s⊩isSymmetricEqX x₂' x₁ _ x₂'~x₁) (s⊩isSymmetricEqX x₂ x₂' _ x₂~x₂')))))) diff --git a/src/Realizability/Topos/MorphismProps.agda b/src/Realizability/Topos/MorphismProps.agda index 3da3e00..e46dfbd 100644 --- a/src/Realizability/Topos/MorphismProps.agda +++ b/src/Realizability/Topos/MorphismProps.agda @@ -17,8 +17,8 @@ open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category open import Cubical.Categories.Morphism open import Cubical.Relation.Binary - -module Realizability.Topos.MorphismProps +-- Functional relations that represent monics +module Realizability.Topos.MonicReprFuncRel {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) @@ -184,8 +184,18 @@ module _ {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : Partial (stDF , stDF⊩isStrictDomainF) ← F .isStrictDomain (s , s⊩isSymmetricEqX) ← perX .isSymmetric (t , t⊩isTransitiveEqX) ← perX .isTransitive + let + cursed : ApplStrTerm as 2 + cursed = + (` pair ̇ + (` pair ̇ + (` pair ̇ (` stDF ̇ # fzero) ̇ (` stDF ̇ # fone)) ̇ + (` pair ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # fzero) ̇ (` stDF ̇ # fone)) ̇ # fzero) ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # fone) ̇ (` stDF ̇ # fzero)) ̇ # fone))) ̇ + (` pair ̇ (` stDF ̇ # fzero) ̇ (` stDF ̇ # fone))) + realizer : ApplStrTerm as 2 + realizer = ` t ̇ (` s ̇ (` pr₁ ̇ (` pr₂ ̇ (` p ̇ cursed)))) ̇ (` s ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₁ ̇ (` p ̇ cursed))))) return - ({!!} , + (λ* realizer , (λ x₁ x₂ y r₁ r₂ r₁⊩Fx₁y r₂⊩Fx₂y → let x₁~x₁ = stDF⊩isStrictDomainF x₁ y r₁ r₁⊩Fx₁y @@ -238,13 +248,25 @@ module _ {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : Partial cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₁pxy≡x _ _)) x₂~x₂ , - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym ({!!} ∙ {!!} ∙ {!!} ∙ {!!} ∙ {!!})) x₁~x₁) , - subst (λ r' → r' ⊩ ∣ F .relation ∣ (x₂ , y)) (sym {!!}) r₂⊩Fx₂y))) , - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym {!!}) x₁~x₁ , - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) (sym {!!}) x₂~x₂)) + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) + (sym + (cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₂ ⨾ (pr₂ ⨾ x)))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ (pr₂ ⨾ x))) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₂pxy≡y _ _)) + x₁~x₁) , + subst (λ r' → r' ⊩ ∣ F .relation ∣ (x₂ , y)) + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ (pr₂ ⨾ x))) (pr₁pxy≡x _ _) ∙ + cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (pr₂pxy≡y _ _) ∙ + cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ + pr₂pxy≡y _ _)) r₂⊩Fx₂y))) , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₁)) (sym (cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) x₁~x₁ , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₂)) (sym (cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) x₂~x₂)) in transport - (propTruncIdempotent {!!}) + (propTruncIdempotent (perX .equality .isPropValued _ _)) (do ((x₁' , x₂') , ((x₁~x₁' , x₂~x₂') , kp2) , p2) ← foo (y' , bar1 , bar2) ← kp2 @@ -255,5 +277,5 @@ module _ {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : Partial return (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₂)) - {!!} + (sym (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ []))) (t⊩isTransitiveEqX x₁ x₂' x₂ _ _ (s⊩isSymmetricEqX x₂' x₁ _ x₂'~x₁) (s⊩isSymmetricEqX x₂ x₂' _ x₂~x₂')))))) diff --git a/src/Realizability/Topos/Object.agda b/src/Realizability/Topos/Object.agda index 0f0bac6..e6727cf 100644 --- a/src/Realizability/Topos/Object.agda +++ b/src/Realizability/Topos/Object.agda @@ -4,6 +4,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Powerset open import Cubical.Foundations.Equiv open import Cubical.Data.Vec open import Cubical.Data.Nat @@ -18,9 +19,7 @@ module Realizability.Topos.Object (ca : CombinatoryAlgebra A) (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) where - -open import Realizability.Tripos.Logic.Syntax {ℓ = ℓ'} -open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca + open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) diff --git a/src/Realizability/Topos/StrictRelation.agda b/src/Realizability/Topos/StrictRelation.agda new file mode 100644 index 0000000..cdeab3b --- /dev/null +++ b/src/Realizability/Topos/StrictRelation.agda @@ -0,0 +1,635 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Vec +open import Cubical.Data.Nat +open import Cubical.Data.FinData +open import Cubical.Data.Fin hiding (Fin; _/_) +open import Cubical.Data.Sigma +open import Cubical.Data.Empty +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Morphism +open import Cubical.Categories.Constructions.SubObject +open import Cubical.Categories.Constructions.Slice +open import Cubical.Relation.Binary + +module Realizability.Topos.StrictRelation + {ℓ ℓ' ℓ''} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.Equalizer {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.BinProducts {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial +open import Realizability.Topos.MonicReprFuncRel {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open Morphism +open PartialEquivalenceRelation +open FunctionalRelation +open Category RT + +private λ*ComputationRule = `λ*ComputationRule as fefermanStructure +private λ* = `λ* as fefermanStructure + +record isStrictRelation {X : Type ℓ'} (perX : PartialEquivalenceRelation X) (ϕ : Predicate X) : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) where + field + isStrict : ∃[ st ∈ A ] (∀ x r → r ⊩ ∣ ϕ ∣ x → (st ⨾ r) ⊩ ∣ perX .equality ∣ (x , x)) + isRelational : ∃[ rel ∈ A ] (∀ x x' r s → r ⊩ ∣ ϕ ∣ x → s ⊩ ∣ perX .equality ∣ (x , x') → (rel ⨾ r ⨾ s) ⊩ ∣ ϕ ∣ x') + +record StrictRelation {X : Type ℓ'} (perX : PartialEquivalenceRelation X) : Type (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) where + field + predicate : Predicate X + isStrictRelationPredicate : isStrictRelation perX predicate + open isStrictRelation isStrictRelationPredicate public + +-- Every strict relation induces a subobject +module InducedSubobject {X : Type ℓ'} (perX : PartialEquivalenceRelation X) (ϕ : StrictRelation perX) where + open StrictRelation + -- the subobject induced by ϕ + {-# TERMINATING #-} + subPer : PartialEquivalenceRelation X + Predicate.isSetX (equality subPer) = isSet× (perX .isSetX) (perX .isSetX) + ∣ equality subPer ∣ (x , x') r = (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ ϕ .predicate ∣ x + isPropValued (equality subPer) (x , x') r = isProp× (perX .equality .isPropValued _ _) (ϕ .predicate .isPropValued _ _) + isPartialEquivalenceRelation.isSetX (isPerEquality subPer) = perX .isSetX + isPartialEquivalenceRelation.isSymmetric (isPerEquality subPer) = + do + -- Trivial : use symmetry of ~X and relationality of ϕ + (s , s⊩isSymmetricX) ← perX .isSymmetric + (relϕ , relϕ⊩isRelationalϕ) ← ϕ .isRelational + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` relϕ ̇ (` pr₂ ̇ # fzero) ̇ (` pr₁ ̇ # fzero)) + return + (λ* realizer , + (λ { x x' r (pr₁r⊩x~x' , pr₂r⊩ϕx) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (s⊩isSymmetricX x x' _ pr₁r⊩x~x') , + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x') + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (relϕ⊩isRelationalϕ x x' _ _ pr₂r⊩ϕx pr₁r⊩x~x') })) + isPartialEquivalenceRelation.isTransitive (isPerEquality subPer) = + do + (t , t⊩isTransitiveX) ← perX .isTransitive + (relϕ , relϕ⊩isRelationalϕ) ← ϕ .isRelational + let + realizer : ApplStrTerm as 2 + realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone)) ̇ (` pr₂ ̇ # fzero) + return + (λ* realizer , + (λ { x₁ x₂ x₃ a b (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₂~x₃ , ⊩ϕx₂) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₃)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (a ∷ b ∷ [])) ∙ pr₁pxy≡x _ _)) + (t⊩isTransitiveX x₁ x₂ x₃ _ _ ⊩x₁~x₂ ⊩x₂~x₃) , + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x₁) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (a ∷ b ∷ [])) ∙ pr₂pxy≡y _ _)) + ⊩ϕx₁ })) + + opaque + unfolding idFuncRel + {-# TERMINATING #-} + incFuncRel : FunctionalRelation subPer perX + Predicate.isSetX (relation incFuncRel) = isSet× (perX .isSetX) (perX .isSetX) + Predicate.∣ relation incFuncRel ∣ (x , x') r = (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × (pr₂ ⨾ r) ⊩ ∣ ϕ .predicate ∣ x + Predicate.isPropValued (relation incFuncRel) (x , x') r = isProp× (perX .equality .isPropValued _ _) (ϕ .predicate .isPropValued _ _) + isFunctionalRelation.isStrictDomain (isFuncRel incFuncRel) = + do + (stD , stD⊩isStrictDomain) ← idFuncRel perX .isStrictDomain + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` stD ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero) + return + (λ* realizer , + (λ { x x' r (⊩x~x' , ⊩ϕx) → + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) (stD⊩isStrictDomain x x' _ ⊩x~x')) , + (subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) ⊩ϕx) })) + isFunctionalRelation.isStrictCodomain (isFuncRel incFuncRel) = + do + (stC , stC⊩isStrictCodomain) ← idFuncRel perX .isStrictCodomain + let + realizer : ApplStrTerm as 1 + realizer = ` stC ̇ (` pr₁ ̇ # fzero) + return + (λ* realizer , + (λ { x x' r (⊩x~x' , ⊩ϕx) → subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) (sym (λ*ComputationRule realizer (r ∷ []))) (stC⊩isStrictCodomain x x' _ ⊩x~x')})) + isFunctionalRelation.isRelational (isFuncRel incFuncRel) = + do + (relX , relX⊩isRelationalX) ← idFuncRel perX .isRelational + (relϕ , relϕ⊩isRelationalϕ) ← ϕ .isRelational + let + realizer : ApplStrTerm as 3 + realizer = ` pair ̇ (` relX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone) ̇ # (fsuc fone)) ̇ (` relϕ ̇ (` pr₂ ̇ # fzero) ̇ (` pr₁ ̇ # fzero)) + return + (λ* realizer , + (λ { x₁ x₂ x₃ x₄ a b c (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₁~x₃ , ⊩ϕx₁') c⊩x₃~x₄ → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₄)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) + (relX⊩isRelationalX x₁ x₂ x₃ x₄ _ _ _ ⊩x₁~x₂ ⊩x₁~x₃ c⊩x₃~x₄) , + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x₂) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (a ∷ b ∷ c ∷ [])) ∙ pr₂pxy≡y _ _)) + (relϕ⊩isRelationalϕ x₁ x₂ _ _ ⊩ϕx₁ ⊩x₁~x₂) })) + isFunctionalRelation.isSingleValued (isFuncRel incFuncRel) = + do + (sv , sv⊩isSingleValuedX) ← idFuncRel perX .isSingleValued + let + realizer : ApplStrTerm as 2 + realizer = ` sv ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone) + return + (λ* realizer , + (λ { x x' x'' r₁ r₂ (⊩x~x' , ⊩ϕx) (⊩x~x'' , ⊩ϕx') → + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x'')) (sym (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ []))) (sv⊩isSingleValuedX x x' x'' _ _ ⊩x~x' ⊩x~x'') })) + isFunctionalRelation.isTotal (isFuncRel incFuncRel) = + do + return + (Id , + (λ { x r (pr₁r⊩x~x , pr₂r⊩ϕx) → + return + (x , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (cong (λ x → pr₁ ⨾ x) (sym (Ida≡a _))) pr₁r⊩x~x , + subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (cong (λ x → pr₂ ⨾ x) (sym (Ida≡a _))) pr₂r⊩ϕx) })) + + opaque + unfolding isInjectiveFuncRel + unfolding incFuncRel + isInjectiveIncFuncRel : isInjectiveFuncRel subPer perX incFuncRel + isInjectiveIncFuncRel = + do + (t , t⊩isTransitiveX) ← perX .isTransitive + (s , s⊩isSymmetricX) ← perX .isSymmetric + let + realizer : ApplStrTerm as 2 + realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # fzero) ̇ (` s ̇ (` pr₁ ̇ # fone))) ̇ (` pr₂ ̇ # fzero) + return + (λ* realizer , + (λ x₁ x₂ x₃ r₁ r₂ (⊩x₁~x₃ , ⊩ϕx₁) (⊩x₂~x₃ , ⊩ϕx₂) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₁ , x₂)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ [])) ∙ pr₁pxy≡x _ _)) + (t⊩isTransitiveX x₁ x₃ x₂ _ _ ⊩x₁~x₃ (s⊩isSymmetricX x₂ x₃ _ ⊩x₂~x₃)) , + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x₁) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ [])) ∙ pr₂pxy≡y _ _)) + ⊩ϕx₁)) + + isMonicInc : isMonic RT [ incFuncRel ] + isMonicInc = isInjectiveFuncRel→isMonic subPer perX incFuncRel isInjectiveIncFuncRel + +-- Every subobject representing functional relation is isomorphic (as a subobject) to a subobject induced by a strict relation +module SubobjectIsoMonicFuncRel + {X Y : Type ℓ'} + (perX : PartialEquivalenceRelation X) + (perY : PartialEquivalenceRelation Y) + (F : FunctionalRelation perY perX) + (isMonicF : isMonic RT [ F ]) where + + {-# TERMINATING #-} + ψ : StrictRelation perX + Predicate.isSetX (StrictRelation.predicate ψ) = perX .isSetX + Predicate.∣ StrictRelation.predicate ψ ∣ x r = ∃[ y ∈ Y ] r ⊩ ∣ F .relation ∣ (y , x) + Predicate.isPropValued (StrictRelation.predicate ψ) x r = isPropPropTrunc + isStrictRelation.isStrict (StrictRelation.isStrictRelationPredicate ψ) = + do + (stCF , stCF⊩isStrictCodomainF) ← F .isStrictCodomain + return + (stCF , + (λ x r r⊩∃y → + transport + (propTruncIdempotent (perX .equality .isPropValued _ _)) + (do + (y , ⊩Fyx) ← r⊩∃y + return (stCF⊩isStrictCodomainF y x _ ⊩Fyx)))) + isStrictRelation.isRelational (StrictRelation.isStrictRelationPredicate ψ) = + do + (relF , relF⊩isRelationalF) ← F .isRelational + (stDF , stDF⊩isStrictDomainF) ← F .isStrictDomain + let + realizer : ApplStrTerm as 2 + realizer = ` relF ̇ (` stDF ̇ # fzero) ̇ # fzero ̇ # fone + return + (λ* realizer , + (λ x x' r s r⊩∃y s⊩x~x' → + do + (y , ⊩Fyx) ← r⊩∃y + return + (y , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (y , x')) + (sym (λ*ComputationRule realizer (r ∷ s ∷ []))) + (relF⊩isRelationalF y y x x' _ _ _ (stDF⊩isStrictDomainF y x _ ⊩Fyx) ⊩Fyx s⊩x~x')))) + + perψ : PartialEquivalenceRelation X + perψ = InducedSubobject.subPer perX ψ + + -- ≤ as subobjects + -- TODO : formalise the preorder category of subobjects + {-# TERMINATING #-} + perY≤perψFuncRel : FunctionalRelation perY perψ + Predicate.isSetX (relation perY≤perψFuncRel) = isSet× (perY .isSetX) (perX .isSetX) + Predicate.∣ relation perY≤perψFuncRel ∣ = ∣ F .relation ∣ + Predicate.isPropValued (relation perY≤perψFuncRel) = F .relation .isPropValued + isFunctionalRelation.isStrictDomain (isFuncRel perY≤perψFuncRel) = + isFunctionalRelation.isStrictDomain (F .isFuncRel) + isFunctionalRelation.isStrictCodomain (isFuncRel perY≤perψFuncRel) = + do + (stCF , stCF⊩isStrictCodomain) ← F .isStrictCodomain + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` stCF ̇ # fzero) ̇ # fzero + return + (λ* realizer , + (λ y x r ⊩Fyx → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (stCF⊩isStrictCodomain y x _ ⊩Fyx) , + ∣ y , + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (y , x)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) + ⊩Fyx ∣₁)) + isFunctionalRelation.isRelational (isFuncRel perY≤perψFuncRel) = + do + (relF , relF⊩isRelationalF) ← F .isRelational + let + realizer : ApplStrTerm as 3 + realizer = ` relF ̇ # fzero ̇ # fone ̇ (` pr₁ ̇ # (fsuc fone)) + return + (λ* realizer , + (λ { y y' x x' a b c ⊩y~y' ⊩Fyx (⊩x~x' , ⊩Fy''x) → + subst (λ r' → r' ⊩ ∣ F .relation ∣ (y' , x')) (sym (λ*ComputationRule realizer (a ∷ b ∷ c ∷ []))) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') })) + isFunctionalRelation.isSingleValued (isFuncRel perY≤perψFuncRel) = + do + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + let + realizer : ApplStrTerm as 2 + realizer = ` pair ̇ (` svF ̇ # fzero ̇ # fone) ̇ # fzero + return + (λ* realizer , + (λ y x x' r₁ r₂ ⊩Fyx ⊩Fyx' → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ [])) ∙ pr₁pxy≡x _ _)) + (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') , + ∣ y , + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (y , x)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ [])) ∙ pr₂pxy≡y _ _)) + ⊩Fyx) ∣₁)) + isFunctionalRelation.isTotal (isFuncRel perY≤perψFuncRel) = + do + (tlF , tlF⊩isTotalF) ← F .isTotal + return + (tlF , + (λ y r ⊩y~y → + do + (x , ⊩Fyx) ← tlF⊩isTotalF y _ ⊩y~y + return (x , ⊩Fyx))) + + -- perY truly is ≤ perψ + opaque + unfolding composeRTMorphism + unfolding InducedSubobject.incFuncRel + perY≤perψCommutes : [ perY≤perψFuncRel ] ⋆ [ InducedSubobject.incFuncRel perX ψ ] ≡ [ F ] + perY≤perψCommutes = + let + answer = + do + (stDF , stDF⊩isStrictDomainF) ← F .isStrictDomain + (relF , relF⊩isRelationalF) ← F .isRelational + let + realizer : ApplStrTerm as 1 + realizer = ` relF ̇ (` stDF ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) + return + (λ* realizer , + (λ y x r r⊩∃x' → + transport + (propTruncIdempotent (F .relation .isPropValued _ _)) + (do + (x' , ⊩Fyx' , ⊩x'~x , ⊩ψx') ← r⊩∃x' + return + (subst + (λ r → r ⊩ ∣ F .relation ∣ (y , x)) + (sym (λ*ComputationRule realizer (r ∷ []))) + (relF⊩isRelationalF y y x' x _ _ _ (stDF⊩isStrictDomainF y x' _ ⊩Fyx') ⊩Fyx' ⊩x'~x))))) + in + eq/ _ _ (answer , F≤G→G≤F perY perX (composeFuncRel _ _ _ perY≤perψFuncRel (InducedSubobject.incFuncRel perX ψ)) F answer) + + opaque + unfolding isInjectiveFuncRel + {-# TERMINATING #-} + perψ≤perYFuncRel : FunctionalRelation perψ perY + Predicate.isSetX (relation perψ≤perYFuncRel) = isSet× (perX .isSetX) (perY .isSetX) + Predicate.∣ relation perψ≤perYFuncRel ∣ (x , y) r = r ⊩ ∣ F .relation ∣ (y , x) + Predicate.isPropValued (relation perψ≤perYFuncRel) (x , y) r = F .relation .isPropValued _ _ + isFunctionalRelation.isStrictDomain (isFuncRel perψ≤perYFuncRel) = + do + (stCF , stCF⊩isStrictCodomainF) ← F .isStrictCodomain + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` stCF ̇ # fzero) ̇ # fzero + return + (λ* realizer , + (λ x y r ⊩Fyx → + (subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (stCF⊩isStrictCodomainF y x _ ⊩Fyx)) , + (return + (y , + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (y , x)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) + ⊩Fyx))))) + isFunctionalRelation.isStrictCodomain (isFuncRel perψ≤perYFuncRel) = + do + (stDF , stDF⊩isStrictDomainF) ← F .isStrictDomain + return + (stDF , + (λ x y r ⊩Fyx → stDF⊩isStrictDomainF y x _ ⊩Fyx)) + isFunctionalRelation.isRelational (isFuncRel perψ≤perYFuncRel) = + do + (relF , relF⊩isRelationalF) ← F .isRelational + let + realizer : ApplStrTerm as 3 + realizer = ` relF ̇ # (fsuc fone) ̇ # fone ̇ (` pr₁ ̇ # fzero) + return + (λ* realizer , + (λ { x x' y y' a b c (⊩x~x' , ⊩ψx) ⊩Fyx ⊩y~y' → + subst (λ r' → r' ⊩ ∣ F .relation ∣ (y' , x')) (sym (λ*ComputationRule realizer (a ∷ b ∷ c ∷ []))) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') })) + isFunctionalRelation.isSingleValued (isFuncRel perψ≤perYFuncRel) = + let + isInjectiveFuncRelF = isMonic→isInjectiveFuncRel perY perX F isMonicF + in + do + (injF , injF⊩isInjectiveF) ← isInjectiveFuncRelF + return + (injF , + (λ x y y' r₁ r₂ ⊩Fyx ⊩Fy'x → + injF⊩isInjectiveF y y' x _ _ ⊩Fyx ⊩Fy'x)) + isFunctionalRelation.isTotal (isFuncRel perψ≤perYFuncRel) = + return + (pr₂ , + (λ { x r (⊩x~x , ⊩ψx) → ⊩ψx })) + + opaque + unfolding composeRTMorphism + unfolding composeFuncRel + unfolding InducedSubobject.incFuncRel + unfolding perψ≤perYFuncRel + perψ≤perYCommutes : [ perψ≤perYFuncRel ] ⋆ [ F ] ≡ [ InducedSubobject.incFuncRel perX ψ ] + perψ≤perYCommutes = + let + answer = + do + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` svF ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero) + return + (λ* realizer , + (λ x x' r r⊩∃y → + transport + (propTruncIdempotent (isProp× (perX .equality .isPropValued _ _) isPropPropTrunc)) + (do + (y , ⊩Fyx , ⊩Fyx') ← r⊩∃y + return + (subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') , + return + (y , + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (y , x)) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) + ⊩Fyx)))))) + in + eq/ _ _ (answer , F≤G→G≤F perψ perX (composeFuncRel _ _ _ perψ≤perYFuncRel F) (InducedSubobject.incFuncRel perX ψ) answer) + +-- For strict relations, subobject inclusion is identified with pointwise entailment +module InclusionEntailment + {X : Type ℓ'} + (perX : PartialEquivalenceRelation X) + (ϕ ψ : StrictRelation perX) where + open StrictRelation + open PredicateProperties X + SubObjX = SubObjCat RT (X , perX) + SubObjHom = Category.Hom[_,_] SubObjX + + perϕ = InducedSubobject.subPer perX ϕ + perψ = InducedSubobject.subPer perX ψ + + incϕ = InducedSubobject.incFuncRel perX ϕ + incψ = InducedSubobject.incFuncRel perX ψ + + ϕsubObject : Category.ob SubObjX + ϕsubObject = sliceob [ InducedSubobject.incFuncRel perX ϕ ] , InducedSubobject.isMonicInc perX ϕ + + ψsubObject : Category.ob SubObjX + ψsubObject = sliceob [ InducedSubobject.incFuncRel perX ψ ] , InducedSubobject.isMonicInc perX ψ + + opaque + unfolding composeRTMorphism + unfolding composeFuncRel + unfolding InducedSubobject.incFuncRel + SubObjHom→ϕ≤ψ : SubObjHom ϕsubObject ψsubObject → (ϕ .predicate ≤ ψ .predicate) + SubObjHom→ϕ≤ψ (slicehom f f⋆incψ≡incϕ) = + SQ.elimProp + {P = λ f → (f ⋆ [ incψ ] ≡ [ incϕ ]) → ϕ .predicate ≤ ψ .predicate} + (λ f → isPropΠ λ f⋆incψ≡incϕ → isProp≤ (ϕ .predicate) (ψ .predicate)) + (λ F F⋆incψ≡incϕ → + let + (p , q) = + SQ.effective + (isPropValuedBientailment (InducedSubobject.subPer perX ϕ) perX) + (isEquivRelBientailment (InducedSubobject.subPer perX ϕ) perX) + (composeFuncRel _ _ _ F incψ) + incϕ + F⋆incψ≡incϕ + in + do + (stϕ , stϕ⊩isStrictϕ) ← ϕ .isStrict + (relψ , relψ⊩isRelationalψ) ← ψ .isRelational + (q , q⊩incϕ≤F⋆incψ) ← q + let + realizer : ApplStrTerm as 1 + realizer = ` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # fzero) ̇ # fzero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # fzero) ̇ # fzero)))) + return + (λ* realizer , + (λ x a a⊩ϕx → + transport + (propTruncIdempotent (ψ .predicate .isPropValued _ _)) + (do + (x' , ⊩Fxx' , ⊩x'~x , ⊩ψx') ← + q⊩incϕ≤F⋆incψ + x x + (pair ⨾ (stϕ ⨾ a) ⨾ a) + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (pr₁pxy≡x _ _)) (stϕ⊩isStrictϕ x a a⊩ϕx)) , + (subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (sym (pr₂pxy≡y _ _)) a⊩ϕx)) + return (subst (λ r' → r' ⊩ ∣ ψ .predicate ∣ x) (sym (λ*ComputationRule realizer (a ∷ []))) (relψ⊩isRelationalψ x' x _ _ ⊩ψx' ⊩x'~x)))))) + f + f⋆incψ≡incϕ + + module _ (ϕ≤ψ : ϕ .predicate ≤ ψ .predicate) where opaque + unfolding idFuncRel + unfolding composeRTMorphism + unfolding composeFuncRel + unfolding InducedSubobject.incFuncRel + + {-# TERMINATING #-} + funcRel : FunctionalRelation perϕ perψ + Predicate.isSetX (relation funcRel) = isSet× (perX .isSetX) (perX .isSetX) + Predicate.∣ relation funcRel ∣ (x , x') r = (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × ((pr₁ ⨾ (pr₂ ⨾ r)) ⊩ ∣ ϕ .predicate ∣ x) × ((pr₂ ⨾ (pr₂ ⨾ r)) ⊩ ∣ ψ .predicate ∣ x) + Predicate.isPropValued (relation funcRel) (x , x') r = isProp× (perX .equality .isPropValued _ _) (isProp× (ϕ .predicate .isPropValued _ _) (ψ .predicate .isPropValued _ _)) + isFunctionalRelation.isStrictDomain (isFuncRel funcRel) = + do + (stϕ , stϕ⊩isStrictϕ) ← ϕ .isStrict + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` stϕ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)) + return + (λ* realizer , + (λ { x x' r (⊩x~x' , ⊩ϕx , ⊩ψx) → + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) (stϕ⊩isStrictϕ x _ ⊩ϕx) , + subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) ⊩ϕx})) + isFunctionalRelation.isStrictCodomain (isFuncRel funcRel) = + do + (stCX , stCX⊩isStrictCodomainX) ← idFuncRel perX .isStrictCodomain + (relψ , relψ⊩isRelationalψ) ← ψ .isRelational + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` stCX ̇ (` pr₁ ̇ # fzero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero)) + return + (λ* realizer , + (λ { x x' r (⊩x~x' , ⊩ϕx , ⊩ψx) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (stCX⊩isStrictCodomainX x x' _ ⊩x~x') , + subst + (λ r' → r' ⊩ ∣ ψ .predicate ∣ x') + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) + (relψ⊩isRelationalψ x x' _ _ ⊩ψx ⊩x~x')})) + isFunctionalRelation.isRelational (isFuncRel funcRel) = + do + (relX , relX⊩isRelationalX) ← idFuncRel perX .isRelational + (relϕ , relϕ⊩isRelationalϕ) ← ϕ .isRelational + (relψ , relψ⊩isRelationalψ) ← ψ .isRelational + let + realizer : ApplStrTerm as 3 + realizer = + ` pair ̇ (` relX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone) ̇ (` pr₁ ̇ # (fsuc fone))) ̇ (` pair ̇ (` relϕ ̇ (` pr₂ ̇ # fzero) ̇ (` pr₁ ̇ # fzero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # fone)) ̇ (` pr₁ ̇ # fzero))) + return + (λ* realizer , + λ { x₁ x₂ x₃ x₄ a b c (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₁~x₃ , ⊩'ϕx₁ , ⊩ψx₁) (⊩x₃~x₄ , ⊩ψx₃) → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₄)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (a ∷ b ∷ c ∷ [])) ∙ pr₁pxy≡x _ _)) + (relX⊩isRelationalX x₁ x₂ x₃ x₄ _ _ _ ⊩x₁~x₂ ⊩x₁~x₃ ⊩x₃~x₄) , + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x₂) + (sym (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule realizer (a ∷ b ∷ c ∷ [])) ∙ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) + (relϕ⊩isRelationalϕ x₁ x₂ _ _ ⊩ϕx₁ ⊩x₁~x₂) , + subst + (λ r' → r' ⊩ ∣ ψ .predicate ∣ x₂) + (sym (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule realizer (a ∷ b ∷ c ∷ [])) ∙ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) + (relψ⊩isRelationalψ x₁ x₂ _ _ ⊩ψx₁ ⊩x₁~x₂)}) + isFunctionalRelation.isSingleValued (isFuncRel funcRel) = + do + (svX , svX⊩isSingleValuedX) ← idFuncRel perX .isSingleValued + (relψ , relψ⊩isRelationalψ) ← ψ .isRelational + let + realizer : ApplStrTerm as 2 + realizer = ` pair ̇ (` svX ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ # fone)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero)) + return + (λ* realizer , + (λ { x₁ x₂ x₃ r₁ r₂ (⊩x₁~x₂ , ⊩ϕx , ⊩ψx) (⊩x₁~x₃ , ⊩'ϕx , ⊩'ψx) → + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x₂ , x₃)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ [])) ∙ pr₁pxy≡x _ _)) (svX⊩isSingleValuedX x₁ x₂ x₃ _ _ ⊩x₁~x₂ ⊩x₁~x₃)) , + subst (λ r' → r' ⊩ ∣ ψ .predicate ∣ x₂) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r₁ ∷ r₂ ∷ [])) ∙ pr₂pxy≡y _ _)) (relψ⊩isRelationalψ x₁ x₂ _ _ ⊩ψx ⊩x₁~x₂)})) + isFunctionalRelation.isTotal (isFuncRel funcRel) = + do + (tl , tl⊩isTotalIncψ) ← incψ .isTotal + (s , s⊩ϕ≤ψ) ← ϕ≤ψ + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ (` pr₂ ̇ # fzero) ̇ (` s ̇ (` pr₂ ̇ # fzero))) + return + (λ* realizer , + (λ { x r (⊩x~x , ⊩ϕx) → + return + (x , + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) + ⊩x~x , + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) + (sym (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule realizer (r ∷ [])) ∙ cong (λ x → pr₁ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₁pxy≡x _ _)) + ⊩ϕx , + subst + (λ r' → r' ⊩ ∣ ψ .predicate ∣ x) + (sym (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x)) (λ*ComputationRule realizer (r ∷ [])) ∙ cong (λ x → pr₂ ⨾ x) (pr₂pxy≡y _ _) ∙ pr₂pxy≡y _ _)) + (s⊩ϕ≤ψ x _ ⊩ϕx))})) + + funcRel⋆incψ≡incϕ : [ funcRel ] ⋆ [ incψ ] ≡ [ incϕ ] + funcRel⋆incψ≡incϕ = + let + answer = + do + (t , t⊩isTransitiveX) ← perX .isTransitive + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))) + return + (λ* realizer , + (λ { x x' r ⊩∃x'' → + transport + (propTruncIdempotent (isPropΣ (perX .equality .isPropValued _ _) λ _ → ϕ .predicate .isPropValued _ _)) + (do + (x'' , (⊩x~x'' , ⊩ϕx , ⊩ψx) , (⊩x''~x' , ⊩'ψx)) ← ⊩∃x'' + return + ((subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₁pxy≡x _ _)) + (t⊩isTransitiveX x x'' x' _ _ ⊩x~x'' ⊩x''~x')) , + (subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) + (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer (r ∷ [])) ∙ pr₂pxy≡y _ _)) + ⊩ϕx)))})) + in + eq/ _ _ (answer , F≤G→G≤F perϕ perX (composeFuncRel _ _ _ funcRel incψ) incϕ answer) + + ϕ≤ψ→SubObjHom : SubObjHom ϕsubObject ψsubObject + ϕ≤ψ→SubObjHom = + slicehom [ funcRel ] funcRel⋆incψ≡incϕ + + SubObjHom≃ϕ≤ψ : SubObjHom ϕsubObject ψsubObject ≃ (ϕ .predicate ≤ ψ .predicate) + SubObjHom≃ϕ≤ψ = + propBiimpl→Equiv + (isPropSubObjMor RT (X , perX) ϕsubObject ψsubObject) + (isProp≤ (ϕ .predicate) (ψ .predicate)) + SubObjHom→ϕ≤ψ + ϕ≤ψ→SubObjHom From a4b867910e4bffccb959c83dcbc4c79c2ba274fc Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 28 Mar 2024 15:54:34 +0530 Subject: [PATCH 35/61] Formulate propositional resizing --- src/Realizability/PropResizing.agda | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) create mode 100644 src/Realizability/PropResizing.agda diff --git a/src/Realizability/PropResizing.agda b/src/Realizability/PropResizing.agda new file mode 100644 index 0000000..67e63fe --- /dev/null +++ b/src/Realizability/PropResizing.agda @@ -0,0 +1,17 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +module Realizability.PropResizing where + +-- Formulation of propositional resizing inspired by the corresponding formulation +-- in TypeTopology +-- https://www.cs.bham.ac.uk/~mhe/TypeTopology/UF.Size.html + +copyOf : ∀ {ℓ} → Type ℓ → (ℓ' : Level) → Type _ +copyOf {ℓ} X ℓ' = Σ[ copy ∈ Type ℓ' ] X ≃ copy + +copy = fst +copyEquiv = snd + +propResizing : ∀ ℓ ℓ' → Type _ +propResizing ℓ ℓ' = ∀ (X : Type ℓ) → isProp X → copyOf X ℓ' From 4f8298442f7ac2540c9800b893f3592c33fbb7d9 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 1 Apr 2024 13:32:30 +0530 Subject: [PATCH 36/61] Define resized predicates and subobject classifier --- src/Realizability/ApplicativeStructure.agda | 2 +- src/Realizability/PropResizing.agda | 10 +- src/Realizability/Topos/PowerObject.agda | 91 ++++++ src/Realizability/Topos/ResizedPredicate.agda | 74 +++++ .../Topos/SubobjectClassifier.agda | 277 ++++++++++++++++++ .../Tripos/Prealgebra/Predicate/Base.agda | 5 + 6 files changed, 456 insertions(+), 3 deletions(-) create mode 100644 src/Realizability/Topos/PowerObject.agda create mode 100644 src/Realizability/Topos/ResizedPredicate.agda create mode 100644 src/Realizability/Topos/SubobjectClassifier.agda diff --git a/src/Realizability/ApplicativeStructure.agda b/src/Realizability/ApplicativeStructure.agda index 9bece44..f78bdfa 100644 --- a/src/Realizability/ApplicativeStructure.agda +++ b/src/Realizability/ApplicativeStructure.agda @@ -104,7 +104,7 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where ... | yes _ = ` s ̇ ` k ̇ ` k ... | no ¬y≡zero with (y .fst) ... | zero = ⊥elim (¬y≡zero refl) - ... | (suc m) = # (m , pred-≤-pred (subst (λ y' → suc y' ≤ suc n) (fromYes fsty≡sucm (discreteℕ (y .fst) (suc m))) (y .snd))) where postulate fsty≡sucm : fst y ≡ suc m + ... | (suc m) = ` k ̇ # (m , pred-≤-pred (subst (λ y' → suc y' ≤ suc n) (fromYes fsty≡sucm (discreteℕ (y .fst) (suc m))) (y .snd))) where postulate fsty≡sucm : fst y ≡ suc m λ*-chainTerm : ∀ n → Term n → Term zero λ*-chainTerm zero t = t diff --git a/src/Realizability/PropResizing.agda b/src/Realizability/PropResizing.agda index 67e63fe..5ba4076 100644 --- a/src/Realizability/PropResizing.agda +++ b/src/Realizability/PropResizing.agda @@ -1,5 +1,8 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure +open import Cubical.Data.Sigma module Realizability.PropResizing where @@ -13,5 +16,8 @@ copyOf {ℓ} X ℓ' = Σ[ copy ∈ Type ℓ' ] X ≃ copy copy = fst copyEquiv = snd -propResizing : ∀ ℓ ℓ' → Type _ -propResizing ℓ ℓ' = ∀ (X : Type ℓ) → isProp X → copyOf X ℓ' +-- We need the principle that TypeTopology calls Ω resizing +-- that the universe of props in a universe 𝓤 has a copy in 𝓤 +-- This we call hPropResizing +hPropResizing : ∀ ℓ → Type _ +hPropResizing ℓ = copyOf (hProp ℓ) ℓ diff --git a/src/Realizability/Topos/PowerObject.agda b/src/Realizability/Topos/PowerObject.agda new file mode 100644 index 0000000..55aa56d --- /dev/null +++ b/src/Realizability/Topos/PowerObject.agda @@ -0,0 +1,91 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*; λ* to `λ*abst) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Fin +open import Cubical.Data.Vec +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Categories.Category +open import Realizability.PropResizing +open import Realizability.CombinatoryAlgebra + +module Realizability.Topos.PowerObject + {ℓ} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + (resizing : hPropResizing ℓ) + where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ} {ℓ'' = ℓ} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ} {ℓ'' = ℓ} ca +open import Realizability.Topos.Object {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.Equalizer {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.BinProducts {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.MonicReprFuncRel {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.ResizedPredicate ca isNonTrivial resizing + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open Morphism +open PartialEquivalenceRelation +open FunctionalRelation +open Category RT + + +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + λ*abst = `λ*abst as fefermanStructure + +-- Power object of X + +module _ (X : Type ℓ) (perX : PartialEquivalenceRelation X) where + 𝓟 : PartialEquivalenceRelation (ResizedPredicate X) + Predicate.isSetX (equality 𝓟) = isSet× isSetResizedPredicate isSetResizedPredicate + Predicate.∣ equality 𝓟 ∣ (ϕ , ψ) r = + (pr₁ ⨾ (pr₁ ⨾ r)) ⊩ isStrictResizedPredicate ϕ × + (pr₂ ⨾ (pr₁ ⨾ r)) ⊩ isRelationalResizedPredicate ϕ × + (pr₁ ⨾ (pr₂ ⨾ r)) ⊩ entailmentResizedPredicate ϕ ψ × + (pr₂ ⨾ (pr₂ ⨾ r)) ⊩ entailmentResizedPredicate ψ ϕ + Predicate.isPropValued (equality 𝓟) (ϕ , ψ) r = + isProp× (isPropIsStrictResizedPredicate ϕ (pr₁ ⨾ (pr₁ ⨾ r))) + (isProp× (isPropIsRelationalResizedPredicate ϕ (pr₂ ⨾ (pr₁ ⨾ r))) + (isProp× + (isPropEntailmentResizedPredicate ϕ ψ (pr₁ ⨾ (pr₂ ⨾ r))) + (isPropEntailmentResizedPredicate ψ ϕ (pr₂ ⨾ (pr₂ ⨾ r))))) + isPartialEquivalenceRelation.isSetX (isPerEquality 𝓟) = isSetResizedPredicate + isPartialEquivalenceRelation.isSymmetric (isPerEquality 𝓟) = + do + let + strictness : ApplStrTerm as 2 + strictness = ` pr₁ ̇ (` pr₁ ̇ # fone) ̇ (` pr₂ ̇ (` pr₂ ̇ # fone) ̇ # fzero) + + proj : ApplStrTerm as 3 + proj = # (fsuc fone) + + proj' : ApplStrTerm as 2 + proj' = ` pr₂ ̇ (` pr₂ ̇ # fzero) ̇ # fone + + proj'' : ApplStrTerm as 2 + proj'' = ` pr₁ ̇ (` pr₂ ̇ # fzero) ̇ # fone + + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` pair ̇ λ*abst strictness ̇ λ*abst (λ*abst proj)) ̇ (` pair ̇ λ*abst proj' ̇ λ*abst proj'') + return + (λ* realizer , + (λ { ϕ ψ r (⊩isStrictϕ , ⊩isRelationalϕ , ⊩ϕ≤ψ , ⊩ψ≤ϕ) → + (λ x b ⊩ψx → + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym {!!}) (⊩isStrictϕ x _ (⊩ψ≤ϕ x b ⊩ψx))) , + (λ x x' b c ⊩x~x' ⊩ψx → + subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x) (sym {!!}) ⊩ψx) , + (λ x a ⊩ψx → + subst (λ r' → r' ⊩ ∣ toPredicate ϕ ∣ x) (sym {!!}) (⊩ψ≤ϕ x _ ⊩ψx)) , + λ x a ⊩ϕx → + subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x) (sym {!!}) (⊩ϕ≤ψ x _ ⊩ϕx) })) + isPartialEquivalenceRelation.isTransitive (isPerEquality 𝓟) = {!!} diff --git a/src/Realizability/Topos/ResizedPredicate.agda b/src/Realizability/Topos/ResizedPredicate.agda new file mode 100644 index 0000000..94634b1 --- /dev/null +++ b/src/Realizability/Topos/ResizedPredicate.agda @@ -0,0 +1,74 @@ +-- Before we can talk about power objects in RT +-- we need to use propositional resizing to get +-- a copy of A-valued predicates in Type ℓ' + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Realizability.PropResizing +open import Realizability.CombinatoryAlgebra + +module Realizability.Topos.ResizedPredicate + {ℓ} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + (resizing : hPropResizing ℓ) + where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ} {ℓ'' = ℓ} ca + +open CombinatoryAlgebra ca +open Predicate renaming (isSetX to isSetPredicateBase) + +smallHProp = resizing .fst +smallHProp≃hProp = resizing .snd + +ResizedPredicate : Type ℓ → Type ℓ +ResizedPredicate X = Σ[ rel ∈ (X → A → smallHProp) ] isSet X + +PredicateΣ≃ResizedPredicate : ∀ X → PredicateΣ X ≃ ResizedPredicate X +PredicateΣ≃ResizedPredicate X = + Σ-cong-equiv-prop + (equivΠ + (idEquiv X) + (λ x → + equivΠ + (idEquiv A) + λ a → + smallHProp≃hProp)) + (λ _ → isPropIsSet) + (λ _ → isPropIsSet) + (λ _ answer → answer) + (λ _ answer → answer) + +Predicate≃ResizedPredicate : ∀ X → Predicate X ≃ ResizedPredicate X +Predicate≃ResizedPredicate X = compEquiv (Predicate≃PredicateΣ X) (PredicateΣ≃ResizedPredicate X) + +isSetResizedPredicate : ∀ {X} → isSet (ResizedPredicate X) +isSetResizedPredicate {X} = isOfHLevelRespectEquiv 2 (Predicate≃ResizedPredicate X) (isSetPredicate X) + +ResizedPredicate≃Predicate : ∀ X → ResizedPredicate X ≃ Predicate X +ResizedPredicate≃Predicate X = invEquiv (Predicate≃ResizedPredicate X) + +toPredicate : ∀ {X} → ResizedPredicate X → Predicate X +toPredicate {X} ϕ = equivFun (ResizedPredicate≃Predicate X) ϕ + +fromPredicate : ∀ {X} → Predicate X → ResizedPredicate X +fromPredicate {X} ϕ = equivFun (Predicate≃ResizedPredicate X) ϕ + +compIsIdEquiv : ∀ X → compEquiv (Predicate≃ResizedPredicate X) (ResizedPredicate≃Predicate X) ≡ idEquiv (Predicate X) +compIsIdEquiv X = invEquiv-is-rinv (Predicate≃ResizedPredicate X) + +compIsIdFunc : ∀ {X} → (p : Predicate X) → toPredicate (fromPredicate p) ≡ p +compIsIdFunc {X} p i = equivFun (compIsIdEquiv X i) p + +module _ {X} where + entailmentResizedPredicate : ∀ (ϕ ψ : ResizedPredicate X) → A → Type ℓ + entailmentResizedPredicate ϕ ψ r = ∀ (x : X) (a : A) (⊩ϕx : a ⊩ ∣ toPredicate ϕ ∣ x) → (r ⨾ a) ⊩ ∣ toPredicate ψ ∣ x + + isPropEntailmentResizedPredicate : ∀ ϕ ψ a → isProp (entailmentResizedPredicate ϕ ψ a) + isPropEntailmentResizedPredicate ϕ ψ a = + isPropΠ λ x → isPropΠ λ b → isPropΠ λ _ → (toPredicate ψ) .isPropValued _ _ diff --git a/src/Realizability/Topos/SubobjectClassifier.agda b/src/Realizability/Topos/SubobjectClassifier.agda new file mode 100644 index 0000000..1334397 --- /dev/null +++ b/src/Realizability/Topos/SubobjectClassifier.agda @@ -0,0 +1,277 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*; λ* to `λ*abst) +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Data.Empty +open import Cubical.Data.Sigma +open import Cubical.Data.Fin +open import Cubical.Data.Vec +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.Pullback +open import Realizability.PropResizing +open import Realizability.CombinatoryAlgebra + +module Realizability.Topos.SubobjectClassifier + {ℓ} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + (resizing : hPropResizing ℓ) + where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ} {ℓ'' = ℓ} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ} {ℓ'' = ℓ} ca +open import Realizability.Topos.Object {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.Equalizer {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.BinProducts {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.MonicReprFuncRel {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.ResizedPredicate ca isNonTrivial resizing +open import Realizability.Topos.TerminalObject {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.StrictRelation {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open Morphism +open PartialEquivalenceRelation +open FunctionalRelation +open Category RT + +private + λ*ComputationRule = `λ*ComputationRule as fefermanStructure + λ* = `λ* as fefermanStructure + λ*abst = `λ*abst as fefermanStructure + +Ωper : PartialEquivalenceRelation (ResizedPredicate Unit*) +Predicate.isSetX (equality Ωper) = isSet× isSetResizedPredicate isSetResizedPredicate +Predicate.∣ equality Ωper ∣ (α , β) r = + (∀ (a : A) (⊩α : a ⊩ ∣ toPredicate α ∣ tt*) → ((pr₁ ⨾ r) ⨾ a) ⊩ ∣ toPredicate β ∣ tt*) × + (∀ (a : A) (⊩β : a ⊩ ∣ toPredicate β ∣ tt*) → ((pr₂ ⨾ r) ⨾ a) ⊩ ∣ toPredicate α ∣ tt*) +Predicate.isPropValued (equality Ωper) (α , β) r = + isProp× + (isPropΠ (λ _ → isPropΠ λ _ → (toPredicate β) .isPropValued _ _)) + (isPropΠ (λ _ → isPropΠ λ _ → (toPredicate α) .isPropValued _ _)) +isPartialEquivalenceRelation.isSetX (isPerEquality Ωper) = isSetResizedPredicate +isPartialEquivalenceRelation.isSymmetric (isPerEquality Ωper) = + do + let + ent₁ : ApplStrTerm as 2 + ent₁ = ` pr₂ ̇ # fone ̇ # fzero + + ent₂ : ApplStrTerm as 2 + ent₂ = ` pr₁ ̇ # fone ̇ # fzero + + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (λ*abst ent₁) ̇ (λ*abst ent₂) + return + (λ* realizer , + λ { α β r (pr₁r⊩α≤β , pr₂r⊩β≤α) → + (λ a a⊩β → + let + eq : pr₁ ⨾ (λ* realizer ⨾ r) ⨾ a ≡ pr₂ ⨾ r ⨾ a + eq = + pr₁ ⨾ (λ* realizer ⨾ r) ⨾ a + ≡⟨ cong (λ x → pr₁ ⨾ x ⨾ a) (λ*ComputationRule realizer (r ∷ [])) ⟩ + pr₁ ⨾ (pair ⨾ (s ⨾ (s ⨾ (k ⨾ pr₂) ⨾ (k ⨾ r)) ⨾ (s ⨾ k ⨾ k)) ⨾ (s ⨾ (s ⨾ (k ⨾ pr₁) ⨾ (k ⨾ r)) ⨾ (s ⨾ k ⨾ k))) ⨾ a + ≡⟨ cong (λ x → x ⨾ a) (pr₁pxy≡x _ _) ⟩ + s ⨾ (s ⨾ (k ⨾ pr₂) ⨾ (k ⨾ r)) ⨾ (s ⨾ k ⨾ k) ⨾ a + ≡⟨ sabc≡ac_bc _ _ _ ⟩ + s ⨾ (k ⨾ pr₂) ⨾ (k ⨾ r) ⨾ a ⨾ (s ⨾ k ⨾ k ⨾ a) + ≡⟨ cong (λ x → x ⨾ (s ⨾ k ⨾ k ⨾ a)) (sabc≡ac_bc _ _ _) ⟩ + k ⨾ pr₂ ⨾ a ⨾ (k ⨾ r ⨾ a) ⨾ (s ⨾ k ⨾ k ⨾ a) + ≡⟨ cong (λ x → x ⨾ (k ⨾ r ⨾ a) ⨾ (s ⨾ k ⨾ k ⨾ a)) (kab≡a _ _) ⟩ + pr₂ ⨾ (k ⨾ r ⨾ a) ⨾ (s ⨾ k ⨾ k ⨾ a) + ≡⟨ cong (λ x → pr₂ ⨾ x ⨾ (s ⨾ k ⨾ k ⨾ a)) (kab≡a _ _) ⟩ + pr₂ ⨾ r ⨾ (s ⨾ k ⨾ k ⨾ a) + ≡⟨ cong (λ x → pr₂ ⨾ r ⨾ x) (Ida≡a _) ⟩ + pr₂ ⨾ r ⨾ a + ∎ + in + subst (λ r' → r' ⊩ ∣ toPredicate α ∣ tt*) (sym eq) (pr₂r⊩β≤α a a⊩β)) , + (λ a a⊩α → subst (λ r' → r' ⊩ ∣ toPredicate β ∣ tt*) (sym {!!}) (pr₁r⊩α≤β a a⊩α)) }) +isPartialEquivalenceRelation.isTransitive (isPerEquality Ωper) = + do + return + ({!!} , + (λ { x y z a b (⊩x≤y , ⊩y≤x) (⊩y≤z , ⊩z≤y) → + (λ r r⊩x → subst (λ r' → r' ⊩ ∣ toPredicate z ∣ tt*) {!!} (⊩y≤z _ (⊩x≤y r r⊩x))) , + (λ r r⊩z → subst (λ r' → r' ⊩ ∣ toPredicate x ∣ tt*) {!!} (⊩y≤x _ (⊩z≤y r r⊩z))) })) + +opaque + unfolding terminalPer + trueFuncRel : FunctionalRelation terminalPer Ωper + Predicate.isSetX (relation trueFuncRel) = isSet× isSetUnit* isSetResizedPredicate + Predicate.∣ relation trueFuncRel ∣ (tt* , p) r = ∀ (a : A) → (r ⨾ a) ⊩ ∣ toPredicate p ∣ tt* + Predicate.isPropValued (relation trueFuncRel) (tt* , p) r = isPropΠ λ a → (toPredicate p) .isPropValued _ _ + isFunctionalRelation.isStrictDomain (isFuncRel trueFuncRel) = + do + return + (k , + (λ { tt* y r r⊩⊤≤y → tt*})) + isFunctionalRelation.isStrictCodomain (isFuncRel trueFuncRel) = + do + return + ({!!} , + (λ { tt* y r r⊩⊤≤y → + (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y) , + (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y)})) + isFunctionalRelation.isRelational (isFuncRel trueFuncRel) = + do + return + ({!!} , + (λ { tt* tt* x y a b c tt* b⊩⊤≤x (pr₁c⊩x≤y , pr₂c⊩y≤x) r → + subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} (pr₁c⊩x≤y (b ⨾ k) (b⊩⊤≤x k))})) + isFunctionalRelation.isSingleValued (isFuncRel trueFuncRel) = + do + return + ({!!} , + (λ { tt* x y r₁ r₂ r₁⊩⊤≤x r₂⊩⊤≤y → + (λ a a⊩x → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} (r₂⊩⊤≤y k)) , + (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate x ∣ tt*) {!!} (r₁⊩⊤≤x k))})) + isFunctionalRelation.isTotal (isFuncRel trueFuncRel) = + do + return + (k , + (λ { tt* r tt* → + let + ⊤ = pre1 Unit* isSetUnit* isNonTrivial + in + ∣ + fromPredicate ⊤ , + (λ a → + subst (λ p → (k ⨾ r ⨾ a) ⊩ ∣ p ∣ tt*) (sym (compIsIdFunc ⊤)) tt*) + ∣₁ })) + +opaque + unfolding isInjectiveFuncRel + unfolding terminalPer + isInjectiveTrueFuncRel : isInjectiveFuncRel _ _ trueFuncRel + isInjectiveTrueFuncRel = + do + return + (k , + (λ { tt* tt* p r₁ r₂ r₁⊩⊤≤p r₂⊩⊤≤p → tt* })) + +-- The subobject classifier indeed classifies subobjects +module _ + (X : Type ℓ) + (perX : PartialEquivalenceRelation X) + (ϕ : StrictRelation perX) where + + open InducedSubobject perX ϕ + open StrictRelation + resizedϕ = fromPredicate (ϕ .predicate) + + -- the functional relation that represents the unique indicator map + theFuncRel : FunctionalRelation perX Ωper + Predicate.isSetX (relation theFuncRel) = isSet× (perX .isSetX) isSetResizedPredicate + Predicate.∣ relation theFuncRel ∣ (x , p) r = + (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x) × + (∀ (b : A) (b⊩ϕx : b ⊩ ∣ ϕ .predicate ∣ x) → (pr₁ ⨾ (pr₂ ⨾ r) ⨾ b) ⊩ ∣ toPredicate p ∣ tt*) × + (∀ (b : A) (b⊩px : b ⊩ ∣ toPredicate p ∣ tt*) → (pr₂ ⨾ (pr₂ ⨾ r) ⨾ b) ⊩ ∣ ϕ .predicate ∣ x) + Predicate.isPropValued (relation theFuncRel) (x , p) r = + isProp× + (perX .equality .isPropValued _ _) + (isProp× + (isPropΠ (λ _ → isPropΠ λ _ → (toPredicate p) .isPropValued _ _)) + (isPropΠ λ _ → isPropΠ λ _ → ϕ .predicate .isPropValued _ _)) + isFunctionalRelation.isStrictDomain (isFuncRel theFuncRel) = + do + return + (pr₁ , + (λ { x p r (pr₁r⊩x~x , ⊩ϕx≤p , ⊩p≤ϕx) → pr₁r⊩x~x})) + isFunctionalRelation.isStrictCodomain (isFuncRel theFuncRel) = + do + return + ({!!} , + (λ x y r x₁ → + (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y) , + (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y))) + isFunctionalRelation.isRelational (isFuncRel theFuncRel) = + do + return + ({!!} , + (λ { x x' p p' a b c a⊩x~x' (pr₁b⊩x~x , pr₁pr₂b⊩ϕx≤y , pr₂pr₂b⊩y≤ϕx) (pr₁c⊩y≤y' , pr₂c⊩y'≤y) → + {!!} , + (λ r r⊩ϕx' → {!!}) , + λ r r⊩p' → {!!} })) + isFunctionalRelation.isSingleValued (isFuncRel theFuncRel) = + do + return + ({!!} , + (λ { x y y' r₁ r₂ (pr₁r₁⊩x~x , pr₁pr₂r₁⊩ϕx≤y , pr₂pr₂r₁⊩y≤ϕx) (pr₂r₂⊩x~x , pr₁pr₂r₂⊩ϕx≤y' , pr₂pr₂r₂⊩y'≤ϕx) → + {!!} })) + isFunctionalRelation.isTotal (isFuncRel theFuncRel) = + do + return + ({!!} , + (λ x r r⊩x~x → + let + resultPredicate : Predicate Unit* + resultPredicate = + consPredicate + isSetUnit* + (λ { tt* s → s ⊩ ∣ ϕ .predicate ∣ x }) + (λ { tt* s → ϕ .predicate .isPropValued _ _ }) + in + return + (fromPredicate resultPredicate , + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym {!!}) r⊩x~x , + (λ b b⊩ϕx → {!compIsIdFunc!}) , + λ b b⊩'ϕx → {!!}))) + + + subobjectCospan : Cospan RT + Cospan.l subobjectCospan = X , perX + Cospan.m subobjectCospan = ResizedPredicate Unit* , Ωper + Cospan.r subobjectCospan = Unit* , terminalPer + Cospan.s₁ subobjectCospan = [ theFuncRel ] + Cospan.s₂ subobjectCospan = [ trueFuncRel ] + + opaque + unfolding composeRTMorphism + unfolding composeFuncRel + unfolding terminalFuncRel + unfolding trueFuncRel + unfolding incFuncRel + subobjectSquareCommutes : [ incFuncRel ] ⋆ [ theFuncRel ] ≡ [ terminalFuncRel subPer ] ⋆ [ trueFuncRel ] + subobjectSquareCommutes = + let + answer = + do + (stX , stX⊩isStrictDomainX) ← idFuncRel perX .isStrictDomain + (relϕ , relϕ⊩isRelationalϕ) ← StrictRelation.isRelational ϕ + let + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` pair ̇ (` stX ̇ (` pr₁ ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero))) ̇ λ*abst (` pr₁ ̇ (` pr₂ ̇ (` pr₂ ̇ # fone)) ̇ (` relϕ ̇ (` pr₂ ̇ (` pr₁ ̇ # fone)) ̇ (` pr₁ ̇ (` pr₁ ̇ # fone)))) + return + (λ* realizer , + (λ { x p r r⊩∃x' → + do + (x' , (⊩x~x' , ⊩ϕx) , ⊩x'~x' , ⊩ϕx'≤p , ⊩p≤ϕx') ← r⊩∃x' + return + (tt* , + ((subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (λ*ComputationRule realizer (r ∷ [])) ∙ cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₁pxy≡x _ _)) (stX⊩isStrictDomainX x x' _ ⊩x~x')) , + (subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (sym (cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (λ*ComputationRule realizer (r ∷ [])) ∙ cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₂pxy≡y _ _)) ⊩ϕx)) , + λ r' → + subst (λ r' → r' ⊩ ∣ toPredicate p ∣ tt*) (sym (cong (λ x → pr₂ ⨾ x ⨾ r') (λ*ComputationRule realizer (r ∷ [])) ∙ cong (λ x → x ⨾ r') (pr₂pxy≡y _ _) ∙ sabc≡ac_bc _ _ _ ∙ {!!})) (⊩ϕx'≤p _ (relϕ⊩isRelationalϕ x x' _ _ ⊩ϕx ⊩x~x'))) })) + in + eq/ _ _ (answer , {!!}) + + classifies : isPullback RT subobjectCospan [ incFuncRel ] [ terminalFuncRel subPer ] subobjectSquareCommutes + classifies {Y , perY} f g f⋆char≡g⋆true = + SQ.elimProp2 + {P = λ f g → ∀ (commutes : f ⋆ [ theFuncRel ] ≡ g ⋆ [ trueFuncRel ]) → ∃![ hk ∈ RTMorphism perY subPer ] (f ≡ hk ⋆ [ incFuncRel ]) × (g ≡ hk ⋆ [ terminalFuncRel subPer ])} + (λ f g → isPropΠ λ _ → isPropIsContr) + (λ F G F⋆char≡G⋆true → + uniqueExists + {!!} + ({!!} , {!!}) + (λ hk' → isProp× (squash/ _ _) (squash/ _ _)) + {!!}) + f g f⋆char≡g⋆true diff --git a/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda b/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda index 75790d8..847dae5 100644 --- a/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda +++ b/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda @@ -4,6 +4,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma open import Cubical.Functions.FunExtEquiv @@ -13,6 +14,7 @@ open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) record Predicate (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + constructor consPredicate field isSetX : isSet X ∣_∣ : X → A → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') @@ -38,6 +40,9 @@ PredicateIsoΣ X = (λ b → refl) λ a → refl +Predicate≃PredicateΣ : ∀ (X : Type ℓ') → Predicate X ≃ PredicateΣ X +Predicate≃PredicateΣ X = isoToEquiv (PredicateIsoΣ X) + Predicate≡PredicateΣ : ∀ (X : Type ℓ') → Predicate X ≡ PredicateΣ X Predicate≡PredicateΣ X = isoToPath (PredicateIsoΣ X) From 796a4d46661c1dd8ba4692f55804d5bf23008ad1 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 2 Apr 2024 21:07:52 +0530 Subject: [PATCH 37/61] Complete refactor --- src/Realizability/ApplicativeStructure.agda | 222 +++++++++--------- src/Realizability/CombinatoryAlgebra.agda | 41 ++-- src/Realizability/Topos/BinProducts.agda | 146 ++++++------ src/Realizability/Topos/Equalizer.agda | 118 +++++----- src/Realizability/Topos/Everything.agda | 7 + .../Topos/FunctionalRelation.agda | 93 ++++---- src/Realizability/Topos/MonicReprFuncRel.agda | 23 +- src/Realizability/Topos/StrictRelation.agda | 147 ++++++------ .../Topos/SubobjectClassifier.agda | 87 +++---- src/Realizability/Topos/TerminalObject.agda | 18 +- .../Prealgebra/Meets/Commutativity.agda | 19 +- .../Tripos/Prealgebra/Meets/Identity.agda | 15 +- .../Tripos/Prealgebra/Predicate/Base.agda | 2 +- .../Prealgebra/Predicate/Properties.agda | 8 +- src/index.agda | 5 - 15 files changed, 467 insertions(+), 484 deletions(-) diff --git a/src/Realizability/ApplicativeStructure.agda b/src/Realizability/ApplicativeStructure.agda index f78bdfa..afaa19c 100644 --- a/src/Realizability/ApplicativeStructure.agda +++ b/src/Realizability/ApplicativeStructure.agda @@ -1,16 +1,34 @@ -{-# OPTIONS --cubical --without-K --allow-unsolved-metas #-} open import Cubical.Core.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Relation.Nullary open import Cubical.Data.Nat open import Cubical.Data.Nat.Order -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Empty renaming (elim to ⊥elim) +open import Cubical.Tactics.NatSolver module Realizability.ApplicativeStructure where +private module _ {ℓ} {A : Type ℓ} where + -- Taken from Data.Vec.Base from agda-stdlib + foldlOp : ∀ {ℓ'} (B : ℕ → Type ℓ') → Type _ + foldlOp B = ∀ {n} → B n → A → B (suc n) + + opaque + foldl : ∀ {ℓ'} {n : ℕ} (B : ℕ → Type ℓ') → foldlOp B → B zero → Vec A n → B n + foldl {ℓ'} {.zero} B op acc emptyVec = acc + foldl {ℓ'} {.(suc _)} B op acc (x ∷vec vec) = foldl (λ n → B (suc n)) op (op acc x) vec + + opaque + reverse : ∀ {n} → Vec A n → Vec A n + reverse vec = foldl (λ n → Vec A n) (λ acc curr → curr ∷ acc) [] vec + + opaque + chain : ∀ {n} → Vec A (suc n) → (A → A → A) → A + chain {n} (x ∷vec vec) op = foldl (λ _ → A) (λ acc curr → op acc curr) x vec + record ApplicativeStructure {ℓ} (A : Type ℓ) : Type ℓ where infixl 20 _⨾_ field @@ -26,33 +44,24 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where `_ : A → Term n _̇_ : Term n → Term n → Term n - upgrade : ∀ {n m} → n < m → Term n → Term m - upgrade _ (` a) = ` a - upgrade {n} {m} n Date: Tue, 9 Apr 2024 04:34:21 +0530 Subject: [PATCH 38/61] Archival commit --- .../Topos/FunctionalRelation.agda | 2 + src/Realizability/Topos/PowerObject.agda | 146 +++++++++++++++--- src/Realizability/Topos/ResizedPredicate.agda | 20 ++- .../Topos/SubobjectClassifier.agda | 72 ++++++++- 4 files changed, 210 insertions(+), 30 deletions(-) diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index 9c8c5c9..aed4695 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -71,6 +71,7 @@ module _ (∀ x r → r ⊩ ∣ equalityX ∣ (x , x) → ∃[ y ∈ Y ] (tl ⨾ r) ⊩ ∣ relation ∣ (x , y)) record isFunctionalRelation : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + constructor makeIsFunctionalRelation field isStrictDomain : ∃[ stD ∈ A ] (realizesStrictDomain stD) isStrictCodomain : ∃[ stC ∈ A ] (realizesStrictCodomain stC) @@ -79,6 +80,7 @@ module _ isTotal : ∃[ tl ∈ A ] (realizesTotal tl) record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where + constructor makeFunctionalRelation field relation : Predicate (X × Y) isFuncRel : isFunctionalRelation perX perY relation diff --git a/src/Realizability/Topos/PowerObject.agda b/src/Realizability/Topos/PowerObject.agda index 55aa56d..228682c 100644 --- a/src/Realizability/Topos/PowerObject.agda +++ b/src/Realizability/Topos/PowerObject.agda @@ -1,10 +1,10 @@ -open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*; λ* to `λ*abst) +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv open import Cubical.Data.Empty open import Cubical.Data.Sigma -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad @@ -27,6 +27,7 @@ open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ} {ℓ'' = ℓ} ca open import Realizability.Topos.Equalizer {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial open import Realizability.Topos.BinProducts {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial open import Realizability.Topos.MonicReprFuncRel {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.StrictRelation {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial open import Realizability.Topos.ResizedPredicate ca isNonTrivial resizing open CombinatoryAlgebra ca @@ -36,16 +37,13 @@ open Morphism open PartialEquivalenceRelation open FunctionalRelation open Category RT - - -private - λ*ComputationRule = `λ*ComputationRule as fefermanStructure - λ* = `λ* as fefermanStructure - λ*abst = `λ*abst as fefermanStructure +open StrictRelation -- Power object of X module _ (X : Type ℓ) (perX : PartialEquivalenceRelation X) where + open ResizedPredicateProps perX + 𝓟 : PartialEquivalenceRelation (ResizedPredicate X) Predicate.isSetX (equality 𝓟) = isSet× isSetResizedPredicate isSetResizedPredicate Predicate.∣ equality 𝓟 ∣ (ϕ , ψ) r = @@ -64,28 +62,134 @@ module _ (X : Type ℓ) (perX : PartialEquivalenceRelation X) where do let strictness : ApplStrTerm as 2 - strictness = ` pr₁ ̇ (` pr₁ ̇ # fone) ̇ (` pr₂ ̇ (` pr₂ ̇ # fone) ̇ # fzero) + strictness = ` pr₁ ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero) ̇ # one) proj : ApplStrTerm as 3 - proj = # (fsuc fone) + proj = # two proj' : ApplStrTerm as 2 - proj' = ` pr₂ ̇ (` pr₂ ̇ # fzero) ̇ # fone + proj' = ` pr₂ ̇ (` pr₂ ̇ # zero) ̇ # one proj'' : ApplStrTerm as 2 - proj'' = ` pr₁ ̇ (` pr₂ ̇ # fzero) ̇ # fone + proj'' = ` pr₁ ̇ (` pr₂ ̇ # zero) ̇ # one realizer : ApplStrTerm as 1 realizer = ` pair ̇ (` pair ̇ λ*abst strictness ̇ λ*abst (λ*abst proj)) ̇ (` pair ̇ λ*abst proj' ̇ λ*abst proj'') return (λ* realizer , - (λ { ϕ ψ r (⊩isStrictϕ , ⊩isRelationalϕ , ⊩ϕ≤ψ , ⊩ψ≤ϕ) → - (λ x b ⊩ψx → - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym {!!}) (⊩isStrictϕ x _ (⊩ψ≤ϕ x b ⊩ψx))) , - (λ x x' b c ⊩x~x' ⊩ψx → - subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x) (sym {!!}) ⊩ψx) , + (λ { ϕ ψ r (⊩isStrictϕ , isRelationalϕ , ⊩ϕ≤ψ , ⊩ψ≤ϕ) → (λ x a ⊩ψx → - subst (λ r' → r' ⊩ ∣ toPredicate ϕ ∣ x) (sym {!!}) (⊩ψ≤ϕ x _ ⊩ψx)) , - λ x a ⊩ϕx → - subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x) (sym {!!}) (⊩ϕ≤ψ x _ ⊩ϕx) })) - isPartialEquivalenceRelation.isTransitive (isPerEquality 𝓟) = {!!} + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym {!!}) (⊩isStrictϕ x _ (⊩ψ≤ϕ x a ⊩ψx))) , + (λ x x' a b ⊩x~x' ⊩ψx → + subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x') (sym {!!}) (⊩ϕ≤ψ x' _ (isRelationalϕ x x' _ _ ⊩x~x' (⊩ψ≤ϕ x _ ⊩ψx)))) , + (λ x a ⊩ψx → subst (λ r' → r' ⊩ ∣ toPredicate ϕ ∣ x) (sym {!!}) (⊩ψ≤ϕ x a ⊩ψx)) , + (λ x a ⊩ϕx → subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x) (sym {!!}) (⊩ϕ≤ψ x a ⊩ϕx)) })) + isPartialEquivalenceRelation.isTransitive (isPerEquality 𝓟) = + do + return + ({!!} , + (λ { + ϕ ψ θ a b + (⊩isStrictϕ , ⊩isRelationalϕ , ⊩ϕ≤ψ , ⊩ψ≤ϕ) + (⊩isStrictψ , ⊩isRelationalψ , ⊩ψ≤θ , ⊩θ≤ψ) → + subst (λ r' → r' ⊩ isStrictResizedPredicate ϕ) {!!} ⊩isStrictϕ , + subst (λ r' → r' ⊩ isRelationalResizedPredicate ϕ) {!!} ⊩isRelationalϕ , + (λ x r ⊩ϕx → subst (λ r' → r' ⊩ ∣ toPredicate θ ∣ x) {!!} (⊩ψ≤θ x _ (⊩ϕ≤ψ x r ⊩ϕx))) , + (λ x r ⊩θx → subst (λ r' → r' ⊩ ∣ toPredicate ϕ ∣ x) {!!} (⊩ψ≤ϕ x _ (⊩θ≤ψ x r ⊩θx))) })) + + opaque + unfolding binProdObRT + ∈StrictRel : StrictRelation (binProdObRT perX 𝓟) + Predicate.isSetX (StrictRelation.predicate ∈StrictRel) = isSet× (perX .isSetX) isSetResizedPredicate + Predicate.∣ StrictRelation.predicate ∈StrictRel ∣ (x , ϕ) r = (pr₁ ⨾ r) ⊩ ∣ 𝓟 .equality ∣ (ϕ , ϕ) × (pr₂ ⨾ r) ⊩ ∣ toPredicate ϕ ∣ x + Predicate.isPropValued (StrictRelation.predicate ∈StrictRel) (x , ϕ) r = + isProp× (𝓟 .equality .isPropValued (ϕ , ϕ) (pr₁ ⨾ r)) (toPredicate ϕ .isPropValued _ _) + isStrictRelation.isStrict (StrictRelation.isStrictRelationPredicate ∈StrictRel) = + do + return + ({!!} , + λ { (x , ϕ) r ((⊩isStrictϕ , ⊩isRelationalϕ , ⊩ϕ≤ϕ , ⊩'ϕ≤ϕ) , ⊩ϕx) → + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) {!!} (⊩isStrictϕ x _ ⊩ϕx) , + subst (λ r' → r' ⊩ isStrictResizedPredicate ϕ) {!!} ⊩isStrictϕ , + subst (λ r' → r' ⊩ isRelationalResizedPredicate ϕ) {!!} ⊩isRelationalϕ , + subst (λ r' → r' ⊩ entailmentResizedPredicate ϕ ϕ) {!!} ⊩ϕ≤ϕ , + subst (λ r' → r' ⊩ entailmentResizedPredicate ϕ ϕ) {!!} ⊩'ϕ≤ϕ }) + isStrictRelation.isRelational (StrictRelation.isStrictRelationPredicate ∈StrictRel) = + do + return + ({!!} , + λ { (x , ϕ) (x' , ψ) a b ((⊩isStrictϕ , ⊩isRelationalϕ , ⊩ϕ≤ϕ , ⊩'ϕ≤ϕ) , ⊩ϕx) (⊩x~x' , (⊩isStrict'ϕ , ⊩isRelational'ϕ , ⊩ϕ≤ψ , ⊩ψ≤ϕ)) → + ((λ x'' r ⊩ψx'' → + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x'' , x'')) {!!} (⊩isStrictϕ x'' _ (⊩ψ≤ϕ x'' r ⊩ψx''))) , + (λ x'' x''' p q ⊩x''~x''' ⊩ψx'' → + subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x''') (sym {!!}) (⊩ϕ≤ψ x''' _ (⊩isRelationalϕ x'' x''' _ _ ⊩x''~x''' (⊩ψ≤ϕ x'' _ ⊩ψx'')))) , + (λ x'' r ⊩ψx'' → subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x'') {!!} ⊩ψx'') , + λ x'' r ⊩ψx'' → subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x'') {!!} ⊩ψx'') , + subst (λ r' → r' ⊩ ∣ toPredicate ψ ∣ x') {!!} (⊩ϕ≤ψ x' _ (⊩isRelationalϕ x x' _ _ ⊩x~x' ⊩ϕx)) }) + + open InducedSubobject (binProdObRT perX 𝓟) ∈StrictRel + renaming + ( subPer to ∈subPer + ; incFuncRel to ∈incFuncRel + ; isInjectiveIncFuncRel to isInjective∈IncFuncRel + ; isMonicInc to isMonicInc∈) + + {- + We need to show that the following is a pullback diagram and that [ F ] is the *unique* RT arrow that + makes this diagram a pullback. + + [ top ] + (X × Y)ϕ - - - - - - - - - -> ∈ₓ + Γ _| Γ + | | + | | + | | + | [ incϕ ] | [ inc∈ ] + | | + | | + | | + (X × Y) - - - - - - - - - -> (X × 𝓟 X) + idₓ × [ F ] + -} + module PowerObjectUnivProp + {Y : Type ℓ} + (perY : PartialEquivalenceRelation Y) + (ϕ : StrictRelation (binProdObRT perX perY)) where + + open InducedSubobject (binProdObRT perX perY) ϕ + renaming + ( subPer to ϕsubPer + ; incFuncRel to ϕincFuncRel + ; isInjectiveIncFuncRel to isInjectiveϕIncFuncRel + ; isMonicInc to isMonicIncϕ) + + opaque + unfolding binProdObRT + unfolding idFuncRel + {-# TERMINATING #-} + topArrowFuncRel : FunctionalRelation ϕsubPer ∈subPer + Predicate.isSetX (relation topArrowFuncRel) = isSet× (isSet× (perX .isSetX) (perY .isSetX)) (isSet× (perX .isSetX) isSetResizedPredicate) + Predicate.∣ relation topArrowFuncRel ∣ ((x , y) , (x' , p)) r = + (pr₁ ⨾ r) ⊩ ∣ perX .equality ∣ (x , x') × + (pr₁ ⨾ (pr₂ ⨾ r)) ⊩ ∣ toPredicate p ∣ x × + (pr₂ ⨾ (pr₂ ⨾ r)) ⊩ ∣ ϕ .predicate ∣ (x , y) + Predicate.isPropValued (relation topArrowFuncRel) ((x , y) , (x' , p)) r = + isProp× + (perX .equality .isPropValued _ _) + (isProp× + (toPredicate p .isPropValued _ _) + (ϕ .predicate .isPropValued _ _)) + isFunctionalRelation.isStrictDomain (isFuncRel topArrowFuncRel) = + do + (stX , stX⊩isStrictDomainEqX) ← idFuncRel perX .isStrictDomain + (stY , stY⊩isStrictDomainEqY) ← idFuncRel perY .isStrictDomain + return + ({!!} , + (λ { (x , y) (x' , p) r (⊩x~x' , ⊩ϕxy , ⊩px) → + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) {!!} (stX⊩isStrictDomainEqX x x' _ ⊩x~x') , + {!subst !}) , + {!!}})) + isFunctionalRelation.isStrictCodomain (isFuncRel topArrowFuncRel) = {!!} + isFunctionalRelation.isRelational (isFuncRel topArrowFuncRel) = {!!} + isFunctionalRelation.isSingleValued (isFuncRel topArrowFuncRel) = {!!} + isFunctionalRelation.isTotal (isFuncRel topArrowFuncRel) = {!!} diff --git a/src/Realizability/Topos/ResizedPredicate.agda b/src/Realizability/Topos/ResizedPredicate.agda index 94634b1..8eb2d7b 100644 --- a/src/Realizability/Topos/ResizedPredicate.agda +++ b/src/Realizability/Topos/ResizedPredicate.agda @@ -19,6 +19,7 @@ module Realizability.Topos.ResizedPredicate where open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ} {ℓ'' = ℓ} ca +open import Realizability.Topos.Object {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial open CombinatoryAlgebra ca open Predicate renaming (isSetX to isSetPredicateBase) @@ -65,10 +66,27 @@ compIsIdEquiv X = invEquiv-is-rinv (Predicate≃ResizedPredicate X) compIsIdFunc : ∀ {X} → (p : Predicate X) → toPredicate (fromPredicate p) ≡ p compIsIdFunc {X} p i = equivFun (compIsIdEquiv X i) p -module _ {X} where +module ResizedPredicateProps {X} (perX : PartialEquivalenceRelation X) where + open PartialEquivalenceRelation + entailmentResizedPredicate : ∀ (ϕ ψ : ResizedPredicate X) → A → Type ℓ entailmentResizedPredicate ϕ ψ r = ∀ (x : X) (a : A) (⊩ϕx : a ⊩ ∣ toPredicate ϕ ∣ x) → (r ⨾ a) ⊩ ∣ toPredicate ψ ∣ x isPropEntailmentResizedPredicate : ∀ ϕ ψ a → isProp (entailmentResizedPredicate ϕ ψ a) isPropEntailmentResizedPredicate ϕ ψ a = isPropΠ λ x → isPropΠ λ b → isPropΠ λ _ → (toPredicate ψ) .isPropValued _ _ + + isStrictResizedPredicate : ∀ (ϕ : ResizedPredicate X) → A → Type ℓ + isStrictResizedPredicate ϕ r = ∀ (x : X) (a : A) (⊩ϕx : a ⊩ ∣ toPredicate ϕ ∣ x) → (r ⨾ a) ⊩ ∣ perX .equality ∣ (x , x) + + isPropIsStrictResizedPredicate : ∀ ϕ r → isProp (isStrictResizedPredicate ϕ r) + isPropIsStrictResizedPredicate ϕ r = + isPropΠ λ x → isPropΠ λ a → isPropΠ λ _ → perX .equality .isPropValued _ _ + + isRelationalResizedPredicate : ∀ (ϕ : ResizedPredicate X) → A → Type ℓ + isRelationalResizedPredicate ϕ r = + ∀ (x x' : X) (a b : A) (⊩x~x' : a ⊩ ∣ perX .equality ∣ (x , x')) (⊩ϕx : b ⊩ ∣ toPredicate ϕ ∣ x) → (r ⨾ a ⨾ b) ⊩ ∣ toPredicate ϕ ∣ x' + + isPropIsRelationalResizedPredicate : ∀ ϕ r → isProp (isRelationalResizedPredicate ϕ r) + isPropIsRelationalResizedPredicate ϕ r = + isPropΠ λ x → isPropΠ λ x' → isPropΠ λ a → isPropΠ λ b → isPropΠ λ _ → isPropΠ λ _ → toPredicate ϕ .isPropValued _ _ diff --git a/src/Realizability/Topos/SubobjectClassifier.agda b/src/Realizability/Topos/SubobjectClassifier.agda index 49363da..14844b4 100644 --- a/src/Realizability/Topos/SubobjectClassifier.agda +++ b/src/Realizability/Topos/SubobjectClassifier.agda @@ -1,4 +1,4 @@ -open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; ⟦_⟧ to pre⟦_⟧) open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv @@ -42,6 +42,8 @@ open PartialEquivalenceRelation open FunctionalRelation open Category RT +⟦_⟧ = pre⟦_⟧ as + Ωper : PartialEquivalenceRelation (ResizedPredicate Unit*) Predicate.isSetX (equality Ωper) = isSet× isSetResizedPredicate isSetResizedPredicate Predicate.∣ equality Ωper ∣ (α , β) r = @@ -56,10 +58,10 @@ isPartialEquivalenceRelation.isSymmetric (isPerEquality Ωper) = do let ent₁ : ApplStrTerm as 2 - ent₁ = ` pr₂ ̇ # zero ̇ # one + ent₁ = ` pr₂ ̇ # one ̇ # zero ent₂ : ApplStrTerm as 2 - ent₂ = ` pr₁ ̇ # zero ̇ # one + ent₂ = ` pr₁ ̇ # one ̇ # zero realizer : ApplStrTerm as 1 realizer = ` pair ̇ (λ*abst ent₁) ̇ (λ*abst ent₂) @@ -69,10 +71,31 @@ isPartialEquivalenceRelation.isSymmetric (isPerEquality Ωper) = (λ a a⊩β → let eq : pr₁ ⨾ (λ* realizer ⨾ r) ⨾ a ≡ pr₂ ⨾ r ⨾ a - eq = {!!} + eq = + pr₁ ⨾ (λ* realizer ⨾ r) ⨾ a + ≡⟨ cong (λ x → pr₁ ⨾ x ⨾ a) (λ*ComputationRule realizer r) ⟩ + pr₁ ⨾ (pair ⨾ _ ⨾ _) ⨾ a + ≡⟨ cong (λ x → x ⨾ a) (pr₁pxy≡x _ _) ⟩ + ⟦ (λ*abst ent₁) ⟧ (r ∷ []) ⨾ a + ≡⟨ βreduction ent₁ a (r ∷ []) ⟩ + pr₂ ⨾ r ⨾ a + ∎ in subst (λ r' → r' ⊩ ∣ toPredicate α ∣ tt*) (sym eq) (pr₂r⊩β≤α a a⊩β)) , - (λ a a⊩α → subst (λ r' → r' ⊩ ∣ toPredicate β ∣ tt*) (sym {!!}) (pr₁r⊩α≤β a a⊩α)) }) + (λ a a⊩α → + let + eq : pr₂ ⨾ (λ* realizer ⨾ r) ⨾ a ≡ pr₁ ⨾ r ⨾ a + eq = + pr₂ ⨾ (λ* realizer ⨾ r) ⨾ a + ≡⟨ cong (λ x → pr₂ ⨾ x ⨾ a) (λ*ComputationRule realizer r) ⟩ + pr₂ ⨾ (pair ⨾ _ ⨾ _) ⨾ a + ≡⟨ cong (λ x → x ⨾ a) (pr₂pxy≡y _ _) ⟩ + ⟦ λ*abst ent₂ ⟧ (r ∷ []) ⨾ a + ≡⟨ βreduction ent₂ a (r ∷ []) ⟩ + pr₁ ⨾ r ⨾ a + ∎ + in + subst (λ r' → r' ⊩ ∣ toPredicate β ∣ tt*) (sym eq) (pr₁r⊩α≤β a a⊩α)) }) isPartialEquivalenceRelation.isTransitive (isPerEquality Ωper) = do return @@ -193,7 +216,7 @@ module _ let resultPredicate : Predicate Unit* resultPredicate = - consPredicate + makePredicate isSetUnit* (λ { tt* s → s ⊩ ∣ ϕ .predicate ∣ x }) (λ { tt* s → ϕ .predicate .isPropValued _ _ }) @@ -261,7 +284,7 @@ module _ (sym eq) (⊩ϕx'≤p _ (relϕ⊩isRelationalϕ x x' _ _ ⊩ϕx ⊩x~x'))) })) in - eq/ _ _ (answer , {!!}) + eq/ _ _ (answer , F≤G→G≤F subPer Ωper (composeFuncRel _ _ _ incFuncRel charFuncRel) (composeFuncRel _ _ _ (terminalFuncRel subPer) trueFuncRel) answer) classifies : isPullback RT subobjectCospan [ incFuncRel ] [ terminalFuncRel subPer ] subobjectSquareCommutes classifies {Y , perY} f g f⋆char≡g⋆true = @@ -269,8 +292,41 @@ module _ {P = λ f g → ∀ (commutes : f ⋆ [ charFuncRel ] ≡ g ⋆ [ trueFuncRel ]) → ∃![ hk ∈ RTMorphism perY subPer ] (f ≡ hk ⋆ [ incFuncRel ]) × (g ≡ hk ⋆ [ terminalFuncRel subPer ])} (λ f g → isPropΠ λ _ → isPropIsContr) (λ F G F⋆char≡G⋆true → + let + H : FunctionalRelation perY subPer + H = + makeFunctionalRelation + (makePredicate + (isSet× (perY .isSetX) (perX .isSetX)) + (λ { (y , x) r → (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (y , x) × (pr₂ ⨾ r) ⊩ ∣ ϕ .predicate ∣ x }) + λ { (y , x) r → isProp× (F .relation .isPropValued _ _) (ϕ .predicate .isPropValued _ _) }) + (makeIsFunctionalRelation + (do + (stF , stF⊩isStrictDomainF) ← F .isStrictDomain + return + ({!!} , + (λ { y x r (pr₁r⊩Fyx , pr₂r⊩ϕx) → + subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym {!!}) (stF⊩isStrictDomainF y x (pr₁ ⨾ r) pr₁r⊩Fyx)}))) + (do + (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain + return + ({!!} , + (λ { y x r (pr₁r⊩Fyx , pr₂r⊩ϕx) → + {!stFC⊩isStrictCodomainF y x (pr₁ ⨾ r) pr₁r⊩Fyx!} , + {!pr₂r⊩ϕx!} }))) + (do + return {!!}) + (do + return {!!}) + (do + return + ({!!} , + (λ { y r r⊩y~y → + return + ({!!} , {!!} , {!!})})))) + in uniqueExists - {!!} + [ H ] ({!!} , {!!}) (λ hk' → isProp× (squash/ _ _) (squash/ _ _)) From 3d55e7c859d805e24cab4192f9d025dc0fcbfd46 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 9 Apr 2024 04:34:40 +0530 Subject: [PATCH 39/61] Change predicate --- src/Realizability/Tripos/Prealgebra/Predicate/Base.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda b/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda index b8a4963..e89afc3 100644 --- a/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda +++ b/src/Realizability/Tripos/Prealgebra/Predicate/Base.agda @@ -14,7 +14,7 @@ open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) record Predicate (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ℓ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where - constructor consPredicate + constructor makePredicate field isSetX : isSet X ∣_∣ : X → A → Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'') From a16d85cd29947d4c1c2758d92e3d3f67956f323f Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 15 Apr 2024 16:50:03 +0530 Subject: [PATCH 40/61] Add subobject classifier --- docs/Cubical.Categories.Adjoint.html | 728 ++++++------- docs/Cubical.Categories.Category.Base.html | 238 +++-- ...ubical.Categories.Category.Properties.html | 58 +- ...l.Categories.Constructions.BinProduct.html | 12 +- ...cal.Categories.Constructions.Elements.html | 114 ++ ...l.Categories.Constructions.Slice.Base.html | 395 +++++++ ...gories.Constructions.Slice.Properties.html | 66 ++ ...ubical.Categories.Constructions.Slice.html | 18 + ...al.Categories.Constructions.SubObject.html | 93 ++ docs/Cubical.Categories.Equivalence.Base.html | 41 + ...cal.Categories.Equivalence.Properties.html | 203 ++++ docs/Cubical.Categories.Equivalence.html | 9 + docs/Cubical.Categories.Functor.Base.html | 34 +- docs/Cubical.Categories.Functor.Compose.html | 88 +- ...Cubical.Categories.Functor.Properties.html | 461 ++++---- docs/Cubical.Categories.Functor.html | 10 +- docs/Cubical.Categories.Instances.Cospan.html | 2 +- ...Cubical.Categories.Instances.Functors.html | 293 +++--- docs/Cubical.Categories.Instances.Sets.html | 139 +++ docs/Cubical.Categories.Isomorphism.html | 196 ++-- ...ubical.Categories.Limits.BinCoproduct.html | 4 +- .../Cubical.Categories.Limits.BinProduct.html | 105 +- docs/Cubical.Categories.Limits.Initial.html | 168 +-- docs/Cubical.Categories.Limits.Limits.html | 802 +++++++------- docs/Cubical.Categories.Limits.Pullback.html | 306 +++--- docs/Cubical.Categories.Limits.Terminal.html | 166 +-- docs/Cubical.Categories.Limits.html | 11 + docs/Cubical.Categories.Morphism.html | 260 +++-- ...Categories.NaturalTransformation.Base.html | 46 +- ...ries.NaturalTransformation.Properties.html | 386 +++---- ...ical.Categories.NaturalTransformation.html | 8 +- docs/Cubical.Data.Bool.Base.html | 6 +- docs/Cubical.Data.Bool.Properties.html | 44 +- docs/Cubical.Data.Bool.SwitchStatement.html | 44 + docs/Cubical.Data.Empty.Base.html | 5 +- docs/Cubical.Data.Empty.Properties.html | 10 +- docs/Cubical.Data.Fin.Base.html | 238 +++-- docs/Cubical.Data.Fin.Literals.html | 6 +- docs/Cubical.Data.Fin.Properties.html | 284 ++--- docs/Cubical.Data.FinData.Properties.html | 16 +- docs/Cubical.Data.List.Properties.html | 34 +- docs/Cubical.Data.Maybe.Properties.html | 44 +- docs/Cubical.Data.Nat.Order.Recursive.html | 20 +- docs/Cubical.Data.Nat.Order.html | 24 +- docs/Cubical.Data.Nat.Properties.html | 12 +- docs/Cubical.Data.Sigma.Base.html | 2 +- docs/Cubical.Data.Sigma.Properties.html | 60 +- docs/Cubical.Data.Sum.Properties.html | 104 +- docs/Cubical.Data.Unit.Base.html | 4 +- docs/Cubical.Data.Unit.Properties.html | 28 +- docs/Cubical.Data.Vec.Properties.html | 24 +- docs/Cubical.Displayed.Base.html | 18 +- docs/Cubical.Displayed.Function.html | 36 +- docs/Cubical.Displayed.Morphism.html | 2 +- docs/Cubical.Displayed.Prop.html | 16 +- docs/Cubical.Displayed.Properties.html | 22 +- docs/Cubical.Displayed.Record.html | 4 +- docs/Cubical.Displayed.Sigma.html | 4 +- docs/Cubical.Displayed.Subst.html | 20 +- docs/Cubical.Displayed.Unit.html | 2 +- docs/Cubical.Displayed.Universe.html | 2 +- docs/Cubical.Foundations.Equiv.Base.html | 2 +- docs/Cubical.Foundations.Equiv.Dependent.html | 378 +++++++ docs/Cubical.Foundations.Equiv.Fiberwise.html | 22 +- ...Cubical.Foundations.Equiv.HalfAdjoint.html | 16 +- .../Cubical.Foundations.Equiv.Properties.html | 130 +-- docs/Cubical.Foundations.Equiv.html | 523 +++++----- docs/Cubical.Foundations.Function.html | 155 +-- docs/Cubical.Foundations.GroupoidLaws.html | 8 +- docs/Cubical.Foundations.HLevels.html | 883 ++++++++-------- docs/Cubical.Foundations.Isomorphism.html | 26 +- docs/Cubical.Foundations.Path.html | 104 +- docs/Cubical.Foundations.Pointed.Base.html | 10 +- docs/Cubical.Foundations.Pointed.FunExt.html | 4 +- ...bical.Foundations.Pointed.Homogeneous.html | 22 +- .../Cubical.Foundations.Pointed.Homotopy.html | 2 +- ...ubical.Foundations.Pointed.Properties.html | 48 +- docs/Cubical.Foundations.Powerset.html | 19 +- docs/Cubical.Foundations.Prelude.html | 524 +++++----- docs/Cubical.Foundations.SIP.html | 32 +- docs/Cubical.Foundations.Transport.html | 30 +- docs/Cubical.Foundations.Univalence.html | 44 +- docs/Cubical.Functions.Embedding.html | 132 +-- docs/Cubical.Functions.Fibration.html | 14 +- docs/Cubical.Functions.Fixpoint.html | 4 +- docs/Cubical.Functions.FunExtEquiv.html | 14 +- docs/Cubical.Functions.Involution.html | 8 +- docs/Cubical.Functions.Logic.html | 2 +- docs/Cubical.Functions.Surjection.html | 6 +- ...Ts.PropositionalTruncation.MagicTrick.html | 2 +- ...Ts.PropositionalTruncation.Properties.html | 626 ++++++----- .../Cubical.HITs.SetQuotients.Properties.html | 84 +- ...Cubical.HITs.SetTruncation.Properties.html | 632 +++++------ ...Cubical.HITs.TypeQuotients.Properties.html | 6 +- docs/Cubical.Induction.WellFounded.html | 62 +- docs/Cubical.Relation.Binary.Base.html | 409 ++++---- ...l.Relation.Binary.Order.Preorder.Base.html | 20 +- ...tion.Binary.Order.Preorder.Properties.html | 18 +- ...elation.Binary.Order.StrictPoset.Base.html | 22 +- docs/Cubical.Relation.Binary.Properties.html | 89 +- docs/Cubical.Relation.Nullary.Base.html | 2 +- docs/Cubical.Relation.Nullary.Properties.html | 22 +- docs/Cubical.Structures.Axioms.html | 10 +- docs/Cubical.Structures.Pointed.html | 6 +- docs/Cubical.Tactics.NatSolver.EvalHom.html | 198 ++++ ...Cubical.Tactics.NatSolver.HornerForms.html | 102 ++ ...bical.Tactics.NatSolver.NatExpression.html | 30 + .../Cubical.Tactics.NatSolver.Reflection.html | 147 +++ docs/Cubical.Tactics.NatSolver.Solver.html | 125 +++ docs/Cubical.Tactics.NatSolver.html | 14 + .../Cubical.Tactics.Reflection.Utilities.html | 37 + .../Cubical.Tactics.Reflection.Variables.html | 83 ++ docs/Cubical.Tactics.Reflection.html | 116 +++ docs/Realizability.ApplicativeStructure.html | 314 +++--- docs/Realizability.Choice.html | 2 +- docs/Realizability.CombinatoryAlgebra.html | 365 +++---- docs/Realizability.PropResizing.html | 25 + docs/Realizability.Topos.BinProducts.html | 602 +++++++++++ docs/Realizability.Topos.Equalizer.html | 619 +++++++++++ docs/Realizability.Topos.Everything.html | 7 + ...ealizability.Topos.FunctionalRelation.html | 984 +++++++++++------- .../Realizability.Topos.MonicReprFuncRel.html | 280 +++++ docs/Realizability.Topos.Object.html | 109 +- .../Realizability.Topos.ResizedPredicate.html | 94 ++ docs/Realizability.Topos.StrictRelation.html | 634 +++++++++++ ...alizability.Topos.SubobjectClassifier.html | 940 +++++++++++++++++ docs/Realizability.Topos.TerminalObject.html | 112 ++ ...Tripos.Prealgebra.Meets.Commutativity.html | 131 ++- ...lity.Tripos.Prealgebra.Meets.Identity.html | 125 ++- ...lity.Tripos.Prealgebra.Predicate.Base.html | 120 +-- ...ripos.Prealgebra.Predicate.Properties.html | 657 ++++++------ ...izability.Tripos.Prealgebra.Predicate.html | 6 +- docs/index.html | 9 +- src/Realizability/ApplicativeStructure.agda | 16 + .../Topos/FunctionalRelation.agda | 31 + src/Realizability/Topos/PowerObject.agda | 44 + .../Topos/SubobjectClassifier.agda | 756 ++++++++++++-- 137 files changed, 13207 insertions(+), 6394 deletions(-) create mode 100644 docs/Cubical.Categories.Constructions.Elements.html create mode 100644 docs/Cubical.Categories.Constructions.Slice.Base.html create mode 100644 docs/Cubical.Categories.Constructions.Slice.Properties.html create mode 100644 docs/Cubical.Categories.Constructions.Slice.html create mode 100644 docs/Cubical.Categories.Constructions.SubObject.html create mode 100644 docs/Cubical.Categories.Equivalence.Base.html create mode 100644 docs/Cubical.Categories.Equivalence.Properties.html create mode 100644 docs/Cubical.Categories.Equivalence.html create mode 100644 docs/Cubical.Categories.Instances.Sets.html create mode 100644 docs/Cubical.Categories.Limits.html create mode 100644 docs/Cubical.Data.Bool.SwitchStatement.html create mode 100644 docs/Cubical.Foundations.Equiv.Dependent.html create mode 100644 docs/Cubical.Tactics.NatSolver.EvalHom.html create mode 100644 docs/Cubical.Tactics.NatSolver.HornerForms.html create mode 100644 docs/Cubical.Tactics.NatSolver.NatExpression.html create mode 100644 docs/Cubical.Tactics.NatSolver.Reflection.html create mode 100644 docs/Cubical.Tactics.NatSolver.Solver.html create mode 100644 docs/Cubical.Tactics.NatSolver.html create mode 100644 docs/Cubical.Tactics.Reflection.Utilities.html create mode 100644 docs/Cubical.Tactics.Reflection.Variables.html create mode 100644 docs/Cubical.Tactics.Reflection.html create mode 100644 docs/Realizability.PropResizing.html create mode 100644 docs/Realizability.Topos.BinProducts.html create mode 100644 docs/Realizability.Topos.Equalizer.html create mode 100644 docs/Realizability.Topos.MonicReprFuncRel.html create mode 100644 docs/Realizability.Topos.ResizedPredicate.html create mode 100644 docs/Realizability.Topos.StrictRelation.html create mode 100644 docs/Realizability.Topos.SubobjectClassifier.html create mode 100644 docs/Realizability.Topos.TerminalObject.html diff --git a/docs/Cubical.Categories.Adjoint.html b/docs/Cubical.Categories.Adjoint.html index 2c878d2..6007114 100644 --- a/docs/Cubical.Categories.Adjoint.html +++ b/docs/Cubical.Categories.Adjoint.html @@ -1,25 +1,25 @@ -Cubical.Categories.Adjoint
{-# OPTIONS --allow-unsolved-metas #-}
+Cubical.Categories.Adjoint
{-# OPTIONS --safe #-}
 
-module Cubical.Categories.Adjoint where
+module Cubical.Categories.Adjoint where
 
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Equiv
 
-open import Cubical.Data.Sigma
-open import Cubical.Categories.Category
-open import Cubical.Categories.Functor
-open import Cubical.Categories.Instances.Functors
-open import Cubical.Categories.NaturalTransformation
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Univalence
+open import Cubical.Data.Sigma
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Instances.Functors
+open import Cubical.Categories.NaturalTransformation
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Univalence
 
-open Functor
+open Functor
 
-open Iso
-open Category
+open Iso
+open Category
 
-{-
+{-
 ==============================================
                   Overview
 ==============================================
@@ -30,11 +30,11 @@
 equivalence.
 -}
 
-private
-  variable
-    ℓC ℓC' ℓD ℓD' : Level
+private
+  variable
+    ℓC ℓC' ℓD ℓD' : Level
 
-{-
+{-
 ==============================================
              Adjoint definitions
 ==============================================
@@ -45,195 +45,195 @@
 definition.
 -}
 
-module UnitCounit {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (G : Functor D C) where
-  record TriangleIdentities
-    (η : 𝟙⟨ C   (funcComp G F))
-    (ε : (funcComp F G)  𝟙⟨ D )
-    : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
-    where
-    field
-      Δ₁ :  c  F  η  c   ⋆⟨ D  ε  F  c    D .id
-      Δ₂ :  d  η  G  d   ⋆⟨ C  G  ε  d    C .id
-
-  -- Adjoint def 1: unit-counit
-  record _⊣_ : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where
-    field
-      -- unit
-      η : 𝟙⟨ C   (funcComp G F)
-      -- counit
-      ε : (funcComp F G)  𝟙⟨ D 
-      triangleIdentities : TriangleIdentities η ε
-    open TriangleIdentities triangleIdentities public
-
-
-private
-  variable
-    C : Category ℓC ℓC'
-    D : Category ℓC ℓC'
-
-
-module _ {F : Functor C D} {G : Functor D C} where
-  open UnitCounit
-  open _⊣_
-  open NatTrans
-  open TriangleIdentities
-  opositeAdjunction : (F  G)  ((G ^opF)  (F ^opF))
-  N-ob (η (opositeAdjunction x)) = N-ob (ε x)
-  N-hom (η (opositeAdjunction x)) f = sym (N-hom (ε x) f)
-  N-ob (ε (opositeAdjunction x)) = N-ob (η x)
-  N-hom (ε (opositeAdjunction x)) f = sym (N-hom (η x) f)
-  Δ₁ (triangleIdentities (opositeAdjunction x)) =
-    Δ₂ (triangleIdentities x)
-  Δ₂ (triangleIdentities (opositeAdjunction x)) =
-   Δ₁ (triangleIdentities x)
-
-  Iso⊣^opF : Iso (F  G) ((G ^opF)  (F ^opF))
-  fun Iso⊣^opF = opositeAdjunction
-  inv Iso⊣^opF = _
-  rightInv Iso⊣^opF _ = refl
-  leftInv Iso⊣^opF _ = refl
-
-private
-  variable
-    F F' : Functor C D
-    G G' : Functor D C
-
-
-module AdjointUniqeUpToNatIso where
- open UnitCounit
- module Left
-          (F⊣G  : _⊣_ {D = D} F G)
-          (F'⊣G : F'  G) where
-  open NatTrans
-
-  private
-    variable
-      H H' : Functor C D
-
-  _D⋆_ = seq' D
-
-  m : (H⊣G : H  G) (H'⊣G : H'  G) 
-         {x}  D [ H  x  , H'  x  ]
-  m {H = H} H⊣G H'⊣G =
-    H  (H'⊣G .η)  _   ⋆⟨ D  (H⊣G .ε)  _  where open _⊣_
-
-  private
-   s : (H⊣G : H  G) (H'⊣G : H'  G)   {x} 
-           seq' D (m H'⊣G H⊣G {x}) (m H⊣G H'⊣G {x})
-               D .id
-   s {H = H} {H' = H'} H⊣G H'⊣G = by-N-homs  by-Δs
-     where
-      open _⊣_ H⊣G  using (η ; Δ₂)
-      open _⊣_ H'⊣G using (ε ; Δ₁)
-      by-N-homs =
-        AssocCong₂⋆R {C = D} _
-        (AssocCong₂⋆L {C = D} (sym (N-hom ε _)) _)
-           cong₂ _D⋆_
-               (sym (F-seq H' _ _)
-                ∙∙ cong (H' ⟪_⟫) ((sym (N-hom η  _)))
-                ∙∙ F-seq H' _ _)
-               (sym (N-hom ε _))
-
-      by-Δs =
-        ⋆Assoc D _ _ _
-        ∙∙ cong (H'  _  D⋆_)
-             (sym (⋆Assoc D _ _ _)
-              cong (_D⋆ ε  _ )
-                 (  sym (F-seq H' _ _)
-                 ∙∙ cong (H' ⟪_⟫) (Δ₂ (H'  _ ))
-                 ∙∙ F-id H')
-              ⋆IdL D _)
-        ∙∙ Δ₁ _
-
-  open NatIso
-  open isIso
-
-  F≅ᶜF' : F ≅ᶜ F'
-  N-ob (trans F≅ᶜF') _ = _
-  N-hom (trans F≅ᶜF') _ =
-       sym (⋆Assoc D _ _ _)
-    ∙∙ cong (_D⋆ (F⊣G .ε)  _ )
-         (sym (F-seq F _ _)
-         ∙∙ cong (F ⟪_⟫) (N-hom (F'⊣G .η) _)
-         ∙∙ (F-seq F _ _))
-    ∙∙ AssocCong₂⋆R {C = D} _ (N-hom (F⊣G .ε) _)
-   where open _⊣_
-  inv (nIso F≅ᶜF' _) = _
-  sec (nIso F≅ᶜF' _) = s F⊣G F'⊣G
-  ret (nIso F≅ᶜF' _) = s F'⊣G F⊣G
-
-  F≡F' : isUnivalent D  F  F'
-  F≡F' univD =
-   isUnivalent.CatIsoToPath
-    (isUnivalentFUNCTOR _ _ univD)
-     (NatIso→FUNCTORIso _ _ F≅ᶜF')
-
- module Right (F⊣G  : F UnitCounit.⊣ G)
-              (F⊣G' : F UnitCounit.⊣ G') where
-
-  G≅ᶜG' : G ≅ᶜ G'
-  G≅ᶜG' = Iso.inv congNatIso^opFiso
-    (Left.F≅ᶜF' (opositeAdjunction F⊣G')
-                (opositeAdjunction F⊣G))
-
-  open NatIso
-
-  G≡G' : isUnivalent _  G  G'
-  G≡G' univC =
-   isUnivalent.CatIsoToPath
-    (isUnivalentFUNCTOR _ _ univC)
-     (NatIso→FUNCTORIso _ _ G≅ᶜG')
-
-module NaturalBijection where
-  -- Adjoint def 2: natural bijection
-  record _⊣_ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (G : Functor D C) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where
-    field
-      adjIso :  {c d}  Iso (D [ F  c  , d ]) (C [ c , G  d  ])
-
-    infix 40 _♭
-    infix 40 _♯
-    _♭ :  {c d}  (D [ F  c  , d ])  (C [ c , G  d  ])
-    (_♭) {_} {_} = adjIso .fun
-
-    _♯ :  {c d}  (C [ c , G  d  ])  (D [ F  c  , d ])
-    (_♯) {_} {_} = adjIso .inv
-
-    field
-      adjNatInD :  {c : C .ob} {d d'} (f : D [ F  c  , d ]) (k : D [ d , d' ])
-                 (f ⋆⟨ D  k)   f  ⋆⟨ C  G  k 
-
-      adjNatInC :  {c' c d} (g : C [ c , G  d  ]) (h : C [ c' , c ])
-                 (h ⋆⟨ C  g)   F  h  ⋆⟨ D  g 
-
-    adjNatInD' :  {c : C .ob} {d d'} (g : C [ c , G  d  ]) (k : D [ d , d' ])
-                 g  ⋆⟨ D  k  (g ⋆⟨ C  G  k ) 
-    adjNatInD' {c} {d} {d'} g k =
-      g  ⋆⟨ D  k
-        ≡⟨ sym (adjIso .leftInv (g  ⋆⟨ D  k)) 
-      ((g  ⋆⟨ D  k) ) 
-        ≡⟨ cong _♯ (adjNatInD (g ) k) 
-      ((g )  ⋆⟨ C  G  k ) 
-        ≡⟨ cong _♯ (cong  g'  seq' C g' (G  k )) (adjIso .rightInv g)) 
-      (g ⋆⟨ C  G  k )  
-
-    adjNatInC' :  {c' c d} (f : D [ F  c  , d ]) (h : C [ c' , c ])
-                 h ⋆⟨ C  (f )  (F  h  ⋆⟨ D  f) 
-    adjNatInC' {c'} {c} {d} f h =
-      h ⋆⟨ C  (f )
-        ≡⟨ sym (adjIso .rightInv (h ⋆⟨ C  (f ))) 
-      ((h ⋆⟨ C  (f )) ) 
-        ≡⟨ cong _♭ (adjNatInC (f ) h) 
-      ((F  h  ⋆⟨ D  (f ) ) )
-        ≡⟨ cong _♭ (cong  f'  seq' D (F  h ) f') (adjIso .leftInv f)) 
-      (F  h  ⋆⟨ D  f)  
-
-  isLeftAdjoint : {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D)  Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
-  isLeftAdjoint {C = C}{D} F = Σ[ G  Functor D C ] F  G
-
-  isRightAdjoint : {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (G : Functor D C)  Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
-  isRightAdjoint {C = C}{D} G = Σ[ F  Functor C D ] F  G
-
-{-
+module UnitCounit {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (G : Functor D C) where
+  record TriangleIdentities
+    (η : 𝟙⟨ C   (funcComp G F))
+    (ε : (funcComp F G)  𝟙⟨ D )
+    : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
+    where
+    field
+      Δ₁ :  c  F  η  c   ⋆⟨ D  ε  F  c    D .id
+      Δ₂ :  d  η  G  d   ⋆⟨ C  G  ε  d    C .id
+
+  -- Adjoint def 1: unit-counit
+  record _⊣_ : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where
+    field
+      -- unit
+      η : 𝟙⟨ C   (funcComp G F)
+      -- counit
+      ε : (funcComp F G)  𝟙⟨ D 
+      triangleIdentities : TriangleIdentities η ε
+    open TriangleIdentities triangleIdentities public
+
+
+private
+  variable
+    C : Category ℓC ℓC'
+    D : Category ℓC ℓC'
+
+
+module _ {F : Functor C D} {G : Functor D C} where
+  open UnitCounit
+  open _⊣_
+  open NatTrans
+  open TriangleIdentities
+  opositeAdjunction : (F  G)  ((G ^opF)  (F ^opF))
+  N-ob (η (opositeAdjunction x)) = N-ob (ε x)
+  N-hom (η (opositeAdjunction x)) f = sym (N-hom (ε x) f)
+  N-ob (ε (opositeAdjunction x)) = N-ob (η x)
+  N-hom (ε (opositeAdjunction x)) f = sym (N-hom (η x) f)
+  Δ₁ (triangleIdentities (opositeAdjunction x)) =
+    Δ₂ (triangleIdentities x)
+  Δ₂ (triangleIdentities (opositeAdjunction x)) =
+   Δ₁ (triangleIdentities x)
+
+  Iso⊣^opF : Iso (F  G) ((G ^opF)  (F ^opF))
+  fun Iso⊣^opF = opositeAdjunction
+  inv Iso⊣^opF = _
+  rightInv Iso⊣^opF _ = refl
+  leftInv Iso⊣^opF _ = refl
+
+private
+  variable
+    F F' : Functor C D
+    G G' : Functor D C
+
+
+module AdjointUniqeUpToNatIso where
+ open UnitCounit
+ module Left
+          (F⊣G  : _⊣_ {D = D} F G)
+          (F'⊣G : F'  G) where
+  open NatTrans
+
+  private
+    variable
+      H H' : Functor C D
+
+  _D⋆_ = seq' D
+
+  m : (H⊣G : H  G) (H'⊣G : H'  G) 
+         {x}  D [ H  x  , H'  x  ]
+  m {H = H} H⊣G H'⊣G =
+    H  (H'⊣G .η)  _   ⋆⟨ D  (H⊣G .ε)  _  where open _⊣_
+
+  private
+   s : (H⊣G : H  G) (H'⊣G : H'  G)   {x} 
+           seq' D (m H'⊣G H⊣G {x}) (m H⊣G H'⊣G {x})
+               D .id
+   s {H = H} {H' = H'} H⊣G H'⊣G = by-N-homs  by-Δs
+     where
+      open _⊣_ H⊣G  using (η ; Δ₂)
+      open _⊣_ H'⊣G using (ε ; Δ₁)
+      by-N-homs =
+        AssocCong₂⋆R {C = D} _
+        (AssocCong₂⋆L {C = D} (sym (N-hom ε _)) _)
+           cong₂ _D⋆_
+               (sym (F-seq H' _ _)
+                ∙∙ cong (H' ⟪_⟫) ((sym (N-hom η  _)))
+                ∙∙ F-seq H' _ _)
+               (sym (N-hom ε _))
+
+      by-Δs =
+        ⋆Assoc D _ _ _
+        ∙∙ cong (H'  _  D⋆_)
+             (sym (⋆Assoc D _ _ _)
+              cong (_D⋆ ε  _ )
+                 (  sym (F-seq H' _ _)
+                 ∙∙ cong (H' ⟪_⟫) (Δ₂ (H'  _ ))
+                 ∙∙ F-id H')
+              ⋆IdL D _)
+        ∙∙ Δ₁ _
+
+  open NatIso
+  open isIso
+
+  F≅ᶜF' : F ≅ᶜ F'
+  N-ob (trans F≅ᶜF') _ = _
+  N-hom (trans F≅ᶜF') _ =
+       sym (⋆Assoc D _ _ _)
+    ∙∙ cong (_D⋆ (F⊣G .ε)  _ )
+         (sym (F-seq F _ _)
+         ∙∙ cong (F ⟪_⟫) (N-hom (F'⊣G .η) _)
+         ∙∙ (F-seq F _ _))
+    ∙∙ AssocCong₂⋆R {C = D} _ (N-hom (F⊣G .ε) _)
+   where open _⊣_
+  inv (nIso F≅ᶜF' _) = _
+  sec (nIso F≅ᶜF' _) = s F⊣G F'⊣G
+  ret (nIso F≅ᶜF' _) = s F'⊣G F⊣G
+
+  F≡F' : isUnivalent D  F  F'
+  F≡F' univD =
+   isUnivalent.CatIsoToPath
+    (isUnivalentFUNCTOR _ _ univD)
+     (NatIso→FUNCTORIso _ _ F≅ᶜF')
+
+ module Right (F⊣G  : F UnitCounit.⊣ G)
+              (F⊣G' : F UnitCounit.⊣ G') where
+
+  G≅ᶜG' : G ≅ᶜ G'
+  G≅ᶜG' = Iso.inv congNatIso^opFiso
+    (Left.F≅ᶜF' (opositeAdjunction F⊣G')
+                (opositeAdjunction F⊣G))
+
+  open NatIso
+
+  G≡G' : isUnivalent _  G  G'
+  G≡G' univC =
+   isUnivalent.CatIsoToPath
+    (isUnivalentFUNCTOR _ _ univC)
+     (NatIso→FUNCTORIso _ _ G≅ᶜG')
+
+module NaturalBijection where
+  -- Adjoint def 2: natural bijection
+  record _⊣_ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (G : Functor D C) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where
+    field
+      adjIso :  {c d}  Iso (D [ F  c  , d ]) (C [ c , G  d  ])
+
+    infix 40 _♭
+    infix 40 _♯
+    _♭ :  {c d}  (D [ F  c  , d ])  (C [ c , G  d  ])
+    (_♭) {_} {_} = adjIso .fun
+
+    _♯ :  {c d}  (C [ c , G  d  ])  (D [ F  c  , d ])
+    (_♯) {_} {_} = adjIso .inv
+
+    field
+      adjNatInD :  {c : C .ob} {d d'} (f : D [ F  c  , d ]) (k : D [ d , d' ])
+                 (f ⋆⟨ D  k)   f  ⋆⟨ C  G  k 
+
+      adjNatInC :  {c' c d} (g : C [ c , G  d  ]) (h : C [ c' , c ])
+                 (h ⋆⟨ C  g)   F  h  ⋆⟨ D  g 
+
+    adjNatInD' :  {c : C .ob} {d d'} (g : C [ c , G  d  ]) (k : D [ d , d' ])
+                 g  ⋆⟨ D  k  (g ⋆⟨ C  G  k ) 
+    adjNatInD' {c} {d} {d'} g k =
+      g  ⋆⟨ D  k
+        ≡⟨ sym (adjIso .leftInv (g  ⋆⟨ D  k)) 
+      ((g  ⋆⟨ D  k) ) 
+        ≡⟨ cong _♯ (adjNatInD (g ) k) 
+      ((g )  ⋆⟨ C  G  k ) 
+        ≡⟨ cong _♯ (cong  g'  seq' C g' (G  k )) (adjIso .rightInv g)) 
+      (g ⋆⟨ C  G  k )  
+
+    adjNatInC' :  {c' c d} (f : D [ F  c  , d ]) (h : C [ c' , c ])
+                 h ⋆⟨ C  (f )  (F  h  ⋆⟨ D  f) 
+    adjNatInC' {c'} {c} {d} f h =
+      h ⋆⟨ C  (f )
+        ≡⟨ sym (adjIso .rightInv (h ⋆⟨ C  (f ))) 
+      ((h ⋆⟨ C  (f )) ) 
+        ≡⟨ cong _♭ (adjNatInC (f ) h) 
+      ((F  h  ⋆⟨ D  (f ) ) )
+        ≡⟨ cong _♭ (cong  f'  seq' D (F  h ) f') (adjIso .leftInv f)) 
+      (F  h  ⋆⟨ D  f)  
+
+  isLeftAdjoint : {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D)  Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
+  isLeftAdjoint {C = C}{D} F = Σ[ G  Functor D C ] F  G
+
+  isRightAdjoint : {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (G : Functor D C)  Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
+  isRightAdjoint {C = C}{D} G = Σ[ F  Functor C D ] F  G
+
+{-
 ==============================================
             Proofs of equivalence
 ==============================================
@@ -245,160 +245,160 @@
 The second unnamed module does the reverse.
 -}
 
-module _ (F : Functor C D) (G : Functor D C) where
-  open UnitCounit
-  open NaturalBijection renaming (_⊣_ to _⊣²_)
-  module _ (adj : F ⊣² G) where
-    open _⊣²_ adj
-    open _⊣_
-
-    -- Naturality condition implies that a commutative square in C
-    -- appears iff the transpose in D is commutative as well
-    -- Used in adj'→adj
-    adjNat' :  {c c' d d'} {f : D [ F  c  , d ]} {k : D [ d , d' ]}
-             {h : C [ c , c' ]} {g : C [ c' , G  d'  ]}
-            -- commutativity of squares is iff
-             ((f ⋆⟨ D  k  F  h  ⋆⟨ D  g )  (f  ⋆⟨ C  G  k   h ⋆⟨ C  g))
-            × ((f  ⋆⟨ C  G  k   h ⋆⟨ C  g)  (f ⋆⟨ D  k  F  h  ⋆⟨ D  g ))
-    adjNat' {c} {c'} {d} {d'} {f} {k} {h} {g} = D→C , C→D
-      where
-        D→C : (f ⋆⟨ D  k  F  h  ⋆⟨ D  g )  (f  ⋆⟨ C  G  k   h ⋆⟨ C  g)
-        D→C eq = f  ⋆⟨ C  G  k 
-              ≡⟨ sym (adjNatInD _ _) 
-                ((f ⋆⟨ D  k) )
-              ≡⟨ cong _♭ eq 
-                (F  h  ⋆⟨ D  g ) 
-              ≡⟨ sym (cong _♭ (adjNatInC _ _)) 
-                (h ⋆⟨ C  g)  
-              ≡⟨ adjIso .rightInv _ 
-                h ⋆⟨ C  g
-              
-        C→D : (f  ⋆⟨ C  G  k   h ⋆⟨ C  g)  (f ⋆⟨ D  k  F  h  ⋆⟨ D  g )
-        C→D eq = f ⋆⟨ D  k
-              ≡⟨ sym (adjIso .leftInv _) 
-                (f ⋆⟨ D  k)  
-              ≡⟨ cong _♯ (adjNatInD _ _) 
-                (f  ⋆⟨ C  G  k ) 
-              ≡⟨ cong _♯ eq 
-                (h ⋆⟨ C  g) 
-              ≡⟨ adjNatInC _ _ 
-                F  h  ⋆⟨ D  g 
-              
-
-    open NatTrans
-
-    -- note : had to make this record syntax because termination checker was complaining
-    -- due to referencing η and ε from the definitions of Δs
-    adj'→adj : F  G
-    adj'→adj = record
-      { η = η'
-      ; ε = ε'
-      ; triangleIdentities = record
-        {Δ₁ = Δ₁'
-        ; Δ₂ = Δ₂' }}
-
-      where
-        -- ETA
-
-        -- trivial commutative diagram between identities in D
-        commInD :  {x y} (f : C [ x , y ])  D .id ⋆⟨ D  F  f   F  f  ⋆⟨ D  D .id
-        commInD f = (D .⋆IdL _)  sym (D .⋆IdR _)
-
-        sharpen1 :  {x y} (f : C [ x , y ])  F  f  ⋆⟨ D  D .id  F  f  ⋆⟨ D  D .id  
-        sharpen1 f = cong  v  F  f  ⋆⟨ D  v) (sym (adjIso .leftInv _))
-
-        η' : 𝟙⟨ C   G ∘F F
-        η' .N-ob x = D .id 
-        η' .N-hom f = sym (fst (adjNat') (commInD f  sharpen1 f))
-
-        -- EPSILON
-
-        -- trivial commutative diagram between identities in C
-        commInC :  {x y} (g : D [ x , y ])  C .id ⋆⟨ C  G  g   G  g  ⋆⟨ C  C .id
-        commInC g = (C .⋆IdL _)  sym (C .⋆IdR _)
-
-        sharpen2 :  {x y} (g : D [ x , y ])  C .id   ⋆⟨ C  G  g   C .id ⋆⟨ C  G  g 
-        sharpen2 g = cong  v  v ⋆⟨ C  G  g ) (adjIso .rightInv _)
-
-        ε' : F ∘F G  𝟙⟨ D 
-        ε' .N-ob x  = C .id 
-        ε' .N-hom g = sym (snd adjNat' (sharpen2 g  commInC g))
-
-        -- DELTA 1
-
-        Δ₁' :  c  F  η'  c   ⋆⟨ D  ε'  F  c    D .id
-        Δ₁' c =
-            F  η'  c   ⋆⟨ D  ε'  F  c  
-          ≡⟨ sym (snd adjNat' (cong  v  (η'  c ) ⋆⟨ C  v) (G .F-id))) 
-            D .id ⋆⟨ D  D .id
-          ≡⟨ D .⋆IdL _ 
-            D .id
-          
-
-        -- DELTA 2
-
-        Δ₂' :  d  η'  G  d   ⋆⟨ C  G  ε'  d    C .id
-        Δ₂' d =
-            (η'  G  d  ) ⋆⟨ C  (G  ε'  d  )
-          ≡⟨ fst adjNat' (cong  v  v ⋆⟨ D  (ε'  d )) (sym (F .F-id))) 
-            C .id ⋆⟨ C  C .id
-          ≡⟨ C .⋆IdL _ 
-            C .id
-          
-
-
-  module _ (adj : F  G) where
-    open _⊣_ adj
-    open _⊣²_
-    open NatTrans
-
-    adj→adj' : F ⊣² G
-    -- ∀ {c d} → Iso (D [ F ⟅ c ⟆ , d ]) (C [ c , G ⟅ d ⟆ ])
-    -- takes f to Gf precomposed with the unit
-    adj→adj' .adjIso {c = c} .fun f = η  c  ⋆⟨ C  G  f 
-    -- takes g to Fg postcomposed with the counit
-    adj→adj' .adjIso {d = d} .inv g = F  g  ⋆⟨ D  ε  d 
-    -- invertibility follows from the triangle identities
-    adj→adj' .adjIso {c = c} {d} .rightInv g
-      = η  c  ⋆⟨ C  G  F  g  ⋆⟨ D  ε  d  
-      ≡⟨ cong  v  η  c  ⋆⟨ C  v) (G .F-seq _ _) 
-        η  c  ⋆⟨ C  (G  F  g   ⋆⟨ C  G  ε  d  )
-      ≡⟨ sym (C .⋆Assoc _ _ _) 
-        η  c  ⋆⟨ C  G  F  g   ⋆⟨ C  G  ε  d  
-      -- apply naturality
-      ≡⟨ rCatWhisker {C = C} _ _ _ natu 
-        (g ⋆⟨ C  η  G  d  ) ⋆⟨ C  G  ε  d  
-      ≡⟨ C .⋆Assoc _ _ _ 
-        g ⋆⟨ C  (η  G  d   ⋆⟨ C  G  ε  d  )
-      ≡⟨ lCatWhisker {C = C} _ _ _ (Δ₂ d) 
-        g ⋆⟨ C  C .id
-      ≡⟨ C .⋆IdR _ 
-        g
-      
-      where
-        natu : η  c  ⋆⟨ C  G  F  g    g ⋆⟨ C  η  G  d  
-        natu = sym (η .N-hom _)
-    adj→adj' .adjIso {c = c} {d} .leftInv f
-      = F  η  c  ⋆⟨ C  G  f   ⋆⟨ D  ε  d 
-      ≡⟨ cong  v  v ⋆⟨ D  ε  d ) (F .F-seq _ _) 
-        F  η  c   ⋆⟨ D  F  G  f   ⋆⟨ D  ε  d 
-      ≡⟨ D .⋆Assoc _ _ _ 
-        F  η  c   ⋆⟨ D  (F  G  f   ⋆⟨ D  ε  d )
-      -- apply naturality
-      ≡⟨ lCatWhisker {C = D} _ _ _ natu 
-        F  η  c   ⋆⟨ D  (ε  F  c   ⋆⟨ D  f)
-      ≡⟨ sym (D .⋆Assoc _ _ _) 
-        F  η  c   ⋆⟨ D  ε  F  c   ⋆⟨ D  f
-      -- apply triangle identity
-      ≡⟨ rCatWhisker {C = D} _ _ _ (Δ₁ c) 
-        D .id ⋆⟨ D  f
-      ≡⟨ D .⋆IdL _ 
-        f
-      
-      where
-        natu : F  G  f   ⋆⟨ D  ε  d   ε  F  c   ⋆⟨ D  f
-        natu = ε .N-hom _
-    -- follows directly from functoriality
-    adj→adj' .adjNatInD {c = c} f k = cong  v  η  c  ⋆⟨ C  v) (G .F-seq _ _)  (sym (C .⋆Assoc _ _ _))
-    adj→adj' .adjNatInC {d = d} g h = cong  v  v ⋆⟨ D  ε  d ) (F .F-seq _ _)  D .⋆Assoc _ _ _
+module _ (F : Functor C D) (G : Functor D C) where
+  open UnitCounit
+  open NaturalBijection renaming (_⊣_ to _⊣²_)
+  module _ (adj : F ⊣² G) where
+    open _⊣²_ adj
+    open _⊣_
+
+    -- Naturality condition implies that a commutative square in C
+    -- appears iff the transpose in D is commutative as well
+    -- Used in adj'→adj
+    adjNat' :  {c c' d d'} {f : D [ F  c  , d ]} {k : D [ d , d' ]}
+             {h : C [ c , c' ]} {g : C [ c' , G  d'  ]}
+            -- commutativity of squares is iff
+             ((f ⋆⟨ D  k  F  h  ⋆⟨ D  g )  (f  ⋆⟨ C  G  k   h ⋆⟨ C  g))
+            × ((f  ⋆⟨ C  G  k   h ⋆⟨ C  g)  (f ⋆⟨ D  k  F  h  ⋆⟨ D  g ))
+    adjNat' {c} {c'} {d} {d'} {f} {k} {h} {g} = D→C , C→D
+      where
+        D→C : (f ⋆⟨ D  k  F  h  ⋆⟨ D  g )  (f  ⋆⟨ C  G  k   h ⋆⟨ C  g)
+        D→C eq = f  ⋆⟨ C  G  k 
+              ≡⟨ sym (adjNatInD _ _) 
+                ((f ⋆⟨ D  k) )
+              ≡⟨ cong _♭ eq 
+                (F  h  ⋆⟨ D  g ) 
+              ≡⟨ sym (cong _♭ (adjNatInC _ _)) 
+                (h ⋆⟨ C  g)  
+              ≡⟨ adjIso .rightInv _ 
+                h ⋆⟨ C  g
+              
+        C→D : (f  ⋆⟨ C  G  k   h ⋆⟨ C  g)  (f ⋆⟨ D  k  F  h  ⋆⟨ D  g )
+        C→D eq = f ⋆⟨ D  k
+              ≡⟨ sym (adjIso .leftInv _) 
+                (f ⋆⟨ D  k)  
+              ≡⟨ cong _♯ (adjNatInD _ _) 
+                (f  ⋆⟨ C  G  k ) 
+              ≡⟨ cong _♯ eq 
+                (h ⋆⟨ C  g) 
+              ≡⟨ adjNatInC _ _ 
+                F  h  ⋆⟨ D  g 
+              
+
+    open NatTrans
+
+    -- note : had to make this record syntax because termination checker was complaining
+    -- due to referencing η and ε from the definitions of Δs
+    adj'→adj : F  G
+    adj'→adj = record
+      { η = η'
+      ; ε = ε'
+      ; triangleIdentities = record
+        {Δ₁ = Δ₁'
+        ; Δ₂ = Δ₂' }}
+
+      where
+        -- ETA
+
+        -- trivial commutative diagram between identities in D
+        commInD :  {x y} (f : C [ x , y ])  D .id ⋆⟨ D  F  f   F  f  ⋆⟨ D  D .id
+        commInD f = (D .⋆IdL _)  sym (D .⋆IdR _)
+
+        sharpen1 :  {x y} (f : C [ x , y ])  F  f  ⋆⟨ D  D .id  F  f  ⋆⟨ D  D .id  
+        sharpen1 f = cong  v  F  f  ⋆⟨ D  v) (sym (adjIso .leftInv _))
+
+        η' : 𝟙⟨ C   G ∘F F
+        η' .N-ob x = D .id 
+        η' .N-hom f = sym (fst (adjNat') (commInD f  sharpen1 f))
+
+        -- EPSILON
+
+        -- trivial commutative diagram between identities in C
+        commInC :  {x y} (g : D [ x , y ])  C .id ⋆⟨ C  G  g   G  g  ⋆⟨ C  C .id
+        commInC g = (C .⋆IdL _)  sym (C .⋆IdR _)
+
+        sharpen2 :  {x y} (g : D [ x , y ])  C .id   ⋆⟨ C  G  g   C .id ⋆⟨ C  G  g 
+        sharpen2 g = cong  v  v ⋆⟨ C  G  g ) (adjIso .rightInv _)
+
+        ε' : F ∘F G  𝟙⟨ D 
+        ε' .N-ob x  = C .id 
+        ε' .N-hom g = sym (snd adjNat' (sharpen2 g  commInC g))
+
+        -- DELTA 1
+
+        Δ₁' :  c  F  η'  c   ⋆⟨ D  ε'  F  c    D .id
+        Δ₁' c =
+            F  η'  c   ⋆⟨ D  ε'  F  c  
+          ≡⟨ sym (snd adjNat' (cong  v  (η'  c ) ⋆⟨ C  v) (G .F-id))) 
+            D .id ⋆⟨ D  D .id
+          ≡⟨ D .⋆IdL _ 
+            D .id
+          
+
+        -- DELTA 2
+
+        Δ₂' :  d  η'  G  d   ⋆⟨ C  G  ε'  d    C .id
+        Δ₂' d =
+            (η'  G  d  ) ⋆⟨ C  (G  ε'  d  )
+          ≡⟨ fst adjNat' (cong  v  v ⋆⟨ D  (ε'  d )) (sym (F .F-id))) 
+            C .id ⋆⟨ C  C .id
+          ≡⟨ C .⋆IdL _ 
+            C .id
+          
+
+
+  module _ (adj : F  G) where
+    open _⊣_ adj
+    open _⊣²_
+    open NatTrans
+
+    adj→adj' : F ⊣² G
+    -- ∀ {c d} → Iso (D [ F ⟅ c ⟆ , d ]) (C [ c , G ⟅ d ⟆ ])
+    -- takes f to Gf precomposed with the unit
+    adj→adj' .adjIso {c = c} .fun f = η  c  ⋆⟨ C  G  f 
+    -- takes g to Fg postcomposed with the counit
+    adj→adj' .adjIso {d = d} .inv g = F  g  ⋆⟨ D  ε  d 
+    -- invertibility follows from the triangle identities
+    adj→adj' .adjIso {c = c} {d} .rightInv g
+      = η  c  ⋆⟨ C  G  F  g  ⋆⟨ D  ε  d  
+      ≡⟨ cong  v  η  c  ⋆⟨ C  v) (G .F-seq _ _) 
+        η  c  ⋆⟨ C  (G  F  g   ⋆⟨ C  G  ε  d  )
+      ≡⟨ sym (C .⋆Assoc _ _ _) 
+        η  c  ⋆⟨ C  G  F  g   ⋆⟨ C  G  ε  d  
+      -- apply naturality
+      ≡⟨ rCatWhisker {C = C} _ _ _ natu 
+        (g ⋆⟨ C  η  G  d  ) ⋆⟨ C  G  ε  d  
+      ≡⟨ C .⋆Assoc _ _ _ 
+        g ⋆⟨ C  (η  G  d   ⋆⟨ C  G  ε  d  )
+      ≡⟨ lCatWhisker {C = C} _ _ _ (Δ₂ d) 
+        g ⋆⟨ C  C .id
+      ≡⟨ C .⋆IdR _ 
+        g
+      
+      where
+        natu : η  c  ⋆⟨ C  G  F  g    g ⋆⟨ C  η  G  d  
+        natu = sym (η .N-hom _)
+    adj→adj' .adjIso {c = c} {d} .leftInv f
+      = F  η  c  ⋆⟨ C  G  f   ⋆⟨ D  ε  d 
+      ≡⟨ cong  v  v ⋆⟨ D  ε  d ) (F .F-seq _ _) 
+        F  η  c   ⋆⟨ D  F  G  f   ⋆⟨ D  ε  d 
+      ≡⟨ D .⋆Assoc _ _ _ 
+        F  η  c   ⋆⟨ D  (F  G  f   ⋆⟨ D  ε  d )
+      -- apply naturality
+      ≡⟨ lCatWhisker {C = D} _ _ _ natu 
+        F  η  c   ⋆⟨ D  (ε  F  c   ⋆⟨ D  f)
+      ≡⟨ sym (D .⋆Assoc _ _ _) 
+        F  η  c   ⋆⟨ D  ε  F  c   ⋆⟨ D  f
+      -- apply triangle identity
+      ≡⟨ rCatWhisker {C = D} _ _ _ (Δ₁ c) 
+        D .id ⋆⟨ D  f
+      ≡⟨ D .⋆IdL _ 
+        f
+      
+      where
+        natu : F  G  f   ⋆⟨ D  ε  d   ε  F  c   ⋆⟨ D  f
+        natu = ε .N-hom _
+    -- follows directly from functoriality
+    adj→adj' .adjNatInD {c = c} f k = cong  v  η  c  ⋆⟨ C  v) (G .F-seq _ _)  (sym (C .⋆Assoc _ _ _))
+    adj→adj' .adjNatInC {d = d} g h = cong  v  v ⋆⟨ D  ε  d ) (F .F-seq _ _)  D .⋆Assoc _ _ _
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Category.Base.html b/docs/Cubical.Categories.Category.Base.html index 1163e41..bd0d8b8 100644 --- a/docs/Cubical.Categories.Category.Base.html +++ b/docs/Cubical.Categories.Category.Base.html @@ -24,7 +24,7 @@ ⋆IdR : {x y} (f : Hom[ x , y ]) f id f ⋆Assoc : {x y z w} (f : Hom[ x , y ]) (g : Hom[ y , z ]) (h : Hom[ z , w ]) (f g) h f (g h) - isSetHom : {x y} isSet Hom[ x , y ] + isSetHom : {x y} isSet Hom[ x , y ] -- composition: alternative to diagramatic order _∘_ : {x y z} (g : Hom[ y , z ]) (f : Hom[ x , y ]) Hom[ x , z ] @@ -39,119 +39,127 @@ _[_,_] : (C : Category ℓ') (x y : C .ob) Type ℓ' _[_,_] = Hom[_,_] --- Needed to define this in order to be able to make the subsequence syntax declaration -seq' : (C : Category ℓ') {x y z} (f : C [ x , y ]) (g : C [ y , z ]) C [ x , z ] -seq' = _⋆_ +_End[_] : (C : Category ℓ') (x : C .ob) Type ℓ' +C End[ x ] = C [ x , x ] -infixl 15 seq' -syntax seq' C f g = f ⋆⟨ C g - --- composition -comp' : (C : Category ℓ') {x y z} (g : C [ y , z ]) (f : C [ x , y ]) C [ x , z ] -comp' = _∘_ - -infixr 16 comp' -syntax comp' C g f = g ∘⟨ C f - - --- Isomorphisms and paths in categories - -record isIso (C : Category ℓ'){x y : C .ob}(f : C [ x , y ]) : Type ℓ' where - constructor isiso - field - inv : C [ y , x ] - sec : inv ⋆⟨ C f C .id - ret : f ⋆⟨ C inv C .id - -open isIso - -isPropIsIso : {C : Category ℓ'}{x y : C .ob}(f : C [ x , y ]) isProp (isIso C f) -isPropIsIso {C = C} f p q i .inv = - (sym (C .⋆IdL _) - i q .sec (~ i) ⋆⟨ C p .inv) - C .⋆Assoc _ _ _ - i q .inv ⋆⟨ C p .ret i) - C .⋆IdR _) i -isPropIsIso {C = C} f p q i .sec j = - isSet→SquareP i j C .isSetHom) - (p .sec) (q .sec) i isPropIsIso {C = C} f p q i .inv ⋆⟨ C f) refl i j -isPropIsIso {C = C} f p q i .ret j = - isSet→SquareP i j C .isSetHom) - (p .ret) (q .ret) i f ⋆⟨ C isPropIsIso {C = C} f p q i .inv) refl i j - -CatIso : (C : Category ℓ') (x y : C .ob) Type ℓ' -CatIso C x y = Σ[ f C [ x , y ] ] isIso C f - -CatIso≡ : {C : Category ℓ'}{x y : C .ob}(f g : CatIso C x y) f .fst g .fst f g -CatIso≡ f g = Σ≡Prop isPropIsIso - --- `constructor` of CatIso -catiso : {C : Category ℓ'}{x y : C .ob} - (mor : C [ x , y ]) - (inv : C [ y , x ]) - (sec : inv ⋆⟨ C mor C .id) - (ret : mor ⋆⟨ C inv C .id) - CatIso C x y -catiso mor inv sec ret = mor , isiso inv sec ret - - -idCatIso : {C : Category ℓ'} {x : C .ob} CatIso C x x -idCatIso {C = C} = C .id , isiso (C .id) (C .⋆IdL (C .id)) (C .⋆IdL (C .id)) - -isSet-CatIso : {C : Category ℓ'} x y isSet (CatIso C x y) -isSet-CatIso {C = C} x y = isOfHLevelΣ 2 (C .isSetHom) f isProp→isSet (isPropIsIso f)) - - -pathToIso : {C : Category ℓ'} {x y : C .ob} (p : x y) CatIso C x y -pathToIso {C = C} p = J z _ CatIso C _ z) idCatIso p - -pathToIso-refl : {C : Category ℓ'} {x : C .ob} pathToIso {C = C} {x} refl idCatIso -pathToIso-refl {C = C} {x} = JRefl z _ CatIso C x z) (idCatIso) - - --- Univalent Categories -record isUnivalent (C : Category ℓ') : Type (ℓ-max ℓ') where - field - univ : (x y : C .ob) isEquiv (pathToIso {C = C} {x = x} {y = y}) - - -- package up the univalence equivalence - univEquiv : (x y : C .ob) (x y) (CatIso _ x y) - univEquiv x y = pathToIso , univ x y - - -- The function extracting paths from category-theoretic isomorphisms. - CatIsoToPath : {x y : C .ob} (p : CatIso _ x y) x y - CatIsoToPath = invEq (univEquiv _ _) - - isGroupoid-ob : isGroupoid (C .ob) - isGroupoid-ob = isOfHLevelPath'⁻ 2 _ _ isOfHLevelRespectEquiv 2 (invEquiv (univEquiv _ _)) (isSet-CatIso _ _)) - - --- Opposite category -_^op : Category ℓ' Category ℓ' -ob (C ^op) = ob C -Hom[_,_] (C ^op) x y = C [ y , x ] -id (C ^op) = id C -_⋆_ (C ^op) f g = g ⋆⟨ C f -⋆IdL (C ^op) = C .⋆IdR -⋆IdR (C ^op) = C .⋆IdL -⋆Assoc (C ^op) f g h = sym (C .⋆Assoc _ _ _) -isSetHom (C ^op) = C .isSetHom - -ΣPropCat : (C : Category ℓ') (P : (ob C)) Category ℓ' -ob (ΣPropCat C P) = Σ[ x ob C ] x P -Hom[_,_] (ΣPropCat C P) x y = C [ fst x , fst y ] -id (ΣPropCat C P) = id C -_⋆_ (ΣPropCat C P) = _⋆_ C -⋆IdL (ΣPropCat C P) = ⋆IdL C -⋆IdR (ΣPropCat C P) = ⋆IdR C -⋆Assoc (ΣPropCat C P) = ⋆Assoc C -isSetHom (ΣPropCat C P) = isSetHom C - -isIsoΣPropCat : {C : Category ℓ'} {P : (ob C)} - {x y : ob C} (p : x P) (q : y P) - (f : C [ x , y ]) - isIso C f isIso (ΣPropCat C P) {x , p} {y , q} f -inv (isIsoΣPropCat p q f isIsoF) = isIsoF .inv -sec (isIsoΣPropCat p q f isIsoF) = isIsoF .sec -ret (isIsoΣPropCat p q f isIsoF) = isIsoF .ret + +-- Needed to define this in order to be able to make the subsequence syntax declaration +seq' : (C : Category ℓ') {x y z} (f : C [ x , y ]) (g : C [ y , z ]) C [ x , z ] +seq' = _⋆_ + +infixl 15 seq' +syntax seq' C f g = f ⋆⟨ C g + +-- composition +comp' : (C : Category ℓ') {x y z} (g : C [ y , z ]) (f : C [ x , y ]) C [ x , z ] +comp' = _∘_ + +infixr 16 comp' +syntax comp' C g f = g ∘⟨ C f + + +-- Isomorphisms and paths in categories + +record isIso (C : Category ℓ'){x y : C .ob}(f : C [ x , y ]) : Type ℓ' where + constructor isiso + field + inv : C [ y , x ] + sec : inv ⋆⟨ C f C .id + ret : f ⋆⟨ C inv C .id + +open isIso + +isPropIsIso : {C : Category ℓ'}{x y : C .ob}(f : C [ x , y ]) isProp (isIso C f) +isPropIsIso {C = C} f p q i .inv = + (sym (C .⋆IdL _) + i q .sec (~ i) ⋆⟨ C p .inv) + C .⋆Assoc _ _ _ + i q .inv ⋆⟨ C p .ret i) + C .⋆IdR _) i +isPropIsIso {C = C} f p q i .sec j = + isSet→SquareP i j C .isSetHom) + (p .sec) (q .sec) i isPropIsIso {C = C} f p q i .inv ⋆⟨ C f) refl i j +isPropIsIso {C = C} f p q i .ret j = + isSet→SquareP i j C .isSetHom) + (p .ret) (q .ret) i f ⋆⟨ C isPropIsIso {C = C} f p q i .inv) refl i j + +CatIso : (C : Category ℓ') (x y : C .ob) Type ℓ' +CatIso C x y = Σ[ f C [ x , y ] ] isIso C f + +CatIso≡ : {C : Category ℓ'}{x y : C .ob}(f g : CatIso C x y) f .fst g .fst f g +CatIso≡ f g = Σ≡Prop isPropIsIso + +-- `constructor` of CatIso +catiso : {C : Category ℓ'}{x y : C .ob} + (mor : C [ x , y ]) + (inv : C [ y , x ]) + (sec : inv ⋆⟨ C mor C .id) + (ret : mor ⋆⟨ C inv C .id) + CatIso C x y +catiso mor inv sec ret = mor , isiso inv sec ret + + +idCatIso : {C : Category ℓ'} {x : C .ob} CatIso C x x +idCatIso {C = C} = C .id , isiso (C .id) (C .⋆IdL (C .id)) (C .⋆IdL (C .id)) + +isSet-CatIso : {C : Category ℓ'} x y isSet (CatIso C x y) +isSet-CatIso {C = C} x y = isOfHLevelΣ 2 (C .isSetHom) f isProp→isSet (isPropIsIso f)) + + +pathToIso : {C : Category ℓ'} {x y : C .ob} (p : x y) CatIso C x y +pathToIso {C = C} p = J z _ CatIso C _ z) idCatIso p + +pathToIso-refl : {C : Category ℓ'} {x : C .ob} pathToIso {C = C} {x} refl idCatIso +pathToIso-refl {C = C} {x} = JRefl z _ CatIso C x z) (idCatIso) + + +-- Univalent Categories +record isUnivalent (C : Category ℓ') : Type (ℓ-max ℓ') where + field + univ : (x y : C .ob) isEquiv (pathToIso {C = C} {x = x} {y = y}) + + -- package up the univalence equivalence + univEquiv : (x y : C .ob) (x y) (CatIso _ x y) + univEquiv x y = pathToIso , univ x y + + -- The function extracting paths from category-theoretic isomorphisms. + CatIsoToPath : {x y : C .ob} (p : CatIso _ x y) x y + CatIsoToPath = invEq (univEquiv _ _) + + isGroupoid-ob : isGroupoid (C .ob) + isGroupoid-ob = isOfHLevelPath'⁻ 2 _ _ isOfHLevelRespectEquiv 2 (invEquiv (univEquiv _ _)) (isSet-CatIso _ _)) + +isPropIsUnivalent : {C : Category ℓ'} isProp (isUnivalent C) +isPropIsUnivalent = + isPropRetract isUnivalent.univ _ _ refl) + (isPropΠ2 λ _ _ isPropIsEquiv _ ) + +-- Opposite category +_^op : Category ℓ' Category ℓ' +ob (C ^op) = ob C +Hom[_,_] (C ^op) x y = C [ y , x ] +id (C ^op) = id C +_⋆_ (C ^op) f g = g ⋆⟨ C f +⋆IdL (C ^op) = C .⋆IdR +⋆IdR (C ^op) = C .⋆IdL +⋆Assoc (C ^op) f g h = sym (C .⋆Assoc _ _ _) +isSetHom (C ^op) = C .isSetHom + +ΣPropCat : (C : Category ℓ') (P : (ob C)) Category ℓ' +ob (ΣPropCat C P) = Σ[ x ob C ] x P +Hom[_,_] (ΣPropCat C P) x y = C [ fst x , fst y ] +id (ΣPropCat C P) = id C +_⋆_ (ΣPropCat C P) = _⋆_ C +⋆IdL (ΣPropCat C P) = ⋆IdL C +⋆IdR (ΣPropCat C P) = ⋆IdR C +⋆Assoc (ΣPropCat C P) = ⋆Assoc C +isSetHom (ΣPropCat C P) = isSetHom C + +isIsoΣPropCat : {C : Category ℓ'} {P : (ob C)} + {x y : ob C} (p : x P) (q : y P) + (f : C [ x , y ]) + isIso C f isIso (ΣPropCat C P) {x , p} {y , q} f +inv (isIsoΣPropCat p q f isIsoF) = isIsoF .inv +sec (isIsoΣPropCat p q f isIsoF) = isIsoF .sec +ret (isIsoΣPropCat p q f isIsoF) = isIsoF .ret
\ No newline at end of file diff --git a/docs/Cubical.Categories.Category.Properties.html b/docs/Cubical.Categories.Category.Properties.html index 71ac7f0..92726db 100644 --- a/docs/Cubical.Categories.Category.Properties.html +++ b/docs/Cubical.Categories.Category.Properties.html @@ -18,21 +18,21 @@ -- isSet where your allowed to compare paths where one side is only -- equal up to path. Is there a generalization of this? isSetHomP1 : {x y : C .ob} {n : C [ x , y ]} - isOfHLevelDep 1 m m n) - isSetHomP1 {x = x} {y} {n} = isOfHLevel→isOfHLevelDep 1 m isSetHom C m n) + isOfHLevelDep 1 m m n) + isSetHomP1 {x = x} {y} {n} = isOfHLevel→isOfHLevelDep 1 m isSetHom C m n) -- isSet where the arrows can be between non-definitionally equal obs isSetHomP2l : {y : C .ob} - isOfHLevelDep 2 x C [ x , y ]) - isSetHomP2l = isOfHLevel→isOfHLevelDep 2 a isSetHom C {x = a}) + isOfHLevelDep 2 x C [ x , y ]) + isSetHomP2l = isOfHLevel→isOfHLevelDep 2 a isSetHom C {x = a}) isSetHomP2r : {x : C .ob} - isOfHLevelDep 2 y C [ x , y ]) - isSetHomP2r = isOfHLevel→isOfHLevelDep 2 a isSetHom C {y = a}) + isOfHLevelDep 2 y C [ x , y ]) + isSetHomP2r = isOfHLevel→isOfHLevelDep 2 a isSetHom C {y = a}) -- opposite of opposite is definitionally equal to itself -involutiveOp : {C : Category ℓ'} C ^op ^op C +involutiveOp : {C : Category ℓ'} C ^op ^op C involutiveOp = refl module _ {C : Category ℓ'} where @@ -40,12 +40,12 @@ -- whisker the parallel morphisms g and g' with f lCatWhisker : {x y z : C .ob} (f : C [ x , y ]) (g g' : C [ y , z ]) (p : g g') - f ⋆⟨ C g f ⋆⟨ C g' + f ⋆⟨ C g f ⋆⟨ C g' lCatWhisker f _ _ p = cong (_⋆_ C f) p rCatWhisker : {x y z : C .ob} (f f' : C [ x , y ]) (g : C [ y , z ]) (p : f f') - f ⋆⟨ C g f' ⋆⟨ C g - rCatWhisker _ _ g p = cong v v ⋆⟨ C g) p + f ⋆⟨ C g f' ⋆⟨ C g + rCatWhisker _ _ g p = cong v v ⋆⟨ C g) p -- working with equal objects idP : {x x'} {p : x x'} C [ x , x' ] @@ -55,65 +55,65 @@ seqP : {x y y' z} {p : y y'} (f : C [ x , y ]) (g : C [ y' , z ]) C [ x , z ] - seqP {x} {_} {_} {z} {p} f g = f ⋆⟨ C (subst a C [ a , z ]) (sym p) g) + seqP {x} {_} {_} {z} {p} f g = f ⋆⟨ C (subst a C [ a , z ]) (sym p) g) -- also heterogeneous seq, but substituting on the left seqP' : {x y y' z} {p : y y'} (f : C [ x , y ]) (g : C [ y' , z ]) C [ x , z ] - seqP' {x} {_} {_} {z} {p} f g = subst a C [ x , a ]) p f ⋆⟨ C g + seqP' {x} {_} {_} {z} {p} f g = subst a C [ x , a ]) p f ⋆⟨ C g -- show that they're equal seqP≡seqP' : {x y y' z} {p : y y'} (f : C [ x , y ]) (g : C [ y' , z ]) seqP {p = p} f g seqP' {p = p} f g seqP≡seqP' {x = x} {z = z} {p = p} f g i = - (toPathP {A = λ i' C [ x , p i' ]} {f} refl i) - ⋆⟨ C - (toPathP {A = λ i' C [ p (~ i') , z ]} {x = g} (sym refl) (~ i)) + (toPathP {A = λ i' C [ x , p i' ]} {f} refl i) + ⋆⟨ C + (toPathP {A = λ i' C [ p (~ i') , z ]} {x = g} (sym refl) (~ i)) -- seqP is equal to normal seq when y ≡ y' seqP≡seq : {x y z} (f : C [ x , y ]) (g : C [ y , z ]) - seqP {p = refl} f g f ⋆⟨ C g - seqP≡seq {y = y} {z} f g i = f ⋆⟨ C toPathP {A = λ _ C [ y , z ]} {x = g} refl (~ i) + seqP {p = refl} f g f ⋆⟨ C g + seqP≡seq {y = y} {z} f g i = f ⋆⟨ C toPathP {A = λ _ C [ y , z ]} {x = g} refl (~ i) -- whiskering with heterogenous seq -- (maybe should let z be heterogeneous too) lCatWhiskerP : {x y z y' : C .ob} {p : y y'} (f : C [ x , y ]) (g : C [ y , z ]) (g' : C [ y' , z ]) (r : PathP i C [ p i , z ]) g g') - f ⋆⟨ C g seqP {p = p} f g' + f ⋆⟨ C g seqP {p = p} f g' lCatWhiskerP {z = z} {p = p} f g g' r = - cong v f ⋆⟨ C v) (sym (fromPathP (symP {A = λ i C [ p (~ i) , z ]} r))) + cong v f ⋆⟨ C v) (sym (fromPathP (symP {A = λ i C [ p (~ i) , z ]} r))) rCatWhiskerP : {x y' y z : C .ob} {p : y' y} (f' : C [ x , y' ]) (f : C [ x , y ]) (g : C [ y , z ]) (r : PathP i C [ x , p i ]) f' f) - f ⋆⟨ C g seqP' {p = p} f' g - rCatWhiskerP f' f g r = cong v v ⋆⟨ C g) (sym (fromPathP r)) + f ⋆⟨ C g seqP' {p = p} f' g + rCatWhiskerP f' f g r = cong v v ⋆⟨ C g) (sym (fromPathP r)) AssocCong₂⋆L : {x y' y z w : C .ob} {f' : C [ x , y' ]} {f : C [ x , y ]} {g' : C [ y' , z ]} {g : C [ y , z ]} - f ⋆⟨ C g f' ⋆⟨ C g' (h : C [ z , w ]) - f ⋆⟨ C (g ⋆⟨ C h) - f' ⋆⟨ C (g' ⋆⟨ C h) + f ⋆⟨ C g f' ⋆⟨ C g' (h : C [ z , w ]) + f ⋆⟨ C (g ⋆⟨ C h) + f' ⋆⟨ C (g' ⋆⟨ C h) AssocCong₂⋆L p h = sym (⋆Assoc C _ _ h) - ∙∙ i p i ⋆⟨ C h) ∙∙ + ∙∙ i p i ⋆⟨ C h) ∙∙ ⋆Assoc C _ _ h AssocCong₂⋆R : {x y z z' w : C .ob} (f : C [ x , y ]) {g' : C [ y , z' ]} {g : C [ y , z ]} {h' : C [ z' , w ]} {h : C [ z , w ]} - g ⋆⟨ C h g' ⋆⟨ C h' - (f ⋆⟨ C g) ⋆⟨ C h - (f ⋆⟨ C g') ⋆⟨ C h' + g ⋆⟨ C h g' ⋆⟨ C h' + (f ⋆⟨ C g) ⋆⟨ C h + (f ⋆⟨ C g') ⋆⟨ C h' AssocCong₂⋆R f p = ⋆Assoc C f _ _ - ∙∙ i f ⋆⟨ C p i) ∙∙ + ∙∙ i f ⋆⟨ C p i) ∙∙ sym (⋆Assoc C _ _ _) \ No newline at end of file diff --git a/docs/Cubical.Categories.Constructions.BinProduct.html b/docs/Cubical.Categories.Constructions.BinProduct.html index 1106138..b1cae3e 100644 --- a/docs/Cubical.Categories.Constructions.BinProduct.html +++ b/docs/Cubical.Categories.Constructions.BinProduct.html @@ -24,7 +24,7 @@ (C ×C D) .ob = (ob C) × (ob D) (C ×C D) .Hom[_,_] (c , d) (c' , d') = (C [ c , c' ]) × (D [ d , d' ]) (C ×C D) .id = (id C , id D) -(C ×C D) ._⋆_ _ _ = (_ ⋆⟨ C _ , _ ⋆⟨ D _) +(C ×C D) ._⋆_ _ _ = (_ ⋆⟨ C _ , _ ⋆⟨ D _) (C ×C D) .⋆IdL _ = ≡-× (⋆IdL C _) (⋆IdL D _) (C ×C D) .⋆IdR _ = ≡-× (⋆IdR C _) (⋆IdR D _) (C ×C D) .⋆Assoc _ _ _ = ≡-× (⋆Assoc C _ _ _) (⋆Assoc D _ _ _) @@ -95,11 +95,11 @@ -- The isomorphisms in product category - open isIso + open isIso - CatIso× : {x y : C .ob}{z w : D .ob} CatIso C x y CatIso D z w CatIso (C ×C D) (x , z) (y , w) + CatIso× : {x y : C .ob}{z w : D .ob} CatIso C x y CatIso D z w CatIso (C ×C D) (x , z) (y , w) CatIso× f g .fst = f .fst , g .fst - CatIso× f g .snd .inv = f .snd .inv , g .snd .inv - CatIso× f g .snd .sec i = f .snd .sec i , g .snd .sec i - CatIso× f g .snd .ret i = f .snd .ret i , g .snd .ret i + CatIso× f g .snd .inv = f .snd .inv , g .snd .inv + CatIso× f g .snd .sec i = f .snd .sec i , g .snd .sec i + CatIso× f g .snd .ret i = f .snd .ret i , g .snd .ret i \ No newline at end of file diff --git a/docs/Cubical.Categories.Constructions.Elements.html b/docs/Cubical.Categories.Constructions.Elements.html new file mode 100644 index 0000000..50f610e --- /dev/null +++ b/docs/Cubical.Categories.Constructions.Elements.html @@ -0,0 +1,114 @@ + +Cubical.Categories.Constructions.Elements
{-# OPTIONS --safe #-}
+
+-- The Category of Elements
+
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Data.Sigma
+
+open import Cubical.Categories.Category
+import      Cubical.Categories.Constructions.Slice.Base as Slice
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Instances.Sets
+open import Cubical.Categories.Isomorphism
+import      Cubical.Categories.Morphism as Morphism
+
+
+
+module Cubical.Categories.Constructions.Elements where
+
+-- some issues
+-- * always need to specify objects during composition because can't infer isSet
+open Category
+open Functor
+
+module Covariant { ℓ'} {C : Category  ℓ'} where
+    getIsSet :  {ℓS} (F : Functor C (SET ℓS))  (c : C .ob)  isSet (fst (F  c ))
+    getIsSet F c = snd (F  c )
+
+    Element :  {ℓS} (F : Functor C (SET ℓS))  Type (ℓ-max  ℓS)
+    Element F = Σ[ c  C .ob ] fst (F  c )
+
+    infix 50 ∫_
+    ∫_ :  {ℓS}  Functor C (SET ℓS)  Category (ℓ-max  ℓS) (ℓ-max ℓ' ℓS)
+    -- objects are (c , x) pairs where c ∈ C and x ∈ F c
+    ( F) .ob = Element F
+    -- morphisms are f : c → c' which take x to x'
+    ( F) .Hom[_,_] (c , x) (c' , x')  = fiber  (f : C [ c , c' ])  (F  f ) x) x'
+    ( F) .id {x = (c , x)} = C .id , funExt⁻ (F .F-id) x
+    ( F) ._⋆_ {c , x} {c₁ , x₁} {c₂ , x₂} (f , p) (g , q)
+      = (f ⋆⟨ C  g) , ((F  f ⋆⟨ C  g ) x
+                ≡⟨ funExt⁻ (F .F-seq _ _) _ 
+                  (F  g ) ((F  f ) x)
+                ≡⟨ cong (F  g ) p 
+                  (F  g ) x₁
+                ≡⟨ q 
+                  x₂
+                )
+    ( F) .⋆IdL o@{c , x} o1@{c' , x'} f'@(f , p) i
+      = (cIdL i) , isOfHLevel→isOfHLevelDep 1  a  isS' ((F  a ) x) x') p' p cIdL i
+        where
+          isS = getIsSet F c
+          isS' = getIsSet F c'
+          cIdL = C .⋆IdL f
+
+          -- proof from composition with id
+          p' : (F  C .id ⋆⟨ C  f ) x  x'
+          p' = snd (( F) ._⋆_ (( F) .id) f')
+    ( F) .⋆IdR o@{c , x} o1@{c' , x'} f'@(f , p) i
+        = (cIdR i) , isOfHLevel→isOfHLevelDep 1  a  isS' ((F  a ) x) x') p' p cIdR i
+          where
+            cIdR = C .⋆IdR f
+            isS' = getIsSet F c'
+
+            p' : (F  f ⋆⟨ C  C .id ) x  x'
+            p' = snd (( F) ._⋆_ f' (( F) .id))
+    ( F) .⋆Assoc o@{c , x} o1@{c₁ , x₁} o2@{c₂ , x₂} o3@{c₃ , x₃} f'@(f , p) g'@(g , q) h'@(h , r) i
+      = (cAssoc i) , isOfHLevel→isOfHLevelDep 1  a  isS₃ ((F  a ) x) x₃) p1 p2 cAssoc i
+        where
+          cAssoc = C .⋆Assoc f g h
+          isS₃ = getIsSet F c₃
+
+          p1 : (F  (f ⋆⟨ C  g) ⋆⟨ C  h ) x  x₃
+          p1 = snd (( F) ._⋆_ (( F) ._⋆_ {o} {o1} {o2} f' g') h')
+
+          p2 : (F  f ⋆⟨ C  (g ⋆⟨ C  h) ) x  x₃
+          p2 = snd (( F) ._⋆_ f' (( F) ._⋆_ {o1} {o2} {o3} g' h'))
+    ( F) .isSetHom {x} {y} = isSetΣSndProp (C .isSetHom) λ _  (F  fst y ) .snd _ _
+
+    ForgetElements :  {ℓS}  (F : Functor C (SET ℓS))  Functor ( F) C
+    F-ob (ForgetElements F) = fst
+    F-hom (ForgetElements F) = fst
+    F-id (ForgetElements F) = refl
+    F-seq (ForgetElements F) _ _ = refl
+
+module Contravariant { ℓ'} {C : Category  ℓ'} where
+    open Covariant {C = C ^op}
+
+    -- same thing but for presheaves
+    ∫ᴾ_ :  {ℓS}  Functor (C ^op) (SET ℓS)  Category (ℓ-max  ℓS) (ℓ-max ℓ' ℓS)
+    ∫ᴾ F = ( F) ^op
+
+    Elementᴾ :  {ℓS}  Functor (C ^op) (SET ℓS)  Type (ℓ-max  ℓS)
+    Elementᴾ F = (∫ᴾ F) .ob
+
+    -- helpful results
+
+    module _ {ℓS} {F : Functor (C ^op) (SET ℓS)} where
+
+      -- morphisms are equal as long as the morphisms in C are equal
+      ∫ᴾhomEq :  {o1 o1' o2 o2'} (f : (∫ᴾ F) [ o1 , o2 ]) (g : (∫ᴾ F) [ o1' , o2' ])
+               (p : o1  o1') (q : o2  o2')
+               (eqInC : PathP  i  C [ fst (p i) , fst (q i) ]) (fst f) (fst g))
+               PathP  i  (∫ᴾ F) [ p i , q i ]) f g
+      ∫ᴾhomEq _ _ _ _ = ΣPathPProp  f  snd (F  _ ) _ _)
+
+      ∫ᴾhomEqSimpl :  {o1 o2} (f g : (∫ᴾ F) [ o1 , o2 ])
+                    fst f  fst g  f  g
+      ∫ᴾhomEqSimpl f g p = ∫ᴾhomEq f g refl refl p
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Constructions.Slice.Base.html b/docs/Cubical.Categories.Constructions.Slice.Base.html new file mode 100644 index 0000000..1c65f6c --- /dev/null +++ b/docs/Cubical.Categories.Constructions.Slice.Base.html @@ -0,0 +1,395 @@ + +Cubical.Categories.Constructions.Slice.Base
{-# OPTIONS --safe #-}
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Transport using (transpFill)
+
+open import Cubical.Categories.Category renaming (isIso to isIsoC)
+open import Cubical.Categories.Morphism
+
+open import Cubical.Data.Sigma
+
+open Category
+open isUnivalent
+open Iso
+
+module Cubical.Categories.Constructions.Slice.Base { ℓ' : Level} (C : Category  ℓ') (c : C .ob) where
+
+-- just a helper to prevent redundency
+TypeC : Type (ℓ-suc (ℓ-max  ℓ'))
+TypeC = Type (ℓ-max  ℓ')
+
+-- Components of a slice category
+
+record SliceOb : TypeC where
+  constructor sliceob
+  field
+    {S-ob} : C .ob
+    S-arr : C [ S-ob , c ]
+
+open SliceOb public
+
+record SliceHom (a b : SliceOb) : Type ℓ' where
+  constructor slicehom
+  field
+    S-hom : C [ S-ob a , S-ob b ]
+    -- commutative diagram
+    S-comm : S-hom ⋆⟨ C  (S-arr b)  S-arr a
+
+open SliceHom public
+
+-- Helpers for working with equality
+-- can probably replace these by showing that SliceOb is isomorphic to Sigma and
+-- that paths are isomorphic to Sigma? But sounds like that would need a lot of transp
+
+SliceOb-≡-intro :  {a b} {f g}
+                  (p : a  b)
+                  PathP  i  C [ p i , c ]) f g
+                  sliceob {a} f  sliceob {b} g
+SliceOb-≡-intro p q = λ i  sliceob {p i} (q i)
+
+module _ {xf yg : SliceOb} where
+  private
+    x = xf .S-ob
+    f = xf .S-arr
+    y = yg .S-ob
+    g = yg .S-arr
+
+  -- a path between slice objects is the "same" as a pair of paths between C obs and C arrows
+  SOPathIsoPathΣ : Iso (xf  yg) (Σ[ p  x  y ] PathP  i  C [ p i , c ]) f g)
+  SOPathIsoPathΣ .fun p =  i  (p i) .S-ob) ,  i  (p i) .S-arr)
+  SOPathIsoPathΣ .inv (p , q) i = sliceob {p i} (q i)
+  SOPathIsoPathΣ .rightInv _ = refl
+  SOPathIsoPathΣ .leftInv _ = refl
+
+  SOPath≃PathΣ = isoToEquiv SOPathIsoPathΣ
+
+  SOPath≡PathΣ = ua (isoToEquiv SOPathIsoPathΣ)
+
+-- If the type of objects of C forms a set then so does the type of objects of the slice cat
+module _ (isSetCOb : isSet (C .ob)) where
+  isSetSliceOb : isSet SliceOb
+  isSetSliceOb x y =
+    subst
+       t  isProp t)
+      (sym (SOPath≡PathΣ {xf = x} {yg = y}))
+      (isPropΣ
+        (isSetCOb (x .S-ob) (y .S-ob))
+        λ x≡y 
+          subst
+             t  isProp t)
+            (sym (ua (PathP≃Path  i  C [ x≡y i , c ]) (x .S-arr) (y .S-arr))))
+            (C .isSetHom (transport  i  C [ x≡y i , c ]) (x .S-arr)) (y .S-arr)))
+
+-- intro and elim for working with SliceHom equalities (is there a better way to do this?)
+SliceHom-≡-intro :  {a b} {f g} {c₁} {c₂}
+                 (p : f  g)
+                 PathP  i  (p i) ⋆⟨ C  (S-arr b)  S-arr a) c₁ c₂
+                 slicehom f c₁  slicehom g c₂
+SliceHom-≡-intro p q = λ i  slicehom (p i) (q i)
+
+SliceHom-≡-elim :  {a b} {f g} {c₁} {c₂}
+                 slicehom f c₁  slicehom g c₂
+                 Σ[ p  f  g ] PathP  i  (p i) ⋆⟨ C  (S-arr b)  S-arr a) c₁ c₂
+SliceHom-≡-elim r =  i  S-hom (r i)) , λ i  S-comm (r i)
+
+
+SliceHom-≡-intro' :  {a b} {f g : C [ a .S-ob , b .S-ob ]} {c₁} {c₂}
+                   (p : f  g)
+                   slicehom f c₁  slicehom g c₂
+SliceHom-≡-intro' {a} {b} {f} {g} {c₁} {c₂} p i = slicehom (p i) (c₁≡c₂ i)
+  where
+    c₁≡c₂ : PathP  i  (p i) ⋆⟨ C  (b .S-arr)  a .S-arr) c₁ c₂
+    c₁≡c₂ = isOfHLevel→isOfHLevelDep 1  _  C .isSetHom _ _) c₁ c₂ p
+
+-- SliceHom is isomorphic to the Sigma type with the same components
+SliceHom-Σ-Iso :  {a b}
+             Iso (SliceHom a b) (Σ[ h  C [ S-ob a , S-ob b ] ] h ⋆⟨ C  (S-arr b)  S-arr a)
+SliceHom-Σ-Iso .fun (slicehom h c) = h , c
+SliceHom-Σ-Iso .inv (h , c) = slicehom h c
+SliceHom-Σ-Iso .rightInv = λ x  refl
+SliceHom-Σ-Iso .leftInv = λ x  refl
+
+
+-- Category definition
+
+SliceCat : Category (ℓ-max  ℓ') ℓ'
+ob SliceCat = SliceOb
+Hom[_,_] SliceCat = SliceHom
+id SliceCat = slicehom (C .id) (C .⋆IdL _)
+_⋆_ SliceCat {sliceob j} {sliceob k} {sliceob l} (slicehom f p) (slicehom g p') =
+  slicehom
+    (f ⋆⟨ C  g)
+    ( f ⋆⟨ C  g ⋆⟨ C  l
+    ≡⟨ C .⋆Assoc _ _ _ 
+      f ⋆⟨ C  (g ⋆⟨ C  l)
+    ≡⟨ cong  v  f ⋆⟨ C  v) p' 
+      f ⋆⟨ C  k
+    ≡⟨ p 
+      j
+    )
+⋆IdL SliceCat (slicehom S-hom S-comm) =
+  SliceHom-≡-intro (⋆IdL C _) (toPathP (C .isSetHom _ _ _ _))
+⋆IdR SliceCat (slicehom S-hom S-comm) =
+  SliceHom-≡-intro (⋆IdR C _) (toPathP (C .isSetHom _ _ _ _))
+⋆Assoc SliceCat f g h =
+  SliceHom-≡-intro (⋆Assoc C _ _ _) (toPathP (C .isSetHom _ _ _ _))
+isSetHom SliceCat {a} {b} (slicehom f c₁) (slicehom g c₂) p q = cong isoP p'≡q'
+    where
+      -- paths between SliceHoms are equivalent to the projection paths
+      p' : Σ[ p  f  g ] PathP  i  (p i) ⋆⟨ C  (S-arr b)  S-arr a) c₁ c₂
+      p' = SliceHom-≡-elim p
+      q' : Σ[ p  f  g ] PathP  i  (p i) ⋆⟨ C  (S-arr b)  S-arr a) c₁ c₂
+      q' = SliceHom-≡-elim q
+
+      -- we want all paths between (dependent) paths of this type to be equal
+      B = λ v  v ⋆⟨ C  (S-arr b)  S-arr a
+
+      -- need the groupoidness for dependent paths
+      isGroupoidDepHom : isOfHLevelDep 2 B
+      isGroupoidDepHom = isOfHLevel→isOfHLevelDep 2  v x y  isSet→isGroupoid (C .isSetHom) _ _ x y)
+
+      -- we first prove that the projected paths are equal
+      p'≡q' : p'  q'
+      p'≡q' = ΣPathP (C .isSetHom _ _ _ _ , toPathP (isGroupoidDepHom _ _ _ _ _))
+
+      -- and then we can use equivalence to lift these paths up
+      -- to actual SliceHom paths
+      isoP = λ g  cong (inv SliceHom-Σ-Iso) (fun (ΣPathIsoPathΣ) g)
+
+-- SliceCat is univalent if C is univalent
+
+module _  isU : isUnivalent C  where
+  open isIsoC
+  open Iso
+
+  module _ { xf yg : SliceOb } where
+    private
+      x = xf .S-ob
+      y = yg .S-ob
+
+    -- names for the equivalences/isos
+
+    pathIsoEquiv : (x  y)  (CatIso _ x y)
+    pathIsoEquiv = univEquiv isU x y
+
+    isoPathEquiv : (CatIso _ x y)  (x  y)
+    isoPathEquiv = invEquiv pathIsoEquiv
+
+    pToIIso' : Iso (x  y) (CatIso _ x y)
+    pToIIso' = equivToIso pathIsoEquiv
+
+    -- the iso in SliceCat we're given induces an iso in C between x and y
+    module _ ( cIso@(kc , isiso lc s r) : CatIso SliceCat xf yg ) where
+      extractIso' : CatIso C x y
+      extractIso' .fst = kc .S-hom
+      extractIso' .snd .inv = lc .S-hom
+      extractIso' .snd .sec i = (s i) .S-hom
+      extractIso' .snd .ret i = (r i) .S-hom
+
+  instance
+    preservesUnivalenceSlice : isUnivalent SliceCat
+    -- we prove the equivalence by going through Iso
+    preservesUnivalenceSlice .univ xf@(sliceob {x} f) yg@(sliceob {y} g) = isoToIsEquiv sIso
+      where
+        -- this is just here because the type checker can't seem to infer xf and yg
+        pToIIso : Iso (x  y) (CatIso _ x y)
+        pToIIso = pToIIso' {xf = xf} {yg}
+
+        -- the meat of the proof
+        sIso : Iso (xf  yg) (CatIso _ xf yg)
+        sIso .fun p = pathToIso p -- we use the normal pathToIso via path induction to get an isomorphism
+        sIso .inv is@(kc , isiso lc s r) = SliceOb-≡-intro x≡y (symP (sym (lc .S-comm)  lf≡f))
+          where
+            -- we get a path between xf and yg by combining paths between
+            -- x and y, and f and g
+            -- 1. x≡y follows from univalence of C
+            -- 2. f≡g is more tricky; by commutativity, we know that g ≡ l ⋆ f
+              -- so we want l to be id; we get this by showing: id ≡ pathToIso x y x≡y ≡ l
+              -- where the first step follows from path induction, and the second from univalence of C
+
+            -- morphisms in C from kc and lc
+            k = kc .S-hom
+            l = lc .S-hom
+
+            -- extract out the iso between x and y
+            extractIso : CatIso C x y
+            extractIso = extractIso' is
+
+            -- and we can use univalence of C to get x ≡ y
+            x≡y : x  y
+            x≡y = pToIIso .inv extractIso
+
+            -- to show that f ≡ g, we show that l ≡ id
+            -- by using C's isomorphism
+            pToI≡id : PathP  i  C [ x≡y (~ i) , x ]) (pathToIso {C = C} x≡y .snd .inv) (C .id)
+            pToI≡id = J  y p  PathP  i  C [ p (~ i) , x ]) (pathToIso {C = C} p .snd .inv) (C .id))
+                         j  JRefl pToIFam pToIBase j .snd .inv)
+                        x≡y
+              where
+                idx = C .id
+                pToIFam =  z _  CatIso C x z)
+                pToIBase = catiso (C .id) idx (C .⋆IdL idx) (C .⋆IdL idx)
+
+            l≡pToI : l  pathToIso {C = C} x≡y .snd .inv
+            l≡pToI i = pToIIso .rightInv extractIso (~ i) .snd .inv
+
+            l≡id : PathP  i  C [ x≡y (~ i) , x ]) l (C .id)
+            l≡id = l≡pToI  pToI≡id
+
+            lf≡f : PathP  i  C [ x≡y (~ i) , c ]) (l ⋆⟨ C  f) f
+            lf≡f =  i  (l≡id i) ⋆⟨ C  f)  C .⋆IdL _
+
+        sIso .rightInv is@(kc , isiso lc s r) i = catiso (kc'≡kc i) (lc'≡lc i) (s'≡s i) (r'≡r i)
+          -- we prove rightInv using a combination of univalence and the fact that homs are an h-set
+          where
+            kc' = (sIso .fun) (sIso .inv is) .fst
+            lc' = (sIso .fun) (sIso .inv is) .snd .inv
+            k' = kc' .S-hom
+            l' = lc' .S-hom
+            k = kc .S-hom
+            l = lc .S-hom
+
+            extractIso : CatIso C x y
+            extractIso = extractIso' is
+
+            -- we do the equality component wise
+
+            -- mor
+
+            k'≡k : k'  k
+            k'≡k i = (pToIIso .rightInv extractIso) i .fst
+
+            kcom'≡kcom : PathP  j  (k'≡k j) ⋆⟨ C  g  f) (kc' .S-comm) (kc .S-comm)
+            kcom'≡kcom = isSetHomP1 {C = C} _ _ λ i  (k'≡k i) ⋆⟨ C  g
+            kc'≡kc : kc'  kc
+            kc'≡kc i = slicehom (k'≡k i) (kcom'≡kcom i)
+
+            -- inv
+
+            l'≡l : l'  l
+            l'≡l i = (pToIIso .rightInv extractIso) i .snd .inv
+
+            lcom'≡lcom : PathP  j  (l'≡l j) ⋆⟨ C  f  g) (lc' .S-comm) (lc .S-comm)
+            lcom'≡lcom = isSetHomP1 {C = C} _ _ λ i  (l'≡l i) ⋆⟨ C  f
+
+            lc'≡lc : lc'  lc
+            lc'≡lc i = slicehom (l'≡l i) (lcom'≡lcom i)
+
+            -- sec
+
+            s' = (sIso .fun) (sIso .inv is) .snd .sec
+            s'≡s : PathP  i  lc'≡lc i ⋆⟨ SliceCat  kc'≡kc i  SliceCat .id) s' s
+            s'≡s = isSetHomP1 {C = SliceCat} _ _ λ i  lc'≡lc i ⋆⟨ SliceCat  kc'≡kc i
+
+            -- ret
+
+            r' = (sIso .fun) (sIso .inv is) .snd .ret
+            r'≡r : PathP  i  kc'≡kc i ⋆⟨ SliceCat  lc'≡lc i  SliceCat .id) r' r
+            r'≡r = isSetHomP1 {C = SliceCat} _ _ λ i  kc'≡kc i ⋆⟨ SliceCat  lc'≡lc i
+
+        sIso .leftInv p = p'≡p
+          -- to show that the round trip is equivalent to the identity
+          -- we show that this is true for each component (S-ob, S-arr)
+          -- and then combine
+          -- specifically, we show that p'Ob≡pOb and p'Mor≡pMor
+          -- and it follows that p'≡p
+          where
+            p' = (sIso .inv) (sIso .fun p)
+
+            pOb : x  y
+            pOb i = (p i) .S-ob
+
+            p'Ob : x  y
+            p'Ob i = (p' i) .S-ob
+
+
+            pMor : PathP  i  C [ pOb i , c ]) f g
+            pMor i = (p i) .S-arr
+
+            p'Mor : PathP  i  C [ p'Ob i , c ]) f g
+            p'Mor i = (p' i) .S-arr
+
+            -- we first show that it's equivalent to use sIso first then extract, or to extract first than use pToIIso
+            extractCom : extractIso' (sIso .fun p)  pToIIso .fun pOb
+            extractCom = J  yg'   extractIso' (pathToIso )  pToIIso' {xf = xf} {yg'} .fun  i  ( i) .S-ob))
+                           (cong extractIso' (JRefl pToIFam' pToIBase')  sym (JRefl pToIFam pToIBase))
+                           p
+               where
+                 idx = C .id
+                 pToIFam =  z _  CatIso C x z)
+                 pToIBase = catiso (C .id) idx (C .⋆IdL idx) (C .⋆IdL idx)
+
+                 idxf = SliceCat .id
+                 pToIFam' =  z _  CatIso SliceCat xf z)
+                 pToIBase' = catiso (SliceCat .id) idxf (SliceCat .⋆IdL idxf) (SliceCat .⋆IdL idxf)
+
+            -- why does this not follow definitionally?
+            -- from extractCom, we get that performing the roundtrip on pOb gives us back p'Ob
+            ppp : p'Ob  (pToIIso .inv) (pToIIso .fun pOb)
+            ppp = cong (pToIIso .inv) extractCom
+
+            -- apply univalence of C
+            -- this gives us the first component that we want
+            p'Ob≡pOb : p'Ob  pOb
+            p'Ob≡pOb = ppp  pToIIso .leftInv pOb
+
+            -- isSetHom gives us the second component, path between morphisms
+            p'Mor≡pMor : PathP  j  PathP  i  C [ (p'Ob≡pOb j) i , c ]) f g) p'Mor pMor
+            p'Mor≡pMor = isSetHomP2l {C = C} _ _ p'Mor pMor p'Ob≡pOb
+
+            -- we can use the above paths to show that p' ≡ p
+            p'≡p : p'  p
+            p'≡p i = comp  i'  SOPath≡PathΣ {xf = xf} {yg} (~ i'))
+                             j  λ { (i = i0)  left (~ j) ; (i = i1)  right (~ j) })
+                            (p'Σ≡pΣ i)
+              where
+                -- we break up p' and p into their constituent paths
+                -- first via transport and then via our component definitions from before
+                -- we show that p'ΣT ≡ p'Σ (and same for p) via univalence
+                -- and p'Σ≡pΣ follows from our work from above
+                p'ΣT : Σ[ p  x  y ] PathP  i  C [ p i , c ]) f g
+                p'ΣT = transport SOPath≡PathΣ p'
+                p'Σ : Σ[ p  x  y ] PathP  i  C [ p i , c ]) f g
+                p'Σ = (p'Ob , p'Mor)
+
+                pΣT : Σ[ p  x  y ] PathP  i  C [ p i , c ]) f g
+                pΣT = transport SOPath≡PathΣ p
+                 : Σ[ p  x  y ] PathP  i  C [ p i , c ]) f g
+                 = (pOb , pMor)-- transport SOPathP≡PathPSO p
+
+                -- using the computation rule to ua
+                p'ΣT≡p'Σ : p'ΣT  p'Σ
+                p'ΣT≡p'Σ = uaβ SOPath≃PathΣ p'
+
+                pΣT≡pΣ : pΣT  
+                pΣT≡pΣ = uaβ SOPath≃PathΣ p
+
+                p'Σ≡pΣ : p'Σ  
+                p'Σ≡pΣ = ΣPathP (p'Ob≡pOb , p'Mor≡pMor)
+
+                -- two sides of the square we're connecting
+                left : PathP  i  SOPath≡PathΣ {xf = xf} {yg} i) p' p'Σ
+                left = transport-filler SOPath≡PathΣ p'  p'ΣT≡p'Σ
+
+                right : PathP  i  SOPath≡PathΣ {xf = xf} {yg} i) p 
+                right = transport-filler SOPath≡PathΣ p  pΣT≡pΣ
+
+-- properties
+-- TODO: move to own file
+
+open isIsoC renaming (inv to invC)
+
+-- make a slice isomorphism from just the hom
+sliceIso :  {a b} (f : C [ a .S-ob , b .S-ob ]) (c : (f ⋆⟨ C  b .S-arr)  a .S-arr)
+          isIsoC C f
+          isIsoC SliceCat (slicehom f c)
+sliceIso f c isof .invC = slicehom (isof .invC) (sym (invMoveL (isIso→areInv isof) c))
+sliceIso f c isof .sec = SliceHom-≡-intro' (isof .sec)
+sliceIso f c isof .ret = SliceHom-≡-intro' (isof .ret)
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Constructions.Slice.Properties.html b/docs/Cubical.Categories.Constructions.Slice.Properties.html new file mode 100644 index 0000000..994ee10 --- /dev/null +++ b/docs/Cubical.Categories.Constructions.Slice.Properties.html @@ -0,0 +1,66 @@ + +Cubical.Categories.Constructions.Slice.Properties
{-# OPTIONS --safe #-}
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+
+open import Cubical.Data.Sigma
+
+open import Cubical.HITs.PropositionalTruncation using (∣_∣₁)
+
+open import Cubical.Categories.Category
+open import Cubical.Categories.Constructions.Slice.Base
+import Cubical.Categories.Constructions.Elements as Elements
+open import Cubical.Categories.Equivalence
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Instances.Sets
+open import Cubical.Categories.NaturalTransformation
+
+open Category
+
+module Cubical.Categories.Constructions.Slice.Properties
+  {ℓC ℓ'C : Level}
+  (C : Category ℓC ℓ'C)
+  (c : C .ob)
+  where
+
+open Elements.Contravariant {C = C}
+
+open _≃ᶜ_
+open Functor
+open WeakInverse
+
+slice→el : Functor (SliceCat C c) (∫ᴾ (C [-, c ]))
+slice→el .F-ob s = s .S-ob , s .S-arr
+slice→el .F-hom f = f .S-hom , f .S-comm
+slice→el .F-id = ΣPathP (refl , (isOfHLevelPath' 1 (C .isSetHom) _ _ _ _))
+slice→el .F-seq _ _ = ΣPathP (refl , (isOfHLevelPath' 1 (C .isSetHom) _ _ _ _))
+
+el→slice : Functor (∫ᴾ (C [-, c ])) (SliceCat C c)
+el→slice .F-ob (_ , s) = sliceob s
+el→slice .F-hom (f , comm) = slicehom f comm
+el→slice .F-id = SliceHom-≡-intro C c refl (isOfHLevelPath' 1 (C .isSetHom) _ _ _ _)
+el→slice .F-seq _ _ = SliceHom-≡-intro C c refl (isOfHLevelPath' 1 (C .isSetHom) _ _ _ _)
+
+open NatTrans
+open NatIso
+
+sliceIsElementsOfRep : SliceCat C c ≃ᶜ (∫ᴾ (C [-, c ]))
+sliceIsElementsOfRep .func = slice→el
+sliceIsElementsOfRep .isEquiv  =  w-inv ∣₁
+  where
+    w-inv : WeakInverse slice→el
+    w-inv .invFunc = el→slice
+    w-inv .η .trans .N-ob s = SliceCat C c .id
+    w-inv .η .trans .N-hom f = (SliceCat C c .⋆IdR f)
+                              sym (SliceCat C c .⋆IdL f)
+    w-inv .η .nIso x = isiso (SliceCat C c .id)
+                             (SliceCat C c .⋆IdR _)
+                             (SliceCat C c .⋆IdR _)
+    w-inv .ε .trans .N-ob s = (∫ᴾ (C [-, c ])) .id
+    w-inv .ε .trans .N-hom f = ((∫ᴾ (C [-, c ])) .⋆IdR f)
+                              sym ((∫ᴾ (C [-, c ])) .⋆IdL f)
+    w-inv .ε .nIso x = isiso ((∫ᴾ (C [-, c ])) .id)
+                             ((∫ᴾ (C [-, c ])) .⋆IdR _)
+                             ((∫ᴾ (C [-, c ])) .⋆IdR _)
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Constructions.Slice.html b/docs/Cubical.Categories.Constructions.Slice.html new file mode 100644 index 0000000..d1aed79 --- /dev/null +++ b/docs/Cubical.Categories.Constructions.Slice.html @@ -0,0 +1,18 @@ + +Cubical.Categories.Constructions.Slice
{-# OPTIONS --safe #-}
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Categories.Category
+
+open Category
+
+module Cubical.Categories.Constructions.Slice
+  { ℓ' : Level}
+  (C : Category  ℓ')
+  (c : C .ob)
+  where
+
+open import Cubical.Categories.Constructions.Slice.Base C c public
+open import Cubical.Categories.Constructions.Slice.Properties C c public
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Constructions.SubObject.html b/docs/Cubical.Categories.Constructions.SubObject.html new file mode 100644 index 0000000..6cec691 --- /dev/null +++ b/docs/Cubical.Categories.Constructions.SubObject.html @@ -0,0 +1,93 @@ + +Cubical.Categories.Constructions.SubObject
{-# OPTIONS --safe #-}
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Univalence
+
+open import Cubical.Data.Sigma
+
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Morphism
+
+open import Cubical.Relation.Binary.Order.Preorder
+
+open Category
+
+module Cubical.Categories.Constructions.SubObject
+  { ℓ' : Level}
+  (C : Category  ℓ')
+  (c : C .ob)
+  where
+
+open import Cubical.Categories.Constructions.Slice C c
+
+isSubObj :  (SliceCat .ob)
+isSubObj (sliceob arr) = isMonic C arr , isPropIsMonic C arr
+
+SubObjCat : Category (ℓ-max  ℓ') ℓ'
+SubObjCat = ΣPropCat SliceCat isSubObj
+
+SubObjCat→SliceCat : Functor SubObjCat SliceCat
+SubObjCat→SliceCat = forgetΣPropCat SliceCat isSubObj
+
+subObjMorIsMonic : {s t : SubObjCat .ob} (f : SubObjCat [ s , t ])
+   isMonic C (S-hom f)
+subObjMorIsMonic {s = s} {t = t} f = postcompCreatesMonic C
+  (S-hom f) (S-arr (t .fst)) comp-is-monic
+  where comp-is-monic = subst (isMonic C) (sym (S-comm f)) (s .snd)
+
+isPropSubObjMor : (s t : SubObjCat .ob)  isProp (SubObjCat [ s , t ])
+SliceHom.S-hom
+  (isPropSubObjMor
+    (sliceob incS , isMonicIncS)
+    (sliceob incT , isMonicIncT)
+    (slicehom x xComm)
+    (slicehom y yComm) i) =
+    isMonicIncT {a = x} {a' = y} (xComm  sym  yComm) i
+SliceHom.S-comm
+  (isPropSubObjMor
+    (sliceob incS , isMonicIncS)
+    (sliceob incT , isMonicIncT)
+    (slicehom x xComm)
+    (slicehom y yComm) i) =
+    isProp→PathP  i  C .isSetHom (seq' C (isMonicIncT (xComm  sym yComm) i) incT) incS) xComm yComm i
+
+isReflSubObjMor : (x : SubObjCat .ob)  SubObjCat [ x , x ]
+SliceHom.S-hom (isReflSubObjMor (sliceob inc , isMonicInc)) = C .id
+SliceHom.S-comm (isReflSubObjMor (sliceob inc , isMonicInc)) = C .⋆IdL inc
+
+isTransSubObjMor : (x y z : SubObjCat .ob)  SubObjCat [ x , y ]  SubObjCat [ y , z ]  SubObjCat [ x , z ]
+SliceHom.S-hom
+  (isTransSubObjMor
+    (sliceob incX , isMonicIncX)
+    (sliceob incY , isMonicIncY)
+    (sliceob incZ , isMonicIncZ)
+    (slicehom f fComm)
+    (slicehom g gComm)) = seq' C f g
+SliceHom.S-comm
+  (isTransSubObjMor
+    (sliceob incX , isMonicIncX)
+    (sliceob incY , isMonicIncY)
+    (sliceob incZ , isMonicIncZ)
+    (slicehom f fComm)
+    (slicehom g gComm)) =
+  seq' C (seq' C f g) incZ
+    ≡⟨ C .⋆Assoc f g incZ 
+  seq' C f (seq' C g incZ)
+    ≡⟨ cong  x  seq' C f x) gComm 
+  seq' C f incY
+    ≡⟨ fComm 
+  incX
+    
+
+module _ (isSetCOb : isSet (C .ob)) where
+  subObjCatPreorderStr : PreorderStr _ (SubObjCat .ob)
+  PreorderStr._≲_ subObjCatPreorderStr x y = SubObjCat [ x , y ]
+  IsPreorder.is-set (PreorderStr.isPreorder subObjCatPreorderStr) = isSetΣ (isSetSliceOb isSetCOb) λ x  isProp→isSet (∈-isProp isSubObj x)
+  IsPreorder.is-prop-valued (PreorderStr.isPreorder subObjCatPreorderStr) = isPropSubObjMor
+  IsPreorder.is-refl (PreorderStr.isPreorder subObjCatPreorderStr) = isReflSubObjMor
+  IsPreorder.is-trans (PreorderStr.isPreorder subObjCatPreorderStr) = isTransSubObjMor
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Equivalence.Base.html b/docs/Cubical.Categories.Equivalence.Base.html new file mode 100644 index 0000000..c8e799e --- /dev/null +++ b/docs/Cubical.Categories.Equivalence.Base.html @@ -0,0 +1,41 @@ + +Cubical.Categories.Equivalence.Base
{-# OPTIONS --safe #-}
+module Cubical.Categories.Equivalence.Base where
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.HITs.PropositionalTruncation
+
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.NaturalTransformation
+
+open Category
+open Functor
+
+private
+  variable
+    ℓC ℓC' ℓD ℓD' : Level
+
+record WeakInverse {C : Category ℓC ℓC'} {D : Category ℓD ℓD'}
+                     (func : Functor C D) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where
+  field
+    invFunc : Functor D C
+
+    η : 𝟙⟨ C  ≅ᶜ invFunc ∘F func
+    ε : func ∘F invFunc ≅ᶜ 𝟙⟨ D 
+
+-- I don't know of a good alternative representation of isEquivalence that
+-- avoids truncation in the general case.  If the categories are univalent, then
+-- an adjoint equivalence might be enough.
+isEquivalence : {C : Category ℓC ℓC'} {D : Category ℓD ℓD'}
+               (func : Functor C D)  Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
+isEquivalence func =  WeakInverse func ∥₁
+
+record _≃ᶜ_ (C : Category ℓC ℓC') (D : Category ℓD ℓD') :
+               Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where
+  constructor equivᶜ
+  field
+    func : Functor C D
+    isEquiv : isEquivalence func
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Equivalence.Properties.html b/docs/Cubical.Categories.Equivalence.Properties.html new file mode 100644 index 0000000..66f319a --- /dev/null +++ b/docs/Cubical.Categories.Equivalence.Properties.html @@ -0,0 +1,203 @@ + +Cubical.Categories.Equivalence.Properties
{-# OPTIONS --safe #-}
+
+module Cubical.Categories.Equivalence.Properties where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Equiv
+  renaming (isEquiv to isEquivMap)
+open import Cubical.Foundations.Equiv.Dependent
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Powerset
+open import Cubical.Data.Sigma
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.NaturalTransformation
+open import Cubical.Categories.Morphism
+open import Cubical.Categories.Equivalence.Base
+open import Cubical.HITs.PropositionalTruncation
+
+open Category
+open Functor
+open NatIso
+open isIso
+open WeakInverse
+
+private
+  variable
+    ℓC ℓC' ℓD ℓD' : Level
+
+-- Equivalence implies Full, Faithul, and Essentially Surjective
+
+module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where
+  symWeakInverse :  {F : Functor C D}  (e : WeakInverse F)  WeakInverse (e .invFunc)
+  symWeakInverse {F} record { invFunc = G ; η = η ; ε = ε } = record { invFunc = F ; η = symNatIso ε ; ε = symNatIso η }
+
+  isEquiv→Faithful :  {F : Functor C D}  isEquivalence F  isFaithful F
+  isEquiv→Faithful {F} = rec (isPropΠ5  _ _ _ _ _  C .isSetHom _ _)) lifted
+    where
+      lifted : WeakInverse F  isFaithful F
+      lifted record { invFunc = G
+                              ; η = η
+                              ; ε = _ }
+                   c c' f g p = f
+                              ≡⟨ sqRL η 
+                                cIso .fst ⋆⟨ C  G  F  f   ⋆⟨ C  c'Iso .snd .inv
+                              ≡⟨ cong  v  cIso .fst ⋆⟨ C  (G  v ) ⋆⟨ C  c'Iso .snd .inv) p 
+                                cIso .fst ⋆⟨ C  G  F  g   ⋆⟨ C  c'Iso .snd .inv
+                              ≡⟨ sym (sqRL η) 
+                                g
+                              
+         where
+           -- isomorphism between c and GFc
+          cIso = isIso→CatIso (η .nIso c)
+          -- isomorphism between c' and GFc'
+          c'Iso = isIso→CatIso (η .nIso c')
+
+module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where
+  isEquiv→Full :  {F : Functor C D}  isEquivalence F  isFull F
+  isEquiv→Full {F} = rec (isPropΠ3  _ _ _  isPropPropTrunc)) lifted
+    where
+      lifted : WeakInverse F  isFull F
+      lifted eq@record { invFunc = G
+                             ; η = η
+                             ; ε = _ }
+        c c' g =  h , isEquiv→Faithful  symWeakInverse eq ∣₁ _ _ _ _ GFh≡Gg ∣₁ -- apply faithfulness of G
+        where
+          -- isomorphism between c and GFc
+          cIso = isIso→CatIso (η .nIso c)
+          -- isomorphism between c' and GFc'
+          c'Iso = isIso→CatIso (η .nIso c')
+
+          -- reverses
+          cIso⁻ = symCatIso cIso
+          c'Iso⁻ = symCatIso c'Iso
+
+          h = cIso .fst ⋆⟨ C  G  g  ⋆⟨ C  c'Iso .snd .inv
+
+          -- we show that both `G ⟪ g ⟫` and `G ⟪ F ⟪ h ⟫ ⟫`
+          -- are equal to the same thing
+          -- namely : cIso .inv ⋆⟨ C ⟩ h ⋆⟨ C ⟩ c'Iso .mor
+          Gg≡ηhη : G  g   cIso .snd .inv ⋆⟨ C  h ⋆⟨ C  c'Iso .fst
+          Gg≡ηhη = invMoveL cAreInv move-c'  sym (C .⋆Assoc _ _ _)
+            where
+              cAreInv : areInv _ (cIso .fst) (cIso .snd .inv)
+              cAreInv = CatIso→areInv cIso
+
+              c'AreInv : areInv _ (c'Iso .fst) (c'Iso .snd .inv)
+              c'AreInv = CatIso→areInv c'Iso
+
+              move-c' : cIso .fst ⋆⟨ C  G  g   h ⋆⟨ C  c'Iso .fst
+              move-c' = invMoveR (symAreInv c'AreInv) refl
+
+          GFh≡Gg : G  F  h    G  g 
+          GFh≡Gg = G  F  h  
+                 ≡⟨ sqLR η 
+                   cIso .snd .inv ⋆⟨ C  h ⋆⟨ C  c'Iso .fst
+                 ≡⟨ sym Gg≡ηhη 
+                   G  g 
+                 
+
+  isEquiv→FullyFaithful :   {F : Functor C D}  isEquivalence F  isFullyFaithful F
+  isEquiv→FullyFaithful {F = F} h = isFull+Faithful→isFullyFaithful {F = F} (isEquiv→Full h) (isEquiv→Faithful h)
+
+  isEquiv→Surj :  {F : Functor C D}  isEquivalence F  isEssentiallySurj F
+  isEquiv→Surj = rec (isPropΠ  _  isPropPropTrunc))
+     wkInv d   (wkInv .invFunc  d ) , isIso→CatIso ((wkInv .ε .nIso) d) ∣₁)
+
+
+-- A fully-faithful functor that induces equivalence on objects is an equivalence
+
+module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'}
+  {F : Functor C D} where
+
+  isFullyFaithful+isEquivF-ob→isEquiv : isFullyFaithful F  isEquivMap (F .F-ob)  isEquivalence F
+  isFullyFaithful+isEquivF-ob→isEquiv fullfaith isequiv =  w ∣₁
+    where
+    open Iso
+    open IsoOver
+
+    MorC : C .ob × C .ob  Type _
+    MorC (x , y) = C [ x , y ]
+
+    MorD : D .ob × D .ob  Type _
+    MorD (x , y) = D [ x , y ]
+
+    F-Mor : ((x , y) : C .ob × C .ob)  C [ x , y ]  D [ F .F-ob x , F .F-ob y ]
+    F-Mor _ = F .F-hom
+
+    equiv-ob² : C .ob × C .ob  D .ob × D .ob
+    equiv-ob² = ≃-× (_ , isequiv) (_ , isequiv)
+
+    iso-ob  = equivToIso (_ , isequiv)
+    iso-hom = equivOver→IsoOver {P = MorC} {Q = MorD} equiv-ob² F-Mor  (x , y)  fullfaith x y)
+
+    w-inv : Functor D C
+    w-inv .F-ob = iso-ob .inv
+    w-inv .F-hom = iso-hom .inv _
+    w-inv .F-id {x = x} = isFullyFaithful→Faithful {F = F} fullfaith _ _ _ _ (p  sym (F .F-id))
+      where
+      p : _
+      p i =
+        comp
+         j  D [ iso-ob .rightInv x (~ j) , iso-ob .rightInv x (~ j) ])
+         j  λ
+          { (i = i0)  iso-hom .rightInv _ (D .id {x = x}) (~ j)
+          ; (i = i1)  D .id {x = iso-ob .rightInv x (~ j)} })
+        (D .id {x = x})
+    w-inv .F-seq {x = x} {z = z} f g = isFullyFaithful→Faithful {F = F} fullfaith _ _ _ _ (p  sym (F .F-seq _ _))
+      where
+      p : _
+      p i =
+        comp
+         j  D [ iso-ob .rightInv x (~ j) , iso-ob .rightInv z (~ j) ])
+         j  λ
+          { (i = i0)  iso-hom .rightInv _ (f ⋆⟨ D  g) (~ j)
+          ; (i = i1)  iso-hom .rightInv _ f (~ j) ⋆⟨ D  iso-hom .rightInv _ g (~ j) })
+        (f ⋆⟨ D  g)
+
+    w-η-path : 𝟙⟨ C   w-inv ∘F F
+    w-η-path = Functor≡  x  sym (retIsEq isequiv x))  {x} {y} f   i  iso-hom .leftInv (x , y) f (~ i)))
+
+    w-ε-path : F ∘F w-inv  𝟙⟨ D 
+    w-ε-path = Functor≡  x  secIsEq isequiv x)  {x} {y} f i  iso-hom .rightInv (x , y) f i)
+
+    w : WeakInverse F
+    w .invFunc = w-inv
+    w .η = pathToNatIso w-η-path
+    w .ε = pathToNatIso w-ε-path
+
+
+
+-- equivalence on full subcategories defined by propositions
+module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) (invF : WeakInverse F) where
+
+  open NatTrans
+  open _≃ᶜ_
+
+  private
+    F⁻¹ = invF .invFunc
+    ηᴱ = invF .η
+    εᴱ = invF .ε
+
+
+  ΣPropCatEquiv : {P :  (ob C)} {Q :  (ob D)}
+                 (presF :  c  c  P  F .F-ob c  Q)
+                 (∀ d  d  Q  F⁻¹ .F-ob d  P)
+                 WeakInverse (ΣPropCatFunc {P = P} {Q = Q} F presF)
+
+  invFunc (ΣPropCatEquiv {P} {Q} _ presF⁻¹) = ΣPropCatFunc {P = Q} {Q = P} F⁻¹ presF⁻¹
+
+  N-ob (trans (η (ΣPropCatEquiv _ _))) (x , _) = ηᴱ .trans .N-ob x
+  N-hom (trans (η (ΣPropCatEquiv _ _))) f = ηᴱ .trans .N-hom f
+  inv (nIso (η (ΣPropCatEquiv _ _)) (x , _)) = ηᴱ .nIso x .inv
+  sec (nIso (η (ΣPropCatEquiv _ _)) (x , _)) = ηᴱ .nIso x .sec
+  ret (nIso (η (ΣPropCatEquiv _ _)) (x , _)) = ηᴱ .nIso x .ret
+
+  N-ob (trans (ε (ΣPropCatEquiv _ _))) (x , _) = εᴱ .trans .N-ob x
+  N-hom (trans (ε (ΣPropCatEquiv _ _))) f = εᴱ .trans .N-hom f
+  inv (nIso (ε (ΣPropCatEquiv _ _)) (x , _)) = εᴱ .nIso x .inv
+  sec (nIso (ε (ΣPropCatEquiv _ _)) (x , _)) = εᴱ .nIso x .sec
+  ret (nIso (ε (ΣPropCatEquiv _ _)) (x , _)) = εᴱ .nIso x .ret
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Equivalence.html b/docs/Cubical.Categories.Equivalence.html new file mode 100644 index 0000000..35b86d2 --- /dev/null +++ b/docs/Cubical.Categories.Equivalence.html @@ -0,0 +1,9 @@ + +Cubical.Categories.Equivalence
+{-# OPTIONS --safe #-}
+
+module Cubical.Categories.Equivalence where
+
+open import Cubical.Categories.Equivalence.Base public
+open import Cubical.Categories.Equivalence.Properties public
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Functor.Base.html b/docs/Cubical.Categories.Functor.Base.html index 9606cd6..0e922bb 100644 --- a/docs/Cubical.Categories.Functor.Base.html +++ b/docs/Cubical.Categories.Functor.Base.html @@ -28,25 +28,25 @@ F-hom : {x y : C .ob} C [ x , y ] D [ F-ob x , F-ob y ] F-id : {x : C .ob} F-hom (C .id) D .id {x = F-ob x} F-seq : {x y z : C .ob} (f : C [ x , y ]) (g : C [ y , z ]) - F-hom (f ⋆⟨ C g) (F-hom f) ⋆⟨ D (F-hom g) + F-hom (f ⋆⟨ C g) (F-hom f) ⋆⟨ D (F-hom g) isFull = (x y : _) (F[f] : D [ F-ob x , F-ob y ]) ∃[ f C [ x , y ] ] F-hom f F[f] isFaithful = (x y : _) (f g : C [ x , y ]) F-hom f F-hom g f g isFullyFaithful = (x y : _) isEquiv (F-hom {x = x} {y = y}) - isEssentiallySurj = (d : D .ob) ∃[ c C .ob ] CatIso D (F-ob c) d + isEssentiallySurj = (d : D .ob) ∃[ c C .ob ] CatIso D (F-ob c) d -- preservation of commuting squares and triangles F-square : {x y u v : C .ob} {f : C [ x , y ]} {g : C [ x , u ]} {h : C [ u , v ]} {k : C [ y , v ]} - f ⋆⟨ C k g ⋆⟨ C h - (F-hom f) ⋆⟨ D (F-hom k) (F-hom g) ⋆⟨ D (F-hom h) + f ⋆⟨ C k g ⋆⟨ C h + (F-hom f) ⋆⟨ D (F-hom k) (F-hom g) ⋆⟨ D (F-hom h) F-square Csquare = sym (F-seq _ _) ∙∙ cong F-hom Csquare ∙∙ F-seq _ _ F-triangle : {x y z : C .ob} {f : C [ x , y ]} {g : C [ y , z ]} {h : C [ x , z ]} - f ⋆⟨ C g h - (F-hom f) ⋆⟨ D (F-hom g) (F-hom h) + f ⋆⟨ C g h + (F-hom f) ⋆⟨ D (F-hom g) (F-hom h) F-triangle Ctriangle = sym (F-seq _ _) cong F-hom Ctriangle private @@ -64,16 +64,16 @@ F-ob (Functor≡ hOb hHom i) c = hOb c i F-hom (Functor≡ hOb hHom i) f = hHom f i F-id (Functor≡ {C = C} {D = D} {F = F} {G = G} hOb hHom i) = - isProp→PathP j isSetHom D (hHom (C .id) j) (D .id)) (F-id F) (F-id G) i + isProp→PathP j isSetHom D (hHom (C .id) j) (D .id)) (F-id F) (F-id G) i F-seq (Functor≡ {C = C} {D = D} {F = F} {G = G} hOb hHom i) f g = - isProp→PathP j isSetHom D (hHom (f ⋆⟨ C g) j) ((hHom f j) ⋆⟨ D (hHom g j))) (F-seq F f g) (F-seq G f g) i + isProp→PathP j isSetHom D (hHom (f ⋆⟨ C g) j) ((hHom f j) ⋆⟨ D (hHom g j))) (F-seq F f g) (F-seq G f g) i FunctorSquare : {F₀₀ F₀₁ F₁₀ F₁₁ : Functor C D} (F₀₋ : F₀₀ F₀₁) (F₁₋ : F₁₀ F₁₁) (F₋₀ : F₀₀ F₁₀) (F₋₁ : F₀₁ F₁₁) - Square (cong F-ob F₀₋) (cong F-ob F₁₋) (cong F-ob F₋₀) (cong F-ob F₋₁) - Square F₀₋ F₁₋ F₋₀ F₋₁ + Square (cong F-ob F₀₋) (cong F-ob F₁₋) (cong F-ob F₋₀) (cong F-ob F₋₁) + Square F₀₋ F₁₋ F₋₀ F₋₁ FunctorSquare {C = C} {D = D} F₀₋ F₁₋ F₋₀ F₋₁ r = sqr where sqr : _ @@ -82,11 +82,11 @@ isSet→SquareP i j D .isSetHom {x = sqr i j .F-ob x} {y = sqr i j .F-ob y}) i F₀₋ i .F-hom f) i F₁₋ i .F-hom f) i F₋₀ i .F-hom f) i F₋₁ i .F-hom f) i j sqr i j .F-id {x = x} = - isSet→SquareP i j isProp→isSet (D .isSetHom (sqr i j .F-hom (C .id)) (D .id))) + isSet→SquareP i j isProp→isSet (D .isSetHom (sqr i j .F-hom (C .id)) (D .id))) i F₀₋ i .F-id) i F₁₋ i .F-id) i F₋₀ i .F-id) i F₋₁ i .F-id) i j sqr i j .F-seq f g = isSet→SquareP i j - isProp→isSet (D .isSetHom (sqr i j .F-hom (f ⋆⟨ C g)) ((sqr i j .F-hom f) ⋆⟨ D (sqr i j .F-hom g)))) + isProp→isSet (D .isSetHom (sqr i j .F-hom (f ⋆⟨ C g)) ((sqr i j .F-hom f) ⋆⟨ D (sqr i j .F-hom g)))) i F₀₋ i .F-seq f g) i F₁₋ i .F-seq f g) i F₋₀ i .F-seq f g) i F₋₁ i .F-seq f g) i j FunctorPath≡ : {F G : Functor C D}{p q : F G} cong F-ob p cong F-ob q p q @@ -122,7 +122,7 @@ Id : {C : Category ℓ'} Functor C C Id = 𝟙⟨ _ -forgetΣPropCat : (C : Category ℓ') (prop : (C .ob)) Functor (ΣPropCat C prop) C +forgetΣPropCat : (C : Category ℓ') (prop : (C .ob)) Functor (ΣPropCat C prop) C forgetΣPropCat _ _ .F-ob x = x .fst forgetΣPropCat _ _ .F-hom f = f forgetΣPropCat _ _ .F-id = refl @@ -145,14 +145,14 @@ funcCompOb≡ G F c = refl -_^opF : Functor C D Functor (C ^op) (D ^op) +_^opF : Functor C D Functor (C ^op) (D ^op) (F ^opF) .F-ob = F .F-ob (F ^opF) .F-hom = F .F-hom (F ^opF) .F-id = F .F-id (F ^opF) .F-seq f g = F .F-seq g f open Iso -Iso^opF : Iso (Functor C D) (Functor (C ^op) (D ^op)) +Iso^opF : Iso (Functor C D) (Functor (C ^op) (D ^op)) fun Iso^opF = _^opF inv Iso^opF = _^opF F-ob (rightInv Iso^opF b i) = F-ob b @@ -164,13 +164,13 @@ F-id (leftInv Iso^opF a i) = F-id a F-seq (leftInv Iso^opF a i) = F-seq a -^opFEquiv : Functor C D Functor (C ^op) (D ^op) +^opFEquiv : Functor C D Functor (C ^op) (D ^op) ^opFEquiv = isoToEquiv Iso^opF -- Functoriality on full subcategories defined by propositions ΣPropCatFunc : {P : (ob C)} {Q : (ob D)} (F : Functor C D) (∀ c c P F .F-ob c Q) - Functor (ΣPropCat C P) (ΣPropCat D Q) + Functor (ΣPropCat C P) (ΣPropCat D Q) F-ob (ΣPropCatFunc F FPres) (c , c∈P) = F .F-ob c , FPres c c∈P F-hom (ΣPropCatFunc F FPres) = F .F-hom F-id (ΣPropCatFunc F FPres) = F .F-id diff --git a/docs/Cubical.Categories.Functor.Compose.html b/docs/Cubical.Categories.Functor.Compose.html index 617cf53..100f6bd 100644 --- a/docs/Cubical.Categories.Functor.Compose.html +++ b/docs/Cubical.Categories.Functor.Compose.html @@ -1,46 +1,46 @@ -Cubical.Categories.Functor.Compose
{-# OPTIONS --allow-unsolved-metas #-}
-
-module Cubical.Categories.Functor.Compose where
-
-open import Cubical.Foundations.Prelude
-
-open import Cubical.Categories.Category
-open import Cubical.Categories.Functor.Base
-open import Cubical.Categories.NaturalTransformation.Base
-
-open import Cubical.Categories.Instances.Functors
-
-module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'}
-  {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (E : Category ℓE ℓE')
-  (F : Functor C D)
-  where
-
-  open Functor
-  open NatTrans
-
-  precomposeF : Functor (FUNCTOR D E) (FUNCTOR C E)
-  precomposeF .F-ob G = funcComp G F
-  precomposeF .F-hom α .N-ob c = α .N-ob (F .F-ob c)
-  precomposeF .F-hom α .N-hom f = α .N-hom (F .F-hom f)
-  precomposeF .F-id = refl
-  precomposeF .F-seq f g = refl
-
-module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'}
-  (C : Category ℓC ℓC') {D : Category ℓD ℓD'} {E : Category ℓE ℓE'}
-  (G : Functor D E)
-  where
-
-  open Functor
-  open NatTrans
-
-  postcomposeF : Functor (FUNCTOR C D) (FUNCTOR C E)
-  postcomposeF .F-ob F = funcComp G F
-  postcomposeF .F-hom α .N-ob c = G. F-hom (α .N-ob c)
-  postcomposeF .F-hom {F₀} {F₁} α .N-hom {x} {y} f =
-    sym (G .F-seq (F₀  f ) (α  y ))
-    ∙∙ cong (G ⟪_⟫) (α .N-hom f)
-    ∙∙ G .F-seq (α  x ) (F₁  f )
-  postcomposeF .F-id = makeNatTransPath (funExt λ _  G .F-id)
-  postcomposeF .F-seq f g = makeNatTransPath (funExt λ _  G .F-seq _ _)
+Cubical.Categories.Functor.Compose
{-# OPTIONS --safe #-}
+
+module Cubical.Categories.Functor.Compose where
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor.Base
+open import Cubical.Categories.NaturalTransformation.Base
+
+open import Cubical.Categories.Instances.Functors
+
+module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'}
+  {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (E : Category ℓE ℓE')
+  (F : Functor C D)
+  where
+
+  open Functor
+  open NatTrans
+
+  precomposeF : Functor (FUNCTOR D E) (FUNCTOR C E)
+  precomposeF .F-ob G = funcComp G F
+  precomposeF .F-hom α .N-ob c = α .N-ob (F .F-ob c)
+  precomposeF .F-hom α .N-hom f = α .N-hom (F .F-hom f)
+  precomposeF .F-id = refl
+  precomposeF .F-seq f g = refl
+
+module _ {ℓC ℓC' ℓD ℓD' ℓE ℓE'}
+  (C : Category ℓC ℓC') {D : Category ℓD ℓD'} {E : Category ℓE ℓE'}
+  (G : Functor D E)
+  where
+
+  open Functor
+  open NatTrans
+
+  postcomposeF : Functor (FUNCTOR C D) (FUNCTOR C E)
+  postcomposeF .F-ob F = funcComp G F
+  postcomposeF .F-hom α .N-ob c = G. F-hom (α .N-ob c)
+  postcomposeF .F-hom {F₀} {F₁} α .N-hom {x} {y} f =
+    sym (G .F-seq (F₀  f ) (α  y ))
+    ∙∙ cong (G ⟪_⟫) (α .N-hom f)
+    ∙∙ G .F-seq (α  x ) (F₁  f )
+  postcomposeF .F-id = makeNatTransPath (funExt λ _  G .F-id)
+  postcomposeF .F-seq f g = makeNatTransPath (funExt λ _  G .F-seq _ _)
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Functor.Properties.html b/docs/Cubical.Categories.Functor.Properties.html index 5cd7e6e..875aae4 100644 --- a/docs/Cubical.Categories.Functor.Properties.html +++ b/docs/Cubical.Categories.Functor.Properties.html @@ -4,233 +4,264 @@ module Cubical.Categories.Functor.Properties where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Equiv.Properties -open import Cubical.Foundations.Function hiding (_∘_) -open import Cubical.Foundations.GroupoidLaws using (lUnit; rUnit; assoc; cong-∙) -open import Cubical.Foundations.HLevels -open import Cubical.Functions.Surjection -open import Cubical.Functions.Embedding -open import Cubical.HITs.PropositionalTruncation as Prop -open import Cubical.Data.Sigma -open import Cubical.Categories.Category -open import Cubical.Categories.Isomorphism -open import Cubical.Categories.Morphism -open import Cubical.Categories.Functor.Base - - -private - variable - ℓ' ℓ'' : Level - B C D E : Category ℓ' - -open Category -open Functor - -F-assoc : {F : Functor B C} {G : Functor C D} {H : Functor D E} - H ∘F (G ∘F F) (H ∘F G) ∘F F -F-assoc = Functor≡ _ refl) _ refl) - - --- Results about functors - -module _ {F : Functor C D} where - - -- the identity is the identity - F-lUnit : F ∘F 𝟙⟨ C F - F-lUnit i .F-ob x = F x - F-lUnit i .F-hom f = F f - F-lUnit i .F-id {x} = lUnit (F .F-id) (~ i) - F-lUnit i .F-seq f g = lUnit (F .F-seq f g) (~ i) - - F-rUnit : 𝟙⟨ D ∘F F F - F-rUnit i .F-ob x = F x - F-rUnit i .F-hom f = F f - F-rUnit i .F-id {x} = rUnit (F .F-id) (~ i) - F-rUnit i .F-seq f g = rUnit (F .F-seq f g) (~ i) - - -- functors preserve commutative diagrams (specificallysqures here) - preserveCommF : {x y z w} {f : C [ x , y ]} {g : C [ y , w ]} {h : C [ x , z ]} {k : C [ z , w ]} - f ⋆⟨ C g h ⋆⟨ C k - (F f ) ⋆⟨ D (F g ) (F h ) ⋆⟨ D (F k ) - preserveCommF {f = f} {g = g} {h = h} {k = k} eq - = (F f ) ⋆⟨ D (F g ) - ≡⟨ sym (F .F-seq _ _) - F f ⋆⟨ C g - ≡⟨ cong (F ⟪_⟫) eq - F h ⋆⟨ C k - ≡⟨ F .F-seq _ _ - (F h ) ⋆⟨ D (F k ) - - - -- functors preserve isomorphisms - preserveIsosF : {x y} CatIso C x y CatIso D (F x ) (F y ) - preserveIsosF {x} {y} (f , isiso f⁻¹ sec' ret') = - catiso - g g⁻¹ - -- sec - ( (g⁻¹ ⋆⟨ D g) - ≡⟨ sym (F .F-seq f⁻¹ f) - F f⁻¹ ⋆⟨ C f - ≡⟨ cong (F .F-hom) sec' - F C .id - ≡⟨ F .F-id - D .id - ) - -- ret - ( (g ⋆⟨ D g⁻¹) - ≡⟨ sym (F .F-seq f f⁻¹) - F f ⋆⟨ C f⁻¹ - ≡⟨ cong (F .F-hom) ret' - F C .id - ≡⟨ F .F-id - D .id - ) - - where - x' : D .ob - x' = F x - y' : D .ob - y' = F y - - g : D [ x' , y' ] - g = F f - g⁻¹ : D [ y' , x' ] - g⁻¹ = F f⁻¹ - - -- hacky lemma helping with type inferences - functorCongP : {x y v w : ob C} {p : x y} {q : v w} {f : C [ x , v ]} {g : C [ y , w ]} - PathP i C [ p i , q i ]) f g - PathP i D [ F .F-ob (p i) , F. F-ob (q i) ]) (F .F-hom f) (F .F-hom g) - functorCongP r i = F .F-hom (r i) - -isSetFunctor : isSet (D .ob) isSet (Functor C D) -isSetFunctor {D = D} {C = C} isSet-D-ob F G p q = w - where - w : _ - F-ob (w i i₁) = isSetΠ _ isSet-D-ob) _ _ (cong F-ob p) (cong F-ob q) i i₁ - F-hom (w i i₁) z = - isSet→SquareP - i i₁ D .isSetHom {(F-ob (w i i₁) _)} {(F-ob (w i i₁) _)}) - i₁ F-hom (p i₁) z) i₁ F-hom (q i₁) z) refl refl i i₁ - - F-id (w i i₁) = - isSet→SquareP - i i₁ isProp→isSet (D .isSetHom (F-hom (w i i₁) _) (D .id))) - i₁ F-id (p i₁)) i₁ F-id (q i₁)) refl refl i i₁ - - F-seq (w i i₁) _ _ = - isSet→SquareP - i i₁ isProp→isSet (D .isSetHom (F-hom (w i i₁) _) ((F-hom (w i i₁) _) ⋆⟨ D (F-hom (w i i₁) _)))) - i₁ F-seq (p i₁) _ _) i₁ F-seq (q i₁) _ _) refl refl i i₁ - - --- Conservative Functor, --- namely if a morphism f is mapped to an isomorphism, --- the morphism f is itself isomorphism. - -isConservative : (F : Functor C D) Type _ -isConservative {C = C} {D = D} F = {x y : C .ob}{f : C [ x , y ]} isIso D (F .F-hom f) isIso C f - - --- Fully-faithfulness of functors - -module _ {F : Functor C D} where - - isFullyFaithful→Full : isFullyFaithful F isFull F - isFullyFaithful→Full fullfaith x y = isEquiv→isSurjection (fullfaith x y) - - isFullyFaithful→Faithful : isFullyFaithful F isFaithful F - isFullyFaithful→Faithful fullfaith x y = isEmbedding→Inj (isEquiv→isEmbedding (fullfaith x y)) - - isFull+Faithful→isFullyFaithful : isFull F isFaithful F isFullyFaithful F - isFull+Faithful→isFullyFaithful full faith x y = isEmbedding×isSurjection→isEquiv - (injEmbedding (D .isSetHom) (faith x y _ _) , full x y) - - isFaithful→reflectsMono : isFaithful F {x y : C .ob} (f : C [ x , y ]) - isMonic D (F f ) isMonic C f - isFaithful→reflectsMono F-faithful f Ff-mon {a = a} {a' = a'} a⋆f≡a'⋆f = - let Fa⋆Ff≡Fa'⋆Ff = sym (F .F-seq a f) - cong (F ⟪_⟫) a⋆f≡a'⋆f - F .F-seq a' f - in F-faithful _ _ _ _ (Ff-mon Fa⋆Ff≡Fa'⋆Ff) - - - -- Fully-faithful functor is conservative - - open isIso - - isFullyFaithful→Conservative : isFullyFaithful F isConservative F - isFullyFaithful→Conservative fullfaith {x = x} {y = y} {f = f} isoFf = w - where - w : isIso C f - w .inv = invIsEq (fullfaith _ _) (isoFf .inv) - w .sec = isFullyFaithful→Faithful fullfaith _ _ _ _ - (F .F-seq _ _ - i secIsEq (fullfaith _ _) (isoFf .inv) i ⋆⟨ D F .F-hom f) - isoFf .sec - sym (F .F-id)) - w .ret = isFullyFaithful→Faithful fullfaith _ _ _ _ - (F .F-seq _ _ - i F .F-hom f ⋆⟨ D secIsEq (fullfaith _ _) (isoFf .inv) i) - isoFf .ret - sym (F .F-id)) +import Cubical.Foundations.Isomorphism as Iso +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Function hiding (_∘_) +open import Cubical.Foundations.GroupoidLaws using (lUnit; rUnit; assoc; cong-∙) +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Functions.Surjection +open import Cubical.Functions.Embedding +open import Cubical.HITs.PropositionalTruncation as Prop +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (_+_) +open import Cubical.Categories.Category +open import Cubical.Categories.Isomorphism +open import Cubical.Categories.Morphism +open import Cubical.Categories.Functor.Base + + +private + variable + ℓ' ℓ'' : Level + B C D E : Category ℓ' + +open Category +open Functor + +F-assoc : {F : Functor B C} {G : Functor C D} {H : Functor D E} + H ∘F (G ∘F F) (H ∘F G) ∘F F +F-assoc = Functor≡ _ refl) _ refl) + + +-- Results about functors + +module _ {F : Functor C D} where + + -- the identity is the identity + F-lUnit : F ∘F 𝟙⟨ C F + F-lUnit i .F-ob x = F x + F-lUnit i .F-hom f = F f + F-lUnit i .F-id {x} = lUnit (F .F-id) (~ i) + F-lUnit i .F-seq f g = lUnit (F .F-seq f g) (~ i) + + F-rUnit : 𝟙⟨ D ∘F F F + F-rUnit i .F-ob x = F x + F-rUnit i .F-hom f = F f + F-rUnit i .F-id {x} = rUnit (F .F-id) (~ i) + F-rUnit i .F-seq f g = rUnit (F .F-seq f g) (~ i) + + -- functors preserve commutative diagrams (specificallysqures here) + preserveCommF : {x y z w} {f : C [ x , y ]} {g : C [ y , w ]} {h : C [ x , z ]} {k : C [ z , w ]} + f ⋆⟨ C g h ⋆⟨ C k + (F f ) ⋆⟨ D (F g ) (F h ) ⋆⟨ D (F k ) + preserveCommF {f = f} {g = g} {h = h} {k = k} eq + = (F f ) ⋆⟨ D (F g ) + ≡⟨ sym (F .F-seq _ _) + F f ⋆⟨ C g + ≡⟨ cong (F ⟪_⟫) eq + F h ⋆⟨ C k + ≡⟨ F .F-seq _ _ + (F h ) ⋆⟨ D (F k ) + + + -- functors preserve isomorphisms + preserveIsosF : {x y} CatIso C x y CatIso D (F x ) (F y ) + preserveIsosF {x} {y} (f , isiso f⁻¹ sec' ret') = + catiso + g g⁻¹ + -- sec + ( (g⁻¹ ⋆⟨ D g) + ≡⟨ sym (F .F-seq f⁻¹ f) + F f⁻¹ ⋆⟨ C f + ≡⟨ cong (F .F-hom) sec' + F C .id + ≡⟨ F .F-id + D .id + ) + -- ret + ( (g ⋆⟨ D g⁻¹) + ≡⟨ sym (F .F-seq f f⁻¹) + F f ⋆⟨ C f⁻¹ + ≡⟨ cong (F .F-hom) ret' + F C .id + ≡⟨ F .F-id + D .id + ) + + where + x' : D .ob + x' = F x + y' : D .ob + y' = F y + + g : D [ x' , y' ] + g = F f + g⁻¹ : D [ y' , x' ] + g⁻¹ = F f⁻¹ + + -- hacky lemma helping with type inferences + functorCongP : {x y v w : ob C} {p : x y} {q : v w} {f : C [ x , v ]} {g : C [ y , w ]} + PathP i C [ p i , q i ]) f g + PathP i D [ F .F-ob (p i) , F. F-ob (q i) ]) (F .F-hom f) (F .F-hom g) + functorCongP r i = F .F-hom (r i) + +isEquivFunctor≡ : {F} {G} isEquiv (uncurry (Functor≡ {C = C} {D = D} {F = F} {G = G})) +isEquivFunctor≡ {C = C} {D = D} = Iso.isoToIsEquiv isom + where + open Iso.Iso + isom : Iso.Iso _ _ + fun isom = _ + inv isom x = c i F-ob (x i) c) , λ {c} {c'} f i F-hom (x i) {c} {c'} f + F-ob (rightInv isom b _ i₁) = F-ob (b i₁) + F-hom (rightInv isom b _ i₁) = F-hom (b i₁) + F-id (rightInv isom b i i₁) = isProp→SquareP + i i₁ D .isSetHom (F-hom (b i₁) (C .id)) (D .id)) refl refl + (isProp→PathP j isSetHom D _ _) _ _) i₁ F-id (b i₁)) i i₁ + F-seq (rightInv isom b i i₁) f g = isProp→SquareP + i i₁ D .isSetHom (F-hom (b i₁) _) (seq' D (F-hom (b i₁) f) _)) + refl refl (isProp→PathP j isSetHom D _ _) _ _) i₁ F-seq (b i₁) f g) i i₁ + leftInv isom _ = refl + +isOfHLevelFunctor : hLevel isOfHLevel (2 + hLevel) (D .ob) + isOfHLevel (2 + hLevel) (Functor C D) +isOfHLevelFunctor {D = D} {C = C} hLevel x _ _ = + isOfHLevelRespectEquiv (1 + hLevel) (_ , isEquivFunctor≡) + (isOfHLevelΣ (1 + hLevel) (isOfHLevelΠ (1 + hLevel) _ x _ _)) + λ _ isOfHLevelPlus' 1 (isPropImplicitΠ2 + λ _ _ isPropΠ λ _ isOfHLevelPathP' 1 _ _ D .isSetHom _ _) _ _ )) + +isSetFunctor : isSet (D .ob) isSet (Functor C D) +isSetFunctor = isOfHLevelFunctor 0 + +-- Conservative Functor, +-- namely if a morphism f is mapped to an isomorphism, +-- the morphism f is itself isomorphism. + +isConservative : (F : Functor C D) Type _ +isConservative {C = C} {D = D} F = {x y : C .ob}{f : C [ x , y ]} isIso D (F .F-hom f) isIso C f + + +-- Fully-faithfulness of functors + +module _ {F : Functor C D} where + + isFullyFaithful→Full : isFullyFaithful F isFull F + isFullyFaithful→Full fullfaith x y = isEquiv→isSurjection (fullfaith x y) + + isFullyFaithful→Faithful : isFullyFaithful F isFaithful F + isFullyFaithful→Faithful fullfaith x y = isEmbedding→Inj (isEquiv→isEmbedding (fullfaith x y)) + + isFull+Faithful→isFullyFaithful : isFull F isFaithful F isFullyFaithful F + isFull+Faithful→isFullyFaithful full faith x y = isEmbedding×isSurjection→isEquiv + (injEmbedding (D .isSetHom) (faith x y _ _) , full x y) + + isFaithful→reflectsMono : isFaithful F {x y : C .ob} (f : C [ x , y ]) + isMonic D (F f ) isMonic C f + isFaithful→reflectsMono F-faithful f Ff-mon {a = a} {a' = a'} a⋆f≡a'⋆f = + let Fa⋆Ff≡Fa'⋆Ff = sym (F .F-seq a f) + cong (F ⟪_⟫) a⋆f≡a'⋆f + F .F-seq a' f + in F-faithful _ _ _ _ (Ff-mon Fa⋆Ff≡Fa'⋆Ff) + + + -- Fully-faithful functor is conservative + + open isIso + + isFullyFaithful→Conservative : isFullyFaithful F isConservative F + isFullyFaithful→Conservative fullfaith {x = x} {y = y} {f = f} isoFf = w + where + w : isIso C f + w .inv = invIsEq (fullfaith _ _) (isoFf .inv) + w .sec = isFullyFaithful→Faithful fullfaith _ _ _ _ + (F .F-seq _ _ + i secIsEq (fullfaith _ _) (isoFf .inv) i ⋆⟨ D F .F-hom f) + isoFf .sec + sym (F .F-id)) + w .ret = isFullyFaithful→Faithful fullfaith _ _ _ _ + (F .F-seq _ _ + i F .F-hom f ⋆⟨ D secIsEq (fullfaith _ _) (isoFf .inv) i) + isoFf .ret + sym (F .F-id)) + + -- Lifting isomorphism upwards a fully faithful functor + + module _ (fullfaith : isFullyFaithful F) where + + liftIso : {x y : C .ob} CatIso D (F .F-ob x) (F .F-ob y) CatIso C x y + liftIso f .fst = invEq (_ , fullfaith _ _) (f .fst) + liftIso f .snd = isFullyFaithful→Conservative fullfaith (subst (isIso D) (sym (secEq (_ , fullfaith _ _) (f .fst))) (f .snd)) + + liftIso≡ : {x y : C .ob} (f : CatIso D (F .F-ob x) (F .F-ob y)) F-Iso {F = F} (liftIso f) f + liftIso≡ f = CatIso≡ _ _ (secEq (_ , fullfaith _ _) (f .fst)) + + +-- Functors inducing surjection on objects is essentially surjective + +isSurj-ob→isSurj : {F : Functor C D} isSurjection (F .F-ob) isEssentiallySurj F +isSurj-ob→isSurj surj y = Prop.map (x , p) x , pathToIso p) (surj y) + + +-- Fully-faithful functors induce equivalence on isomorphisms - -- Lifting isomorphism upwards a fully faithful functor +isFullyFaithful→isEquivF-Iso : {F : Functor C D} + isFullyFaithful F x y isEquiv (F-Iso {F = F} {x = x} {y = y}) +isFullyFaithful→isEquivF-Iso {F = F} fullfaith x y = + Σ-cong-equiv-prop (_ , fullfaith x y) isPropIsIso isPropIsIso _ + f isFullyFaithful→Conservative {F = F} fullfaith {f = f}) .snd - module _ (fullfaith : isFullyFaithful F) where - liftIso : {x y : C .ob} CatIso D (F .F-ob x) (F .F-ob y) CatIso C x y - liftIso f .fst = invEq (_ , fullfaith _ _) (f .fst) - liftIso f .snd = isFullyFaithful→Conservative fullfaith (subst (isIso D) (sym (secEq (_ , fullfaith _ _) (f .fst))) (f .snd)) +-- Functors involving univalent categories - liftIso≡ : {x y : C .ob} (f : CatIso D (F .F-ob x) (F .F-ob y)) F-Iso {F = F} (liftIso f) f - liftIso≡ f = CatIso≡ _ _ (secEq (_ , fullfaith _ _) (f .fst)) +module _ + (isUnivD : isUnivalent D) + where + open isUnivalent isUnivD --- Functors inducing surjection on objects is essentially surjective + -- Essentially surjective functor with univalent target induces surjection on objects -isSurj-ob→isSurj : {F : Functor C D} isSurjection (F .F-ob) isEssentiallySurj F -isSurj-ob→isSurj surj y = Prop.map (x , p) x , pathToIso p) (surj y) + isSurj→isSurj-ob : {F : Functor C D} isEssentiallySurj F isSurjection (F .F-ob) + isSurj→isSurj-ob surj y = Prop.map (x , f) x , CatIsoToPath f) (surj y) --- Fully-faithful functors induce equivalence on isomorphisms +module _ + (isUnivC : isUnivalent C) + (isUnivD : isUnivalent D) + {F : Functor C D} + where -isFullyFaithful→isEquivF-Iso : {F : Functor C D} - isFullyFaithful F x y isEquiv (F-Iso {F = F} {x = x} {y = y}) -isFullyFaithful→isEquivF-Iso {F = F} fullfaith x y = - Σ-cong-equiv-prop (_ , fullfaith x y) isPropIsIso isPropIsIso _ - f isFullyFaithful→Conservative {F = F} fullfaith {f = f}) .snd + open isUnivalent + -- Fully-faithful functor between univalent target induces embedding on objects --- Functors involving univalent categories + isFullyFaithful→isEmbd-ob : isFullyFaithful F isEmbedding (F .F-ob) + isFullyFaithful→isEmbd-ob fullfaith x y = + isEquiv[equivFunA≃B∘f]→isEquiv[f] _ (_ , isUnivD .univ _ _) + (subst isEquiv (F-pathToIso-∘ {F = F}) + (compEquiv (_ , isUnivC .univ _ _) + (_ , isFullyFaithful→isEquivF-Iso {F = F} fullfaith x y) .snd)) -module _ - (isUnivD : isUnivalent D) - where +TransportFunctor : (C D) Functor C D +F-ob (TransportFunctor p) = subst ob p +F-hom (TransportFunctor p) {x} {y} = + transport λ i cong Hom[_,_] p i + (transport-filler (cong ob p) x i) + (transport-filler (cong ob p) y i) +F-id (TransportFunctor p) {x} i = + transp jj Hom[ p (i jj) , transport-filler i₁ ob (p i₁)) x (i jj) ] + (transport-filler i₁ ob (p i₁)) x (i jj))) i + (id (p i) {(transport-filler (cong ob p) x i)}) - open isUnivalent isUnivD - - -- Essentially surjective functor with univalent target induces surjection on objects - - isSurj→isSurj-ob : {F : Functor C D} isEssentiallySurj F isSurjection (F .F-ob) - isSurj→isSurj-ob surj y = Prop.map (x , f) x , CatIsoToPath f) (surj y) - - -module _ - (isUnivC : isUnivalent C) - (isUnivD : isUnivalent D) - {F : Functor C D} - where - - open isUnivalent - - -- Fully-faithful functor between univalent target induces embedding on objects - - isFullyFaithful→isEmbd-ob : isFullyFaithful F isEmbedding (F .F-ob) - isFullyFaithful→isEmbd-ob fullfaith x y = - isEquiv[equivFunA≃B∘f]→isEquiv[f] _ (_ , isUnivD .univ _ _) - (subst isEquiv (F-pathToIso-∘ {F = F}) - (compEquiv (_ , isUnivC .univ _ _) - (_ , isFullyFaithful→isEquivF-Iso {F = F} fullfaith x y) .snd)) +F-seq (TransportFunctor p) {x} {y} {z} f g i = + let q : {x y} _ _ + q = λ {x y} λ i₁ + Hom[ p i₁ , transport-filler i₂ ob (p i₂)) x i₁ ] + (transport-filler i₂ ob (p i₂)) y i₁) + in transp jj Hom[ p (i jj) + , transport-filler i₁ ob (p i₁)) x (i jj) ] + (transport-filler i₁ ob (p i₁)) z (i jj))) i + (_⋆_ (p i) (transport-filler q f i) (transport-filler q g i))
\ No newline at end of file diff --git a/docs/Cubical.Categories.Functor.html b/docs/Cubical.Categories.Functor.html index 369e874..a696234 100644 --- a/docs/Cubical.Categories.Functor.html +++ b/docs/Cubical.Categories.Functor.html @@ -1,9 +1,9 @@ -Cubical.Categories.Functor
{-# OPTIONS --allow-unsolved-metas #-}
+Cubical.Categories.Functor
{-# OPTIONS --safe #-}
 
-module Cubical.Categories.Functor where
+module Cubical.Categories.Functor where
 
-open import Cubical.Categories.Functor.Base public
-open import Cubical.Categories.Functor.Compose public
-open import Cubical.Categories.Functor.Properties public
+open import Cubical.Categories.Functor.Base public
+open import Cubical.Categories.Functor.Compose public
+open import Cubical.Categories.Functor.Properties public
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Instances.Cospan.html b/docs/Cubical.Categories.Instances.Cospan.html index 9c9b56f..42b70a6 100644 --- a/docs/Cubical.Categories.Instances.Cospan.html +++ b/docs/Cubical.Categories.Instances.Cospan.html @@ -66,7 +66,7 @@ -- makes it easier to write functors into CospanCat -isPropHomCospanCat : (x y : ob CospanCat) isProp (CospanCat [ x , y ]) +isPropHomCospanCat : (x y : ob CospanCat) isProp (CospanCat [ x , y ]) isPropHomCospanCat = isPropUnit isPropHomCospanCat = isPropUnit isPropHomCospanCat = isProp⊥ diff --git a/docs/Cubical.Categories.Instances.Functors.html b/docs/Cubical.Categories.Instances.Functors.html index d87eccc..f5a25ee 100644 --- a/docs/Cubical.Categories.Instances.Functors.html +++ b/docs/Cubical.Categories.Instances.Functors.html @@ -1,173 +1,142 @@ -Cubical.Categories.Instances.Functors
{-# OPTIONS --allow-unsolved-metas #-}
+Cubical.Categories.Instances.Functors
{-# OPTIONS --safe #-}
 
-{-
+{-
    Category whose objects are functors and morphisms are natural transformations.
 
    Includes the following
    - isos in FUNCTOR are precisely the pointwise isos
    - FUNCTOR C D is univalent when D is
-   - currying of functors
 
-   TODO: show that currying of functors is an isomorphism.
 -}
 
-module Cubical.Categories.Instances.Functors where
-
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Equiv.Properties
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Isomorphism
-
-open import Cubical.Categories.Category renaming (isIso to isIsoC)
-open import Cubical.Categories.Constructions.BinProduct
-open import Cubical.Categories.Functor.Base
-open import Cubical.Categories.Morphism
-open import Cubical.Categories.NaturalTransformation.Base
-open import Cubical.Categories.NaturalTransformation.Properties
-
-private
-  variable
-    ℓC ℓC' ℓD ℓD' ℓE ℓE' : Level
-
-module _ (C : Category ℓC ℓC') (D : Category ℓD ℓD') where
-  open Category
-  open NatTrans
-  open Functor
-
-  FUNCTOR : Category (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) (ℓ-max (ℓ-max ℓC ℓC') ℓD')
-  ob FUNCTOR           = Functor C D
-  Hom[_,_] FUNCTOR     = NatTrans
-  id FUNCTOR {F}       = idTrans F
-  _⋆_ FUNCTOR          = seqTrans
-  ⋆IdL FUNCTOR α       = makeNatTransPath λ i x  D .⋆IdL (α .N-ob x) i
-  ⋆IdR FUNCTOR α       = makeNatTransPath λ i x  D .⋆IdR (α .N-ob x) i
-  ⋆Assoc FUNCTOR α β γ = makeNatTransPath λ i x  D .⋆Assoc (α .N-ob x) (β .N-ob x) (γ .N-ob x) i
-  isSetHom FUNCTOR     = isSetNatTrans
-
-  open isIsoC renaming (inv to invC)
-  -- componentwise iso is an iso in Functor
-  FUNCTORIso :  {F G : Functor C D} (α : F  G)
-              (∀ (c : C .ob)  isIsoC D (α  c ))
-              isIsoC FUNCTOR α
-  FUNCTORIso α is .invC .N-ob c = (is c) .invC
-  FUNCTORIso {F} {G} α is .invC .N-hom {c} {d} f
-    = invMoveL areInv-αc
-               ( α  c  ⋆⟨ D  (G  f  ⋆⟨ D  is d .invC)
-               ≡⟨ sym (D .⋆Assoc _ _ _) 
-                 (α  c  ⋆⟨ D  G  f ) ⋆⟨ D  is d .invC
-               ≡⟨ sym (invMoveR areInv-αd (α .N-hom f)) 
-                 F  f 
-                )
-    where
-      areInv-αc : areInv _ (α  c ) ((is c) .invC)
-      areInv-αc = isIso→areInv (is c)
-
-      areInv-αd : areInv _ (α  d ) ((is d) .invC)
-      areInv-αd = isIso→areInv (is d)
-  FUNCTORIso α is .sec = makeNatTransPath (funExt  c  (is c) .sec))
-  FUNCTORIso α is .ret = makeNatTransPath (funExt  c  (is c) .ret))
-
-  -- iso is componentwise iso in Functor
-  FUNCTORIso' :  {F G : Functor C D} (α : F  G)
-              isIsoC FUNCTOR α
-              ((c : C .ob)  isIsoC D (α  c ))
-  FUNCTORIso' α isom c .invC = isom .invC .N-ob c
-  FUNCTORIso' α isom c .sec i = isom .sec i .N-ob c
-  FUNCTORIso' α isom c .ret i = isom .ret i .N-ob c
-
-  open Iso
-  open NatIso
-
-  FUNCTORIso→NatIso : {F G : Functor C D}  CatIso FUNCTOR F G  NatIso F G
-  FUNCTORIso→NatIso α .trans = α .fst
-  FUNCTORIso→NatIso α .nIso = FUNCTORIso' _ (α .snd)
-
-  NatIso→FUNCTORIso : {F G : Functor C D}  NatIso F G  CatIso FUNCTOR F G
-  NatIso→FUNCTORIso α = α .trans , FUNCTORIso _ (α .nIso)
-
-  Path→FUNCTORIso→NatIso : {F G : Functor C D}  (p : F  G)  pathToNatIso p  FUNCTORIso→NatIso (pathToIso p)
-  Path→FUNCTORIso→NatIso {F = F} p = J  _ p  pathToNatIso p  FUNCTORIso→NatIso (pathToIso p)) (NatIso≡ refl-helper) p
-    where
-    refl-helper : _
-    refl-helper i x = ((λ i  pathToIso-refl {C = D} {x = F .F-ob x} i .fst)
-        i  pathToIso-refl {C = FUNCTOR} {x = F} (~ i) .fst .N-ob x)) i
-
-  Iso-FUNCTORIso-NatIso : {F G : Functor C D}  Iso (CatIso FUNCTOR F G) (NatIso F G)
-  Iso-FUNCTORIso-NatIso .fun = FUNCTORIso→NatIso
-  Iso-FUNCTORIso-NatIso .inv = NatIso→FUNCTORIso
-  Iso-FUNCTORIso-NatIso .rightInv α i .trans = α .trans
-  Iso-FUNCTORIso-NatIso .rightInv α i .nIso =
-    isProp→PathP  i  isPropΠ  _  isPropIsIso _)) (FUNCTORIso' (α .trans) (FUNCTORIso _ (α .nIso))) (α .nIso) i
-  Iso-FUNCTORIso-NatIso .leftInv α i .fst = α .fst
-  Iso-FUNCTORIso-NatIso .leftInv α i .snd =
-    isProp→PathP  i  isPropIsIso _) (FUNCTORIso _ (FUNCTORIso' _ (α .snd))) (α .snd) i
-
-  FUNCTORIso≃NatIso : {F G : Functor C D}  CatIso FUNCTOR F G  NatIso F G
-  FUNCTORIso≃NatIso = isoToEquiv Iso-FUNCTORIso-NatIso
-
-
-  -- Functor Category is Univalent if the Target Category is Univalent
-
-  open isUnivalent
-
-  isUnivalentFUNCTOR : isUnivalent D  isUnivalent FUNCTOR
-  isUnivalentFUNCTOR isUnivD .univ _ _ =
-    isEquiv[equivFunA≃B∘f]→isEquiv[f] _ FUNCTORIso≃NatIso
-      (subst isEquiv  i p  Path→FUNCTORIso→NatIso p i) (Path≃NatIso isUnivD .snd))
-
-  appF : Functor (FUNCTOR ×C C) D
-  appF .F-ob (F , c) = F  c 
-  appF .F-hom {F , c} {G , d} (α , f) = α .N-ob d ∘⟨ D  F .F-hom f
-  appF .F-id {F , c} =
-    D .id ∘⟨ D  F .F-hom (C .id) ≡⟨ D .⋆IdR (F .F-hom (C .id)) 
-    F .F-hom (C .id) ≡⟨ F .F-id 
-    D .id 
-  appF .F-seq {F , c}{G , d}{H , e} (α , f) (β , g ) =
-    (β .N-ob e ∘⟨ D  α .N-ob e) ∘⟨ D  F .F-hom (g ∘⟨ C  f)
-      ≡⟨  i  (β .N-ob e ∘⟨ D  α .N-ob e) ∘⟨ D  F .F-seq f g i) 
-    (β .N-ob e ∘⟨ D  α .N-ob e) ∘⟨ D  (F .F-hom g ∘⟨ D  F .F-hom f)
-      ≡⟨ sym (D .⋆Assoc _ _ _) 
-    β .N-ob e ∘⟨ D  (α .N-ob e ∘⟨ D  (F .F-hom g ∘⟨ D  F .F-hom f))
-      ≡⟨  i  β .N-ob e
-                ∘⟨ D  D .⋆Assoc (F .F-hom f) (F .F-hom g) (α .N-ob e) i) 
-    β .N-ob e ∘⟨ D  ((α .N-ob e ∘⟨ D  F .F-hom g) ∘⟨ D  F .F-hom f)
-      ≡⟨  i  β .N-ob e ∘⟨ D  α .N-hom g i ∘⟨ D  F .F-hom f) 
-    β .N-ob e ∘⟨ D  ((G .F-hom g ∘⟨ D  α .N-ob d) ∘⟨ D  F .F-hom f)
-      ≡⟨  i  β .N-ob e
-                ∘⟨ D  D .⋆Assoc (F .F-hom f) (α .N-ob d) (G .F-hom g) (~ i) ) 
-    β .N-ob e ∘⟨ D  (G .F-hom g ∘⟨ D  (α .N-ob d ∘⟨ D  F .F-hom f))
-      ≡⟨ D .⋆Assoc _ _ _ 
-    (β .N-ob e ∘⟨ D  G .F-hom g) ∘⟨ D  (α .N-ob d ∘⟨ D  F .F-hom f) 
-  module _ (E : Category ℓE ℓE') where
-    λF : Functor (E ×C C) D  Functor E FUNCTOR
-    λF F .F-ob e .F-ob c = F  e , c 
-    λF F .F-ob e .F-hom f = F  (E .id) , f 
-    λF F .F-ob e .F-id = F .F-id
-    λF F .F-ob e .F-seq f g =
-      F  E .id , g ∘⟨ C  f 
-        ≡⟨  i  F  (E .⋆IdL (E .id) (~ i)) , (g ∘⟨ C  f) ) 
-      (F  (E .id ∘⟨ E  E .id) , g ∘⟨ C  f )
-        ≡⟨ F .F-seq (E .id , f) (E .id , g) 
-      (F  E .id , g  ∘⟨ D  F  E .id , f ) 
-    λF F .F-hom h .N-ob c = F  h , (C .id) 
-    λF F .F-hom h .N-hom f =
-      F  h , C .id  ∘⟨ D  F  E .id , f  ≡⟨ sym (F .F-seq _ _) 
-      F  h ∘⟨ E  E .id , C .id ∘⟨ C  f 
-        ≡⟨  i  F  E .⋆IdL h i , C .⋆IdR f i  ) 
-      F  h , f  ≡⟨  i  F  (E .⋆IdR h (~ i)) , (C .⋆IdL f (~ i)) ) 
-      F  E .id ∘⟨ E  h , f ∘⟨ C  C .id  ≡⟨ F .F-seq _ _ 
-      F  E .id , f  ∘⟨ D  F  h , C .id  
-    λF F .F-id = makeNatTransPath (funExt λ c  F .F-id)
-    λF F .F-seq f g = makeNatTransPath (funExt lem) where
-      lem : (c : C .ob) 
-            F  g ∘⟨ E  f , C .id  
-            F  g , C .id  ∘⟨ D  F  f , C .id 
-      lem c =
-        F  g ∘⟨ E  f , C .id 
-          ≡⟨  i  F  (g ∘⟨ E  f) , (C .⋆IdR (C .id) (~ i)) ) 
-        F  g ∘⟨ E  f , C .id ∘⟨ C  C .id 
-          ≡⟨ F .F-seq (f , C .id) (g , C .id) 
-        (F  g , C .id ) ∘⟨ D  (F  f , C .id ) 
+module Cubical.Categories.Instances.Functors where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Equiv.Properties
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+
+open import Cubical.Categories.Category renaming (isIso to isIsoC)
+open import Cubical.Categories.Constructions.BinProduct
+open import Cubical.Categories.Functor.Base
+open import Cubical.Categories.Morphism
+open import Cubical.Categories.NaturalTransformation.Base
+open import Cubical.Categories.NaturalTransformation.Properties
+
+private
+  variable
+    ℓC ℓC' ℓD ℓD' ℓE ℓE' : Level
+
+module _ (C : Category ℓC ℓC') (D : Category ℓD ℓD') where
+  open Category
+  open NatTrans
+  open Functor
+
+  FUNCTOR : Category (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) (ℓ-max (ℓ-max ℓC ℓC') ℓD')
+  ob FUNCTOR           = Functor C D
+  Hom[_,_] FUNCTOR     = NatTrans
+  id FUNCTOR {F}       = idTrans F
+  _⋆_ FUNCTOR          = seqTrans
+  ⋆IdL FUNCTOR α       = makeNatTransPath λ i x  D .⋆IdL (α .N-ob x) i
+  ⋆IdR FUNCTOR α       = makeNatTransPath λ i x  D .⋆IdR (α .N-ob x) i
+  ⋆Assoc FUNCTOR α β γ = makeNatTransPath λ i x  D .⋆Assoc (α .N-ob x) (β .N-ob x) (γ .N-ob x) i
+  isSetHom FUNCTOR     = isSetNatTrans
+
+  open isIsoC renaming (inv to invC)
+  -- componentwise iso is an iso in Functor
+  FUNCTORIso :  {F G : Functor C D} (α : F  G)
+              (∀ (c : C .ob)  isIsoC D (α  c ))
+              isIsoC FUNCTOR α
+  FUNCTORIso α is .invC .N-ob c = (is c) .invC
+  FUNCTORIso {F} {G} α is .invC .N-hom {c} {d} f
+    = invMoveL areInv-αc
+               ( α  c  ⋆⟨ D  (G  f  ⋆⟨ D  is d .invC)
+               ≡⟨ sym (D .⋆Assoc _ _ _) 
+                 (α  c  ⋆⟨ D  G  f ) ⋆⟨ D  is d .invC
+               ≡⟨ sym (invMoveR areInv-αd (α .N-hom f)) 
+                 F  f 
+                )
+    where
+      areInv-αc : areInv _ (α  c ) ((is c) .invC)
+      areInv-αc = isIso→areInv (is c)
+
+      areInv-αd : areInv _ (α  d ) ((is d) .invC)
+      areInv-αd = isIso→areInv (is d)
+  FUNCTORIso α is .sec = makeNatTransPath (funExt  c  (is c) .sec))
+  FUNCTORIso α is .ret = makeNatTransPath (funExt  c  (is c) .ret))
+
+  -- iso is componentwise iso in Functor
+  FUNCTORIso' :  {F G : Functor C D} (α : F  G)
+              isIsoC FUNCTOR α
+              ((c : C .ob)  isIsoC D (α  c ))
+  FUNCTORIso' α isom c .invC = isom .invC .N-ob c
+  FUNCTORIso' α isom c .sec i = isom .sec i .N-ob c
+  FUNCTORIso' α isom c .ret i = isom .ret i .N-ob c
+
+  open Iso
+  open NatIso
+
+  FUNCTORIso→NatIso : {F G : Functor C D}  CatIso FUNCTOR F G  NatIso F G
+  FUNCTORIso→NatIso α .trans = α .fst
+  FUNCTORIso→NatIso α .nIso = FUNCTORIso' _ (α .snd)
+
+  NatIso→FUNCTORIso : {F G : Functor C D}  NatIso F G  CatIso FUNCTOR F G
+  NatIso→FUNCTORIso α = α .trans , FUNCTORIso _ (α .nIso)
+
+  Path→FUNCTORIso→NatIso : {F G : Functor C D}  (p : F  G)  pathToNatIso p  FUNCTORIso→NatIso (pathToIso p)
+  Path→FUNCTORIso→NatIso {F = F} p = J  _ p  pathToNatIso p  FUNCTORIso→NatIso (pathToIso p)) (NatIso≡ refl-helper) p
+    where
+    refl-helper : _
+    refl-helper i x = ((λ i  pathToIso-refl {C = D} {x = F .F-ob x} i .fst)
+        i  pathToIso-refl {C = FUNCTOR} {x = F} (~ i) .fst .N-ob x)) i
+
+  Iso-FUNCTORIso-NatIso : {F G : Functor C D}  Iso (CatIso FUNCTOR F G) (NatIso F G)
+  Iso-FUNCTORIso-NatIso .fun = FUNCTORIso→NatIso
+  Iso-FUNCTORIso-NatIso .inv = NatIso→FUNCTORIso
+  Iso-FUNCTORIso-NatIso .rightInv α i .trans = α .trans
+  Iso-FUNCTORIso-NatIso .rightInv α i .nIso =
+    isProp→PathP  i  isPropΠ  _  isPropIsIso _)) (FUNCTORIso' (α .trans) (FUNCTORIso _ (α .nIso))) (α .nIso) i
+  Iso-FUNCTORIso-NatIso .leftInv α i .fst = α .fst
+  Iso-FUNCTORIso-NatIso .leftInv α i .snd =
+    isProp→PathP  i  isPropIsIso _) (FUNCTORIso _ (FUNCTORIso' _ (α .snd))) (α .snd) i
+
+  FUNCTORIso≃NatIso : {F G : Functor C D}  CatIso FUNCTOR F G  NatIso F G
+  FUNCTORIso≃NatIso = isoToEquiv Iso-FUNCTORIso-NatIso
+
+
+  -- Functor Category is Univalent if the Target Category is Univalent
+
+  open isUnivalent
+
+  isUnivalentFUNCTOR : isUnivalent D  isUnivalent FUNCTOR
+  isUnivalentFUNCTOR isUnivD .univ _ _ =
+    isEquiv[equivFunA≃B∘f]→isEquiv[f] _ FUNCTORIso≃NatIso
+      (subst isEquiv  i p  Path→FUNCTORIso→NatIso p i) (Path≃NatIso isUnivD .snd))
+
+  appF : Functor (FUNCTOR ×C C) D
+  appF .F-ob (F , c) = F  c 
+  appF .F-hom {F , c} {G , d} (α , f) = α .N-ob d ∘⟨ D  F .F-hom f
+  appF .F-id {F , c} =
+    D .id ∘⟨ D  F .F-hom (C .id) ≡⟨ D .⋆IdR (F .F-hom (C .id)) 
+    F .F-hom (C .id) ≡⟨ F .F-id 
+    D .id 
+  appF .F-seq {F , c}{G , d}{H , e} (α , f) (β , g ) =
+    (β .N-ob e ∘⟨ D  α .N-ob e) ∘⟨ D  F .F-hom (g ∘⟨ C  f)
+      ≡⟨  i  (β .N-ob e ∘⟨ D  α .N-ob e) ∘⟨ D  F .F-seq f g i) 
+    (β .N-ob e ∘⟨ D  α .N-ob e) ∘⟨ D  (F .F-hom g ∘⟨ D  F .F-hom f)
+      ≡⟨ sym (D .⋆Assoc _ _ _) 
+    β .N-ob e ∘⟨ D  (α .N-ob e ∘⟨ D  (F .F-hom g ∘⟨ D  F .F-hom f))
+      ≡⟨  i  β .N-ob e
+                ∘⟨ D  D .⋆Assoc (F .F-hom f) (F .F-hom g) (α .N-ob e) i) 
+    β .N-ob e ∘⟨ D  ((α .N-ob e ∘⟨ D  F .F-hom g) ∘⟨ D  F .F-hom f)
+      ≡⟨  i  β .N-ob e ∘⟨ D  α .N-hom g i ∘⟨ D  F .F-hom f) 
+    β .N-ob e ∘⟨ D  ((G .F-hom g ∘⟨ D  α .N-ob d) ∘⟨ D  F .F-hom f)
+      ≡⟨  i  β .N-ob e
+                ∘⟨ D  D .⋆Assoc (F .F-hom f) (α .N-ob d) (G .F-hom g) (~ i) ) 
+    β .N-ob e ∘⟨ D  (G .F-hom g ∘⟨ D  (α .N-ob d ∘⟨ D  F .F-hom f))
+      ≡⟨ D .⋆Assoc _ _ _ 
+    (β .N-ob e ∘⟨ D  G .F-hom g) ∘⟨ D  (α .N-ob d ∘⟨ D  F .F-hom f) 
+
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Instances.Sets.html b/docs/Cubical.Categories.Instances.Sets.html new file mode 100644 index 0000000..327af33 --- /dev/null +++ b/docs/Cubical.Categories.Instances.Sets.html @@ -0,0 +1,139 @@ + +Cubical.Categories.Instances.Sets
{-# OPTIONS --safe #-}
+
+module Cubical.Categories.Instances.Sets where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Equiv.Properties
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Data.Unit
+open import Cubical.Data.Sigma
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.NaturalTransformation
+
+open import Cubical.Categories.Limits
+
+open Category
+
+module _  where
+  SET : Category (ℓ-suc ) 
+  ob SET = hSet 
+  Hom[_,_] SET (A , _) (B , _) = A  B
+  id SET x = x
+  _⋆_ SET f g x = g (f x)
+  ⋆IdL SET f = refl
+  ⋆IdR SET f = refl
+  ⋆Assoc SET f g h = refl
+  isSetHom SET {A} {B} = isSetΠ  _  snd B)
+
+private
+  variable
+     ℓ' : Level
+
+open Functor
+
+-- Hom functors
+_[-,_] : (C : Category  ℓ')  (c : C .ob)  Functor (C ^op) (SET ℓ')
+(C [-, c ]) .F-ob x    = (C [ x , c ]) , C .isSetHom
+(C [-, c ]) .F-hom f k = f ⋆⟨ C  k
+(C [-, c ]) .F-id      = funExt λ _  C .⋆IdL _
+(C [-, c ]) .F-seq _ _ = funExt λ _  C .⋆Assoc _ _ _
+
+_[_,-] : (C : Category  ℓ')  (c : C .ob)→ Functor C (SET ℓ')
+(C [ c ,-]) .F-ob x    = (C [ c , x ]) , C .isSetHom
+(C [ c ,-]) .F-hom f k = k ⋆⟨ C  f
+(C [ c ,-]) .F-id      = funExt λ _  C .⋆IdR _
+(C [ c ,-]) .F-seq _ _ = funExt λ _  sym (C .⋆Assoc _ _ _)
+
+-- Lift functor
+LiftF : Functor (SET ) (SET (ℓ-max  ℓ'))
+LiftF {}{ℓ'} .F-ob A = (Lift {}{ℓ'} (A .fst)) , isOfHLevelLift 2 (A .snd)
+LiftF .F-hom f x = lift (f (x .lower))
+LiftF .F-id = refl
+LiftF .F-seq f g = funExt λ x  refl
+
+module _ {C : Category  ℓ'} {F : Functor C (SET ℓ')} where
+  open NatTrans
+
+  -- natural transformations by pre/post composition
+  preComp : {x y : C .ob}
+           (f : C [ x , y ])
+           C [ x ,-]  F
+           C [ y ,-]  F
+  preComp f α .N-ob c k = (α  c ) (f ⋆⟨ C  k)
+  preComp f α .N-hom {x = c} {d} k
+    =  l  (α  d ) (f ⋆⟨ C  (l ⋆⟨ C  k)))
+    ≡[ i ]⟨  l  (α  d ) (⋆Assoc C f l k (~ i))) 
+       l  (α  d ) (f ⋆⟨ C  l ⋆⟨ C  k))
+    ≡[ i ]⟨  l  (α .N-hom k) i (f ⋆⟨ C  l)) 
+       l  (F  k ) ((α  c ) (f ⋆⟨ C  l)))
+    
+
+-- properties
+-- TODO: move to own file
+open isIso renaming (inv to cInv)
+open Iso
+
+module _ {A B : (SET ) .ob } where
+
+  Iso→CatIso : Iso (fst A) (fst B)
+              CatIso (SET ) A B
+  Iso→CatIso is .fst = is .fun
+  Iso→CatIso is .snd .cInv = is .inv
+  Iso→CatIso is .snd .sec = funExt λ b  is .rightInv b -- is .rightInv
+  Iso→CatIso is .snd .ret = funExt λ b  is .leftInv b -- is .rightInv
+
+  CatIso→Iso : CatIso (SET ) A B
+              Iso (fst A) (fst B)
+  CatIso→Iso cis .fun = cis .fst
+  CatIso→Iso cis .inv = cis .snd .cInv
+  CatIso→Iso cis .rightInv = funExt⁻ λ b  cis .snd .sec b
+  CatIso→Iso cis .leftInv  = funExt⁻ λ b  cis .snd .ret b
+
+
+  Iso-Iso-CatIso : Iso (Iso (fst A) (fst B)) (CatIso (SET ) A B)
+  fun Iso-Iso-CatIso = Iso→CatIso
+  inv Iso-Iso-CatIso = CatIso→Iso
+  rightInv Iso-Iso-CatIso b = refl
+  fun (leftInv Iso-Iso-CatIso a i) = fun a
+  inv (leftInv Iso-Iso-CatIso a i) = inv a
+  rightInv (leftInv Iso-Iso-CatIso a i) = rightInv a
+  leftInv (leftInv Iso-Iso-CatIso a i) = leftInv a
+
+  Iso-CatIso-≡ : Iso (CatIso (SET ) A B) (A  B)
+  Iso-CatIso-≡ = compIso (invIso Iso-Iso-CatIso) (hSet-Iso-Iso-≡ _ _)
+
+-- SET is univalent
+
+isUnivalentSET : isUnivalent {ℓ' = } (SET _)
+isUnivalent.univ isUnivalentSET (A , isSet-A) (B , isSet-B)  =
+   precomposesToId→Equiv
+      pathToIso _ (funExt w) (isoToIsEquiv Iso-CatIso-≡)
+   where
+     w : _
+     w ci =
+       invEq
+         (congEquiv (isoToEquiv (invIso Iso-Iso-CatIso)))
+         (SetsIso≡-ext isSet-A isSet-B
+             x i  transp  _  B) i (ci .fst (transp  _  A) i x)))
+             x i  transp  _  A) i (ci .snd .cInv (transp  _  B) i x))))
+
+-- SET is complete
+
+open LimCone
+open Cone
+
+completeSET :  {ℓJ ℓJ'}  Limits {ℓJ} {ℓJ'} (SET (ℓ-max ℓJ ℓJ'))
+lim (completeSET J D) = Cone D (Unit* , isOfHLevelLift 2 isSetUnit) , isSetCone D _
+coneOut (limCone (completeSET J D)) j e = coneOut e j tt*
+coneOutCommutes (limCone (completeSET J D)) j i e = coneOutCommutes e j i tt*
+univProp (completeSET J D) c cc =
+  uniqueExists
+     x  cone  v _  coneOut cc v x)  e i _  coneOutCommutes cc e i x))
+     _  funExt  _  refl))
+     x  isPropIsConeMor cc (limCone (completeSET J D)) x)
+     x hx  funExt  d  cone≡ λ u  funExt  _  sym (funExt⁻ (hx u) d))))
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Isomorphism.html b/docs/Cubical.Categories.Isomorphism.html index 03fb5f2..9361388 100644 --- a/docs/Cubical.Categories.Isomorphism.html +++ b/docs/Cubical.Categories.Isomorphism.html @@ -17,228 +17,228 @@ module _ {C : Category ℓC ℓC'} where open Category C - open isIso + open isIso - invIso : {x y : ob} CatIso C x y CatIso C y x - invIso f .fst = f .snd .inv - invIso f .snd .inv = f .fst - invIso f .snd .sec = f .snd .ret - invIso f .snd .ret = f .snd .sec + invIso : {x y : ob} CatIso C x y CatIso C y x + invIso f .fst = f .snd .inv + invIso f .snd .inv = f .fst + invIso f .snd .sec = f .snd .ret + invIso f .snd .ret = f .snd .sec - invIsoIdem : {x y : ob} (f : CatIso C x y) invIso (invIso f) f + invIsoIdem : {x y : ob} (f : CatIso C x y) invIso (invIso f) f invIsoIdem f = refl - ⋆Iso : {x y z : ob} (f : CatIso C x y)(g : CatIso C y z) CatIso C x z + ⋆Iso : {x y z : ob} (f : CatIso C x y)(g : CatIso C y z) CatIso C x z ⋆Iso f g .fst = f .fst g .fst - ⋆Iso f g .snd .inv = g .snd .inv f .snd .inv - ⋆Iso f g .snd .sec = sym (⋆Assoc _ _ _) - i ⋆Assoc (g .snd .inv) (f .snd .inv) (f .fst) i g .fst) - i (g .snd .inv f .snd .sec i) g .fst) - i ⋆IdR (g .snd .inv) i g .fst) - g .snd .sec - ⋆Iso f g .snd .ret = sym (⋆Assoc _ _ _) - i ⋆Assoc (f .fst) (g .fst) (g .snd .inv) i f .snd .inv) - i (f .fst g .snd .ret i) f .snd .inv) - i ⋆IdR (f .fst) i f .snd .inv) - f .snd .ret - - compIso : {x y z : ob} (g : CatIso C y z)(f : CatIso C x y) CatIso C x z + ⋆Iso f g .snd .inv = g .snd .inv f .snd .inv + ⋆Iso f g .snd .sec = sym (⋆Assoc _ _ _) + i ⋆Assoc (g .snd .inv) (f .snd .inv) (f .fst) i g .fst) + i (g .snd .inv f .snd .sec i) g .fst) + i ⋆IdR (g .snd .inv) i g .fst) + g .snd .sec + ⋆Iso f g .snd .ret = sym (⋆Assoc _ _ _) + i ⋆Assoc (f .fst) (g .fst) (g .snd .inv) i f .snd .inv) + i (f .fst g .snd .ret i) f .snd .inv) + i ⋆IdR (f .fst) i f .snd .inv) + f .snd .ret + + compIso : {x y z : ob} (g : CatIso C y z)(f : CatIso C x y) CatIso C x z compIso g f = ⋆Iso f g - ⋆IsoIdL : {x y : ob} (f : CatIso C x y) ⋆Iso idCatIso f f - ⋆IsoIdL _ = CatIso≡ _ _ (⋆IdL _) + ⋆IsoIdL : {x y : ob} (f : CatIso C x y) ⋆Iso idCatIso f f + ⋆IsoIdL _ = CatIso≡ _ _ (⋆IdL _) - ⋆IsoIdR : {x y : ob} (f : CatIso C x y) ⋆Iso f idCatIso f - ⋆IsoIdR _ = CatIso≡ _ _ (⋆IdR _) + ⋆IsoIdR : {x y : ob} (f : CatIso C x y) ⋆Iso f idCatIso f + ⋆IsoIdR _ = CatIso≡ _ _ (⋆IdR _) - ⋆IsoInvRev : {x y z : ob} (f : CatIso C x y)(g : CatIso C y z) invIso (⋆Iso f g) ⋆Iso (invIso g) (invIso f) + ⋆IsoInvRev : {x y z : ob} (f : CatIso C x y)(g : CatIso C y z) invIso (⋆Iso f g) ⋆Iso (invIso g) (invIso f) ⋆IsoInvRev _ _ = refl - pathToIso-∙ : {x y z : ob}(p : x y)(q : y z) pathToIso (p q) ⋆Iso (pathToIso p) (pathToIso q) - pathToIso-∙ p q = J2 y p z q pathToIso (p q) ⋆Iso (pathToIso p) (pathToIso q)) pathToIso-∙-refl p q + pathToIso-∙ : {x y z : ob}(p : x y)(q : y z) pathToIso (p q) ⋆Iso (pathToIso p) (pathToIso q) + pathToIso-∙ p q = J2 y p z q pathToIso (p q) ⋆Iso (pathToIso p) (pathToIso q)) pathToIso-∙-refl p q where - pathToIso-∙-refl : {x : ob} pathToIso {x = x} (refl refl) ⋆Iso (pathToIso refl) (pathToIso refl) - pathToIso-∙-refl = cong pathToIso (sym compPathRefl) + pathToIso-∙-refl : {x : ob} pathToIso {x = x} (refl refl) ⋆Iso (pathToIso refl) (pathToIso refl) + pathToIso-∙-refl = cong pathToIso (sym compPathRefl) sym (⋆IsoIdL _) - i ⋆Iso (pathToIso-refl (~ i)) (pathToIso refl)) + i ⋆Iso (pathToIso-refl (~ i)) (pathToIso refl)) - transportPathToIso : {x y z : ob}(f : C [ x , y ])(p : y z) PathP i C [ x , p i ]) f (f pathToIso {C = C} p .fst) - transportPathToIso {x = x} f = J _ p PathP i C [ x , p i ]) f (f pathToIso {C = C} p .fst)) - (sym (⋆IdR _) cong x f x) (sym (cong fst (pathToIso-refl {C = C})))) + transportPathToIso : {x y z : ob}(f : C [ x , y ])(p : y z) PathP i C [ x , p i ]) f (f pathToIso {C = C} p .fst) + transportPathToIso {x = x} f = J _ p PathP i C [ x , p i ]) f (f pathToIso {C = C} p .fst)) + (sym (⋆IdR _) cong x f x) (sym (cong fst (pathToIso-refl {C = C})))) pathToIso-Comm : {x y z w : ob} (p : x y)(q : z w) (f : Hom[ x , z ])(g : Hom[ y , w ]) PathP i Hom[ p i , q i ]) f g - f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g + f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g pathToIso-Comm {x = x} {z = z} p q = - J2 y p w q + J2 y p w q (f : Hom[ x , z ])(g : Hom[ y , w ]) PathP i Hom[ p i , q i ]) f g - f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g) + f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g) sqr-refl p q where sqr-refl : (f g : Hom[ x , z ]) f g - f pathToIso {C = C} refl .fst pathToIso {C = C} refl .fst g - sqr-refl f g p = i f pathToIso-refl {C = C} i .fst) + f pathToIso {C = C} refl .fst pathToIso {C = C} refl .fst g + sqr-refl f g p = i f pathToIso-refl {C = C} i .fst) ⋆IdR _ p sym (⋆IdL _) - i pathToIso-refl {C = C} (~ i) .fst g) + i pathToIso-refl {C = C} (~ i) .fst g) pathToIso-Square : {x y z w : ob} (p : x y)(q : z w) (f : Hom[ x , z ])(g : Hom[ y , w ]) - f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g + f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g PathP i Hom[ p i , q i ]) f g pathToIso-Square {x = x} {z = z} p q = - J2 y p w q + J2 y p w q (f : Hom[ x , z ])(g : Hom[ y , w ]) - f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g + f pathToIso {C = C} q .fst pathToIso {C = C} p .fst g PathP i Hom[ p i , q i ]) f g) sqr-refl p q where sqr-refl : (f g : Hom[ x , z ]) - f pathToIso {C = C} refl .fst pathToIso {C = C} refl .fst g + f pathToIso {C = C} refl .fst pathToIso {C = C} refl .fst g f g sqr-refl f g p = sym (⋆IdR _) - i f pathToIso-refl {C = C} (~ i) .fst) + i f pathToIso-refl {C = C} (~ i) .fst) p - i pathToIso-refl {C = C} i .fst g) + i pathToIso-refl {C = C} i .fst g) ⋆IdL _ - module _ (isUnivC : isUnivalent C) where + module _ (isUnivC : isUnivalent C) where - open isUnivalent isUnivC + open isUnivalent isUnivC - transportIsoToPath : {x y z : ob}(f : C [ x , y ])(p : CatIso C y z) - PathP i C [ x , CatIsoToPath p i ]) f (f p .fst) + transportIsoToPath : {x y z : ob}(f : C [ x , y ])(p : CatIso C y z) + PathP i C [ x , CatIsoToPath p i ]) f (f p .fst) transportIsoToPath f p i = hcomp j λ { (i = i0) f - ; (i = i1) f secEq (univEquiv _ _) p j .fst }) - (transportPathToIso f (CatIsoToPath p) i) + ; (i = i1) f secEq (univEquiv _ _) p j .fst }) + (transportPathToIso f (CatIsoToPath p) i) - transportIsoToPathIso : {x y z : ob}(f : CatIso C x y)(p : CatIso C y z) - PathP i CatIso C x (CatIsoToPath p i)) f (⋆Iso f p) + transportIsoToPathIso : {x y z : ob}(f : CatIso C x y)(p : CatIso C y z) + PathP i CatIso C x (CatIsoToPath p i)) f (⋆Iso f p) transportIsoToPathIso f p i .fst = transportIsoToPath (f .fst) p i transportIsoToPathIso f p i .snd = - isProp→PathP i isPropIsIso (transportIsoToPath (f .fst) p i)) (f .snd) (⋆Iso f p .snd) i + isProp→PathP i isPropIsIso (transportIsoToPath (f .fst) p i)) (f .snd) (⋆Iso f p .snd) i isoToPath-Square : {x y z w : ob} - (p : CatIso C x y)(q : CatIso C z w) + (p : CatIso C x y)(q : CatIso C z w) (f : Hom[ x , z ])(g : Hom[ y , w ]) f q .fst p .fst g - PathP i Hom[ CatIsoToPath p i , CatIsoToPath q i ]) f g + PathP i Hom[ CatIsoToPath p i , CatIsoToPath q i ]) f g isoToPath-Square p q f g comm = - pathToIso-Square (CatIsoToPath p) (CatIsoToPath q) _ _ - ((λ i f secEq (univEquiv _ _) q i .fst) comm i secEq (univEquiv _ _) p (~ i) .fst g)) + pathToIso-Square (CatIsoToPath p) (CatIsoToPath q) _ _ + ((λ i f secEq (univEquiv _ _) q i .fst) comm i secEq (univEquiv _ _) p (~ i) .fst g)) module _ {C : Category ℓC ℓC'} where open Category C - open isIso + open isIso ⋆InvLMove : {x y z : ob} - (f : CatIso C x y) + (f : CatIso C x y) {g : Hom[ y , z ]}{h : Hom[ x , z ]} f .fst g h - g f .snd .inv h + g f .snd .inv h ⋆InvLMove f {g = g} p = sym (⋆IdL _) - i f .snd .sec (~ i) g) + i f .snd .sec (~ i) g) ⋆Assoc _ _ _ - i f .snd .inv p i) + i f .snd .inv p i) ⋆InvRMove : {x y z : ob} - (f : CatIso C y z) + (f : CatIso C y z) {g : Hom[ x , y ]}{h : Hom[ x , z ]} g f .fst h - g h f .snd .inv + g h f .snd .inv ⋆InvRMove f {g = g} p = sym (⋆IdR _) - i g f .snd .ret (~ i)) + i g f .snd .ret (~ i)) sym (⋆Assoc _ _ _) - i p i f .snd .inv) + i p i f .snd .inv) ⋆CancelL : {x y z : ob} - (f : CatIso C x y){g h : Hom[ y , z ]} + (f : CatIso C x y){g h : Hom[ y , z ]} f .fst g f .fst h g h ⋆CancelL f {g = g} {h = h} p = sym (⋆IdL _) - i f .snd .sec (~ i) g) + i f .snd .sec (~ i) g) ⋆Assoc _ _ _ - i f .snd .inv p i) + i f .snd .inv p i) sym (⋆Assoc _ _ _) - i f .snd .sec i h) + i f .snd .sec i h) ⋆IdL _ ⋆CancelR : {x y z : ob} - (f : CatIso C y z){g h : Hom[ x , y ]} + (f : CatIso C y z){g h : Hom[ x , y ]} g f .fst h f .fst g h ⋆CancelR f {g = g} {h = h} p = sym (⋆IdR _) - i g f .snd .ret (~ i)) + i g f .snd .ret (~ i)) sym (⋆Assoc _ _ _) - i p i f .snd .inv) + i p i f .snd .inv) ⋆Assoc _ _ _ - i h f .snd .ret i) + i h f .snd .ret i) ⋆IdR _ module _ {C : Category ℓC ℓC'} where open Category - open isIso + open isIso - op-Iso : {x y : C .ob} CatIso C x y CatIso (C ^op) x y - op-Iso f .fst = f .snd .inv - op-Iso f .snd .inv = f .fst - op-Iso f .snd .sec = f .snd .sec - op-Iso f .snd .ret = f .snd .ret + op-Iso : {x y : C .ob} CatIso C x y CatIso (C ^op) x y + op-Iso f .fst = f .snd .inv + op-Iso f .snd .inv = f .fst + op-Iso f .snd .sec = f .snd .sec + op-Iso f .snd .ret = f .snd .ret module _ {C : Category ℓC ℓC'}{D : Category ℓD ℓD'}{F : Functor C D} where open Category hiding (_∘_) - open isIso + open isIso open Functor - F-PresIsIso : {x y : C .ob}{f : C [ x , y ]} isIso C f isIso D (F .F-hom f) - F-PresIsIso f .inv = F . F-hom (f .inv) - F-PresIsIso f .sec = sym (F .F-seq (f .inv) _) cong (F .F-hom) (f .sec) F .F-id - F-PresIsIso f .ret = sym (F .F-seq _ (f .inv)) cong (F .F-hom) (f .ret) F .F-id + F-PresIsIso : {x y : C .ob}{f : C [ x , y ]} isIso C f isIso D (F .F-hom f) + F-PresIsIso f .inv = F . F-hom (f .inv) + F-PresIsIso f .sec = sym (F .F-seq (f .inv) _) cong (F .F-hom) (f .sec) F .F-id + F-PresIsIso f .ret = sym (F .F-seq _ (f .inv)) cong (F .F-hom) (f .ret) F .F-id - F-Iso : {x y : C .ob} CatIso C x y CatIso D (F .F-ob x) (F .F-ob y) + F-Iso : {x y : C .ob} CatIso C x y CatIso D (F .F-ob x) (F .F-ob y) F-Iso f .fst = F . F-hom (f .fst) F-Iso f .snd = F-PresIsIso (f .snd) - F-Iso-PresId : {x : C .ob} F-Iso (idCatIso {x = x}) idCatIso - F-Iso-PresId = CatIso≡ _ _ (F .F-id) + F-Iso-PresId : {x : C .ob} F-Iso (idCatIso {x = x}) idCatIso + F-Iso-PresId = CatIso≡ _ _ (F .F-id) - F-Iso-Pres⋆ : {x y z : C .ob} (f : CatIso C x y)(g : CatIso C y z) F-Iso (⋆Iso f g) ⋆Iso (F-Iso f) (F-Iso g) - F-Iso-Pres⋆ _ _ = CatIso≡ _ _ (F .F-seq _ _) + F-Iso-Pres⋆ : {x y z : C .ob} (f : CatIso C x y)(g : CatIso C y z) F-Iso (⋆Iso f g) ⋆Iso (F-Iso f) (F-Iso g) + F-Iso-Pres⋆ _ _ = CatIso≡ _ _ (F .F-seq _ _) - F-pathToIso : {x y : C .ob} (p : x y) F-Iso (pathToIso p) pathToIso (cong (F .F-ob) p) - F-pathToIso p = J y p F-Iso (pathToIso p) pathToIso (cong (F .F-ob) p)) F-pathToIso-refl p + F-pathToIso : {x y : C .ob} (p : x y) F-Iso (pathToIso p) pathToIso (cong (F .F-ob) p) + F-pathToIso p = J y p F-Iso (pathToIso p) pathToIso (cong (F .F-ob) p)) F-pathToIso-refl p where - F-pathToIso-refl : {x : C .ob} F-Iso (pathToIso {x = x} refl) pathToIso (cong (F .F-ob) refl) - F-pathToIso-refl = cong F-Iso pathToIso-refl + F-pathToIso-refl : {x : C .ob} F-Iso (pathToIso {x = x} refl) pathToIso (cong (F .F-ob) refl) + F-pathToIso-refl = cong F-Iso pathToIso-refl F-Iso-PresId - sym pathToIso-refl + sym pathToIso-refl - F-pathToIso-∘ : {x y : C .ob} F-Iso pathToIso {x = x} {y = y} pathToIso cong (F .F-ob) + F-pathToIso-∘ : {x y : C .ob} F-Iso pathToIso {x = x} {y = y} pathToIso cong (F .F-ob) F-pathToIso-∘ i p = F-pathToIso p i
\ No newline at end of file diff --git a/docs/Cubical.Categories.Limits.BinCoproduct.html b/docs/Cubical.Categories.Limits.BinCoproduct.html index efe694f..258ce54 100644 --- a/docs/Cubical.Categories.Limits.BinCoproduct.html +++ b/docs/Cubical.Categories.Limits.BinCoproduct.html @@ -25,8 +25,8 @@ isBinCoproduct = {z : ob} (f₁ : Hom[ x , z ]) (f₂ : Hom[ y , z ]) ∃![ f Hom[ x+y , z ] ] (i₁ f f₁) × (i₂ f f₂) - isPropIsBinCoproduct : isProp (isBinCoproduct) - isPropIsBinCoproduct = isPropImplicitΠ _ isPropΠ2 _ _ isPropIsContr)) + isPropIsBinCoproduct : isProp (isBinCoproduct) + isPropIsBinCoproduct = isPropImplicitΠ _ isPropΠ2 _ _ isPropIsContr)) record BinCoproduct (x y : ob) : Type (ℓ-max ℓ') where diff --git a/docs/Cubical.Categories.Limits.BinProduct.html b/docs/Cubical.Categories.Limits.BinProduct.html index 2e8f4d3..fc8bdd2 100644 --- a/docs/Cubical.Categories.Limits.BinProduct.html +++ b/docs/Cubical.Categories.Limits.BinProduct.html @@ -1,90 +1,45 @@ Cubical.Categories.Limits.BinProduct
-- Binary products
-{-# OPTIONS --allow-unsolved-metas #-}
+{-# OPTIONS --safe #-}
 
-module Cubical.Categories.Limits.BinProduct where
+module Cubical.Categories.Limits.BinProduct where
 
-open import Cubical.Categories.Category.Base
-open import Cubical.Categories.Constructions.BinProduct
-open import Cubical.Categories.Functor.Base
-open import Cubical.Data.Sigma.Base
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Prelude
-open import Cubical.HITs.PropositionalTruncation.Base
+open import Cubical.Categories.Category.Base
+open import Cubical.Data.Sigma.Base
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Prelude
+open import Cubical.HITs.PropositionalTruncation.Base
 
-private
-  variable
-     ℓ' : Level
+private
+  variable
+     ℓ' : Level
 
-module _ (C : Category  ℓ') where
-  open Category C
+module _ (C : Category  ℓ') where
+  open Category C
 
-  module _ {x y x×y : ob}
-           (π₁ : Hom[ x×y , x ])
-           (π₂ : Hom[ x×y , y ]) where
+  module _ {x y x×y : ob}
+           (π₁ : Hom[ x×y , x ])
+           (π₂ : Hom[ x×y , y ]) where
 
-    isBinProduct : Type (ℓ-max  ℓ')
-    isBinProduct =  {z : ob} (f₁ : Hom[ z , x ]) (f₂ : Hom[ z , y ]) 
-        ∃![ f  Hom[ z , x×y ] ] (f  π₁  f₁) × (f  π₂  f₂)
+    isBinProduct : Type (ℓ-max  ℓ')
+    isBinProduct =  {z : ob} (f₁ : Hom[ z , x ]) (f₂ : Hom[ z , y ]) 
+        ∃![ f  Hom[ z , x×y ] ] (f  π₁  f₁) × (f  π₂  f₂)
 
-    isPropIsBinProduct : isProp (isBinProduct)
-    isPropIsBinProduct = isPropImplicitΠ  _  isPropΠ2  _ _  isPropIsContr))
+    isPropIsBinProduct : isProp (isBinProduct)
+    isPropIsBinProduct = isPropImplicitΠ  _  isPropΠ2  _ _  isPropIsContr))
 
 
-  record BinProduct (x y : ob) : Type (ℓ-max  ℓ') where
-    field
-      binProdOb : ob
-      binProdPr₁ : Hom[ binProdOb , x ]
-      binProdPr₂ : Hom[ binProdOb , y ]
-      univProp : isBinProduct binProdPr₁ binProdPr₂
+  record BinProduct (x y : ob) : Type (ℓ-max  ℓ') where
+    field
+      binProdOb : ob
+      binProdPr₁ : Hom[ binProdOb , x ]
+      binProdPr₂ : Hom[ binProdOb , y ]
+      univProp : isBinProduct binProdPr₁ binProdPr₂
 
 
-  BinProducts : Type (ℓ-max  ℓ')
-  BinProducts = (x y : ob)  BinProduct x y
+  BinProducts : Type (ℓ-max  ℓ')
+  BinProducts = (x y : ob)  BinProduct x y
 
-  hasBinProducts : Type (ℓ-max  ℓ')
-  hasBinProducts =  BinProducts ∥₁
-
-  module _ (binProducts : BinProducts) where
-    {-
-      Given morphisms f : Hom[ a , b ] and g : Hom[ c , d ]
-      we can construct a morphism f×g : Hom[ a × c , b × d ].
-      This is given by f , g ⟨ a , c ⟩ ⊢ ⟨ f a , g c ⟩
-      Here the ⊢ and ⟨⟩ notation have their obvious meanings
-    -}
-    open BinProduct
-    private variable
-      a b c d : ob
-    infix 20 _⊗₀_
-    _⊗₀_ : ob  ob  ob
-    x ⊗₀ y = binProducts x y .binProdOb
-    
-    _⊗₁_ : (f : Hom[ a , b ]) (g : Hom[ c , d ])  Hom[ a ⊗₀ c , b ⊗₀ d ]
-    _⊗₁_ {a} {b} {c} {d} f g = b⊗d .univProp (a⊗c .binProdPr₁  f) (a⊗c .binProdPr₂  g) .fst .fst where
-         b⊗d = binProducts b d
-         a⊗c = binProducts a c
-
-    open Functor
-    prodFunctor : Functor (C ×C C) C
-    prodFunctor .F-ob (x , y) = x ⊗₀ y
-    prodFunctor .F-hom {x , y} {x' , y'} (f , g) = f ⊗₁ g
-    prodFunctor .F-id {x , y} i =
-      (isContr→isProp
-        (x×y .univProp
-          (x×y .binProdPr₁  id)
-          (x×y .binProdPr₂  id))
-        ( id {x} ⊗₁ id {y}
-        , x×y .univProp (x×y .binProdPr₁  id) (x×y .binProdPr₂  id) .fst .snd .fst
-        , x×y .univProp (x×y .binProdPr₁  id) (x×y .binProdPr₂  id) .fst .snd .snd)
-        ( id {x ⊗₀ y}
-        , (id  (x×y .binProdPr₁) ≡⟨ ⋆IdL _ 
-          x×y .binProdPr₁ ≡⟨ sym (⋆IdR _) 
-          x×y .binProdPr₁  id )
-        , (id  (x×y .binProdPr₂) ≡⟨ ⋆IdL _ 
-          x×y .binProdPr₂ ≡⟨ sym (⋆IdR _) 
-          x×y .binProdPr₂  id )))
-        i .fst where
-          x×y = binProducts x y
-    prodFunctor .F-seq {a , b} {c , d} {e , f} (α , β) (δ , γ) =
-      {!!}
+  hasBinProducts : Type (ℓ-max  ℓ')
+  hasBinProducts =  BinProducts ∥₁
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Limits.Initial.html b/docs/Cubical.Categories.Limits.Initial.html index 0a638be..3bcf23e 100644 --- a/docs/Cubical.Categories.Limits.Initial.html +++ b/docs/Cubical.Categories.Limits.Initial.html @@ -1,86 +1,86 @@ -Cubical.Categories.Limits.Initial
{-# OPTIONS --allow-unsolved-metas #-}
-module Cubical.Categories.Limits.Initial where
-
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Isomorphism renaming (Iso to _≅_)
-open import Cubical.HITs.PropositionalTruncation.Base
-
-open import Cubical.Data.Sigma
-
-open import Cubical.Categories.Category
-open import Cubical.Categories.Functor
-open import Cubical.Categories.Adjoint
-
-private
-  variable
-     ℓ' : Level
-    ℓC ℓC' ℓD ℓD' : Level
-
-module _ (C : Category  ℓ') where
-  open Category C
-
-  isInitial : (x : ob)  Type (ℓ-max  ℓ')
-  isInitial x =  (y : ob)  isContr (C [ x , y ])
-
-  Initial : Type (ℓ-max  ℓ')
-  Initial = Σ[ x  ob ] isInitial x
-
-  initialOb : Initial  ob
-  initialOb = fst
-
-  initialArrow : (T : Initial) (y : ob)  C [ initialOb T , y ]
-  initialArrow T y = T .snd y .fst
-
-  initialArrowUnique : {T : Initial} {y : ob} (f : C [ initialOb T , y ])
-                       initialArrow T y  f
-  initialArrowUnique {T} {y} f = T .snd y .snd f
-
-  initialEndoIsId : (T : Initial) (f : C [ initialOb T , initialOb T ])
-                    f  id
-  initialEndoIsId T f = isContr→isProp (T .snd (initialOb T)) f id
-
-  hasInitial : Type (ℓ-max  ℓ')
-  hasInitial =  Initial ∥₁
-
-  -- Initiality of an object is a proposition.
-  isPropIsInitial : (x : ob)  isProp (isInitial x)
-  isPropIsInitial _ = isPropΠ λ _  isPropIsContr
-
-  open isIso
-
-  -- Objects that are initial are isomorphic.
-  initialToIso : (x y : Initial)  CatIso C (initialOb x) (initialOb y)
-  initialToIso x y .fst = initialArrow x (initialOb y)
-  initialToIso x y .snd .inv = initialArrow y (initialOb x)
-  initialToIso x y .snd .sec = initialEndoIsId y _
-  initialToIso x y .snd .ret = initialEndoIsId x _
-
-  open isUnivalent
-
-  -- The type of initial objects of a univalent category is a proposition,
-  -- i.e. all initial objects are equal.
-  isPropInitial : (hC : isUnivalent C)  isProp Initial
-  isPropInitial hC x y =
-    Σ≡Prop isPropIsInitial (CatIsoToPath hC (initialToIso x y))
-
-module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) where
-  open Category
-  open Functor
-  open NaturalBijection
-  open _⊣_
-  open _≅_
-
-  preservesInitial : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
-  preservesInitial =  (x : ob C)  isInitial C x  isInitial D (F-ob F x)
-
-  isLeftAdjoint→preservesInitial : isLeftAdjoint F  preservesInitial
-  fst (isLeftAdjoint→preservesInitial (G , F⊣G) x initX y) = _♯ F⊣G (fst (initX (F-ob G y)))
-  snd (isLeftAdjoint→preservesInitial (G , F⊣G) x initX y) ψ =
-    _♯ F⊣G (fst (initX (F-ob G y)))
-      ≡⟨ cong (F⊣G ) (snd (initX (F-ob G y)) (_♭ F⊣G ψ)) 
-    _♯ F⊣G (_♭ F⊣G ψ)
-      ≡⟨ leftInv (adjIso F⊣G) ψ 
-    ψ 
+Cubical.Categories.Limits.Initial
{-# OPTIONS --safe #-}
+module Cubical.Categories.Limits.Initial where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism renaming (Iso to _≅_)
+open import Cubical.HITs.PropositionalTruncation.Base
+
+open import Cubical.Data.Sigma
+
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Adjoint
+
+private
+  variable
+     ℓ' : Level
+    ℓC ℓC' ℓD ℓD' : Level
+
+module _ (C : Category  ℓ') where
+  open Category C
+
+  isInitial : (x : ob)  Type (ℓ-max  ℓ')
+  isInitial x =  (y : ob)  isContr (C [ x , y ])
+
+  Initial : Type (ℓ-max  ℓ')
+  Initial = Σ[ x  ob ] isInitial x
+
+  initialOb : Initial  ob
+  initialOb = fst
+
+  initialArrow : (T : Initial) (y : ob)  C [ initialOb T , y ]
+  initialArrow T y = T .snd y .fst
+
+  initialArrowUnique : {T : Initial} {y : ob} (f : C [ initialOb T , y ])
+                       initialArrow T y  f
+  initialArrowUnique {T} {y} f = T .snd y .snd f
+
+  initialEndoIsId : (T : Initial) (f : C [ initialOb T , initialOb T ])
+                    f  id
+  initialEndoIsId T f = isContr→isProp (T .snd (initialOb T)) f id
+
+  hasInitial : Type (ℓ-max  ℓ')
+  hasInitial =  Initial ∥₁
+
+  -- Initiality of an object is a proposition.
+  isPropIsInitial : (x : ob)  isProp (isInitial x)
+  isPropIsInitial _ = isPropΠ λ _  isPropIsContr
+
+  open isIso
+
+  -- Objects that are initial are isomorphic.
+  initialToIso : (x y : Initial)  CatIso C (initialOb x) (initialOb y)
+  initialToIso x y .fst = initialArrow x (initialOb y)
+  initialToIso x y .snd .inv = initialArrow y (initialOb x)
+  initialToIso x y .snd .sec = initialEndoIsId y _
+  initialToIso x y .snd .ret = initialEndoIsId x _
+
+  open isUnivalent
+
+  -- The type of initial objects of a univalent category is a proposition,
+  -- i.e. all initial objects are equal.
+  isPropInitial : (hC : isUnivalent C)  isProp Initial
+  isPropInitial hC x y =
+    Σ≡Prop isPropIsInitial (CatIsoToPath hC (initialToIso x y))
+
+module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} (F : Functor C D) where
+  open Category
+  open Functor
+  open NaturalBijection
+  open _⊣_
+  open _≅_
+
+  preservesInitial : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
+  preservesInitial =  (x : ob C)  isInitial C x  isInitial D (F-ob F x)
+
+  isLeftAdjoint→preservesInitial : isLeftAdjoint F  preservesInitial
+  fst (isLeftAdjoint→preservesInitial (G , F⊣G) x initX y) = _♯ F⊣G (fst (initX (F-ob G y)))
+  snd (isLeftAdjoint→preservesInitial (G , F⊣G) x initX y) ψ =
+    _♯ F⊣G (fst (initX (F-ob G y)))
+      ≡⟨ cong (F⊣G ) (snd (initX (F-ob G y)) (_♭ F⊣G ψ)) 
+    _♯ F⊣G (_♭ F⊣G ψ)
+      ≡⟨ leftInv (adjIso F⊣G) ψ 
+    ψ 
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Limits.Limits.html b/docs/Cubical.Categories.Limits.Limits.html index eba4c32..72d598d 100644 --- a/docs/Cubical.Categories.Limits.Limits.html +++ b/docs/Cubical.Categories.Limits.Limits.html @@ -3,405 +3,405 @@ -- Heavily inspired by https://github.com/UniMath/UniMath/blob/master/UniMath/CategoryTheory/limits/graphs/limits.v -- (except we're using categories instead of graphs to index limits) -{-# OPTIONS #-} -module Cubical.Categories.Limits.Limits where - -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Function - -open import Cubical.Data.Sigma - -open import Cubical.Categories.Category -open import Cubical.Categories.Isomorphism -open import Cubical.Categories.Functor -open import Cubical.Categories.NaturalTransformation - -open import Cubical.Categories.Limits.Initial - -open import Cubical.HITs.PropositionalTruncation.Base - -module _ {ℓJ ℓJ' ℓC ℓC' : Level} {J : Category ℓJ ℓJ'} {C : Category ℓC ℓC'} where - open Category - open Functor - open NatTrans - - private - = ℓ-max (ℓ-max (ℓ-max ℓJ ℓJ') ℓC) ℓC' - - record Cone (D : Functor J C) (c : ob C) : Type (ℓ-max (ℓ-max ℓJ ℓJ') ℓC') where - constructor cone - - field - coneOut : (v : ob J) C [ c , F-ob D v ] - coneOutCommutes : {u v : ob J} (e : J [ u , v ]) - coneOut u ⋆⟨ C D .F-hom e coneOut v - - open Cone - - cone≡ : {D : Functor J C} {c : ob C} {c1 c2 : Cone D c} - ((v : ob J) coneOut c1 v coneOut c2 v) c1 c2 - coneOut (cone≡ h i) v = h v i - coneOutCommutes (cone≡ {D} {_} {c1} {c2} h i) {u} {v} e = - isProp→PathP j isSetHom C (h u j ⋆⟨ C D .F-hom e) (h v j)) - (coneOutCommutes c1 e) (coneOutCommutes c2 e) i - - -- dependent versions - conePathP : {D₁ D₂ : Functor J C} {c₁ c₂ : ob C} {cc₁ : Cone D₁ c₁} {cc₂ : Cone D₂ c₂} - {p : D₁ D₂} {q : c₁ c₂} - (∀ v PathP i C [ q i , p i .F-ob v ]) (cc₁ .coneOut v) (cc₂ .coneOut v)) - PathP i Cone (p i) (q i)) cc₁ cc₂ - coneOut (conePathP coneOutPathP i) v = coneOutPathP v i - coneOutCommutes (conePathP {cc₁ = cc₁} {cc₂ = cc₂} {p = p} coneOutPathP i) {u} {v} e = - isProp→PathP j isSetHom C (coneOutPathP u j ⋆⟨ C p j .F-hom e) - (coneOutPathP v j)) - (coneOutCommutes cc₁ e) (coneOutCommutes cc₂ e) i - - conePathPOb : {D : Functor J C} {c c' : ob C} {c1 : Cone D c} {c2 : Cone D c'} {p : c c'} - (∀ (v : ob J) PathP i C [ p i , F-ob D v ]) (coneOut c1 v) (coneOut c2 v)) - PathP i Cone D (p i)) c1 c2 - conePathPOb coneOutPathP = conePathP {p = refl} coneOutPathP - - conePathPDiag : {D₁ D₂ : Functor J C} {c : ob C} {cc₁ : Cone D₁ c} {cc₂ : Cone D₂ c} {p : D₁ D₂} - (∀ v PathP i C [ c , p i .F-ob v ]) (cc₁ .coneOut v) (cc₂ .coneOut v)) - PathP i Cone (p i) c) cc₁ cc₂ - conePathPDiag coneOutPathP = conePathP {q = refl} coneOutPathP - - - -- TODO: can we automate this a bit? - isSetCone : (D : Functor J C) (c : ob C) isSet (Cone D c) - isSetCone D c = isSetRetract toConeΣ fromConeΣ _ refl) - (isSetΣSndProp (isSetΠ _ isSetHom C)) - _ isPropImplicitΠ2 _ _ isPropΠ _ isSetHom C _ _)))) - where - ConeΣ = Σ[ f ((v : ob J) C [ c , F-ob D v ]) ] - ({u v : ob J} (e : J [ u , v ]) f u ⋆⟨ C D .F-hom e f v) - - toConeΣ : Cone D c ConeΣ - fst (toConeΣ x) = coneOut x - snd (toConeΣ x) = coneOutCommutes x - - fromConeΣ : ConeΣ Cone D c - coneOut (fromConeΣ x) = fst x - coneOutCommutes (fromConeΣ x) = snd x - - preCompCone : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} - Cone D c2 Cone D c1 - coneOut (preCompCone f cc) v = f ⋆⟨ C coneOut cc v - coneOutCommutes (preCompCone f cc) e = ⋆Assoc C _ _ _ - cong x f ⋆⟨ C x) (coneOutCommutes cc e) - - _★_ : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} Cone D c2 Cone D c1 - _★_ = preCompCone - - natTransPostCompCone : {c : ob C} {D₁ D₂ : Functor J C} (α : NatTrans D₁ D₂) - Cone D₁ c Cone D₂ c - coneOut (natTransPostCompCone α cc) u = cc .coneOut u ⋆⟨ C α .N-ob u - coneOutCommutes (natTransPostCompCone {D₁ = D₁} {D₂ = D₂} α cc) {u = u} {v = v} e = - cc .coneOut u ⋆⟨ C α .N-ob u ⋆⟨ C D₂ .F-hom e - ≡⟨ ⋆Assoc C _ _ _ - cc .coneOut u ⋆⟨ C (α .N-ob u ⋆⟨ C D₂ .F-hom e) - ≡⟨ cong x cc .coneOut u ⋆⟨ C x) (sym (α .N-hom e)) - cc .coneOut u ⋆⟨ C (D₁ .F-hom e ⋆⟨ C α .N-ob v) - ≡⟨ sym (⋆Assoc C _ _ _) - cc .coneOut u ⋆⟨ C D₁ .F-hom e ⋆⟨ C α .N-ob v - ≡⟨ cong x x ⋆⟨ C α .N-ob v) (cc .coneOutCommutes e) - cc .coneOut v ⋆⟨ C α .N-ob v - - _★ₙₜ_ : {c : ob C} {D₁ D₂ : Functor J C} Cone D₁ c NatTrans D₁ D₂ Cone D₂ c - _★ₙₜ_ = flip natTransPostCompCone - - isConeMor : {c1 c2 : ob C} {D : Functor J C} - (cc1 : Cone D c1) (cc2 : Cone D c2) - C [ c1 , c2 ] Type (ℓ-max ℓJ ℓC') - isConeMor cc1 cc2 f = (v : ob J) f ⋆⟨ C coneOut cc2 v coneOut cc1 v - - isPropIsConeMor : {c1 c2 : ob C} {D : Functor J C} - (cc1 : Cone D c1) (cc2 : Cone D c2) (f : C [ c1 , c2 ]) - isProp (isConeMor cc1 cc2 f) - isPropIsConeMor cc1 cc2 f = isPropΠ _ isSetHom C _ _) - - isConeMorId : {c : ob C} {D : Functor J C} (cc : Cone D c) - isConeMor cc cc (id C) - isConeMorId _ _ = ⋆IdL C _ - - isConeMorSeq : {c1 c2 c3 : ob C} {D : Functor J C} - (cc1 : Cone D c1) (cc2 : Cone D c2) (cc3 : Cone D c3) - {f : C [ c1 , c2 ]} {g : C [ c2 , c3 ]} - isConeMor cc1 cc2 f isConeMor cc2 cc3 g isConeMor cc1 cc3 (f ⋆⟨ C g) - isConeMorSeq cc1 cc2 cc3 {f} {g} isConeMorF isConeMorG v = - ⋆Assoc C _ _ _ ∙∙ cong x f ⋆⟨ C x) (isConeMorG v) ∙∙ isConeMorF v - - preCompConeMor : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} (cc : Cone D c2) - isConeMor (f cc) cc f - preCompConeMor f cc v = refl - - isLimCone : (D : Functor J C) (c0 : ob C) Cone D c0 Type - isLimCone D c0 cc0 = (c : ob C) (cc : Cone D c) - ∃![ f C [ c , c0 ] ] isConeMor cc cc0 f - - isPropIsLimCone : (D : Functor J C) (c0 : ob C) (cc0 : Cone D c0) - isProp (isLimCone D c0 cc0) - isPropIsLimCone D c0 cc0 = isPropΠ2 λ _ _ isProp∃! - - preCompUnique : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} (cc : Cone D c2) - isLimCone D c2 cc - (g : C [ c1 , c2 ]) isConeMor (f cc) cc g g f - preCompUnique f cc ccIsLimCone g gIsConeMor = - cong fst (isContr→isProp (ccIsLimCone _ (f cc)) (g , gIsConeMor) (f , preCompConeMor f cc)) - - record LimCone (D : Functor J C) : Type where - constructor climCone - - field - lim : ob C - limCone : Cone D lim - univProp : isLimCone D lim limCone - - limOut : (v : ob J) C [ lim , D .F-ob v ] - limOut = coneOut limCone - - limOutCommutes : {u v : ob J} (e : J [ u , v ]) - limOut u ⋆⟨ C D .F-hom e limOut v - limOutCommutes = coneOutCommutes limCone - - limArrow : (c : ob C) (cc : Cone D c) C [ c , lim ] - limArrow c cc = univProp c cc .fst .fst - - limArrowCommutes : (c : ob C) (cc : Cone D c) (u : ob J) - limArrow c cc ⋆⟨ C limOut u coneOut cc u - limArrowCommutes c cc = univProp c cc .fst .snd - - limArrowUnique : (c : ob C) (cc : Cone D c) (k : C [ c , lim ]) - isConeMor cc limCone k limArrow c cc k - limArrowUnique c cc k hk = cong fst (univProp c cc .snd (k , hk)) - - open LimCone - limOfArrowsCone : {D₁ D₂ : Functor J C} - (CC₁ : LimCone D₁) - NatTrans D₁ D₂ - Cone D₂ (CC₁ .lim) - coneOut (limOfArrowsCone {D₁} {D₂} CC₁ α) v = limOut CC₁ v ⋆⟨ C α .N-ob v - coneOutCommutes (limOfArrowsCone {D₁} {D₂} CC₁ α) {u = u} {v = v} e = - limOut CC₁ u ⋆⟨ C α .N-ob u ⋆⟨ C D₂ .F-hom e ≡⟨ ⋆Assoc C _ _ _ - limOut CC₁ u ⋆⟨ C (α .N-ob u ⋆⟨ C D₂ .F-hom e) ≡⟨ cong x seq' C (limOut CC₁ u) x) (sym (α .N-hom e)) - limOut CC₁ u ⋆⟨ C (D₁ .F-hom e ⋆⟨ C α .N-ob v) ≡⟨ sym (⋆Assoc C _ _ _) - limOut CC₁ u ⋆⟨ C D₁ .F-hom e ⋆⟨ C α .N-ob v ≡⟨ cong x x ⋆⟨ C α .N-ob v) (limOutCommutes CC₁ e) - limOut CC₁ v ⋆⟨ C α .N-ob v - - limOfArrows : {D₁ D₂ : Functor J C} - (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) - NatTrans D₁ D₂ - C [ CC₁ .lim , CC₂ .lim ] - limOfArrows {D₁} {D₂} CC₁ CC₂ α = limArrow CC₂ (CC₁ .lim) (limOfArrowsCone CC₁ α) - - limOfArrowsOut : {D₁ D₂ : Functor J C} - (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) - (α : NatTrans D₁ D₂) (u : ob J) - limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u limOut CC₁ u ⋆⟨ C α .N-ob u - limOfArrowsOut _ CC₂ _ _ = limArrowCommutes CC₂ _ _ _ - - limOfArrowsId : {D : Functor J C} (CC : LimCone D) - limOfArrows CC CC (idTrans D) id C - limOfArrowsId CC = limArrowUnique CC _ _ _ λ v ⋆IdL C _ sym (⋆IdR C _) - - limOfArrowsSeq : {D₁ D₂ D₃ : Functor J C} - (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) (CC₃ : LimCone D₃) - (α : NatTrans D₁ D₂) (β : NatTrans D₂ D₃) - limOfArrows CC₁ CC₃ (α ●ᵛ β) - limOfArrows CC₁ CC₂ α ⋆⟨ C limOfArrows CC₂ CC₃ β - limOfArrowsSeq CC₁ CC₂ CC₃ α β = limArrowUnique CC₃ _ _ _ path - where - path : u - (limOfArrows CC₁ CC₂ α ⋆⟨ C limOfArrows CC₂ CC₃ β) ⋆⟨ C limOut CC₃ u - limOut CC₁ u ⋆⟨ C (α .N-ob u ⋆⟨ C β .N-ob u) - path u = (limOfArrows CC₁ CC₂ α ⋆⟨ C limOfArrows CC₂ CC₃ β) ⋆⟨ C limOut CC₃ u - ≡⟨ ⋆Assoc C _ _ _ - limOfArrows CC₁ CC₂ α ⋆⟨ C (limOfArrows CC₂ CC₃ β ⋆⟨ C limOut CC₃ u) - ≡⟨ cong x limOfArrows CC₁ CC₂ α ⋆⟨ C x) (limOfArrowsOut CC₂ CC₃ β u) - limOfArrows CC₁ CC₂ α ⋆⟨ C (limOut CC₂ u ⋆⟨ C β .N-ob u) - ≡⟨ sym (⋆Assoc C _ _ _) - (limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u) ⋆⟨ C β .N-ob u - ≡⟨ cong x x ⋆⟨ C β .N-ob u) (limOfArrowsOut CC₁ CC₂ α u) - (limOut CC₁ u ⋆⟨ C α .N-ob u) ⋆⟨ C β .N-ob u - ≡⟨ ⋆Assoc C _ _ _ - limOut CC₁ u ⋆⟨ C (α .N-ob u ⋆⟨ C β .N-ob u) - - limArrowCompLimOfArrows : {D₁ D₂ : Functor J C} - (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) - (α : NatTrans D₁ D₂) - (c : ob C) (cc : Cone D₁ c) - limArrow CC₂ _ (cc ★ₙₜ α) limArrow CC₁ _ cc ⋆⟨ C limOfArrows CC₁ CC₂ α - limArrowCompLimOfArrows CC₁ CC₂ α c cc = limArrowUnique CC₂ _ _ _ isConeMorComp - where - isConeMorComp : (u : ob J) - limArrow CC₁ _ cc ⋆⟨ C limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u - cc .coneOut u ⋆⟨ C α .N-ob u - isConeMorComp u = - limArrow CC₁ _ cc ⋆⟨ C limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u - ≡⟨ ⋆Assoc C _ _ _ - limArrow CC₁ _ cc ⋆⟨ C (limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u) - ≡⟨ cong x limArrow CC₁ _ cc ⋆⟨ C x) (limOfArrowsOut CC₁ CC₂ α u) - limArrow CC₁ _ cc ⋆⟨ C (limOut CC₁ u ⋆⟨ C α .N-ob u) - ≡⟨ sym (⋆Assoc C _ _ _) - limArrow CC₁ _ cc ⋆⟨ C limOut CC₁ u ⋆⟨ C α .N-ob u - ≡⟨ cong x x ⋆⟨ C α .N-ob u) (limArrowCommutes CC₁ _ cc u) - cc .coneOut u ⋆⟨ C α .N-ob u - - -- any two limits are isomorphic up to a unique cone isomorphism - open isIso - LimIso : (D : Functor J C) (L₁ L₂ : LimCone D) - ∃![ e CatIso C (lim L₁) (lim L₂) ] isConeMor (limCone L₁) (limCone L₂) (fst e) - fst (fst (fst (LimIso D L₁ L₂))) = limArrow L₂ _ (limCone L₁) - inv (snd (fst (fst (LimIso D L₁ L₂)))) = limArrow L₁ _ (limCone L₂) - sec (snd (fst (fst (LimIso D L₁ L₂)))) = cong fst (isContr→isProp (univProp L₂ _ (limCone L₂)) - (_ , isConeMorComp) (id C , isConeMorId (limCone L₂))) - where - isConeMorComp : isConeMor (limCone L₂) (limCone L₂) - (limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C limArrow L₂ (lim L₁) (limCone L₁)) - isConeMorComp v = - (limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C limArrow L₂ (lim L₁) (limCone L₁)) - ⋆⟨ C coneOut (limCone L₂) v - ≡⟨ ⋆Assoc C _ _ _ - limArrow L₁ (lim L₂) (limCone L₂) - ⋆⟨ C (limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C coneOut (limCone L₂) v) - ≡⟨ cong x limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C x) - (limArrowCommutes L₂ _ (limCone L₁) v) - limArrow L₁ (lim L₂) (limCone L₂) - ⋆⟨ C (coneOut (limCone L₁) v) - ≡⟨ limArrowCommutes L₁ _ (limCone L₂) v - coneOut (limCone L₂) v - ret (snd (fst (fst (LimIso D L₁ L₂)))) = cong fst (isContr→isProp (univProp L₁ _ (limCone L₁)) - (_ , isConeMorComp) (id C , isConeMorId (limCone L₁))) - where - isConeMorComp : isConeMor (limCone L₁) (limCone L₁) - (limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C limArrow L₁ (lim L₂) (limCone L₂)) - isConeMorComp v = - (limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C limArrow L₁ (lim L₂) (limCone L₂)) - ⋆⟨ C coneOut (limCone L₁) v - ≡⟨ ⋆Assoc C _ _ _ - limArrow L₂ (lim L₁) (limCone L₁) - ⋆⟨ C (limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C coneOut (limCone L₁) v) - ≡⟨ cong x limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C x) - (limArrowCommutes L₁ _ (limCone L₂) v) - limArrow L₂ (lim L₁) (limCone L₁) - ⋆⟨ C (coneOut (limCone L₂) v) - ≡⟨ limArrowCommutes L₂ _ (limCone L₁) v - coneOut (limCone L₁) v - - snd (fst (LimIso D L₁ L₂)) = limArrowCommutes L₂ _ _ - snd (LimIso D L₁ L₂) (e , isConeMorE) = Σ≡Prop - _ isPropIsConeMor (limCone L₁) (limCone L₂) _) - (CatIso≡ _ _ (limArrowUnique L₂ _ _ (fst e) isConeMorE)) - - -- if the index category of the diagram has an initial object, - -- we get a canonical limiting cone - Initial→LimCone : (D : Functor J C) Initial J LimCone D - lim (Initial→LimCone D (j , isInitJ)) = D .F-ob j - coneOut (limCone (Initial→LimCone D (j , isInitJ))) k = D .F-hom (isInitJ k .fst) - coneOutCommutes (limCone (Initial→LimCone D (j , isInitJ))) f = - sym (D .F-seq _ _) cong (D .F-hom) (sym (isInitJ _ .snd _)) - fst (fst (univProp (Initial→LimCone D (j , isInitJ)) c cc)) = cc .coneOut j - -- canonical arrow c → D(j) - snd (fst (univProp (Initial→LimCone D (j , isInitJ)) c cc)) k = - cc .coneOutCommutes (isInitJ k .fst) - -- is a cone morphism - snd (univProp (Initial→LimCone D (j , isInitJ)) c cc) (f , isConeMorF) = -- and indeed unique - Σ≡Prop - _ isPropIsConeMor cc (limCone (Initial→LimCone D (j , isInitJ))) _) - (sym (isConeMorF j) ∙∙ cong x f ⋆⟨ C x) (subst x D .F-hom x id C) - (sym (isInitJ j .snd _)) (D .F-id)) - ∙∙ ⋆IdR C f) - - -- cones that respect isomorphisms are limiting cones - Iso→LimCone : {D : Functor J C} {c₁ c₂ : ob C} (cc₁ : Cone D c₁) (e : CatIso C c₁ c₂) - isLimCone _ _ cc₁ - (cc₂ : Cone D c₂) - isConeMor cc₁ cc₂ (e .fst) - isLimCone _ _ cc₂ - fst (fst (Iso→LimCone cc₁ e isLimConeCC₁ cc₂ isConeMorE d cd)) = - isLimConeCC₁ d cd .fst .fst ⋆⟨ C e .fst - snd (fst (Iso→LimCone cc₁ e isLimConeCC₁ cc₂ isConeMorE d cd)) = - isConeMorSeq cd cc₁ cc₂ (isLimConeCC₁ d cd .fst .snd) isConeMorE - snd (Iso→LimCone cc₁ e isLimConeCC₁ cc₂ isConeMorE d cd) (f , isConeMorF) = - Σ≡Prop (isPropIsConeMor cd cc₂) path - where - isConeMorE⁻¹ : isConeMor cc₂ cc₁ (e .snd .inv) - isConeMorE⁻¹ v = - e .snd .inv ⋆⟨ C coneOut cc₁ v ≡⟨ cong x e .snd .inv ⋆⟨ C x) (sym (isConeMorE v)) - e .snd .inv ⋆⟨ C (e .fst ⋆⟨ C coneOut cc₂ v) ≡⟨ sym (⋆Assoc C _ _ _) - (e .snd .inv ⋆⟨ C e .fst) ⋆⟨ C coneOut cc₂ v ≡⟨ cong x x ⋆⟨ C coneOut cc₂ v) - (e .snd .sec) - id C ⋆⟨ C coneOut cc₂ v ≡⟨ ⋆IdL C _ - coneOut cc₂ v - - p : isLimConeCC₁ d cd .fst .fst f ⋆⟨ C e .snd .inv - p = cong fst (isLimConeCC₁ d cd .snd ( f ⋆⟨ C e .snd .inv - , isConeMorSeq cd cc₂ cc₁ isConeMorF isConeMorE⁻¹)) - - path : isLimConeCC₁ d cd .fst .fst ⋆⟨ C e .fst f - path = isLimConeCC₁ d cd .fst .fst ⋆⟨ C e .fst ≡⟨ cong x x ⋆⟨ C e .fst) p - (f ⋆⟨ C e .snd .inv) ⋆⟨ C e .fst ≡⟨ ⋆Assoc C _ _ _ - f ⋆⟨ C (e .snd .inv ⋆⟨ C e .fst) ≡⟨ cong x f ⋆⟨ C x) (e .snd .sec) - f ⋆⟨ C id C ≡⟨ ⋆IdR C _ - f - --- A category is complete if it has all limits -Limits : {ℓJ ℓJ' ℓC ℓC' : Level} Category ℓC ℓC' Type _ -Limits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') (D : Functor J C) LimCone D - -hasLimits : {ℓJ ℓJ' ℓC ℓC' : Level} Category ℓC ℓC' Type _ -hasLimits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') (D : Functor J C) LimCone D ∥₁ - --- Limits of a specific shape J in a category C -LimitsOfShape : {ℓJ ℓJ' ℓC ℓC' : Level} Category ℓJ ℓJ' Category ℓC ℓC' Type _ -LimitsOfShape J C = (D : Functor J C) LimCone D - - --- precomposition with a functor and preservation of limits -module _ {ℓJ ℓJ' ℓC1 ℓC1' ℓC2 ℓC2' : Level} - {C1 : Category ℓC1 ℓC1'} {C2 : Category ℓC2 ℓC2'} - (F : Functor C1 C2) where - open Category - open Functor - open Cone - - -- same as F-cone!!! - F-cone : {J : Category ℓJ ℓJ'} {D : Functor J C1} {x : ob C1} - Cone D x Cone (funcComp F D) (F .F-ob x) - coneOut (F-cone ccx) v = F .F-hom (coneOut ccx v) - coneOutCommutes (F-cone ccx) e = - sym (F-seq F (coneOut ccx _) _) cong (F .F-hom) (coneOutCommutes ccx e) - - preservesLimits : Type _ - preservesLimits = {J : Category ℓJ ℓJ'} {D : Functor J C1} {L : ob C1} - (ccL : Cone D L) - isLimCone _ _ ccL isLimCone (funcComp F D) (F .F-ob L) (F-cone ccL) - - -- characterizing limit preserving functors - open LimCone - - preservesLimitsChar : (L₁ : Limits {ℓJ} {ℓJ'} C1) (L₂ : Limits {ℓJ} {ℓJ'} C2) - (e : J D CatIso C2 (L₂ J (F ∘F D) .lim) (F .F-ob (L₁ J D .lim))) - (∀ J D isConeMor (L₂ J (F ∘F D) .limCone) (F-cone (L₁ J D .limCone)) (e J D .fst)) - preservesLimits - preservesLimitsChar L₁ L₂ e isConeMorE {J = J} {D = D} {L = c} cc isLimConeCC = - Iso→LimCone (L₂ J (F ∘F D) .limCone) f (L₂ J (F ∘F D) .univProp) (F-cone cc) isConeMorF - where - theCLimCone : LimCone D - lim theCLimCone = c - limCone theCLimCone = cc - univProp theCLimCone = isLimConeCC - - f : CatIso C2 (lim (L₂ J (funcComp F D))) (F .F-ob c) - f = ⋆Iso (e J D) (F-Iso {F = F} (LimIso D (L₁ J D) theCLimCone .fst .fst)) - - isConeMorF : isConeMor (L₂ J (funcComp F D) .limCone) (F-cone cc) (f .fst) - isConeMorF = isConeMorSeq (L₂ J (funcComp F D) .limCone) - (F-cone (L₁ J D .limCone)) - (F-cone cc) - (isConeMorE J D) - v F-triangle F (limArrowCommutes theCLimCone _ _ _)) - - -- TODO: prove that right adjoints preserve limits +{-# OPTIONS --safe #-} +module Cubical.Categories.Limits.Limits where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Function + +open import Cubical.Data.Sigma + +open import Cubical.Categories.Category +open import Cubical.Categories.Isomorphism +open import Cubical.Categories.Functor +open import Cubical.Categories.NaturalTransformation + +open import Cubical.Categories.Limits.Initial + +open import Cubical.HITs.PropositionalTruncation.Base + +module _ {ℓJ ℓJ' ℓC ℓC' : Level} {J : Category ℓJ ℓJ'} {C : Category ℓC ℓC'} where + open Category + open Functor + open NatTrans + + private + = ℓ-max (ℓ-max (ℓ-max ℓJ ℓJ') ℓC) ℓC' + + record Cone (D : Functor J C) (c : ob C) : Type (ℓ-max (ℓ-max ℓJ ℓJ') ℓC') where + constructor cone + + field + coneOut : (v : ob J) C [ c , F-ob D v ] + coneOutCommutes : {u v : ob J} (e : J [ u , v ]) + coneOut u ⋆⟨ C D .F-hom e coneOut v + + open Cone + + cone≡ : {D : Functor J C} {c : ob C} {c1 c2 : Cone D c} + ((v : ob J) coneOut c1 v coneOut c2 v) c1 c2 + coneOut (cone≡ h i) v = h v i + coneOutCommutes (cone≡ {D} {_} {c1} {c2} h i) {u} {v} e = + isProp→PathP j isSetHom C (h u j ⋆⟨ C D .F-hom e) (h v j)) + (coneOutCommutes c1 e) (coneOutCommutes c2 e) i + + -- dependent versions + conePathP : {D₁ D₂ : Functor J C} {c₁ c₂ : ob C} {cc₁ : Cone D₁ c₁} {cc₂ : Cone D₂ c₂} + {p : D₁ D₂} {q : c₁ c₂} + (∀ v PathP i C [ q i , p i .F-ob v ]) (cc₁ .coneOut v) (cc₂ .coneOut v)) + PathP i Cone (p i) (q i)) cc₁ cc₂ + coneOut (conePathP coneOutPathP i) v = coneOutPathP v i + coneOutCommutes (conePathP {cc₁ = cc₁} {cc₂ = cc₂} {p = p} coneOutPathP i) {u} {v} e = + isProp→PathP j isSetHom C (coneOutPathP u j ⋆⟨ C p j .F-hom e) + (coneOutPathP v j)) + (coneOutCommutes cc₁ e) (coneOutCommutes cc₂ e) i + + conePathPOb : {D : Functor J C} {c c' : ob C} {c1 : Cone D c} {c2 : Cone D c'} {p : c c'} + (∀ (v : ob J) PathP i C [ p i , F-ob D v ]) (coneOut c1 v) (coneOut c2 v)) + PathP i Cone D (p i)) c1 c2 + conePathPOb coneOutPathP = conePathP {p = refl} coneOutPathP + + conePathPDiag : {D₁ D₂ : Functor J C} {c : ob C} {cc₁ : Cone D₁ c} {cc₂ : Cone D₂ c} {p : D₁ D₂} + (∀ v PathP i C [ c , p i .F-ob v ]) (cc₁ .coneOut v) (cc₂ .coneOut v)) + PathP i Cone (p i) c) cc₁ cc₂ + conePathPDiag coneOutPathP = conePathP {q = refl} coneOutPathP + + + -- TODO: can we automate this a bit? + isSetCone : (D : Functor J C) (c : ob C) isSet (Cone D c) + isSetCone D c = isSetRetract toConeΣ fromConeΣ _ refl) + (isSetΣSndProp (isSetΠ _ isSetHom C)) + _ isPropImplicitΠ2 _ _ isPropΠ _ isSetHom C _ _)))) + where + ConeΣ = Σ[ f ((v : ob J) C [ c , F-ob D v ]) ] + ({u v : ob J} (e : J [ u , v ]) f u ⋆⟨ C D .F-hom e f v) + + toConeΣ : Cone D c ConeΣ + fst (toConeΣ x) = coneOut x + snd (toConeΣ x) = coneOutCommutes x + + fromConeΣ : ConeΣ Cone D c + coneOut (fromConeΣ x) = fst x + coneOutCommutes (fromConeΣ x) = snd x + + preCompCone : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} + Cone D c2 Cone D c1 + coneOut (preCompCone f cc) v = f ⋆⟨ C coneOut cc v + coneOutCommutes (preCompCone f cc) e = ⋆Assoc C _ _ _ + cong x f ⋆⟨ C x) (coneOutCommutes cc e) + + _★_ : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} Cone D c2 Cone D c1 + _★_ = preCompCone + + natTransPostCompCone : {c : ob C} {D₁ D₂ : Functor J C} (α : NatTrans D₁ D₂) + Cone D₁ c Cone D₂ c + coneOut (natTransPostCompCone α cc) u = cc .coneOut u ⋆⟨ C α .N-ob u + coneOutCommutes (natTransPostCompCone {D₁ = D₁} {D₂ = D₂} α cc) {u = u} {v = v} e = + cc .coneOut u ⋆⟨ C α .N-ob u ⋆⟨ C D₂ .F-hom e + ≡⟨ ⋆Assoc C _ _ _ + cc .coneOut u ⋆⟨ C (α .N-ob u ⋆⟨ C D₂ .F-hom e) + ≡⟨ cong x cc .coneOut u ⋆⟨ C x) (sym (α .N-hom e)) + cc .coneOut u ⋆⟨ C (D₁ .F-hom e ⋆⟨ C α .N-ob v) + ≡⟨ sym (⋆Assoc C _ _ _) + cc .coneOut u ⋆⟨ C D₁ .F-hom e ⋆⟨ C α .N-ob v + ≡⟨ cong x x ⋆⟨ C α .N-ob v) (cc .coneOutCommutes e) + cc .coneOut v ⋆⟨ C α .N-ob v + + _★ₙₜ_ : {c : ob C} {D₁ D₂ : Functor J C} Cone D₁ c NatTrans D₁ D₂ Cone D₂ c + _★ₙₜ_ = flip natTransPostCompCone + + isConeMor : {c1 c2 : ob C} {D : Functor J C} + (cc1 : Cone D c1) (cc2 : Cone D c2) + C [ c1 , c2 ] Type (ℓ-max ℓJ ℓC') + isConeMor cc1 cc2 f = (v : ob J) f ⋆⟨ C coneOut cc2 v coneOut cc1 v + + isPropIsConeMor : {c1 c2 : ob C} {D : Functor J C} + (cc1 : Cone D c1) (cc2 : Cone D c2) (f : C [ c1 , c2 ]) + isProp (isConeMor cc1 cc2 f) + isPropIsConeMor cc1 cc2 f = isPropΠ _ isSetHom C _ _) + + isConeMorId : {c : ob C} {D : Functor J C} (cc : Cone D c) + isConeMor cc cc (id C) + isConeMorId _ _ = ⋆IdL C _ + + isConeMorSeq : {c1 c2 c3 : ob C} {D : Functor J C} + (cc1 : Cone D c1) (cc2 : Cone D c2) (cc3 : Cone D c3) + {f : C [ c1 , c2 ]} {g : C [ c2 , c3 ]} + isConeMor cc1 cc2 f isConeMor cc2 cc3 g isConeMor cc1 cc3 (f ⋆⟨ C g) + isConeMorSeq cc1 cc2 cc3 {f} {g} isConeMorF isConeMorG v = + ⋆Assoc C _ _ _ ∙∙ cong x f ⋆⟨ C x) (isConeMorG v) ∙∙ isConeMorF v + + preCompConeMor : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} (cc : Cone D c2) + isConeMor (f cc) cc f + preCompConeMor f cc v = refl + + isLimCone : (D : Functor J C) (c0 : ob C) Cone D c0 Type + isLimCone D c0 cc0 = (c : ob C) (cc : Cone D c) + ∃![ f C [ c , c0 ] ] isConeMor cc cc0 f + + isPropIsLimCone : (D : Functor J C) (c0 : ob C) (cc0 : Cone D c0) + isProp (isLimCone D c0 cc0) + isPropIsLimCone D c0 cc0 = isPropΠ2 λ _ _ isProp∃! + + preCompUnique : {c1 c2 : ob C} (f : C [ c1 , c2 ]) {D : Functor J C} (cc : Cone D c2) + isLimCone D c2 cc + (g : C [ c1 , c2 ]) isConeMor (f cc) cc g g f + preCompUnique f cc ccIsLimCone g gIsConeMor = + cong fst (isContr→isProp (ccIsLimCone _ (f cc)) (g , gIsConeMor) (f , preCompConeMor f cc)) + + record LimCone (D : Functor J C) : Type where + constructor climCone + + field + lim : ob C + limCone : Cone D lim + univProp : isLimCone D lim limCone + + limOut : (v : ob J) C [ lim , D .F-ob v ] + limOut = coneOut limCone + + limOutCommutes : {u v : ob J} (e : J [ u , v ]) + limOut u ⋆⟨ C D .F-hom e limOut v + limOutCommutes = coneOutCommutes limCone + + limArrow : (c : ob C) (cc : Cone D c) C [ c , lim ] + limArrow c cc = univProp c cc .fst .fst + + limArrowCommutes : (c : ob C) (cc : Cone D c) (u : ob J) + limArrow c cc ⋆⟨ C limOut u coneOut cc u + limArrowCommutes c cc = univProp c cc .fst .snd + + limArrowUnique : (c : ob C) (cc : Cone D c) (k : C [ c , lim ]) + isConeMor cc limCone k limArrow c cc k + limArrowUnique c cc k hk = cong fst (univProp c cc .snd (k , hk)) + + open LimCone + limOfArrowsCone : {D₁ D₂ : Functor J C} + (CC₁ : LimCone D₁) + NatTrans D₁ D₂ + Cone D₂ (CC₁ .lim) + coneOut (limOfArrowsCone {D₁} {D₂} CC₁ α) v = limOut CC₁ v ⋆⟨ C α .N-ob v + coneOutCommutes (limOfArrowsCone {D₁} {D₂} CC₁ α) {u = u} {v = v} e = + limOut CC₁ u ⋆⟨ C α .N-ob u ⋆⟨ C D₂ .F-hom e ≡⟨ ⋆Assoc C _ _ _ + limOut CC₁ u ⋆⟨ C (α .N-ob u ⋆⟨ C D₂ .F-hom e) ≡⟨ cong x seq' C (limOut CC₁ u) x) (sym (α .N-hom e)) + limOut CC₁ u ⋆⟨ C (D₁ .F-hom e ⋆⟨ C α .N-ob v) ≡⟨ sym (⋆Assoc C _ _ _) + limOut CC₁ u ⋆⟨ C D₁ .F-hom e ⋆⟨ C α .N-ob v ≡⟨ cong x x ⋆⟨ C α .N-ob v) (limOutCommutes CC₁ e) + limOut CC₁ v ⋆⟨ C α .N-ob v + + limOfArrows : {D₁ D₂ : Functor J C} + (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) + NatTrans D₁ D₂ + C [ CC₁ .lim , CC₂ .lim ] + limOfArrows {D₁} {D₂} CC₁ CC₂ α = limArrow CC₂ (CC₁ .lim) (limOfArrowsCone CC₁ α) + + limOfArrowsOut : {D₁ D₂ : Functor J C} + (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) + (α : NatTrans D₁ D₂) (u : ob J) + limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u limOut CC₁ u ⋆⟨ C α .N-ob u + limOfArrowsOut _ CC₂ _ _ = limArrowCommutes CC₂ _ _ _ + + limOfArrowsId : {D : Functor J C} (CC : LimCone D) + limOfArrows CC CC (idTrans D) id C + limOfArrowsId CC = limArrowUnique CC _ _ _ λ v ⋆IdL C _ sym (⋆IdR C _) + + limOfArrowsSeq : {D₁ D₂ D₃ : Functor J C} + (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) (CC₃ : LimCone D₃) + (α : NatTrans D₁ D₂) (β : NatTrans D₂ D₃) + limOfArrows CC₁ CC₃ (α ●ᵛ β) + limOfArrows CC₁ CC₂ α ⋆⟨ C limOfArrows CC₂ CC₃ β + limOfArrowsSeq CC₁ CC₂ CC₃ α β = limArrowUnique CC₃ _ _ _ path + where + path : u + (limOfArrows CC₁ CC₂ α ⋆⟨ C limOfArrows CC₂ CC₃ β) ⋆⟨ C limOut CC₃ u + limOut CC₁ u ⋆⟨ C (α .N-ob u ⋆⟨ C β .N-ob u) + path u = (limOfArrows CC₁ CC₂ α ⋆⟨ C limOfArrows CC₂ CC₃ β) ⋆⟨ C limOut CC₃ u + ≡⟨ ⋆Assoc C _ _ _ + limOfArrows CC₁ CC₂ α ⋆⟨ C (limOfArrows CC₂ CC₃ β ⋆⟨ C limOut CC₃ u) + ≡⟨ cong x limOfArrows CC₁ CC₂ α ⋆⟨ C x) (limOfArrowsOut CC₂ CC₃ β u) + limOfArrows CC₁ CC₂ α ⋆⟨ C (limOut CC₂ u ⋆⟨ C β .N-ob u) + ≡⟨ sym (⋆Assoc C _ _ _) + (limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u) ⋆⟨ C β .N-ob u + ≡⟨ cong x x ⋆⟨ C β .N-ob u) (limOfArrowsOut CC₁ CC₂ α u) + (limOut CC₁ u ⋆⟨ C α .N-ob u) ⋆⟨ C β .N-ob u + ≡⟨ ⋆Assoc C _ _ _ + limOut CC₁ u ⋆⟨ C (α .N-ob u ⋆⟨ C β .N-ob u) + + limArrowCompLimOfArrows : {D₁ D₂ : Functor J C} + (CC₁ : LimCone D₁) (CC₂ : LimCone D₂) + (α : NatTrans D₁ D₂) + (c : ob C) (cc : Cone D₁ c) + limArrow CC₂ _ (cc ★ₙₜ α) limArrow CC₁ _ cc ⋆⟨ C limOfArrows CC₁ CC₂ α + limArrowCompLimOfArrows CC₁ CC₂ α c cc = limArrowUnique CC₂ _ _ _ isConeMorComp + where + isConeMorComp : (u : ob J) + limArrow CC₁ _ cc ⋆⟨ C limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u + cc .coneOut u ⋆⟨ C α .N-ob u + isConeMorComp u = + limArrow CC₁ _ cc ⋆⟨ C limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u + ≡⟨ ⋆Assoc C _ _ _ + limArrow CC₁ _ cc ⋆⟨ C (limOfArrows CC₁ CC₂ α ⋆⟨ C limOut CC₂ u) + ≡⟨ cong x limArrow CC₁ _ cc ⋆⟨ C x) (limOfArrowsOut CC₁ CC₂ α u) + limArrow CC₁ _ cc ⋆⟨ C (limOut CC₁ u ⋆⟨ C α .N-ob u) + ≡⟨ sym (⋆Assoc C _ _ _) + limArrow CC₁ _ cc ⋆⟨ C limOut CC₁ u ⋆⟨ C α .N-ob u + ≡⟨ cong x x ⋆⟨ C α .N-ob u) (limArrowCommutes CC₁ _ cc u) + cc .coneOut u ⋆⟨ C α .N-ob u + + -- any two limits are isomorphic up to a unique cone isomorphism + open isIso + LimIso : (D : Functor J C) (L₁ L₂ : LimCone D) + ∃![ e CatIso C (lim L₁) (lim L₂) ] isConeMor (limCone L₁) (limCone L₂) (fst e) + fst (fst (fst (LimIso D L₁ L₂))) = limArrow L₂ _ (limCone L₁) + inv (snd (fst (fst (LimIso D L₁ L₂)))) = limArrow L₁ _ (limCone L₂) + sec (snd (fst (fst (LimIso D L₁ L₂)))) = cong fst (isContr→isProp (univProp L₂ _ (limCone L₂)) + (_ , isConeMorComp) (id C , isConeMorId (limCone L₂))) + where + isConeMorComp : isConeMor (limCone L₂) (limCone L₂) + (limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C limArrow L₂ (lim L₁) (limCone L₁)) + isConeMorComp v = + (limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C limArrow L₂ (lim L₁) (limCone L₁)) + ⋆⟨ C coneOut (limCone L₂) v + ≡⟨ ⋆Assoc C _ _ _ + limArrow L₁ (lim L₂) (limCone L₂) + ⋆⟨ C (limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C coneOut (limCone L₂) v) + ≡⟨ cong x limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C x) + (limArrowCommutes L₂ _ (limCone L₁) v) + limArrow L₁ (lim L₂) (limCone L₂) + ⋆⟨ C (coneOut (limCone L₁) v) + ≡⟨ limArrowCommutes L₁ _ (limCone L₂) v + coneOut (limCone L₂) v + ret (snd (fst (fst (LimIso D L₁ L₂)))) = cong fst (isContr→isProp (univProp L₁ _ (limCone L₁)) + (_ , isConeMorComp) (id C , isConeMorId (limCone L₁))) + where + isConeMorComp : isConeMor (limCone L₁) (limCone L₁) + (limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C limArrow L₁ (lim L₂) (limCone L₂)) + isConeMorComp v = + (limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C limArrow L₁ (lim L₂) (limCone L₂)) + ⋆⟨ C coneOut (limCone L₁) v + ≡⟨ ⋆Assoc C _ _ _ + limArrow L₂ (lim L₁) (limCone L₁) + ⋆⟨ C (limArrow L₁ (lim L₂) (limCone L₂) ⋆⟨ C coneOut (limCone L₁) v) + ≡⟨ cong x limArrow L₂ (lim L₁) (limCone L₁) ⋆⟨ C x) + (limArrowCommutes L₁ _ (limCone L₂) v) + limArrow L₂ (lim L₁) (limCone L₁) + ⋆⟨ C (coneOut (limCone L₂) v) + ≡⟨ limArrowCommutes L₂ _ (limCone L₁) v + coneOut (limCone L₁) v + + snd (fst (LimIso D L₁ L₂)) = limArrowCommutes L₂ _ _ + snd (LimIso D L₁ L₂) (e , isConeMorE) = Σ≡Prop + _ isPropIsConeMor (limCone L₁) (limCone L₂) _) + (CatIso≡ _ _ (limArrowUnique L₂ _ _ (fst e) isConeMorE)) + + -- if the index category of the diagram has an initial object, + -- we get a canonical limiting cone + Initial→LimCone : (D : Functor J C) Initial J LimCone D + lim (Initial→LimCone D (j , isInitJ)) = D .F-ob j + coneOut (limCone (Initial→LimCone D (j , isInitJ))) k = D .F-hom (isInitJ k .fst) + coneOutCommutes (limCone (Initial→LimCone D (j , isInitJ))) f = + sym (D .F-seq _ _) cong (D .F-hom) (sym (isInitJ _ .snd _)) + fst (fst (univProp (Initial→LimCone D (j , isInitJ)) c cc)) = cc .coneOut j + -- canonical arrow c → D(j) + snd (fst (univProp (Initial→LimCone D (j , isInitJ)) c cc)) k = + cc .coneOutCommutes (isInitJ k .fst) + -- is a cone morphism + snd (univProp (Initial→LimCone D (j , isInitJ)) c cc) (f , isConeMorF) = -- and indeed unique + Σ≡Prop + _ isPropIsConeMor cc (limCone (Initial→LimCone D (j , isInitJ))) _) + (sym (isConeMorF j) ∙∙ cong x f ⋆⟨ C x) (subst x D .F-hom x id C) + (sym (isInitJ j .snd _)) (D .F-id)) + ∙∙ ⋆IdR C f) + + -- cones that respect isomorphisms are limiting cones + Iso→LimCone : {D : Functor J C} {c₁ c₂ : ob C} (cc₁ : Cone D c₁) (e : CatIso C c₁ c₂) + isLimCone _ _ cc₁ + (cc₂ : Cone D c₂) + isConeMor cc₁ cc₂ (e .fst) + isLimCone _ _ cc₂ + fst (fst (Iso→LimCone cc₁ e isLimConeCC₁ cc₂ isConeMorE d cd)) = + isLimConeCC₁ d cd .fst .fst ⋆⟨ C e .fst + snd (fst (Iso→LimCone cc₁ e isLimConeCC₁ cc₂ isConeMorE d cd)) = + isConeMorSeq cd cc₁ cc₂ (isLimConeCC₁ d cd .fst .snd) isConeMorE + snd (Iso→LimCone cc₁ e isLimConeCC₁ cc₂ isConeMorE d cd) (f , isConeMorF) = + Σ≡Prop (isPropIsConeMor cd cc₂) path + where + isConeMorE⁻¹ : isConeMor cc₂ cc₁ (e .snd .inv) + isConeMorE⁻¹ v = + e .snd .inv ⋆⟨ C coneOut cc₁ v ≡⟨ cong x e .snd .inv ⋆⟨ C x) (sym (isConeMorE v)) + e .snd .inv ⋆⟨ C (e .fst ⋆⟨ C coneOut cc₂ v) ≡⟨ sym (⋆Assoc C _ _ _) + (e .snd .inv ⋆⟨ C e .fst) ⋆⟨ C coneOut cc₂ v ≡⟨ cong x x ⋆⟨ C coneOut cc₂ v) + (e .snd .sec) + id C ⋆⟨ C coneOut cc₂ v ≡⟨ ⋆IdL C _ + coneOut cc₂ v + + p : isLimConeCC₁ d cd .fst .fst f ⋆⟨ C e .snd .inv + p = cong fst (isLimConeCC₁ d cd .snd ( f ⋆⟨ C e .snd .inv + , isConeMorSeq cd cc₂ cc₁ isConeMorF isConeMorE⁻¹)) + + path : isLimConeCC₁ d cd .fst .fst ⋆⟨ C e .fst f + path = isLimConeCC₁ d cd .fst .fst ⋆⟨ C e .fst ≡⟨ cong x x ⋆⟨ C e .fst) p + (f ⋆⟨ C e .snd .inv) ⋆⟨ C e .fst ≡⟨ ⋆Assoc C _ _ _ + f ⋆⟨ C (e .snd .inv ⋆⟨ C e .fst) ≡⟨ cong x f ⋆⟨ C x) (e .snd .sec) + f ⋆⟨ C id C ≡⟨ ⋆IdR C _ + f + +-- A category is complete if it has all limits +Limits : {ℓJ ℓJ' ℓC ℓC' : Level} Category ℓC ℓC' Type _ +Limits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') (D : Functor J C) LimCone D + +hasLimits : {ℓJ ℓJ' ℓC ℓC' : Level} Category ℓC ℓC' Type _ +hasLimits {ℓJ} {ℓJ'} C = (J : Category ℓJ ℓJ') (D : Functor J C) LimCone D ∥₁ + +-- Limits of a specific shape J in a category C +LimitsOfShape : {ℓJ ℓJ' ℓC ℓC' : Level} Category ℓJ ℓJ' Category ℓC ℓC' Type _ +LimitsOfShape J C = (D : Functor J C) LimCone D + + +-- precomposition with a functor and preservation of limits +module _ {ℓJ ℓJ' ℓC1 ℓC1' ℓC2 ℓC2' : Level} + {C1 : Category ℓC1 ℓC1'} {C2 : Category ℓC2 ℓC2'} + (F : Functor C1 C2) where + open Category + open Functor + open Cone + + -- same as F-cone!!! + F-cone : {J : Category ℓJ ℓJ'} {D : Functor J C1} {x : ob C1} + Cone D x Cone (funcComp F D) (F .F-ob x) + coneOut (F-cone ccx) v = F .F-hom (coneOut ccx v) + coneOutCommutes (F-cone ccx) e = + sym (F-seq F (coneOut ccx _) _) cong (F .F-hom) (coneOutCommutes ccx e) + + preservesLimits : Type _ + preservesLimits = {J : Category ℓJ ℓJ'} {D : Functor J C1} {L : ob C1} + (ccL : Cone D L) + isLimCone _ _ ccL isLimCone (funcComp F D) (F .F-ob L) (F-cone ccL) + + -- characterizing limit preserving functors + open LimCone + + preservesLimitsChar : (L₁ : Limits {ℓJ} {ℓJ'} C1) (L₂ : Limits {ℓJ} {ℓJ'} C2) + (e : J D CatIso C2 (L₂ J (F ∘F D) .lim) (F .F-ob (L₁ J D .lim))) + (∀ J D isConeMor (L₂ J (F ∘F D) .limCone) (F-cone (L₁ J D .limCone)) (e J D .fst)) + preservesLimits + preservesLimitsChar L₁ L₂ e isConeMorE {J = J} {D = D} {L = c} cc isLimConeCC = + Iso→LimCone (L₂ J (F ∘F D) .limCone) f (L₂ J (F ∘F D) .univProp) (F-cone cc) isConeMorF + where + theCLimCone : LimCone D + lim theCLimCone = c + limCone theCLimCone = cc + univProp theCLimCone = isLimConeCC + + f : CatIso C2 (lim (L₂ J (funcComp F D))) (F .F-ob c) + f = ⋆Iso (e J D) (F-Iso {F = F} (LimIso D (L₁ J D) theCLimCone .fst .fst)) + + isConeMorF : isConeMor (L₂ J (funcComp F D) .limCone) (F-cone cc) (f .fst) + isConeMorF = isConeMorSeq (L₂ J (funcComp F D) .limCone) + (F-cone (L₁ J D .limCone)) + (F-cone cc) + (isConeMorE J D) + v F-triangle F (limArrowCommutes theCLimCone _ _ _)) + + -- TODO: prove that right adjoints preserve limits
\ No newline at end of file diff --git a/docs/Cubical.Categories.Limits.Pullback.html b/docs/Cubical.Categories.Limits.Pullback.html index 02501cb..4152123 100644 --- a/docs/Cubical.Categories.Limits.Pullback.html +++ b/docs/Cubical.Categories.Limits.Pullback.html @@ -1,155 +1,155 @@ -Cubical.Categories.Limits.Pullback
{-# OPTIONS --allow-unsolved-metas #-}
-module Cubical.Categories.Limits.Pullback where
-
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.HITs.PropositionalTruncation.Base
-
-open import Cubical.Data.Sigma
-open import Cubical.Data.Unit
-
-open import Cubical.Categories.Category
-open import Cubical.Categories.Functor
-open import Cubical.Categories.Instances.Cospan
-open import Cubical.Categories.Limits.Limits
-
-private
-  variable
-     ℓ' : Level
-
-module _ (C : Category  ℓ') where
-
-  open Category C
-  open Functor
-
-  record Cospan : Type (ℓ-max  ℓ') where
-    constructor cospan
-    field
-      l m r : ob
-      s₁ : C [ l , m ]
-      s₂ : C [ r , m ]
-
-  open Cospan
-
-  isPullback : (cspn : Cospan) 
-    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
-    (H : p₁  cspn .s₁  p₂  cspn .s₂)  Type (ℓ-max  ℓ')
-  isPullback cspn {c} p₁ p₂ H =
-     {d} (h : C [ d , cspn .l ]) (k : C [ d , cspn .r ])
-          (H' : h  cspn .s₁  k  cspn .s₂)
-     ∃![ hk  C [ d , c ] ] (h  hk  p₁) × (k  hk  p₂)
-
-  isPropIsPullback : (cspn : Cospan) 
-    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
-    (H : p₁  cspn .s₁  p₂  cspn .s₂)  isProp (isPullback cspn p₁ p₂ H)
-  isPropIsPullback cspn p₁ p₂ H =
-    isPropImplicitΠ  x  isPropΠ3 λ h k H'  isPropIsContr)
-
-  record Pullback (cspn : Cospan) : Type (ℓ-max  ℓ') where
-    field
-      pbOb  : ob
-      pbPr₁ : C [ pbOb , cspn .l ]
-      pbPr₂ : C [ pbOb , cspn .r ]
-      pbCommutes : pbPr₁  cspn .s₁  pbPr₂  cspn .s₂
-      univProp : isPullback cspn pbPr₁ pbPr₂ pbCommutes
-
-  open Pullback
-
-  pullbackArrow :
-    {cspn : Cospan} (pb : Pullback cspn)
-    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
-    (H : p₁  cspn .s₁  p₂  cspn .s₂)  C [ c , pb . pbOb ]
-  pullbackArrow pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .fst
-
-  pullbackArrowPr₁ :
-    {cspn : Cospan} (pb : Pullback cspn)
-    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
-    (H : p₁  cspn .s₁  p₂  cspn .s₂) 
-    p₁  pullbackArrow pb p₁ p₂ H  pbPr₁ pb
-  pullbackArrowPr₁ pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .snd .fst
-
-  pullbackArrowPr₂ :
-    {cspn : Cospan} (pb : Pullback cspn)
-    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
-    (H : p₁  cspn .s₁  p₂  cspn .s₂) 
-    p₂  pullbackArrow pb p₁ p₂ H  pbPr₂ pb
-  pullbackArrowPr₂ pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .snd .snd
-
-  pullbackArrowUnique :
-    {cspn : Cospan} (pb : Pullback cspn)
-    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
-    (H : p₁  cspn .s₁  p₂  cspn .s₂)
-    (pbArrow' : C [ c , pb .pbOb ])
-    (H₁ : p₁  pbArrow'  pbPr₁ pb) (H₂ : p₂  pbArrow'  pbPr₂ pb)
-     pullbackArrow pb p₁ p₂ H  pbArrow'
-  pullbackArrowUnique pb p₁ p₂ H pbArrow' H₁ H₂ =
-    cong fst (pb .univProp p₁ p₂ H .snd (pbArrow' , (H₁ , H₂)))
-
-  Pullbacks : Type (ℓ-max  ℓ')
-  Pullbacks = (cspn : Cospan)  Pullback cspn
-
-  hasPullbacks : Type (ℓ-max  ℓ')
-  hasPullbacks =  Pullbacks ∥₁
-
-
--- Pullbacks from limits
-module _ {C : Category  ℓ'} where
-  open Category C
-  open Functor
-  open Pullback
-  open LimCone
-  open Cone
-  open Cospan
-
-  Cospan→Func : Cospan C  Functor CospanCat C
-  Cospan→Func (cospan l m r f g) .F-ob  = l
-  Cospan→Func (cospan l m r f g) .F-ob  = m
-  Cospan→Func (cospan l m r f g) .F-ob  = r
-  Cospan→Func (cospan l m r f g) .F-hom {} {} k = f
-  Cospan→Func (cospan l m r f g) .F-hom {} {} k = g
-  Cospan→Func (cospan l m r f g) .F-hom {} {} k = id
-  Cospan→Func (cospan l m r f g) .F-hom {} {} k = id
-  Cospan→Func (cospan l m r f g) .F-hom {} {} k = id
-  Cospan→Func (cospan l m r f g) .F-id {} = refl
-  Cospan→Func (cospan l m r f g) .F-id {} = refl
-  Cospan→Func (cospan l m r f g) .F-id {} = refl
-  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
-  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
-  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdR _)
-  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
-  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
-  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
-  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdR _)
-
-  LimitsOfShapeCospanCat→Pullbacks : LimitsOfShape CospanCat C  Pullbacks C
-  pbOb (LimitsOfShapeCospanCat→Pullbacks H cspn) = lim (H (Cospan→Func cspn))
-  pbPr₁ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) 
-  pbPr₂ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) 
-  pbCommutes (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOutCommutes (H (Cospan→Func cspn)) {v = } tt
-                           sym (limOutCommutes (H (Cospan→Func cspn)) tt)
-  univProp (LimitsOfShapeCospanCat→Pullbacks H cspn) {d = d} h k H' =
-    uniqueExists (limArrow (H (Cospan→Func cspn)) d cc)
-                 ( sym (limArrowCommutes (H (Cospan→Func cspn)) d cc )
-                 , sym (limArrowCommutes (H (Cospan→Func cspn)) d cc ))
-                  _  isProp× (isSetHom _ _) (isSetHom _ _))
-                 λ a' ha'  limArrowUnique (H (Cospan→Func cspn)) d cc a'
-                                {   sym (ha' .fst)
-                                  ;   cong (a' ⋆_) (sym (limOutCommutes (H (Cospan→Func cspn)) {} {} tt))
-                                      ∙∙ sym (⋆Assoc _ _ _)
-                                      ∙∙ cong (_⋆ cspn .s₁) (sym (ha' .fst))
-                                  ;   sym (ha' .snd) })
-    where
-    cc : Cone (Cospan→Func cspn) d
-    coneOut cc  = h
-    coneOut cc  = h  cspn .s₁
-    coneOut cc  = k
-    coneOutCommutes cc {} {} e = ⋆IdR h
-    coneOutCommutes cc {} {} e = refl
-    coneOutCommutes cc {} {} e = ⋆IdR _
-    coneOutCommutes cc {} {} e = sym H'
-    coneOutCommutes cc {} {} e = ⋆IdR k
-
-  Limits→Pullbacks : Limits C  Pullbacks C
-  Limits→Pullbacks H = LimitsOfShapeCospanCat→Pullbacks (H CospanCat)
+Cubical.Categories.Limits.Pullback
{-# OPTIONS --safe #-}
+module Cubical.Categories.Limits.Pullback where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.HITs.PropositionalTruncation.Base
+
+open import Cubical.Data.Sigma
+open import Cubical.Data.Unit
+
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Instances.Cospan
+open import Cubical.Categories.Limits.Limits
+
+private
+  variable
+     ℓ' : Level
+
+module _ (C : Category  ℓ') where
+
+  open Category C
+  open Functor
+
+  record Cospan : Type (ℓ-max  ℓ') where
+    constructor cospan
+    field
+      l m r : ob
+      s₁ : C [ l , m ]
+      s₂ : C [ r , m ]
+
+  open Cospan
+
+  isPullback : (cspn : Cospan) 
+    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
+    (H : p₁  cspn .s₁  p₂  cspn .s₂)  Type (ℓ-max  ℓ')
+  isPullback cspn {c} p₁ p₂ H =
+     {d} (h : C [ d , cspn .l ]) (k : C [ d , cspn .r ])
+          (H' : h  cspn .s₁  k  cspn .s₂)
+     ∃![ hk  C [ d , c ] ] (h  hk  p₁) × (k  hk  p₂)
+
+  isPropIsPullback : (cspn : Cospan) 
+    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
+    (H : p₁  cspn .s₁  p₂  cspn .s₂)  isProp (isPullback cspn p₁ p₂ H)
+  isPropIsPullback cspn p₁ p₂ H =
+    isPropImplicitΠ  x  isPropΠ3 λ h k H'  isPropIsContr)
+
+  record Pullback (cspn : Cospan) : Type (ℓ-max  ℓ') where
+    field
+      pbOb  : ob
+      pbPr₁ : C [ pbOb , cspn .l ]
+      pbPr₂ : C [ pbOb , cspn .r ]
+      pbCommutes : pbPr₁  cspn .s₁  pbPr₂  cspn .s₂
+      univProp : isPullback cspn pbPr₁ pbPr₂ pbCommutes
+
+  open Pullback
+
+  pullbackArrow :
+    {cspn : Cospan} (pb : Pullback cspn)
+    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
+    (H : p₁  cspn .s₁  p₂  cspn .s₂)  C [ c , pb . pbOb ]
+  pullbackArrow pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .fst
+
+  pullbackArrowPr₁ :
+    {cspn : Cospan} (pb : Pullback cspn)
+    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
+    (H : p₁  cspn .s₁  p₂  cspn .s₂) 
+    p₁  pullbackArrow pb p₁ p₂ H  pbPr₁ pb
+  pullbackArrowPr₁ pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .snd .fst
+
+  pullbackArrowPr₂ :
+    {cspn : Cospan} (pb : Pullback cspn)
+    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
+    (H : p₁  cspn .s₁  p₂  cspn .s₂) 
+    p₂  pullbackArrow pb p₁ p₂ H  pbPr₂ pb
+  pullbackArrowPr₂ pb p₁ p₂ H = pb .univProp p₁ p₂ H .fst .snd .snd
+
+  pullbackArrowUnique :
+    {cspn : Cospan} (pb : Pullback cspn)
+    {c : ob} (p₁ : C [ c , cspn .l ]) (p₂ : C [ c , cspn .r ])
+    (H : p₁  cspn .s₁  p₂  cspn .s₂)
+    (pbArrow' : C [ c , pb .pbOb ])
+    (H₁ : p₁  pbArrow'  pbPr₁ pb) (H₂ : p₂  pbArrow'  pbPr₂ pb)
+     pullbackArrow pb p₁ p₂ H  pbArrow'
+  pullbackArrowUnique pb p₁ p₂ H pbArrow' H₁ H₂ =
+    cong fst (pb .univProp p₁ p₂ H .snd (pbArrow' , (H₁ , H₂)))
+
+  Pullbacks : Type (ℓ-max  ℓ')
+  Pullbacks = (cspn : Cospan)  Pullback cspn
+
+  hasPullbacks : Type (ℓ-max  ℓ')
+  hasPullbacks =  Pullbacks ∥₁
+
+
+-- Pullbacks from limits
+module _ {C : Category  ℓ'} where
+  open Category C
+  open Functor
+  open Pullback
+  open LimCone
+  open Cone
+  open Cospan
+
+  Cospan→Func : Cospan C  Functor CospanCat C
+  Cospan→Func (cospan l m r f g) .F-ob  = l
+  Cospan→Func (cospan l m r f g) .F-ob  = m
+  Cospan→Func (cospan l m r f g) .F-ob  = r
+  Cospan→Func (cospan l m r f g) .F-hom {} {} k = f
+  Cospan→Func (cospan l m r f g) .F-hom {} {} k = g
+  Cospan→Func (cospan l m r f g) .F-hom {} {} k = id
+  Cospan→Func (cospan l m r f g) .F-hom {} {} k = id
+  Cospan→Func (cospan l m r f g) .F-hom {} {} k = id
+  Cospan→Func (cospan l m r f g) .F-id {} = refl
+  Cospan→Func (cospan l m r f g) .F-id {} = refl
+  Cospan→Func (cospan l m r f g) .F-id {} = refl
+  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
+  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
+  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdR _)
+  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
+  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
+  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdL _)
+  Cospan→Func (cospan l m r f g) .F-seq {} {} {} φ ψ = sym (⋆IdR _)
+
+  LimitsOfShapeCospanCat→Pullbacks : LimitsOfShape CospanCat C  Pullbacks C
+  pbOb (LimitsOfShapeCospanCat→Pullbacks H cspn) = lim (H (Cospan→Func cspn))
+  pbPr₁ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) 
+  pbPr₂ (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOut (H (Cospan→Func cspn)) 
+  pbCommutes (LimitsOfShapeCospanCat→Pullbacks H cspn) = limOutCommutes (H (Cospan→Func cspn)) {v = } tt
+                           sym (limOutCommutes (H (Cospan→Func cspn)) tt)
+  univProp (LimitsOfShapeCospanCat→Pullbacks H cspn) {d = d} h k H' =
+    uniqueExists (limArrow (H (Cospan→Func cspn)) d cc)
+                 ( sym (limArrowCommutes (H (Cospan→Func cspn)) d cc )
+                 , sym (limArrowCommutes (H (Cospan→Func cspn)) d cc ))
+                  _  isProp× (isSetHom _ _) (isSetHom _ _))
+                 λ a' ha'  limArrowUnique (H (Cospan→Func cspn)) d cc a'
+                                {   sym (ha' .fst)
+                                  ;   cong (a' ⋆_) (sym (limOutCommutes (H (Cospan→Func cspn)) {} {} tt))
+                                      ∙∙ sym (⋆Assoc _ _ _)
+                                      ∙∙ cong (_⋆ cspn .s₁) (sym (ha' .fst))
+                                  ;   sym (ha' .snd) })
+    where
+    cc : Cone (Cospan→Func cspn) d
+    coneOut cc  = h
+    coneOut cc  = h  cspn .s₁
+    coneOut cc  = k
+    coneOutCommutes cc {} {} e = ⋆IdR h
+    coneOutCommutes cc {} {} e = refl
+    coneOutCommutes cc {} {} e = ⋆IdR _
+    coneOutCommutes cc {} {} e = sym H'
+    coneOutCommutes cc {} {} e = ⋆IdR k
+
+  Limits→Pullbacks : Limits C  Pullbacks C
+  Limits→Pullbacks H = LimitsOfShapeCospanCat→Pullbacks (H CospanCat)
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Limits.Terminal.html b/docs/Cubical.Categories.Limits.Terminal.html index 13b3edd..1cc508c 100644 --- a/docs/Cubical.Categories.Limits.Terminal.html +++ b/docs/Cubical.Categories.Limits.Terminal.html @@ -1,85 +1,85 @@ -Cubical.Categories.Limits.Terminal
{-# OPTIONS #-}
-module Cubical.Categories.Limits.Terminal where
-
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.HITs.PropositionalTruncation.Base
-open import Cubical.Data.Sigma
-
-open import Cubical.Categories.Category
-open import Cubical.Categories.Functor
-open import Cubical.Categories.Isomorphism
-
-private
-  variable
-     ℓ' ℓc ℓc' ℓd ℓd' : Level
-
-module _ (C : Category  ℓ') where
-  open Category C
-
-  isTerminal : (x : ob)  Type (ℓ-max  ℓ')
-  isTerminal x =  (y : ob)  isContr (C [ y , x ])
-
-  Terminal : Type (ℓ-max  ℓ')
-  Terminal = Σ[ x  ob ] isTerminal x
-
-  terminalOb : Terminal  ob
-  terminalOb = fst
-
-  terminalArrow : (T : Terminal) (y : ob)  C [ y , terminalOb T ]
-  terminalArrow T y = T .snd y .fst
-
-  terminalArrowUnique : {T : Terminal} {y : ob} (f : C [ y , terminalOb T ])
-                       terminalArrow T y  f
-  terminalArrowUnique {T} {y} f = T .snd y .snd f
-
-  terminalEndoIsId : (T : Terminal) (f : C [ terminalOb T , terminalOb T ])
-                    f  id
-  terminalEndoIsId T f = isContr→isProp (T .snd (terminalOb T)) f id
-
-  hasTerminal : Type (ℓ-max  ℓ')
-  hasTerminal =  Terminal ∥₁
-
-  -- Terminality of an object is a proposition.
-  isPropIsTerminal : (x : ob)  isProp (isTerminal x)
-  isPropIsTerminal _ = isPropΠ λ _  isPropIsContr
-
-  open isIso
-
-  -- Objects that are terminal are isomorphic.
-  terminalToIso : (x y : Terminal)  CatIso C (terminalOb x) (terminalOb y)
-  terminalToIso x y .fst = terminalArrow y (terminalOb x)
-  terminalToIso x y .snd .inv = terminalArrow x (terminalOb y)
-  terminalToIso x y .snd .sec = terminalEndoIsId y _
-  terminalToIso x y .snd .ret = terminalEndoIsId x _
-
-  isoToTerminal :  (x : Terminal) y  CatIso C (terminalOb x) y  isTerminal y
-  isoToTerminal x y x≅y y' .fst = x≅y .fst ∘⟨ C  terminalArrow x y'
-  isoToTerminal x y x≅y y' .snd f =
-    sym (⋆InvRMove
-          (invIso x≅y)
-          (sym (terminalArrowUnique {T = x} (invIso x≅y .fst ∘⟨ C  f))))
-
-  open isUnivalent
-
-  -- The type of terminal objects of a univalent category is a proposition,
-  -- i.e. all terminal objects are equal.
-  isPropTerminal : (hC : isUnivalent C)  isProp Terminal
-  isPropTerminal hC x y =
-    Σ≡Prop isPropIsTerminal (CatIsoToPath hC (terminalToIso x y))
-
-preservesTerminals :  (C : Category ℓc ℓc')(D : Category ℓd ℓd')
-                    Functor C D
-                    Type (ℓ-max (ℓ-max (ℓ-max ℓc ℓc') ℓd) ℓd')
-preservesTerminals C D F =  (term : Terminal C)  isTerminal D (F  term .fst )
-
-preserveAnyTerminal→PreservesTerminals :  (C : Category ℓc ℓc')(D : Category ℓd ℓd')
-   (F : Functor C D)
-   (term : Terminal C)  isTerminal D (F  term .fst )
-   preservesTerminals C D F
-preserveAnyTerminal→PreservesTerminals C D F term D-preserves-term term' =
-  isoToTerminal D
-                ((F  term .fst ) , D-preserves-term) (F  term' .fst )
-                (F-Iso {F = F} (terminalToIso C term term'))
+Cubical.Categories.Limits.Terminal
{-# OPTIONS --safe #-}
+module Cubical.Categories.Limits.Terminal where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.HITs.PropositionalTruncation.Base
+open import Cubical.Data.Sigma
+
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Isomorphism
+
+private
+  variable
+     ℓ' ℓc ℓc' ℓd ℓd' : Level
+
+module _ (C : Category  ℓ') where
+  open Category C
+
+  isTerminal : (x : ob)  Type (ℓ-max  ℓ')
+  isTerminal x =  (y : ob)  isContr (C [ y , x ])
+
+  Terminal : Type (ℓ-max  ℓ')
+  Terminal = Σ[ x  ob ] isTerminal x
+
+  terminalOb : Terminal  ob
+  terminalOb = fst
+
+  terminalArrow : (T : Terminal) (y : ob)  C [ y , terminalOb T ]
+  terminalArrow T y = T .snd y .fst
+
+  terminalArrowUnique : {T : Terminal} {y : ob} (f : C [ y , terminalOb T ])
+                       terminalArrow T y  f
+  terminalArrowUnique {T} {y} f = T .snd y .snd f
+
+  terminalEndoIsId : (T : Terminal) (f : C [ terminalOb T , terminalOb T ])
+                    f  id
+  terminalEndoIsId T f = isContr→isProp (T .snd (terminalOb T)) f id
+
+  hasTerminal : Type (ℓ-max  ℓ')
+  hasTerminal =  Terminal ∥₁
+
+  -- Terminality of an object is a proposition.
+  isPropIsTerminal : (x : ob)  isProp (isTerminal x)
+  isPropIsTerminal _ = isPropΠ λ _  isPropIsContr
+
+  open isIso
+
+  -- Objects that are terminal are isomorphic.
+  terminalToIso : (x y : Terminal)  CatIso C (terminalOb x) (terminalOb y)
+  terminalToIso x y .fst = terminalArrow y (terminalOb x)
+  terminalToIso x y .snd .inv = terminalArrow x (terminalOb y)
+  terminalToIso x y .snd .sec = terminalEndoIsId y _
+  terminalToIso x y .snd .ret = terminalEndoIsId x _
+
+  isoToTerminal :  (x : Terminal) y  CatIso C (terminalOb x) y  isTerminal y
+  isoToTerminal x y x≅y y' .fst = x≅y .fst ∘⟨ C  terminalArrow x y'
+  isoToTerminal x y x≅y y' .snd f =
+    sym (⋆InvRMove
+          (invIso x≅y)
+          (sym (terminalArrowUnique {T = x} (invIso x≅y .fst ∘⟨ C  f))))
+
+  open isUnivalent
+
+  -- The type of terminal objects of a univalent category is a proposition,
+  -- i.e. all terminal objects are equal.
+  isPropTerminal : (hC : isUnivalent C)  isProp Terminal
+  isPropTerminal hC x y =
+    Σ≡Prop isPropIsTerminal (CatIsoToPath hC (terminalToIso x y))
+
+preservesTerminals :  (C : Category ℓc ℓc')(D : Category ℓd ℓd')
+                    Functor C D
+                    Type (ℓ-max (ℓ-max (ℓ-max ℓc ℓc') ℓd) ℓd')
+preservesTerminals C D F =  (term : Terminal C)  isTerminal D (F  term .fst )
+
+preserveAnyTerminal→PreservesTerminals :  (C : Category ℓc ℓc')(D : Category ℓd ℓd')
+   (F : Functor C D)
+   (term : Terminal C)  isTerminal D (F  term .fst )
+   preservesTerminals C D F
+preserveAnyTerminal→PreservesTerminals C D F term D-preserves-term term' =
+  isoToTerminal D
+                ((F  term .fst ) , D-preserves-term) (F  term' .fst )
+                (F-Iso {F = F} (terminalToIso C term term'))
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.Limits.html b/docs/Cubical.Categories.Limits.html new file mode 100644 index 0000000..04e0ccf --- /dev/null +++ b/docs/Cubical.Categories.Limits.html @@ -0,0 +1,11 @@ + +Cubical.Categories.Limits
{-# OPTIONS --safe #-}
+module Cubical.Categories.Limits where
+
+open import Cubical.Categories.Limits.Limits public
+open import Cubical.Categories.Limits.BinProduct public
+open import Cubical.Categories.Limits.BinCoproduct public
+open import Cubical.Categories.Limits.Initial public
+open import Cubical.Categories.Limits.Terminal public
+open import Cubical.Categories.Limits.Pullback public
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Morphism.html b/docs/Cubical.Categories.Morphism.html index a2b32e9..84f91c2 100644 --- a/docs/Cubical.Categories.Morphism.html +++ b/docs/Cubical.Categories.Morphism.html @@ -24,9 +24,9 @@ isMonic {x} {y} f = {z} {a a' : Hom[ z , x ]} f a f a' a a' - isPropIsMonic : (f : Hom[ x , y ]) isProp (isMonic f) + isPropIsMonic : (f : Hom[ x , y ]) isProp (isMonic f) isPropIsMonic _ = isPropImplicitΠ _ (isPropImplicitΠ2 - a a' (isProp→ (isOfHLevelPath' 1 isSetHom a a'))))) + a a' (isProp→ (isOfHLevelPath' 1 isSetHom a a'))))) postcompCreatesMonic : (f : Hom[ x , y ]) (g : Hom[ y , z ]) isMonic (f g) isMonic f @@ -41,135 +41,129 @@ monicId : {x : ob} isMonic (id {x}) monicId {a = a} {a' = a'} eq = sym (⋆IdR a) eq ⋆IdR a' - isJointMono : {ℓ''} {a : ob} (I : Type ℓ'') (b : I ob) ((i : I) Hom[ a , b i ]) Type _ - isJointMono {a = a} I b f = {c} (g₁ g₂ : Hom[ c , a ]) ((i : I) (f i) g₁ (f i) g₂) g₁ g₂ - - retraction⇒monic : (f : Hom[ x , y ]) (lInv : Hom[ y , x ]) - (lInv f id) isMonic f - retraction⇒monic f lInv eq = - postcompCreatesMonic f lInv (subst isMonic (sym eq) monicId) - - isEpic : (Hom[ x , y ]) Type (ℓ-max ℓ') - isEpic {x} {y} g = - {z} {b b' : Hom[ y , z ]} b g b' g b b' - - isPropIsEpic : (f : Hom[ x , y ]) isProp (isEpic f) - isPropIsEpic _ = isPropImplicitΠ _ (isPropImplicitΠ2 - a a' (isProp→ (isOfHLevelPath' 1 isSetHom a a'))))) - - precompCreatesEpic : (f : Hom[ x , y ]) (g : Hom[ z , x ]) - isEpic (g f) isEpic f - precompCreatesEpic f g epic {b = b} {b' = b'} bf≡b'f = - epic (⋆Assoc g f b cong (g ⋆_) bf≡b'f sym (⋆Assoc g f b')) - - isJointEpic : {ℓ''} {b : ob} (I : Type ℓ'') (a : I ob) ((i : I) Hom[ a i , b ]) Type _ - isJointEpic {b = b} I a f = {c : ob} (g₁ g₂ : Hom[ b , c ]) ((i : I) g₁ (f i) g₂ (f i)) g₁ g₂ - - -- A morphism is split monic if it has a *retraction* - isSplitMon : (Hom[ x , y ]) Type ℓ' - isSplitMon {x} {y} f = ∃[ r Hom[ y , x ] ] r f id - - -- A morphism is split epic if it has a *section* - isSplitEpi : (Hom[ x , y ]) Type ℓ' - isSplitEpi {x} {y} g = ∃[ s Hom[ y , x ] ] g s id - - record areInv (f : Hom[ x , y ]) (g : Hom[ y , x ]) : Type ℓ' where - field - sec : g f id - ret : f g id - - --- C can be implicit here -module _ {C : Category ℓ'} where - open Category C - - private - variable - x y z v w : ob - - open areInv - - symAreInv : {f : Hom[ x , y ]} {g : Hom[ y , x ]} areInv C f g areInv C g f - sec (symAreInv x) = ret x - ret (symAreInv x) = sec x - - -- equational reasoning with inverses - invMoveR : {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ z , x ]} {k : Hom[ z , y ]} - areInv C f g - h f k - h k g - invMoveR {f = f} {g} {h} {k} ai p - = h - ≡⟨ sym (⋆IdR _) - (h id) - ≡⟨ cong (h ⋆_) (sym (ai .ret)) - (h (f g)) - ≡⟨ sym (⋆Assoc _ _ _) - ((h f) g) - ≡⟨ cong (_⋆ g) p - k g - - - invMoveL : {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ y , z ]} {k : Hom[ x , z ]} - areInv C f g - f h k - h g k - invMoveL {f = f} {g} {h} {k} ai p - = h - ≡⟨ sym (⋆IdL _) - id h - ≡⟨ cong (_⋆ h) (sym (ai .sec)) - (g f) h - ≡⟨ ⋆Assoc _ _ _ - g (f h) - ≡⟨ cong (g ⋆_) p - (g k) - - - invFlipSq : {f₁⁻¹ : Hom[ x , y ]} {f₁ : Hom[ y , x ]} - {f₂⁻¹ : Hom[ v , w ]} {f₂ : Hom[ w , v ]} - {g : Hom[ x , v ]} {h : Hom[ y , w ]} - areInv C f₁ f₁⁻¹ areInv C f₂ f₂⁻¹ - h f₂ f₁ g - g f₂⁻¹ f₁⁻¹ h - invFlipSq {f₁⁻¹ = f₁⁻¹} {f₁} {f₂⁻¹} {f₂} {g} {h} inv₁ inv₂ sq = - g f₂⁻¹ ≡⟨ cong (_⋆ f₂⁻¹) (sym (⋆IdL _)) - (id g) f₂⁻¹ ≡⟨ cong m (m g) f₂⁻¹) (sym (inv₁ .sec)) - ((f₁⁻¹ f₁) g) f₂⁻¹ ≡⟨ cong (_⋆ f₂⁻¹) (⋆Assoc _ _ _) - (f₁⁻¹ (f₁ g)) f₂⁻¹ ≡⟨ ⋆Assoc _ _ _ - f₁⁻¹ ((f₁ g) f₂⁻¹) ≡⟨ cong m f₁⁻¹ (m f₂⁻¹)) (sym sq) - f₁⁻¹ ((h f₂) f₂⁻¹) ≡⟨ cong (f₁⁻¹ ⋆_) (⋆Assoc _ _ _) - f₁⁻¹ (h (f₂ f₂⁻¹)) ≡⟨ cong m f₁⁻¹ (h m)) (inv₂ .ret) - f₁⁻¹ (h id) ≡⟨ cong (f₁⁻¹ ⋆_) (⋆IdR _) - f₁⁻¹ h - - open isIso - - isIso→areInv : {f : Hom[ x , y ]} - (isI : isIso C f) - areInv C f (isI .inv) - sec (isIso→areInv isI) = sec isI - ret (isIso→areInv isI) = ret isI - - - -- Back and forth between isIso and CatIso - - isIso→CatIso : {f : C [ x , y ]} - isIso C f - CatIso C x y - isIso→CatIso x = _ , x - - CatIso→isIso : (cIso : CatIso C x y) - isIso C (cIso .fst) - CatIso→isIso = snd - - CatIso→areInv : (cIso : CatIso C x y) - areInv C (cIso .fst) (cIso .snd .inv) - CatIso→areInv cIso = isIso→areInv (CatIso→isIso cIso) - - -- reverse of an iso is also an iso - symCatIso : {x y} - CatIso C x y - CatIso C y x - symCatIso (mor , isiso inv sec ret) = inv , isiso mor ret sec + retraction⇒monic : (f : Hom[ x , y ]) (lInv : Hom[ y , x ]) + (lInv f id) isMonic f + retraction⇒monic f lInv eq = + postcompCreatesMonic f lInv (subst isMonic (sym eq) monicId) + + isEpic : (Hom[ x , y ]) Type (ℓ-max ℓ') + isEpic {x} {y} g = + {z} {b b' : Hom[ y , z ]} b g b' g b b' + + isPropIsEpic : (f : Hom[ x , y ]) isProp (isEpic f) + isPropIsEpic _ = isPropImplicitΠ _ (isPropImplicitΠ2 + a a' (isProp→ (isOfHLevelPath' 1 isSetHom a a'))))) + + precompCreatesEpic : (f : Hom[ x , y ]) (g : Hom[ z , x ]) + isEpic (g f) isEpic f + precompCreatesEpic f g epic {b = b} {b' = b'} bf≡b'f = + epic (⋆Assoc g f b cong (g ⋆_) bf≡b'f sym (⋆Assoc g f b')) + + -- A morphism is split monic if it has a *retraction* + isSplitMon : (Hom[ x , y ]) Type ℓ' + isSplitMon {x} {y} f = ∃[ r Hom[ y , x ] ] r f id + + -- A morphism is split epic if it has a *section* + isSplitEpi : (Hom[ x , y ]) Type ℓ' + isSplitEpi {x} {y} g = ∃[ s Hom[ y , x ] ] g s id + + record areInv (f : Hom[ x , y ]) (g : Hom[ y , x ]) : Type ℓ' where + field + sec : g f id + ret : f g id + + +-- C can be implicit here +module _ {C : Category ℓ'} where + open Category C + + private + variable + x y z v w : ob + + open areInv + + symAreInv : {f : Hom[ x , y ]} {g : Hom[ y , x ]} areInv C f g areInv C g f + sec (symAreInv x) = ret x + ret (symAreInv x) = sec x + + -- equational reasoning with inverses + invMoveR : {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ z , x ]} {k : Hom[ z , y ]} + areInv C f g + h f k + h k g + invMoveR {f = f} {g} {h} {k} ai p + = h + ≡⟨ sym (⋆IdR _) + (h id) + ≡⟨ cong (h ⋆_) (sym (ai .ret)) + (h (f g)) + ≡⟨ sym (⋆Assoc _ _ _) + ((h f) g) + ≡⟨ cong (_⋆ g) p + k g + + + invMoveL : {f : Hom[ x , y ]} {g : Hom[ y , x ]} {h : Hom[ y , z ]} {k : Hom[ x , z ]} + areInv C f g + f h k + h g k + invMoveL {f = f} {g} {h} {k} ai p + = h + ≡⟨ sym (⋆IdL _) + id h + ≡⟨ cong (_⋆ h) (sym (ai .sec)) + (g f) h + ≡⟨ ⋆Assoc _ _ _ + g (f h) + ≡⟨ cong (g ⋆_) p + (g k) + + + invFlipSq : {f₁⁻¹ : Hom[ x , y ]} {f₁ : Hom[ y , x ]} + {f₂⁻¹ : Hom[ v , w ]} {f₂ : Hom[ w , v ]} + {g : Hom[ x , v ]} {h : Hom[ y , w ]} + areInv C f₁ f₁⁻¹ areInv C f₂ f₂⁻¹ + h f₂ f₁ g + g f₂⁻¹ f₁⁻¹ h + invFlipSq {f₁⁻¹ = f₁⁻¹} {f₁} {f₂⁻¹} {f₂} {g} {h} inv₁ inv₂ sq = + g f₂⁻¹ ≡⟨ cong (_⋆ f₂⁻¹) (sym (⋆IdL _)) + (id g) f₂⁻¹ ≡⟨ cong m (m g) f₂⁻¹) (sym (inv₁ .sec)) + ((f₁⁻¹ f₁) g) f₂⁻¹ ≡⟨ cong (_⋆ f₂⁻¹) (⋆Assoc _ _ _) + (f₁⁻¹ (f₁ g)) f₂⁻¹ ≡⟨ ⋆Assoc _ _ _ + f₁⁻¹ ((f₁ g) f₂⁻¹) ≡⟨ cong m f₁⁻¹ (m f₂⁻¹)) (sym sq) + f₁⁻¹ ((h f₂) f₂⁻¹) ≡⟨ cong (f₁⁻¹ ⋆_) (⋆Assoc _ _ _) + f₁⁻¹ (h (f₂ f₂⁻¹)) ≡⟨ cong m f₁⁻¹ (h m)) (inv₂ .ret) + f₁⁻¹ (h id) ≡⟨ cong (f₁⁻¹ ⋆_) (⋆IdR _) + f₁⁻¹ h + + open isIso + + isIso→areInv : {f : Hom[ x , y ]} + (isI : isIso C f) + areInv C f (isI .inv) + sec (isIso→areInv isI) = sec isI + ret (isIso→areInv isI) = ret isI + + + -- Back and forth between isIso and CatIso + + isIso→CatIso : {f : C [ x , y ]} + isIso C f + CatIso C x y + isIso→CatIso x = _ , x + + CatIso→isIso : (cIso : CatIso C x y) + isIso C (cIso .fst) + CatIso→isIso = snd + + CatIso→areInv : (cIso : CatIso C x y) + areInv C (cIso .fst) (cIso .snd .inv) + CatIso→areInv cIso = isIso→areInv (CatIso→isIso cIso) + + -- reverse of an iso is also an iso + symCatIso : {x y} + CatIso C x y + CatIso C y x + symCatIso (mor , isiso inv sec ret) = inv , isiso mor ret sec
\ No newline at end of file diff --git a/docs/Cubical.Categories.NaturalTransformation.Base.html b/docs/Cubical.Categories.NaturalTransformation.Base.html index 84a9837..086c4d9 100644 --- a/docs/Cubical.Categories.NaturalTransformation.Base.html +++ b/docs/Cubical.Categories.NaturalTransformation.Base.html @@ -7,7 +7,7 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism renaming (iso to iIso) open import Cubical.Data.Sigma -open import Cubical.Categories.Category renaming (isIso to isIsoC) +open import Cubical.Categories.Category renaming (isIso to isIsoC) open import Cubical.Categories.Functor.Base open import Cubical.Categories.Functor.Properties open import Cubical.Categories.Commutativity @@ -23,7 +23,7 @@ infixl 15 _⋆ᴰ_ private _⋆ᴰ_ : {x y z} (f : D [ x , y ]) (g : D [ y , z ]) D [ x , z ] - f ⋆ᴰ g = f ⋆⟨ D g + f ⋆ᴰ g = f ⋆⟨ D g open Category open Functor @@ -55,19 +55,19 @@ -- the three other commuting squares sqRL : {x y : C .ob} {f : C [ x , y ]} - F f (N-ob x) ⋆ᴰ G f ⋆ᴰ (nIso y) .inv - sqRL {x} {y} {f} = invMoveR (isIso→areInv (nIso y)) (N-hom f) + F f (N-ob x) ⋆ᴰ G f ⋆ᴰ (nIso y) .inv + sqRL {x} {y} {f} = invMoveR (isIso→areInv (nIso y)) (N-hom f) sqLL : {x y : C .ob} {f : C [ x , y ]} - G f ⋆ᴰ (nIso y) .inv (nIso x) .inv ⋆ᴰ F f - sqLL {x} {y} {f} = invMoveL (isIso→areInv (nIso x)) (sym sqRL') + G f ⋆ᴰ (nIso y) .inv (nIso x) .inv ⋆ᴰ F f + sqLL {x} {y} {f} = invMoveL (isIso→areInv (nIso x)) (sym sqRL') where - sqRL' : F f (N-ob x) ⋆ᴰ ( G f ⋆ᴰ (nIso y) .inv ) + sqRL' : F f (N-ob x) ⋆ᴰ ( G f ⋆ᴰ (nIso y) .inv ) sqRL' = sqRL (D .⋆Assoc _ _ _) sqLR : {x y : C .ob} {f : C [ x , y ]} - G f (nIso x) .inv ⋆ᴰ F f ⋆ᴰ (N-ob y) - sqLR {x} {y} {f} = invMoveR (symAreInv (isIso→areInv (nIso y))) sqLL + G f (nIso x) .inv ⋆ᴰ F f ⋆ᴰ (N-ob y) + sqLR {x} {y} {f} = invMoveR (symAreInv (isIso→areInv (nIso y))) sqLL open NatTrans open NatIso @@ -101,19 +101,19 @@ idNatIso : (F : Functor C D) NatIso F F idNatIso F .trans = idTrans F - idNatIso F .nIso _ = idCatIso .snd + idNatIso F .nIso _ = idCatIso .snd -- Natural isomorphism induced by path of functors pathToNatTrans : {F G : Functor C D} F G NatTrans F G - pathToNatTrans p .N-ob x = pathToIso {C = D} i p i .F-ob x) .fst + pathToNatTrans p .N-ob x = pathToIso {C = D} i p i .F-ob x) .fst pathToNatTrans {F = F} {G = G} p .N-hom {x = x} {y = y} f = pathToIso-Comm {C = D} _ _ _ _ i p i .F-hom f) pathToNatIso : {F G : Functor C D} F G NatIso F G pathToNatIso p .trans = pathToNatTrans p - pathToNatIso p .nIso x = pathToIso {C = D} _ .snd + pathToNatIso p .nIso x = pathToIso {C = D} _ .snd -- vertical sequencing @@ -174,21 +174,21 @@ -- extensions are equal to originals βy'≡βy : PathP i D [ Gy≡G'y i , H y ]) βy' (β y ) - βy'≡βy = symP (toPathP {A = λ i D [ Gy≡G'y (~ i) , H y ]} refl) + βy'≡βy = symP (toPathP {A = λ i D [ Gy≡G'y (~ i) , H y ]} refl) βx≡βx' : PathP i D [ Gx≡G'x (~ i) , H x ]) (β x ) βx' - βx≡βx' = toPathP refl + βx≡βx' = toPathP refl -- left wall of square - left : PathP i D [ Gx≡G'x i , H y ]) (G f ⋆⟨ D βy') (G' f ⋆⟨ D β y ) - left i = Gf≡G'f i ⋆⟨ D βy'≡βy i + left : PathP i D [ Gx≡G'x i , H y ]) (G f ⋆⟨ D βy') (G' f ⋆⟨ D β y ) + left i = Gf≡G'f i ⋆⟨ D βy'≡βy i -- right wall of square - right : PathP i D [ Gx≡G'x (~ i) , H y ]) (β x ⋆⟨ D H f ) (βx' ⋆⟨ D H f ) - right i = βx≡βx' i ⋆⟨ D refl {x = H f } i + right : PathP i D [ Gx≡G'x (~ i) , H y ]) (β x ⋆⟨ D H f ) (βx' ⋆⟨ D H f ) + right i = βx≡βx' i ⋆⟨ D refl {x = H f } i -- putting it all together - βSq : G f ⋆⟨ D βy' βx' ⋆⟨ D H f + βSq : G f ⋆⟨ D βy' βx' ⋆⟨ D H f βSq i = comp k D [ Gx≡G'x (~ k) , H y ]) j λ { (i = i0) left (~ j) ; (i = i1) right j }) @@ -205,14 +205,14 @@ where rem : PathP i (F .F-hom f) ⋆ᴰ (p i _) (p i _) ⋆ᴰ (G .F-hom f)) (α .N-hom f) (β .N-hom f) - rem = toPathP (D .isSetHom _ _ _ _) + rem = toPathP (D .isSetHom _ _ _ _) -- `constructor` for path of natural isomorphisms NatIso≡ : {F G : Functor C D}{f g : NatIso F G} f .trans .N-ob g .trans .N-ob f g NatIso≡ {f = f} {g} p i .trans = makeNatTransPath {α = f .trans} {β = g .trans} p i NatIso≡ {f = f} {g} p i .nIso x = - isProp→PathP i isPropIsIso (NatIso≡ {f = f} {g} p i .trans .N-ob x)) (f .nIso _) (g .nIso _) i + isProp→PathP i isPropIsIso (NatIso≡ {f = f} {g} p i .trans .N-ob x)) (f .nIso _) (g .nIso _) i module _ {F F' G G' : Functor C D} {α : NatTrans F G} {β : NatTrans F' G'} where @@ -228,7 +228,7 @@ where rem : PathP i ((p i) .F-hom f) ⋆ᴰ (P i _) (P i _) ⋆ᴰ ((q i) .F-hom f)) (α .N-hom f) (β .N-hom f) - rem = toPathP (D .isSetHom _ _ _ _) + rem = toPathP (D .isSetHom _ _ _ _) module _ {B : Category ℓB ℓB'} {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where open NatTrans @@ -243,7 +243,7 @@ _∘ʳ_ : (K : Functor C D) {G H : Functor B C} (β : NatTrans G H) NatTrans (K ∘F G) (K ∘F H) (_∘ʳ_ K {G} {H} β) .N-ob x = K β x - (_∘ʳ_ K {G} {H} β) .N-hom f = preserveCommF {C = C} {D = D} {K} (β .N-hom f) + (_∘ʳ_ K {G} {H} β) .N-hom f = preserveCommF {C = C} {D = D} {K} (β .N-hom f) whiskerTrans : {F F' : Functor B C} {G G' : Functor C D} (β : NatTrans G G') (α : NatTrans F F') NatTrans (G ∘F F) (G' ∘F F') diff --git a/docs/Cubical.Categories.NaturalTransformation.Properties.html b/docs/Cubical.Categories.NaturalTransformation.Properties.html index 4d4f5b0..63e6cad 100644 --- a/docs/Cubical.Categories.NaturalTransformation.Properties.html +++ b/docs/Cubical.Categories.NaturalTransformation.Properties.html @@ -1,197 +1,197 @@ Cubical.Categories.NaturalTransformation.Properties
-{-# OPTIONS --allow-unsolved-metas #-}
-
-module Cubical.Categories.NaturalTransformation.Properties where
-
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Univalence
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Isomorphism renaming (iso to iIso)
-open import Cubical.Data.Sigma
-open import Cubical.Categories.Category renaming (isIso to isIsoC)
-open import Cubical.Categories.Functor.Base
-open import Cubical.Categories.Functor.Properties
-open import Cubical.Categories.Morphism
-open import Cubical.Categories.Isomorphism
-open import Cubical.Categories.NaturalTransformation.Base
-
-private
-  variable
-    ℓB ℓB' ℓC ℓC' ℓD ℓD' ℓE ℓE' : Level
-    C : Category ℓC ℓC'
-    D : Category ℓD ℓD'
-    F F' : Functor C D
-
-open isIsoC
-open NatIso
-open NatTrans
-open Category
-open Functor
-open Iso
-
-module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where
-  private
-    _⋆ᴰ_ :  {x y z} (f : D [ x , y ]) (g : D [ y , z ])  D [ x , z ]
-    f ⋆ᴰ g = f ⋆⟨ D  g
-
-  -- natural isomorphism is symmetric
-  symNatIso :  {F G : Functor C D}
-             F ≅ᶜ G
-             G ≅ᶜ F
-  symNatIso η .trans .N-ob x = η .nIso x .inv
-  symNatIso η .trans .N-hom _ = sqLL η
-  symNatIso η .nIso x .inv = η .trans .N-ob x
-  symNatIso η .nIso x .sec = η .nIso x .ret
-  symNatIso η .nIso x .ret = η .nIso x .sec
-
-  -- Properties
-
-  -- path helpers
-  module NatTransP where
-
-    module _ {F G : Functor C D} where
-
-      -- same as Sigma version
-      NatTransΣ : Type (ℓ-max (ℓ-max ℓC ℓC') ℓD')
-      NatTransΣ = Σ[ ob  ((x : C .ob)  D [(F .F-ob x) , (G .F-ob x)]) ]
-                     ({x y : _ } (f : C [ x , y ])  (F .F-hom f) ⋆ᴰ (ob y)  (ob x) ⋆ᴰ (G .F-hom f))
-
-      NatTransIsoΣ : Iso (NatTrans F G) NatTransΣ
-      NatTransIsoΣ .fun (natTrans N-ob N-hom) = N-ob , N-hom
-      NatTransIsoΣ .inv (N-ob , N-hom) = (natTrans N-ob N-hom)
-      NatTransIsoΣ .rightInv _ = refl
-      NatTransIsoΣ .leftInv _ = refl
-
-      NatTrans≡Σ : NatTrans F G  NatTransΣ
-      NatTrans≡Σ = isoToPath NatTransIsoΣ
-
-      -- introducing paths
-      NatTrans-≡-intro :  {αo βo : N-ob-Type F G}
-                           {αh : N-hom-Type F G αo}
-                           {βh : N-hom-Type F G βo}
-                        (p : αo  βo)
-                        PathP  i  ({x y : C .ob} (f : C [ x , y ])  (F .F-hom f) ⋆ᴰ (p i y)  (p i x) ⋆ᴰ (G .F-hom f))) αh βh
-                        natTrans {F = F} {G} αo αh  natTrans βo βh
-      NatTrans-≡-intro p q i = natTrans (p i) (q i)
-
-  module _ {F G : Functor C D} {α β : NatTrans F G} where
-      open Iso
-      private
-        αOb = α .N-ob
-        βOb = β .N-ob
-        αHom = α .N-hom
-        βHom = β .N-hom
-      -- path between natural transformations is the same as a pair of paths (between ob and hom)
-      NTPathIsoPathΣ : Iso (α  β)
-                         (Σ[ p  (αOb  βOb) ]
-                              (PathP  i  ({x y : _} (f : _)  F  f  ⋆ᴰ (p i y)  (p i x) ⋆ᴰ G  f ))
-                                  αHom
-                                  βHom))
-      NTPathIsoPathΣ .fun p =  i  p i .N-ob) ,  i  p i .N-hom)
-      NTPathIsoPathΣ .inv (po , ph) i = record { N-ob = po i ; N-hom = ph i }
-      NTPathIsoPathΣ .rightInv  = refl
-      NTPathIsoPathΣ .leftInv p = refl
-
-      NTPath≃PathΣ = isoToEquiv NTPathIsoPathΣ
-
-      NTPath≡PathΣ = ua NTPath≃PathΣ
-
-  module _ where
-    open NatTransP
-
-    isSetNatTrans : {F G : Functor C D}  isSet (NatTrans F G)
-    isSetNatTrans =
-      isSetRetract (fun NatTransIsoΣ) (inv NatTransIsoΣ) (leftInv NatTransIsoΣ)
-                   (isSetΣSndProp (isSetΠ  _  isSetHom D))
-                                   _  isPropImplicitΠ2  _ _  isPropΠ  _  isSetHom D _ _))))
-
-
--- Natural isomorphism is path when the target category is univalent.
-
-module _
-  (isUnivD : isUnivalent D)
-  {F G : Functor C D} where
-
-  open isUnivalent isUnivD
-
-  NatIsoToPath : NatIso F G  F  G
-  NatIsoToPath niso =
-    Functor≡  x  CatIsoToPath (_ , niso .nIso x))
-       f  isoToPath-Square isUnivD _ _ _ _ (niso .trans .N-hom f))
-
-  NatIso→Path→NatIso : (niso : NatIso F G)  pathToNatIso (NatIsoToPath niso)  niso
-  NatIso→Path→NatIso niso = NatIso≡  i x  secEq (univEquiv _ _) (_ , niso .nIso x) i .fst)
-
-  Path→NatIso→Path : (p : F  G)  NatIsoToPath (pathToNatIso p)  p
-  Path→NatIso→Path p = FunctorPath≡  i j x  retEq (univEquiv _ _)  i  p i .F-ob x) i j)
-
-  Iso-Path-NatIso : Iso (F  G) (NatIso F G)
-  Iso-Path-NatIso = iso pathToNatIso NatIsoToPath NatIso→Path→NatIso Path→NatIso→Path
-
-  Path≃NatIso : (F  G)  NatIso F G
-  Path≃NatIso = isoToEquiv Iso-Path-NatIso
-module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where
-  seqNatIso : {F G H : Functor C D}  NatIso F G  NatIso G H  NatIso F H
-  seqNatIso ı ı' .trans = seqTrans (ı .trans) (ı' .trans)
-  seqNatIso ı ı' .nIso x .inv = ı' .nIso x .inv ⋆⟨ D  ı .nIso x .inv
-  seqNatIso ı ı' .nIso x .sec =
-    D .⋆Assoc _ _ _
-     cong (_⋆_ D (ı' .nIso x .inv))
-      (sym (D .⋆Assoc _ _ _)
-       cong (D  ı' .trans .N-ob x) (ı .nIso x .sec)
-       D .⋆IdL (ı' .trans .N-ob x))
-     ı' .nIso x .sec
-  seqNatIso ı ı' .nIso x .ret =
-    (sym (D .⋆Assoc _ _ _))
-     cong (_∘_ D (ı .nIso x .inv))
-      (D .⋆Assoc _ _ _
-       cong (D  ı .trans .N-ob x) (ı' .nIso x .ret)
-       D .⋆IdR (ı .trans .N-ob x))
-     ı .nIso x .ret
-
-  CAT⋆IdR : {F : Functor C D}  NatIso (Id ∘F F) F
-  CAT⋆IdR {F} .trans .N-ob = idTrans F .N-ob
-  CAT⋆IdR {F} .trans .N-hom = idTrans F .N-hom
-  CAT⋆IdR {F} .nIso = idNatIso F .nIso
-
-module _ {B : Category ℓB ℓB'}{C : Category ℓC ℓC'}{D : Category ℓD ℓD'} where
-  _∘ʳi_ :  (K : Functor C D)  {G H : Functor B C} (β : NatIso G H)
-        NatIso (K ∘F G) (K ∘F H)
-  _∘ʳi_ K β .trans = K ∘ʳ β .trans
-  _∘ʳi_ K β .nIso x = preserveIsosF {F = K} (β .trans .N-ob _ , β .nIso x) .snd
-
-  open Functor
-  _∘ˡi_ :  (K : Functor B C)  {G H : Functor C D} (β : NatIso G H)
-        NatIso (G ∘F K) (H ∘F K)
-  _∘ˡi_ K β .trans = β .trans ∘ˡ K
-  _∘ˡi_ K β .nIso b  = β .nIso (K  b )
-
-  CAT⋆Assoc : {E : Category ℓE ℓE'}
-            (F : Functor B C)(G : Functor C D)(H : Functor D E)
-             NatIso (H ∘F (G ∘F F)) ((H ∘F G) ∘F F)
-  CAT⋆Assoc F G H .trans .N-ob = idTrans ((H ∘F G) ∘F F) .N-ob
-  CAT⋆Assoc F G H .trans .N-hom = idTrans ((H ∘F G) ∘F F) .N-hom
-  CAT⋆Assoc F G H .nIso = idNatIso ((H ∘F G) ∘F F) .nIso
-
-
-
-⇒^opFiso : Iso (F  F') (_^opF {C = C} {D = D} F'  F ^opF )
-N-ob (fun ⇒^opFiso x) = N-ob x
-N-hom (fun ⇒^opFiso x) f = sym (N-hom x f)
-inv ⇒^opFiso = _
-rightInv ⇒^opFiso _ = refl
-leftInv ⇒^opFiso _ = refl
-
-congNatIso^opFiso : Iso (F ≅ᶜ F') (_^opF  {C = C} {D = D} F'  ≅ᶜ F ^opF )
-trans (fun congNatIso^opFiso x) = Iso.fun ⇒^opFiso (trans x)
-inv (nIso (fun congNatIso^opFiso x) x₁) = _
-sec (nIso (fun congNatIso^opFiso x) x₁) = ret (nIso x x₁)
-ret (nIso (fun congNatIso^opFiso x) x₁) = sec (nIso x x₁)
-inv congNatIso^opFiso = _
-rightInv congNatIso^opFiso _ = refl
-leftInv congNatIso^opFiso _ = refl
+{-# OPTIONS --safe #-}
+
+module Cubical.Categories.NaturalTransformation.Properties where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism renaming (iso to iIso)
+open import Cubical.Data.Sigma
+open import Cubical.Categories.Category renaming (isIso to isIsoC)
+open import Cubical.Categories.Functor.Base
+open import Cubical.Categories.Functor.Properties
+open import Cubical.Categories.Morphism
+open import Cubical.Categories.Isomorphism
+open import Cubical.Categories.NaturalTransformation.Base
+
+private
+  variable
+    ℓB ℓB' ℓC ℓC' ℓD ℓD' ℓE ℓE' : Level
+    C : Category ℓC ℓC'
+    D : Category ℓD ℓD'
+    F F' : Functor C D
+
+open isIsoC
+open NatIso
+open NatTrans
+open Category
+open Functor
+open Iso
+
+module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where
+  private
+    _⋆ᴰ_ :  {x y z} (f : D [ x , y ]) (g : D [ y , z ])  D [ x , z ]
+    f ⋆ᴰ g = f ⋆⟨ D  g
+
+  -- natural isomorphism is symmetric
+  symNatIso :  {F G : Functor C D}
+             F ≅ᶜ G
+             G ≅ᶜ F
+  symNatIso η .trans .N-ob x = η .nIso x .inv
+  symNatIso η .trans .N-hom _ = sqLL η
+  symNatIso η .nIso x .inv = η .trans .N-ob x
+  symNatIso η .nIso x .sec = η .nIso x .ret
+  symNatIso η .nIso x .ret = η .nIso x .sec
+
+  -- Properties
+
+  -- path helpers
+  module NatTransP where
+
+    module _ {F G : Functor C D} where
+
+      -- same as Sigma version
+      NatTransΣ : Type (ℓ-max (ℓ-max ℓC ℓC') ℓD')
+      NatTransΣ = Σ[ ob  ((x : C .ob)  D [(F .F-ob x) , (G .F-ob x)]) ]
+                     ({x y : _ } (f : C [ x , y ])  (F .F-hom f) ⋆ᴰ (ob y)  (ob x) ⋆ᴰ (G .F-hom f))
+
+      NatTransIsoΣ : Iso (NatTrans F G) NatTransΣ
+      NatTransIsoΣ .fun (natTrans N-ob N-hom) = N-ob , N-hom
+      NatTransIsoΣ .inv (N-ob , N-hom) = (natTrans N-ob N-hom)
+      NatTransIsoΣ .rightInv _ = refl
+      NatTransIsoΣ .leftInv _ = refl
+
+      NatTrans≡Σ : NatTrans F G  NatTransΣ
+      NatTrans≡Σ = isoToPath NatTransIsoΣ
+
+      -- introducing paths
+      NatTrans-≡-intro :  {αo βo : N-ob-Type F G}
+                           {αh : N-hom-Type F G αo}
+                           {βh : N-hom-Type F G βo}
+                        (p : αo  βo)
+                        PathP  i  ({x y : C .ob} (f : C [ x , y ])  (F .F-hom f) ⋆ᴰ (p i y)  (p i x) ⋆ᴰ (G .F-hom f))) αh βh
+                        natTrans {F = F} {G} αo αh  natTrans βo βh
+      NatTrans-≡-intro p q i = natTrans (p i) (q i)
+
+  module _ {F G : Functor C D} {α β : NatTrans F G} where
+      open Iso
+      private
+        αOb = α .N-ob
+        βOb = β .N-ob
+        αHom = α .N-hom
+        βHom = β .N-hom
+      -- path between natural transformations is the same as a pair of paths (between ob and hom)
+      NTPathIsoPathΣ : Iso (α  β)
+                         (Σ[ p  (αOb  βOb) ]
+                              (PathP  i  ({x y : _} (f : _)  F  f  ⋆ᴰ (p i y)  (p i x) ⋆ᴰ G  f ))
+                                  αHom
+                                  βHom))
+      NTPathIsoPathΣ .fun p =  i  p i .N-ob) ,  i  p i .N-hom)
+      NTPathIsoPathΣ .inv (po , ph) i = record { N-ob = po i ; N-hom = ph i }
+      NTPathIsoPathΣ .rightInv  = refl
+      NTPathIsoPathΣ .leftInv p = refl
+
+      NTPath≃PathΣ = isoToEquiv NTPathIsoPathΣ
+
+      NTPath≡PathΣ = ua NTPath≃PathΣ
+
+  module _ where
+    open NatTransP
+
+    isSetNatTrans : {F G : Functor C D}  isSet (NatTrans F G)
+    isSetNatTrans =
+      isSetRetract (fun NatTransIsoΣ) (inv NatTransIsoΣ) (leftInv NatTransIsoΣ)
+                   (isSetΣSndProp (isSetΠ  _  isSetHom D))
+                                   _  isPropImplicitΠ2  _ _  isPropΠ  _  isSetHom D _ _))))
+
+
+-- Natural isomorphism is path when the target category is univalent.
+
+module _
+  (isUnivD : isUnivalent D)
+  {F G : Functor C D} where
+
+  open isUnivalent isUnivD
+
+  NatIsoToPath : NatIso F G  F  G
+  NatIsoToPath niso =
+    Functor≡  x  CatIsoToPath (_ , niso .nIso x))
+       f  isoToPath-Square isUnivD _ _ _ _ (niso .trans .N-hom f))
+
+  NatIso→Path→NatIso : (niso : NatIso F G)  pathToNatIso (NatIsoToPath niso)  niso
+  NatIso→Path→NatIso niso = NatIso≡  i x  secEq (univEquiv _ _) (_ , niso .nIso x) i .fst)
+
+  Path→NatIso→Path : (p : F  G)  NatIsoToPath (pathToNatIso p)  p
+  Path→NatIso→Path p = FunctorPath≡  i j x  retEq (univEquiv _ _)  i  p i .F-ob x) i j)
+
+  Iso-Path-NatIso : Iso (F  G) (NatIso F G)
+  Iso-Path-NatIso = iso pathToNatIso NatIsoToPath NatIso→Path→NatIso Path→NatIso→Path
+
+  Path≃NatIso : (F  G)  NatIso F G
+  Path≃NatIso = isoToEquiv Iso-Path-NatIso
+module _ {C : Category ℓC ℓC'} {D : Category ℓD ℓD'} where
+  seqNatIso : {F G H : Functor C D}  NatIso F G  NatIso G H  NatIso F H
+  seqNatIso ı ı' .trans = seqTrans (ı .trans) (ı' .trans)
+  seqNatIso ı ı' .nIso x .inv = ı' .nIso x .inv ⋆⟨ D  ı .nIso x .inv
+  seqNatIso ı ı' .nIso x .sec =
+    D .⋆Assoc _ _ _
+     cong (_⋆_ D (ı' .nIso x .inv))
+      (sym (D .⋆Assoc _ _ _)
+       cong (D  ı' .trans .N-ob x) (ı .nIso x .sec)
+       D .⋆IdL (ı' .trans .N-ob x))
+     ı' .nIso x .sec
+  seqNatIso ı ı' .nIso x .ret =
+    (sym (D .⋆Assoc _ _ _))
+     cong (_∘_ D (ı .nIso x .inv))
+      (D .⋆Assoc _ _ _
+       cong (D  ı .trans .N-ob x) (ı' .nIso x .ret)
+       D .⋆IdR (ı .trans .N-ob x))
+     ı .nIso x .ret
+
+  CAT⋆IdR : {F : Functor C D}  NatIso (Id ∘F F) F
+  CAT⋆IdR {F} .trans .N-ob = idTrans F .N-ob
+  CAT⋆IdR {F} .trans .N-hom = idTrans F .N-hom
+  CAT⋆IdR {F} .nIso = idNatIso F .nIso
+
+module _ {B : Category ℓB ℓB'}{C : Category ℓC ℓC'}{D : Category ℓD ℓD'} where
+  _∘ʳi_ :  (K : Functor C D)  {G H : Functor B C} (β : NatIso G H)
+        NatIso (K ∘F G) (K ∘F H)
+  _∘ʳi_ K β .trans = K ∘ʳ β .trans
+  _∘ʳi_ K β .nIso x = preserveIsosF {F = K} (β .trans .N-ob _ , β .nIso x) .snd
+
+  open Functor
+  _∘ˡi_ :  (K : Functor B C)  {G H : Functor C D} (β : NatIso G H)
+        NatIso (G ∘F K) (H ∘F K)
+  _∘ˡi_ K β .trans = β .trans ∘ˡ K
+  _∘ˡi_ K β .nIso b  = β .nIso (K  b )
+
+  CAT⋆Assoc : {E : Category ℓE ℓE'}
+            (F : Functor B C)(G : Functor C D)(H : Functor D E)
+             NatIso (H ∘F (G ∘F F)) ((H ∘F G) ∘F F)
+  CAT⋆Assoc F G H .trans .N-ob = idTrans ((H ∘F G) ∘F F) .N-ob
+  CAT⋆Assoc F G H .trans .N-hom = idTrans ((H ∘F G) ∘F F) .N-hom
+  CAT⋆Assoc F G H .nIso = idNatIso ((H ∘F G) ∘F F) .nIso
+
+
+
+⇒^opFiso : Iso (F  F') (_^opF {C = C} {D = D} F'  F ^opF )
+N-ob (fun ⇒^opFiso x) = N-ob x
+N-hom (fun ⇒^opFiso x) f = sym (N-hom x f)
+inv ⇒^opFiso = _
+rightInv ⇒^opFiso _ = refl
+leftInv ⇒^opFiso _ = refl
+
+congNatIso^opFiso : Iso (F ≅ᶜ F') (_^opF  {C = C} {D = D} F'  ≅ᶜ F ^opF )
+trans (fun congNatIso^opFiso x) = Iso.fun ⇒^opFiso (trans x)
+inv (nIso (fun congNatIso^opFiso x) x₁) = _
+sec (nIso (fun congNatIso^opFiso x) x₁) = ret (nIso x x₁)
+ret (nIso (fun congNatIso^opFiso x) x₁) = sec (nIso x x₁)
+inv congNatIso^opFiso = _
+rightInv congNatIso^opFiso _ = refl
+leftInv congNatIso^opFiso _ = refl
 
 
\ No newline at end of file diff --git a/docs/Cubical.Categories.NaturalTransformation.html b/docs/Cubical.Categories.NaturalTransformation.html index e6c830b..dae3d9d 100644 --- a/docs/Cubical.Categories.NaturalTransformation.html +++ b/docs/Cubical.Categories.NaturalTransformation.html @@ -1,8 +1,8 @@ -Cubical.Categories.NaturalTransformation
{-# OPTIONS --allow-unsolved-metas #-}
+Cubical.Categories.NaturalTransformation
{-# OPTIONS --safe #-}
 
-module Cubical.Categories.NaturalTransformation where
+module Cubical.Categories.NaturalTransformation where
 
-open import Cubical.Categories.NaturalTransformation.Base public
-open import Cubical.Categories.NaturalTransformation.Properties public
+open import Cubical.Categories.NaturalTransformation.Base public
+open import Cubical.Categories.NaturalTransformation.Properties public
 
\ No newline at end of file diff --git a/docs/Cubical.Data.Bool.Base.html b/docs/Cubical.Data.Bool.Base.html index 2094232..fe4a9cb 100644 --- a/docs/Cubical.Data.Bool.Base.html +++ b/docs/Cubical.Data.Bool.Base.html @@ -88,13 +88,13 @@ -- Universe lifted booleans Bool* : {} Type -Bool* = Lift Bool +Bool* = Lift Bool true* : {} Bool* {} -true* = lift true +true* = lift true false* : {} Bool* {} -false* = lift false +false* = lift false -- Pointed version Bool*∙ : {} Σ[ X Type ] X diff --git a/docs/Cubical.Data.Bool.Properties.html b/docs/Cubical.Data.Bool.Properties.html index 85e36f9..b7ee834 100644 --- a/docs/Cubical.Data.Bool.Properties.html +++ b/docs/Cubical.Data.Bool.Properties.html @@ -70,11 +70,11 @@ : (P : {b : Bool} b b Type ) (∀{b} P {b} refl) ∀{b} (q : b b) P q -K-Bool P Pr {false} = J (λ{ false q P q ; true _ Lift }) Pr -K-Bool P Pr {true} = J (λ{ true q P q ; false _ Lift }) Pr +K-Bool P Pr {false} = J (λ{ false q P q ; true _ Lift }) Pr +K-Bool P Pr {true} = J (λ{ true q P q ; false _ Lift }) Pr -isSetBool : isSet Bool -isSetBool a b = J _ p q p q) (K-Bool (refl ≡_) refl) +isSetBool : isSet Bool +isSetBool a b = J _ p q p q) (K-Bool (refl ≡_) refl) true≢false : ¬ true false true≢false p = subst b if b then Bool else ) p true @@ -200,7 +200,7 @@ Table A P table : P Table Bool P - table = J Table (inspect false true refl refl) + table = J Table (inspect false true refl refl) reflLemma : (P : Bool Bool) transport P false false @@ -252,24 +252,24 @@ false true = _ _ = Unit -isProp≤ : b c isProp (b c) +isProp≤ : b c isProp (b c) isProp≤ true false = isProp⊥ isProp≤ true true = isPropUnit isProp≤ false false = isPropUnit isProp≤ false true = isPropUnit -isProp≥ : b c isProp (b c) +isProp≥ : b c isProp (b c) isProp≥ false true = isProp⊥ isProp≥ true true = isPropUnit isProp≥ false false = isPropUnit isProp≥ true false = isPropUnit -isProp-Bool→Type : b isProp (Bool→Type b) +isProp-Bool→Type : b isProp (Bool→Type b) isProp-Bool→Type false = isProp⊥ isProp-Bool→Type true = isPropUnit -isPropDep-Bool→Type : isPropDep Bool→Type -isPropDep-Bool→Type = isOfHLevel→isOfHLevelDep 1 isProp-Bool→Type +isPropDep-Bool→Type : isPropDep Bool→Type +isPropDep-Bool→Type = isOfHLevel→isOfHLevelDep 1 isProp-Bool→Type IsoBool→∙ : {} {A : Pointed } Iso ((Bool , true) →∙ A) (typ A) Iso.fun IsoBool→∙ f = fst f false @@ -293,20 +293,20 @@ Bool→TypeInj : (a b : Bool) Bool→Type a Bool→Type b a b Bool→TypeInj true true _ = refl Bool→TypeInj true false p = Empty.rec (p .fst tt) -Bool→TypeInj false true p = Empty.rec (invEq p tt) +Bool→TypeInj false true p = Empty.rec (invEq p tt) Bool→TypeInj false false _ = refl Bool→TypeInj* : (a b : Bool) Bool→Type* {} a Bool→Type* {} b a b Bool→TypeInj* true true _ = refl Bool→TypeInj* true false p = Empty.rec* (p .fst tt*) -Bool→TypeInj* false true p = Empty.rec* (invEq p tt*) +Bool→TypeInj* false true p = Empty.rec* (invEq p tt*) Bool→TypeInj* false false _ = refl -isPropBool→Type : {a : Bool} isProp (Bool→Type a) +isPropBool→Type : {a : Bool} isProp (Bool→Type a) isPropBool→Type {a = true} = isPropUnit isPropBool→Type {a = false} = isProp⊥ -isPropBool→Type* : {a : Bool} isProp (Bool→Type* {} a) +isPropBool→Type* : {a : Bool} isProp (Bool→Type* {} a) isPropBool→Type* {a = true} = isPropUnit* isPropBool→Type* {a = false} = isProp⊥* @@ -316,7 +316,7 @@ DecBool→Type* : {a : Bool} Dec (Bool→Type* {} a) DecBool→Type* {a = true} = yes tt* -DecBool→Type* {a = false} = no (lift x) x) +DecBool→Type* {a = false} = no (lift x) x) Dec→DecBool : {P : Type } (dec : Dec P) P Bool→Type (Dec→Bool dec) Dec→DecBool (yes p) _ = tt @@ -325,8 +325,8 @@ DecBool→Dec : {P : Type } (dec : Dec P) Bool→Type (Dec→Bool dec) P DecBool→Dec (yes p) _ = p -Dec≃DecBool : {P : Type } (h : isProp P)(dec : Dec P) P Bool→Type (Dec→Bool dec) -Dec≃DecBool h dec = propBiimpl→Equiv h isPropBool→Type (Dec→DecBool dec) (DecBool→Dec dec) +Dec≃DecBool : {P : Type } (h : isProp P)(dec : Dec P) P Bool→Type (Dec→Bool dec) +Dec≃DecBool h dec = propBiimpl→Equiv h isPropBool→Type (Dec→DecBool dec) (DecBool→Dec dec) Bool≡BoolDec : {a : Bool} a Dec→Bool (DecBool→Type {a = a}) Bool≡BoolDec {a = true} = refl @@ -339,8 +339,8 @@ DecBool→Dec* : {P : Type } (dec : Dec P) Bool→Type* {} (Dec→Bool dec) P DecBool→Dec* (yes p) _ = p -Dec≃DecBool* : {P : Type } (h : isProp P)(dec : Dec P) P Bool→Type* {} (Dec→Bool dec) -Dec≃DecBool* h dec = propBiimpl→Equiv h isPropBool→Type* (Dec→DecBool* dec) (DecBool→Dec* dec) +Dec≃DecBool* : {P : Type } (h : isProp P)(dec : Dec P) P Bool→Type* {} (Dec→Bool dec) +Dec≃DecBool* h dec = propBiimpl→Equiv h isPropBool→Type* (Dec→DecBool* dec) (DecBool→Dec* dec) Bool≡BoolDec* : {a : Bool} a Dec→Bool (DecBool→Type* {} {a = a}) Bool≡BoolDec* {a = true} = refl @@ -360,7 +360,7 @@ Bool→Type×≃ : (a b : Bool) Bool→Type a × Bool→Type b Bool→Type (a and b) Bool→Type×≃ a b = - propBiimpl→Equiv (isProp× isPropBool→Type isPropBool→Type) isPropBool→Type + propBiimpl→Equiv (isProp× isPropBool→Type isPropBool→Type) isPropBool→Type (Bool→Type×' a b) (Bool→Type× a b) Bool→Type⊎ : (a b : Bool) Bool→Type (a or b) Bool→Type a Bool→Type b @@ -390,10 +390,10 @@ Bool≡ false false = true Bool≡≃ : (a b : Bool) (a b) Bool→Type (Bool≡ a b) -Bool≡≃ true true = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) +Bool≡≃ true true = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) Bool≡≃ true false = uninhabEquiv true≢false Empty.rec Bool≡≃ false true = uninhabEquiv false≢true Empty.rec -Bool≡≃ false false = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) +Bool≡≃ false false = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) open Iso -- Bool is equivalent to bi-point type diff --git a/docs/Cubical.Data.Bool.SwitchStatement.html b/docs/Cubical.Data.Bool.SwitchStatement.html new file mode 100644 index 0000000..64fc02b --- /dev/null +++ b/docs/Cubical.Data.Bool.SwitchStatement.html @@ -0,0 +1,44 @@ + +Cubical.Data.Bool.SwitchStatement
{-# OPTIONS --safe #-}
+module Cubical.Data.Bool.SwitchStatement where
+
+open import Cubical.Core.Everything
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Data.Bool.Base
+open import Cubical.Data.Nat
+
+{-
+  Switch-case:
+
+    _==_ : A → A → Bool
+
+    _ : B
+    _ = switch (λ x → x == fixedValue) cases
+           case value1 ⇒ result1 break
+           case value2 ⇒ result2 break
+           ...
+           case valueN ⇒ resultN break
+           default⇒ defaultResult
+-}
+
+
+private
+  variable
+     ℓ' : Level
+
+
+infixr 6 default⇒_
+infixr 5 case_⇒_break_
+infixr 4 switch_cases_
+
+switch_cases_ : {A : Type } {B : Type ℓ'}  (A  Bool)  ((A  Bool)  B)  B
+switch caseIndicator cases caseData = caseData caseIndicator
+
+case_⇒_break_ : {A : Type } {B : Type ℓ'}  A  B  (otherCases : (A  Bool)  B)  (A  Bool)  B
+case forValue  result break otherCases = λ caseIndicator  if (caseIndicator forValue) then result else (otherCases caseIndicator)
+
+default⇒_ : {A : Type } {B : Type ℓ'}  B  (A  Bool)  B
+default⇒_ value caseIndicator = value
+
\ No newline at end of file diff --git a/docs/Cubical.Data.Empty.Base.html b/docs/Cubical.Data.Empty.Base.html index ac0fa90..488b385 100644 --- a/docs/Cubical.Data.Empty.Base.html +++ b/docs/Cubical.Data.Empty.Base.html @@ -11,7 +11,7 @@ data : Type₀ where ⊥* : Type -⊥* = Lift +⊥* = Lift rec : {A : Type } A rec () @@ -21,4 +21,7 @@ elim : {A : Type } (x : ) A x elim () + +elim* : {A : ⊥* {ℓ'} Type } (x : ⊥* {ℓ'}) A x +elim* ()
\ No newline at end of file diff --git a/docs/Cubical.Data.Empty.Properties.html b/docs/Cubical.Data.Empty.Properties.html index b9e7ccf..51261d3 100644 --- a/docs/Cubical.Data.Empty.Properties.html +++ b/docs/Cubical.Data.Empty.Properties.html @@ -8,21 +8,21 @@ open import Cubical.Data.Empty.Base -isProp⊥ : isProp +isProp⊥ : isProp isProp⊥ () -isProp⊥* : {} isProp {} ⊥* +isProp⊥* : {} isProp {} ⊥* isProp⊥* _ () -isContr⊥→A : {} {A : Type } isContr ( A) +isContr⊥→A : {} {A : Type } isContr ( A) fst isContr⊥→A () snd isContr⊥→A f i () -isContrΠ⊥ : {} {A : Type } isContr ((x : ) A x) +isContrΠ⊥ : {} {A : Type } isContr ((x : ) A x) fst isContrΠ⊥ () snd isContrΠ⊥ f i () -isContrΠ⊥* : { ℓ'} {A : ⊥* {} Type ℓ'} isContr ((x : ⊥*) A x) +isContrΠ⊥* : { ℓ'} {A : ⊥* {} Type ℓ'} isContr ((x : ⊥*) A x) fst isContrΠ⊥* () snd isContrΠ⊥* f i () diff --git a/docs/Cubical.Data.Fin.Base.html b/docs/Cubical.Data.Fin.Base.html index 5bc2121..b0bca88 100644 --- a/docs/Cubical.Data.Fin.Base.html +++ b/docs/Cubical.Data.Fin.Base.html @@ -6,108 +6,138 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Function open import Cubical.Foundations.HLevels - -import Cubical.Data.Empty as -open import Cubical.Data.Nat using ( ; zero ; suc ; _+_ ; znots) -open import Cubical.Data.Nat.Order -open import Cubical.Data.Nat.Order.Recursive using () renaming (_≤_ to _≤′_) -open import Cubical.Data.Sigma -open import Cubical.Data.Sum using (_⊎_; _⊎?_; inl; inr) - -open import Cubical.Relation.Nullary - --- Finite types. --- --- Currently it is most convenient to define these as a subtype of the --- natural numbers, because indexed inductive definitions don't behave --- well with cubical Agda. This definition also has some more general --- attractive properties, of course, such as easy conversion back to --- ℕ. -Fin : Type₀ -Fin n = Σ[ k ] k < n - -private - variable - : Level - k : - -fzero : Fin (suc k) -fzero = (0 , suc-≤-suc zero-≤) - -fone : Fin (suc (suc k)) -fone = (1 , suc-≤-suc (suc-≤-suc zero-≤)) - -fzero≠fone : ¬ fzero {k = suc k} fone -fzero≠fone p = znots (cong fst p) - --- It is easy, using this representation, to take the successor of a --- number as a number in the next largest finite type. -fsuc : Fin k Fin (suc k) -fsuc (k , l) = (suc k , suc-≤-suc l) - --- Conversion back to ℕ is trivial... -toℕ : Fin k -toℕ = fst - --- ... and injective. -toℕ-injective : ∀{fj fk : Fin k} toℕ fj toℕ fk fj fk -toℕ-injective {fj = fj} {fk} = Σ≡Prop _ isProp≤) - --- Conversion from ℕ with a recursive definition of ≤ - -fromℕ≤ : (m n : ) m ≤′ n Fin (suc n) -fromℕ≤ zero _ _ = fzero -fromℕ≤ (suc m) (suc n) m≤n = fsuc (fromℕ≤ m n m≤n) - --- A case analysis helper for induction. -fsplit - : ∀(fj : Fin (suc k)) - (fzero fj) (Σ[ fk Fin k ] fsuc fk fj) -fsplit (0 , k<sn) = inl (toℕ-injective refl) -fsplit (suc k , k<sn) = inr ((k , pred-≤-pred k<sn) , toℕ-injective refl) - -inject< : {m n} (m<n : m < n) Fin m Fin n -inject< m<n (k , k<m) = k , <-trans k<m m<n - -flast : Fin (suc k) -flast {k = k} = k , suc-≤-suc ≤-refl - --- Fin 0 is empty -¬Fin0 : ¬ Fin 0 -¬Fin0 (k , k<0) = ¬-<-zero k<0 - --- The full inductive family eliminator for finite types. -elim - : ∀(P : ∀{k} Fin k Type ) - (∀{k} P {suc k} fzero) - (∀{k} {fn : Fin k} P fn P (fsuc fn)) - {k : } (fn : Fin k) P fn -elim P fz fs {zero} = ⊥.rec ¬Fin0 -elim P fz fs {suc k} fj - = case fsplit fj return _ P fj) of λ - { (inl p) subst P p fz - ; (inr (fk , p)) subst P p (fs (elim P fz fs fk)) - } - -any? : {n} {P : Fin n Type } (∀ i Dec (P i)) Dec (Σ (Fin n) P) -any? {n = zero} {P = _} P? = no (x , _) ¬Fin0 x) -any? {n = suc n} {P = P} P? = - mapDec - - { (inl P0) fzero , P0 - ; (inr (x , Px)) fsuc x , Px - } - ) - n h n (helper h)) - (P? fzero ⊎? any? (P? fsuc)) - where - helper : Σ (Fin (suc n)) P P fzero Σ (Fin n) λ z P (fsuc z) - helper (x , Px) with fsplit x - ... | inl x≡0 = inl (subst P (sym x≡0) Px) - ... | inr (k , x≡sk) = inr (k , subst P (sym x≡sk) Px) - -FinPathℕ : {n : } (x : Fin n) (y : ) fst x y Σ[ p _ ] (x (y , p)) -FinPathℕ {n = n} x y p = - ((fst (snd x)) , (cong y fst (snd x) + y) (cong suc (sym p)) snd (snd x))) - , (Σ≡Prop _ isProp≤) p) +open import Cubical.Foundations.Pointed + +import Cubical.Data.Empty as +open import Cubical.Data.Nat using ( ; zero ; suc ; _+_ ; znots) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Nat.Order.Recursive using () renaming (_≤_ to _≤′_) +open import Cubical.Data.Sigma +open import Cubical.Data.Sum using (_⊎_; _⊎?_; inl; inr) + +open import Cubical.Relation.Nullary + +-- Finite types. +-- +-- Currently it is most convenient to define these as a subtype of the +-- natural numbers, because indexed inductive definitions don't behave +-- well with cubical Agda. This definition also has some more general +-- attractive properties, of course, such as easy conversion back to +-- ℕ. +Fin : Type₀ +Fin n = Σ[ k ] k < n + +private + variable + : Level + k : + +fzero : Fin (suc k) +fzero = (0 , suc-≤-suc zero-≤) + +fone : Fin (suc (suc k)) +fone = (1 , suc-≤-suc (suc-≤-suc zero-≤)) + +fzero≠fone : ¬ fzero {k = suc k} fone +fzero≠fone p = znots (cong fst p) + +-- It is easy, using this representation, to take the successor of a +-- number as a number in the next largest finite type. +fsuc : Fin k Fin (suc k) +fsuc (k , l) = (suc k , suc-≤-suc l) + +finj : Fin k Fin (suc k) +finj (k , l) = k , ≤-trans l (1 , refl) + +-- Conversion back to ℕ is trivial... +toℕ : Fin k +toℕ = fst + +-- ... and injective. +toℕ-injective : ∀{fj fk : Fin k} toℕ fj toℕ fk fj fk +toℕ-injective {fj = fj} {fk} = Σ≡Prop _ isProp≤) + +-- Conversion from ℕ with a recursive definition of ≤ + +fromℕ≤ : (m n : ) m ≤′ n Fin (suc n) +fromℕ≤ zero _ _ = fzero +fromℕ≤ (suc m) (suc n) m≤n = fsuc (fromℕ≤ m n m≤n) + +-- A case analysis helper for induction. +fsplit + : ∀(fj : Fin (suc k)) + (fzero fj) (Σ[ fk Fin k ] fsuc fk fj) +fsplit (0 , k<sn) = inl (toℕ-injective refl) +fsplit (suc k , k<sn) = inr ((k , pred-≤-pred k<sn) , toℕ-injective refl) + +inject< : {m n} (m<n : m < n) Fin m Fin n +inject< m<n (k , k<m) = k , <-trans k<m m<n + +flast : Fin (suc k) +flast {k = k} = k , suc-≤-suc ≤-refl + +-- Fin 0 is empty +¬Fin0 : ¬ Fin 0 +¬Fin0 (k , k<0) = ¬-<-zero k<0 + +-- The full inductive family eliminator for finite types. +elim + : ∀(P : ∀{k} Fin k Type ) + (∀{k} P {suc k} fzero) + (∀{k} {fn : Fin k} P fn P (fsuc fn)) + {k : } (fn : Fin k) P fn +elim P fz fs {zero} = ⊥.rec ¬Fin0 +elim P fz fs {suc k} fj + = case fsplit fj return _ P fj) of λ + { (inl p) subst P p fz + ; (inr (fk , p)) subst P p (fs (elim P fz fs fk)) + } + +any? : {n} {P : Fin n Type } (∀ i Dec (P i)) Dec (Σ (Fin n) P) +any? {n = zero} {P = _} P? = no (x , _) ¬Fin0 x) +any? {n = suc n} {P = P} P? = + mapDec + + { (inl P0) fzero , P0 + ; (inr (x , Px)) fsuc x , Px + } + ) + n h n (helper h)) + (P? fzero ⊎? any? (P? fsuc)) + where + helper : Σ (Fin (suc n)) P P fzero Σ (Fin n) λ z P (fsuc z) + helper (x , Px) with fsplit x + ... | inl x≡0 = inl (subst P (sym x≡0) Px) + ... | inr (k , x≡sk) = inr (k , subst P (sym x≡sk) Px) + +FinPathℕ : {n : } (x : Fin n) (y : ) fst x y Σ[ p _ ] (x (y , p)) +FinPathℕ {n = n} x y p = + ((fst (snd x)) , (cong y fst (snd x) + y) (cong suc (sym p)) snd (snd x))) + , (Σ≡Prop _ isProp≤) p) + +FinVec : (A : Type ) (n : ) Type +FinVec A n = Fin n A + +FinFamily : (n : ) ( : Level) Type (ℓ-suc ) +FinFamily n = FinVec (Type ) n + +FinFamily∙ : (n : ) ( : Level) Type (ℓ-suc ) +FinFamily∙ n = FinVec (Pointed ) n + +predFinFamily : {n : } FinFamily (suc n) FinFamily n +predFinFamily A n = A (finj n) + +predFinFamily∙ : {n : } FinFamily∙ (suc n) FinFamily∙ n +fst (predFinFamily∙ A x) = predFinFamily (fst A) x +snd (predFinFamily∙ A x) = snd (A _) + +prodFinFamily : (n : ) FinFamily (suc n) Type +prodFinFamily zero A = A fzero +prodFinFamily (suc n) A = prodFinFamily n (predFinFamily A) × A flast + +prodFinFamily∙ : (n : ) FinFamily∙ (suc n) Pointed +fst (prodFinFamily∙ n A) = prodFinFamily n (fst A) +snd (prodFinFamily∙ zero A) = snd (A fzero) +snd (prodFinFamily∙ (suc n) A) = + snd (prodFinFamily∙ n (predFinFamily∙ A)) , snd (A flast)
\ No newline at end of file diff --git a/docs/Cubical.Data.Fin.Literals.html b/docs/Cubical.Data.Fin.Literals.html index 724f3f2..ae4578a 100644 --- a/docs/Cubical.Data.Fin.Literals.html +++ b/docs/Cubical.Data.Fin.Literals.html @@ -7,14 +7,14 @@ open import Agda.Builtin.FromNat renaming (Number to HasFromNat) open import Cubical.Data.Fin.Base - using (Fin; fromℕ≤) + using (Fin; fromℕ≤) open import Cubical.Data.Nat.Order.Recursive using (_≤_) instance - fromNatFin : {n : _} HasFromNat (Fin (suc n)) + fromNatFin : {n : _} HasFromNat (Fin (suc n)) fromNatFin {n} = record { Constraint = λ m m n - ; fromNat = λ m m≤n fromℕ≤ m n m≤n + ; fromNat = λ m m≤n fromℕ≤ m n m≤n }
\ No newline at end of file diff --git a/docs/Cubical.Data.Fin.Properties.html b/docs/Cubical.Data.Fin.Properties.html index def4b12..e1292fe 100644 --- a/docs/Cubical.Data.Fin.Properties.html +++ b/docs/Cubical.Data.Fin.Properties.html @@ -39,71 +39,71 @@ A : Type a -- Fin 0 is empty, and thus a proposition. -isPropFin0 : isProp (Fin 0) -isPropFin0 = Empty.rec ¬Fin0 +isPropFin0 : isProp (Fin 0) +isPropFin0 = Empty.rec ¬Fin0 -- Fin 1 has only one value. -isContrFin1 : isContr (Fin 1) +isContrFin1 : isContr (Fin 1) isContrFin1 - = fzero , λ - { (zero , _) toℕ-injective refl + = fzero , λ + { (zero , _) toℕ-injective refl ; (suc k , sk<1) Empty.rec (¬-<-zero (pred-≤-pred sk<1)) } -Unit≃Fin1 : Unit Fin 1 +Unit≃Fin1 : Unit Fin 1 Unit≃Fin1 = isoToEquiv (iso - (const fzero) + (const fzero) (const tt) (isContrFin1 .snd) (isContrUnit .snd) ) -- Regardless of k, Fin k is a set. -isSetFin : ∀{k} isSet (Fin k) -isSetFin {k} = isSetΣ isSetℕ _ isProp→isSet isProp≤) +isSetFin : ∀{k} isSet (Fin k) +isSetFin {k} = isSetΣ isSetℕ _ isProp→isSet isProp≤) -discreteFin : {n} Discrete (Fin n) +discreteFin : {n} Discrete (Fin n) discreteFin {n} (x , hx) (y , hy) with discreteℕ x y ... | yes prf = yes (Σ≡Prop _ isProp≤) prf) ... | no prf = no λ h prf (cong fst h) -inject<-ne : {n} (i : Fin n) ¬ inject< ≤-refl i (n , ≤-refl) +inject<-ne : {n} (i : Fin n) ¬ inject< ≤-refl i (n , ≤-refl) inject<-ne {n} (k , k<n) p = <→≢ k<n (cong fst p) -Fin-fst-≡ : {n} {i j : Fin n} fst i fst j i j +Fin-fst-≡ : {n} {i j : Fin n} fst i fst j i j Fin-fst-≡ = Σ≡Prop _ isProp≤) private subst-app : (B : A Type b) (f : (x : A) B x) {x y : A} (x≡y : x y) subst B x≡y (f x) f y subst-app B f {x = x} = - J y e subst B e (f x) f y) (substRefl {B = B} (f x)) + J y e subst B e (f x) f y) (substRefl {B = B} (f x)) -- Computation rules for the eliminator. -module _ (P : {k} Fin k Type ) - (fz : {k} P {suc k} fzero) - (fs : {k} {fn : Fin k} P fn P (fsuc fn)) +module _ (P : {k} Fin k Type ) + (fz : {k} P {suc k} fzero) + (fs : {k} {fn : Fin k} P fn P (fsuc fn)) {k : } where - elim-fzero : Fin.elim P fz fs {k = suc k} fzero fz + elim-fzero : Fin.elim P fz fs {k = suc k} fzero fz elim-fzero = - subst P (toℕ-injective _) fz ≡⟨ cong p subst P p fz) (isSetFin _ _ _ _) + subst P (toℕ-injective _) fz ≡⟨ cong p subst P p fz) (isSetFin _ _ _ _) subst P refl fz ≡⟨ substRefl {B = P} fz fz - elim-fsuc : (fk : Fin k) Fin.elim P fz fs (fsuc fk) fs (Fin.elim P fz fs fk) + elim-fsuc : (fk : Fin k) Fin.elim P fz fs (fsuc fk) fs (Fin.elim P fz fs fk) elim-fsuc fk = - subst P (toℕ-injective _ toℕ (fsuc fk′))) (fs (Fin.elim P fz fs fk′)) - ≡⟨ cong p subst P p (fs (Fin.elim P fz fs fk′)) ) (isSetFin _ _ _ _) - subst P (cong fsuc fk′≡fk) (fs (Fin.elim P fz fs fk′)) - ≡⟨ subst-app _ fj fs (Fin.elim P fz fs fj)) fk′≡fk - fs (Fin.elim P fz fs fk) + subst P (toℕ-injective _ toℕ (fsuc fk′))) (fs (Fin.elim P fz fs fk′)) + ≡⟨ cong p subst P p (fs (Fin.elim P fz fs fk′)) ) (isSetFin _ _ _ _) + subst P (cong fsuc fk′≡fk) (fs (Fin.elim P fz fs fk′)) + ≡⟨ subst-app _ fj fs (Fin.elim P fz fs fj)) fk′≡fk + fs (Fin.elim P fz fs fk) where - fk′ = fst fk , pred-≤-pred (snd (fsuc fk)) + fk′ = fst fk , pred-≤-pred (snd (fsuc fk)) fk′≡fk : fk′ fk - fk′≡fk = toℕ-injective refl + fk′≡fk = toℕ-injective refl -- Helper function for the reduction procedure below. -- @@ -119,16 +119,16 @@ -- Expand a pair. This is useful because the whole function is -- injective. -expand× : ∀{k} (Fin k × ) -expand× {k} (f , o) = expand o k (toℕ f) +expand× : ∀{k} (Fin k × ) +expand× {k} (f , o) = expand o k (toℕ f) private lemma₀ : ∀{k m n r} r n k + m n k r lemma₀ {k = k} {m} p q = m , +-comm m k q sym p - expand×Inj : k {t1 t2 : Fin (suc k) × } expand× t1 expand× t2 t1 t2 + expand×Inj : k {t1 t2 : Fin (suc k) × } expand× t1 expand× t2 t1 t2 expand×Inj k {f1 , zero} {f2 , zero} p i - = toℕ-injective {fj = f1} {f2} p i , zero + = toℕ-injective {fj = f1} {f2} p i , zero expand×Inj k {f1 , suc o1} {(r , r<sk) , zero} p = Empty.rec (<-asym r<sk (lemma₀ refl p)) expand×Inj k {(r , r<sk) , zero} {f2 , suc o2} p @@ -139,22 +139,22 @@ inj-m+ {suc k} expand×Emb : k isEmbedding (expand× {k}) - expand×Emb 0 = Empty.rec ¬Fin0 fst + expand×Emb 0 = Empty.rec ¬Fin0 fst expand×Emb (suc k) = injEmbedding isSetℕ (expand×Inj k) -- A Residue is a family of types representing evidence that a -- natural is congruent to a value of a finite type. Residue : Type₀ -Residue k n = Σ[ tup Fin k × ] expand× tup n +Residue k n = Σ[ tup Fin k × ] expand× tup n -- There is at most one canonical finite value congruent to each -- natural. -isPropResidue : k n isProp (Residue k n) +isPropResidue : k n isProp (Residue k n) isPropResidue k = isEmbedding→hasPropFibers (expand×Emb k) -- A value of a finite type is its own residue. -Fin→Residue : ∀{k} (f : Fin k) Residue k (toℕ f) +Fin→Residue : ∀{k} (f : Fin k) Residue k (toℕ f) Fin→Residue f = (f , 0) , refl -- Fibers of numbers that differ by k are equivalent in a more obvious @@ -212,11 +212,11 @@ reduceP : n PathP i Residue≡ k n i) (reduce n) (reduce (k + n)) - reduceP n = toPathP (reduce≡ n) + reduceP n = toPathP (reduce≡ n) open Reduce using (reduce; reduce≡) public -extract : ∀{k n} Residue k n Fin k +extract : ∀{k n} Residue k n Fin k extract = fst fst private @@ -230,14 +230,14 @@ extract≡ k n = lemma₅ (suc k) n (reduce k n) cong extract (Reduce.reduce≡ k n) -isContrResidue : ∀{k n} isContr (Residue (suc k) n) -isContrResidue {k} {n} = inhProp→isContr (reduce k n) (isPropResidue (suc k) n) +isContrResidue : ∀{k n} isContr (Residue (suc k) n) +isContrResidue {k} {n} = inhProp→isContr (reduce k n) (isPropResidue (suc k) n) -- the modulo operator on ℕ _%_ : n % zero = n -n % (suc k) = toℕ (extract (reduce k n)) +n % (suc k) = toℕ (extract (reduce k n)) _/_ : n / zero = zero @@ -253,20 +253,20 @@ n%sk<sk : (n k : ) (n % suc k) < suc k n%sk<sk n k = extract (reduce k n) .snd -fznotfs : {m : } {k : Fin m} ¬ fzero fsuc k +fznotfs : {m : } {k : Fin m} ¬ fzero fsuc k fznotfs {m} p = subst F p tt where - F : Fin (suc m) Type₀ + F : Fin (suc m) Type₀ F (zero , _) = Unit F (suc _ , _) = -fsuc-inj : {fj fk : Fin n} fsuc fj fsuc fk fj fk -fsuc-inj = toℕ-injective injSuc cong toℕ +fsuc-inj : {fj fk : Fin n} fsuc fj fsuc fk fj fk +fsuc-inj = toℕ-injective injSuc cong toℕ -punchOut : {m} {i j : Fin (suc m)} (¬ i j) Fin m -punchOut {_} {i} {j} p with fsplit i | fsplit j +punchOut : {m} {i j : Fin (suc m)} (¬ i j) Fin m +punchOut {_} {i} {j} p with fsplit i | fsplit j punchOut {_} {i} {j} p | inl prfi | inl prfj = - Empty.elim (p (i ≡⟨ sym prfi fzero ≡⟨ prfj j )) + Empty.elim (p (i ≡⟨ sym prfi fzero ≡⟨ prfj j )) punchOut {_} {i} {j} p | inl prfi | inr (kj , prfj) = kj punchOut {zero} {i} {j} p | inr (ki , prfi) | inl prfj = @@ -277,7 +277,7 @@ )) where c = isContrFin1 .fst punchOut {suc _} {i} {j} p | inr (ki , prfi) | inl prfj = - fzero + fzero punchOut {zero} {i} {j} p | inr (ki , prfi) | inr (kj , prfj) = Empty.elim ((p ( i ≡⟨ sym (isContrFin1 .snd i) @@ -286,38 +286,38 @@ )) where c = isContrFin1 .fst punchOut {suc _} {i} {j} p | inr (ki , prfi) | inr (kj , prfj) = - fsuc (punchOut {i = ki} {j = kj} - q p (i ≡⟨ sym prfi fsuc ki ≡⟨ cong fsuc q fsuc kj ≡⟨ prfj j )) + fsuc (punchOut {i = ki} {j = kj} + q p (i ≡⟨ sym prfi fsuc ki ≡⟨ cong fsuc q fsuc kj ≡⟨ prfj j )) ) punchOut-inj - : {m} {i j k : Fin (suc m)} (i≢j : ¬ i j) (i≢k : ¬ i k) + : {m} {i j k : Fin (suc m)} (i≢j : ¬ i j) (i≢k : ¬ i k) punchOut i≢j punchOut i≢k j k -punchOut-inj {_} {i} {j} {k} i≢j i≢k p with fsplit i | fsplit j | fsplit k +punchOut-inj {_} {i} {j} {k} i≢j i≢k p with fsplit i | fsplit j | fsplit k punchOut-inj {zero} {i} {j} {k} i≢j i≢k p | _ | _ | _ = Empty.elim (i≢j (i ≡⟨ sym (isContrFin1 .snd i) c ≡⟨ isContrFin1 .snd j j )) where c = isContrFin1 .fst punchOut-inj {suc _} {i} {j} {k} i≢j i≢k p | inl prfi | inl prfj | _ = - Empty.elim (i≢j (i ≡⟨ sym prfi fzero ≡⟨ prfj j )) + Empty.elim (i≢j (i ≡⟨ sym prfi fzero ≡⟨ prfj j )) punchOut-inj {suc _} {i} {j} {k} i≢j i≢k p | inl prfi | _ | inl prfk = - Empty.elim (i≢k (i ≡⟨ sym prfi fzero ≡⟨ prfk k )) + Empty.elim (i≢k (i ≡⟨ sym prfi fzero ≡⟨ prfk k )) punchOut-inj {suc _} {i} {j} {k} i≢j i≢k p | inl prfi | inr (kj , prfj) | inr (kk , prfk) = j ≡⟨ sym prfj - fsuc kj ≡⟨ cong fsuc p - fsuc kk ≡⟨ prfk + fsuc kj ≡⟨ cong fsuc p + fsuc kk ≡⟨ prfk k punchOut-inj {suc _} {i} {j} {k} i≢j i≢k p | inr (ki , prfi) | inl prfj | inl prfk = j ≡⟨ sym prfj - fzero ≡⟨ prfk + fzero ≡⟨ prfk k punchOut-inj {suc _} {i} {j} {k} i≢j i≢k p | inr (ki , prfi) | inr (kj , prfj) | inr (kk , prfk) = j ≡⟨ sym prfj - fsuc kj ≡⟨ cong fsuc lemma4 - fsuc kk ≡⟨ prfk + fsuc kj ≡⟨ cong fsuc lemma4 + fsuc kk ≡⟨ prfk k where - lemma1 = λ q i≢j (i ≡⟨ sym prfi fsuc ki ≡⟨ cong fsuc q fsuc kj ≡⟨ prfj j ) - lemma2 = λ q i≢k (i ≡⟨ sym prfi fsuc ki ≡⟨ cong fsuc q fsuc kk ≡⟨ prfk k ) + lemma1 = λ q i≢j (i ≡⟨ sym prfi fsuc ki ≡⟨ cong fsuc q fsuc kj ≡⟨ prfj j ) + lemma2 = λ q i≢k (i ≡⟨ sym prfi fsuc ki ≡⟨ cong fsuc q fsuc kk ≡⟨ prfk k ) lemma3 = fsuc-inj p lemma4 = punchOut-inj lemma1 lemma2 lemma3 punchOut-inj {suc m} {i} {j} {k} i≢j i≢k p | inr (ki , prfi) | inl prfj | inr (kk , prfk) = @@ -327,35 +327,35 @@ pigeonhole-special : {n} - (f : Fin (suc n) Fin n) - Σ[ i Fin (suc n) ] Σ[ j Fin (suc n) ] (¬ i j) × (f i f j) -pigeonhole-special {zero} f = Empty.rec (¬Fin0 (f fzero)) + (f : Fin (suc n) Fin n) + Σ[ i Fin (suc n) ] Σ[ j Fin (suc n) ] (¬ i j) × (f i f j) +pigeonhole-special {zero} f = Empty.rec (¬Fin0 (f fzero)) pigeonhole-special {suc n} f = - proof (any? - (i : Fin (suc n)) - discreteFin (f (inject< ≤-refl i)) (f (suc n , ≤-refl)) + proof (any? + (i : Fin (suc n)) + discreteFin (f (inject< ≤-refl i)) (f (suc n , ≤-refl)) )) where proof - : Dec (Σ (Fin (suc n)) z f (inject< ≤-refl z) f (suc n , ≤-refl))) - Σ[ i Fin (suc (suc n)) ] Σ[ j Fin (suc (suc n)) ] (¬ i j) × (f i f j) - proof (yes (i , prf)) = inject< ≤-refl i , (suc n , ≤-refl) , inject<-ne i , prf + : Dec (Σ (Fin (suc n)) z f (inject< ≤-refl z) f (suc n , ≤-refl))) + Σ[ i Fin (suc (suc n)) ] Σ[ j Fin (suc (suc n)) ] (¬ i j) × (f i f j) + proof (yes (i , prf)) = inject< ≤-refl i , (suc n , ≤-refl) , inject<-ne i , prf proof (no h) = let - g : Fin (suc n) Fin n + g : Fin (suc n) Fin n g k = punchOut {i = f (suc n , ≤-refl)} - {j = f (inject< ≤-refl k)} + {j = f (inject< ≤-refl k)} p h (k , Fin-fst-≡ (sym (cong fst p)))) i , j , i≢j , p = pigeonhole-special g in - inject< ≤-refl i - , inject< ≤-refl j + inject< ≤-refl i + , inject< ≤-refl j , q i≢j (Fin-fst-≡ (cong fst q))) , punchOut-inj {i = f (suc n , ≤-refl)} - {j = f (inject< ≤-refl i)} - {k = f (inject< ≤-refl j)} + {j = f (inject< ≤-refl i)} + {k = f (inject< ≤-refl j)} q h (i , Fin-fst-≡ (sym (cong fst q)))) q h (j , Fin-fst-≡ (sym (cong fst q)))) (Fin-fst-≡ (cong fst p)) @@ -363,26 +363,26 @@ pigeonhole : {m n} m < n - (f : Fin n Fin m) - Σ[ i Fin n ] Σ[ j Fin n ] (¬ i j) × (f i f j) + (f : Fin n Fin m) + Σ[ i Fin n ] Σ[ j Fin n ] (¬ i j) × (f i f j) pigeonhole {m} {n} (zero , sm≡n) f = transport transport-prf (pigeonhole-special f′) where - f′ : Fin (suc m) Fin m - f′ = subst h Fin h Fin m) (sym sm≡n) f + f′ : Fin (suc m) Fin m + f′ = subst h Fin h Fin m) (sym sm≡n) f - f′≡f : PathP i Fin (sm≡n i) Fin m) f′ f - f′≡f i = transport-fillerExt (cong h Fin h Fin m) (sym sm≡n)) (~ i) f + f′≡f : PathP i Fin (sm≡n i) Fin m) f′ f + f′≡f i = transport-fillerExt (cong h Fin h Fin m) (sym sm≡n)) (~ i) f transport-prf - : (Σ[ i Fin (suc m) ] Σ[ j Fin (suc m) ] (¬ i j) × (f′ i f′ j)) - (Σ[ i Fin n ] Σ[ j Fin n ] (¬ i j) × (f i f j)) + : (Σ[ i Fin (suc m) ] Σ[ j Fin (suc m) ] (¬ i j) × (f′ i f′ j)) + (Σ[ i Fin n ] Σ[ j Fin n ] (¬ i j) × (f i f j)) transport-prf φ = - Σ[ i Fin (sm≡n φ) ] Σ[ j Fin (sm≡n φ) ] + Σ[ i Fin (sm≡n φ) ] Σ[ j Fin (sm≡n φ) ] (¬ i j) × (f′≡f φ i f′≡f φ j) pigeonhole {m} {n} (suc k , prf) f = let - g : Fin (suc n′) Fin n′ + g : Fin (suc n′) Fin n′ g k = fst (f′ k) , <-trans (snd (f′ k)) m<n′ i , j , ¬q , r = pigeonhole-special g in transport transport-prf (i , j , ¬q , Σ≡Prop _ isProp≤) (cong fst r)) @@ -399,20 +399,20 @@ m<n′ : m < n′ m<n′ = k , injSuc (suc (k + suc m) ≡⟨ prf n ≡⟨ n≡sn′ suc n′ ) - f′ : Fin (suc n′) Fin m - f′ = subst h Fin h Fin m) n≡sn′ f + f′ : Fin (suc n′) Fin m + f′ = subst h Fin h Fin m) n≡sn′ f - f′≡f : PathP i Fin (n≡sn′ (~ i)) Fin m) f′ f - f′≡f i = transport-fillerExt (cong h Fin h Fin m) n≡sn′) (~ i) f + f′≡f : PathP i Fin (n≡sn′ (~ i)) Fin m) f′ f + f′≡f i = transport-fillerExt (cong h Fin h Fin m) n≡sn′) (~ i) f transport-prf - : (Σ[ i Fin (suc n′) ] Σ[ j Fin (suc n′) ] (¬ i j) × (f′ i f′ j)) - (Σ[ i Fin n ] Σ[ j Fin n ] (¬ i j) × (f i f j)) + : (Σ[ i Fin (suc n′) ] Σ[ j Fin (suc n′) ] (¬ i j) × (f′ i f′ j)) + (Σ[ i Fin n ] Σ[ j Fin n ] (¬ i j) × (f i f j)) transport-prf φ = - Σ[ i Fin (n≡sn′ (~ φ)) ] Σ[ j Fin (n≡sn′ (~ φ)) ] + Σ[ i Fin (n≡sn′ (~ φ)) ] Σ[ j Fin (n≡sn′ (~ φ)) ] (¬ i j) × (f′≡f φ i f′≡f φ j) -Fin-inj′ : {n m : } n < m ¬ Fin m Fin n +Fin-inj′ : {n m : } n < m ¬ Fin m Fin n Fin-inj′ n<m p = let i , j , i≢j , q = pigeonhole n<m (transport p) @@ -433,7 +433,7 @@ transport (sym p) (transport p y) ≡⟨ transport⁻Transport p y y -Fin-inj : (n m : ) Fin n Fin m n m +Fin-inj : (n m : ) Fin n Fin m n m Fin-inj n m p with n m ... | eq prf = prf ... | lt n<m = Empty.rec (Fin-inj′ n<m (sym p)) @@ -452,11 +452,11 @@ d + m · suc k ≡⟨ p n · suc k - residuek·n : k n (r : Residue (suc k) (n · suc k)) ((fzero , n) , expand≡ (suc k) 0 n +-zero _) r - residuek·n _ _ = isContr→isProp isContrResidue _ + residuek·n : k n (r : Residue (suc k) (n · suc k)) ((fzero , n) , expand≡ (suc k) 0 n +-zero _) r + residuek·n _ _ = isContr→isProp isContrResidue _ r≡0 : r 0 - r≡0 = cong (toℕ extract) (sym (residuek·n k n resn·k)) + r≡0 = cong (toℕ extract) (sym (residuek·n k n resn·k)) d≡o·sk : d o · suc k d≡o·sk = sym (moddiv d (suc k)) ∙∙ cong (o · suc k +_) r≡0 ∙∙ +-zero _ goal : (o + m) · suc k n · suc k @@ -472,37 +472,37 @@ ; (inr e) Empty.rec (¬m<m (subst m m · suc k < n · suc k) e p)) } -factorEquiv : {n} {m} Fin n × Fin m Fin (n · m) -factorEquiv {zero} {m} = uninhabEquiv (¬Fin0 fst) ¬Fin0 +factorEquiv : {n} {m} Fin n × Fin m Fin (n · m) +factorEquiv {zero} {m} = uninhabEquiv (¬Fin0 fst) ¬Fin0 factorEquiv {suc n} {m} = intro , isEmbedding×isSurjection→isEquiv (isEmbeddingIntro , isSurjectionIntro) where - intro : Fin (suc n) × Fin m Fin (suc n · m) - intro (nn , mm) = nm , subst nm₁ nm₁ < suc n · m) (sym (expand≡ _ (toℕ nn) (toℕ mm))) nm<n·m where + intro : Fin (suc n) × Fin m Fin (suc n · m) + intro (nn , mm) = nm , subst nm₁ nm₁ < suc n · m) (sym (expand≡ _ (toℕ nn) (toℕ mm))) nm<n·m where nm : - nm = expand× (nn , toℕ mm) - nm<n·m : toℕ mm · suc n + toℕ nn < suc n · m + nm = expand× (nn , toℕ mm) + nm<n·m : toℕ mm · suc n + toℕ nn < suc n · m nm<n·m = - toℕ mm · suc n + toℕ nn <≤⟨ <-k+ (snd nn) - toℕ mm · suc n + suc n ≡≤⟨ +-comm _ (suc n) - suc (toℕ mm) · suc n ≤≡⟨ ≤-·k (snd mm) + toℕ mm · suc n + toℕ nn <≤⟨ <-k+ (snd nn) + toℕ mm · suc n + suc n ≡≤⟨ +-comm _ (suc n) + suc (toℕ mm) · suc n ≤≡⟨ ≤-·k (snd mm) m · suc n ≡⟨ ·-comm _ (suc n) suc n · m where open <-Reasoning intro-injective : {o} {p} intro o intro p o p - intro-injective {o} {p} io≡ip = λ i io′≡ip′ i .fst , toℕ-injective {fj = snd o} {fk = snd p} (cong snd io′≡ip′) i where - io′≡ip′ : (fst o , toℕ (snd o)) (fst p , toℕ (snd p)) + intro-injective {o} {p} io≡ip = λ i io′≡ip′ i .fst , toℕ-injective {fj = snd o} {fk = snd p} (cong snd io′≡ip′) i where + io′≡ip′ : (fst o , toℕ (snd o)) (fst p , toℕ (snd p)) io′≡ip′ = expand×Inj _ (cong fst io≡ip) isEmbeddingIntro : isEmbedding intro isEmbeddingIntro = injEmbedding isSetFin intro-injective elimF : nm fiber intro nm - elimF nm = ((nn , nn<n) , (mm , mm<m)) , toℕ-injective (reduce n (toℕ nm) .snd) where - mm = toℕ nm / suc n - nn = toℕ nm % suc n + elimF nm = ((nn , nn<n) , (mm , mm<m)) , toℕ-injective (reduce n (toℕ nm) .snd) where + mm = toℕ nm / suc n + nn = toℕ nm % suc n - nmmoddiv : mm · suc n + nn toℕ nm + nmmoddiv : mm · suc n + nn toℕ nm nmmoddiv = moddiv _ (suc n) nn<n : nn < suc n - nn<n = n%sk<sk (toℕ nm) _ + nn<n = n%sk<sk (toℕ nm) _ nmsnd : mm · suc n + nn < suc n · m nmsnd = subst l l < suc n · m) (sym nmmoddiv) (snd nm) @@ -562,16 +562,16 @@ suc (m + (k m)) ≡⟨ cong suc (∸-lemma (pred-≤-pred m≤k)) suc k -Fin+≅Fin⊎Fin : (m n : ) Iso (Fin (m + n)) (Fin m Fin n) +Fin+≅Fin⊎Fin : (m n : ) Iso (Fin (m + n)) (Fin m Fin n) Iso.fun (Fin+≅Fin⊎Fin m n) = f where - f : Fin (m + n) Fin m Fin n + f : Fin (m + n) Fin m Fin n f (k , k<m+n) with k ≤? m f (k , k<m+n) | inl k<m = inl (k , k<m) f (k , k<m+n) | inr k≥m = inr (k m , ∸-<-lemma m n k k<m+n k≥m) Iso.inv (Fin+≅Fin⊎Fin m n) = g where - g : Fin m Fin n Fin (m + n) + g : Fin m Fin n Fin (m + n) g (inl (k , k<m)) = k , o<m→o<m+n m n k k<m g (inr (k , k<n)) = m + k , <-k+ k<n Iso.rightInv (Fin+≅Fin⊎Fin m n) = sec-f-g @@ -593,20 +593,20 @@ ret-f-g (k , k<m+n) | inl _ = Σ≡Prop _ isProp≤) refl ret-f-g (k , k<m+n) | inr m≥k = Σ≡Prop _ isProp≤) (∸-lemma m≥k) -Fin+≡Fin⊎Fin : (m n : ) Fin (m + n) Fin m Fin n +Fin+≡Fin⊎Fin : (m n : ) Fin (m + n) Fin m Fin n Fin+≡Fin⊎Fin m n = isoToPath (Fin+≅Fin⊎Fin m n) -- Equivalence between FinData and Fin -sucFin : {N : } Fin N Fin (suc N) +sucFin : {N : } Fin N Fin (suc N) sucFin (k , n , p) = suc k , n , (+-suc _ _ cong suc p) -FinData→Fin : (N : ) FinData N Fin N +FinData→Fin : (N : ) FinData N Fin N FinData→Fin zero () FinData→Fin (suc N) zero = 0 , suc-≤-suc zero-≤ FinData→Fin (suc N) (suc k) = sucFin (FinData→Fin N k) -Fin→FinData : (N : ) Fin N FinData N +Fin→FinData : (N : ) Fin N FinData N Fin→FinData zero (k , n , p) = Empty.rec (snotz (sym (+-suc n k) p)) Fin→FinData (suc N) (0 , n , p) = zero Fin→FinData (suc N) ((suc k) , n , p) = suc (Fin→FinData N (k , n , p')) where @@ -625,60 +625,60 @@ retFin (suc N) zero = refl retFin (suc N) (suc k) = cong FinData.suc (cong (Fin→FinData N) (Fin-fst-≡ refl) retFin N k) -FinDataIsoFin : (N : ) Iso (FinData N) (Fin N) +FinDataIsoFin : (N : ) Iso (FinData N) (Fin N) Iso.fun (FinDataIsoFin N) = FinData→Fin N Iso.inv (FinDataIsoFin N) = Fin→FinData N Iso.rightInv (FinDataIsoFin N) = secFin N Iso.leftInv (FinDataIsoFin N) = retFin N -FinData≃Fin : (N : ) FinData N Fin N +FinData≃Fin : (N : ) FinData N Fin N FinData≃Fin N = isoToEquiv (FinDataIsoFin N) -FinData≡Fin : (N : ) FinData N Fin N +FinData≡Fin : (N : ) FinData N Fin N FinData≡Fin N = ua (FinData≃Fin N) -- decidability of Fin -DecFin : (n : ) Dec (Fin n) -DecFin 0 = no ¬Fin0 -DecFin (suc n) = yes fzero +DecFin : (n : ) Dec (Fin n) +DecFin 0 = no ¬Fin0 +DecFin (suc n) = yes fzero -- propositional truncation of Fin -Dec∥Fin∥ : (n : ) Dec Fin n ∥₁ +Dec∥Fin∥ : (n : ) Dec Fin n ∥₁ Dec∥Fin∥ n = Dec∥∥ (DecFin n) -- some properties about cardinality -Fin>0→isInhab : (n : ) 0 < n Fin n +Fin>0→isInhab : (n : ) 0 < n Fin n Fin>0→isInhab 0 p = Empty.rec (¬-<-zero p) -Fin>0→isInhab (suc n) p = fzero +Fin>0→isInhab (suc n) p = fzero -Fin>1→hasNonEqualTerm : (n : ) 1 < n Σ[ i Fin n ] Σ[ j Fin n ] ¬ i j +Fin>1→hasNonEqualTerm : (n : ) 1 < n Σ[ i Fin n ] Σ[ j Fin n ] ¬ i j Fin>1→hasNonEqualTerm 0 p = Empty.rec (snotz (≤0→≡0 p)) Fin>1→hasNonEqualTerm 1 p = Empty.rec (snotz (≤0→≡0 (pred-≤-pred p))) -Fin>1→hasNonEqualTerm (suc (suc n)) _ = fzero , fone , fzero≠fone +Fin>1→hasNonEqualTerm (suc (suc n)) _ = fzero , fone , fzero≠fone -isEmpty→Fin≡0 : (n : ) ¬ Fin n 0 n +isEmpty→Fin≡0 : (n : ) ¬ Fin n 0 n isEmpty→Fin≡0 0 _ = refl -isEmpty→Fin≡0 (suc n) p = Empty.rec (p fzero) +isEmpty→Fin≡0 (suc n) p = Empty.rec (p fzero) -isInhab→Fin>0 : (n : ) Fin n 0 < n -isInhab→Fin>0 0 i = Empty.rec (¬Fin0 i) +isInhab→Fin>0 : (n : ) Fin n 0 < n +isInhab→Fin>0 0 i = Empty.rec (¬Fin0 i) isInhab→Fin>0 (suc n) _ = suc-≤-suc zero-≤ -hasNonEqualTerm→Fin>1 : (n : ) (i j : Fin n) ¬ i j 1 < n -hasNonEqualTerm→Fin>1 0 i _ _ = Empty.rec (¬Fin0 i) -hasNonEqualTerm→Fin>1 1 i j p = Empty.rec (p (isContr→isProp isContrFin1 i j)) +hasNonEqualTerm→Fin>1 : (n : ) (i j : Fin n) ¬ i j 1 < n +hasNonEqualTerm→Fin>1 0 i _ _ = Empty.rec (¬Fin0 i) +hasNonEqualTerm→Fin>1 1 i j p = Empty.rec (p (isContr→isProp isContrFin1 i j)) hasNonEqualTerm→Fin>1 (suc (suc n)) _ _ _ = suc-≤-suc (suc-≤-suc zero-≤) -Fin≤1→isProp : (n : ) n 1 isProp (Fin n) +Fin≤1→isProp : (n : ) n 1 isProp (Fin n) Fin≤1→isProp 0 _ = isPropFin0 -Fin≤1→isProp 1 _ = isContr→isProp isContrFin1 +Fin≤1→isProp 1 _ = isContr→isProp isContrFin1 Fin≤1→isProp (suc (suc n)) p = Empty.rec (¬-<-zero (pred-≤-pred p)) -isProp→Fin≤1 : (n : ) isProp (Fin n) n 1 +isProp→Fin≤1 : (n : ) isProp (Fin n) n 1 isProp→Fin≤1 0 _ = ≤-solver 0 1 isProp→Fin≤1 1 _ = ≤-solver 1 1 -isProp→Fin≤1 (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) +isProp→Fin≤1 (suc (suc n)) p = Empty.rec (fzero≠fone (p fzero fone)) \ No newline at end of file diff --git a/docs/Cubical.Data.FinData.Properties.html b/docs/Cubical.Data.FinData.Properties.html index aad12b8..e798c23 100644 --- a/docs/Cubical.Data.FinData.Properties.html +++ b/docs/Cubical.Data.FinData.Properties.html @@ -75,10 +75,10 @@ inj-cong : {n : } {k l : Fin n} (p : toℕ k toℕ l) cong toℕ (inj-toℕ p) p inj-cong p = isSetℕ _ _ _ _ -isPropFin0 : isProp (Fin 0) +isPropFin0 : isProp (Fin 0) isPropFin0 = ⊥.rec ¬Fin0 -isContrFin1 : isContr (Fin 1) +isContrFin1 : isContr (Fin 1) isContrFin1 .fst = zero isContrFin1 .snd zero = refl @@ -114,7 +114,7 @@ ... | yes p = yes (cong suc p) ... | no ¬p = no q ¬p (injSucFin q)) -isSetFin : ∀{k} isSet (Fin k) +isSetFin : ∀{k} isSet (Fin k) isSetFin = Discrete→isSet discreteFin isWeaken? : {n} (p : Fin (ℕsuc n)) Dec (Σ[ q Fin n ] p weakenFin q) @@ -196,7 +196,7 @@ toℕ∘enum {n = ℕsuc n} (ℕsuc m) p i = ℕsuc (toℕ∘enum m (pred-≤-pred p) i) enumExt : {m m' : }(p : m < n)(p' : m' < n) m m' enum m p enum m' p' -enumExt p p' q i = enum (q i) (isProp→PathP i isProp≤ {m = ℕsuc (q i)}) p p' i) +enumExt p p' q i = enum (q i) (isProp→PathP i isProp≤ {m = ℕsuc (q i)}) p p' i) enumInj : (p : m < k) (q : n < k) enum m p enum n q m n enumInj p q path = sym (toℕ∘enum _ p) cong toℕ path toℕ∘enum _ q @@ -336,10 +336,10 @@ Equiv : (n m : ) (Fin n × Fin m) Fin (n · m) Equiv ℕzero m = uninhabEquiv x ¬Fin0 (fst x)) ¬Fin0 - Equiv (ℕsuc n) m = Fin (ℕsuc n) × Fin m ≃⟨ isoToEquiv (sucProdToSumIso n m) - Fin m (Fin n × Fin m) ≃⟨ isoToEquiv (⊎Iso idIso (equivToIso (Equiv n m))) - Fin m Fin (n · m) ≃⟨ FinSumChar.Equiv m (n · m) - Fin (m + n · m) + Equiv (ℕsuc n) m = Fin (ℕsuc n) × Fin m ≃⟨ isoToEquiv (sucProdToSumIso n m) + Fin m (Fin n × Fin m) ≃⟨ isoToEquiv (⊎Iso idIso (equivToIso (Equiv n m))) + Fin m Fin (n · m) ≃⟨ FinSumChar.Equiv m (n · m) + Fin (m + n · m) -- Exhaustion of decidable predicate diff --git a/docs/Cubical.Data.List.Properties.html b/docs/Cubical.Data.List.Properties.html index a17298c..90eb2d7 100644 --- a/docs/Cubical.Data.List.Properties.html +++ b/docs/Cubical.Data.List.Properties.html @@ -40,7 +40,7 @@ rev-rev (x xs) = rev-snoc (rev xs) x cong (_∷_ x) (rev-rev xs) rev-rev-snoc : (xs : List A) (y : A) - Square (rev-rev (xs ++ [ y ])) (cong (_++ [ y ]) (rev-rev xs)) (cong rev (rev-snoc xs y)) refl + Square (rev-rev (xs ++ [ y ])) (cong (_++ [ y ]) (rev-rev xs)) (cong rev (rev-snoc xs y)) refl rev-rev-snoc [] y = sym (lUnit refl) rev-rev-snoc (x xs) y i j = hcomp @@ -66,25 +66,25 @@ module ListPath {} {A : Type } where Cover : List A List A Type - Cover [] [] = Lift Unit - Cover [] (_ _) = Lift - Cover (_ _) [] = Lift + Cover [] [] = Lift Unit + Cover [] (_ _) = Lift + Cover (_ _) [] = Lift Cover (x xs) (y ys) = (x y) × Cover xs ys reflCode : xs Cover xs xs - reflCode [] = lift tt + reflCode [] = lift tt reflCode (_ xs) = refl , reflCode xs encode : xs ys (p : xs ys) Cover xs ys - encode xs _ = J ys _ Cover xs ys) (reflCode xs) + encode xs _ = J ys _ Cover xs ys) (reflCode xs) encodeRefl : xs encode xs xs refl reflCode xs - encodeRefl xs = JRefl ys _ Cover xs ys) (reflCode xs) + encodeRefl xs = JRefl ys _ Cover xs ys) (reflCode xs) decode : xs ys Cover xs ys xs ys decode [] [] _ = refl - decode [] (_ _) (lift ()) - decode (x xs) [] (lift ()) + decode [] (_ _) (lift ()) + decode (x xs) [] (lift ()) decode (x xs) (y ys) (p , c) = cong₂ _∷_ p (decode xs ys c) decodeRefl : xs decode xs xs (reflCode xs) refl @@ -93,17 +93,17 @@ decodeEncode : xs ys (p : xs ys) decode xs ys (encode xs ys p) p decodeEncode xs _ = - J ys p decode xs ys (encode xs ys p) p) + J ys p decode xs ys (encode xs ys p) p) (cong (decode xs xs) (encodeRefl xs) decodeRefl xs) isOfHLevelCover : (n : HLevel) (p : isOfHLevel (suc (suc n)) A) (xs ys : List A) isOfHLevel (suc n) (Cover xs ys) isOfHLevelCover n p [] [] = - isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isPropUnit) + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isPropUnit) isOfHLevelCover n p [] (y ys) = - isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) isOfHLevelCover n p (x xs) [] = - isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) isOfHLevelCover n p (x xs) (y ys) = isOfHLevelΣ (suc n) (p x y) (\ _ isOfHLevelCover n p xs ys) @@ -140,10 +140,10 @@ cons-inj₂ = cong safe-tail ¬cons≡nil : {x : A} {xs} ¬ (x xs []) -¬cons≡nil {A = A} p = lower (subst (caseList (Lift ) (List A)) p []) +¬cons≡nil {A = A} p = lower (subst (caseList (Lift ) (List A)) p []) ¬nil≡cons : {x : A} {xs} ¬ ([] x xs) -¬nil≡cons {A = A} p = lower (subst (caseList (List A) (Lift )) p []) +¬nil≡cons {A = A} p = lower (subst (caseList (List A) (Lift )) p []) ¬snoc≡nil : {x : A} {xs} ¬ (xs ∷ʳ x []) ¬snoc≡nil {xs = []} contra = ¬cons≡nil contra @@ -156,10 +156,10 @@ cons≡rev-snoc _ [] = refl cons≡rev-snoc x (y ys) = λ i cons≡rev-snoc x ys i ++ y [] -isContr[]≡[] : isContr (Path (List A) [] []) +isContr[]≡[] : isContr (Path (List A) [] []) isContr[]≡[] = refl , ListPath.decodeEncode [] [] -isPropXs≡[] : {xs : List A} isProp (xs []) +isPropXs≡[] : {xs : List A} isProp (xs []) isPropXs≡[] {xs = []} = isOfHLevelSuc 0 isContr[]≡[] isPropXs≡[] {xs = x xs} = λ p _ ⊥.rec (¬cons≡nil p) diff --git a/docs/Cubical.Data.Maybe.Properties.html b/docs/Cubical.Data.Maybe.Properties.html index 76ab4dd..7b52822 100644 --- a/docs/Cubical.Data.Maybe.Properties.html +++ b/docs/Cubical.Data.Maybe.Properties.html @@ -46,20 +46,20 @@ -- Path space of Maybe type module MaybePath {} {A : Type } where Cover : Maybe A Maybe A Type - Cover nothing nothing = Lift Unit - Cover nothing (just _) = Lift - Cover (just _) nothing = Lift + Cover nothing nothing = Lift Unit + Cover nothing (just _) = Lift + Cover (just _) nothing = Lift Cover (just a) (just a') = a a' reflCode : (c : Maybe A) Cover c c - reflCode nothing = lift tt + reflCode nothing = lift tt reflCode (just b) = refl encode : c c' c c' Cover c c' - encode c _ = J c' _ Cover c c') (reflCode c) + encode c _ = J c' _ Cover c c') (reflCode c) encodeRefl : c encode c c refl reflCode c - encodeRefl c = JRefl c' _ Cover c c') (reflCode c) + encodeRefl c = JRefl c' _ Cover c c') (reflCode c) decode : c c' Cover c c' c c' decode nothing nothing _ = refl @@ -71,13 +71,13 @@ decodeEncode : c c' (p : c c') decode c c' (encode c c' p) p decodeEncode c _ = - J c' p decode c c' (encode c c' p) p) + J c' p decode c c' (encode c c' p) p) (cong (decode c c) (encodeRefl c) decodeRefl c) encodeDecode : c c' (d : Cover c c') encode c c' (decode c c' d) d encodeDecode nothing nothing _ = refl encodeDecode (just a) (just a') = - J a' p encode (just a) (just a') (cong just p) p) (encodeRefl (just a)) + J a' p encode (just a) (just a') (cong just p) p) (encodeRefl (just a)) Cover≃Path : c c' Cover c c' (c c') Cover≃Path c c' = isoToEquiv @@ -90,9 +90,9 @@ isOfHLevelCover : (n : HLevel) isOfHLevel (suc (suc n)) A c c' isOfHLevel (suc n) (Cover c c') - isOfHLevelCover n p nothing nothing = isOfHLevelLift (suc n) (isOfHLevelUnit (suc n)) - isOfHLevelCover n p nothing (just a') = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) - isOfHLevelCover n p (just a) nothing = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p nothing nothing = isOfHLevelLift (suc n) (isOfHLevelUnit (suc n)) + isOfHLevelCover n p nothing (just a') = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (just a) nothing = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) isOfHLevelCover n p (just a) (just a') = p a a' isOfHLevelMaybe : {} (n : HLevel) {A : Type } @@ -121,23 +121,23 @@ isEmbedding-just w z = MaybePath.Cover≃Path (just w) (just z) .snd ¬nothing≡just : {x : A} ¬ (nothing just x) -¬nothing≡just {A = A} {x = x} p = lower (subst (caseMaybe (Maybe A) (Lift )) p (just x)) +¬nothing≡just {A = A} {x = x} p = lower (subst (caseMaybe (Maybe A) (Lift )) p (just x)) ¬just≡nothing : {x : A} ¬ (just x nothing) -¬just≡nothing {A = A} {x = x} p = lower (subst (caseMaybe (Lift ) (Maybe A)) p (just x)) +¬just≡nothing {A = A} {x = x} p = lower (subst (caseMaybe (Lift ) (Maybe A)) p (just x)) -isProp-x≡nothing : (x : Maybe A) isProp (x nothing) +isProp-x≡nothing : (x : Maybe A) isProp (x nothing) isProp-x≡nothing nothing x w = - subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w isProp-x≡nothing (just _) p _ = ⊥.rec (¬just≡nothing p) -isProp-nothing≡x : (x : Maybe A) isProp (nothing x) +isProp-nothing≡x : (x : Maybe A) isProp (nothing x) isProp-nothing≡x nothing x w = - subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w isProp-nothing≡x (just _) p _ = ⊥.rec (¬nothing≡just p) -isContr-nothing≡nothing : isContr (nothing {A = A} nothing) -isContr-nothing≡nothing = inhProp→isContr refl (isProp-x≡nothing _) +isContr-nothing≡nothing : isContr (nothing {A = A} nothing) +isContr-nothing≡nothing = inhProp→isContr refl (isProp-x≡nothing _) discreteMaybe : Discrete A Discrete (Maybe A) discreteMaybe eqA nothing nothing = yes refl @@ -175,9 +175,9 @@ open Iso isom : Iso _ _ isom .fun = map-Maybe (equivFun e) - isom .inv = map-Maybe (invEq e) + isom .inv = map-Maybe (invEq e) isom .rightInv nothing = refl - isom .rightInv (just b) = cong just (secEq e b) + isom .rightInv (just b) = cong just (secEq e b) isom .leftInv nothing = refl - isom .leftInv (just a) = cong just (retEq e a) + isom .leftInv (just a) = cong just (retEq e a) \ No newline at end of file diff --git a/docs/Cubical.Data.Nat.Order.Recursive.html b/docs/Cubical.Data.Nat.Order.Recursive.html index b244adf..8e8fe28 100644 --- a/docs/Cubical.Data.Nat.Order.Recursive.html +++ b/docs/Cubical.Data.Nat.Order.Recursive.html @@ -46,7 +46,7 @@ P : Type k l m n : -isProp≤ : isProp (m n) +isProp≤ : isProp (m n) isProp≤ {zero} = isPropUnit isProp≤ {suc m} {zero} = isProp⊥ isProp≤ {suc m} {suc n} = isProp≤ {m} {n} @@ -127,17 +127,17 @@ = Sum.map (idfun _) (cong suc) (≤-split {m} {n} m≤n) module WellFounded where - wf-< : WellFounded _<_ - wf-rec-< : n WFRec _<_ (Acc _<_) n + wf-< : WellFounded _<_ + wf-rec-< : n WFRec _<_ (Acc _<_) n - wf-< n = acc (wf-rec-< n) + wf-< n = acc (wf-rec-< n) wf-rec-< (suc n) m m≤n with ≤-split {m} {n} m≤n ... | inl m<n = wf-rec-< n m m<n - ... | inr m≡n = subst⁻ (Acc _<_) m≡n (wf-< n) + ... | inr m≡n = subst⁻ (Acc _<_) m≡n (wf-< n) wf-elim : (∀ n (∀ m m < n P m) P n) n P n -wf-elim = WFI.induction WellFounded.wf-< +wf-elim = WFI.induction WellFounded.wf-< wf-rec : (∀ n (∀ m m < n R) R) R wf-rec {R = R} = wf-elim {P = λ _ R} @@ -146,7 +146,7 @@ Least : ∀{} ( Type ) ( Type ) Least P m = P m × (∀ n n < m ¬ P n) - isPropLeast : (∀ m isProp (P m)) m isProp (Least P m) + isPropLeast : (∀ m isProp (P m)) m isProp (Least P m) isPropLeast pP m = isPropΣ (pP m) _ isPropΠ3 λ _ _ _ isProp⊥) @@ -177,15 +177,15 @@ ... | eq m≡n = m≡n ... | gt n<m = Empty.rec (¬P<m n n<m Pn) - isPropΣLeast : (∀ m isProp (P m)) isProp (Σ _ (Least P)) + isPropΣLeast : (∀ m isProp (P m)) isProp (Σ _ (Least P)) isPropΣLeast pP (m , LPm) (n , LPn) = ΣPathP λ where .fst Least-unique m n LPm LPn - .snd isOfHLevel→isOfHLevelDep 1 (isPropLeast pP) + .snd isOfHLevel→isOfHLevelDep 1 (isPropLeast pP) LPm LPn (Least-unique m n LPm LPn) Decidable→Collapsible - : (∀ m isProp (P m)) (∀ m Dec (P m)) Collapsible (Σ P) + : (∀ m isProp (P m)) (∀ m Dec (P m)) Collapsible (Σ P) Decidable→Collapsible pP dP = λ where .fst Least→ →Least dP .snd x y cong Least→ (isPropΣLeast pP (→Least dP x) (→Least dP y)) diff --git a/docs/Cubical.Data.Nat.Order.html b/docs/Cubical.Data.Nat.Order.html index 6055c47..88e1038 100644 --- a/docs/Cubical.Data.Nat.Order.html +++ b/docs/Cubical.Data.Nat.Order.html @@ -46,10 +46,10 @@ k l m n : private - witness-prop : j isProp (j + m n) + witness-prop : j isProp (j + m n) witness-prop {m} {n} j = isSetℕ (j + m) n -isProp≤ : isProp (m n) +isProp≤ : isProp (m n) isProp≤ {m} {n} (k , p) (l , q) = Σ≡Prop witness-prop lemma where @@ -335,16 +335,16 @@ <-asym' = <-asym'-case (_≟_ _ _) private - acc-suc : Acc _<_ n Acc _<_ (suc n) + acc-suc : Acc _<_ n Acc _<_ (suc n) acc-suc a - = acc λ y y<sn + = acc λ y y<sn case <-split y<sn of λ - { (inl y<n) access a y y<n + { (inl y<n) access a y y<n ; (inr y≡n) subst _ (sym y≡n) a } -<-wellfounded : WellFounded _<_ -<-wellfounded zero = acc λ _ ⊥.rec ¬-<-zero +<-wellfounded : WellFounded _<_ +<-wellfounded zero = acc λ _ ⊥.rec ¬-<-zero <-wellfounded (suc n) = acc-suc (<-wellfounded n) <→≢ : n < m ¬ n m @@ -356,7 +356,7 @@ (base : n n < suc b₀ P n) (step : n P n P (suc b₀ + n)) where - open WFI (<-wellfounded) + open WFI (<-wellfounded) private dichotomy : b n (n < b) (Σ[ m ] n b + m) @@ -404,13 +404,13 @@ transportRefl _ +induction : n P n - +induction = induction wfStep + +induction = induction wfStep +inductionBase : n (l : n < b) +induction n base n l - +inductionBase n l = induction-compute wfStep n wfStepLemma₀ n l _ + +inductionBase n l = induction-compute wfStep n wfStepLemma₀ n l _ +inductionStep : n +induction (b + n) step n (+induction n) - +inductionStep n = induction-compute wfStep (b + n) wfStepLemma₁ n _ + +inductionStep n = induction-compute wfStep (b + n) wfStepLemma₁ n _ module <-Reasoning where -- TODO: would it be better to mirror the way it is done in the agda-stdlib? @@ -503,7 +503,7 @@ ... | yes m≤'n = yes (s≤s m≤'n) ... | no m≰'n = no λ { (s≤s m≤'n) m≰'n m≤'n } -≤'IsPropValued : m n isProp (m ≤' n) +≤'IsPropValued : m n isProp (m ≤' n) ≤'IsPropValued zero n z≤ z≤ = refl ≤'IsPropValued (suc m) zero () ≤'IsPropValued (suc m) (suc n) (s≤s x) (s≤s y) = cong s≤s (≤'IsPropValued m n x y) diff --git a/docs/Cubical.Data.Nat.Properties.html b/docs/Cubical.Data.Nat.Properties.html index 85a2d95..33cb8c5 100644 --- a/docs/Cubical.Data.Nat.Properties.html +++ b/docs/Cubical.Data.Nat.Properties.html @@ -96,7 +96,7 @@ reflRetr (suc n) i = cong suc (reflRetr n i) retr : (n m : ) (p : n m) (decodeℕ n m (encodeℕ n m p) p) - retr n m p = J m p decodeℕ n m (encodeℕ n m p) p) (reflRetr n) p + retr n m p = J m p decodeℕ n m (encodeℕ n m p) p) (reflRetr n) p ≡ℕ≃Codeℕ' : (n m : ) (n m) codeℕ n m @@ -117,7 +117,7 @@ reflRetr (suc n) i = cong suc (reflRetr n i) retr : (n m : ) (p : n m) decodeℕ n m (compute-eqℕ n m p) p - retr n m p = J m p decodeℕ n m (compute-eqℕ n m p) p) (reflRetr n) p + retr n m p = J m p decodeℕ n m (compute-eqℕ n m p) p) (reflRetr n) p discreteℕ : Discrete @@ -131,7 +131,7 @@ separatedℕ : Separated separatedℕ = Discrete→Separated discreteℕ -isSetℕ : isSet +isSetℕ : isSet isSetℕ = Discrete→isSet discreteℕ -- Arithmetic facts about predℕ @@ -284,15 +284,15 @@ ¬evenAndOdd (suc zero) () ¬evenAndOdd (suc (suc n)) = ¬evenAndOdd n -isPropIsEvenT : (n : ) isProp (isEvenT n) +isPropIsEvenT : (n : ) isProp (isEvenT n) isPropIsEvenT zero x y = refl isPropIsEvenT (suc zero) = isProp⊥ isPropIsEvenT (suc (suc n)) = isPropIsEvenT n -isPropIsOddT : (n : ) isProp (isOddT n) +isPropIsOddT : (n : ) isProp (isOddT n) isPropIsOddT n = isPropIsEvenT (suc n) -isPropEvenOrOdd : (n : ) isProp (isEvenT n isOddT n) +isPropEvenOrOdd : (n : ) isProp (isEvenT n isOddT n) isPropEvenOrOdd n (inl x) (inl x₁) = cong inl (isPropIsEvenT n x x₁) isPropEvenOrOdd n (inl x) (inr x₁) = ⊥.rec (¬evenAndOdd n (x , x₁)) isPropEvenOrOdd n (inr x) (inl x₁) = ⊥.rec (¬evenAndOdd (suc n) (x , x₁)) diff --git a/docs/Cubical.Data.Sigma.Base.html b/docs/Cubical.Data.Sigma.Base.html index 2c7ca4d..e29f596 100644 --- a/docs/Cubical.Data.Sigma.Base.html +++ b/docs/Cubical.Data.Sigma.Base.html @@ -43,7 +43,7 @@ -- Unique existence ∃! : { ℓ'} (A : Type ) (B : A Type ℓ') Type (ℓ-max ℓ') -∃! A B = isContr (Σ A B) +∃! A B = isContr (Σ A B) infix 2 ∃!-syntax diff --git a/docs/Cubical.Data.Sigma.Properties.html b/docs/Cubical.Data.Sigma.Properties.html index 1ace07e..18e494a 100644 --- a/docs/Cubical.Data.Sigma.Properties.html +++ b/docs/Cubical.Data.Sigma.Properties.html @@ -89,13 +89,13 @@ (PathP i Σ (A i) (B i)) x y) ΣPath≡PathΣ = ua ΣPath≃PathΣ -×≡Prop : isProp A' {u v : A × A'} u .fst v .fst u v +×≡Prop : isProp A' {u v : A × A'} u .fst v .fst u v ×≡Prop pB {u} {v} p i = (p i) , (pB (u .snd) (v .snd) i) -- Useful lemma to prove unique existence -uniqueExists : (a : A) (b : B a) (h : (a' : A) isProp (B a')) (H : (a' : A) B a' a a') ∃![ a A ] B a +uniqueExists : (a : A) (b : B a) (h : (a' : A) isProp (B a')) (H : (a' : A) B a' a a') ∃![ a A ] B a fst (uniqueExists a b h H) = (a , b) -snd (uniqueExists a b h H) (a' , b') = ΣPathP (H a' b' , isProp→PathP i h (H a' b' i)) b b') +snd (uniqueExists a b h H) (a' , b') = ΣPathP (H a' b' , isProp→PathP i h (H a' b' i)) b b') -- Characterization of dependent paths in Σ @@ -123,7 +123,7 @@ discreteΣ {B = B} Adis Bdis (a0 , b0) (a1 , b1) = discreteΣ' (Adis a0 a1) where discreteΣ' : Dec (a0 a1) Dec ((a0 , b0) (a1 , b1)) - discreteΣ' (yes p) = J a1 p b1 Dec ((a0 , b0) (a1 , b1))) (discreteΣ'') p b1 + discreteΣ' (yes p) = J a1 p b1 Dec ((a0 , b0) (a1 , b1))) (discreteΣ'') p b1 where discreteΣ'' : (b1 : B a0) Dec ((a0 , b0) (a0 , b1)) discreteΣ'' b1 with Bdis a0 b0 b1 @@ -195,14 +195,14 @@ inv (Σ-cong-iso-fst {B = B} isom) x = inv isom (x .fst) , subst B (sym (ε (x .fst))) (x .snd) where ε = isHAEquiv.rinv (snd (iso→HAEquiv isom)) -rightInv (Σ-cong-iso-fst {B = B} isom) (x , y) = ΣPathP (ε x , toPathP goal) +rightInv (Σ-cong-iso-fst {B = B} isom) (x , y) = ΣPathP (ε x , toPathP goal) where ε = isHAEquiv.rinv (snd (iso→HAEquiv isom)) goal : subst B (ε x) (subst B (sym (ε x)) y) y goal = sym (substComposite B (sym (ε x)) (ε x) y) ∙∙ cong x subst B x y) (lCancel (ε x)) ∙∙ substRefl {B = B} y -leftInv (Σ-cong-iso-fst {A = A} {B = B} isom) (x , y) = ΣPathP (leftInv isom x , toPathP goal) +leftInv (Σ-cong-iso-fst {A = A} {B = B} isom) (x , y) = ΣPathP (leftInv isom x , toPathP goal) where ε = isHAEquiv.rinv (snd (iso→HAEquiv isom)) γ = isHAEquiv.com (snd (iso→HAEquiv isom)) @@ -264,7 +264,7 @@ leftInv (Σ-cong-iso-snd isom) (x , y') = ΣPathP (refl , leftInv (isom x) y') Σ-cong-equiv-snd : (∀ a B a B' a) Σ A B Σ A B' -Σ-cong-equiv-snd h = isoToEquiv (Σ-cong-iso-snd (equivToIso h)) +Σ-cong-equiv-snd h = isoToEquiv (Σ-cong-iso-snd (equivToIso h)) Σ-cong-snd : ((x : A) B x B' x) Σ A B Σ A B' Σ-cong-snd {A = A} p i = Σ[ x A ] (p x i) @@ -277,20 +277,20 @@ Σ-cong-equiv : (e : A A') ((x : A) B x B' (equivFun e x)) Σ A B Σ A' B' -Σ-cong-equiv e e' = isoToEquiv (Σ-cong-iso (equivToIso e) (equivToIso e')) +Σ-cong-equiv e e' = isoToEquiv (Σ-cong-iso (equivToIso e) (equivToIso e')) Σ-cong' : (p : A A') PathP i p i Type ℓ') B B' Σ A B Σ A' B' Σ-cong' p p' = cong₂ (A : Type _) (B : A Type _) Σ A B) p p' Σ-cong-equiv-prop : (e : A A') - ((x : A ) isProp (B x)) - ((x : A') isProp (B' x)) + ((x : A ) isProp (B x)) + ((x : A') isProp (B' x)) ((x : A) B x B' (equivFun e x)) ((x : A) B' (equivFun e x) B x) Σ A B Σ A' B' Σ-cong-equiv-prop e prop prop' prop→ prop← = - Σ-cong-equiv e x propBiimpl→Equiv (prop x) (prop' (equivFun e x)) (prop→ x) (prop← x)) + Σ-cong-equiv e x propBiimpl→Equiv (prop x) (prop' (equivFun e x)) (prop→ x) (prop← x)) -- Alternative version for path in Σ-types, as in the HoTT book @@ -314,16 +314,16 @@ ΣPathTransport≡PathΣ : (a b : Σ A B) ΣPathTransport a b (a b) ΣPathTransport≡PathΣ a b = ua (ΣPathTransport≃PathΣ a b) -Σ-contractFstIso : (c : isContr A) Iso (Σ A B) (B (c .fst)) +Σ-contractFstIso : (c : isContr A) Iso (Σ A B) (B (c .fst)) fun (Σ-contractFstIso {B = B} c) p = subst B (sym (c .snd (fst p))) (snd p) inv (Σ-contractFstIso {B = B} c) b = _ , b rightInv (Σ-contractFstIso {B = B} c) b = - cong p subst B p b) (isProp→isSet (isContr→isProp c) _ _ _ _) transportRefl _ + cong p subst B p b) (isProp→isSet (isContr→isProp c) _ _ _ _) transportRefl _ fst (leftInv (Σ-contractFstIso {B = B} c) p j) = c .snd (fst p) j snd (leftInv (Σ-contractFstIso {B = B} c) p j) = transp i B (c .snd (fst p) (~ i j))) j (snd p) -Σ-contractFst : (c : isContr A) Σ A B B (c .fst) +Σ-contractFst : (c : isContr A) Σ A B B (c .fst) Σ-contractFst {B = B} c = isoToEquiv (Σ-contractFstIso c) -- a special case of the above @@ -331,7 +331,7 @@ ΣUnit : Σ Unit A A tt unquoteDef ΣUnit = defStrictEquiv ΣUnit snd { x (tt , x) }) -Σ-contractSnd : ((a : A) isContr (B a)) Σ A B A +Σ-contractSnd : ((a : A) isContr (B a)) Σ A B A Σ-contractSnd c = isoToEquiv isom where isom : Iso _ _ @@ -340,44 +340,44 @@ isom .rightInv _ = refl isom .leftInv (a , b) = cong (a ,_) (c a .snd b) -isEmbeddingFstΣProp : ((x : A) isProp (B x)) +isEmbeddingFstΣProp : ((x : A) isProp (B x)) {u v : Σ A B} isEquiv (p : u v) cong fst p) isEmbeddingFstΣProp {B = B} pB {u = u} {v = v} .equiv-proof x = ctr , isCtr where ctrP : u v - ctrP = ΣPathP (x , isProp→PathP _ pB _) _ _) + ctrP = ΣPathP (x , isProp→PathP _ pB _) _ _) ctr : fiber (p : u v) cong fst p) x ctr = ctrP , refl isCtr : z ctr z isCtr (z , p) = ΣPathP (ctrP≡ , cong (sym snd) fzsingl) where - fzsingl : Path (singl x) (x , refl) (cong fst z , sym p) - fzsingl = isContrSingl x .snd (cong fst z , sym p) - ctrSnd : SquareP i j B (fzsingl i .fst j)) (cong snd ctrP) (cong snd z) _ _ + fzsingl : Path (singl x) (x , refl) (cong fst z , sym p) + fzsingl = isContrSingl x .snd (cong fst z , sym p) + ctrSnd : SquareP i j B (fzsingl i .fst j)) (cong snd ctrP) (cong snd z) _ _ ctrSnd = isProp→SquareP _ _ pB _) _ _ _ _ ctrP≡ : ctrP z ctrP≡ i = ΣPathP (fzsingl i .fst , ctrSnd i) -Σ≡PropEquiv : ((x : A) isProp (B x)) {u v : Σ A B} +Σ≡PropEquiv : ((x : A) isProp (B x)) {u v : Σ A B} (u .fst v .fst) (u v) -Σ≡PropEquiv pB = invEquiv (_ , isEmbeddingFstΣProp pB) +Σ≡PropEquiv pB = invEquiv (_ , isEmbeddingFstΣProp pB) -Σ≡Prop : ((x : A) isProp (B x)) {u v : Σ A B} +Σ≡Prop : ((x : A) isProp (B x)) {u v : Σ A B} (p : u .fst v .fst) u v Σ≡Prop pB p = equivFun (Σ≡PropEquiv pB) p -- dependent version ΣPathPProp : { ℓ'} {A : I Type } {B : (i : I) A i Type ℓ'} {u : Σ (A i0) (B i0)} {v : Σ (A i1) (B i1)} - ((a : A (i1)) isProp (B i1 a)) + ((a : A (i1)) isProp (B i1 a)) PathP i A i) (fst u) (fst v) PathP i Σ (A i) (B i)) u v fst (ΣPathPProp {u = u} {v = v} pB p i) = p i snd (ΣPathPProp {B = B} {u = u} {v = v} pB p i) = lem i where lem : PathP i B i (p i)) (snd u) (snd v) - lem = toPathP (pB _ _ _) + lem = toPathP (pB _ _ _) ≃-× : {ℓ'' ℓ'''} {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} A C B D A × B C × D ≃-× eq1 eq2 = @@ -409,11 +409,11 @@ prodEquivToIso : {ℓ'' ℓ'''} {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} (e : A C)(e' : B D) - prodIso (equivToIso e) (equivToIso e') equivToIso (≃-× e e') -Iso.fun (prodEquivToIso e e' i) = Iso.fun (equivToIso (≃-× e e')) -Iso.inv (prodEquivToIso e e' i) = Iso.inv (equivToIso (≃-× e e')) -Iso.rightInv (prodEquivToIso e e' i) = Iso.rightInv (equivToIso (≃-× e e')) -Iso.leftInv (prodEquivToIso e e' i) = Iso.leftInv (equivToIso (≃-× e e')) + prodIso (equivToIso e) (equivToIso e') equivToIso (≃-× e e') +Iso.fun (prodEquivToIso e e' i) = Iso.fun (equivToIso (≃-× e e')) +Iso.inv (prodEquivToIso e e' i) = Iso.inv (equivToIso (≃-× e e')) +Iso.rightInv (prodEquivToIso e e' i) = Iso.rightInv (equivToIso (≃-× e e')) +Iso.leftInv (prodEquivToIso e e' i) = Iso.leftInv (equivToIso (≃-× e e')) toProdIso : {B C : A Type } Iso ((a : A) B a × C a) (((a : A) B a) × ((a : A) C a)) diff --git a/docs/Cubical.Data.Sum.Properties.html b/docs/Cubical.Data.Sum.Properties.html index ff43422..93e0e40 100644 --- a/docs/Cubical.Data.Sum.Properties.html +++ b/docs/Cubical.Data.Sum.Properties.html @@ -33,26 +33,26 @@ module ⊎Path { ℓ'} {A : Type } {B : Type ℓ'} where Cover : A B A B Type (ℓ-max ℓ') - Cover (inl a) (inl a') = Lift {j = ℓ-max ℓ'} (a a') - Cover (inl _) (inr _) = Lift - Cover (inr _) (inl _) = Lift - Cover (inr b) (inr b') = Lift {j = ℓ-max ℓ'} (b b') + Cover (inl a) (inl a') = Lift {j = ℓ-max ℓ'} (a a') + Cover (inl _) (inr _) = Lift + Cover (inr _) (inl _) = Lift + Cover (inr b) (inr b') = Lift {j = ℓ-max ℓ'} (b b') reflCode : (c : A B) Cover c c - reflCode (inl a) = lift refl - reflCode (inr b) = lift refl + reflCode (inl a) = lift refl + reflCode (inr b) = lift refl encode : c c' c c' Cover c c' - encode c _ = J c' _ Cover c c') (reflCode c) + encode c _ = J c' _ Cover c c') (reflCode c) encodeRefl : c encode c c refl reflCode c - encodeRefl c = JRefl c' _ Cover c c') (reflCode c) + encodeRefl c = JRefl c' _ Cover c c') (reflCode c) decode : c c' Cover c c' c c' - decode (inl a) (inl a') (lift p) = cong inl p + decode (inl a) (inl a') (lift p) = cong inl p decode (inl a) (inr b') () decode (inr b) (inl a') () - decode (inr b) (inr b') (lift q) = cong inr q + decode (inr b) (inr b') (lift q) = cong inr q decodeRefl : c decode c c (reflCode c) refl decodeRefl (inl a) = refl @@ -60,14 +60,14 @@ decodeEncode : c c' (p : c c') decode c c' (encode c c' p) p decodeEncode c _ = - J c' p decode c c' (encode c c' p) p) + J c' p decode c c' (encode c c' p) p) (cong (decode c c) (encodeRefl c) decodeRefl c) encodeDecode : c c' (d : Cover c c') encode c c' (decode c c' d) d - encodeDecode (inl a) (inl _) (lift d) = - J a' p encode (inl a) (inl a') (cong inl p) lift p) (encodeRefl (inl a)) d - encodeDecode (inr a) (inr _) (lift d) = - J a' p encode (inr a) (inr a') (cong inr p) lift p) (encodeRefl (inr a)) d + encodeDecode (inl a) (inl _) (lift d) = + J a' p encode (inl a) (inl a') (cong inl p) lift p) (encodeRefl (inl a)) d + encodeDecode (inr a) (inr _) (lift d) = + J a' p encode (inr a) (inr a') (cong inr p) lift p) (encodeRefl (inr a)) d Cover≃Path : c c' Cover c c' (c c') Cover≃Path c c' = @@ -77,18 +77,18 @@ isOfHLevel (suc (suc n)) A isOfHLevel (suc (suc n)) B c c' isOfHLevel (suc n) (Cover c c') - isOfHLevelCover n p q (inl a) (inl a') = isOfHLevelLift (suc n) (p a a') + isOfHLevelCover n p q (inl a) (inl a') = isOfHLevelLift (suc n) (p a a') isOfHLevelCover n p q (inl a) (inr b') = - isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) isOfHLevelCover n p q (inr b) (inl a') = - isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) - isOfHLevelCover n p q (inr b) (inr b') = isOfHLevelLift (suc n) (q b b') + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p q (inr b) (inr b') = isOfHLevelLift (suc n) (q b b') isEmbedding-inl : isEmbedding (inl {A = A} {B = B}) -isEmbedding-inl w z = snd (compEquiv LiftEquiv (⊎Path.Cover≃Path (inl w) (inl z))) +isEmbedding-inl w z = snd (compEquiv LiftEquiv (⊎Path.Cover≃Path (inl w) (inl z))) isEmbedding-inr : isEmbedding (inr {A = A} {B = B}) -isEmbedding-inr w z = snd (compEquiv LiftEquiv (⊎Path.Cover≃Path (inr w) (inr z))) +isEmbedding-inr w z = snd (compEquiv LiftEquiv (⊎Path.Cover≃Path (inr w) (inr z))) isOfHLevel⊎ : (n : HLevel) isOfHLevel (suc (suc n)) A @@ -101,26 +101,26 @@ (⊎Path.decodeEncode c c') (⊎Path.isOfHLevelCover n lA lB c c') -isProp⊎ : isProp A isProp B (A B ) isProp (A B) +isProp⊎ : isProp A isProp B (A B ) isProp (A B) isProp⊎ propA _ _ (inl x) (inl y) i = inl (propA x y i) isProp⊎ _ _ AB⊥ (inl x) (inr y) = ⊥.rec (AB⊥ x y) isProp⊎ _ _ AB⊥ (inr x) (inl y) = ⊥.rec (AB⊥ y x) isProp⊎ _ propB _ (inr x) (inr y) i = inr (propB x y i) -isSet⊎ : isSet A isSet B isSet (A B) +isSet⊎ : isSet A isSet B isSet (A B) isSet⊎ = isOfHLevel⊎ 0 -isGroupoid⊎ : isGroupoid A isGroupoid B isGroupoid (A B) +isGroupoid⊎ : isGroupoid A isGroupoid B isGroupoid (A B) isGroupoid⊎ = isOfHLevel⊎ 1 -is2Groupoid⊎ : is2Groupoid A is2Groupoid B is2Groupoid (A B) +is2Groupoid⊎ : is2Groupoid A is2Groupoid B is2Groupoid (A B) is2Groupoid⊎ = isOfHLevel⊎ 2 discrete⊎ : Discrete A Discrete B Discrete (A B) discrete⊎ decA decB (inl a) (inl a') = mapDec (cong inl) p q p (isEmbedding→Inj isEmbedding-inl _ _ q)) (decA a a') -discrete⊎ decA decB (inl a) (inr b') = no p lower (⊎Path.encode (inl a) (inr b') p)) -discrete⊎ decA decB (inr b) (inl a') = no ((λ p lower (⊎Path.encode (inr b) (inl a') p))) +discrete⊎ decA decB (inl a) (inr b') = no p lower (⊎Path.encode (inl a) (inr b') p)) +discrete⊎ decA decB (inr b) (inl a') = no ((λ p lower (⊎Path.encode (inr b) (inl a') p))) discrete⊎ decA decB (inr b) (inr b') = mapDec (cong inr) p q p (isEmbedding→Inj isEmbedding-inr _ _ q)) (decB b b') @@ -135,7 +135,7 @@ leftInv (⊎Iso iac ibd) (inr x) = cong inr (ibd .leftInv x) ⊎-equiv : A C B D (A B) (C D) -⊎-equiv p q = isoToEquiv (⊎Iso (equivToIso p) (equivToIso q)) +⊎-equiv p q = isoToEquiv (⊎Iso (equivToIso p) (equivToIso q)) ⊎-swap-Iso : Iso (A B) (B A) fun ⊎-swap-Iso (inl x) = inr x @@ -223,15 +223,15 @@ leftInv Σ⊎Iso (inl a , ea) = refl leftInv Σ⊎Iso (inr b , eb) = refl -×DistL⊎Iso : Iso (A × (B C)) ((A × B) (A × C)) -fun ×DistL⊎Iso (a , inl b) = inl (a , b) -fun ×DistL⊎Iso (a , inr c) = inr (a , c) -inv ×DistL⊎Iso (inl (a , b)) = a , inl b -inv ×DistL⊎Iso (inr (a , c)) = a , inr c -rightInv ×DistL⊎Iso (inl (a , b)) = refl -rightInv ×DistL⊎Iso (inr (a , c)) = refl -leftInv ×DistL⊎Iso (a , inl b) = refl -leftInv ×DistL⊎Iso (a , inr c) = refl +×DistR⊎Iso : Iso (A × (B C)) ((A × B) (A × C)) +fun ×DistR⊎Iso (a , inl b) = inl (a , b) +fun ×DistR⊎Iso (a , inr c) = inr (a , c) +inv ×DistR⊎Iso (inl (a , b)) = a , inl b +inv ×DistR⊎Iso (inr (a , c)) = a , inr c +rightInv ×DistR⊎Iso (inl (a , b)) = refl +rightInv ×DistR⊎Iso (inr (a , c)) = refl +leftInv ×DistR⊎Iso (a , inl b) = refl +leftInv ×DistR⊎Iso (a , inr c) = refl Π⊎≃ : ((x : A B) E x) ((a : A) E (inl a)) × ((b : B) E (inr b)) Π⊎≃ = isoToEquiv Π⊎Iso @@ -243,20 +243,20 @@ ⊎Monotone↪ (f , embf) (g , embg) = (map f g) , emb where coverToMap : x y ⊎Path.Cover x y ⊎Path.Cover (map f g x) (map f g y) - coverToMap (inl _) (inl _) cover = lift (cong f (lower cover)) - coverToMap (inr _) (inr _) cover = lift (cong g (lower cover)) + coverToMap (inl _) (inl _) cover = lift (cong f (lower cover)) + coverToMap (inr _) (inr _) cover = lift (cong g (lower cover)) equiv : x y isEquiv (coverToMap x y) equiv (inl a₀) (inl a₁) - = ((invEquiv LiftEquiv) - ∙ₑ ((cong f) , (embf a₀ a₁)) - ∙ₑ LiftEquiv) .snd + = ((invEquiv LiftEquiv) + ∙ₑ ((cong f) , (embf a₀ a₁)) + ∙ₑ LiftEquiv) .snd equiv (inl a₀) (inr b₁) .equiv-proof () equiv (inr b₀) (inl a₁) .equiv-proof () equiv (inr b₀) (inr b₁) - = ((invEquiv LiftEquiv) - ∙ₑ ((cong g) , (embg b₀ b₁)) - ∙ₑ LiftEquiv) .snd + = ((invEquiv LiftEquiv) + ∙ₑ ((cong g) , (embg b₀ b₁)) + ∙ₑ LiftEquiv) .snd lemma : x y (p : x y) @@ -266,7 +266,7 @@ (map f g y) (coverToMap x y (⊎Path.encode x y p)) lemma (inl a₀) _ - = J y p + = J y p cong (map f g) p ⊎Path.decode (map f g (inl a₀)) (map f g y) @@ -274,7 +274,7 @@ (⊎Path.encode (inl a₀) y p))) (sym $ cong (cong inl) (cong (cong f) (transportRefl _))) lemma (inr b₀) _ - = J y p + = J y p cong (map f g) p ⊎Path.decode (map f g (inr b₀)) @@ -285,12 +285,12 @@ emb : isEmbedding (map f g) emb x y = subst eq isEquiv eq) (sym (funExt (lemma x y))) - ((x y ≃⟨ invEquiv (⊎Path.Cover≃Path x y) - ⊎Path.Cover x y ≃⟨ (coverToMap x y) , (equiv x y) + ((x y ≃⟨ invEquiv (⊎Path.Cover≃Path x y) + ⊎Path.Cover x y ≃⟨ (coverToMap x y) , (equiv x y) ⊎Path.Cover (map f g x) - (map f g y) ≃⟨ ⊎Path.Cover≃Path + (map f g y) ≃⟨ ⊎Path.Cover≃Path (map f g x) - (map f g y) - map f g x map f g y ) .snd) + (map f g y) + map f g x map f g y ) .snd) \ No newline at end of file diff --git a/docs/Cubical.Data.Unit.Base.html b/docs/Cubical.Data.Unit.Base.html index a883e44..cc667c4 100644 --- a/docs/Cubical.Data.Unit.Base.html +++ b/docs/Cubical.Data.Unit.Base.html @@ -10,9 +10,9 @@ -- Universe polymorphic version Unit* : { : Level} Type -Unit* = Lift Unit +Unit* = Lift Unit -pattern tt* = lift tt +pattern tt* = lift tt -- Pointed version Unit*∙ : {} Σ[ X Type ] X diff --git a/docs/Cubical.Data.Unit.Properties.html b/docs/Cubical.Data.Unit.Properties.html index 624a6f6..5d4fe6c 100644 --- a/docs/Cubical.Data.Unit.Properties.html +++ b/docs/Cubical.Data.Unit.Properties.html @@ -26,14 +26,14 @@ variable ℓ' : Level -isContrUnit : isContr Unit +isContrUnit : isContr Unit isContrUnit = tt , λ {tt refl} -isPropUnit : isProp Unit +isPropUnit : isProp Unit isPropUnit _ _ i = tt -- definitionally equal to: isContr→isProp isContrUnit -isSetUnit : isSet Unit -isSetUnit = isProp→isSet isPropUnit +isSetUnit : isSet Unit +isSetUnit = isProp→isSet isPropUnit isOfHLevelUnit : (n : HLevel) isOfHLevel n Unit isOfHLevelUnit n = isContr→isOfHLevel n isContrUnit @@ -77,7 +77,7 @@ rightInv fiberUnitIso _ = refl leftInv fiberUnitIso _ = refl -isContr→Iso2 : {A : Type } {B : Type ℓ'} isContr A Iso (A B) B +isContr→Iso2 : {A : Type } {B : Type ℓ'} isContr A Iso (A B) B fun (isContr→Iso2 iscontr) f = f (fst iscontr) inv (isContr→Iso2 iscontr) b _ = b rightInv (isContr→Iso2 iscontr) _ = refl @@ -91,19 +91,19 @@ where unquoteDecl e = declStrictEquiv e fst a a , refl) -isContr→≃Unit : {A : Type } isContr A A Unit +isContr→≃Unit : {A : Type } isContr A A Unit isContr→≃Unit contr = isoToEquiv (iso _ tt) _ fst contr) _ refl) λ _ snd contr _) -isContr→≡Unit : {A : Type₀} isContr A A Unit +isContr→≡Unit : {A : Type₀} isContr A A Unit isContr→≡Unit contr = ua (isContr→≃Unit contr) -isContrUnit* : {} isContr (Unit* {}) +isContrUnit* : {} isContr (Unit* {}) isContrUnit* = tt* , λ _ refl -isPropUnit* : {} isProp (Unit* {}) +isPropUnit* : {} isProp (Unit* {}) isPropUnit* _ _ = refl -isSetUnit* : {} isSet (Unit* {}) +isSetUnit* : {} isSet (Unit* {}) isSetUnit* _ _ _ _ = refl isOfHLevelUnit* : {} (n : HLevel) isOfHLevel n (Unit* {}) @@ -113,11 +113,11 @@ isOfHLevelUnit* (suc (suc (suc n))) = isOfHLevelPlus 3 (isOfHLevelUnit* n) Unit≃Unit* : {} Unit Unit* {} -Unit≃Unit* = invEquiv (isContr→≃Unit isContrUnit*) +Unit≃Unit* = invEquiv (isContr→≃Unit isContrUnit*) -isContr→≃Unit* : {A : Type } isContr A A Unit* {} -isContr→≃Unit* contr = compEquiv (isContr→≃Unit contr) Unit≃Unit* +isContr→≃Unit* : {A : Type } isContr A A Unit* {} +isContr→≃Unit* contr = compEquiv (isContr→≃Unit contr) Unit≃Unit* -isContr→≡Unit* : {A : Type } isContr A A Unit* +isContr→≡Unit* : {A : Type } isContr A A Unit* isContr→≡Unit* contr = ua (isContr→≃Unit* contr) \ No newline at end of file diff --git a/docs/Cubical.Data.Vec.Properties.html b/docs/Cubical.Data.Vec.Properties.html index 3bb2ebd..6f5baa4 100644 --- a/docs/Cubical.Data.Vec.Properties.html +++ b/docs/Cubical.Data.Vec.Properties.html @@ -63,7 +63,7 @@ FinVec≡Vec : (n : ) FinVec A n Vec A n FinVec≡Vec n = ua (FinVec≃Vec n) -isContrVec0 : isContr (Vec A 0) +isContrVec0 : isContr (Vec A 0) isContrVec0 = [] , λ { [] refl } -- encode - decode Vec @@ -80,10 +80,10 @@ reflEncode (a v) = refl , refl encode : {n : } (v v' : Vec A n) (v v') code v v' - encode v v' p = J v' _ code v v') (reflEncode v) p + encode v v' p = J v' _ code v v') (reflEncode v) p encodeRefl : {n : } (v : Vec A n) encode v v refl reflEncode v - encodeRefl v = JRefl v' _ code v v') (reflEncode v) + encodeRefl v = JRefl v' _ code v v') (reflEncode v) -- decode decode : {n : } (v v' : Vec A n) (r : code v v') (v v') @@ -106,32 +106,32 @@ sect : {n : } (v v' : Vec A n) (r : code v v') encode v v' (decode v v' r) r sect [] [] tt* = encodeRefl [] - sect (a v) (a' v') (p , q) = J a' p encode (a v) (a' v') (decode (a v) (a' v') (p , q)) (p , q)) - (J v' q encode (a v) (a v') (decode (a v) (a v') (refl , q)) (refl , q)) + sect (a v) (a' v') (p , q) = J a' p encode (a v) (a' v') (decode (a v) (a' v') (p , q)) (p , q)) + (J v' q encode (a v) (a v') (decode (a v) (a v') (refl , q)) (refl , q)) (encodeRefl (a v)) q) p leftInv is = retr v v' where retr : {n : } (v v' : Vec A n) (p : v v') decode v v' (encode v v' p) p - retr v v' p = J v' p decode v v' (encode v v' p) p) + retr v v' p = J v' p decode v v' (encode v v' p) p) (cong (decode v v) (encodeRefl v) decodeRefl v) p isOfHLevelVec : (h : HLevel) (n : ) isOfHLevel (suc (suc h)) A isOfHLevel (suc (suc h)) (Vec A n) - isOfHLevelVec h zero ofLevelA [] [] = isOfHLevelRespectEquiv (suc h) (invEquiv (≡Vec≃codeVec [] [])) + isOfHLevelVec h zero ofLevelA [] [] = isOfHLevelRespectEquiv (suc h) (invEquiv (≡Vec≃codeVec [] [])) (isOfHLevelUnit* (suc h)) - isOfHLevelVec h (suc n) ofLevelA (x v) (x' v') = isOfHLevelRespectEquiv (suc h) (invEquiv (≡Vec≃codeVec _ _)) + isOfHLevelVec h (suc n) ofLevelA (x v) (x' v') = isOfHLevelRespectEquiv (suc h) (invEquiv (≡Vec≃codeVec _ _)) (isOfHLevelΣ (suc h) (ofLevelA x x') _ isOfHLevelVec h n ofLevelA v v')) discreteA→discreteVecA : Discrete A (n : ) Discrete (Vec A n) discreteA→discreteVecA DA zero [] [] = yes refl discreteA→discreteVecA DA (suc n) (a v) (a' v') with (DA a a') | (discreteA→discreteVecA DA n v v') - ... | yes p | yes q = yes (invIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) (p , q)) - ... | yes p | no ¬q = no r ¬q (snd (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) - ... | no ¬p | yes q = no r ¬p (fst (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) - ... | no ¬p | no ¬q = no r ¬q (snd (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) + ... | yes p | yes q = yes (invIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) (p , q)) + ... | yes p | no ¬q = no r ¬q (snd (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) + ... | no ¬p | yes q = no r ¬p (fst (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) + ... | no ¬p | no ¬q = no r ¬q (snd (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) ≢-∷ : {m : } (Discrete A) (a : A) (v : Vec A m) (a' : A) (v' : Vec A m) (a v a' v' ⊥.⊥) (a a' ⊥.⊥) (v v' ⊥.⊥) diff --git a/docs/Cubical.Displayed.Base.html b/docs/Cubical.Displayed.Base.html index 403e9ba..db0daec 100644 --- a/docs/Cubical.Displayed.Base.html +++ b/docs/Cubical.Displayed.Base.html @@ -28,7 +28,7 @@ ua : (a a' : A) (a a') (a a') uaIso : (a a' : A) Iso (a a') (a a') - uaIso a a' = equivToIso (ua a a') + uaIso a a' = equivToIso (ua a a') ≅→≡ : {a a' : A} (p : a a') a a' ≅→≡ {a} {a'} = Iso.fun (uaIso a a') @@ -38,13 +38,13 @@ ρ : (a : A) a a ρ a = ≡→≅ refl -open BinaryRelation +open BinaryRelation -- another constructor for UARel using contractibility of relational singletons make-𝒮 : {A : Type ℓA} {_≅_ : A A Type ℓ≅A} - (ρ : isRefl _≅_) (contrTotal : contrRelSingl _≅_) UARel A ℓ≅A + (ρ : isRefl _≅_) (contrTotal : contrRelSingl _≅_) UARel A ℓ≅A UARel._≅_ (make-𝒮 {_≅_ = _≅_} _ _) = _≅_ -UARel.ua (make-𝒮 {_≅_ = _≅_} ρ c) = contrRelSingl→isUnivalent _≅_ ρ c +UARel.ua (make-𝒮 {_≅_ = _≅_} ρ c) = contrRelSingl→isUnivalent _≅_ ρ c record DUARel {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (B : A Type ℓB) (ℓ≅B : Level) : Type (ℓ-max (ℓ-max ℓA ℓB) (ℓ-max ℓ≅A (ℓ-suc ℓ≅B))) where @@ -56,17 +56,17 @@ _≅ᴰ⟨_⟩_ : {a a' : A} B a a a' B a' Type ℓ≅B uaᴰ : {a a' : A} (b : B a) (p : a a') (b' : B a') (b ≅ᴰ⟨ p b') PathP i B (≅→≡ p i)) b b' - fiberRel : (a : A) Rel (B a) (B a) ℓ≅B + fiberRel : (a : A) Rel (B a) (B a) ℓ≅B fiberRel a = _≅ᴰ⟨ ρ a ⟩_ uaᴰρ : {a : A} (b b' : B a) b ≅ᴰ⟨ ρ a b' (b b') uaᴰρ {a} b b' = - compEquiv + compEquiv (uaᴰ b (ρ _) b') - (substEquiv q PathP i B (q i)) b b') (secEq (ua a a) refl)) + (substEquiv q PathP i B (q i)) b b') (secEq (ua a a) refl)) ρᴰ : {a : A} (b : B a) b ≅ᴰ⟨ ρ a b - ρᴰ {a} b = invEq (uaᴰρ b b) refl + ρᴰ {a} b = invEq (uaᴰρ b b) refl -- total UARel induced by a DUARel @@ -82,7 +82,7 @@ : UARel (Σ A B) (ℓ-max ℓ≅A ℓ≅B) UARel._≅_ (a , b) (a' , b') = Σ[ p a a' ] (b ≅ᴰ⟨ p b') UARel.ua (a , b) (a' , b') = - compEquiv + compEquiv (Σ-cong-equiv (ua a a') p uaᴰ b p b')) ΣPath≃PathΣ diff --git a/docs/Cubical.Displayed.Function.html b/docs/Cubical.Displayed.Function.html index bde1c04..2018770 100644 --- a/docs/Cubical.Displayed.Function.html +++ b/docs/Cubical.Displayed.Function.html @@ -37,8 +37,8 @@ 𝒮-Π : UARel ((a : A) B a) (ℓ-max ℓA ℓ≅B) UARel._≅_ 𝒮-Π f f' = a f a ≅ᴰ⟨ ρ a f' a UARel.ua 𝒮-Π f f' = - compEquiv - (equivΠCod λ a uaᴰρ (f a) (f' a)) + compEquiv + (equivΠCod λ a uaᴰρ (f a) (f' a)) funExtEquiv -- Parameterize UARel by type @@ -46,8 +46,8 @@ _→𝒮_ : (A : Type ℓA) {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) UARel (A B) (ℓ-max ℓA ℓ≅B) (A →𝒮 𝒮-B) .UARel._≅_ f f' = a 𝒮-B .UARel._≅_ (f a) (f' a) (A →𝒮 𝒮-B) .UARel.ua f f' = - compEquiv - (equivΠCod λ a 𝒮-B .UARel.ua (f a) (f' a)) + compEquiv + (equivΠCod λ a 𝒮-B .UARel.ua (f a) (f' a)) funExtEquiv 𝒮-app : {A : Type ℓA} {B : Type ℓB} {𝒮-B : UARel B ℓ≅B} @@ -73,10 +73,10 @@ DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Π f p f' = {b b'} (q : b B.≅ᴰ⟨ p b') f b C.≅ᴰ⟨ p , q f' b' DUARel.uaᴰ 𝒮ᴰ-Π f p f' = - compEquiv - (equivImplicitΠCod λ {b} - (equivImplicitΠCod λ {b'} - (equivΠ (B.uaᴰ b p b') q C.uaᴰ (f b) (p , q) (f' b'))))) + compEquiv + (equivImplicitΠCod λ {b} + (equivImplicitΠCod λ {b'} + (equivΠ (B.uaᴰ b p b') q C.uaᴰ (f b) (p , q) (f' b'))))) funExtDepEquiv _→𝒮ᴰ_ : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} @@ -102,10 +102,10 @@ DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Πˢ f p f' = (b : B _) f b ≅ᴰ⟨ p , refl f' (act p .fst b) DUARel.uaᴰ 𝒮ᴰ-Πˢ f p f' = - compEquiv - (compEquiv - (equivΠCod λ b Jequiv b' q f b ≅ᴰ⟨ p , q f' b')) - (invEquiv implicit≃Explicit)) + compEquiv + (compEquiv + (equivΠCod λ b Jequiv b' q f b ≅ᴰ⟨ p , q f' b')) + (invEquiv implicit≃Explicit)) (DUARel.uaᴰ (𝒮ᴰ-Π (Subst→DUA 𝒮ˢ-B) 𝒮ᴰ-C) f p f') -- SubstRel on a dependent function type @@ -123,15 +123,15 @@ module C = SubstRel 𝒮ˢ-C 𝒮ˢ-Π : SubstRel 𝒮-A a (b : B a) C (a , b)) - 𝒮ˢ-Π .act p = equivΠ' (B.act p) q C.act (p , q)) + 𝒮ˢ-Π .act p = equivΠ' (B.act p) q C.act (p , q)) 𝒮ˢ-Π .uaˢ p f = - fromPathP + fromPathP (DUARel.uaᴰ (𝒮ᴰ-Π (Subst→DUA 𝒮ˢ-B) (Subst→DUA 𝒮ˢ-C)) f p (equivFun (𝒮ˢ-Π .act p) f) .fst {b} - J b' q + J b' q equivFun (C.act (p , q)) (f b) - equivFun (equivΠ' (𝒮ˢ-B .act p) q C.act (p , q))) f b') + equivFun (equivΠ' (𝒮ˢ-B .act p) q C.act (p , q))) f b') i - C.act (p , λ j commSqIsEq (𝒮ˢ-B .act p .snd) b (~ i) j) .fst - (f (retEq (𝒮ˢ-B .act p) b (~ i)))))) + C.act (p , λ j commSqIsEq (𝒮ˢ-B .act p .snd) b (~ i) j) .fst + (f (retEq (𝒮ˢ-B .act p) b (~ i)))))) \ No newline at end of file diff --git a/docs/Cubical.Displayed.Morphism.html b/docs/Cubical.Displayed.Morphism.html index e6a8453..5114c9c 100644 --- a/docs/Cubical.Displayed.Morphism.html +++ b/docs/Cubical.Displayed.Morphism.html @@ -54,7 +54,7 @@ DUARel 𝒮-A (C fun f) ℓ≅C 𝒮ᴰ-reindex f 𝒮ᴰ-C .DUARel._≅ᴰ⟨_⟩_ c p c' = 𝒮ᴰ-C .DUARel._≅ᴰ⟨_⟩_ c (f .rel p) c' 𝒮ᴰ-reindex {C = C} f 𝒮ᴰ-C .DUARel.uaᴰ c p c' = - compEquiv + compEquiv (𝒮ᴰ-C .DUARel.uaᴰ c (f .rel p) c') (substEquiv q PathP i C (q i)) c c') (sym (f .ua p))) diff --git a/docs/Cubical.Displayed.Prop.html b/docs/Cubical.Displayed.Prop.html index 4240431..2be0919 100644 --- a/docs/Cubical.Displayed.Prop.html +++ b/docs/Cubical.Displayed.Prop.html @@ -24,31 +24,31 @@ 𝒮-prop : (P : hProp ℓP) UARel (P .fst) ℓ-zero 𝒮-prop P .UARel._≅_ _ _ = Unit 𝒮-prop P .UARel.ua _ _ = - invEquiv (isContr→≃Unit (isOfHLevelPath' 0 (P .snd) _ _)) + invEquiv (isContr→≃Unit (isOfHLevelPath' 0 (P .snd) _ _)) 𝒮ᴰ-prop : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) (P : A hProp ℓP) DUARel 𝒮-A a P a .fst) ℓ-zero 𝒮ᴰ-prop 𝒮-A P .DUARel._≅ᴰ⟨_⟩_ _ _ _ = Unit 𝒮ᴰ-prop 𝒮-A P .DUARel.uaᴰ _ _ _ = - invEquiv (isContr→≃Unit (isOfHLevelPathP' 0 (P _ .snd) _ _)) + invEquiv (isContr→≃Unit (isOfHLevelPathP' 0 (P _ .snd) _ _)) 𝒮-subtype : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {P : A Type ℓP} - (∀ a isProp (P a)) + (∀ a isProp (P a)) UARel (Σ A P) ℓ≅A 𝒮-subtype 𝒮-A propP .UARel._≅_ (a , _) (a' , _) = 𝒮-A .UARel._≅_ a a' 𝒮-subtype 𝒮-A propP .UARel.ua (a , _) (a' , _) = - compEquiv (𝒮-A .UARel.ua a a') (Σ≡PropEquiv propP) + compEquiv (𝒮-A .UARel.ua a a') (Σ≡PropEquiv propP) 𝒮ᴰ-subtype : {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} {B : A Type ℓB} (𝒮ᴰ-B : DUARel 𝒮-A B ℓ≅B) {P : (a : A) B a Type ℓP} - (∀ a b isProp (P a b)) + (∀ a b isProp (P a b)) DUARel 𝒮-A a Σ[ b B a ] (P a b)) ℓ≅B 𝒮ᴰ-subtype 𝒮ᴰ-B propP .DUARel._≅ᴰ⟨_⟩_ (b , _) p (b' , _) = 𝒮ᴰ-B .DUARel._≅ᴰ⟨_⟩_ b p b' 𝒮ᴰ-subtype 𝒮ᴰ-B propP .DUARel.uaᴰ (b , _) p (b' , _) = - compEquiv + compEquiv (𝒮ᴰ-B .DUARel.uaᴰ b p b') - (compEquiv - (invEquiv (Σ-contractSnd λ _ isOfHLevelPathP' 0 (propP _ b') _ _)) + (compEquiv + (invEquiv (Σ-contractSnd λ _ isOfHLevelPathP' 0 (propP _ b') _ _)) ΣPath≃PathΣ) \ No newline at end of file diff --git a/docs/Cubical.Displayed.Properties.html b/docs/Cubical.Displayed.Properties.html index e1059bf..ed52540 100644 --- a/docs/Cubical.Displayed.Properties.html +++ b/docs/Cubical.Displayed.Properties.html @@ -12,7 +12,7 @@ open import Cubical.Data.Sigma open import Cubical.Relation.Binary -open BinaryRelation +open BinaryRelation open import Cubical.Displayed.Base @@ -31,7 +31,7 @@ (p : a a') P a' (≅→≡ p) 𝒮-J {a} P d {a'} p - = J y q P y q) + = J y q P y q) d (≅→≡ p) @@ -72,7 +72,7 @@ (sym (Iso.rightInv (uaIso a a) refl)) refl uni' : (b' : B a) b ≅ᴰ⟨ ρ a b' PathP i B (≅→≡ (ρ a) i)) b b' - uni' b' = compEquiv (uni b b') (pathToEquiv (g b')) + uni' b' = compEquiv (uni b b') (pathToEquiv (g b')) 𝒮ᴰ-make-1 : (uni : {a : A} (b b' : B a) b ≅ᴰ⟨ ρ a b' (b b')) DUARel 𝒮-A B ℓ≅B @@ -81,19 +81,19 @@ -- constructor that reduces univalence further to contractibility of relational singletons - 𝒮ᴰ-make-2 : (ρᴰ : {a : A} isRefl _≅ᴰ⟨ ρ a ⟩_) - (contrTotal : (a : A) contrRelSingl _≅ᴰ⟨ ρ a ⟩_) + 𝒮ᴰ-make-2 : (ρᴰ : {a : A} isRefl _≅ᴰ⟨ ρ a ⟩_) + (contrTotal : (a : A) contrRelSingl _≅ᴰ⟨ ρ a ⟩_) DUARel 𝒮-A B ℓ≅B - 𝒮ᴰ-make-2 ρᴰ contrTotal = 𝒮ᴰ-make-1 (contrRelSingl→isUnivalent _ ρᴰ (contrTotal _)) + 𝒮ᴰ-make-2 ρᴰ contrTotal = 𝒮ᴰ-make-1 (contrRelSingl→isUnivalent _ ρᴰ (contrTotal _)) -- relational isomorphisms 𝒮-iso→iso : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) {B : Type ℓB} (𝒮-B : UARel B ℓ≅B) - (F : RelIso (UARel._≅_ 𝒮-A) (UARel._≅_ 𝒮-B)) + (F : RelIso (UARel._≅_ 𝒮-A) (UARel._≅_ 𝒮-B)) Iso A B 𝒮-iso→iso 𝒮-A 𝒮-B F - = RelIso→Iso (UARel._≅_ 𝒮-A) + = RelIso→Iso (UARel._≅_ 𝒮-A) (UARel._≅_ 𝒮-B) (UARel.≅→≡ 𝒮-A) (UARel.≅→≡ 𝒮-B) @@ -123,14 +123,14 @@ -- the following can of course be done slightly more generally -- for fiberwise binary relations - F*fiberRelB' : (a : A) Rel (B' (f a)) (B' (f a)) ℓ≅B' + F*fiberRelB' : (a : A) Rel (B' (f a)) (B' (f a)) ℓ≅B' F*fiberRelB' a = fiberRelB' (f a) - module _ (G : (a : A) RelIso (fiberRelB a) (F*fiberRelB' a)) where + module _ (G : (a : A) RelIso (fiberRelB a) (F*fiberRelB' a)) where private fiberIsoOver : (a : A) Iso (B a) (B' (f a)) fiberIsoOver a - = RelIso→Iso (fiberRelB a) + = RelIso→Iso (fiberRelB a) (F*fiberRelB' a) (equivFun (uaᴰρB _ _)) (equivFun (uaᴰρB' _ _)) diff --git a/docs/Cubical.Displayed.Record.html b/docs/Cubical.Displayed.Record.html index 60d6a1b..292938b 100644 --- a/docs/Cubical.Displayed.Record.html +++ b/docs/Cubical.Displayed.Record.html @@ -91,7 +91,7 @@ DUAFields 𝒮-A R _≅R⟨_⟩_ πS 𝒮ᴰ-S πS≅ {ℓF} {F : (a : A) S a Type ℓF} (πF : {a} (r : R a) F a (πS r)) - (propF : a s isProp (F a s)) + (propF : a s isProp (F a s)) DUAFields 𝒮-A R _≅R⟨_⟩_ r πS r , πF r) (𝒮ᴰ-subtype 𝒮ᴰ-S propF) p πS≅ p) module _ {ℓA ℓ≅A} {A : Type ℓA} {𝒮-A : UARel A ℓ≅A} @@ -115,7 +115,7 @@ (compIso (e≅ _ _ r p r') (compIso - (equivToIso (uaᴰ (e _ .Iso.fun r) p (e _ .Iso.fun r'))) + (equivToIso (uaᴰ (e _ .Iso.fun r) p (e _ .Iso.fun r'))) (invIso (congPathIso λ i isoToEquiv (e _))))) module DisplayedRecordMacro where diff --git a/docs/Cubical.Displayed.Sigma.html b/docs/Cubical.Displayed.Sigma.html index 285a520..4fc99e5 100644 --- a/docs/Cubical.Displayed.Sigma.html +++ b/docs/Cubical.Displayed.Sigma.html @@ -68,7 +68,7 @@ DUARel._≅ᴰ⟨_⟩_ 𝒮ᴰ-Σ (b , c) p (b' , c') = Σ[ q b B.≅ᴰ⟨ p b' ] (c C.≅ᴰ⟨ p , q c') DUARel.uaᴰ 𝒮ᴰ-Σ (b , c) p (b' , c') = - compEquiv + compEquiv (Σ-cong-equiv (B.uaᴰ b p b') q C.uaᴰ c (p , q) c')) ΣPath≃PathΣ @@ -98,7 +98,7 @@ 𝒮ˢ-Σ : SubstRel 𝒮-A a Σ[ b B a ] C (a , b)) 𝒮ˢ-Σ .act p = Σ-cong-equiv (B.act p) b C.act (p , refl)) 𝒮ˢ-Σ .uaˢ p _ = - fromPathP + fromPathP (DUARel.uaᴰ (𝒮ᴰ-Σ (Subst→DUA 𝒮ˢ-B) (Subst→DUA 𝒮ˢ-C)) _ p _ .fst (refl , refl)) -- SubstRel on a non-dependent product diff --git a/docs/Cubical.Displayed.Subst.html b/docs/Cubical.Displayed.Subst.html index 3792aa8..467c7ae 100644 --- a/docs/Cubical.Displayed.Subst.html +++ b/docs/Cubical.Displayed.Subst.html @@ -35,15 +35,15 @@ act : {a a' : A} a a' B a B a' uaˢ : {a a' : A} (p : a a') (b : B a) subst B (≅→≡ p) b equivFun (act p) b - uaˢ⁻ : {a a' : A} (p : a a') (b : B a') subst B (sym (≅→≡ p)) b invEq (act p) b + uaˢ⁻ : {a a' : A} (p : a a') (b : B a') subst B (sym (≅→≡ p)) b invEq (act p) b uaˢ⁻ p b = subst B (sym (≅→≡ p)) b - ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (secEq (act p) b)) - subst B (sym (≅→≡ p)) (equivFun (act p) (invEq (act p) b)) - ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (uaˢ p (invEq (act p) b))) - subst B (sym (≅→≡ p)) (subst B (≅→≡ p) (invEq (act p) b)) - ≡⟨ pathToIso (cong B (≅→≡ p)) .Iso.leftInv (invEq (act p) b) - invEq (act p) b + ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (secEq (act p) b)) + subst B (sym (≅→≡ p)) (equivFun (act p) (invEq (act p) b)) + ≡⟨ cong (subst B (sym (≅→≡ p))) (sym (uaˢ p (invEq (act p) b))) + subst B (sym (≅→≡ p)) (subst B (≅→≡ p) (invEq (act p) b)) + ≡⟨ pathToIso (cong B (≅→≡ p)) .Iso.leftInv (invEq (act p) b) + invEq (act p) b Subst→DUA : {A : Type ℓA} {ℓ≅A : Level} {𝒮-A : UARel A ℓ≅A} {B : A Type ℓB} @@ -52,11 +52,11 @@ equivFun (SubstRel.act 𝒮ˢ-B p) b b' DUARel.uaᴰ (Subst→DUA {𝒮-A = 𝒮-A} {B = B} 𝒮ˢ-B) b p b' = equivFun (SubstRel.act 𝒮ˢ-B p) b b' - ≃⟨ invEquiv (compPathlEquiv (sym (SubstRel.uaˢ 𝒮ˢ-B p b))) + ≃⟨ invEquiv (compPathlEquiv (sym (SubstRel.uaˢ 𝒮ˢ-B p b))) subst B (≅→≡ p) b b' - ≃⟨ invEquiv (PathP≃Path i B (≅→≡ p i)) b b') + ≃⟨ invEquiv (PathP≃Path i B (≅→≡ p i)) b b') PathP i B (≅→≡ p i)) b b' - + where open UARel 𝒮-A \ No newline at end of file diff --git a/docs/Cubical.Displayed.Unit.html b/docs/Cubical.Displayed.Unit.html index 7b7e7a0..c9e3dc3 100644 --- a/docs/Cubical.Displayed.Unit.html +++ b/docs/Cubical.Displayed.Unit.html @@ -22,7 +22,7 @@ 𝒮-Unit : UARel Unit ℓ-zero 𝒮-Unit .UARel._≅_ _ _ = Unit -𝒮-Unit .UARel.ua _ _ = invEquiv (isContr→≃Unit (isProp→isContrPath isPropUnit _ _)) +𝒮-Unit .UARel.ua _ _ = invEquiv (isContr→≃Unit (isProp→isContrPath isPropUnit _ _)) 𝒮ᴰ-Unit : {A : Type ℓA} (𝒮-A : UARel A ℓ≅A) DUARel 𝒮-A _ Unit) ℓ-zero 𝒮ᴰ-Unit 𝒮-A = 𝒮ᴰ-const 𝒮-A 𝒮-Unit diff --git a/docs/Cubical.Displayed.Universe.html b/docs/Cubical.Displayed.Universe.html index 6823ab9..5f9eaf6 100644 --- a/docs/Cubical.Displayed.Universe.html +++ b/docs/Cubical.Displayed.Universe.html @@ -30,5 +30,5 @@ 𝒮ᴰ-El : DUARel (𝒮-Univ ) X X) 𝒮ᴰ-El .DUARel._≅ᴰ⟨_⟩_ a e a' = e .fst a a' -𝒮ᴰ-El .DUARel.uaᴰ a e a' = invEquiv (ua-ungluePath-Equiv e) +𝒮ᴰ-El .DUARel.uaᴰ a e a' = invEquiv (ua-ungluePath-Equiv e) \ No newline at end of file diff --git a/docs/Cubical.Foundations.Equiv.Base.html b/docs/Cubical.Foundations.Equiv.Base.html index 24b1497..61fbdb3 100644 --- a/docs/Cubical.Foundations.Equiv.Base.html +++ b/docs/Cubical.Foundations.Equiv.Base.html @@ -29,7 +29,7 @@ -- the definition using Π-type isEquiv' : {}{ℓ'}{A : Type }{B : Type ℓ'} (A B) Type (ℓ-max ℓ') -isEquiv' {B = B} f = (y : B) isContr (fiber f y) +isEquiv' {B = B} f = (y : B) isContr (fiber f y) -- Transport along a line of types A, constant on some extent φ, is an equivalence. isEquivTransp : { : I Level} (A : (i : I) Type ( i)) (φ : I) isEquiv (transp j A (φ j)) φ) diff --git a/docs/Cubical.Foundations.Equiv.Dependent.html b/docs/Cubical.Foundations.Equiv.Dependent.html new file mode 100644 index 0000000..c2a35e0 --- /dev/null +++ b/docs/Cubical.Foundations.Equiv.Dependent.html @@ -0,0 +1,378 @@ + +Cubical.Foundations.Equiv.Dependent
{-
+
+Dependent version of isomorphisms and equivalences
+
+Extremely useful if one wants to construct explicit isomorphisms between record types
+with fields dependent on each other.
+
+This can be generalize in inumerable ways.
+Maybe one day someone will find a common scheme and then computer could automatically generate them.
+
+-}
+{-# OPTIONS --safe #-}
+module Cubical.Foundations.Equiv.Dependent where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Equiv.HalfAdjoint
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Transport
+
+private
+  variable
+     ℓ' ℓ'' ℓ''' : Level
+    A : Type 
+    B : Type ℓ'
+    P : A  Type ℓ''
+    Q : B  Type ℓ'''
+
+
+-- Relative version of maps and their composition
+
+mapOver :
+  (f : A  B)
+  (P : A  Type ℓ'')(Q : B  Type ℓ''')
+   Type _
+mapOver {A = A} f P Q = (a : A)  P a  Q (f a)
+
+compMapOver :
+  {ℓA ℓB ℓC ℓP ℓQ ℓR : Level}
+  {A : Type ℓA}{B : Type ℓB}{C : Type ℓC}
+  {P : A  Type ℓP}{Q : B  Type ℓQ}{R : C  Type ℓR}
+  {f : A  B}{g : B  C}
+   mapOver f P Q  mapOver g Q R
+   mapOver (g  f) P R
+compMapOver f g _ p = g _ (f _ p)
+
+
+-- Fiberwise equivalence
+
+isEquivOver :
+  {f : A  B}
+  (F : mapOver f P Q)
+   Type _
+isEquivOver {A = A} F = (a : A)  isEquiv (F a)
+
+isPropIsEquivOver :
+  {f : A  B}
+  (F : mapOver f P Q)
+   isProp (isEquivOver {Q = Q} F)
+isPropIsEquivOver F = isPropΠ  a  isPropIsEquiv (F a))
+
+-- Relative version of section and retract
+
+sectionOver :
+  {f : A  B}{g : B  A}
+  (sec : section f g)
+  (F : mapOver f P Q)(G : mapOver g Q P)
+   Type _
+sectionOver {B = B} {Q = Q} sec F G =
+  (b : B)(q : Q b)  PathP  i  Q (sec b i)) (F _ (G _ q)) q
+
+retractOver :
+  {f : A  B}{g : B  A}
+  (ret : retract f g)
+  (F : mapOver f P Q)(G : mapOver g Q P)
+   Type _
+retractOver {A = A} {P = P} ret F G =
+  (a : A)(p : P a)  PathP  i  P (ret a i)) (G _ (F _ p)) p
+
+
+-- Relative version of isomorphism
+
+open Iso
+
+record IsoOver { ℓ'} {A : Type }{B : Type ℓ'}
+  (isom : Iso A B)(P : A  Type ℓ'')(Q : B  Type ℓ''')
+  : Type (ℓ-max (ℓ-max  ℓ') (ℓ-max ℓ'' ℓ''')) where
+  no-eta-equality
+  constructor isoover
+  field
+    fun : mapOver (isom .fun) P Q
+    inv : mapOver (isom .inv) Q P
+    rightInv : sectionOver (isom .rightInv) fun inv
+    leftInv  : retractOver (isom .leftInv ) fun inv
+
+record isIsoOver { ℓ'} {A : Type }{B : Type ℓ'}
+  (isom : Iso A B)(P : A  Type ℓ'')(Q : B  Type ℓ''')
+  (fun : mapOver (isom .fun) P Q)
+  : Type (ℓ-max (ℓ-max  ℓ') (ℓ-max ℓ'' ℓ''')) where
+  no-eta-equality
+  constructor isisoover
+  field
+    inv : mapOver (isom .inv) Q P
+    rightInv : sectionOver (isom .rightInv) fun inv
+    leftInv  : retractOver (isom .leftInv ) fun inv
+
+open IsoOver
+open isIsoOver
+
+
+isIsoOver→IsoOver :
+  {isom : Iso A B}
+  {fun : mapOver (isom .fun) P Q}
+   isIsoOver isom P Q fun
+   IsoOver isom P Q
+isIsoOver→IsoOver {fun = fun} isom .fun = fun
+isIsoOver→IsoOver {fun = fun} isom .inv = isom .inv
+isIsoOver→IsoOver {fun = fun} isom .rightInv = isom .rightInv
+isIsoOver→IsoOver {fun = fun} isom .leftInv  = isom .leftInv
+
+IsoOver→isIsoOver :
+  {isom : Iso A B}
+   (isom' : IsoOver isom P Q)
+   isIsoOver isom P Q (isom' .fun)
+IsoOver→isIsoOver isom .inv = isom .inv
+IsoOver→isIsoOver isom .rightInv = isom .rightInv
+IsoOver→isIsoOver isom .leftInv  = isom .leftInv
+
+invIsoOver : {isom : Iso A B}  IsoOver isom P Q  IsoOver (invIso isom) Q P
+invIsoOver {isom = isom} isom' .fun = isom' .inv
+invIsoOver {isom = isom} isom' .inv = isom' .fun
+invIsoOver {isom = isom} isom' .rightInv = isom' .leftInv
+invIsoOver {isom = isom} isom' .leftInv = isom' .rightInv
+
+compIsoOver :
+  {ℓA ℓB ℓC ℓP ℓQ ℓR : Level}
+  {A : Type ℓA}{B : Type ℓB}{C : Type ℓC}
+  {P : A  Type ℓP}{Q : B  Type ℓQ}{R : C  Type ℓR}
+  {isom₁ : Iso A B}{isom₂ : Iso B C}
+   IsoOver isom₁ P Q  IsoOver isom₂ Q R
+   IsoOver (compIso isom₁ isom₂) P R
+compIsoOver {A = A} {B} {C} {P} {Q} {R} {isom₁} {isom₂} isoover₁ isoover₂ = w
+  where
+  w : IsoOver _ _ _
+  w .fun _ = isoover₂ .fun _  isoover₁ .fun _
+  w .inv _ = isoover₁ .inv _  isoover₂ .inv _
+  w .rightInv b q i =
+    comp
+     j  R (compPath-filler (cong (isom₂ .fun) (isom₁ .rightInv _)) (isom₂ .rightInv b) j i))
+     j  λ
+      { (i = i0)  w .fun _ (w .inv _ q)
+      ; (i = i1)  isoover₂ .rightInv _ q j })
+    (isoover₂ .fun _ (isoover₁ .rightInv _ (isoover₂ .inv _ q) i))
+  w .leftInv a p i =
+    comp
+     j  P (compPath-filler (cong (isom₁ .inv) (isom₂ .leftInv _)) (isom₁ .leftInv a) j i))
+     j  λ
+      { (i = i0)  w .inv _ (w .fun _ p)
+      ; (i = i1)  isoover₁ .leftInv _ p j })
+    (isoover₁ .inv _ (isoover₂ .leftInv _ (isoover₁ .fun _ p) i))
+
+
+-- Special cases
+
+fiberIso→IsoOver :
+  {ℓA ℓP ℓQ : Level}
+  {A : Type ℓA}
+  {P : A  Type ℓP}{Q : A  Type ℓQ}
+   ((a : A)  Iso (P a) (Q a))
+   IsoOver idIso P Q
+fiberIso→IsoOver isom .fun a = isom a .fun
+fiberIso→IsoOver isom .inv b = isom b .inv
+fiberIso→IsoOver isom .rightInv b = isom b .rightInv
+fiberIso→IsoOver isom .leftInv  a = isom a .leftInv
+
+-- Only half-adjoint equivalence can be lifted.
+-- This is another clue that HAE is more natural than isomorphism.
+
+open isHAEquiv renaming (g to inv)
+
+pullbackIsoOver :
+  {ℓA ℓB ℓP : Level}
+  {A : Type ℓA}{B : Type ℓB}
+  {P : B  Type ℓP}
+  (f : A  B)
+  (hae : isHAEquiv f)
+   IsoOver (isHAEquiv→Iso hae) (P  f) P
+pullbackIsoOver {A = A} {B} {P} f hae = w
+  where
+  isom = isHAEquiv→Iso hae
+
+  w : IsoOver _ _ _
+  w .fun a = idfun _
+  w .inv b = subst P (sym (isom .rightInv b))
+  w .rightInv b p i = subst-filler P (sym (isom .rightInv b)) p (~ i)
+  w .leftInv  a p i =
+    comp
+     j  P (hae .com a (~ j) i))
+     j  λ
+      { (i = i0)  w .inv _ (w .fun _ p)
+      ; (i = i1)  p })
+    (w .rightInv _ p i)
+
+
+-- Lifting isomorphism of bases to isomorphism of families
+
+-- Since there is no regularity for transport (also no-eta-equality),
+-- we have to fix one field manually to make it invariant under transportation.
+liftHAEToIsoOver :
+  (f : A  B)
+  (hae : isHAEquiv f)
+   ((a : A)  Iso (P a) (Q (f a)))
+   IsoOver (isHAEquiv→Iso hae) P Q
+liftHAEToIsoOver {P = P} {Q = Q} f hae isom =
+  isIsoOver→IsoOver
+    (transport  i  isIsoOver (compIsoIdL (isHAEquiv→Iso hae) i) P Q  a x  isom a .fun x))
+      (IsoOver→isIsoOver (compIsoOver (fiberIso→IsoOver isom) (pullbackIsoOver f hae))))
+
+equivOver→IsoOver :
+  (e : A  B)
+  (f : mapOver (e .fst) P Q)
+   isEquivOver {P = P} {Q = Q} f
+   IsoOver (equivToIso e) P Q
+equivOver→IsoOver {P = P} {Q = Q} e f equiv = w
+  where
+  isom = liftHAEToIsoOver _ (equiv→HAEquiv e .snd)  a  equivToIso (_ , equiv a))
+
+  -- no-eta-equality for Iso, so we have to fill in fields manually
+  w : IsoOver (equivToIso e) P Q
+  w .fun = isom .fun
+  w .inv = isom .inv
+  w .rightInv = isom .rightInv
+  w .leftInv  = isom .leftInv
+
+
+-- Turn isomorphism over HAE into relative equivalence,
+-- i.e. the inverse of the previous precedure.
+
+isoToEquivOver :
+  {A : Type  } {P : A  Type ℓ'' }
+  {B : Type ℓ'} {Q : B  Type ℓ'''}
+  (f : A  B) (hae : isHAEquiv f)
+  (isom' : IsoOver (isHAEquiv→Iso hae) P Q)
+   isEquivOver {Q = Q} (isom' .fun)
+isoToEquivOver {A = A} {P} {Q = Q} f hae isom' a = isoToEquiv (fibiso a) .snd
+  where
+  isom = isHAEquiv→Iso hae
+  finv = isom .inv
+
+  fibiso : (a : A)  Iso (P a) (Q (f a))
+  fibiso a .fun = isom' .fun a
+  fibiso a .inv x = transport  i  P (isom .leftInv a i)) (isom' .inv (f a) x)
+  fibiso a .leftInv  x = fromPathP (isom' .leftInv _ _)
+  fibiso a .rightInv x =
+    sym (substCommSlice _ _ (isom' .fun) _ _)
+     cong  p  subst Q p (isom' .fun _ (isom' .inv _ x))) (hae .com a)
+     fromPathP (isom' .rightInv _ _)
+
+
+-- Half-adjoint equivalence over half-adjoint equivalence
+
+record isHAEquivOver { ℓ'} {A : Type }{B : Type ℓ'}
+  (hae : HAEquiv A B)(P : A  Type ℓ'')(Q : B  Type ℓ''')
+  (fun : mapOver (hae .fst) P Q)
+  : Type (ℓ-max (ℓ-max  ℓ') (ℓ-max ℓ'' ℓ''')) where
+  field
+    inv  : mapOver (hae .snd .inv) Q P
+    rinv : sectionOver (hae .snd .rinv) fun inv
+    linv : retractOver (hae .snd .linv) fun inv
+    com  :  {a} b  SquareP  i j  Q (hae .snd .com a i j))
+       i  fun _ (linv _ b i)) (rinv _ (fun _ b))
+      (refl {x = fun _ (inv _ (fun a b))}) (refl {x = fun a b})
+
+open isHAEquivOver
+
+HAEquivOver : (P : A  Type ℓ'')(Q : B  Type ℓ''')(hae : HAEquiv A B)  Type _
+HAEquivOver P Q hae = Σ[ f  mapOver (hae .fst) P Q ] isHAEquivOver hae P Q f
+
+
+-- forget the coherence square to get an dependent isomorphism
+
+isHAEquivOver→isIsoOver :
+  {hae : HAEquiv A B} (hae' : HAEquivOver P Q hae)
+   IsoOver (isHAEquiv→Iso (hae .snd)) P Q
+isHAEquivOver→isIsoOver hae' .fun = hae' .fst
+isHAEquivOver→isIsoOver hae' .inv = hae' .snd .inv
+isHAEquivOver→isIsoOver hae' .leftInv  = hae' .snd .linv
+isHAEquivOver→isIsoOver hae' .rightInv = hae' .snd .rinv
+
+
+-- A dependent version of `isoToHAEquiv`
+
+IsoOver→HAEquivOver :
+  {isom : Iso A B}
+   (isom' : IsoOver isom P Q)
+   isHAEquivOver (iso→HAEquiv isom) P Q (isom' .fun)
+IsoOver→HAEquivOver {A = A} {P = P} {Q = Q} {isom = isom} isom' = w
+  where
+  f = isom .fun
+  g = isom .inv
+  ε = isom .rightInv
+  η = isom .leftInv
+
+  f' = isom' .fun
+  g' = isom' .inv
+  ε' = isom' .rightInv
+  η' = isom' .leftInv
+
+  sq : _  I  I  _
+  sq b i j =
+    hfill  j  λ
+      { (i = i0)  ε (f (g b)) j
+      ; (i = i1)  ε b j })
+    (inS (f (η (g b) i))) j
+
+  Hfa≡fHa : (f : A  A) (H :  a  f a  a)   a  H (f a)  cong f (H a)
+  Hfa≡fHa f H =
+    J  f p   a  funExt⁻ (sym p) (f a)  cong f (funExt⁻ (sym p) a))
+       a  refl) (sym (funExt H))
+
+  Hfa≡fHaRefl : Hfa≡fHa (idfun _)  _  refl)  λ _  refl
+  Hfa≡fHaRefl =
+    JRefl {x = idfun _}
+       f p   a  funExt⁻ (sym p) (f a)  cong f (funExt⁻ (sym p) a))
+       a  refl)
+
+  Hfa≡fHaOver : (f : A  A) (H :  a  f a  a)
+    (f' : mapOver f P P) (H' :  a b  PathP  i  P (H a i)) (f' _ b) b)
+      a b  SquareP  i j 
+      P (Hfa≡fHa f H a i j)) (H' _ (f' _ b))  i  f' _ (H' _ b i)) refl refl
+  Hfa≡fHaOver f H f' H' =
+   JDep {B = λ f  mapOver f P P}
+     f p f' p'   a b
+       SquareP  i j  P (Hfa≡fHa f (funExt⁻ (sym p)) a i j))
+           i  p' (~ i) _ (f' _ b))  i  f' _ (p' (~ i) _ b))
+          (refl {x = f' _ (f' _ b)}) (refl {x = f' _ b}))
+     a b 
+      transport  t  SquareP  i j  P (Hfa≡fHaRefl (~ t) a i j))
+        (refl {x = b}) refl refl refl) refl)
+    (sym (funExt H))  i a b  H' a b (~ i))
+
+  cube : _  I  I  I  _
+  cube a i j k =
+    hfill  k  λ
+      { (i = i0)  ε (f (η a j)) k
+      ; (j = i0)  ε (f (g (f a))) k
+      ; (j = i1)  ε (f a) k})
+    (inS (f (Hfa≡fHa  x  g (f x)) η a (~ i) j))) k
+
+  w : isHAEquivOver _ _ _ _
+  w .inv  = isom' .inv
+  w .linv = isom' .leftInv
+  w .rinv b x i =
+    comp  j  Q (sq b i j))
+     j  λ
+      { (i = i0)  ε' _ (f' _ (g' _ x)) j
+      ; (i = i1)  ε' _ x j })
+    (f' _ (η' _ (g' _ x) i))
+  w. com {a} b i j =
+    comp  k  Q (cube a i j k))
+     k  λ
+      { (i = i0)  ε' _ (f' _ (η' _ b j)) k
+      ; (j = i0)  ε' _ (f' _ (g' _ (f' _ b))) k
+      ; (j = i1)  ε' _ (f' _ b) k})
+    (f' _ (Hfa≡fHaOver _ η _ η' a b (~ i) j))
+
+
+-- transform an isomorphism over some isomorphism to an isomorphism over its half-adjoint-ization
+
+IsoOverIso→IsoOverHAEquiv :
+  {isom : Iso A B} (isom' : IsoOver isom P Q)
+   IsoOver (isHAEquiv→Iso (iso→HAEquiv isom .snd)) P Q
+IsoOverIso→IsoOverHAEquiv isom' =
+  isHAEquivOver→isIsoOver (_ , IsoOver→HAEquivOver isom')
+
\ No newline at end of file diff --git a/docs/Cubical.Foundations.Equiv.Fiberwise.html b/docs/Cubical.Foundations.Equiv.Fiberwise.html index 77b11da..a11baca 100644 --- a/docs/Cubical.Foundations.Equiv.Fiberwise.html +++ b/docs/Cubical.Foundations.Equiv.Fiberwise.html @@ -27,14 +27,14 @@ fibers-total {xv} = iso h g h-g g-h where h : {xv} fiber total xv fiber (f (xv .fst)) (xv .snd) - h {xv} (p , eq) = J (\ xv eq fiber (f (xv .fst)) (xv .snd)) ((p .snd) , refl) eq + h {xv} (p , eq) = J (\ xv eq fiber (f (xv .fst)) (xv .snd)) ((p .snd) , refl) eq g : {xv} fiber (f (xv .fst)) (xv .snd) fiber total xv g {xv} (p , eq) = (xv .fst , p) , (\ i _ , eq i) h-g : {xv} y h {xv} (g {xv} y) y - h-g {x , v} (p , eq) = J _ eq₁ h (g (p , eq₁)) (p , eq₁)) (JRefl xv₁ eq₁ fiber (f (xv₁ .fst)) (xv₁ .snd)) ((p , refl))) (eq) + h-g {x , v} (p , eq) = J _ eq₁ h (g (p , eq₁)) (p , eq₁)) (JRefl xv₁ eq₁ fiber (f (xv₁ .fst)) (xv₁ .snd)) ((p , refl))) (eq) g-h : {xv} y g {xv} (h {xv} y) y - g-h {xv} ((a , p) , eq) = J _ eq₁ g (h ((a , p) , eq₁)) ((a , p) , eq₁)) - (cong g (JRefl xv₁ eq₁ fiber (f (xv₁ .fst)) (xv₁ .snd)) (p , refl))) + g-h {xv} ((a , p) , eq) = J _ eq₁ g (h ((a , p) , eq₁)) ((a , p) , eq₁)) + (cong g (JRefl xv₁ eq₁ fiber (f (xv₁ .fst)) (xv₁ .snd)) (p , refl))) eq -- Thm 4.7.7 (fiberwise equivalences) fiberEquiv : ([tf] : isEquiv total) @@ -57,7 +57,7 @@ isContrToUniv {A} {B} = fiberEquiv z A z) z A ~ z) (\ B idTo~ {A} {B}) { .equiv-proof y - isContrΣ (isContrSingl _) + isContrΣ (isContrSingl _) \ a isContr→isContrPath (c A) _ _ }) B @@ -69,33 +69,33 @@ -} recognizeId : {A : Type } {a : A} (Eq : A Type ℓ') Eq a - isContr (Σ _ Eq) + isContr (Σ _ Eq) (x : A) (a x) (Eq x) recognizeId {A = A} {a = a} Eq eqRefl eqContr x = (fiberMap x) , (isEquivFiberMap x) where fiberMap : (x : A) a x Eq x - fiberMap x = J x p Eq x) eqRefl + fiberMap x = J x p Eq x) eqRefl mapOnSigma : Σ[ x A ] a x Σ _ Eq mapOnSigma pair = fst pair , fiberMap (fst pair) (snd pair) equivOnSigma : (x : A) isEquiv mapOnSigma - equivOnSigma x = isEquivFromIsContr mapOnSigma (isContrSingl a) eqContr + equivOnSigma x = isEquivFromIsContr mapOnSigma (isContrSingl a) eqContr isEquivFiberMap : (x : A) isEquiv (fiberMap x) isEquivFiberMap = fiberEquiv x a x) Eq fiberMap (equivOnSigma x) fundamentalTheoremOfId : {A : Type } (Eq : A A Type ℓ') ((x : A) Eq x x) - ((x : A) isContr (Σ[ y A ] Eq x y)) + ((x : A) isContr (Σ[ y A ] Eq x y)) (x y : A) (x y) (Eq x y) fundamentalTheoremOfId Eq eqRefl eqContr x = recognizeId (Eq x) (eqRefl x) (eqContr x) fundamentalTheoremOfIdβ : {A : Type } (Eq : A A Type ℓ') (eqRefl : (x : A) Eq x x) - (eqContr : (x : A) isContr (Σ[ y A ] Eq x y)) + (eqContr : (x : A) isContr (Σ[ y A ] Eq x y)) (x : A) fst (fundamentalTheoremOfId Eq eqRefl eqContr x x) refl eqRefl x -fundamentalTheoremOfIdβ Eq eqRefl eqContr x = JRefl y p Eq x y) (eqRefl x) +fundamentalTheoremOfIdβ Eq eqRefl eqContr x = JRefl y p Eq x y) (eqRefl x) \ No newline at end of file diff --git a/docs/Cubical.Foundations.Equiv.HalfAdjoint.html b/docs/Cubical.Foundations.Equiv.HalfAdjoint.html index 42585ec..34279a4 100644 --- a/docs/Cubical.Foundations.Equiv.HalfAdjoint.html +++ b/docs/Cubical.Foundations.Equiv.HalfAdjoint.html @@ -52,10 +52,10 @@ gy≡x : g y x gy≡x = sym (cong g p) ∙∙ refl ∙∙ linv x - lem0 : Square (cong f (linv x)) p (cong f (linv x)) p + lem0 : Square (cong f (linv x)) p (cong f (linv x)) p lem0 i j = invSides-filler p (sym (cong f (linv x))) (~ i) j - ry≡p : Square (rinv y) p (cong f gy≡x) refl + ry≡p : Square (rinv y) p (cong f gy≡x) refl ry≡p i j = hcomp k λ { (i = i0) cong rinv p k j ; (i = i1) lem0 k j ; (j = i0) f (doubleCompPath-filler (sym (cong g p)) refl (linv x) k i) @@ -77,7 +77,7 @@ η = Iso.leftInv e Hfa≡fHa : (f : A A) (H : a f a a) a H (f a) cong f (H a) - Hfa≡fHa f H = J f p a funExt⁻ (sym p) (f a) cong f (funExt⁻ (sym p) a)) + Hfa≡fHa f H = J f p a funExt⁻ (sym p) (f a) cong f (funExt⁻ (sym p) a)) a refl) (sym (funExt H)) @@ -96,10 +96,10 @@ equiv→HAEquiv : A B HAEquiv A B equiv→HAEquiv e = e .fst , λ where - .isHAEquiv.g invIsEq (snd e) - .isHAEquiv.linv retIsEq (snd e) - .isHAEquiv.rinv secIsEq (snd e) - .isHAEquiv.com a sym (commPathIsEq (snd e) a) + .isHAEquiv.g invIsEq (snd e) + .isHAEquiv.linv retIsEq (snd e) + .isHAEquiv.rinv secIsEq (snd e) + .isHAEquiv.com a sym (commPathIsEq (snd e) a) congIso : {x y : A} (e : Iso A B) Iso (x y) (Iso.fun e x Iso.fun e y) congIso {x = x} {y} e = goal @@ -130,7 +130,7 @@ helper : {x : A} {y : B} (f : A B) (r : f x y) (p q : x x) (sym r ∙∙ cong f (p q) ∙∙ r) (sym r ∙∙ cong f p ∙∙ r) (sym r ∙∙ cong f q ∙∙ r) helper {x = x} f = - J y r (p q : x x) + J y r (p q : x x) (sym r ∙∙ cong f (p q) ∙∙ r) (sym r ∙∙ cong f p ∙∙ r) (sym r ∙∙ cong f q ∙∙ r)) λ p q i rUnit (congFunct f p q i) (~ i)) λ i rUnit (cong f p) i rUnit (cong f q) i diff --git a/docs/Cubical.Foundations.Equiv.Properties.html b/docs/Cubical.Foundations.Equiv.Properties.html index 8225c96..cd3e02e 100644 --- a/docs/Cubical.Foundations.Equiv.Properties.html +++ b/docs/Cubical.Foundations.Equiv.Properties.html @@ -33,38 +33,38 @@ ℓ' ℓ'' : Level A B C : Type -isEquivInvEquiv : isEquiv (e : A B) invEquiv e) +isEquivInvEquiv : isEquiv (e : A B) invEquiv e) isEquivInvEquiv = isoToIsEquiv goal where open Iso goal : Iso (A B) (B A) - goal .fun = invEquiv - goal .inv = invEquiv - goal .rightInv g = equivEq refl - goal .leftInv f = equivEq refl + goal .fun = invEquiv + goal .inv = invEquiv + goal .rightInv g = equivEq refl + goal .leftInv f = equivEq refl invEquivEquiv : (A B) (B A) invEquivEquiv = _ , isEquivInvEquiv isEquivCong : {x y : A} (e : A B) isEquiv (p : x y) cong (equivFun e) p) -isEquivCong e = isoToIsEquiv (congIso (equivToIso e)) +isEquivCong e = isoToIsEquiv (congIso (equivToIso e)) congEquiv : {x y : A} (e : A B) (x y) (equivFun e x equivFun e y) -congEquiv e = isoToEquiv (congIso (equivToIso e)) +congEquiv e = isoToEquiv (congIso (equivToIso e)) -equivAdjointEquiv : (e : A B) {a b} (a invEq e b) (equivFun e a b) -equivAdjointEquiv e = compEquiv (congEquiv e) (compPathrEquiv (secEq e _)) +equivAdjointEquiv : (e : A B) {a b} (a invEq e b) (equivFun e a b) +equivAdjointEquiv e = compEquiv (congEquiv e) (compPathrEquiv (secEq e _)) -invEq≡→equivFun≡ : (e : A B) {a b} invEq e b a equivFun e a b +invEq≡→equivFun≡ : (e : A B) {a b} invEq e b a equivFun e a b invEq≡→equivFun≡ e = equivFun (equivAdjointEquiv e) sym isEquivPreComp : (e : A B) isEquiv (φ : B C) φ equivFun e) -isEquivPreComp e = snd (equiv→ (invEquiv e) (idEquiv _)) +isEquivPreComp e = snd (equiv→ (invEquiv e) (idEquiv _)) preCompEquiv : (e : A B) (B C) (A C) preCompEquiv e = φ φ fst e) , isEquivPreComp e isEquivPostComp : (e : A B) isEquiv (φ : C A) e .fst φ) -isEquivPostComp e = snd (equivΠCod _ e)) +isEquivPostComp e = snd (equivΠCod _ e)) postCompEquiv : (e : A B) (C A) (C B) postCompEquiv e = _ , isEquivPostComp e @@ -77,43 +77,43 @@ hasRetract : (A B) Type _ hasRetract {A = A} {B = B} f = Σ[ g (B A) ] retract f g -isEquiv→isContrHasSection : {f : A B} isEquiv f isContr (hasSection f) -fst (isEquiv→isContrHasSection isEq) = invIsEq isEq , secIsEq isEq +isEquiv→isContrHasSection : {f : A B} isEquiv f isContr (hasSection f) +fst (isEquiv→isContrHasSection isEq) = invIsEq isEq , secIsEq isEq snd (isEquiv→isContrHasSection isEq) (f , ε) i = b fst (p b i)) , b snd (p b i)) - where p : b (invIsEq isEq b , secIsEq isEq b) (f b , ε b) + where p : b (invIsEq isEq b , secIsEq isEq b) (f b , ε b) p b = isEq .equiv-proof b .snd (f b , ε b) isEquiv→hasSection : {f : A B} isEquiv f hasSection f isEquiv→hasSection = fst isEquiv→isContrHasSection -isContr-hasSection : (e : A B) isContr (hasSection (fst e)) +isContr-hasSection : (e : A B) isContr (hasSection (fst e)) isContr-hasSection e = isEquiv→isContrHasSection (snd e) -isEquiv→isContrHasRetract : {f : A B} isEquiv f isContr (hasRetract f) -fst (isEquiv→isContrHasRetract isEq) = invIsEq isEq , retIsEq isEq +isEquiv→isContrHasRetract : {f : A B} isEquiv f isContr (hasRetract f) +fst (isEquiv→isContrHasRetract isEq) = invIsEq isEq , retIsEq isEq snd (isEquiv→isContrHasRetract {f = f} isEq) (g , η) = λ i b p b i) , a q a i) - where p : b invIsEq isEq b g b - p b = sym (η (invIsEq isEq b)) ∙' cong g (secIsEq isEq b) + where p : b invIsEq isEq b g b + p b = sym (η (invIsEq isEq b)) ∙' cong g (secIsEq isEq b) -- one square from the definition of invIsEq - ieSq : a Square (cong g (secIsEq isEq (f a))) + ieSq : a Square (cong g (secIsEq isEq (f a))) refl - (cong (g f) (retIsEq isEq a)) + (cong (g f) (retIsEq isEq a)) refl - ieSq a k j = g (commSqIsEq isEq a k j) + ieSq a k j = g (commSqIsEq isEq a k j) -- one square from η - ηSq : a Square (η (invIsEq isEq (f a))) + ηSq : a Square (η (invIsEq isEq (f a))) (η a) - (cong (g f) (retIsEq isEq a)) - (retIsEq isEq a) - ηSq a i j = η (retIsEq isEq a i) j + (cong (g f) (retIsEq isEq a)) + (retIsEq isEq a) + ηSq a i j = η (retIsEq isEq a i) j -- and one last square from the definition of p - pSq : b Square (η (invIsEq isEq b)) + pSq : b Square (η (invIsEq isEq b)) refl - (cong g (secIsEq isEq b)) + (cong g (secIsEq isEq b)) (p b) - pSq b i j = compPath'-filler (sym (η (invIsEq isEq b))) (cong g (secIsEq isEq b)) j i - q : a Square (retIsEq isEq a) (η a) (p (f a)) refl + pSq b i j = compPath'-filler (sym (η (invIsEq isEq b))) (cong g (secIsEq isEq b)) j i + q : a Square (retIsEq isEq a) (η a) (p (f a)) refl q a i j = hcomp k λ { (i = i0) ηSq a j k ; (i = i1) η a (j k) ; (j = i0) pSq (f a) i k @@ -124,12 +124,12 @@ isEquiv→hasRetract : {f : A B} isEquiv f hasRetract f isEquiv→hasRetract = fst isEquiv→isContrHasRetract -isContr-hasRetract : (e : A B) isContr (hasRetract (fst e)) +isContr-hasRetract : (e : A B) isContr (hasRetract (fst e)) isContr-hasRetract e = isEquiv→isContrHasRetract (snd e) isEquiv→retractIsEquiv : {f : A B} {g : B A} isEquiv f retract f g isEquiv g isEquiv→retractIsEquiv {f = f} {g = g} isEquiv-f retract-g = subst isEquiv f⁻¹≡g (snd f⁻¹) - where f⁻¹ = invEquiv (f , isEquiv-f) + where f⁻¹ = invEquiv (f , isEquiv-f) retract-f⁻¹ : retract f (fst f⁻¹) retract-f⁻¹ = snd (isEquiv→hasRetract isEquiv-f) @@ -137,14 +137,14 @@ f⁻¹≡g : fst f⁻¹ g f⁻¹≡g = cong fst - (isContr→isProp (isEquiv→isContrHasRetract isEquiv-f) + (isContr→isProp (isEquiv→isContrHasRetract isEquiv-f) (fst f⁻¹ , retract-f⁻¹) (g , retract-g)) isEquiv→sectionIsEquiv : {f : A B} {g : B A} isEquiv f section f g isEquiv g isEquiv→sectionIsEquiv {f = f} {g = g} isEquiv-f section-g = subst isEquiv f⁻¹≡g (snd f⁻¹) - where f⁻¹ = invEquiv (f , isEquiv-f) + where f⁻¹ = invEquiv (f , isEquiv-f) section-f⁻¹ : section f (fst f⁻¹) section-f⁻¹ = snd (isEquiv→hasSection isEquiv-f) @@ -152,7 +152,7 @@ f⁻¹≡g : fst f⁻¹ g f⁻¹≡g = cong fst - (isContr→isProp (isEquiv→isContrHasSection isEquiv-f) + (isContr→isProp (isEquiv→isContrHasSection isEquiv-f) (fst f⁻¹ , section-f⁻¹) (g , section-g)) @@ -167,7 +167,7 @@ pathToEquiv refl ≡⟨ pathToEquivRefl idEquiv (F A) -isPropIsHAEquiv : {f : A B} isProp (isHAEquiv f) +isPropIsHAEquiv : {f : A B} isProp (isHAEquiv f) isPropIsHAEquiv {f = f} ishaef = goal ishaef where equivF : isEquiv f equivF = isHAEquiv→isEquiv ishaef @@ -176,10 +176,10 @@ rCoh1 (g , ε) = Σ[ η retract f g ] x cong f (η x) ε (f x) rCoh2 : (sec : hasSection f) Type _ - rCoh2 (g , ε) = Σ[ η retract f g ] x Square (ε (f x)) refl (cong f (η x)) refl + rCoh2 (g , ε) = Σ[ η retract f g ] x Square (ε (f x)) refl (cong f (η x)) refl rCoh3 : (sec : hasSection f) Type _ - rCoh3 (g , ε) = x Σ[ ηx g (f x) x ] Square (ε (f x)) refl (cong f ηx) refl + rCoh3 (g , ε) = x Σ[ ηx g (f x) x ] Square (ε (f x)) refl (cong f ηx) refl rCoh4 : (sec : hasSection f) Type _ rCoh4 (g , ε) = x Path (fiber f (f x)) (g (f x) , ε (f x)) (x , refl) @@ -188,31 +188,31 @@ characterization = isHAEquiv f -- first convert between Σ and record - ≃⟨ isoToEquiv (iso e (e .g , e .rinv) , (e .linv , e .com)) + ≃⟨ isoToEquiv (iso e (e .g , e .rinv) , (e .linv , e .com)) e record { g = e .fst .fst ; rinv = e .fst .snd ; linv = e .snd .fst ; com = e .snd .snd }) - _ refl) λ _ refl) + _ refl) λ _ refl) Σ _ rCoh1 -- secondly, convert the path into a dependent path for later convenience - ≃⟨ Σ-cong-equiv-snd s Σ-cong-equiv-snd - λ η equivΠCod - λ x compEquiv (flipSquareEquiv {a₀₀ = f x}) (invEquiv slideSquareEquiv)) + ≃⟨ Σ-cong-equiv-snd s Σ-cong-equiv-snd + λ η equivΠCod + λ x compEquiv (flipSquareEquiv {a₀₀ = f x}) (invEquiv slideSquareEquiv)) Σ _ rCoh2 - ≃⟨ Σ-cong-equiv-snd s invEquiv Σ-Π-≃) + ≃⟨ Σ-cong-equiv-snd s invEquiv Σ-Π-≃) Σ _ rCoh3 - ≃⟨ Σ-cong-equiv-snd s equivΠCod λ x ΣPath≃PathΣ) + ≃⟨ Σ-cong-equiv-snd s equivΠCod λ x ΣPath≃PathΣ) Σ _ rCoh4 - + where open isHAEquiv - goal : isProp (isHAEquiv f) - goal = subst isProp (sym (ua characterization)) - (isPropΣ (isContr→isProp (isEquiv→isContrHasSection equivF)) - λ s isPropΠ λ x isProp→isSet (isContr→isProp (equivF .equiv-proof (f x))) _ _) + goal : isProp (isHAEquiv f) + goal = subst isProp (sym (ua characterization)) + (isPropΣ (isContr→isProp (isEquiv→isContrHasSection equivF)) + λ s isPropΠ λ x isProp→isSet (isContr→isProp (equivF .equiv-proof (f x))) _ _) -- loop spaces connected by a path are equivalent conjugatePathEquiv : {A : Type } {a b : A} (p : a b) (a a) (b b) -conjugatePathEquiv p = compEquiv (compPathrEquiv p) (compPathlEquiv (sym p)) +conjugatePathEquiv p = compEquiv (compPathrEquiv p) (compPathlEquiv (sym p)) -- composition on the right induces an equivalence of path types compr≡Equiv : {A : Type } {a b c : A} (p q : a b) (r : b c) (p q) (p r q r) @@ -223,37 +223,37 @@ compl≡Equiv p q r = congEquiv ((λ s p s) , (compPathl-isEquiv p)) isEquivFromIsContr : {A : Type } {B : Type ℓ'} - (f : A B) isContr A isContr B + (f : A B) isContr A isContr B isEquiv f isEquivFromIsContr f isContrA isContrB = - subst isEquiv i x isContr→isProp isContrB (fst B≃A x) (f x) i) (snd B≃A) - where B≃A = isContr→Equiv isContrA isContrB + subst isEquiv i x isContr→isProp isContrB (fst B≃A x) (f x) i) (snd B≃A) + where B≃A = isContr→Equiv isContrA isContrB isEquiv[f∘equivFunA≃B]→isEquiv[f] : {A : Type } {B : Type ℓ'} {C : Type ℓ''} (f : B C) (A≃B : A B) isEquiv (f equivFun A≃B) isEquiv f isEquiv[f∘equivFunA≃B]→isEquiv[f] f (g , gIsEquiv) f∘gIsEquiv = - precomposesToId→Equiv f _ w w' + precomposesToId→Equiv f _ w w' where - w : f g equivFun (invEquiv (_ , f∘gIsEquiv)) idfun _ - w = (cong fst (invEquiv-is-linv (_ , f∘gIsEquiv))) + w : f g equivFun (invEquiv (_ , f∘gIsEquiv)) idfun _ + w = (cong fst (invEquiv-is-linv (_ , f∘gIsEquiv))) - w' : isEquiv (g equivFun (invEquiv (_ , f∘gIsEquiv))) - w' = (snd (compEquiv (invEquiv (_ , f∘gIsEquiv) ) (_ , gIsEquiv))) + w' : isEquiv (g equivFun (invEquiv (_ , f∘gIsEquiv))) + w' = (snd (compEquiv (invEquiv (_ , f∘gIsEquiv) ) (_ , gIsEquiv))) isEquiv[equivFunA≃B∘f]→isEquiv[f] : {A : Type } {B : Type ℓ'} {C : Type ℓ''} (f : C A) (A≃B : A B) isEquiv (equivFun A≃B f) isEquiv f isEquiv[equivFunA≃B∘f]→isEquiv[f] f (g , gIsEquiv) g∘fIsEquiv = - composesToId→Equiv _ f w w' + composesToId→Equiv _ f w w' where - w : equivFun (invEquiv (_ , g∘fIsEquiv)) g f idfun _ - w = (cong fst (invEquiv-is-rinv (_ , g∘fIsEquiv))) + w : equivFun (invEquiv (_ , g∘fIsEquiv)) g f idfun _ + w = (cong fst (invEquiv-is-rinv (_ , g∘fIsEquiv))) - w' : isEquiv (equivFun (invEquiv (_ , g∘fIsEquiv)) g) - w' = snd (compEquiv (_ , gIsEquiv) (invEquiv (_ , g∘fIsEquiv))) + w' : isEquiv (equivFun (invEquiv (_ , g∘fIsEquiv)) g) + w' = snd (compEquiv (_ , gIsEquiv) (invEquiv (_ , g∘fIsEquiv))) isPointedTarget→isEquiv→isEquiv : {A B : Type } (f : A B) (B isEquiv f) isEquiv f diff --git a/docs/Cubical.Foundations.Equiv.html b/docs/Cubical.Foundations.Equiv.html index 37a371d..482462e 100644 --- a/docs/Cubical.Foundations.Equiv.html +++ b/docs/Cubical.Foundations.Equiv.html @@ -30,7 +30,7 @@ ℓ' ℓ'' : Level A B C D : Type -infixr 30 _∙ₑ_ +infixr 30 _∙ₑ_ equivIsEquiv : (e : A B) isEquiv (equivFun e) equivIsEquiv e = snd e @@ -45,15 +45,15 @@ -- Proof using isPropIsContr. This is slow and the direct proof below is better -isPropIsEquiv' : (f : A B) isProp (isEquiv f) +isPropIsEquiv' : (f : A B) isProp (isEquiv f) equiv-proof (isPropIsEquiv' f u0 u1 i) y = - isPropIsContr (u0 .equiv-proof y) (u1 .equiv-proof y) i + isPropIsContr (u0 .equiv-proof y) (u1 .equiv-proof y) i -- Direct proof that computes quite ok (can be optimized further if -- necessary, see: -- https://github.com/mortberg/cubicaltt/blob/pi4s3_dimclosures/examples/brunerie2.ctt#L562 -isPropIsEquiv : (f : A B) isProp (isEquiv f) +isPropIsEquiv : (f : A B) isProp (isEquiv f) equiv-proof (isPropIsEquiv f p q i) y = let p2 = p .equiv-proof y .snd q2 = q .equiv-proof y .snd @@ -64,260 +64,265 @@ ; (j = i1) w }) (p2 w (i j)) -equivEq : {e f : A B} (h : e .fst f .fst) e f -equivEq {e = e} {f = f} h = λ i (h i) , isProp→PathP i isPropIsEquiv (h i)) (e .snd) (f .snd) i - -module _ {f : A B} (equivF : isEquiv f) where - funIsEq : A B - funIsEq = f - - invIsEq : B A - invIsEq y = equivF .equiv-proof y .fst .fst - - secIsEq : section f invIsEq - secIsEq y = equivF .equiv-proof y .fst .snd - - retIsEq : retract f invIsEq - retIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .fst - - commSqIsEq : a Square (secIsEq (f a)) refl (cong f (retIsEq a)) refl - commSqIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .snd - - commPathIsEq : a secIsEq (f a) cong f (retIsEq a) - commPathIsEq a i j = - hcomp - k λ - { (i = i0) secIsEq (f a) j - ; (i = i1) f (retIsEq a (j ~ k)) - ; (j = i0) f (retIsEq a (i ~ k)) - ; (j = i1) f a - }) - (commSqIsEq a i j) - -module _ (w : A B) where - invEq : B A - invEq = invIsEq (snd w) - - retEq : retract (w .fst) invEq - retEq = retIsEq (snd w) - - secEq : section (w .fst) invEq - secEq = secIsEq (snd w) - -open Iso - -equivToIso : { ℓ'} {A : Type } {B : Type ℓ'} A B Iso A B -fun (equivToIso e) = e .fst -inv (equivToIso e) = invEq e -rightInv (equivToIso e) = secEq e -leftInv (equivToIso e) = retEq e - --- TODO: there should be a direct proof of this that doesn't use equivToIso -invEquiv : A B B A -invEquiv e = isoToEquiv (invIso (equivToIso e)) - -invEquivIdEquiv : (A : Type ) invEquiv (idEquiv A) idEquiv A -invEquivIdEquiv _ = equivEq refl - -compEquiv : A B B C A C -compEquiv f g .fst = g .fst f .fst -compEquiv {A = A} {C = C} f g .snd .equiv-proof c = contr - where - contractG = g .snd .equiv-proof c .snd - - secFiller : (a : A) (p : g .fst (f .fst a) c) _ {- square in A -} - secFiller a p = compPath-filler (cong (invEq f fst) (contractG (_ , p))) (retEq f a) - - contr : isContr (fiber (g .fst f .fst) c) - contr .fst .fst = invEq f (invEq g c) - contr .fst .snd = cong (g .fst) (secEq f (invEq g c)) secEq g c - contr .snd (a , p) i .fst = secFiller a p i1 i - contr .snd (a , p) i .snd j = - hcomp - k λ - { (i = i1) fSquare k - ; (j = i0) g .fst (f .fst (secFiller a p k i)) - ; (j = i1) contractG (_ , p) i .snd k - }) - (g .fst (secEq f (contractG (_ , p) i .fst) j)) - where - fSquare : I C - fSquare k = - hcomp - l λ - { (j = i0) g .fst (f .fst (retEq f a k)) - ; (j = i1) p (k l) - ; (k = i0) g .fst (secEq f (f .fst a) j) - ; (k = i1) p (j l) - }) - (g .fst (f .snd .equiv-proof (f .fst a) .snd (a , refl) k .snd j)) - -_∙ₑ_ = compEquiv - -compEquivIdEquiv : (e : A B) compEquiv (idEquiv A) e e -compEquivIdEquiv e = equivEq refl - -compEquivEquivId : (e : A B) compEquiv e (idEquiv B) e -compEquivEquivId e = equivEq refl - -invEquiv-is-rinv : (e : A B) compEquiv e (invEquiv e) idEquiv A -invEquiv-is-rinv e = equivEq (funExt (retEq e)) - -invEquiv-is-linv : (e : A B) compEquiv (invEquiv e) e idEquiv B -invEquiv-is-linv e = equivEq (funExt (secEq e)) - -compEquiv-assoc : (f : A B) (g : B C) (h : C D) - compEquiv f (compEquiv g h) compEquiv (compEquiv f g) h -compEquiv-assoc f g h = equivEq refl - -LiftEquiv : A Lift {i = } {j = ℓ'} A -LiftEquiv .fst a .lower = a -LiftEquiv .snd .equiv-proof = strictContrFibers lower - -Lift≃Lift : (e : A B) Lift {j = ℓ'} A Lift {j = ℓ''} B -Lift≃Lift e .fst a .lower = e .fst (a .lower) -Lift≃Lift e .snd .equiv-proof b .fst .fst .lower = invEq e (b .lower) -Lift≃Lift e .snd .equiv-proof b .fst .snd i .lower = - e .snd .equiv-proof (b .lower) .fst .snd i -Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .fst .lower = - e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .fst -Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .snd j .lower = - e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .snd j - -isContr→Equiv : isContr A isContr B A B -isContr→Equiv Actr Bctr = isoToEquiv (isContr→Iso Actr Bctr) - -propBiimpl→Equiv : (Aprop : isProp A) (Bprop : isProp B) (f : A B) (g : B A) A B -propBiimpl→Equiv Aprop Bprop f g = f , hf - where - hf : isEquiv f - hf .equiv-proof y .fst = (g y , Bprop (f (g y)) y) - hf .equiv-proof y .snd h i .fst = Aprop (g y) (h .fst) i - hf .equiv-proof y .snd h i .snd = isProp→isSet' Bprop (Bprop (f (g y)) y) (h .snd) - (cong f (Aprop (g y) (h .fst))) refl i - -isEquivPropBiimpl→Equiv : isProp A isProp B - ((A B) × (B A)) (A B) -isEquivPropBiimpl→Equiv {A = A} {B = B} Aprop Bprop = isoToEquiv isom where - isom : Iso (Σ (A B) _ B A)) (A B) - isom .fun (f , g) = propBiimpl→Equiv Aprop Bprop f g - isom .inv e = equivFun e , invEq e - isom .rightInv e = equivEq refl - isom .leftInv _ = refl - -equivΠCod : {F : A Type } {G : A Type ℓ'} - ((x : A) F x G x) ((x : A) F x) ((x : A) G x) -equivΠCod k .fst f x = k x .fst (f x) -equivΠCod k .snd .equiv-proof f .fst .fst x = equivCtr (k x) (f x) .fst -equivΠCod k .snd .equiv-proof f .fst .snd i x = equivCtr (k x) (f x) .snd i -equivΠCod k .snd .equiv-proof f .snd (g , p) i .fst x = - equivCtrPath (k x) (f x) (g x , λ j p j x) i .fst -equivΠCod k .snd .equiv-proof f .snd (g , p) i .snd j x = - equivCtrPath (k x) (f x) (g x , λ k p k x) i .snd j - -equivImplicitΠCod : {F : A Type } {G : A Type ℓ'} - ({x : A} F x G x) ({x : A} F x) ({x : A} G x) -equivImplicitΠCod k .fst f {x} = k {x} .fst (f {x}) -equivImplicitΠCod k .snd .equiv-proof f .fst .fst {x} = equivCtr (k {x}) (f {x}) .fst -equivImplicitΠCod k .snd .equiv-proof f .fst .snd i {x} = equivCtr (k {x}) (f {x}) .snd i -equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .fst {x} = - equivCtrPath (k {x}) (f {x}) (g {x} , λ j p j {x}) i .fst -equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .snd j {x} = - equivCtrPath (k {x}) (f {x}) (g {x} , λ k p k {x}) i .snd j - -equiv→Iso : (A B) (C D) Iso (A C) (B D) -equiv→Iso h k .Iso.fun f b = equivFun k (f (invEq h b)) -equiv→Iso h k .Iso.inv g a = invEq k (g (equivFun h a)) -equiv→Iso h k .Iso.rightInv g = funExt λ b secEq k _ cong g (secEq h b) -equiv→Iso h k .Iso.leftInv f = funExt λ a retEq k _ cong f (retEq h a) - -equiv→ : (A B) (C D) (A C) (B D) -equiv→ h k = isoToEquiv (equiv→Iso h k) - - -equivΠ' : {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} - {B : A Type ℓB} {B' : A' Type ℓB'} - (eA : A A') - (eB : {a : A} {a' : A'} eA .fst a a' B a B' a') - ((a : A) B a) ((a' : A') B' a') -equivΠ' {B' = B'} eA eB = isoToEquiv isom - where - open Iso - - isom : Iso _ _ - isom .fun f a' = - eB (secEq eA a') .fst (f (invEq eA a')) - isom .inv f' a = - invEq (eB refl) (f' (eA .fst a)) - isom .rightInv f' = - funExt λ a' - J a'' p eB p .fst (invEq (eB refl) (f' (p i0))) f' a'') - (secEq (eB refl) (f' (eA .fst (invEq eA a')))) - (secEq eA a') - isom .leftInv f = - funExt λ a - subst - p invEq (eB refl) (eB p .fst (f (invEq eA (eA .fst a)))) f a) - (sym (commPathIsEq (eA .snd) a)) - (J a'' p invEq (eB refl) (eB (cong (eA .fst) p) .fst (f (invEq eA (eA .fst a)))) f a'') - (retEq (eB refl) (f (invEq eA (eA .fst a)))) - (retEq eA a)) - -equivΠ : {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} - {B : A Type ℓB} {B' : A' Type ℓB'} - (eA : A A') - (eB : (a : A) B a B' (eA .fst a)) - ((a : A) B a) ((a' : A') B' a') -equivΠ {B = B} {B' = B'} eA eB = equivΠ' eA {a = a} p J a' p B a B' a') (eB a) p) - - -equivCompIso : (A B) (C D) Iso (A C) (B D) -equivCompIso h k .Iso.fun f = compEquiv (compEquiv (invEquiv h) f) k -equivCompIso h k .Iso.inv g = compEquiv (compEquiv h g) (invEquiv k) -equivCompIso h k .Iso.rightInv g = equivEq (equiv→Iso h k .Iso.rightInv (equivFun g)) -equivCompIso h k .Iso.leftInv f = equivEq (equiv→Iso h k .Iso.leftInv (equivFun f)) - -equivComp : (A B) (C D) (A C) (B D) -equivComp h k = isoToEquiv (equivCompIso h k) - --- Some helpful notation: -_≃⟨_⟩_ : (X : Type ) (X B) (B C) (X C) -_ ≃⟨ f g = compEquiv f g - -_■ : (X : Type ) (X X) -_■ = idEquiv - -infixr 0 _≃⟨_⟩_ -infix 1 _■ - -composesToId→Equiv : (f : A B) (g : B A) f g idfun B isEquiv f isEquiv g -composesToId→Equiv f g id iseqf = - isoToIsEquiv - (iso g f - b i equiv-proof iseqf (f b) .snd (g (f b) , cong h h (f b)) id) (~ i) .fst) - ∙∙ cong x equiv-proof iseqf (f b) .fst .fst) id - ∙∙ λ i equiv-proof iseqf (f b) .snd (b , refl) i .fst) - λ a i id i a) - -precomposesToId→Equiv : (f : A B) (g : B A) f g idfun B isEquiv g isEquiv f -precomposesToId→Equiv f g id iseqg = subst isEquiv (sym f-≡-g⁻) (snd (invEquiv (_ , iseqg))) - where - g⁻ = invEq (g , iseqg) - - f-≡-g⁻ : _ - f-≡-g⁻ = cong (f ∘_ ) (cong fst (sym (invEquiv-is-linv (g , iseqg)))) cong (_∘ g⁻) id - --- equivalence between isEquiv and isEquiv' - -isEquiv-isEquiv'-Iso : (f : A B) Iso (isEquiv f) (isEquiv' f) -isEquiv-isEquiv'-Iso f .fun p = p .equiv-proof -isEquiv-isEquiv'-Iso f .inv q .equiv-proof = q -isEquiv-isEquiv'-Iso f .rightInv q = refl -isEquiv-isEquiv'-Iso f .leftInv p i .equiv-proof = p .equiv-proof - -isEquiv≃isEquiv' : (f : A B) isEquiv f isEquiv' f -isEquiv≃isEquiv' f = isoToEquiv (isEquiv-isEquiv'-Iso f) - --- The fact that funExt is an equivalence can be found in Cubical.Functions.FunExtEquiv +equivPathP : {A : I Type } {B : I Type ℓ'} {e : A i0 B i0} {f : A i1 B i1} + (h : PathP i A i B i) (e .fst) (f .fst)) PathP i A i B i) e f +equivPathP {e = e} {f = f} h = + λ i (h i) , isProp→PathP i isPropIsEquiv (h i)) (e .snd) (f .snd) i + +equivEq : {e f : A B} (h : e .fst f .fst) e f +equivEq = equivPathP + +module _ {f : A B} (equivF : isEquiv f) where + funIsEq : A B + funIsEq = f + + invIsEq : B A + invIsEq y = equivF .equiv-proof y .fst .fst + + secIsEq : section f invIsEq + secIsEq y = equivF .equiv-proof y .fst .snd + + retIsEq : retract f invIsEq + retIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .fst + + commSqIsEq : a Square (secIsEq (f a)) refl (cong f (retIsEq a)) refl + commSqIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .snd + + commPathIsEq : a secIsEq (f a) cong f (retIsEq a) + commPathIsEq a i j = + hcomp + k λ + { (i = i0) secIsEq (f a) j + ; (i = i1) f (retIsEq a (j ~ k)) + ; (j = i0) f (retIsEq a (i ~ k)) + ; (j = i1) f a + }) + (commSqIsEq a i j) + +module _ (w : A B) where + invEq : B A + invEq = invIsEq (snd w) + + retEq : retract (w .fst) invEq + retEq = retIsEq (snd w) + + secEq : section (w .fst) invEq + secEq = secIsEq (snd w) + +open Iso + +equivToIso : { ℓ'} {A : Type } {B : Type ℓ'} A B Iso A B +fun (equivToIso e) = e .fst +inv (equivToIso e) = invEq e +rightInv (equivToIso e) = secEq e +leftInv (equivToIso e) = retEq e + +-- TODO: there should be a direct proof of this that doesn't use equivToIso +invEquiv : A B B A +invEquiv e = isoToEquiv (invIso (equivToIso e)) + +invEquivIdEquiv : (A : Type ) invEquiv (idEquiv A) idEquiv A +invEquivIdEquiv _ = equivEq refl + +compEquiv : A B B C A C +compEquiv f g .fst = g .fst f .fst +compEquiv {A = A} {C = C} f g .snd .equiv-proof c = contr + where + contractG = g .snd .equiv-proof c .snd + + secFiller : (a : A) (p : g .fst (f .fst a) c) _ {- square in A -} + secFiller a p = compPath-filler (cong (invEq f fst) (contractG (_ , p))) (retEq f a) + + contr : isContr (fiber (g .fst f .fst) c) + contr .fst .fst = invEq f (invEq g c) + contr .fst .snd = cong (g .fst) (secEq f (invEq g c)) secEq g c + contr .snd (a , p) i .fst = secFiller a p i1 i + contr .snd (a , p) i .snd j = + hcomp + k λ + { (i = i1) fSquare k + ; (j = i0) g .fst (f .fst (secFiller a p k i)) + ; (j = i1) contractG (_ , p) i .snd k + }) + (g .fst (secEq f (contractG (_ , p) i .fst) j)) + where + fSquare : I C + fSquare k = + hcomp + l λ + { (j = i0) g .fst (f .fst (retEq f a k)) + ; (j = i1) p (k l) + ; (k = i0) g .fst (secEq f (f .fst a) j) + ; (k = i1) p (j l) + }) + (g .fst (f .snd .equiv-proof (f .fst a) .snd (a , refl) k .snd j)) + +_∙ₑ_ = compEquiv + +compEquivIdEquiv : (e : A B) compEquiv (idEquiv A) e e +compEquivIdEquiv e = equivEq refl + +compEquivEquivId : (e : A B) compEquiv e (idEquiv B) e +compEquivEquivId e = equivEq refl + +invEquiv-is-rinv : (e : A B) compEquiv e (invEquiv e) idEquiv A +invEquiv-is-rinv e = equivEq (funExt (retEq e)) + +invEquiv-is-linv : (e : A B) compEquiv (invEquiv e) e idEquiv B +invEquiv-is-linv e = equivEq (funExt (secEq e)) + +compEquiv-assoc : (f : A B) (g : B C) (h : C D) + compEquiv f (compEquiv g h) compEquiv (compEquiv f g) h +compEquiv-assoc f g h = equivEq refl + +LiftEquiv : A Lift {i = } {j = ℓ'} A +LiftEquiv .fst a .lower = a +LiftEquiv .snd .equiv-proof = strictContrFibers lower + +Lift≃Lift : (e : A B) Lift {j = ℓ'} A Lift {j = ℓ''} B +Lift≃Lift e .fst a .lower = e .fst (a .lower) +Lift≃Lift e .snd .equiv-proof b .fst .fst .lower = invEq e (b .lower) +Lift≃Lift e .snd .equiv-proof b .fst .snd i .lower = + e .snd .equiv-proof (b .lower) .fst .snd i +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .fst .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .fst +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .snd j .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .snd j + +isContr→Equiv : isContr A isContr B A B +isContr→Equiv Actr Bctr = isoToEquiv (isContr→Iso Actr Bctr) + +propBiimpl→Equiv : (Aprop : isProp A) (Bprop : isProp B) (f : A B) (g : B A) A B +propBiimpl→Equiv Aprop Bprop f g = f , hf + where + hf : isEquiv f + hf .equiv-proof y .fst = (g y , Bprop (f (g y)) y) + hf .equiv-proof y .snd h i .fst = Aprop (g y) (h .fst) i + hf .equiv-proof y .snd h i .snd = isProp→isSet' Bprop (Bprop (f (g y)) y) (h .snd) + (cong f (Aprop (g y) (h .fst))) refl i + +isEquivPropBiimpl→Equiv : isProp A isProp B + ((A B) × (B A)) (A B) +isEquivPropBiimpl→Equiv {A = A} {B = B} Aprop Bprop = isoToEquiv isom where + isom : Iso (Σ (A B) _ B A)) (A B) + isom .fun (f , g) = propBiimpl→Equiv Aprop Bprop f g + isom .inv e = equivFun e , invEq e + isom .rightInv e = equivEq refl + isom .leftInv _ = refl + +equivΠCod : {F : A Type } {G : A Type ℓ'} + ((x : A) F x G x) ((x : A) F x) ((x : A) G x) +equivΠCod k .fst f x = k x .fst (f x) +equivΠCod k .snd .equiv-proof f .fst .fst x = equivCtr (k x) (f x) .fst +equivΠCod k .snd .equiv-proof f .fst .snd i x = equivCtr (k x) (f x) .snd i +equivΠCod k .snd .equiv-proof f .snd (g , p) i .fst x = + equivCtrPath (k x) (f x) (g x , λ j p j x) i .fst +equivΠCod k .snd .equiv-proof f .snd (g , p) i .snd j x = + equivCtrPath (k x) (f x) (g x , λ k p k x) i .snd j + +equivImplicitΠCod : {F : A Type } {G : A Type ℓ'} + ({x : A} F x G x) ({x : A} F x) ({x : A} G x) +equivImplicitΠCod k .fst f {x} = k {x} .fst (f {x}) +equivImplicitΠCod k .snd .equiv-proof f .fst .fst {x} = equivCtr (k {x}) (f {x}) .fst +equivImplicitΠCod k .snd .equiv-proof f .fst .snd i {x} = equivCtr (k {x}) (f {x}) .snd i +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .fst {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ j p j {x}) i .fst +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .snd j {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ k p k {x}) i .snd j + +equiv→Iso : (A B) (C D) Iso (A C) (B D) +equiv→Iso h k .Iso.fun f b = equivFun k (f (invEq h b)) +equiv→Iso h k .Iso.inv g a = invEq k (g (equivFun h a)) +equiv→Iso h k .Iso.rightInv g = funExt λ b secEq k _ cong g (secEq h b) +equiv→Iso h k .Iso.leftInv f = funExt λ a retEq k _ cong f (retEq h a) + +equiv→ : (A B) (C D) (A C) (B D) +equiv→ h k = isoToEquiv (equiv→Iso h k) + + +equivΠ' : {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A Type ℓB} {B' : A' Type ℓB'} + (eA : A A') + (eB : {a : A} {a' : A'} eA .fst a a' B a B' a') + ((a : A) B a) ((a' : A') B' a') +equivΠ' {B' = B'} eA eB = isoToEquiv isom + where + open Iso + + isom : Iso _ _ + isom .fun f a' = + eB (secEq eA a') .fst (f (invEq eA a')) + isom .inv f' a = + invEq (eB refl) (f' (eA .fst a)) + isom .rightInv f' = + funExt λ a' + J a'' p eB p .fst (invEq (eB refl) (f' (p i0))) f' a'') + (secEq (eB refl) (f' (eA .fst (invEq eA a')))) + (secEq eA a') + isom .leftInv f = + funExt λ a + subst + p invEq (eB refl) (eB p .fst (f (invEq eA (eA .fst a)))) f a) + (sym (commPathIsEq (eA .snd) a)) + (J a'' p invEq (eB refl) (eB (cong (eA .fst) p) .fst (f (invEq eA (eA .fst a)))) f a'') + (retEq (eB refl) (f (invEq eA (eA .fst a)))) + (retEq eA a)) + +equivΠ : {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A Type ℓB} {B' : A' Type ℓB'} + (eA : A A') + (eB : (a : A) B a B' (eA .fst a)) + ((a : A) B a) ((a' : A') B' a') +equivΠ {B = B} {B' = B'} eA eB = equivΠ' eA {a = a} p J a' p B a B' a') (eB a) p) + + +equivCompIso : (A B) (C D) Iso (A C) (B D) +equivCompIso h k .Iso.fun f = compEquiv (compEquiv (invEquiv h) f) k +equivCompIso h k .Iso.inv g = compEquiv (compEquiv h g) (invEquiv k) +equivCompIso h k .Iso.rightInv g = equivEq (equiv→Iso h k .Iso.rightInv (equivFun g)) +equivCompIso h k .Iso.leftInv f = equivEq (equiv→Iso h k .Iso.leftInv (equivFun f)) + +equivComp : (A B) (C D) (A C) (B D) +equivComp h k = isoToEquiv (equivCompIso h k) + +-- Some helpful notation: +_≃⟨_⟩_ : (X : Type ) (X B) (B C) (X C) +_ ≃⟨ f g = compEquiv f g + +_■ : (X : Type ) (X X) +_■ = idEquiv + +infixr 0 _≃⟨_⟩_ +infix 1 _■ + +composesToId→Equiv : (f : A B) (g : B A) f g idfun B isEquiv f isEquiv g +composesToId→Equiv f g id iseqf = + isoToIsEquiv + (iso g f + b i equiv-proof iseqf (f b) .snd (g (f b) , cong h h (f b)) id) (~ i) .fst) + ∙∙ cong x equiv-proof iseqf (f b) .fst .fst) id + ∙∙ λ i equiv-proof iseqf (f b) .snd (b , refl) i .fst) + λ a i id i a) + +precomposesToId→Equiv : (f : A B) (g : B A) f g idfun B isEquiv g isEquiv f +precomposesToId→Equiv f g id iseqg = subst isEquiv (sym f-≡-g⁻) (snd (invEquiv (_ , iseqg))) + where + g⁻ = invEq (g , iseqg) + + f-≡-g⁻ : _ + f-≡-g⁻ = cong (f ∘_ ) (cong fst (sym (invEquiv-is-linv (g , iseqg)))) cong (_∘ g⁻) id + +-- equivalence between isEquiv and isEquiv' + +isEquiv-isEquiv'-Iso : (f : A B) Iso (isEquiv f) (isEquiv' f) +isEquiv-isEquiv'-Iso f .fun p = p .equiv-proof +isEquiv-isEquiv'-Iso f .inv q .equiv-proof = q +isEquiv-isEquiv'-Iso f .rightInv q = refl +isEquiv-isEquiv'-Iso f .leftInv p i .equiv-proof = p .equiv-proof + +isEquiv≃isEquiv' : (f : A B) isEquiv f isEquiv' f +isEquiv≃isEquiv' f = isoToEquiv (isEquiv-isEquiv'-Iso f) + +-- The fact that funExt is an equivalence can be found in Cubical.Functions.FunExtEquiv \ No newline at end of file diff --git a/docs/Cubical.Foundations.Function.html b/docs/Cubical.Foundations.Function.html index a40ec2f..b9a3c60 100644 --- a/docs/Cubical.Foundations.Function.html +++ b/docs/Cubical.Foundations.Function.html @@ -83,80 +83,83 @@ curry : ((p : Σ A B) C (fst p) (snd p)) (x : A) (y : B x) C x y curry f x y = f (x , y) -module _ { ℓ'} {A : Type } {B : Type ℓ'} where - -- Notions of 'coherently constant' functions for low dimensions. - -- These are the properties of functions necessary to e.g. eliminate - -- from the propositional truncation. - - -- 2-Constant functions are coherently constant if B is a set. - 2-Constant : (A B) Type _ - 2-Constant f = x y f x f y - - 2-Constant-isProp : isSet B (f : A B) isProp (2-Constant f) - 2-Constant-isProp Bset f link1 link2 i x y j - = Bset (f x) (f y) (link1 x y) (link2 x y) i j - - -- 3-Constant functions are coherently constant if B is a groupoid. - record 3-Constant (f : A B) : Type (ℓ-max ℓ') where - field - link : 2-Constant f - coh₁ : x y z Square (link x y) (link x z) refl (link y z) - - coh₂ : x y z Square (link x z) (link y z) (link x y) refl - coh₂ x y z i j - = hcomp k λ - { (j = i0) link x y i - ; (i = i0) link x z (j k) - ; (j = i1) link x z (i k) - ; (i = i1) link y z j - }) - (coh₁ x y z j i) - - link≡refl : x refl link x x - link≡refl x i j - = hcomp k λ - { (i = i0) link x x (j ~ k) - ; (i = i1) link x x j - ; (j = i0) f x - ; (j = i1) link x x (~ i ~ k) - }) - (coh₁ x x x (~ i) j) - - downleft : x y Square (link x y) refl refl (link y x) - downleft x y i j - = hcomp k λ - { (i = i0) link x y j - ; (i = i1) link≡refl x (~ k) j - ; (j = i0) f x - ; (j = i1) link y x i - }) - (coh₁ x y x i j) - - link≡symlink : x y link x y sym (link y x) - link≡symlink x y i j - = hcomp k λ - { (i = i0) link x y j - ; (i = i1) link y x (~ j ~ k) - ; (j = i0) f x - ; (j = i1) link y x (i ~ k) - }) - (downleft x y i j) - -homotopyNatural : {B : Type ℓ'} {f g : A B} (H : a f a g a) {x y : A} (p : x y) - H x cong g p cong f p H y -homotopyNatural {f = f} {g = g} H {x} {y} p i j = - hcomp k λ { (i = i0) compPath-filler (H x) (cong g p) k j - ; (i = i1) compPath-filler' (cong f p) (H y) k j - ; (j = i0) cong f p (i ~ k) - ; (j = i1) cong g p (i k) }) - (H (p i) j) - -homotopySymInv : {f : A A} (H : a f a a) (a : A) - Path (f a f a) i H (H a (~ i)) i) refl -homotopySymInv {f = f} H a j i = - hcomp k λ { (i = i0) f a - ; (i = i1) H a (j ~ k) - ; (j = i0) H (H a (~ i)) i - ; (j = i1) H a (i ~ k)}) - (H (H a (~ i j)) i) +∘diag : {B : (x y : A) Type } (∀ x y B x y) x B x x +∘diag f x = f x x + +module _ { ℓ'} {A : Type } {B : Type ℓ'} where + -- Notions of 'coherently constant' functions for low dimensions. + -- These are the properties of functions necessary to e.g. eliminate + -- from the propositional truncation. + + -- 2-Constant functions are coherently constant if B is a set. + 2-Constant : (A B) Type _ + 2-Constant f = x y f x f y + + 2-Constant-isProp : isSet B (f : A B) isProp (2-Constant f) + 2-Constant-isProp Bset f link1 link2 i x y j + = Bset (f x) (f y) (link1 x y) (link2 x y) i j + + -- 3-Constant functions are coherently constant if B is a groupoid. + record 3-Constant (f : A B) : Type (ℓ-max ℓ') where + field + link : 2-Constant f + coh₁ : x y z Square (link x y) (link x z) refl (link y z) + + coh₂ : x y z Square (link x z) (link y z) (link x y) refl + coh₂ x y z i j + = hcomp k λ + { (j = i0) link x y i + ; (i = i0) link x z (j k) + ; (j = i1) link x z (i k) + ; (i = i1) link y z j + }) + (coh₁ x y z j i) + + link≡refl : x refl link x x + link≡refl x i j + = hcomp k λ + { (i = i0) link x x (j ~ k) + ; (i = i1) link x x j + ; (j = i0) f x + ; (j = i1) link x x (~ i ~ k) + }) + (coh₁ x x x (~ i) j) + + downleft : x y Square (link x y) refl refl (link y x) + downleft x y i j + = hcomp k λ + { (i = i0) link x y j + ; (i = i1) link≡refl x (~ k) j + ; (j = i0) f x + ; (j = i1) link y x i + }) + (coh₁ x y x i j) + + link≡symlink : x y link x y sym (link y x) + link≡symlink x y i j + = hcomp k λ + { (i = i0) link x y j + ; (i = i1) link y x (~ j ~ k) + ; (j = i0) f x + ; (j = i1) link y x (i ~ k) + }) + (downleft x y i j) + +homotopyNatural : {B : Type ℓ'} {f g : A B} (H : a f a g a) {x y : A} (p : x y) + H x cong g p cong f p H y +homotopyNatural {f = f} {g = g} H {x} {y} p i j = + hcomp k λ { (i = i0) compPath-filler (H x) (cong g p) k j + ; (i = i1) compPath-filler' (cong f p) (H y) k j + ; (j = i0) cong f p (i ~ k) + ; (j = i1) cong g p (i k) }) + (H (p i) j) + +homotopySymInv : {f : A A} (H : a f a a) (a : A) + Path (f a f a) i H (H a (~ i)) i) refl +homotopySymInv {f = f} H a j i = + hcomp k λ { (i = i0) f a + ; (i = i1) H a (j ~ k) + ; (j = i0) H (H a (~ i)) i + ; (j = i1) H a (i ~ k)}) + (H (H a (~ i j)) i) \ No newline at end of file diff --git a/docs/Cubical.Foundations.GroupoidLaws.html b/docs/Cubical.Foundations.GroupoidLaws.html index 1e09c8a..33b8fa6 100644 --- a/docs/Cubical.Foundations.GroupoidLaws.html +++ b/docs/Cubical.Foundations.GroupoidLaws.html @@ -169,7 +169,7 @@ -- some exchange law for doubleCompPath and refl -invSides-filler : {x y z : A} (p : x y) (q : x z) Square p (sym q) q (sym p) +invSides-filler : {x y z : A} (p : x y) (q : x z) Square p (sym q) q (sym p) invSides-filler {x = x} p q i j = hcomp k λ { (i = i0) p (k j) ; (i = i1) q (~ j k) @@ -438,7 +438,7 @@ compPath-filler-in-filler : (p : _ y) (q : _ _ ) - _≡_ {A = Square (p q) (p q) _ x) _ z)} + _≡_ {A = Square (p q) (p q) _ x) _ z)} i j hcomp i₂ λ { (j = i0) x @@ -473,7 +473,7 @@ cube-comp₋₀₋ : (c : I I I A) - {a' : Square _ _ _ _} + {a' : Square _ _ _ _} i i₁ c i i0 i₁) a' (I I I A) cube-comp₋₀₋ c p i j k = @@ -490,7 +490,7 @@ cube-comp₀₋₋ : (c : I I I A) - {a' : Square _ _ _ _} + {a' : Square _ _ _ _} i i₁ c i0 i i₁) a' (I I I A) cube-comp₀₋₋ c p i j k = diff --git a/docs/Cubical.Foundations.HLevels.html b/docs/Cubical.Foundations.HLevels.html index 411eb21..ff42a1e 100644 --- a/docs/Cubical.Foundations.HLevels.html +++ b/docs/Cubical.Foundations.HLevels.html @@ -43,8 +43,8 @@ n : HLevel isOfHLevel : HLevel Type Type -isOfHLevel 0 A = isContr A -isOfHLevel 1 A = isProp A +isOfHLevel 0 A = isContr A +isOfHLevel 1 A = isProp A isOfHLevel (suc (suc n)) A = (x y : A) isOfHLevel (suc n) (x y) isOfHLevelFun : (n : HLevel) {A : Type } {B : Type ℓ'} (f : A B) Type (ℓ-max ℓ') @@ -54,9 +54,9 @@ {} {A : Type } (n : ) ((x : A) isOfHLevel (suc n) (x x)) isOfHLevel (2 + n) A isOfHLevelΩ→isOfHLevel zero x y = - J y p (q : x y) p q) ( x refl) + J y p (q : x y) p q) ( x refl) isOfHLevelΩ→isOfHLevel (suc n) x y = - J y p (q : x y) isOfHLevel (suc n) (p q)) ( x refl) + J y p (q : x y) isOfHLevel (suc n) (p q)) ( x refl) TypeOfHLevel : HLevel Type (ℓ-suc ) TypeOfHLevel n = TypeWithStr (isOfHLevel n) @@ -70,25 +70,25 @@ -- lower h-levels imply higher h-levels isOfHLevelSuc : (n : HLevel) isOfHLevel n A isOfHLevel (suc n) A -isOfHLevelSuc 0 = isContr→isProp -isOfHLevelSuc 1 = isProp→isSet +isOfHLevelSuc 0 = isContr→isProp +isOfHLevelSuc 1 = isProp→isSet isOfHLevelSuc (suc (suc n)) h a b = isOfHLevelSuc (suc n) (h a b) -isSet→isGroupoid : isSet A isGroupoid A +isSet→isGroupoid : isSet A isGroupoid A isSet→isGroupoid = isOfHLevelSuc 2 -isGroupoid→is2Groupoid : isGroupoid A is2Groupoid A +isGroupoid→is2Groupoid : isGroupoid A is2Groupoid A isGroupoid→is2Groupoid = isOfHLevelSuc 3 isOfHLevelPlus : (m : HLevel) isOfHLevel n A isOfHLevel (m + n) A isOfHLevelPlus zero hA = hA isOfHLevelPlus (suc m) hA = isOfHLevelSuc _ (isOfHLevelPlus m hA) -isContr→isOfHLevel : (n : HLevel) isContr A isOfHLevel n A +isContr→isOfHLevel : (n : HLevel) isContr A isOfHLevel n A isContr→isOfHLevel zero cA = cA isContr→isOfHLevel (suc n) cA = isOfHLevelSuc _ (isContr→isOfHLevel n cA) -isProp→isOfHLevelSuc : (n : HLevel) isProp A isOfHLevel (suc n) A +isProp→isOfHLevelSuc : (n : HLevel) isProp A isOfHLevel (suc n) A isProp→isOfHLevelSuc zero pA = pA isProp→isOfHLevelSuc (suc n) pA = isOfHLevelSuc _ (isProp→isOfHLevelSuc n pA) @@ -99,11 +99,11 @@ -- hlevel of path types -isProp→isContrPath : isProp A (x y : A) isContr (x y) -isProp→isContrPath h x y = h x y , isProp→isSet h x y _ +isProp→isContrPath : isProp A (x y : A) isContr (x y) +isProp→isContrPath h x y = h x y , isProp→isSet h x y _ -isContr→isContrPath : isContr A (x y : A) isContr (x y) -isContr→isContrPath cA = isProp→isContrPath (isContr→isProp cA) +isContr→isContrPath : isContr A (x y : A) isContr (x y) +isContr→isContrPath cA = isProp→isContrPath (isContr→isProp cA) isOfHLevelPath' : (n : HLevel) isOfHLevel (suc n) A (x y : A) isOfHLevel n (x y) isOfHLevelPath' 0 = isProp→isContrPath @@ -119,19 +119,19 @@ -- h-level of isOfHLevel -isPropIsOfHLevel : (n : HLevel) isProp (isOfHLevel n A) -isPropIsOfHLevel 0 = isPropIsContr -isPropIsOfHLevel 1 = isPropIsProp +isPropIsOfHLevel : (n : HLevel) isProp (isOfHLevel n A) +isPropIsOfHLevel 0 = isPropIsContr +isPropIsOfHLevel 1 = isPropIsProp isPropIsOfHLevel (suc (suc n)) f g i a b = isPropIsOfHLevel (suc n) (f a b) (g a b) i -isPropIsSet : isProp (isSet A) +isPropIsSet : isProp (isSet A) isPropIsSet = isPropIsOfHLevel 2 -isPropIsGroupoid : isProp (isGroupoid A) +isPropIsGroupoid : isProp (isGroupoid A) isPropIsGroupoid = isPropIsOfHLevel 3 -isPropIs2Groupoid : isProp (is2Groupoid A) +isPropIs2Groupoid : isProp (is2Groupoid A) isPropIs2Groupoid = isPropIsOfHLevel 4 TypeOfHLevel≡ : (n : HLevel) {X Y : TypeOfHLevel n} X Y X Y @@ -143,7 +143,7 @@ : {B : Type } (f : A B) (g : B A) (h : retract f g) - (v : isContr B) isContr A + (v : isContr B) isContr A fst (isContrRetract f g h (b , p)) = g b snd (isContrRetract f g h (b , p)) x = (cong g (p (f x))) (h x) @@ -151,7 +151,7 @@ : {B : Type } (f : A B) (g : B A) (h : (x : A) g (f x) x) - isProp B isProp A + isProp B isProp A isPropRetract f g h p x y i = hcomp j λ @@ -163,7 +163,7 @@ : {B : Type } (f : A B) (g : B A) (h : (x : A) g (f x) x) - isSet B isSet A + isSet B isSet A isSetRetract f g h set x y p q i j = hcomp k λ { (i = i0) h (p j) k ; (i = i1) h (q j) k @@ -176,7 +176,7 @@ : {B : Type } (f : A B) (g : B A) (h : (x : A) g (f x) x) - isGroupoid B isGroupoid A + isGroupoid B isGroupoid A isGroupoidRetract f g h grp x y p q P Q i j k = hcomp ((λ l λ { (i = i0) h (P j k) l ; (i = i1) h (Q j k) l @@ -191,7 +191,7 @@ : {B : Type } (f : A B) (g : B A) (h : (x : A) g (f x) x) - is2Groupoid B is2Groupoid A + is2Groupoid B is2Groupoid A is2GroupoidRetract f g h grp x y p q P Q R S i j k l = hcomp r λ { (i = i0) h (R j k l) r ; (i = i1) h (S j k l) r @@ -247,11 +247,11 @@ isOfHLevelRetractFromIso n e hlev = isOfHLevelRetract n (Iso.fun e) (Iso.inv e) (Iso.leftInv e) hlev isOfHLevelRespectEquiv : {A : Type } {B : Type ℓ'} (n : HLevel) A B isOfHLevel n A isOfHLevel n B -isOfHLevelRespectEquiv n eq = isOfHLevelRetract n (invEq eq) (eq .fst) (secEq eq) +isOfHLevelRespectEquiv n eq = isOfHLevelRetract n (invEq eq) (eq .fst) (secEq eq) isContrRetractOfConstFun : {A : Type } {B : Type ℓ'} (b₀ : B) Σ[ f (B A) ] ((x : A) (f _ b₀)) x x) - isContr A + isContr A fst (isContrRetractOfConstFun b₀ ret) = ret .fst b₀ snd (isContrRetractOfConstFun b₀ ret) y = ret .snd y @@ -273,35 +273,35 @@ isSet→SquareP : {A : I I Type } - (isSet : (i j : I) isSet (A i j)) + (isSet : (i j : I) isSet (A i j)) {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP j A i0 j) a₀₀ a₀₁) {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP j A i1 j) a₁₀ a₁₁) (a₋₀ : PathP i A i i0) a₀₀ a₁₀) (a₋₁ : PathP i A i i1) a₀₁ a₁₁) - SquareP A a₀₋ a₁₋ a₋₀ a₋₁ + SquareP A a₀₋ a₁₋ a₋₀ a₋₁ isSet→SquareP isset a₀₋ a₁₋ a₋₀ a₋₁ = PathPIsoPath _ _ _ .Iso.inv (isOfHLevelPathP' 1 (isset _ _) _ _ _ _ ) -isGroupoid→isGroupoid' : isGroupoid A isGroupoid' A +isGroupoid→isGroupoid' : isGroupoid A isGroupoid' A isGroupoid→isGroupoid' {A = A} Agpd a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = - PathPIsoPath i Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ .Iso.inv + PathPIsoPath i Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ .Iso.inv (isGroupoid→isPropSquare _ _ _ _ _ _) where isGroupoid→isPropSquare : {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) - isProp (Square a₀₋ a₁₋ a₋₀ a₋₁) + isProp (Square a₀₋ a₁₋ a₋₀ a₋₁) isGroupoid→isPropSquare a₀₋ a₁₋ a₋₀ a₋₁ = isOfHLevelRetractFromIso 1 (PathPIsoPath i a₋₀ i a₋₁ i) a₀₋ a₁₋) (Agpd _ _ _ _) -isGroupoid'→isGroupoid : isGroupoid' A isGroupoid A +isGroupoid'→isGroupoid : isGroupoid' A isGroupoid A isGroupoid'→isGroupoid Agpd' x y p q r s = Agpd' r s refl refl refl refl -- h-level of Σ-types -isProp∃! : isProp (∃! A B) -isProp∃! = isPropIsContr +isProp∃! : isProp (∃! A B) +isProp∃! = isPropIsContr -isContrΣ : isContr A ((x : A) isContr (B x)) isContr (Σ A B) +isContrΣ : isContr A ((x : A) isContr (B x)) isContr (Σ A B) isContrΣ {A = A} {B = B} (a , p) q = let h : (x : A) (y : B x) (q x) .fst y h x y = (q x) .snd y @@ -309,24 +309,24 @@ , ( λ x i p (x .fst) i , h (p (x .fst) i) (transp j B (p (x .fst) (i ~ j))) i (x .snd)) i)) -isContrΣ' : (ca : isContr A) isContr (B (fst ca)) isContr (Σ A B) +isContrΣ' : (ca : isContr A) isContr (B (fst ca)) isContr (Σ A B) isContrΣ' ca cb = isContrΣ ca x subst _ (snd ca x) cb) section-Σ≡Prop - : (pB : (x : A) isProp (B x)) {u v : Σ A B} + : (pB : (x : A) isProp (B x)) {u v : Σ A B} section (Σ≡Prop pB {u} {v}) (cong fst) section-Σ≡Prop {A = A} pB {u} {v} p j i = - (p i .fst) , isProp→PathP i isOfHLevelPath 1 (pB (fst (p i))) + (p i .fst) , isProp→PathP i isOfHLevelPath 1 (pB (fst (p i))) (Σ≡Prop pB {u} {v} (cong fst p) i .snd) (p i .snd) ) refl refl i j isEquiv-Σ≡Prop - : (pB : (x : A) isProp (B x)) {u v : Σ A B} + : (pB : (x : A) isProp (B x)) {u v : Σ A B} isEquiv (Σ≡Prop pB {u} {v}) isEquiv-Σ≡Prop {A = A} pB {u} {v} = isoToIsEquiv (iso (Σ≡Prop pB) (cong fst) (section-Σ≡Prop pB) _ refl)) -isPropΣ : isProp A ((x : A) isProp (B x)) isProp (Σ A B) +isPropΣ : isProp A ((x : A) isProp (B x)) isProp (Σ A B) isPropΣ pA pB t u = Σ≡Prop pB (pA (t .fst) (u .fst)) isOfHLevelΣ : n isOfHLevel n A ((x : A) isOfHLevel n (B x)) @@ -338,39 +338,39 @@ (invIso (IsoΣPathTransportPathΣ _ _)) (isOfHLevelΣ (suc n) (h1 (fst x) (fst y)) λ x h2 _ _ _) -isSetΣ : isSet A ((x : A) isSet (B x)) isSet (Σ A B) +isSetΣ : isSet A ((x : A) isSet (B x)) isSet (Σ A B) isSetΣ = isOfHLevelΣ 2 -- Useful consequence -isSetΣSndProp : isSet A ((x : A) isProp (B x)) isSet (Σ A B) -isSetΣSndProp h p = isSetΣ h x isProp→isSet (p x)) +isSetΣSndProp : isSet A ((x : A) isProp (B x)) isSet (Σ A B) +isSetΣSndProp h p = isSetΣ h x isProp→isSet (p x)) -isGroupoidΣ : isGroupoid A ((x : A) isGroupoid (B x)) isGroupoid (Σ A B) +isGroupoidΣ : isGroupoid A ((x : A) isGroupoid (B x)) isGroupoid (Σ A B) isGroupoidΣ = isOfHLevelΣ 3 -is2GroupoidΣ : is2Groupoid A ((x : A) is2Groupoid (B x)) is2Groupoid (Σ A B) +is2GroupoidΣ : is2Groupoid A ((x : A) is2Groupoid (B x)) is2Groupoid (Σ A B) is2GroupoidΣ = isOfHLevelΣ 4 -- h-level of × -isProp× : {A : Type } {B : Type ℓ'} isProp A isProp B isProp (A × B) +isProp× : {A : Type } {B : Type ℓ'} isProp A isProp B isProp (A × B) isProp× pA pB = isPropΣ pA _ pB) isProp×2 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} - isProp A isProp B isProp C isProp (A × B × C) + isProp A isProp B isProp C isProp (A × B × C) isProp×2 pA pB pC = isProp× pA (isProp× pB pC) isProp×3 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} - isProp A isProp B isProp C isProp D isProp (A × B × C × D) + isProp A isProp B isProp C isProp D isProp (A × B × C × D) isProp×3 pA pB pC pD = isProp×2 pA pB (isProp× pC pD) isProp×4 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} - isProp A isProp B isProp C isProp D isProp E isProp (A × B × C × D × E) + isProp A isProp B isProp C isProp D isProp E isProp (A × B × C × D × E) isProp×4 pA pB pC pD pE = isProp×3 pA pB pC (isProp× pD pE) isProp×5 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} {F : Type ℓ'''''} - isProp A isProp B isProp C isProp D isProp E isProp F - isProp (A × B × C × D × E × F) + isProp A isProp B isProp C isProp D isProp E isProp F + isProp (A × B × C × D × E × F) isProp×5 pA pB pC pD pE pF = isProp×4 pA pB pC pD (isProp× pE pF) @@ -378,15 +378,15 @@ isOfHLevel n (A × B) isOfHLevel× n hA hB = isOfHLevelΣ n hA _ hB) -isSet× : {A : Type } {B : Type ℓ'} isSet A isSet B isSet (A × B) +isSet× : {A : Type } {B : Type ℓ'} isSet A isSet B isSet (A × B) isSet× = isOfHLevel× 2 -isGroupoid× : {A : Type } {B : Type ℓ'} isGroupoid A isGroupoid B - isGroupoid (A × B) +isGroupoid× : {A : Type } {B : Type ℓ'} isGroupoid A isGroupoid B + isGroupoid (A × B) isGroupoid× = isOfHLevel× 3 -is2Groupoid× : {A : Type } {B : Type ℓ'} is2Groupoid A is2Groupoid B - is2Groupoid (A × B) +is2Groupoid× : {A : Type } {B : Type ℓ'} is2Groupoid A is2Groupoid B + is2Groupoid (A × B) is2Groupoid× = isOfHLevel× 4 -- h-level of Π-types @@ -417,386 +417,427 @@ isOfHLevel n ((x : A) (y : B x) C x y) isOfHLevelΠ2 n f = isOfHLevelΠ n x isOfHLevelΠ n (f x)) -isContrΠ : (h : (x : A) isContr (B x)) isContr ((x : A) B x) +isContrΠ : (h : (x : A) isContr (B x)) isContr ((x : A) B x) isContrΠ = isOfHLevelΠ 0 -isPropΠ : (h : (x : A) isProp (B x)) isProp ((x : A) B x) +isPropΠ : (h : (x : A) isProp (B x)) isProp ((x : A) B x) isPropΠ = isOfHLevelΠ 1 -isPropΠ2 : (h : (x : A) (y : B x) isProp (C x y)) - isProp ((x : A) (y : B x) C x y) +isPropΠ2 : (h : (x : A) (y : B x) isProp (C x y)) + isProp ((x : A) (y : B x) C x y) isPropΠ2 h = isPropΠ λ x isPropΠ λ y h x y -isPropΠ3 : (h : (x : A) (y : B x) (z : C x y) isProp (D x y z)) - isProp ((x : A) (y : B x) (z : C x y) D x y z) +isPropΠ3 : (h : (x : A) (y : B x) (z : C x y) isProp (D x y z)) + isProp ((x : A) (y : B x) (z : C x y) D x y z) isPropΠ3 h = isPropΠ λ x isPropΠ λ y isPropΠ λ z h x y z -isPropΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isProp (E x y z w)) - isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) E x y z w) +isPropΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isProp (E x y z w)) + isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) E x y z w) isPropΠ4 h = isPropΠ λ _ isPropΠ3 (h _) -isPropΠ5 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) isProp (F x y z w v)) - isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) F x y z w v) +isPropΠ5 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) isProp (F x y z w v)) + isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) F x y z w v) isPropΠ5 h = isPropΠ λ _ isPropΠ4 (h _) -isPropΠ6 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) (u : F x y z w v) isProp (G x y z w v u)) - isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) (u : F x y z w v) G x y z w v u) +isPropΠ6 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) (u : F x y z w v) isProp (G x y z w v u)) + isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) (u : F x y z w v) G x y z w v u) isPropΠ6 h = isPropΠ λ _ isPropΠ5 (h _) -isPropImplicitΠ : (h : (x : A) isProp (B x)) isProp ({x : A} B x) +isPropImplicitΠ : (h : (x : A) isProp (B x)) isProp ({x : A} B x) isPropImplicitΠ h f g i {x} = h x (f {x}) (g {x}) i -isPropImplicitΠ2 : (h : (x : A) (y : B x) isProp (C x y)) isProp ({x : A} {y : B x} C x y) +isPropImplicitΠ2 : (h : (x : A) (y : B x) isProp (C x y)) isProp ({x : A} {y : B x} C x y) isPropImplicitΠ2 h = isPropImplicitΠ x isPropImplicitΠ y h x y)) -isProp→ : {A : Type } {B : Type ℓ'} isProp B isProp (A B) -isProp→ pB = isPropΠ λ _ pB - -isSetΠ : ((x : A) isSet (B x)) isSet ((x : A) B x) -isSetΠ = isOfHLevelΠ 2 - -isSetImplicitΠ : (h : (x : A) isSet (B x)) isSet ({x : A} B x) -isSetImplicitΠ h f g F G i j {x} = h x (f {x}) (g {x}) i F i {x}) i G i {x}) i j - -isSet→ : isSet A' isSet (A A') -isSet→ isSet-A' = isOfHLevelΠ 2 _ isSet-A') - -isSetΠ2 : (h : (x : A) (y : B x) isSet (C x y)) - isSet ((x : A) (y : B x) C x y) -isSetΠ2 h = isSetΠ λ x isSetΠ λ y h x y - -isSetΠ3 : (h : (x : A) (y : B x) (z : C x y) isSet (D x y z)) - isSet ((x : A) (y : B x) (z : C x y) D x y z) -isSetΠ3 h = isSetΠ λ x isSetΠ λ y isSetΠ λ z h x y z - -isGroupoidΠ : ((x : A) isGroupoid (B x)) isGroupoid ((x : A) B x) -isGroupoidΠ = isOfHLevelΠ 3 - -isGroupoidΠ2 : (h : (x : A) (y : B x) isGroupoid (C x y)) isGroupoid ((x : A) (y : B x) C x y) -isGroupoidΠ2 h = isGroupoidΠ λ _ isGroupoidΠ λ _ h _ _ - -isGroupoidΠ3 : (h : (x : A) (y : B x) (z : C x y) isGroupoid (D x y z)) - isGroupoid ((x : A) (y : B x) (z : C x y) D x y z) -isGroupoidΠ3 h = isGroupoidΠ λ _ isGroupoidΠ2 λ _ h _ _ - -isGroupoidΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isGroupoid (E x y z w)) - isGroupoid ((x : A) (y : B x) (z : C x y) (w : D x y z) E x y z w) -isGroupoidΠ4 h = isGroupoidΠ λ _ isGroupoidΠ3 λ _ h _ _ - -is2GroupoidΠ : ((x : A) is2Groupoid (B x)) is2Groupoid ((x : A) B x) -is2GroupoidΠ = isOfHLevelΠ 4 - -isOfHLevelΠ⁻ : {A : Type } {B : Type ℓ'} n - isOfHLevel n (A B) (A isOfHLevel n B) -isOfHLevelΠ⁻ 0 h x = fst h x , λ y funExt⁻ (snd h (const y)) x -isOfHLevelΠ⁻ 1 h x y z = funExt⁻ (h (const y) (const z)) x -isOfHLevelΠ⁻ (suc (suc n)) h x y z = - isOfHLevelΠ⁻ (suc n) (isOfHLevelRetractFromIso (suc n) funExtIso (h _ _)) x - -isOfHLevel→∙ : {A : Pointed } {B : Pointed ℓ'} (n : ) - isOfHLevel n (fst B) isOfHLevel n (A →∙ B) -isOfHLevel→∙ n hlev = - isOfHLevelΣ n (isOfHLevelΠ n _ hlev)) - λ _ isOfHLevelPath n hlev _ _ - --- h-level of A ≃ B and A ≡ B - -isOfHLevel≃ - : n {A : Type } {B : Type ℓ'} - (hA : isOfHLevel n A) (hB : isOfHLevel n B) isOfHLevel n (A B) -isOfHLevel≃ zero {A = A} {B = B} hA hB = isContr→Equiv hA hB , contr - where - contr : (y : A B) isContr→Equiv hA hB y - contr y = Σ≡Prop isPropIsEquiv (funExt a snd hB (fst y a))) - -isOfHLevel≃ (suc n) {A = A} {B = B} hA hB = - isOfHLevelΣ (suc n) (isOfHLevelΠ _ λ _ hB) - f isProp→isOfHLevelSuc n (isPropIsEquiv f)) - -isOfHLevel≡ : n {A B : Type } (hA : isOfHLevel n A) (hB : isOfHLevel n B) - isOfHLevel n (A B) -isOfHLevel≡ n hA hB = isOfHLevelRetractFromIso n univalenceIso (isOfHLevel≃ n hA hB) - -isOfHLevel⁺≃ₗ - : n {A : Type } {B : Type ℓ'} - isOfHLevel (suc n) A isOfHLevel (suc n) (A B) -isOfHLevel⁺≃ₗ zero pA e = isOfHLevel≃ 1 pA (isOfHLevelRespectEquiv 1 e pA) e -isOfHLevel⁺≃ₗ (suc n) hA e = isOfHLevel≃ m hA (isOfHLevelRespectEquiv m e hA) e - where - m = suc (suc n) - -isOfHLevel⁺≃ᵣ - : n {A : Type } {B : Type ℓ'} - isOfHLevel (suc n) B isOfHLevel (suc n) (A B) -isOfHLevel⁺≃ᵣ zero pB e - = isOfHLevel≃ 1 (isPropRetract (e .fst) (invEq e) (retEq e) pB) pB e -isOfHLevel⁺≃ᵣ (suc n) hB e - = isOfHLevel≃ m (isOfHLevelRetract m (e .fst) (invEq e) (retEq e) hB) hB e - where - m = suc (suc n) - -isOfHLevel⁺≡ₗ - : n {A B : Type } - isOfHLevel (suc n) A isOfHLevel (suc n) (A B) -isOfHLevel⁺≡ₗ zero pA P = isOfHLevel≡ 1 pA (subst isProp P pA) P -isOfHLevel⁺≡ₗ (suc n) hA P - = isOfHLevel≡ m hA (subst (isOfHLevel m) P hA) P - where - m = suc (suc n) - -isOfHLevel⁺≡ᵣ - : n {A B : Type } - isOfHLevel (suc n) B isOfHLevel (suc n) (A B) -isOfHLevel⁺≡ᵣ zero pB P = isOfHLevel≡ 1 (subst⁻ isProp P pB) pB P -isOfHLevel⁺≡ᵣ (suc n) hB P - = isOfHLevel≡ m (subst⁻ (isOfHLevel m) P hB) hB P - where - m = suc (suc n) - --- h-level of TypeOfHLevel - -isPropHContr : isProp (TypeOfHLevel 0) -isPropHContr x y = Σ≡Prop _ isPropIsContr) (isOfHLevel≡ 0 (x .snd) (y .snd) .fst) - -isOfHLevelTypeOfHLevel : n isOfHLevel (suc n) (TypeOfHLevel n) -isOfHLevelTypeOfHLevel zero = isPropHContr -isOfHLevelTypeOfHLevel (suc n) (X , a) (Y , b) = - isOfHLevelRetract (suc n) (cong fst) (Σ≡Prop λ _ isPropIsOfHLevel (suc n)) - (section-Σ≡Prop λ _ isPropIsOfHLevel (suc n)) - (isOfHLevel≡ (suc n) a b) - -isSetHProp : isSet (hProp ) -isSetHProp = isOfHLevelTypeOfHLevel 1 - -isGroupoidHSet : isGroupoid (hSet ) -isGroupoidHSet = isOfHLevelTypeOfHLevel 2 - - --- h-level of lifted type - -isOfHLevelLift : { ℓ'} (n : HLevel) {A : Type } isOfHLevel n A isOfHLevel n (Lift {j = ℓ'} A) -isOfHLevelLift n = isOfHLevelRetract n lower lift λ _ refl - -isOfHLevelLower : { ℓ'} (n : HLevel) {A : Type } isOfHLevel n (Lift {j = ℓ'} A) isOfHLevel n A -isOfHLevelLower n = isOfHLevelRetract n lift lower λ _ refl - ----------------------------- - --- More consequences of isProp and isContr - -inhProp→isContr : A isProp A isContr A -inhProp→isContr x h = x , h x - -extend : isContr A (∀ φ (u : Partial φ A) Sub A φ u) -extend (x , p) φ u = inS (hcomp { j (φ = i1) p (u 1=1) j }) x) - -isContrPartial→isContr : {} {A : Type } - (extend : φ Partial φ A A) - (∀ u u (extend i1 λ { _ u})) - isContr A -isContrPartial→isContr {A = A} extend law - = ex , λ y law ex i Aux.v y i) sym (law y) - where ex = extend i0 empty - module Aux (y : A) (i : I) where - φ = ~ i i - u : Partial φ A - u = λ { (i = i0) ex ; (i = i1) y } - v = extend φ u - --- Dependent h-level over a type - -isOfHLevelDep : HLevel {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') -isOfHLevelDep 0 {A = A} B = {a : A} Σ[ b B a ] ({a' : A} (b' : B a') (p : a a') PathP i B (p i)) b b') -isOfHLevelDep 1 {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) (p : a0 a1) PathP i B (p i)) b0 b1 -isOfHLevelDep (suc (suc n)) {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) isOfHLevelDep (suc n) {A = a0 a1} p PathP i B (p i)) b0 b1) - -isContrDep : {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') -isContrDep = isOfHLevelDep 0 - -isPropDep : {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') -isPropDep = isOfHLevelDep 1 - -isContrDep∘ - : {A' : Type } (f : A' A) isContrDep B isContrDep (B f) -isContrDep∘ f cB {a} = λ where - .fst cB .fst - .snd b' p cB .snd b' (cong f p) - -isPropDep∘ : {A' : Type } (f : A' A) isPropDep B isPropDep (B f) -isPropDep∘ f pB b0 b1 = pB b0 b1 cong f - -isOfHLevelDep→isOfHLevel : (n : HLevel) - {A : Type } {B : A Type ℓ'} isOfHLevelDep n {A = A} B (a : A) isOfHLevel n (B a) -isOfHLevelDep→isOfHLevel 0 h a = h .fst , λ b h .snd b refl -isOfHLevelDep→isOfHLevel 1 h a x y = h x y refl -isOfHLevelDep→isOfHLevel (suc (suc n)) h a x y = isOfHLevelDep→isOfHLevel (suc n) (h x y) refl - -isOfHLevel→isOfHLevelDep : (n : HLevel) - {A : Type } {B : A Type ℓ'} (h : (a : A) isOfHLevel n (B a)) isOfHLevelDep n {A = A} B -isOfHLevel→isOfHLevelDep 0 h {a} = - (h a .fst , λ b' p isProp→PathP i isContr→isProp (h (p i))) (h a .fst) b') -isOfHLevel→isOfHLevelDep 1 h = λ b0 b1 p isProp→PathP i h (p i)) b0 b1 -isOfHLevel→isOfHLevelDep (suc (suc n)) {A = A} {B} h {a0} {a1} b0 b1 = - isOfHLevel→isOfHLevelDep (suc n) p helper p) - where - helper : (p : a0 a1) - isOfHLevel (suc n) (PathP i B (p i)) b0 b1) - helper p = J a1 p b1 isOfHLevel (suc n) (PathP i B (p i)) b0 b1)) - _ h _ _ _) p b1 - -isContrDep→isPropDep : isOfHLevelDep 0 B isOfHLevelDep 1 B -isContrDep→isPropDep {B = B} Bctr {a0 = a0} b0 b1 p i - = comp k B (p (i k))) k λ where - (i = i0) Bctr .snd b0 refl k - (i = i1) Bctr .snd b1 p k) - (c0 .fst) - where - c0 = Bctr {a0} - -isPropDep→isSetDep : isOfHLevelDep 1 B isOfHLevelDep 2 B -isPropDep→isSetDep {B = B} Bprp b0 b1 b2 b3 p i j - = comp k B (p (i k) (j k))) k λ where - (j = i0) Bprp b0 b0 refl k - (i = i0) Bprp b0 (b2 j) k p i0 (j k)) k - (i = i1) Bprp b0 (b3 j) k p k (j k)) k - (j = i1) Bprp b0 b1 k p (i k) (j k)) k) - b0 - -isOfHLevelDepSuc : (n : HLevel) isOfHLevelDep n B isOfHLevelDep (suc n) B -isOfHLevelDepSuc 0 = isContrDep→isPropDep -isOfHLevelDepSuc 1 = isPropDep→isSetDep -isOfHLevelDepSuc (suc (suc n)) Blvl b0 b1 = isOfHLevelDepSuc (suc n) (Blvl b0 b1) - -isPropDep→isSetDep' - : isOfHLevelDep 1 B - {p : w x} {q : y z} {r : w y} {s : x z} - {tw : B w} {tx : B x} {ty : B y} {tz : B z} - (sq : Square p q r s) - (tp : PathP i B (p i)) tw tx) - (tq : PathP i B (q i)) ty tz) - (tr : PathP i B (r i)) tw ty) - (ts : PathP i B (s i)) tx tz) - SquareP i j B (sq i j)) tp tq tr ts -isPropDep→isSetDep' {B = B} Bprp {p} {q} {r} {s} {tw} sq tp tq tr ts i j - = comp k B (sq (i k) (j k))) k λ where - (i = i0) Bprp tw (tp j) k p (k j)) k - (i = i1) Bprp tw (tq j) k sq (i k) (j k)) k - (j = i0) Bprp tw (tr i) k r (k i)) k - (j = i1) Bprp tw (ts i) k sq (k i) (j k)) k) - tw - -isOfHLevelΣ' : n isOfHLevel n A isOfHLevelDep n B isOfHLevel n (Σ A B) -isOfHLevelΣ' 0 Actr Bctr .fst = (Actr .fst , Bctr .fst) -isOfHLevelΣ' 0 Actr Bctr .snd (x , y) i - = Actr .snd x i , Bctr .snd y (Actr .snd x) i -isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .fst = Alvl w x i -isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .snd = Blvl y z (Alvl w x) i -isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) - = isOfHLevelRetract (suc n) - p i p i .fst) , λ i p i .snd) - ΣPathP - x refl) - (isOfHLevelΣ' (suc n) (Alvl w x) (Blvl y z)) - -ΣSquareSet : ((x : A) isSet (B x)) {u v w x : Σ A B} - {p : u v} {q : v w} {r : x w} {s : u x} - Square (cong fst p) (cong fst r) - (cong fst s) (cong fst q) - Square p r s q -fst (ΣSquareSet pB sq i j) = sq i j -snd (ΣSquareSet {B = B} pB {p = p} {q = q} {r = r} {s = s} sq i j) = lem i j - where - lem : SquareP i j B (sq i j)) - (cong snd p) (cong snd r) (cong snd s) (cong snd q) - lem = toPathP (isOfHLevelPathP' 1 (pB _) _ _ _ _) - -module _ (isSet-A : isSet A) (isSet-A' : isSet A') where - - - isSet-SetsIso : isSet (Iso A A') - isSet-SetsIso x y p₀ p₁ = h - where - - module X = Iso x - module Y = Iso y - - f-p : i₁ (Iso.fun (p₀ i₁) , Iso.inv (p₀ i₁)) - (Iso.fun (p₁ i₁) , Iso.inv (p₁ i₁)) - fst (f-p i₁ i) a = isSet-A' (X.fun a ) (Y.fun a ) (cong _ p₀) (cong _ p₁) i i₁ - snd (f-p i₁ i) a' = isSet-A (X.inv a') (Y.inv a') (cong _ p₀) (cong _ p₁) i i₁ - - s-p : b _ - s-p b = - isSet→SquareP i j isProp→isSet (isSet-A' _ _)) - refl refl i₁ (Iso.rightInv (p₀ i₁) b)) i₁ (Iso.rightInv (p₁ i₁) b)) - - r-p : a _ - r-p a = - isSet→SquareP i j isProp→isSet (isSet-A _ _)) - refl refl i₁ (Iso.leftInv (p₀ i₁) a)) i₁ (Iso.leftInv (p₁ i₁) a)) - - - h : p₀ p₁ - Iso.fun (h i i₁) = fst (f-p i₁ i) - Iso.inv (h i i₁) = snd (f-p i₁ i) - Iso.rightInv (h i i₁) b = s-p b i₁ i - Iso.leftInv (h i i₁) a = r-p a i₁ i - - - SetsIso≡-ext : {a b : Iso A A'} - (∀ x Iso.fun a x Iso.fun b x) - (∀ x Iso.inv a x Iso.inv b x) - a b - Iso.fun (SetsIso≡-ext {a} {b} fun≡ inv≡ i) x = fun≡ x i - Iso.inv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) x = inv≡ x i - Iso.rightInv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) b₁ = - isSet→SquareP _ _ isSet-A') - (Iso.rightInv a b₁) - (Iso.rightInv b b₁) - i fun≡ (inv≡ b₁ i) i) - refl i - Iso.leftInv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) a₁ = - isSet→SquareP _ _ isSet-A) - (Iso.leftInv a a₁) - (Iso.leftInv b a₁) - i inv≡ (fun≡ a₁ i) i ) - refl i - - SetsIso≡ : {a b : Iso A A'} - (Iso.fun a Iso.fun b) - (Iso.inv a Iso.inv b) - a b - SetsIso≡ p q = - SetsIso≡-ext (funExt⁻ p) (funExt⁻ q) - - isSet→Iso-Iso-≃ : Iso (Iso A A') (A A') - isSet→Iso-Iso-≃ = ww - where - open Iso - - ww : Iso _ _ - fun ww = isoToEquiv - inv ww = equivToIso - rightInv ww b = equivEq refl - leftInv ww a = SetsIso≡ refl refl - - - isSet→isEquiv-isoToPath : isEquiv isoToEquiv - isSet→isEquiv-isoToPath = isoToIsEquiv isSet→Iso-Iso-≃ - - - -isSet→Iso-Iso-≡ : (isSet-A : isSet A) (isSet-A' : isSet A') Iso (Iso A A') (A A') -isSet→Iso-Iso-≡ isSet-A isSet-A' = ww - where - open Iso - - ww : Iso _ _ - fun ww = isoToPath - inv ww = pathToIso - rightInv ww b = isInjectiveTransport (funExt λ _ transportRefl _) - leftInv ww a = SetsIso≡-ext isSet-A isSet-A' _ transportRefl (fun a _)) λ _ cong (inv a) (transportRefl _) - -hSet-Iso-Iso-≡ : (A : hSet ) (A' : hSet ) Iso (Iso (fst A) (fst A')) (A A') -hSet-Iso-Iso-≡ A A' = compIso (isSet→Iso-Iso-≡ (snd A) (snd A')) (equivToIso (_ , isEquiv-Σ≡Prop λ _ isPropIsSet)) +isPropImplicitΠ3 : (h : (x : A) (y : B x) (z : C x y) isProp (D x y z)) + isProp ({x : A} {y : B x} {z : C x y} D x y z) +isPropImplicitΠ3 h = isPropImplicitΠ x isPropImplicitΠ2 y h x y)) + +isPropImplicitΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isProp (E x y z w)) + isProp ({x : A} {y : B x} {z : C x y} {w : D x y z} E x y z w) +isPropImplicitΠ4 h = isPropImplicitΠ x isPropImplicitΠ3 y h x y)) + +isProp→ : {A : Type } {B : Type ℓ'} isProp B isProp (A B) +isProp→ pB = isPropΠ λ _ pB + +isSetΠ : ((x : A) isSet (B x)) isSet ((x : A) B x) +isSetΠ = isOfHLevelΠ 2 + +isSetImplicitΠ : (h : (x : A) isSet (B x)) isSet ({x : A} B x) +isSetImplicitΠ h f g F G i j {x} = h x (f {x}) (g {x}) i F i {x}) i G i {x}) i j + +isSetImplicitΠ2 : (h : (x : A) (y : B x) isSet (C x y)) isSet ({x : A} {y : B x} C x y) +isSetImplicitΠ2 h = isSetImplicitΠ x isSetImplicitΠ y h x y)) + +isSetImplicitΠ3 : (h : (x : A) (y : B x) (z : C x y) isSet (D x y z)) + isSet ({x : A} {y : B x} {z : C x y} D x y z) +isSetImplicitΠ3 h = isSetImplicitΠ x isSetImplicitΠ2 y λ z h x y z)) + +isSet→ : isSet A' isSet (A A') +isSet→ isSet-A' = isOfHLevelΠ 2 _ isSet-A') + +isSetΠ2 : (h : (x : A) (y : B x) isSet (C x y)) + isSet ((x : A) (y : B x) C x y) +isSetΠ2 h = isSetΠ λ x isSetΠ λ y h x y + +isSetΠ3 : (h : (x : A) (y : B x) (z : C x y) isSet (D x y z)) + isSet ((x : A) (y : B x) (z : C x y) D x y z) +isSetΠ3 h = isSetΠ λ x isSetΠ λ y isSetΠ λ z h x y z + +isGroupoidΠ : ((x : A) isGroupoid (B x)) isGroupoid ((x : A) B x) +isGroupoidΠ = isOfHLevelΠ 3 + +isGroupoidΠ2 : (h : (x : A) (y : B x) isGroupoid (C x y)) isGroupoid ((x : A) (y : B x) C x y) +isGroupoidΠ2 h = isGroupoidΠ λ _ isGroupoidΠ λ _ h _ _ + +isGroupoidΠ3 : (h : (x : A) (y : B x) (z : C x y) isGroupoid (D x y z)) + isGroupoid ((x : A) (y : B x) (z : C x y) D x y z) +isGroupoidΠ3 h = isGroupoidΠ λ _ isGroupoidΠ2 λ _ h _ _ + +isGroupoidΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isGroupoid (E x y z w)) + isGroupoid ((x : A) (y : B x) (z : C x y) (w : D x y z) E x y z w) +isGroupoidΠ4 h = isGroupoidΠ λ _ isGroupoidΠ3 λ _ h _ _ + +is2GroupoidΠ : ((x : A) is2Groupoid (B x)) is2Groupoid ((x : A) B x) +is2GroupoidΠ = isOfHLevelΠ 4 + +isOfHLevelΠ⁻ : {A : Type } {B : Type ℓ'} n + isOfHLevel n (A B) (A isOfHLevel n B) +isOfHLevelΠ⁻ 0 h x = fst h x , λ y funExt⁻ (snd h (const y)) x +isOfHLevelΠ⁻ 1 h x y z = funExt⁻ (h (const y) (const z)) x +isOfHLevelΠ⁻ (suc (suc n)) h x y z = + isOfHLevelΠ⁻ (suc n) (isOfHLevelRetractFromIso (suc n) funExtIso (h _ _)) x + +isOfHLevel→∙ : {A : Pointed } {B : Pointed ℓ'} (n : ) + isOfHLevel n (fst B) isOfHLevel n (A →∙ B) +isOfHLevel→∙ n hlev = + isOfHLevelΣ n (isOfHLevelΠ n _ hlev)) + λ _ isOfHLevelPath n hlev _ _ + +-- h-level of A ≃ B and A ≡ B + +isOfHLevel≃ + : n {A : Type } {B : Type ℓ'} + (hA : isOfHLevel n A) (hB : isOfHLevel n B) isOfHLevel n (A B) +isOfHLevel≃ zero {A = A} {B = B} hA hB = isContr→Equiv hA hB , contr + where + contr : (y : A B) isContr→Equiv hA hB y + contr y = Σ≡Prop isPropIsEquiv (funExt a snd hB (fst y a))) + +isOfHLevel≃ (suc n) {A = A} {B = B} hA hB = + isOfHLevelΣ (suc n) (isOfHLevelΠ _ λ _ hB) + f isProp→isOfHLevelSuc n (isPropIsEquiv f)) + +isOfHLevel≡ : n {A B : Type } (hA : isOfHLevel n A) (hB : isOfHLevel n B) + isOfHLevel n (A B) +isOfHLevel≡ n hA hB = isOfHLevelRetractFromIso n univalenceIso (isOfHLevel≃ n hA hB) + +isOfHLevel⁺≃ₗ + : n {A : Type } {B : Type ℓ'} + isOfHLevel (suc n) A isOfHLevel (suc n) (A B) +isOfHLevel⁺≃ₗ zero pA e = isOfHLevel≃ 1 pA (isOfHLevelRespectEquiv 1 e pA) e +isOfHLevel⁺≃ₗ (suc n) hA e = isOfHLevel≃ m hA (isOfHLevelRespectEquiv m e hA) e + where + m = suc (suc n) + +isOfHLevel⁺≃ᵣ + : n {A : Type } {B : Type ℓ'} + isOfHLevel (suc n) B isOfHLevel (suc n) (A B) +isOfHLevel⁺≃ᵣ zero pB e + = isOfHLevel≃ 1 (isPropRetract (e .fst) (invEq e) (retEq e) pB) pB e +isOfHLevel⁺≃ᵣ (suc n) hB e + = isOfHLevel≃ m (isOfHLevelRetract m (e .fst) (invEq e) (retEq e) hB) hB e + where + m = suc (suc n) + +isOfHLevel⁺≡ₗ + : n {A B : Type } + isOfHLevel (suc n) A isOfHLevel (suc n) (A B) +isOfHLevel⁺≡ₗ zero pA P = isOfHLevel≡ 1 pA (subst isProp P pA) P +isOfHLevel⁺≡ₗ (suc n) hA P + = isOfHLevel≡ m hA (subst (isOfHLevel m) P hA) P + where + m = suc (suc n) + +isOfHLevel⁺≡ᵣ + : n {A B : Type } + isOfHLevel (suc n) B isOfHLevel (suc n) (A B) +isOfHLevel⁺≡ᵣ zero pB P = isOfHLevel≡ 1 (subst⁻ isProp P pB) pB P +isOfHLevel⁺≡ᵣ (suc n) hB P + = isOfHLevel≡ m (subst⁻ (isOfHLevel m) P hB) hB P + where + m = suc (suc n) + +-- h-level of TypeOfHLevel + +isPropHContr : isProp (TypeOfHLevel 0) +isPropHContr x y = Σ≡Prop _ isPropIsContr) (isOfHLevel≡ 0 (x .snd) (y .snd) .fst) + +isOfHLevelTypeOfHLevel : n isOfHLevel (suc n) (TypeOfHLevel n) +isOfHLevelTypeOfHLevel zero = isPropHContr +isOfHLevelTypeOfHLevel (suc n) (X , a) (Y , b) = + isOfHLevelRetract (suc n) (cong fst) (Σ≡Prop λ _ isPropIsOfHLevel (suc n)) + (section-Σ≡Prop λ _ isPropIsOfHLevel (suc n)) + (isOfHLevel≡ (suc n) a b) + +isSetHProp : isSet (hProp ) +isSetHProp = isOfHLevelTypeOfHLevel 1 + +isGroupoidHSet : isGroupoid (hSet ) +isGroupoidHSet = isOfHLevelTypeOfHLevel 2 + + +-- h-level of lifted type + +isOfHLevelLift : { ℓ'} (n : HLevel) {A : Type } isOfHLevel n A isOfHLevel n (Lift {j = ℓ'} A) +isOfHLevelLift n = isOfHLevelRetract n lower lift λ _ refl + +isOfHLevelLower : { ℓ'} (n : HLevel) {A : Type } isOfHLevel n (Lift {j = ℓ'} A) isOfHLevel n A +isOfHLevelLower n = isOfHLevelRetract n lift lower λ _ refl + +---------------------------- + +-- More consequences of isProp and isContr + +inhProp→isContr : A isProp A isContr A +inhProp→isContr x h = x , h x + +extend : isContr A (∀ φ (u : Partial φ A) Sub A φ u) +extend (x , p) φ u = inS (hcomp { j (φ = i1) p (u 1=1) j }) x) + +isContrPartial→isContr : {} {A : Type } + (extend : φ Partial φ A A) + (∀ u u (extend i1 λ { _ u})) + isContr A +isContrPartial→isContr {A = A} extend law + = ex , λ y law ex i Aux.v y i) sym (law y) + where ex = extend i0 empty + module Aux (y : A) (i : I) where + φ = ~ i i + u : Partial φ A + u = λ { (i = i0) ex ; (i = i1) y } + v = extend φ u + +-- Dependent h-level over a type + +isOfHLevelDep : HLevel {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') +isOfHLevelDep 0 {A = A} B = {a : A} Σ[ b B a ] ({a' : A} (b' : B a') (p : a a') PathP i B (p i)) b b') +isOfHLevelDep 1 {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) (p : a0 a1) PathP i B (p i)) b0 b1 +isOfHLevelDep (suc (suc n)) {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) isOfHLevelDep (suc n) {A = a0 a1} p PathP i B (p i)) b0 b1) + +isContrDep : {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') +isContrDep = isOfHLevelDep 0 + +isPropDep : {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') +isPropDep = isOfHLevelDep 1 + +isContrDep∘ + : {A' : Type } (f : A' A) isContrDep B isContrDep (B f) +isContrDep∘ f cB {a} = λ where + .fst cB .fst + .snd b' p cB .snd b' (cong f p) + +isPropDep∘ : {A' : Type } (f : A' A) isPropDep B isPropDep (B f) +isPropDep∘ f pB b0 b1 = pB b0 b1 cong f + +isOfHLevelDep→isOfHLevel : (n : HLevel) + {A : Type } {B : A Type ℓ'} isOfHLevelDep n {A = A} B (a : A) isOfHLevel n (B a) +isOfHLevelDep→isOfHLevel 0 h a = h .fst , λ b h .snd b refl +isOfHLevelDep→isOfHLevel 1 h a x y = h x y refl +isOfHLevelDep→isOfHLevel (suc (suc n)) h a x y = isOfHLevelDep→isOfHLevel (suc n) (h x y) refl + +isOfHLevel→isOfHLevelDep : (n : HLevel) + {A : Type } {B : A Type ℓ'} (h : (a : A) isOfHLevel n (B a)) isOfHLevelDep n {A = A} B +isOfHLevel→isOfHLevelDep 0 h {a} = + (h a .fst , λ b' p isProp→PathP i isContr→isProp (h (p i))) (h a .fst) b') +isOfHLevel→isOfHLevelDep 1 h = λ b0 b1 p isProp→PathP i h (p i)) b0 b1 +isOfHLevel→isOfHLevelDep (suc (suc n)) {A = A} {B} h {a0} {a1} b0 b1 = + isOfHLevel→isOfHLevelDep (suc n) p helper p) + where + helper : (p : a0 a1) + isOfHLevel (suc n) (PathP i B (p i)) b0 b1) + helper p = J a1 p b1 isOfHLevel (suc n) (PathP i B (p i)) b0 b1)) + _ h _ _ _) p b1 + +isContrDep→isPropDep : isOfHLevelDep 0 B isOfHLevelDep 1 B +isContrDep→isPropDep {B = B} Bctr {a0 = a0} b0 b1 p i + = comp k B (p (i k))) k λ where + (i = i0) Bctr .snd b0 refl k + (i = i1) Bctr .snd b1 p k) + (c0 .fst) + where + c0 = Bctr {a0} + +isPropDep→isSetDep : isOfHLevelDep 1 B isOfHLevelDep 2 B +isPropDep→isSetDep {B = B} Bprp b0 b1 b2 b3 p i j + = comp k B (p (i k) (j k))) k λ where + (j = i0) Bprp b0 b0 refl k + (i = i0) Bprp b0 (b2 j) k p i0 (j k)) k + (i = i1) Bprp b0 (b3 j) k p k (j k)) k + (j = i1) Bprp b0 b1 k p (i k) (j k)) k) + b0 + +isOfHLevelDepSuc : (n : HLevel) isOfHLevelDep n B isOfHLevelDep (suc n) B +isOfHLevelDepSuc 0 = isContrDep→isPropDep +isOfHLevelDepSuc 1 = isPropDep→isSetDep +isOfHLevelDepSuc (suc (suc n)) Blvl b0 b1 = isOfHLevelDepSuc (suc n) (Blvl b0 b1) + +isPropDep→isSetDep' + : isOfHLevelDep 1 B + {p : w x} {q : y z} {r : w y} {s : x z} + {tw : B w} {tx : B x} {ty : B y} {tz : B z} + (sq : Square p q r s) + (tp : PathP i B (p i)) tw tx) + (tq : PathP i B (q i)) ty tz) + (tr : PathP i B (r i)) tw ty) + (ts : PathP i B (s i)) tx tz) + SquareP i j B (sq i j)) tp tq tr ts +isPropDep→isSetDep' {B = B} Bprp {p} {q} {r} {s} {tw} sq tp tq tr ts i j + = comp k B (sq (i k) (j k))) k λ where + (i = i0) Bprp tw (tp j) k p (k j)) k + (i = i1) Bprp tw (tq j) k sq (i k) (j k)) k + (j = i0) Bprp tw (tr i) k r (k i)) k + (j = i1) Bprp tw (ts i) k sq (k i) (j k)) k) + tw + +isOfHLevelΣ' : n isOfHLevel n A isOfHLevelDep n B isOfHLevel n (Σ A B) +isOfHLevelΣ' 0 Actr Bctr .fst = (Actr .fst , Bctr .fst) +isOfHLevelΣ' 0 Actr Bctr .snd (x , y) i + = Actr .snd x i , Bctr .snd y (Actr .snd x) i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .fst = Alvl w x i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .snd = Blvl y z (Alvl w x) i +isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) + = isOfHLevelRetract (suc n) + p i p i .fst) , λ i p i .snd) + ΣPathP + x refl) + (isOfHLevelΣ' (suc n) (Alvl w x) (Blvl y z)) + +ΣSquareSet : ((x : A) isSet (B x)) {u v w x : Σ A B} + {p : u v} {q : v w} {r : x w} {s : u x} + Square (cong fst p) (cong fst r) + (cong fst s) (cong fst q) + Square p r s q +fst (ΣSquareSet pB sq i j) = sq i j +snd (ΣSquareSet {B = B} pB {p = p} {q = q} {r = r} {s = s} sq i j) = lem i j + where + lem : SquareP i j B (sq i j)) + (cong snd p) (cong snd r) (cong snd s) (cong snd q) + lem = toPathP (isOfHLevelPathP' 1 (pB _) _ _ _ _) + +module _ (isSet-A : isSet A) (isSet-A' : isSet A') where + + + isSet-SetsIso : isSet (Iso A A') + isSet-SetsIso x y p₀ p₁ = h + where + + module X = Iso x + module Y = Iso y + + f-p : i₁ (Iso.fun (p₀ i₁) , Iso.inv (p₀ i₁)) + (Iso.fun (p₁ i₁) , Iso.inv (p₁ i₁)) + fst (f-p i₁ i) a = isSet-A' (X.fun a ) (Y.fun a ) (cong _ p₀) (cong _ p₁) i i₁ + snd (f-p i₁ i) a' = isSet-A (X.inv a') (Y.inv a') (cong _ p₀) (cong _ p₁) i i₁ + + s-p : b _ + s-p b = + isSet→SquareP i j isProp→isSet (isSet-A' _ _)) + refl refl i₁ (Iso.rightInv (p₀ i₁) b)) i₁ (Iso.rightInv (p₁ i₁) b)) + + r-p : a _ + r-p a = + isSet→SquareP i j isProp→isSet (isSet-A _ _)) + refl refl i₁ (Iso.leftInv (p₀ i₁) a)) i₁ (Iso.leftInv (p₁ i₁) a)) + + + h : p₀ p₁ + Iso.fun (h i i₁) = fst (f-p i₁ i) + Iso.inv (h i i₁) = snd (f-p i₁ i) + Iso.rightInv (h i i₁) b = s-p b i₁ i + Iso.leftInv (h i i₁) a = r-p a i₁ i + + + SetsIso≡-ext : {a b : Iso A A'} + (∀ x Iso.fun a x Iso.fun b x) + (∀ x Iso.inv a x Iso.inv b x) + a b + Iso.fun (SetsIso≡-ext {a} {b} fun≡ inv≡ i) x = fun≡ x i + Iso.inv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) x = inv≡ x i + Iso.rightInv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) b₁ = + isSet→SquareP _ _ isSet-A') + (Iso.rightInv a b₁) + (Iso.rightInv b b₁) + i fun≡ (inv≡ b₁ i) i) + refl i + Iso.leftInv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) a₁ = + isSet→SquareP _ _ isSet-A) + (Iso.leftInv a a₁) + (Iso.leftInv b a₁) + i inv≡ (fun≡ a₁ i) i ) + refl i + + SetsIso≡ : {a b : Iso A A'} + (Iso.fun a Iso.fun b) + (Iso.inv a Iso.inv b) + a b + SetsIso≡ p q = + SetsIso≡-ext (funExt⁻ p) (funExt⁻ q) + + isSet→Iso-Iso-≃ : Iso (Iso A A') (A A') + isSet→Iso-Iso-≃ = ww + where + open Iso + + ww : Iso _ _ + fun ww = isoToEquiv + inv ww = equivToIso + rightInv ww b = equivEq refl + leftInv ww a = SetsIso≡ refl refl + + + isSet→isEquiv-isoToPath : isEquiv isoToEquiv + isSet→isEquiv-isoToPath = isoToIsEquiv isSet→Iso-Iso-≃ + + + +isSet→Iso-Iso-≡ : (isSet-A : isSet A) (isSet-A' : isSet A') Iso (Iso A A') (A A') +isSet→Iso-Iso-≡ isSet-A isSet-A' = ww + where + open Iso + + ww : Iso _ _ + fun ww = isoToPath + inv ww = pathToIso + rightInv ww b = isInjectiveTransport (funExt λ _ transportRefl _) + leftInv ww a = SetsIso≡-ext isSet-A isSet-A' _ transportRefl (fun a _)) λ _ cong (inv a) (transportRefl _) + +hSet-Iso-Iso-≡ : (A : hSet ) (A' : hSet ) Iso (Iso (fst A) (fst A')) (A A') +hSet-Iso-Iso-≡ A A' = compIso (isSet→Iso-Iso-≡ (snd A) (snd A')) (equivToIso (_ , isEquiv-Σ≡Prop λ _ isPropIsSet)) + +module _ (B : (i j k : I) Type ) + {c₀₀₀ : B i0 i0 i0} {c₀₀₁ : B i0 i0 i1} {c₀₁₀ : B i0 i1 i0} {c₀₁₁ : B i0 i1 i1} + {c₁₀₀ : B i1 i0 i0} {c₁₀₁ : B i1 i0 i1} {c₁₁₀ : B i1 i1 i0} {c₁₁₁ : B i1 i1 i1} + {c₀₀₋ : PathP k B i0 i0 k) c₀₀₀ c₀₀₁} {c₀₁₋ : PathP k B i0 i1 k) c₀₁₀ c₀₁₁} + {c₀₋₀ : PathP i B i0 i i0) c₀₀₀ c₀₁₀} {c₀₋₁ : PathP i B i0 i i1) c₀₀₁ c₀₁₁} + {c₁₀₋ : PathP k B i1 i0 k) c₁₀₀ c₁₀₁} {c₁₁₋ : PathP k B i1 i1 k) c₁₁₀ c₁₁₁} + {c₁₋₀ : PathP i B i1 i i0) c₁₀₀ c₁₁₀} {c₁₋₁ : PathP i B i1 i i1) c₁₀₁ c₁₁₁} + {c₋₀₀ : PathP i B i i0 i0) c₀₀₀ c₁₀₀} {c₋₀₁ : PathP i B i i0 i1) c₀₀₁ c₁₀₁} + {c₋₁₀ : PathP i B i i1 i0) c₀₁₀ c₁₁₀} {c₋₁₁ : PathP i B i i1 i1) c₀₁₁ c₁₁₁} + (c₀₋₋ : SquareP j k B i0 j k) c₀₀₋ c₀₁₋ c₀₋₀ c₀₋₁) + (c₁₋₋ : SquareP j k B i1 j k) c₁₀₋ c₁₁₋ c₁₋₀ c₁₋₁) + (c₋₀₋ : SquareP i k B i i0 k) c₀₀₋ c₁₀₋ c₋₀₀ c₋₀₁) + (c₋₁₋ : SquareP i k B i i1 k) c₀₁₋ c₁₁₋ c₋₁₀ c₋₁₁) + (c₋₋₀ : SquareP i j B i j i0) c₀₋₀ c₁₋₀ c₋₀₀ c₋₁₀) + (c₋₋₁ : SquareP i j B i j i1) c₀₋₁ c₁₋₁ c₋₀₁ c₋₁₁) where + + CubeP : Type + CubeP = PathP i SquareP j k B i j k) + (c₋₀₋ i) (c₋₁₋ i) + (c₋₋₀ i) (c₋₋₁ i)) + c₀₋₋ c₁₋₋ + + isGroupoid→CubeP : isGroupoid (B i1 i1 i1) CubeP + isGroupoid→CubeP grpd = + isOfHLevelPathP' 0 (isOfHLevelPathP' 1 (isOfHLevelPathP' 2 grpd _ _) _ _) _ _ .fst \ No newline at end of file diff --git a/docs/Cubical.Foundations.Isomorphism.html b/docs/Cubical.Foundations.Isomorphism.html index b9bd2d5..14f72f6 100644 --- a/docs/Cubical.Foundations.Isomorphism.html +++ b/docs/Cubical.Foundations.Isomorphism.html @@ -149,25 +149,25 @@ rightInv (compIsoIdR isom i) b = rUnit (isom .rightInv b) (~ i) leftInv (compIsoIdR isom i) a = lUnit (isom .leftInv a) (~ i) -LiftIso : Iso A (Lift {i = } {j = ℓ'} A) -fun LiftIso = lift -inv LiftIso = lower +LiftIso : Iso A (Lift {i = } {j = ℓ'} A) +fun LiftIso = lift +inv LiftIso = lower rightInv LiftIso _ = refl leftInv LiftIso _ = refl -isContr→Iso : isContr A isContr B Iso A B +isContr→Iso : isContr A isContr B Iso A B fun (isContr→Iso _ Bctr) _ = Bctr .fst inv (isContr→Iso Actr _) _ = Actr .fst rightInv (isContr→Iso _ Bctr) = Bctr .snd leftInv (isContr→Iso Actr _) = Actr .snd -isContr→Iso' : isContr A isContr B (A B) Iso A B +isContr→Iso' : isContr A isContr B (A B) Iso A B fun (isContr→Iso' _ Bctr f) = f inv (isContr→Iso' Actr _ _) _ = Actr .fst -rightInv (isContr→Iso' _ Bctr f) = isContr→isProp Bctr _ +rightInv (isContr→Iso' _ Bctr f) = isContr→isProp Bctr _ leftInv (isContr→Iso' Actr _ _) = Actr .snd -isProp→Iso : (Aprop : isProp A) (Bprop : isProp B) (f : A B) (g : B A) Iso A B +isProp→Iso : (Aprop : isProp A) (Bprop : isProp B) (f : A B) (g : B A) Iso A B fun (isProp→Iso _ _ f _) = f inv (isProp→Iso _ _ _ g) = g rightInv (isProp→Iso _ Bprop f g) b = Bprop (f (g b)) b @@ -208,14 +208,20 @@ binaryOpIso : Iso A B Iso (A A A) (B B B) binaryOpIso is = compIso (domIso is) (codomainIso (endoIso is)) -Iso≡Set : isSet A isSet B (f g : Iso A B) +Iso≡Set : isSet A isSet B (f g : Iso A B) ((x : A) f .fun x g .fun x) ((x : B) f .inv x g .inv x) f g fun (Iso≡Set hA hB f g hfun hinv i) x = hfun x i inv (Iso≡Set hA hB f g hfun hinv i) x = hinv x i rightInv (Iso≡Set hA hB f g hfun hinv i) x j = - isSet→isSet' hB (rightInv f x) (rightInv g x) i hfun (hinv x i) i) refl i j + isSet→isSet' hB (rightInv f x) (rightInv g x) i hfun (hinv x i) i) refl i j leftInv (Iso≡Set hA hB f g hfun hinv i) x j = - isSet→isSet' hA (leftInv f x) (leftInv g x) i hinv (hfun x i) i) refl i j + isSet→isSet' hA (leftInv f x) (leftInv g x) i hinv (hfun x i) i) refl i j + +transportIsoToPath : (f : Iso A B) (x : A) transport (isoToPath f) x f .fun x +transportIsoToPath f x = transportRefl _ + +transportIsoToPath⁻ : (f : Iso A B) (x : B) transport (sym (isoToPath f)) x f .inv x +transportIsoToPath⁻ f x = cong (f .inv) (transportRefl _) \ No newline at end of file diff --git a/docs/Cubical.Foundations.Path.html b/docs/Cubical.Foundations.Path.html index 87ee14d..30857c1 100644 --- a/docs/Cubical.Foundations.Path.html +++ b/docs/Cubical.Foundations.Path.html @@ -19,10 +19,10 @@ module _ {A : I Type } {x : A i0} {y : A i1} where toPathP⁻ : x transport⁻ i A i) y PathP A x y - toPathP⁻ p = symP (toPathP (sym p)) + toPathP⁻ p = symP (toPathP (sym p)) fromPathP⁻ : PathP A x y x transport⁻ i A i) y - fromPathP⁻ p = sym (fromPathP {A = λ i A (~ i)} (symP p)) + fromPathP⁻ p = sym (fromPathP {A = λ i A (~ i)} (symP p)) PathP≡Path : (P : I Type ) (p : P i0) (q : P i1) PathP P p q Path (P i1) (transport i P i) p) q @@ -33,8 +33,8 @@ PathP≡Path⁻ P p q i = PathP j P (~ i j)) p (transport⁻-filler j P j) q i) PathPIsoPath : (A : I Type ) (x : A i0) (y : A i1) Iso (PathP A x y) (transport i A i) x y) -PathPIsoPath A x y .Iso.fun = fromPathP -PathPIsoPath A x y .Iso.inv = toPathP +PathPIsoPath A x y .Iso.fun = fromPathP +PathPIsoPath A x y .Iso.inv = toPathP PathPIsoPath A x y .Iso.rightInv q k i = hcomp j λ @@ -89,11 +89,11 @@ PathP≡compPath p q r k = PathP i p i0 q (i k)) j compPath-filler p q k j) r -- a quick corollary for 3-constant functions -3-ConstantCompChar : {A : Type } {B : Type ℓ'} (f : A B) (link : 2-Constant f) +3-ConstantCompChar : {A : Type } {B : Type ℓ'} (f : A B) (link : 2-Constant f) (∀ x y z link x y link y z link x z) - 3-Constant f -3-Constant.link (3-ConstantCompChar f link coh₂) = link -3-Constant.coh₁ (3-ConstantCompChar f link coh₂) _ _ _ = + 3-Constant f +3-Constant.link (3-ConstantCompChar f link coh₂) = link +3-Constant.coh₁ (3-ConstantCompChar f link coh₂) _ _ _ = transport⁻ (PathP≡compPath _ _ _) (coh₂ _ _ _) PathP≡doubleCompPathˡ : {A : Type } {w x y z : A} (p : w y) (q : w x) (r : y z) (s : x z) @@ -129,11 +129,11 @@ compPathrEquiv p = (_∙ p) , compPathr-isEquiv p -- Variations of isProp→isSet for PathP -isProp→SquareP : {B : I I Type } ((i j : I) isProp (B i j)) +isProp→SquareP : {B : I I Type } ((i j : I) isProp (B i j)) {a : B i0 i0} {b : B i0 i1} {c : B i1 i0} {d : B i1 i1} (r : PathP j B j i0) a c) (s : PathP j B j i1) b d) (t : PathP j B i0 j) a b) (u : PathP j B i1 j) c d) - SquareP B t u r s + SquareP B t u r s isProp→SquareP {B = B} isPropB {a = a} r s t u i j = hcomp { k (i = i0) isPropB i0 j (base i0 j) (t j) k ; k (i = i1) isPropB i1 j (base i1 j) (u j) k @@ -143,32 +143,32 @@ base : (i j : I) B i j base i j = transport k B (i k) (j k)) a -isProp→isPropPathP : {} {B : I Type } ((i : I) isProp (B i)) +isProp→isPropPathP : {} {B : I Type } ((i : I) isProp (B i)) (b0 : B i0) (b1 : B i1) - isProp (PathP i B i) b0 b1) + isProp (PathP i B i) b0 b1) isProp→isPropPathP {B = B} hB b0 b1 = isProp→SquareP _ hB) refl refl -isProp→isContrPathP : {A : I Type } (∀ i isProp (A i)) +isProp→isContrPathP : {A : I Type } (∀ i isProp (A i)) (x : A i0) (y : A i1) - isContr (PathP A x y) -isProp→isContrPathP h x y = isProp→PathP h x y , isProp→isPropPathP h x y _ + isContr (PathP A x y) +isProp→isContrPathP h x y = isProp→PathP h x y , isProp→isPropPathP h x y _ -- Flipping a square along its diagonal flipSquare : {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ a₀₁} {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ a₁₁} {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} - Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ + Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ flipSquare sq i j = sq j i module _ {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ a₀₁} {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ a₁₁} {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} where - flipSquareEquiv : Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ + flipSquareEquiv : Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ unquoteDef flipSquareEquiv = defStrictEquiv flipSquareEquiv flipSquare flipSquare - flipSquarePath : Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ + flipSquarePath : Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ flipSquarePath = ua flipSquareEquiv module _ {a₀₀ a₁₁ : A} {a₋ : a₀₀ a₁₁} @@ -180,12 +180,12 @@ slideSquareFaces i j k (j = i0) = a₋₀ i slideSquareFaces i j k (j = i1) = a₋ (i ~ k) - slideSquare : Square a₋ a₁₋ a₋₀ refl Square refl a₁₋ a₋₀ a₋ + slideSquare : Square a₋ a₁₋ a₋₀ refl Square refl a₁₋ a₋₀ a₋ slideSquare sq i j = hcomp (slideSquareFaces i j) (sq i j) - slideSquareEquiv : (Square a₋ a₁₋ a₋₀ refl) (Square refl a₁₋ a₋₀ a₋) + slideSquareEquiv : (Square a₋ a₁₋ a₋₀ refl) (Square refl a₁₋ a₋₀ a₋) slideSquareEquiv = isoToEquiv (iso slideSquare slideSquareInv fillerTo fillerFrom) where - slideSquareInv : Square refl a₁₋ a₋₀ a₋ Square a₋ a₁₋ a₋₀ refl + slideSquareInv : Square refl a₁₋ a₋₀ a₋ Square a₋ a₁₋ a₋₀ refl slideSquareInv sq i j = hcomp k slideSquareFaces i j (~ k)) (sq i j) fillerTo : p slideSquare (slideSquareInv p) p fillerTo p k i j = hcomp-equivFiller k slideSquareFaces i j (~ k)) (inS (p i j)) (~ k) @@ -196,11 +196,11 @@ Square≃doubleComp : {a₀₀ a₀₁ a₁₀ a₁₁ : A} (a₀₋ : a₀₀ a₀₁) (a₁₋ : a₁₀ a₁₁) (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) - Square a₀₋ a₁₋ a₋₀ a₋₁ (a₋₀ ⁻¹ ∙∙ a₀₋ ∙∙ a₋₁ a₁₋) + Square a₀₋ a₁₋ a₋₀ a₋₁ (a₋₀ ⁻¹ ∙∙ a₀₋ ∙∙ a₋₁ a₁₋) Square≃doubleComp a₀₋ a₁₋ a₋₀ a₋₁ = pathToEquiv (PathP≡doubleCompPathˡ a₋₀ a₀₋ a₁₋ a₋₁) -- Flipping a square in Ω²A is the same as inverting it -sym≡flipSquare : {x : A} (P : Square (refl {x = x}) refl refl refl) +sym≡flipSquare : {x : A} (P : Square (refl {x = x}) refl refl refl) sym P flipSquare P sym≡flipSquare {x = x} P = sym (main refl P) where @@ -208,10 +208,10 @@ B q i = PathP j x q (i j)) k q (i k)) refl main : (q : x x) (p : refl q) PathP i B q i) i j p j i) (sym p) - main q = J q p PathP i B q i) i j p j i) (sym p)) refl + main q = J q p PathP i B q i) i j p j i) (sym p)) refl -- Inverting both interval arguments of a square in Ω²A is the same as doing nothing -sym-cong-sym≡id : {x : A} (P : Square (refl {x = x}) refl refl refl) +sym-cong-sym≡id : {x : A} (P : Square (refl {x = x}) refl refl refl) P λ i j P (~ i) (~ j) sym-cong-sym≡id {x = x} P = sym (main refl P) where @@ -219,15 +219,15 @@ B q i = Path (x q i) j q (i ~ j)) λ j q (i j) main : (q : x x) (p : refl q) PathP i B q i) i j p (~ i) (~ j)) p - main q = J q p PathP i B q i) i j p (~ i) (~ j)) p) refl + main q = J q p PathP i B q i) i j p (~ i) (~ j)) p) refl -- Applying cong sym is the same as flipping a square in Ω²A -flipSquare≡cong-sym : {} {A : Type } {x : A} (P : Square (refl {x = x}) refl refl refl) +flipSquare≡cong-sym : {} {A : Type } {x : A} (P : Square (refl {x = x}) refl refl refl) flipSquare P λ i j P i (~ j) flipSquare≡cong-sym P = sym (sym≡flipSquare P) sym (sym-cong-sym≡id (cong sym P)) -- Applying cong sym is the same as inverting a square in Ω²A -sym≡cong-sym : {} {A : Type } {x : A} (P : Square (refl {x = x}) refl refl refl) +sym≡cong-sym : {} {A : Type } {x : A} (P : Square (refl {x = x}) refl refl refl) sym P cong sym P sym≡cong-sym P = sym-cong-sym≡id (sym P) @@ -247,13 +247,13 @@ where -- "Pointwise" composition - _∙v_ : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) - Square a₀₋ a₂₋ (a₋₀ b₋₀) (a₋₁ b₋₁) + _∙v_ : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) + Square a₀₋ a₂₋ (a₋₀ b₋₀) (a₋₁ b₋₁) (p ∙v q) i j = ((λ i p i j) i q i j)) i -- "Direct" composition - _∙v'_ : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) - Square a₀₋ a₂₋ (a₋₀ b₋₀) (a₋₁ b₋₁) + _∙v'_ : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) + Square a₀₋ a₂₋ (a₋₀ b₋₀) (a₋₁ b₋₁) (p ∙v' q) i = comp k compPath-filler a₋₀ b₋₀ k i compPath-filler a₋₁ b₋₁ k i) where k (i = i0) a₀₋ @@ -262,7 +262,7 @@ -- The two ways of composing squares are equal, because they are -- correct "lids" for the same box - ∙v≡∙v' : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) + ∙v≡∙v' : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) p ∙v q p ∙v' q ∙v≡∙v' p q l i = outS (comp-unique {A = λ k compPath-filler a₋₀ b₋₀ k i compPath-filler a₋₁ b₋₁ k i} @@ -288,13 +288,13 @@ Jequiv P = isoToEquiv isom where isom : Iso _ _ - Iso.fun isom = J P + Iso.fun isom = J P Iso.inv isom f = f refl Iso.rightInv isom f = implicitFunExt λ {_} funExt λ t - J _ t J P (f refl) t f t) (JRefl P (f refl)) t - Iso.leftInv isom = JRefl P + J _ t J P (f refl) t f t) (JRefl P (f refl)) t + Iso.leftInv isom = JRefl P -- Action of PathP on equivalences (without relying on univalence) @@ -305,36 +305,36 @@ congPathIso {A = A} {B} e {a₀} {a₁} .Iso.inv q i = hcomp j λ - { (i = i0) retEq (e i0) a₀ j - ; (i = i1) retEq (e i1) a₁ j + { (i = i0) retEq (e i0) a₀ j + ; (i = i1) retEq (e i1) a₁ j }) - (invEq (e i) (q i)) + (invEq (e i) (q i)) congPathIso {A = A} {B} e {a₀} {a₁} .Iso.rightInv q k i = hcomp j λ - { (i = i0) commSqIsEq (e i0 .snd) a₀ j k - ; (i = i1) commSqIsEq (e i1 .snd) a₁ j k + { (i = i0) commSqIsEq (e i0 .snd) a₀ j k + ; (i = i1) commSqIsEq (e i1 .snd) a₁ j k ; (k = i0) e i .fst (hfill j λ - { (i = i0) retEq (e i0) a₀ j - ; (i = i1) retEq (e i1) a₁ j + { (i = i0) retEq (e i0) a₀ j + ; (i = i1) retEq (e i1) a₁ j }) - (inS (invEq (e i) (q i))) + (inS (invEq (e i) (q i))) j) ; (k = i1) q i }) - (secEq (e i) (q i) k) - where b = commSqIsEq + (secEq (e i) (q i) k) + where b = commSqIsEq congPathIso {A = A} {B} e {a₀} {a₁} .Iso.leftInv p k i = hcomp j λ - { (i = i0) retEq (e i0) a₀ (j k) - ; (i = i1) retEq (e i1) a₁ (j k) + { (i = i0) retEq (e i0) a₀ (j k) + ; (i = i1) retEq (e i1) a₁ (j k) ; (k = i1) p i }) - (retEq (e i) (p i) k) + (retEq (e i) (p i) k) congPathEquiv : { ℓ'} {A : I Type } {B : I Type ℓ'} (e : i A i B i) {a₀ : A i0} {a₁ : A i1} @@ -420,16 +420,16 @@ (j = i1) compPath-filler' r q (~ k) i compPath→Square : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} - p s r q Square r s p q + p s r q Square r s p q compPath→Square {p = p} {q = q} {r = r} {s = s} P i j = hcomp (compPath→Square-faces p q r s i j) (P j i) Square→compPath : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} - Square r s p q p s r q + Square r s p q p s r q Square→compPath {p = p} {q = q} {r = r} {s = s} sq i j = hcomp k compPath→Square-faces p q r s j i (~ k)) (sq j i) -Square→compPathΩ² : {a : A} (sq : Square _ a) refl refl refl) +Square→compPathΩ² : {a : A} (sq : Square _ a) refl refl refl) Square→compPath sq cong (_∙ refl) (flipSquare sq) Square→compPathΩ² {a = a} sq k i j = hcomp r λ {(i = i0) rUnit _ a) r j diff --git a/docs/Cubical.Foundations.Pointed.Base.html b/docs/Cubical.Foundations.Pointed.Base.html index bbdb4dd..3d1120d 100644 --- a/docs/Cubical.Foundations.Pointed.Base.html +++ b/docs/Cubical.Foundations.Pointed.Base.html @@ -48,13 +48,13 @@ snd (≃∙map e) = snd e invEquiv∙ : {A : Pointed } {B : Pointed ℓ'} A ≃∙ B B ≃∙ A -fst (invEquiv∙ x) = invEquiv (fst x) +fst (invEquiv∙ x) = invEquiv (fst x) snd (invEquiv∙ {A = A} x) = - sym (cong (fst (invEquiv (fst x))) (snd x)) retEq (fst x) (pt A) + sym (cong (fst (invEquiv (fst x))) (snd x)) retEq (fst x) (pt A) compEquiv∙ : { ℓ' ℓ''} {A : Pointed } {B : Pointed ℓ'} {C : Pointed ℓ''} A ≃∙ B B ≃∙ C A ≃∙ C -fst (compEquiv∙ e1 e2) = compEquiv (fst e1) (fst e2) +fst (compEquiv∙ e1 e2) = compEquiv (fst e1) (fst e2) snd (compEquiv∙ e1 e2) = cong (fst (fst e2)) (snd e1) snd e2 Equiv∙J : {B : Pointed } (C : (A : Pointed ) A ≃∙ B Type ℓ') @@ -73,7 +73,7 @@ (C : (A : Pointed ) A ≃∙ (fst B , b) Type ℓ') C (fst B , b) (idEquiv (fst B) , refl) C (A , a) (e , p)) - λ a b J b p + λ a b J b p (C : (A : Pointed ) A ≃∙ (fst B , b) Type ℓ') C (fst B , b) (idEquiv (fst B) , refl) @@ -91,7 +91,7 @@ ((f : fst A B) P (f (pt A)) (f , refl)) {b₀ : B} (f : A →∙ (B , b₀)) P b₀ f →∙J {A = A} P ind = - uncurry λ f J b₀ y P b₀ (f , y)) (ind f) + uncurry λ f J b₀ y P b₀ (f , y)) (ind f) {- HIT allowing for pattern matching on pointed types -} data Pointer {} (A : Pointed ) : Type where diff --git a/docs/Cubical.Foundations.Pointed.FunExt.html b/docs/Cubical.Foundations.Pointed.FunExt.html index e47bfe5..ca86d57 100644 --- a/docs/Cubical.Foundations.Pointed.FunExt.html +++ b/docs/Cubical.Foundations.Pointed.FunExt.html @@ -39,12 +39,12 @@ -- funExt∙≃ using the other kind of pointed homotopy funExt∙≃ : (f g : Π∙ A B ptB) (f ∙∼ g) (f g) - funExt∙≃ f g = compEquiv (∙∼≃∙∼P f g) (funExt∙P≃ f g) + funExt∙≃ f g = compEquiv (∙∼≃∙∼P f g) (funExt∙P≃ f g) -- standard pointed function extensionality and its inverse funExt∙ : {f g : Π∙ A B ptB} f ∙∼ g f g funExt∙ {f = f} {g = g} = equivFun (funExt∙≃ f g) funExt∙⁻ : {f g : Π∙ A B ptB} f g f ∙∼ g - funExt∙⁻ {f = f} {g = g} = equivFun (invEquiv (funExt∙≃ f g)) + funExt∙⁻ {f = f} {g = g} = equivFun (invEquiv (funExt∙≃ f g)) \ No newline at end of file diff --git a/docs/Cubical.Foundations.Pointed.Homogeneous.html b/docs/Cubical.Foundations.Pointed.Homogeneous.html index 8c97d49..4aefe52 100644 --- a/docs/Cubical.Foundations.Pointed.Homogeneous.html +++ b/docs/Cubical.Foundations.Pointed.Homogeneous.html @@ -64,7 +64,7 @@ (h : isHomogeneous B∙') PathP i fst (p i) fst (q i)) (f∙ .fst) (g∙ .fst) PathP i p i →∙ q i) f∙ g∙ -→∙HomogeneousPathP p q h r = toPathP (→∙Homogeneous≡ h (fromPathP r)) +→∙HomogeneousPathP p q h r = toPathP (→∙Homogeneous≡ h (fromPathP r)) →∙Homogeneous≡Path : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} (h : isHomogeneous B∙) (p q : f∙ g∙) cong fst p cong fst q p q @@ -76,7 +76,7 @@ (badPath p q r) where newPath : (p q : f∙ g∙) (r : cong fst p cong fst q) - Square (refl {x = b}) refl refl refl + Square (refl {x = b}) refl refl refl newPath p q r i j = hcomp k λ {(i = i0) cong snd p j k ; (i = i1) cong snd q j k @@ -111,15 +111,15 @@ →∙HomogeneousSquare : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ h∙ l∙ : A∙ →∙ B∙} (h : isHomogeneous B∙) (s : f∙ h∙) (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) - Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) - Square p q s t + Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) + Square p q s t →∙HomogeneousSquare {f∙ = f∙} {g∙ = g∙} {h∙ = h∙} {l∙ = l∙} h = - J h∙ s (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) - Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) - Square p q s t) - (J l∙ t (p : f∙ g∙) (q : f∙ l∙) - Square (cong fst p) (cong fst q) refl (cong fst t) - Square p q refl t) + J h∙ s (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) + Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) + Square p q s t) + (J l∙ t (p : f∙ g∙) (q : f∙ l∙) + Square (cong fst p) (cong fst q) refl (cong fst t) + Square p q refl t) (→∙Homogeneous≡Path {f∙ = f∙} {g∙ = g∙} h)) isHomogeneousPi : { ℓ'} {A : Type } {B∙ : A Pointed ℓ'} @@ -160,7 +160,7 @@ T a i = typ (h (f∙ .fst a) i) t₀ : PathP i T (pt A∙) i) (pt B∙) (pt B∙) - t₀ = cong pt (h (f∙ .fst (pt A∙))) f∙ .snd + t₀ = cong pt (h (f∙ .fst (pt A∙))) f∙ .snd isHomogeneousProd : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} isHomogeneous A∙ isHomogeneous B∙ isHomogeneous (A∙ ×∙ B∙) diff --git a/docs/Cubical.Foundations.Pointed.Homotopy.html b/docs/Cubical.Foundations.Pointed.Homotopy.html index 5fd82f5..b634f57 100644 --- a/docs/Cubical.Foundations.Pointed.Homotopy.html +++ b/docs/Cubical.Foundations.Pointed.Homotopy.html @@ -99,7 +99,7 @@ -- inverse of ∙∼→∙∼P extracted from the equivalence ∙∼P→∙∼ : {f g : Π∙ A B ptB} f ∙∼P g f ∙∼ g - ∙∼P→∙∼ {f = f} {g = g} = invEq (∙∼≃∙∼P f g) + ∙∼P→∙∼ {f = f} {g = g} = invEq (∙∼≃∙∼P f g) -- ∙∼≃∙∼P transformed to a path ∙∼≡∙∼P : (f g : Π∙ A B ptB) (f ∙∼ g) (f ∙∼P g) diff --git a/docs/Cubical.Foundations.Pointed.Properties.html b/docs/Cubical.Foundations.Pointed.Properties.html index 87b015b..14ba990 100644 --- a/docs/Cubical.Foundations.Pointed.Properties.html +++ b/docs/Cubical.Foundations.Pointed.Properties.html @@ -61,7 +61,7 @@ post∘∙ X f .snd = ΣPathP ( (funExt λ _ f .snd) - , (sym (lUnit (f .snd)) λ i j f .snd (i j))) + , (sym (lUnit (f .snd)) λ i j f .snd (i j))) -- pointed identity id∙ : (A : Pointed ℓA) (A →∙ A) @@ -121,21 +121,21 @@ Iso.fun (pre∘∙equiv {A = A} {B = B} {C = C} e) = toEq' B C A e Iso.inv (pre∘∙equiv {A = A} {B = B} {C = C} e) = fromEq' B C A e Iso.rightInv (pre∘∙equiv {A = A} {B = B} {C = C} e) = - J ptC p section (toEq' B (fst C , ptC) A (fst e , p)) + J ptC p section (toEq' B (fst C , ptC) A (fst e , p)) (fromEq' B (fst C , ptC) A (fst e , p))) (uncurry f p ΣPathP (funExt x isHAEquiv.rinv (HAe .snd) (f x)) , ((sym (rUnit _) cong (cong (fst (fst e))) - λ i cong (invEq (fst e)) p - (lUnit (retEq (fst e) (pt B)) (~ i))) + λ i cong (invEq (fst e)) p + (lUnit (retEq (fst e) (pt B)) (~ i))) cong-∙ (fst (fst e)) - (cong (invEq (fst e)) p) - (retEq (fst e) (pt B)) + (cong (invEq (fst e)) p) + (retEq (fst e) (pt B)) refl - flipSquare (((λ _ isHAEquiv.rinv (HAe .snd) (f (pt A))) + flipSquare (((λ _ isHAEquiv.rinv (HAe .snd) (f (pt A))) refl) - lem _ _ _ _ (cong (isHAEquiv.rinv (HAe .snd)) p - sym (isHAEquiv.com (HAe .snd) (pt B)))))))) + lem _ _ _ _ (cong (isHAEquiv.rinv (HAe .snd)) p + sym (isHAEquiv.com (HAe .snd) (pt B)))))))) (snd e) where HAe = equiv→HAEquiv (fst e) @@ -149,21 +149,21 @@ ; (j = i1) l i}) (P i j) Iso.leftInv (pre∘∙equiv {A = A} {B = B} {C = C} e) = - J pt p retract (toEq' B (fst C , pt) A (fst e , p)) + J pt p retract (toEq' B (fst C , pt) A (fst e , p)) (fromEq' B (fst C , pt) A (fst e , p))) (uncurry f - J ptB p + J ptB p fromEq' (fst B , ptB) (fst C , fst (fst e) ptB) A (fst e , refl) (toEq' (fst B , ptB) (fst C , fst (fst e) ptB) A (fst e , refl) (f , p)) (f , p)) (ΣPathP - (funExt x retEq (fst e) (f x)) - , ((cong₂ _∙_ ((λ i cong (invEq (fst e)) (lUnit refl (~ i)))) - (sym (lUnit _) λ _ retEq (fst e) (f (pt A))) + (funExt x retEq (fst e) (f x)) + , ((cong₂ _∙_ ((λ i cong (invEq (fst e)) (lUnit refl (~ i)))) + (sym (lUnit _) λ _ retEq (fst e) (f (pt A))) sym (lUnit _)) - λ i j retEq (fst e) (f (pt A)) (i j)))))) + λ i j retEq (fst e) (f (pt A)) (i j)))))) (snd e) post∘∙equiv : {ℓA ℓB ℓC} {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} @@ -171,37 +171,37 @@ Iso.fun (post∘∙equiv {A = A} {B = B} {C = C} e) = toEq A B C e Iso.inv (post∘∙equiv {A = A} {B = B} {C = C} e) = fromEq A B C e Iso.rightInv (post∘∙equiv {A = A}{B = B , ptB} {C = C} e) = - J pt p section (toEq A (B , pt) C (fst e , p)) + J pt p section (toEq A (B , pt) C (fst e , p)) (fromEq A (B , pt) C (fst e , p))) (uncurry f - J ptC p toEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) + J ptC p toEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) (fromEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) (f , p)) (f , p)) (ΣPathP (funExt x cong f (isHAEquiv.rinv (snd HAe) x)) , (cong₂ _∙_ - i cong f (cong (fst (fst e)) (lUnit (retEq (fst e) (pt A)) (~ i)))) + i cong f (cong (fst (fst e)) (lUnit (retEq (fst e) (pt A)) (~ i)))) (sym (rUnit refl)) sym (rUnit _) cong (cong f) (isHAEquiv.com (snd HAe) (pt A))) - λ i j f (isHAEquiv.rinv (snd HAe) (fst HAe (pt A)) (i j)))))) + λ i j f (isHAEquiv.rinv (snd HAe) (fst HAe (pt A)) (i j)))))) (snd e) where HAe = equiv→HAEquiv (fst e) Iso.leftInv (post∘∙equiv {A = A} {B = B , ptB} {C = C} e) = - J pt p retract (toEq A (B , pt) C (fst e , p)) + J pt p retract (toEq A (B , pt) C (fst e , p)) (fromEq A (B , pt) C (fst e , p))) - (uncurry f J ptC y + (uncurry f J ptC y fromEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) (toEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) (f , y)) (f , y)) - (ΣPathP (funExt x cong f (retEq (fst e) x)) + (ΣPathP (funExt x cong f (retEq (fst e) x)) , (sym (lUnit _) sym (rUnit _) cong (cong f) (sym (lUnit _)) - _ cong f (retEq (fst e) (pt A))) - λ i j f (retEq (fst e) (pt A) (i j))))))) + _ cong f (retEq (fst e) (pt A))) + λ i j f (retEq (fst e) (pt A) (i j))))))) (snd e) flip→∙∙ : {A : Pointed } {B : Pointed ℓ'} {C : Pointed ℓA} diff --git a/docs/Cubical.Foundations.Powerset.html b/docs/Cubical.Foundations.Powerset.html index d07a814..1e5bac6 100644 --- a/docs/Cubical.Foundations.Powerset.html +++ b/docs/Cubical.Foundations.Powerset.html @@ -28,8 +28,8 @@ : Type Type (ℓ-suc ) X = X hProp _ -isSetℙ : isSet ( X) -isSetℙ = isSetΠ λ x isSetHProp +isSetℙ : isSet ( X) +isSetℙ = isSetΠ λ x isSetHProp infix 5 _∈_ @@ -39,10 +39,10 @@ _⊆_ : {X : Type } X X Type A B = x x A x B -∈-isProp : (A : X) (x : X) isProp (x A) +∈-isProp : (A : X) (x : X) isProp (x A) ∈-isProp A = snd A -⊆-isProp : (A B : X) isProp (A B) +⊆-isProp : (A B : X) isProp (A B) ⊆-isProp A B = isPropΠ2 x _ ∈-isProp B x) ⊆-refl : (A : X) A A @@ -64,15 +64,4 @@ (⊆-refl-consequence A B) _ isSetℙ A B _ _) _ isPropΣ (⊆-isProp A B) _ ⊆-isProp B A) _ _)) - -module _ where - open import Cubical.Data.Unit - open import Cubical.Data.Empty - - full : X - full = λ _ Unit* , isPropUnit* - - : X - = λ _ ⊥* , isProp⊥* - \ No newline at end of file diff --git a/docs/Cubical.Foundations.Prelude.html b/docs/Cubical.Foundations.Prelude.html index 21f2444..9192fbb 100644 --- a/docs/Cubical.Foundations.Prelude.html +++ b/docs/Cubical.Foundations.Prelude.html @@ -27,11 +27,11 @@ open import Cubical.Core.Primitives public infixr 30 _∙_ -infixr 30 _∙₂_ +infixr 30 _∙₂_ infix 3 _∎ infixr 2 step-≡ _≡⟨⟩_ infixr 2.5 _≡⟨_⟩≡⟨_⟩_ -infixl 4 _≡$_ _≡$S_ +infixl 4 _≡$_ _≡$S_ -- Basic theory about paths. These proofs should typically be -- inlined. This module also makes equational reasoning work with @@ -328,283 +328,291 @@ ((x : A) PathP (B x) (f x) (g x)) funExt⁻ eq x i = eq i x -implicitFunExt⁻ : {B : A I Type ℓ'} - {f : {x : A} B x i0} {g : {x : A} B x i1} - PathP i {x : A} B x i) f g - ({x : A} PathP (B x) (f {x}) (g {x})) -implicitFunExt⁻ eq {x} i = eq i {x} +congP₂$ : {A : I Type } {B : i A i Type ℓ'} + {f : x B i0 x} {g : y B i1 y} + (p : PathP i x B i x) f g) + {x y} (p : PathP A x y) PathP i B i (p i)) (f x) (g y) +congP₂$ eq x i = eq i (x i) -_≡$_ = funExt⁻ +implicitFunExt⁻ : {B : A I Type ℓ'} + {f : {x : A} B x i0} {g : {x : A} B x i1} + PathP i {x : A} B x i) f g + ({x : A} PathP (B x) (f {x}) (g {x})) +implicitFunExt⁻ eq {x} i = eq i {x} -{- `S` stands for simply typed. Using `funExtS⁻` instead of `funExt⁻` +_≡$_ = funExt⁻ + +{- `S` stands for simply typed. Using `funExtS⁻` instead of `funExt⁻` can help Agda to solve metavariables that may otherwise remain unsolved. -} -funExtS⁻ : {B : I Type ℓ'} - {f : (x : A) B i0} {g : (x : A) B i1} - PathP i (x : A) B i) f g - ((x : A) PathP i B i) (f x) (g x)) -funExtS⁻ eq x i = eq i x +funExtS⁻ : {B : I Type ℓ'} + {f : (x : A) B i0} {g : (x : A) B i1} + PathP i (x : A) B i) f g + ((x : A) PathP i B i) (f x) (g x)) +funExtS⁻ eq x i = eq i x -_≡$S_ = funExtS⁻ +_≡$S_ = funExtS⁻ --- J for paths and its computation rule +-- J for paths and its computation rule -module _ (P : y x y Type ℓ') (d : P x refl) where +module _ (P : y x y Type ℓ') (d : P x refl) where - J : (p : x y) P y p - J p = transport i P (p i) j p (i j))) d + J : (p : x y) P y p + J p = transport i P (p i) j p (i j))) d - JRefl : J refl d - JRefl = transportRefl d + JRefl : J refl d + JRefl = transportRefl d - J-∙ : (p : x y) (q : y z) - J (p q) transport i P (q i) j compPath-filler p q i j)) (J p) - J-∙ p q k = - transp - i P (q (i ~ k)) - j compPath-filler p q (i ~ k) j)) (~ k) - (J j compPath-filler p q (~ k) j)) + J-∙ : (p : x y) (q : y z) + J (p q) transport i P (q i) j compPath-filler p q i j)) (J p) + J-∙ p q k = + transp + i P (q (i ~ k)) + j compPath-filler p q (i ~ k) j)) (~ k) + (J j compPath-filler p q (~ k) j)) --- Multi-variable versions of J +-- Multi-variable versions of J -module _ {b : B x} - (P : (y : A) (p : x y) (z : B y) (q : PathP i B (p i)) b z) Type ℓ'') - (d : P _ refl _ refl) where +module _ {b : B x} + (P : (y : A) (p : x y) (z : B y) (q : PathP i B (p i)) b z) Type ℓ'') + (d : P _ refl _ refl) where - JDep : {y : A} (p : x y) {z : B y} (q : PathP i B (p i)) b z) P _ p _ q - JDep _ q = transport i P _ _ _ j q (i j))) d + JDep : {y : A} (p : x y) {z : B y} (q : PathP i B (p i)) b z) P _ p _ q + JDep _ q = transport i P _ _ _ j q (i j))) d - JDepRefl : JDep refl refl d - JDepRefl = transportRefl d + JDepRefl : JDep refl refl d + JDepRefl = transportRefl d -module _ {x : A} - {P : (y : A) x y Type ℓ'} {d : (y : A) (p : x y) P y p} - (Q : (y : A) (p : x y) (z : P y p) d y p z Type ℓ'') - (r : Q _ refl _ refl) where +module _ {x : A} + {P : (y : A) x y Type ℓ'} {d : (y : A) (p : x y) P y p} + (Q : (y : A) (p : x y) (z : P y p) d y p z Type ℓ'') + (r : Q _ refl _ refl) where - private - ΠQ : (y : A) x y _ - ΠQ y p = z q Q y p z q + private + ΠQ : (y : A) x y _ + ΠQ y p = z q Q y p z q - J2 : {y : A} (p : x y) {z : P y p} (q : d y p z) Q _ p _ q - J2 p = J ΠQ _ J (Q x refl) r) p _ + J2 : {y : A} (p : x y) {z : P y p} (q : d y p z) Q _ p _ q + J2 p = J ΠQ _ J (Q x refl) r) p _ - J2Refl : J2 refl refl r - J2Refl = i JRefl ΠQ _ J (Q x refl) r) i _ refl) JRefl (Q x refl) _ + J2Refl : J2 refl refl r + J2Refl = i JRefl ΠQ _ J (Q x refl) r) i _ refl) JRefl (Q x refl) _ --- A prefix operator version of J that is more suitable to be nested +-- A prefix operator version of J that is more suitable to be nested -module _ {P : y x y Type ℓ'} (d : P x refl) where +module _ {P : y x y Type ℓ'} (d : P x refl) where - J>_ : y (p : x y) P y p - J>_ _ p = transport i P (p i) j p (i j))) d + J>_ : y (p : x y) P y p + J>_ _ p = transport i P (p i) j p (i j))) d - infix 10 J>_ + infix 10 J>_ --- Converting to and from a PathP - -module _ {A : I Type } {x : A i0} {y : A i1} where - toPathP : transport i A i) x y PathP A x y - toPathP p i = hcomp j λ { (i = i0) x - ; (i = i1) p j }) - (transp j A (i j)) (~ i) x) - - fromPathP : PathP A x y transport i A i) x y - fromPathP p i = transp j A (i j)) i (p i) - --- Whiskering a dependent path by a path --- Double whiskering -_◁_▷_ : {} {A : I Type } {a₀ a₀' : A i0} {a₁ a₁' : A i1} - a₀ a₀' PathP A a₀' a₁ a₁ a₁' - PathP A a₀ a₁' -(p P q) i = - hcomp j λ {(i = i0) p (~ j) ; (i = i1) q j}) (P i) - -doubleWhiskFiller : - {} {A : I Type } {a₀ a₀' : A i0} {a₁ a₁' : A i1} - (p : a₀ a₀') (pq : PathP A a₀' a₁) (q : a₁ a₁') - PathP i PathP A (p (~ i)) (q i)) - pq - (p pq q) -doubleWhiskFiller p pq q k i = - hfill j λ {(i = i0) p (~ j) ; (i = i1) q j}) - (inS (pq i)) - k - -_◁_ : {} {A : I Type } {a₀ a₀' : A i0} {a₁ : A i1} - a₀ a₀' PathP A a₀' a₁ PathP A a₀ a₁ -(p q) = p q refl - -_▷_ : {} {A : I Type } {a₀ : A i0} {a₁ a₁' : A i1} - PathP A a₀ a₁ a₁ a₁' PathP A a₀ a₁' -p q = refl p q - --- Direct definitions of lower h-levels - -isContr : Type Type -isContr A = Σ[ x A ] (∀ y x y) - -isProp : Type Type -isProp A = (x y : A) x y - -isSet : Type Type -isSet A = (x y : A) isProp (x y) - -isGroupoid : Type Type -isGroupoid A = a b isSet (Path A a b) - -is2Groupoid : Type Type -is2Groupoid A = a b isGroupoid (Path A a b) - --- Contractibility of singletons - -singlP : (A : I Type ) (a : A i0) Type _ -singlP A a = Σ[ x A i1 ] PathP A a x - -singl : (a : A) Type _ -singl {A = A} a = singlP _ A) a - -isContrSingl : (a : A) isContr (singl a) -isContrSingl a .fst = (a , refl) -isContrSingl a .snd p i .fst = p .snd i -isContrSingl a .snd p i .snd j = p .snd (i j) - -isContrSinglP : (A : I Type ) (a : A i0) isContr (singlP A a) -isContrSinglP A a .fst = _ , transport-filler i A i) a -isContrSinglP A a .snd (x , p) i = - _ , λ j fill A j λ {(i = i0) transport-filler i A i) a j; (i = i1) p j}) (inS a) j - --- Higher cube types - -SquareP : - (A : I I Type ) - {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP j A i0 j) a₀₀ a₀₁) - {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP j A i1 j) a₁₀ a₁₁) - (a₋₀ : PathP i A i i0) a₀₀ a₁₀) (a₋₁ : PathP i A i i1) a₀₁ a₁₁) - Type -SquareP A a₀₋ a₁₋ a₋₀ a₋₁ = PathP i PathP j A i j) (a₋₀ i) (a₋₁ i)) a₀₋ a₁₋ - -Square : - {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) - {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) - (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) - Type _ -Square a₀₋ a₁₋ a₋₀ a₋₁ = PathP i a₋₀ i a₋₁ i) a₀₋ a₁₋ - -Cube : - {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ a₀₀₁} - {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ a₀₁₁} - {a₀₋₀ : a₀₀₀ a₀₁₀} {a₀₋₁ : a₀₀₁ a₀₁₁} - (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) - {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ a₁₀₁} - {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ a₁₁₁} - {a₁₋₀ : a₁₀₀ a₁₁₀} {a₁₋₁ : a₁₀₁ a₁₁₁} - (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) - {a₋₀₀ : a₀₀₀ a₁₀₀} {a₋₀₁ : a₀₀₁ a₁₀₁} - (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) - {a₋₁₀ : a₀₁₀ a₁₁₀} {a₋₁₁ : a₀₁₁ a₁₁₁} - (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) - (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) - (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) - Type _ -Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = - PathP i Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ - --- Horizontal composition of squares (along their second dimension) --- See Cubical.Foundations.Path for vertical composition - -_∙₂_ : - {a₀₀ a₀₁ a₀₂ : A} {a₀₋ : a₀₀ a₀₁} {b₀₋ : a₀₁ a₀₂} - {a₁₀ a₁₁ a₁₂ : A} {a₁₋ : a₁₀ a₁₁} {b₁₋ : a₁₁ a₁₂} - {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} {a₋₂ : a₀₂ a₁₂} - (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square b₀₋ b₁₋ a₋₁ a₋₂) - Square (a₀₋ b₀₋) (a₁₋ b₁₋) a₋₀ a₋₂ -_∙₂_ = congP₂ _ _∙_) - --- Alternative (equivalent) definitions of hlevel n that give fillers for n-cubes instead of n-globes - -isSet' : Type Type -isSet' A = - {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) - {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) - (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) - Square a₀₋ a₁₋ a₋₀ a₋₁ - -isSet→isSet' : isSet A isSet' A -isSet→isSet' Aset _ _ _ _ = toPathP (Aset _ _ _ _) - -isSet'→isSet : isSet' A isSet A -isSet'→isSet Aset' x y p q = Aset' p q refl refl - -isGroupoid' : Type Type -isGroupoid' A = - {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ a₀₀₁} - {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ a₀₁₁} - {a₀₋₀ : a₀₀₀ a₀₁₀} {a₀₋₁ : a₀₀₁ a₀₁₁} - (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) - {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ a₁₀₁} - {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ a₁₁₁} - {a₁₋₀ : a₁₀₀ a₁₁₀} {a₁₋₁ : a₁₀₁ a₁₁₁} - (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) - {a₋₀₀ : a₀₀₀ a₁₀₀} {a₋₀₁ : a₀₀₁ a₁₀₁} - (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) - {a₋₁₀ : a₀₁₀ a₁₁₀} {a₋₁₁ : a₀₁₁ a₁₁₁} - (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) - (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) - (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) - Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ - --- Essential properties of isProp and isContr - -isProp→PathP : {B : I Type } ((i : I) isProp (B i)) - (b0 : B i0) (b1 : B i1) - PathP B b0 b1 -isProp→PathP hB b0 b1 = toPathP (hB _ _ _) - -isPropIsContr : isProp (isContr A) -isPropIsContr (c0 , h0) (c1 , h1) j .fst = h0 c1 j -isPropIsContr (c0 , h0) (c1 , h1) j .snd y i = - hcomp k λ { (i = i0) h0 (h0 c1 j) k; - (i = i1) h0 y k; - (j = i0) h0 (h0 y i) k; - (j = i1) h0 (h1 y i) k}) - c0 - -isContr→isProp : isContr A isProp A -isContr→isProp (x , p) a b = sym (p a) p b - -isProp→isSet : isProp A isSet A -isProp→isSet h a b p q j i = - hcomp k λ { (i = i0) h a a k - ; (i = i1) h a b k - ; (j = i0) h a (p i) k - ; (j = i1) h a (q i) k }) a - -isProp→isSet' : isProp A isSet' A -isProp→isSet' h {a} p q r s i j = - hcomp k λ { (i = i0) h a (p j) k - ; (i = i1) h a (q j) k - ; (j = i0) h a (r i) k - ; (j = i1) h a (s i) k}) a - -isPropIsProp : isProp (isProp A) -isPropIsProp f g i a b = isProp→isSet f a b (f a b) (g a b) i - -isPropSingl : {a : A} isProp (singl a) -isPropSingl = isContr→isProp (isContrSingl _) - -isPropSinglP : {A : I Type } {a : A i0} isProp (singlP A a) -isPropSinglP = isContr→isProp (isContrSinglP _ _) - --- Universe lifting - -record Lift {i j} (A : Type i) : Type (ℓ-max i j) where - constructor lift - field - lower : A - -open Lift public - -liftExt : {A : Type } {a b : Lift {} {ℓ'} A} (lower a lower b) a b -liftExt x i = lift (x i) +-- Converting to and from a PathP + +module _ {A : I Type } {x : A i0} {y : A i1} where + toPathP : transport i A i) x y PathP A x y + toPathP p i = hcomp j λ { (i = i0) x + ; (i = i1) p j }) + (transp j A (i j)) (~ i) x) + + fromPathP : PathP A x y transport i A i) x y + fromPathP p i = transp j A (i j)) i (p i) + +-- Whiskering a dependent path by a path +-- Double whiskering +_◁_▷_ : {} {A : I Type } {a₀ a₀' : A i0} {a₁ a₁' : A i1} + a₀ a₀' PathP A a₀' a₁ a₁ a₁' + PathP A a₀ a₁' +(p P q) i = + hcomp j λ {(i = i0) p (~ j) ; (i = i1) q j}) (P i) + +doubleWhiskFiller : + {} {A : I Type } {a₀ a₀' : A i0} {a₁ a₁' : A i1} + (p : a₀ a₀') (pq : PathP A a₀' a₁) (q : a₁ a₁') + PathP i PathP A (p (~ i)) (q i)) + pq + (p pq q) +doubleWhiskFiller p pq q k i = + hfill j λ {(i = i0) p (~ j) ; (i = i1) q j}) + (inS (pq i)) + k + +_◁_ : {} {A : I Type } {a₀ a₀' : A i0} {a₁ : A i1} + a₀ a₀' PathP A a₀' a₁ PathP A a₀ a₁ +(p q) = p q refl + +_▷_ : {} {A : I Type } {a₀ : A i0} {a₁ a₁' : A i1} + PathP A a₀ a₁ a₁ a₁' PathP A a₀ a₁' +p q = refl p q + +-- Direct definitions of lower h-levels + +isContr : Type Type +isContr A = Σ[ x A ] (∀ y x y) + +isProp : Type Type +isProp A = (x y : A) x y + +isSet : Type Type +isSet A = (x y : A) isProp (x y) + +isGroupoid : Type Type +isGroupoid A = a b isSet (Path A a b) + +is2Groupoid : Type Type +is2Groupoid A = a b isGroupoid (Path A a b) + +-- Contractibility of singletons + +singlP : (A : I Type ) (a : A i0) Type _ +singlP A a = Σ[ x A i1 ] PathP A a x + +singl : (a : A) Type _ +singl {A = A} a = singlP _ A) a + +isContrSingl : (a : A) isContr (singl a) +isContrSingl a .fst = (a , refl) +isContrSingl a .snd p i .fst = p .snd i +isContrSingl a .snd p i .snd j = p .snd (i j) + +isContrSinglP : (A : I Type ) (a : A i0) isContr (singlP A a) +isContrSinglP A a .fst = _ , transport-filler i A i) a +isContrSinglP A a .snd (x , p) i = + _ , λ j fill A j λ {(i = i0) transport-filler i A i) a j; (i = i1) p j}) (inS a) j + +-- Higher cube types + +SquareP : + (A : I I Type ) + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP j A i0 j) a₀₀ a₀₁) + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP j A i1 j) a₁₀ a₁₁) + (a₋₀ : PathP i A i i0) a₀₀ a₁₀) (a₋₁ : PathP i A i i1) a₀₁ a₁₁) + Type +SquareP A a₀₋ a₁₋ a₋₀ a₋₁ = PathP i PathP j A i j) (a₋₀ i) (a₋₁ i)) a₀₋ a₁₋ + +Square : + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) + (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) + Type _ +Square a₀₋ a₁₋ a₋₀ a₋₁ = PathP i a₋₀ i a₋₁ i) a₀₋ a₁₋ + +Cube : + {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ a₀₁₁} + {a₀₋₀ : a₀₀₀ a₀₁₀} {a₀₋₁ : a₀₀₁ a₀₁₁} + (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) + {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ a₁₀₁} + {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ a₁₁₁} + {a₁₋₀ : a₁₀₀ a₁₁₀} {a₁₋₁ : a₁₀₁ a₁₁₁} + (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) + {a₋₀₀ : a₀₀₀ a₁₀₀} {a₋₀₁ : a₀₀₁ a₁₀₁} + (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) + {a₋₁₀ : a₀₁₀ a₁₁₀} {a₋₁₁ : a₀₁₁ a₁₁₁} + (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) + (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) + (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) + Type _ +Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = + PathP i Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ + +-- See HLevels.agda for CubeP + +-- Horizontal composition of squares (along their second dimension) +-- See Cubical.Foundations.Path for vertical composition + +_∙₂_ : + {a₀₀ a₀₁ a₀₂ : A} {a₀₋ : a₀₀ a₀₁} {b₀₋ : a₀₁ a₀₂} + {a₁₀ a₁₁ a₁₂ : A} {a₁₋ : a₁₀ a₁₁} {b₁₋ : a₁₁ a₁₂} + {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} {a₋₂ : a₀₂ a₁₂} + (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square b₀₋ b₁₋ a₋₁ a₋₂) + Square (a₀₋ b₀₋) (a₁₋ b₁₋) a₋₀ a₋₂ +_∙₂_ = congP₂ _ _∙_) + +-- Alternative (equivalent) definitions of hlevel n that give fillers for n-cubes instead of n-globes + +isSet' : Type Type +isSet' A = + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) + (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) + Square a₀₋ a₁₋ a₋₀ a₋₁ + +isSet→isSet' : isSet A isSet' A +isSet→isSet' Aset _ _ _ _ = toPathP (Aset _ _ _ _) + +isSet'→isSet : isSet' A isSet A +isSet'→isSet Aset' x y p q = Aset' p q refl refl + +isGroupoid' : Type Type +isGroupoid' A = + {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ a₀₁₁} + {a₀₋₀ : a₀₀₀ a₀₁₀} {a₀₋₁ : a₀₀₁ a₀₁₁} + (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) + {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ a₁₀₁} + {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ a₁₁₁} + {a₁₋₀ : a₁₀₀ a₁₁₀} {a₁₋₁ : a₁₀₁ a₁₁₁} + (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) + {a₋₀₀ : a₀₀₀ a₁₀₀} {a₋₀₁ : a₀₀₁ a₁₀₁} + (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) + {a₋₁₀ : a₀₁₀ a₁₁₀} {a₋₁₁ : a₀₁₁ a₁₁₁} + (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) + (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) + (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) + Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + +-- Essential properties of isProp and isContr + +isProp→PathP : {B : I Type } ((i : I) isProp (B i)) + (b0 : B i0) (b1 : B i1) + PathP B b0 b1 +isProp→PathP hB b0 b1 = toPathP (hB _ _ _) + +isPropIsContr : isProp (isContr A) +isPropIsContr (c0 , h0) (c1 , h1) j .fst = h0 c1 j +isPropIsContr (c0 , h0) (c1 , h1) j .snd y i = + hcomp k λ { (i = i0) h0 (h0 c1 j) k; + (i = i1) h0 y k; + (j = i0) h0 (h0 y i) k; + (j = i1) h0 (h1 y i) k}) + c0 + +isContr→isProp : isContr A isProp A +isContr→isProp (x , p) a b = sym (p a) p b + +isProp→isSet : isProp A isSet A +isProp→isSet h a b p q j i = + hcomp k λ { (i = i0) h a a k + ; (i = i1) h a b k + ; (j = i0) h a (p i) k + ; (j = i1) h a (q i) k }) a + +isProp→isSet' : isProp A isSet' A +isProp→isSet' h {a} p q r s i j = + hcomp k λ { (i = i0) h a (p j) k + ; (i = i1) h a (q j) k + ; (j = i0) h a (r i) k + ; (j = i1) h a (s i) k}) a + +isPropIsProp : isProp (isProp A) +isPropIsProp f g i a b = isProp→isSet f a b (f a b) (g a b) i + +isPropSingl : {a : A} isProp (singl a) +isPropSingl = isContr→isProp (isContrSingl _) + +isPropSinglP : {A : I Type } {a : A i0} isProp (singlP A a) +isPropSinglP = isContr→isProp (isContrSinglP _ _) + +-- Universe lifting + +record Lift {i j} (A : Type i) : Type (ℓ-max i j) where + constructor lift + field + lower : A + +open Lift public + +liftExt : {A : Type } {a b : Lift {} {ℓ'} A} (lower a lower b) a b +liftExt x i = lift (x i) \ No newline at end of file diff --git a/docs/Cubical.Foundations.SIP.html b/docs/Cubical.Foundations.SIP.html index 8b2d24a..9b1ed73 100644 --- a/docs/Cubical.Foundations.SIP.html +++ b/docs/Cubical.Foundations.SIP.html @@ -55,11 +55,11 @@ UnivalentStr S ι SNS S ι UnivalentStr→SNS S ι θ {X = X} s t = ι (X , s) (X , t) (idEquiv X) - ≃⟨ θ (idEquiv X) + ≃⟨ θ (idEquiv X) PathP i S (ua (idEquiv X) i)) s t - ≃⟨ pathToEquiv j PathP i S (uaIdEquiv {A = X} j i)) s t) + ≃⟨ pathToEquiv j PathP i S (uaIdEquiv {A = X} j i)) s t) s t - + SNS→UnivalentStr : (ι : StrEquiv S ℓ₃) SNS S ι UnivalentStr S ι @@ -73,11 +73,11 @@ C : (s t : S Y) ι (Y , s) (Y , t) (idEquiv Y) PathP i S (ua (idEquiv Y) i)) s t C s t = ι (Y , s) (Y , t) (idEquiv Y) - ≃⟨ θ s t + ≃⟨ θ s t s t - ≃⟨ pathToEquiv j PathP i S (uaIdEquiv {A = Y} (~ j) i)) s t) + ≃⟨ pathToEquiv j PathP i S (uaIdEquiv {A = Y} (~ j) i)) s t) PathP i S (ua (idEquiv Y) i)) s t - + TransportStr : {S : Type Type ℓ₁} (α : EquivAction S) Type (ℓ-max (ℓ-suc ) ℓ₁) TransportStr {} {S = S} α = @@ -87,23 +87,23 @@ TransportStr α UnivalentStr S (EquivAction→StrEquiv α) TransportStr→UnivalentStr {S = S} α τ {X , s} {Y , t} e = equivFun (α e) s t - ≃⟨ pathToEquiv (cong (_≡ t) (τ e s)) + ≃⟨ pathToEquiv (cong (_≡ t) (τ e s)) subst S (ua e) s t - ≃⟨ invEquiv (PathP≃Path _ _ _) + ≃⟨ invEquiv (PathP≃Path _ _ _) PathP i S (ua e i)) s t - + UnivalentStr→TransportStr : {S : Type Type ℓ₁} (α : EquivAction S) UnivalentStr S (EquivAction→StrEquiv α) TransportStr α UnivalentStr→TransportStr {S = S} α θ e s = - invEq (θ e) (transport-filler (cong S (ua e)) s) + invEq (θ e) (transport-filler (cong S (ua e)) s) invTransportStr : {S : Type Type ℓ₂} (α : EquivAction S) (τ : TransportStr α) - {X Y : Type } (e : X Y) (t : S Y) invEq (α e) t subst⁻ S (ua e) t + {X Y : Type } (e : X Y) (t : S Y) invEq (α e) t subst⁻ S (ua e) t invTransportStr {S = S} α τ e t = - sym (transport⁻Transport (cong S (ua e)) (invEq (α e) t)) - ∙∙ sym (cong (subst⁻ S (ua e)) (τ e (invEq (α e) t))) - ∙∙ cong (subst⁻ S (ua e)) (secEq (α e) t) + sym (transport⁻Transport (cong S (ua e)) (invEq (α e) t)) + ∙∙ sym (cong (subst⁻ S (ua e)) (τ e (invEq (α e) t))) + ∙∙ cong (subst⁻ S (ua e)) (secEq (α e) t) --- We can now define an invertible function @@ -119,8 +119,8 @@ SIP : A ≃[ ι ] B (A B) SIP = - sip , isoToIsEquiv (compIso (Σ-cong-iso (invIso univalenceIso) (equivToIso θ)) ΣPathIsoPathΣ) + sip , isoToIsEquiv (compIso (Σ-cong-iso (invIso univalenceIso) (equivToIso θ)) ΣPathIsoPathΣ) sip⁻ : A B A ≃[ ι ] B - sip⁻ = invEq SIP + sip⁻ = invEq SIP \ No newline at end of file diff --git a/docs/Cubical.Foundations.Transport.html b/docs/Cubical.Foundations.Transport.html index 91db5c8..ceba929 100644 --- a/docs/Cubical.Foundations.Transport.html +++ b/docs/Cubical.Foundations.Transport.html @@ -104,22 +104,22 @@ isInjectiveTransport {p = p} {q} α i = hcomp j λ - { (i = i0) retEq univalence p j - ; (i = i1) retEq univalence q j + { (i = i0) retEq univalence p j + ; (i = i1) retEq univalence q j }) - (invEq univalence ((λ a α i a) , t i)) + (invEq univalence ((λ a α i a) , t i)) where t : PathP i isEquiv a α i a)) (pathToEquiv p .snd) (pathToEquiv q .snd) - t = isProp→PathP i isPropIsEquiv a α i a)) _ _ + t = isProp→PathP i isPropIsEquiv a α i a)) _ _ -transportUaInv : {} {A B : Type } (e : A B) transport (ua (invEquiv e)) transport (sym (ua e)) +transportUaInv : {} {A B : Type } (e : A B) transport (ua (invEquiv e)) transport (sym (ua e)) transportUaInv e = cong transport (uaInvEquiv e) -- notice that transport (ua e) would reduce, thus an alternative definition using EquivJ can give -- refl for the case of idEquiv: -- transportUaInv e = EquivJ (λ _ e → transport (ua (invEquiv e)) ≡ transport (sym (ua e))) refl e isSet-subst : { ℓ'} {A : Type } {B : A Type ℓ'} - (isSet-A : isSet A) + (isSet-A : isSet A) {a : A} (p : a a) (x : B a) subst B p x x isSet-subst {B = B} isSet-A p x = subst p′ subst B p′ x x) (isSet-A _ _ refl p) (substRefl {B = B} x) @@ -165,7 +165,7 @@ transport i P i P i) (p q) transport i P i P i) p transport i P i P i) q overPathFunct p q = - J y P transport i P i P i) (p q) transport i P i P i) p transport i P i P i) q) + J y P transport i P i P i) (p q) transport i P i P i) p transport i P i P i) q) (transportRefl (p q) cong₂ _∙_ (sym (transportRefl p)) (sym (transportRefl q))) -- substition over families of paths @@ -174,7 +174,7 @@ (f g : A B) (p : a a') (q : f a g a) subst x f x g x) p q sym (cong f p) q cong g p substInPaths {a = a} f g p q = - J x p' (subst y f y g y) p' q) (sym (cong f p') q cong g p')) + J x p' (subst y f y g y) p' q) (sym (cong f p') q cong g p')) p=refl p where p=refl : subst y f y g y) refl q @@ -206,4 +206,18 @@ ≡⟨ assoc (sym p) q refl (sym p q) refl ≡⟨ sym (rUnit (sym p q)) sym p q + +transport-filler-ua : {} {A B : Type } (e : A B) (a : A) + SquareP _ i ua e i) + (transport-filler (ua e) a) + (ua-gluePath e refl) + refl + (transportRefl (fst e a)) +transport-filler-ua {A = A} {B = B} (e , _) a j i = + let b = e a + tr = transportRefl b + z = tr (j ~ i) + in glue { (i = i0) a ; (i = i1) tr j }) + (hcomp k λ { (i = i0) b ; (i = i1) tr (j k) ; (j = i1) tr (~ i k) }) + (hcomp k λ { (i = i0) tr (j k) ; (i = i1) z ; (j = i1) z }) z)) \ No newline at end of file diff --git a/docs/Cubical.Foundations.Univalence.html b/docs/Cubical.Foundations.Univalence.html index 50e0aa3..5a8c73c 100644 --- a/docs/Cubical.Foundations.Univalence.html +++ b/docs/Cubical.Foundations.Univalence.html @@ -42,8 +42,8 @@ uaIdEquiv {A = A} i j = Glue A {φ = i ~ j j} _ A , idEquiv A) -- Propositional extensionality -hPropExt : {A B : Type } isProp A isProp B (A B) (B A) A B -hPropExt Aprop Bprop f g = ua (propBiimpl→Equiv Aprop Bprop f g) +hPropExt : {A B : Type } isProp A isProp B (A B) (B A) A B +hPropExt Aprop Bprop f g = ua (propBiimpl→Equiv Aprop Bprop f g) -- the unglue and glue primitives specialized to the case of ua @@ -164,7 +164,7 @@ contrSinglEquiv : {A B : Type } (e : A B) (B , idEquiv B) (A , e) contrSinglEquiv {A = A} {B = B} e = - isContr→isProp (EquivContr B) (B , idEquiv B) (A , e) + isContr→isProp (EquivContr B) (B , idEquiv B) (A , e) -- Equivalence induction EquivJ : {A B : Type } (P : (A : Type ) (e : A B) Type ℓ') @@ -187,7 +187,7 @@ pathToEquiv p .snd = isEquivTransport p pathToEquivRefl : {A : Type } pathToEquiv refl idEquiv A -pathToEquivRefl {A = A} = equivEq i x transp _ A) i x) +pathToEquivRefl {A = A} = equivEq i x transp _ A) i x) -- The computation rule for ua. Because of "ghcomp" it is now very -- simple compared to cubicaltt: @@ -206,7 +206,7 @@ sides (j = i1) = B , idEquiv B pathToEquiv-ua : {A B : Type } (e : A B) pathToEquiv (ua e) e -pathToEquiv-ua e = equivEq (funExt (uaβ e)) +pathToEquiv-ua e = equivEq (funExt (uaβ e)) ua-pathToEquiv : {A B : Type } (p : A B) ua (pathToEquiv p) p ua-pathToEquiv = uaη @@ -230,7 +230,7 @@ (aurefl : {} {A : Type } au refl idEquiv A) where ua-au : {A B : Type } (p : A B) ua (au p) p - ua-au {B = B} = J _ p ua (au p) p) + ua-au {B = B} = J _ p ua (au p) p) (cong ua aurefl uaIdEquiv) au-ua : {A B : Type } (e : A B) au (ua e) e @@ -248,10 +248,10 @@ -- The original map from UniMath/Foundations eqweqmap : {A B : Type } A B A B -eqweqmap {A = A} e = J X _ A X) (idEquiv A) e +eqweqmap {A = A} e = J X _ A X) (idEquiv A) e eqweqmapid : {A : Type } eqweqmap refl idEquiv A -eqweqmapid {A = A} = JRefl X _ A X) (idEquiv A) +eqweqmapid {A = A} = JRefl X _ A X) (idEquiv A) univalenceStatement : {A B : Type } isEquiv (eqweqmap {} {A} {B}) univalenceStatement = Univalence.thm eqweqmap eqweqmapid @@ -259,8 +259,8 @@ univalenceUAH : {A B : Type } (A B) (A B) univalenceUAH = ( _ , univalenceStatement ) -univalencePath : {A B : Type } (A B) Lift (A B) -univalencePath = ua (compEquiv univalence LiftEquiv) +univalencePath : {A B : Type } (A B) Lift (A B) +univalencePath = ua (compEquiv univalence LiftEquiv) -- Lemmas for constructing and destructing dependent paths in a function type where the domain is ua. ua→ : { ℓ'} {A₀ A₁ : Type } {e : A₀ A₁} {B : (i : I) Type ℓ'} @@ -276,7 +276,7 @@ (h (transp j ua e (~ j i)) (~ i) a) i) where lem : a₁ e .fst (transport (sym (ua e)) a₁) a₁ - lem a₁ = secEq e _ transportRefl _ + lem a₁ = secEq e _ transportRefl _ ua→⁻ : { ℓ'} {A₀ A₁ : Type } {e : A₀ A₁} {B : (i : I) Type ℓ'} {f₀ : A₀ B i0} {f₁ : A₁ B i1} @@ -301,16 +301,16 @@ -- Useful lemma for unfolding a transported function over ua -- If we would have regularity this would be refl transportUAop₁ : {A B : Type } (e : A B) (f : A A) (x : B) - transport i ua e i ua e i) f x equivFun e (f (invEq e x)) -transportUAop₁ e f x i = transportRefl (equivFun e (f (invEq e (transportRefl x i)))) i + transport i ua e i ua e i) f x equivFun e (f (invEq e x)) +transportUAop₁ e f x i = transportRefl (equivFun e (f (invEq e (transportRefl x i)))) i -- Binary version transportUAop₂ : {} {A B : Type } (e : A B) (f : A A A) (x y : B) transport i ua e i ua e i ua e i) f x y - equivFun e (f (invEq e x) (invEq e y)) + equivFun e (f (invEq e x) (invEq e y)) transportUAop₂ e f x y i = - transportRefl (equivFun e (f (invEq e (transportRefl x i)) - (invEq e (transportRefl y i)))) i + transportRefl (equivFun e (f (invEq e (transportRefl x i)) + (invEq e (transportRefl y i)))) i -- Alternative version of EquivJ that only requires a predicate on functions elimEquivFun : {A B : Type } (P : (A : Type ) (A B) Type ℓ') @@ -333,13 +333,13 @@ rem1 f g sfg rfg = elimEquivFun P rem (f , isoToIsEquiv (iso f g sfg rfg)) g sfg rfg -uaInvEquiv : {A B : Type } (e : A B) ua (invEquiv e) sym (ua e) -uaInvEquiv {B = B} = EquivJ _ e ua (invEquiv e) sym (ua e)) - (cong ua (invEquivIdEquiv B)) +uaInvEquiv : {A B : Type } (e : A B) ua (invEquiv e) sym (ua e) +uaInvEquiv {B = B} = EquivJ _ e ua (invEquiv e) sym (ua e)) + (cong ua (invEquivIdEquiv B)) -uaCompEquiv : {A B C : Type } (e : A B) (f : B C) ua (compEquiv e f) ua e ua f -uaCompEquiv {B = B} {C} = EquivJ _ e (f : B C) ua (compEquiv e f) ua e ua f) - f cong ua (compEquivIdEquiv f) +uaCompEquiv : {A B C : Type } (e : A B) (f : B C) ua (compEquiv e f) ua e ua f +uaCompEquiv {B = B} {C} = EquivJ _ e (f : B C) ua (compEquiv e f) ua e ua f) + f cong ua (compEquivIdEquiv f) sym (cong x x ua f) uaIdEquiv sym (lUnit (ua f)))) \ No newline at end of file diff --git a/docs/Cubical.Functions.Embedding.html b/docs/Cubical.Functions.Embedding.html index e82e004..91346ed 100644 --- a/docs/Cubical.Functions.Embedding.html +++ b/docs/Cubical.Functions.Embedding.html @@ -45,7 +45,7 @@ isEmbedding : (A B) Type _ isEmbedding f = w x isEquiv {A = w x} (cong f) -isPropIsEmbedding : isProp (isEmbedding f) +isPropIsEmbedding : isProp (isEmbedding f) isPropIsEmbedding {f = f} = isPropΠ2 λ _ _ isPropIsEquiv (cong f) -- Embedding is injection in the aforementioned sense: @@ -62,24 +62,24 @@ -- If `f` is an embedding, we'd expect the fibers of `f` to be -- propositions, like an injective function. hasPropFibers : (A B) Type _ -hasPropFibers f = y isProp (fiber f y) +hasPropFibers f = y isProp (fiber f y) -- This can be relaxed to having all prop fibers over the image, see [hasPropFibersOfImage→isEmbedding] hasPropFibersOfImage : (A B) Type _ -hasPropFibersOfImage f = x isProp (fiber f (f x)) +hasPropFibersOfImage f = x isProp (fiber f (f x)) -- some notation _↪_ : Type ℓ' Type ℓ'' Type (ℓ-max ℓ' ℓ'') A B = Σ[ f (A B) ] isEmbedding f -hasPropFibersIsProp : isProp (hasPropFibers f) -hasPropFibersIsProp = isPropΠ _ isPropIsProp) +hasPropFibersIsProp : isProp (hasPropFibers f) +hasPropFibersIsProp = isPropΠ _ isPropIsProp) private lemma₀ : (p : y z) fiber f y fiber f z lemma₀ {f = f} p = λ i fiber f (p i) - lemma₁ : isEmbedding f x isContr (fiber f (f x)) + lemma₁ : isEmbedding f x isContr (fiber f (f x)) lemma₁ {f = f} iE x = value , path where value : fiber f (f x) @@ -98,7 +98,7 @@ isEmbedding→hasPropFibers : isEmbedding f hasPropFibers f isEmbedding→hasPropFibers iE y (x , p) - = subst f isProp f) (lemma₀ p) (isContr→isProp (lemma₁ iE x)) (x , p) + = subst f isProp f) (lemma₀ p) (isContr→isProp (lemma₁ iE x)) (x , p) private fibCong→PathP @@ -124,7 +124,7 @@ hasPropFibers→isEmbedding : hasPropFibers f isEmbedding f hasPropFibers→isEmbedding {f = f} iP w x .equiv-proof p - = subst isContr (PathP≡fibCong p) (isProp→isContrPathP i iP (p i)) fw fx) + = subst isContr (PathP≡fibCong p) (isProp→isContrPathP i iP (p i)) fw fx) where fw : fiber f (f w) fw = (w , refl) @@ -134,7 +134,7 @@ hasPropFibersOfImage→hasPropFibers : hasPropFibersOfImage f hasPropFibers f hasPropFibersOfImage→hasPropFibers {f = f} fibImg y a b = - subst y isProp (fiber f y)) (snd a) (fibImg (fst a)) a b + subst y isProp (fiber f y)) (snd a) (fibImg (fst a)) a b hasPropFibersOfImage→isEmbedding : hasPropFibersOfImage f isEmbedding f hasPropFibersOfImage→isEmbedding = hasPropFibers→isEmbedding hasPropFibersOfImage→hasPropFibers @@ -151,7 +151,7 @@ -- implies isEmbedding as long as B is an h-set. module _ {f : A B} - (isSetB : isSet B) + (isSetB : isSet B) where module _ @@ -174,7 +174,7 @@ inj {w = w} {x = x} p = sym (ret w) ∙∙ cong g p ∙∙ ret x isEquiv→hasPropFibers : isEquiv f hasPropFibers f -isEquiv→hasPropFibers e b = isContr→isProp (equiv-proof e b) +isEquiv→hasPropFibers e b = isContr→isProp (equiv-proof e b) isEquiv→isEmbedding : isEquiv f isEmbedding f isEquiv→isEmbedding e = λ _ _ congEquiv (_ , e) .snd @@ -202,29 +202,29 @@ Embedding-into-Discrete→Discrete : A B Discrete B Discrete A Embedding-into-Discrete→Discrete (f , isEmbeddingF) _≟_ x y with f x f y -... | yes p = yes (invIsEq (isEmbeddingF x y) p) +... | yes p = yes (invIsEq (isEmbeddingF x y) p) ... | no ¬p = no (¬p cong f) -Embedding-into-isProp→isProp : A B isProp B isProp A +Embedding-into-isProp→isProp : A B isProp B isProp A Embedding-into-isProp→isProp (f , isEmbeddingF) isProp-B x y - = invIsEq (isEmbeddingF x y) (isProp-B (f x) (f y)) + = invIsEq (isEmbeddingF x y) (isProp-B (f x) (f y)) -Embedding-into-isSet→isSet : A B isSet B isSet A +Embedding-into-isSet→isSet : A B isSet B isSet A Embedding-into-isSet→isSet (f , isEmbeddingF) isSet-B x y p q = - p ≡⟨ sym (retIsEq isEquiv-cong-f p) + p ≡⟨ sym (retIsEq isEquiv-cong-f p) cong-f⁻¹ (cong f p) ≡⟨ cong cong-f⁻¹ cong-f-p≡cong-f-q - cong-f⁻¹ (cong f q) ≡⟨ retIsEq isEquiv-cong-f q + cong-f⁻¹ (cong f q) ≡⟨ retIsEq isEquiv-cong-f q q where cong-f-p≡cong-f-q = isSet-B (f x) (f y) (cong f p) (cong f q) isEquiv-cong-f = isEmbeddingF x y - cong-f⁻¹ = invIsEq isEquiv-cong-f + cong-f⁻¹ = invIsEq isEquiv-cong-f Embedding-into-hLevel→hLevel : n A B isOfHLevel (suc n) B isOfHLevel (suc n) A Embedding-into-hLevel→hLevel zero = Embedding-into-isProp→isProp Embedding-into-hLevel→hLevel (suc n) (f , isEmbeddingF) Blvl x y - = isOfHLevelRespectEquiv (suc n) (invEquiv equiv) subLvl + = isOfHLevelRespectEquiv (suc n) (invEquiv equiv) subLvl where equiv : (x y) (f x f y) equiv .fst = cong f @@ -246,11 +246,11 @@ Ψ w x = isEmbeddingFstΣProp (∈-isProp A) Subset→Embedding→Subset : {X : Type } section (Embedding→Subset {} {X}) (Subset→Embedding {} {X}) -Subset→Embedding→Subset _ = funExt λ x Σ≡Prop _ isPropIsProp) (ua (FiberIso.fiberEquiv _ x)) +Subset→Embedding→Subset _ = funExt λ x Σ≡Prop _ isPropIsProp) (ua (FiberIso.fiberEquiv _ x)) Embedding→Subset→Embedding : {X : Type } retract (Embedding→Subset {} {X}) (Subset→Embedding {} {X}) Embedding→Subset→Embedding { = } {X = X} (A , f , ψ) = - cong (equivFun Σ-assoc-≃) (Σ≡Prop _ isPropIsEmbedding) (retEq (fibrationEquiv X ) (A , f))) + cong (equivFun Σ-assoc-≃) (Σ≡Prop _ isPropIsEmbedding) (retEq (fibrationEquiv X ) (A , f))) Subset≃Embedding : {X : Type } X (Σ[ A Type ] (A X)) Subset≃Embedding = isoToEquiv (iso Subset→Embedding Embedding→Subset @@ -261,7 +261,7 @@ isEmbedding-∘ : isEmbedding f isEmbedding h isEmbedding (f h) isEmbedding-∘ {f = f} {h = h} Embf Embh w x - = compEquiv (cong h , Embh w x) (cong f , Embf (h w) (h x)) .snd + = compEquiv (cong h , Embh w x) (cong f , Embf (h w) (h x)) .snd compEmbedding : (B C) (A B) (A C) (compEmbedding (g , _ ) (f , _ )).fst = g f @@ -269,9 +269,9 @@ isEmbedding→embedsFibersIntoSingl : isEmbedding f - z fiber f z singl z + z fiber f z singl z isEmbedding→embedsFibersIntoSingl {f = f} isE z = e , isEmbE where - e : fiber f z singl z + e : fiber f z singl z e x = f (fst x) , sym (snd x) isEmbE : isEmbedding e @@ -306,7 +306,7 @@ isEmbedding→hasPropFibers′ : isEmbedding f hasPropFibers f isEmbedding→hasPropFibers′ {f = f} iE z = - Embedding-into-isProp→isProp (isEmbedding→embedsFibersIntoSingl iE z) isPropSingl + Embedding-into-isProp→isProp (isEmbedding→embedsFibersIntoSingl iE z) isPropSingl universeEmbedding : { ℓ' : Level} @@ -315,19 +315,19 @@ isEmbedding F universeEmbedding F liftingEquiv = hasPropFibersOfImage→isEmbedding propFibersF where lemma : A B (F A F B) (B A) - lemma A B = (F A F B) ≃⟨ univalence - (F A F B) ≃⟨ equivComp (liftingEquiv A) (liftingEquiv B) - (A B) ≃⟨ invEquivEquiv - (B A) ≃⟨ invEquiv univalence - (B A) - fiberSingl : X fiber F (F X) singl X + lemma A B = (F A F B) ≃⟨ univalence + (F A F B) ≃⟨ equivComp (liftingEquiv A) (liftingEquiv B) + (A B) ≃⟨ invEquivEquiv + (B A) ≃⟨ invEquiv univalence + (B A) + fiberSingl : X fiber F (F X) singl X fiberSingl X = Σ-cong-equiv-snd _ lemma _ _) propFibersF : hasPropFibersOfImage F - propFibersF X = Embedding-into-isProp→isProp (Equiv→Embedding (fiberSingl X)) isPropSingl + propFibersF X = Embedding-into-isProp→isProp (Equiv→Embedding (fiberSingl X)) isPropSingl liftEmbedding : ( ℓ' : Level) - isEmbedding (Lift {i = } {j = ℓ'}) -liftEmbedding ℓ' = universeEmbedding (Lift {j = ℓ'}) _ invEquiv LiftEquiv) + isEmbedding (Lift {i = } {j = ℓ'}) +liftEmbedding ℓ' = universeEmbedding (Lift {j = ℓ'}) _ invEquiv LiftEquiv) module FibrationIdentityPrinciple {B : Type } {ℓ'} where -- note that fibrationEquiv (for good reason) uses ℓ' = ℓ-max ℓ ℓ', so we have to work @@ -343,38 +343,38 @@ Fibration′IP : f≃g′ (f g) Fibration′IP = f≃g′ - ≃⟨ equivΠCod _ invEquiv univalence) + ≃⟨ equivΠCod _ invEquiv univalence) (∀ b fiber (f .snd) b fiber (g .snd) b) - ≃⟨ funExtEquiv + ≃⟨ funExtEquiv fiber (f .snd) fiber (g .snd) - ≃⟨ invEquiv (congEquiv (fibrationEquiv B ℓ')) + ≃⟨ invEquiv (congEquiv (fibrationEquiv B ℓ')) f g - + -- Then embed into the above case by lifting the type L : Type _ Type _ -- local synonym fixing the levels of Lift - L = Lift {i = ℓ'} {j = } + L = Lift {i = ℓ'} {j = } liftFibration : Fibration B ℓ' Fibration′ - liftFibration (A , f) = L A , f lower + liftFibration (A , f) = L A , f lower hasPropFibersLiftFibration : hasPropFibers liftFibration hasPropFibersLiftFibration (A , f) = Embedding-into-isProp→isProp (Equiv→Embedding fiberChar) (isPropΣ (isEmbedding→hasPropFibers (liftEmbedding _ _) A) - λ _ isEquiv→hasPropFibers (snd (invEquiv (preCompEquiv LiftEquiv))) _) + λ _ isEquiv→hasPropFibers (snd (invEquiv (preCompEquiv LiftEquiv))) _) where fiberChar : fiber liftFibration (A , f) - (Σ[ (E , eq) fiber L A ] fiber (_∘ lower) (transport⁻ i eq i B) f)) + (Σ[ (E , eq) fiber L A ] fiber (_∘ lower) (transport⁻ i eq i B) f)) fiberChar = fiber liftFibration (A , f) - ≃⟨ Σ-cong-equiv-snd _ invEquiv ΣPath≃PathΣ) - (Σ[ (E , g) Fibration B ℓ' ] Σ[ eq (L E A) ] PathP i eq i B) (g lower) f) - ≃⟨ boringSwap - (Σ[ (E , eq) fiber L A ] Σ[ g (E B) ] PathP i eq i B) (g lower) f) - ≃⟨ Σ-cong-equiv-snd _ Σ-cong-equiv-snd λ _ pathToEquiv (PathP≡Path⁻ _ _ _)) - (Σ[ (E , eq) fiber L A ] fiber (_∘ lower) (transport⁻ i eq i B) f)) - where + ≃⟨ Σ-cong-equiv-snd _ invEquiv ΣPath≃PathΣ) + (Σ[ (E , g) Fibration B ℓ' ] Σ[ eq (L E A) ] PathP i eq i B) (g lower) f) + ≃⟨ boringSwap + (Σ[ (E , eq) fiber L A ] Σ[ g (E B) ] PathP i eq i B) (g lower) f) + ≃⟨ Σ-cong-equiv-snd _ Σ-cong-equiv-snd λ _ pathToEquiv (PathP≡Path⁻ _ _ _)) + (Σ[ (E , eq) fiber L A ] fiber (_∘ lower) (transport⁻ i eq i B) f)) + where unquoteDecl boringSwap = declStrictEquiv boringSwap ((E , g) , (eq , p)) ((E , eq) , (g , p))) @@ -391,11 +391,11 @@ FibrationIP : f≃g (f g) FibrationIP = - f≃g ≃⟨ equivΠCod b equivComp (Σ-cong-equiv-fst LiftEquiv) - (Σ-cong-equiv-fst LiftEquiv)) - f≃g′ ≃⟨ Fibration′IP - (liftFibration f liftFibration g) ≃⟨ invEquiv (_ , isEmbeddingLiftFibration _ _) - (f g) + f≃g ≃⟨ equivΠCod b equivComp (Σ-cong-equiv-fst LiftEquiv) + (Σ-cong-equiv-fst LiftEquiv)) + f≃g′ ≃⟨ Fibration′IP + (liftFibration f liftFibration g) ≃⟨ invEquiv (_ , isEmbeddingLiftFibration _ _) + (f g) _≃Fib_ : {B : Type } (f g : Fibration B ℓ') Type (ℓ-max ℓ') _≃Fib_ = FibrationIdentityPrinciple.f≃g @@ -421,21 +421,21 @@ isEmbeddingToFibr w x = fullEquiv .snd where -- carefully managed such that (cong toFibr) is the equivalence fullEquiv : (w x) (toFibr w toFibr x) - fullEquiv = compEquiv (congEquiv (invEquiv Σ-assoc-≃)) (invEquiv (Σ≡PropEquiv _ isPropIsEmbedding))) + fullEquiv = compEquiv (congEquiv (invEquiv Σ-assoc-≃)) (invEquiv (Σ≡PropEquiv _ isPropIsEmbedding))) EmbeddingIP : f≃g (f g) EmbeddingIP = f≃g - ≃⟨ strictIsoToEquiv (invIso toProdIso) + ≃⟨ strictIsoToEquiv (invIso toProdIso) (∀ b (fiber ffun b fiber gfun b) × (fiber gfun b fiber ffun b)) - ≃⟨ equivΠCod _ isEquivPropBiimpl→Equiv (isEmbedding→hasPropFibers isEmbF _) - (isEmbedding→hasPropFibers isEmbG _)) + ≃⟨ equivΠCod _ isEquivPropBiimpl→Equiv (isEmbedding→hasPropFibers isEmbF _) + (isEmbedding→hasPropFibers isEmbG _)) (∀ b (fiber (f .snd .fst) b) (fiber (g .snd .fst) b)) - ≃⟨ FibrationIP (toFibr f) (toFibr g) + ≃⟨ FibrationIP (toFibr f) (toFibr g) (toFibr f toFibr g) - ≃⟨ invEquiv (_ , isEmbeddingToFibr _ _) + ≃⟨ invEquiv (_ , isEmbeddingToFibr _ _) f g - + _≃Emb_ : {B : Type } (f g : Embedding B ℓ') Type _ _≃Emb_ = EmbeddingIdentityPrinciple.f≃g @@ -444,7 +444,7 @@ EmbeddingIP = EmbeddingIdentityPrinciple.EmbeddingIP -- Cantor's theorem for sets -Set-Embedding-into-Powerset : {A : Type } isSet A A A +Set-Embedding-into-Powerset : {A : Type } isSet A A A Set-Embedding-into-Powerset {A = A} setA = fun , (injEmbedding isSetℙ y sym (H₃ (H₂ y)))) where fun : A A @@ -465,14 +465,14 @@ apmap x y x≡y = ΣPathP (cong (f fst) x≡y , cong (g snd) x≡y) equiv : x y isEquiv (apmap x y) - equiv x y = ((invEquiv ΣPathP≃PathPΣ) - ∙ₑ (≃-× ((cong f) , (embf (fst x) (fst y))) + equiv x y = ((invEquiv ΣPathP≃PathPΣ) + ∙ₑ (≃-× ((cong f) , (embf (fst x) (fst y))) ((cong g) , (embg (snd x) (snd y)))) - ∙ₑ ΣPathP≃PathPΣ) .snd + ∙ₑ ΣPathP≃PathPΣ) .snd emb : isEmbedding (map-× f g) emb x y = equiv x y -EmbeddingΣProp : {A : Type } {B : A Type ℓ'} (∀ a isProp (B a)) Σ A B A +EmbeddingΣProp : {A : Type } {B : A Type ℓ'} (∀ a isProp (B a)) Σ A B A EmbeddingΣProp f = fst , _ _ isEmbeddingFstΣProp f) \ No newline at end of file diff --git a/docs/Cubical.Functions.Fibration.html b/docs/Cubical.Functions.Fibration.html index 45af7e9..7edb58c 100644 --- a/docs/Cubical.Functions.Fibration.html +++ b/docs/Cubical.Functions.Fibration.html @@ -3,7 +3,7 @@ module Cubical.Functions.Fibration where open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels using (isOfHLevel→isOfHLevelDep) +open import Cubical.Foundations.HLevels using (isOfHLevel→isOfHLevelDep) open import Cubical.Foundations.Function open import Cubical.Foundations.GroupoidLaws open import Cubical.Foundations.Equiv @@ -35,13 +35,13 @@ bwd-fwd : x bwd (fwd x) x bwd-fwd ((x' , y) , q) i = h (r i) - where h : Σ[ s singl x ] p⁻¹ (s .fst) fiber p x + where h : Σ[ s singl x ] p⁻¹ (s .fst) fiber p x h ((x , p) , y) = (x , y) , sym p - r : Path (Σ[ s singl x ] p⁻¹ (s .fst)) + r : Path (Σ[ s singl x ] p⁻¹ (s .fst)) ((x , refl ) , subst p⁻¹ q y) ((x' , sym q) , y ) - r = ΣPathP (isContrSingl x .snd (x' , sym q) - , toPathP (transport⁻Transport i p⁻¹ (q i)) y)) + r = ΣPathP (isContrSingl x .snd (x' , sym q) + , toPathP (transport⁻Transport i p⁻¹ (q i)) y)) -- HoTT Lemma 4.8.1 fiberEquiv : fiber p x p⁻¹ x @@ -75,13 +75,13 @@ where e = totalEquiv p -module ForSets {E : Type } {isSetB : isSet B} (f : E B) where +module ForSets {E : Type } {isSetB : isSet B} (f : E B) where module _ {x x'} {px : x x'} {a' : fiber f x} {b' : fiber f x'} where -- fibers are equal when their representatives are equal fibersEqIfRepsEq : fst a' fst b' PathP i fiber f (px i)) a' b' fibersEqIfRepsEq p = ΣPathP (p , - (isOfHLevel→isOfHLevelDep 1 + (isOfHLevel→isOfHLevelDep 1 (v , w) isSetB (f v) w) (snd a') (snd b') i (p i , px i)))) diff --git a/docs/Cubical.Functions.Fixpoint.html b/docs/Cubical.Functions.Fixpoint.html index bd73a20..721831e 100644 --- a/docs/Cubical.Functions.Fixpoint.html +++ b/docs/Cubical.Functions.Fixpoint.html @@ -26,14 +26,14 @@ -- Kraus' lemma -- a version not using cubical features can be found at -- https://www.cs.bham.ac.uk/~mhe/GeneralizedHedberg/html/GeneralizedHedberg.html#21576 -2-Constant→isPropFixpoint : (f : A A) 2-Constant f isProp (Fixpoint f) +2-Constant→isPropFixpoint : (f : A A) 2-Constant f isProp (Fixpoint f) 2-Constant→isPropFixpoint f fconst (x , p) (y , q) i = s i , t i where noose : x y f x f y noose x y = sym (fconst x x) fconst x y -- the main idea is that for any path p, cong f p does not depend on p -- but only on its endpoints and the structure of 2-Constant f KrausInsight : {x y} (p : x y) noose x y cong f p - KrausInsight {x} = J y p noose x y cong f p) (lCancel (fconst x x)) + KrausInsight {x} = J y p noose x y cong f p) (lCancel (fconst x x)) -- Need to solve for a path s : x ≡ y, such that: -- transport (λ i → cong f s i ≡ s i) p ≡ q s : x y diff --git a/docs/Cubical.Functions.FunExtEquiv.html b/docs/Cubical.Functions.FunExtEquiv.html index beb457b..4077317 100644 --- a/docs/Cubical.Functions.FunExtEquiv.html +++ b/docs/Cubical.Functions.FunExtEquiv.html @@ -168,7 +168,7 @@ lemi→i : PathP m lemi→j i m p i) (coei→i A i (p i)) refl lemi→i = sym (coei→i k coei→j A i k (p i) p k) i (coei→i A i (p i))) - λ m k lemi→j i (m k) + λ m k lemi→j i (m k) heteroHomotopy≃Homotopy : {A : I Type } {B : (i : I) Type ℓ₁} {f : A i0 B i0} {g : A i1 B i1} @@ -178,18 +178,18 @@ where open Iso isom : Iso _ _ - isom .fun h x₀ = h (isContrSinglP A x₀ .fst .snd) + isom .fun h x₀ = h (isContrSinglP A x₀ .fst .snd) isom .inv k {x₀} {x₁} p = - subst fib PathP B (f x₀) (g (fib .fst))) (isContrSinglP A x₀ .snd (x₁ , p)) (k x₀) + subst fib PathP B (f x₀) (g (fib .fst))) (isContrSinglP A x₀ .snd (x₁ , p)) (k x₀) isom .rightInv k = funExt λ x₀ cong α subst fib PathP B (f x₀) (g (fib .fst))) α (k x₀)) - (isProp→isSet isPropSinglP (isContrSinglP A x₀ .fst) _ - (isContrSinglP A x₀ .snd (isContrSinglP A x₀ .fst)) + (isProp→isSet isPropSinglP (isContrSinglP A x₀ .fst) _ + (isContrSinglP A x₀ .snd (isContrSinglP A x₀ .fst)) refl) transportRefl (k x₀) isom .leftInv h j {x₀} {x₁} p = transp - i PathP B (f x₀) (g (isContrSinglP A x₀ .snd (x₁ , p) (i j) .fst))) + i PathP B (f x₀) (g (isContrSinglP A x₀ .snd (x₁ , p) (i j) .fst))) j - (h (isContrSinglP A x₀ .snd (x₁ , p) j .snd)) + (h (isContrSinglP A x₀ .snd (x₁ , p) j .snd)) \ No newline at end of file diff --git a/docs/Cubical.Functions.Involution.html b/docs/Cubical.Functions.Involution.html index 4a8cc37..4ca249d 100644 --- a/docs/Cubical.Functions.Involution.html +++ b/docs/Cubical.Functions.Involution.html @@ -29,16 +29,16 @@ involPath : A A involPath = ua involEquiv - involEquivComp : compEquiv involEquiv involEquiv idEquiv A + involEquivComp : compEquiv involEquiv involEquiv idEquiv A involEquivComp - = equivEq i x invol x i) + = equivEq i x invol x i) involPathComp : involPath involPath refl involPathComp = sym (uaCompEquiv involEquiv involEquiv) ∙∙ cong ua involEquivComp ∙∙ uaIdEquiv - involPath² : Square involPath refl refl involPath + involPath² : Square involPath refl refl involPath involPath² - = subst s Square involPath s refl involPath) + = subst s Square involPath s refl involPath) involPathComp (compPath-filler involPath involPath) \ No newline at end of file diff --git a/docs/Cubical.Functions.Logic.html b/docs/Cubical.Functions.Logic.html index 9b787e9..ff755ea 100644 --- a/docs/Cubical.Functions.Logic.html +++ b/docs/Cubical.Functions.Logic.html @@ -69,7 +69,7 @@ hProp≡ : P Q P Q hProp≡ = TypeOfHLevel≡ 1 -isProp⟨⟩ : (A : hProp ) isProp A +isProp⟨⟩ : (A : hProp ) isProp A isProp⟨⟩ = snd -------------------------------------------------------------------------------- diff --git a/docs/Cubical.Functions.Surjection.html b/docs/Cubical.Functions.Surjection.html index 0baf98b..16e502f 100644 --- a/docs/Cubical.Functions.Surjection.html +++ b/docs/Cubical.Functions.Surjection.html @@ -33,7 +33,7 @@ section→isSurjection : {g : B A} section f g isSurjection f section→isSurjection {g = g} s b = g b , s b ∣₁ -isPropIsSurjection : isProp (isSurjection f) +isPropIsSurjection : isProp (isSurjection f) isPropIsSurjection = isPropΠ λ _ squash₁ isEquiv→isSurjection : isEquiv f isSurjection f @@ -44,7 +44,7 @@ isEmbedding×isSurjection→isEquiv : isEmbedding f × isSurjection f isEquiv f equiv-proof (isEmbedding×isSurjection→isEquiv {f = f} (emb , sur)) b = - inhProp→isContr (PT.rec fib' x x) fib) fib' + inhProp→isContr (PT.rec fib' x x) fib) fib' where hpf : hasPropFibers f hpf = isEmbedding→hasPropFibers emb @@ -52,7 +52,7 @@ fib : fiber f b ∥₁ fib = sur b - fib' : isProp (fiber f b) + fib' : isProp (fiber f b) fib' = hpf b isEquiv≃isEmbedding×isSurjection : isEquiv f isEmbedding f × isSurjection f diff --git a/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html b/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html index 945ae72..81647aa 100644 --- a/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html +++ b/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html @@ -31,7 +31,7 @@ a = pt A∙ toEquivPtd : A ∥₁ Σ[ B∙ Pointed ] (A , a) B∙ - toEquivPtd = rec isPropSingl x (A , x) , h x) + toEquivPtd = rec isPropSingl x (A , x) , h x) private B∙ : A ∥₁ Pointed B∙ tx = fst (toEquivPtd tx) diff --git a/docs/Cubical.HITs.PropositionalTruncation.Properties.html b/docs/Cubical.HITs.PropositionalTruncation.Properties.html index f10c3bf..01a6127 100644 --- a/docs/Cubical.HITs.PropositionalTruncation.Properties.html +++ b/docs/Cubical.HITs.PropositionalTruncation.Properties.html @@ -31,19 +31,19 @@ A B C : Type A′ : Type ℓ' -∥∥-isPropDep : (P : A Type ) isOfHLevelDep 1 x P x ∥₁) -∥∥-isPropDep P = isOfHLevel→isOfHLevelDep 1 _ squash₁) +∥∥-isPropDep : (P : A Type ) isOfHLevelDep 1 x P x ∥₁) +∥∥-isPropDep P = isOfHLevel→isOfHLevelDep 1 _ squash₁) -rec : {P : Type } isProp P (A P) A ∥₁ P +rec : {P : Type } isProp P (A P) A ∥₁ P rec Pprop f x ∣₁ = f x rec Pprop f (squash₁ x y i) = Pprop (rec Pprop f x) (rec Pprop f y) i -rec2 : {P : Type } isProp P (A B P) A ∥₁ B ∥₁ P +rec2 : {P : Type } isProp P (A B P) A ∥₁ B ∥₁ P rec2 Pprop f x ∣₁ y ∣₁ = f x y rec2 Pprop f x ∣₁ (squash₁ y z i) = Pprop (rec2 Pprop f x ∣₁ y) (rec2 Pprop f x ∣₁ z) i rec2 Pprop f (squash₁ x y i) z = Pprop (rec2 Pprop f x z) (rec2 Pprop f y z) i -rec3 : {P : Type } isProp P (A B C P) A ∥₁ B ∥₁ C ∥₁ P +rec3 : {P : Type } isProp P (A B C P) A ∥₁ B ∥₁ C ∥₁ P rec3 Pprop f x ∣₁ y ∣₁ z ∣₁ = f x y z rec3 Pprop f x ∣₁ y ∣₁ (squash₁ z w i) = Pprop (rec3 Pprop f x ∣₁ y ∣₁ z) (rec3 Pprop f x ∣₁ y ∣₁ w) i rec3 Pprop f x ∣₁ (squash₁ y z i) w = Pprop (rec3 Pprop f x ∣₁ y w) (rec3 Pprop f x ∣₁ z w) i @@ -55,7 +55,7 @@ -- n-ary recursor, stated using a dependent FinVec recFin : {m : } {P : Fin m Type } - {B : Type ℓ'} (isPropB : isProp B) + {B : Type ℓ'} (isPropB : isProp B) ((∀ i P i) B) --------------------- ((∀ i P i ∥₁) B) @@ -68,10 +68,10 @@ famSuc untruncHyp { zero p₀ ; (suc i) famSuc i })) curriedishTrunc : P zero ∥₁ (∀ i P (suc i) ∥₁) B - curriedishTrunc = rec (isProp→ isPropB) curriedish + curriedishTrunc = rec (isProp→ isPropB) curriedish recFin2 : {m1 m2 : } {P : Fin m1 Fin m2 Type } - {B : Type ℓ'} (isPropB : isProp B) + {B : Type ℓ'} (isPropB : isProp B) ((∀ i j P i j) B) -------------------------- (∀ i j P i j ∥₁) @@ -86,18 +86,18 @@ truncFamSuc curriedishTrunc : (∀ j P zero j ∥₁) (∀ i j P (suc i) j ∥₁) B - curriedishTrunc = recFin (isProp→ isPropB) curriedish + curriedishTrunc = recFin (isProp→ isPropB) curriedish -elim : {P : A ∥₁ Type } ((a : A ∥₁) isProp (P a)) +elim : {P : A ∥₁ Type } ((a : A ∥₁) isProp (P a)) ((x : A) P x ∣₁) (a : A ∥₁) P a elim Pprop f x ∣₁ = f x elim Pprop f (squash₁ x y i) = - isOfHLevel→isOfHLevelDep 1 Pprop + isOfHLevel→isOfHLevelDep 1 Pprop (elim Pprop f x) (elim Pprop f y) (squash₁ x y) i elim2 : {P : A ∥₁ B ∥₁ Type } - (Pprop : (x : A ∥₁) (y : B ∥₁) isProp (P x y)) + (Pprop : (x : A ∥₁) (y : B ∥₁) isProp (P x y)) (f : (a : A) (b : B) P a ∣₁ b ∣₁) (x : A ∥₁) (y : B ∥₁) P x y elim2 Pprop f = @@ -105,7 +105,7 @@ a elim _ Pprop _ _) (f a)) elim3 : {P : A ∥₁ B ∥₁ C ∥₁ Type } - (Pprop : ((x : A ∥₁) (y : B ∥₁) (z : C ∥₁) isProp (P x y z))) + (Pprop : ((x : A ∥₁) (y : B ∥₁) (z : C ∥₁) isProp (P x y z))) (g : (a : A) (b : B) (c : C) P ( a ∣₁) b ∣₁ c ∣₁) (x : A ∥₁) (y : B ∥₁) (z : C ∥₁) P x y z elim3 Pprop g = elim2 _ _ isPropΠ _ Pprop _ _ _)) @@ -114,7 +114,7 @@ -- n-ary eliminator, stated using a dependent FinVec elimFin : {m : } {P : Fin m Type } {B : (∀ i P i ∥₁) Type ℓ'} - (isPropB : x isProp (B x)) + (isPropB : x isProp (B x)) ((x : i P i) B i x i ∣₁)) ---------------------------------------- ((x : i P i ∥₁) B x) @@ -136,30 +136,30 @@ λ x₀ xₛ subst B (funExt { zero refl ; (suc i) refl})) (curriedish x₀ xₛ) -isPropPropTrunc : isProp A ∥₁ +isPropPropTrunc : isProp A ∥₁ isPropPropTrunc x y = squash₁ x y propTrunc≃ : A B A ∥₁ B ∥₁ propTrunc≃ e = - propBiimpl→Equiv + propBiimpl→Equiv isPropPropTrunc isPropPropTrunc (rec isPropPropTrunc a e .fst a ∣₁)) - (rec isPropPropTrunc b invEq e b ∣₁)) + (rec isPropPropTrunc b invEq e b ∣₁)) -propTruncIdempotent≃ : isProp A A ∥₁ A +propTruncIdempotent≃ : isProp A A ∥₁ A propTruncIdempotent≃ {A = A} hA = isoToEquiv f where f : Iso A ∥₁ A Iso.fun f = rec hA (idfun A) Iso.inv f x = x ∣₁ Iso.rightInv f _ = refl - Iso.leftInv f = elim _ isProp→isSet isPropPropTrunc _ _) _ refl) + Iso.leftInv f = elim _ isProp→isSet isPropPropTrunc _ _) _ refl) -propTruncIdempotent : isProp A A ∥₁ A +propTruncIdempotent : isProp A A ∥₁ A propTruncIdempotent hA = ua (propTruncIdempotent≃ hA) -- We could also define the eliminator using the recursor -elim' : {P : A ∥₁ Type } ((a : A ∥₁) isProp (P a)) +elim' : {P : A ∥₁ Type } ((a : A ∥₁) isProp (P a)) ((x : A) P x ∣₁) (a : A ∥₁) P a elim' {P = P} Pprop f a = rec (Pprop a) x transp i P (squash₁ x ∣₁ a i)) i0 (f x)) a @@ -175,12 +175,12 @@ -- constant.' The details of this can be found in the following paper: -- -- https://arxiv.org/pdf/1411.2682.pdf -module SetElim (Bset : isSet B) where - Bset' : isSet' B - Bset' = isSet→isSet' Bset +module SetElim (Bset : isSet B) where + Bset' : isSet' B + Bset' = isSet→isSet' Bset - rec→Set : (f : A B) (kf : 2-Constant f) A ∥₁ B - helper : (f : A B) (kf : 2-Constant f) (t u : A ∥₁) + rec→Set : (f : A B) (kf : 2-Constant f) A ∥₁ B + helper : (f : A B) (kf : 2-Constant f) (t u : A ∥₁) rec→Set f kf t rec→Set f kf u rec→Set f kf x ∣₁ = f x @@ -192,14 +192,14 @@ helper f kf t (squash₁ u v i) = Bset' (helper f kf t u) (helper f kf t v) refl (helper f kf u v) i - kcomp : (f : A ∥₁ B) 2-Constant (f ∣_∣₁) + kcomp : (f : A ∥₁ B) 2-Constant (f ∣_∣₁) kcomp f x y = cong f (squash₁ x ∣₁ y ∣₁) - Fset : isSet (A B) - Fset = isSetΠ (const Bset) + Fset : isSet (A B) + Fset = isSetΠ (const Bset) - Kset : (f : A B) isSet (2-Constant f) - Kset f = isSetΠ _ isSetΠ _ isProp→isSet (Bset _ _))) + Kset : (f : A B) isSet (2-Constant f) + Kset f = isSetΠ _ isSetΠ _ isProp→isSet (Bset _ _))) setRecLemma : (f : A ∥₁ B) @@ -208,28 +208,28 @@ = elim {P = λ t rec→Set (f ∣_∣₁) (kcomp f) t f t} t Bset _ _) x refl) t i - mkKmap : ( A ∥₁ B) Σ (A B) 2-Constant + mkKmap : ( A ∥₁ B) Σ (A B) 2-Constant mkKmap f = f ∣_∣₁ , kcomp f - fib : (g : Σ (A B) 2-Constant) fiber mkKmap g + fib : (g : Σ (A B) 2-Constant) fiber mkKmap g fib (g , kg) = rec→Set g kg , refl - eqv : (g : Σ (A B) 2-Constant) fi fib g fi + eqv : (g : Σ (A B) 2-Constant) fi fib g fi eqv g (f , p) = Σ≡Prop f isOfHLevelΣ 2 Fset Kset _ _) (cong (uncurry rec→Set) (sym p) setRecLemma f) - trunc→Set≃ : ( A ∥₁ B) (Σ (A B) 2-Constant) + trunc→Set≃ : ( A ∥₁ B) (Σ (A B) 2-Constant) trunc→Set≃ .fst = mkKmap trunc→Set≃ .snd .equiv-proof g = fib g , eqv g -- The strategy of this equivalence proof follows the paper more closely. -- It is used further down for the groupoid version, because the above -- strategy does not generalize so easily. - e : B Σ (A B) 2-Constant + e : B Σ (A B) 2-Constant e b = const b , λ _ _ refl - eval : A (γ : Σ (A B) 2-Constant) B + eval : A (γ : Σ (A B) 2-Constant) B eval a₀ (g , _) = g a₀ e-eval : (a₀ : A) γ e (eval a₀ γ) γ @@ -239,13 +239,13 @@ e-isEquiv : A isEquiv (e {A = A}) e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ refl) - preEquiv₁ : A ∥₁ B Σ (A B) 2-Constant + preEquiv₁ : A ∥₁ B Σ (A B) 2-Constant preEquiv₁ t = e , rec (isPropIsEquiv e) e-isEquiv t - preEquiv₂ : ( A ∥₁ Σ (A B) 2-Constant) Σ (A B) 2-Constant + preEquiv₂ : ( A ∥₁ Σ (A B) 2-Constant) Σ (A B) 2-Constant preEquiv₂ = isoToEquiv (iso to const _ refl) retr) where - to : ( A ∥₁ Σ (A B) 2-Constant) Σ (A B) 2-Constant + to : ( A ∥₁ Σ (A B) 2-Constant) Σ (A B) 2-Constant to f .fst x = f x ∣₁ .fst x to f .snd x y i = f (squash₁ x ∣₁ y ∣₁ i) .snd x y i @@ -259,14 +259,14 @@ j f (squash₁ y ∣₁ t j) .fst y) i - trunc→Set≃₂ : ( A ∥₁ B) Σ (A B) 2-Constant - trunc→Set≃₂ = compEquiv (equivΠCod preEquiv₁) preEquiv₂ + trunc→Set≃₂ : ( A ∥₁ B) Σ (A B) 2-Constant + trunc→Set≃₂ = compEquiv (equivΠCod preEquiv₁) preEquiv₂ open SetElim public using (rec→Set; trunc→Set≃) elim→Set : {P : A ∥₁ Type } - (∀ t isSet (P t)) + (∀ t isSet (P t)) (f : (x : A) P x ∣₁) (kf : x y PathP i P (squash₁ x ∣₁ y ∣₁ i)) (f x) (f y)) (t : A ∥₁) P t @@ -276,20 +276,20 @@ g : A P t g x = transp i P (squash₁ x ∣₁ t i)) i0 (f x) - gk : 2-Constant g + gk : 2-Constant g gk x y i = transp j P (squash₁ (squash₁ x ∣₁ y ∣₁ i) t j)) i0 (kf x y i) elim2→Set : {P : A ∥₁ B ∥₁ Type } - (∀ t u isSet (P t u)) + (∀ t u isSet (P t u)) (f : (x : A) (y : B) P x ∣₁ y ∣₁) (kf₁ : x y v PathP i P (squash₁ x ∣₁ y ∣₁ i) v ∣₁) (f x v) (f y v)) (kf₂ : x v w PathP i P x ∣₁ (squash₁ v ∣₁ w ∣₁ i)) (f x v) (f x w)) - (sf : x y v w SquareP i j P (squash₁ x ∣₁ y ∣₁ i) (squash₁ v ∣₁ w ∣₁ j)) + (sf : x y v w SquareP i j P (squash₁ x ∣₁ y ∣₁ i) (squash₁ v ∣₁ w ∣₁ j)) (kf₂ x v w) (kf₂ y v w) (kf₁ x y v) (kf₁ x y w)) (t : A ∥₁) (u : B ∥₁) P t u elim2→Set {A = A} {B = B} {P = P} Pset f kf₁ kf₂ sf = - elim→Set _ isSetΠ _ Pset _ _)) mapHelper squareHelper + elim→Set _ isSetΠ _ Pset _ _)) mapHelper squareHelper where mapHelper : (x : A) (u : B ∥₁) P x ∣₁ u mapHelper x = elim→Set _ Pset _ _) (f x) (kf₂ x) @@ -299,272 +299,268 @@ squareHelper x y i = elim→Set _ Pset _ _) v kf₁ x y v i) λ v w sf x y v w i RecHProp : (P : A hProp ) (kP : x y P x P y) A ∥₁ hProp -RecHProp P kP = rec→Set isSetHProp P kP - -module GpdElim (Bgpd : isGroupoid B) where - Bgpd' : isGroupoid' B - Bgpd' = isGroupoid→isGroupoid' Bgpd - - module _ (f : A B) (3kf : 3-Constant f) where - open 3-Constant 3kf - - rec→Gpd : A ∥₁ B - pathHelper : (t u : A ∥₁) rec→Gpd t rec→Gpd u - triHelper₁ - : (t u v : A ∥₁) - Square (pathHelper t u) (pathHelper t v) refl (pathHelper u v) - triHelper₂ - : (t u v : A ∥₁) - Square (pathHelper t v) (pathHelper u v) (pathHelper t u) refl - - rec→Gpd x ∣₁ = f x - rec→Gpd (squash₁ t u i) = pathHelper t u i - - pathHelper x ∣₁ y ∣₁ = link x y - pathHelper (squash₁ t u j) v = triHelper₂ t u v j - pathHelper x ∣₁ (squash₁ u v j) = triHelper₁ x ∣₁ u v j - - triHelper₁ x ∣₁ y ∣₁ z ∣₁ = coh₁ x y z - triHelper₁ (squash₁ s t i) u v - = Bgpd' - (triHelper₁ s u v) - (triHelper₁ t u v) - (triHelper₂ s t u) - (triHelper₂ s t v) - i refl) - i pathHelper u v) - i - triHelper₁ x ∣₁ (squash₁ t u i) v - = Bgpd' - (triHelper₁ x ∣₁ t v) - (triHelper₁ x ∣₁ u v) - (triHelper₁ x ∣₁ t u) - i pathHelper x ∣₁ v) - i refl) - (triHelper₂ t u v) - i - triHelper₁ x ∣₁ y ∣₁ (squash₁ u v i) - = Bgpd' - (triHelper₁ x ∣₁ y ∣₁ u) - (triHelper₁ x ∣₁ y ∣₁ v) - i link x y) - (triHelper₁ x ∣₁ u v) - i refl) - (triHelper₁ y ∣₁ u v) - i - - triHelper₂ x ∣₁ y ∣₁ z ∣₁ = coh₂ x y z - triHelper₂ (squash₁ s t i) u v - = Bgpd' - (triHelper₂ s u v) - (triHelper₂ t u v) - (triHelper₂ s t v) - i pathHelper u v) - (triHelper₂ s t u) - i refl) - i - triHelper₂ x ∣₁ (squash₁ t u i) v - = Bgpd' - (triHelper₂ x ∣₁ t v) - (triHelper₂ x ∣₁ u v) - i pathHelper x ∣₁ v) - (triHelper₂ t u v) - (triHelper₁ x ∣₁ t u) - i refl) - i - triHelper₂ x ∣₁ y ∣₁ (squash₁ u v i) - = Bgpd' - (triHelper₂ x ∣₁ y ∣₁ u) - (triHelper₂ x ∣₁ y ∣₁ v) - (triHelper₁ x ∣₁ u v) - (triHelper₁ y ∣₁ u v) - i link x y) - i refl) - i - - preEquiv₁ : ( A ∥₁ Σ (A B) 3-Constant) Σ (A B) 3-Constant - preEquiv₁ = isoToEquiv (iso fn const _ refl) retr) - where - open 3-Constant - - fn : ( A ∥₁ Σ (A B) 3-Constant) Σ (A B) 3-Constant - fn f .fst x = f x ∣₁ .fst x - fn f .snd .link x y i = f (squash₁ x ∣₁ y ∣₁ i) .snd .link x y i - fn f .snd .coh₁ x y z i j - = f (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j) .snd .coh₁ x y z i j - - retr : retract fn const - retr f i t .fst x = f (squash₁ x ∣₁ t i) .fst x - retr f i t .snd .link x y j - = f (squash₁ (squash₁ x ∣₁ y ∣₁ j) t i) .snd .link x y j - retr f i t .snd .coh₁ x y z - = Bgpd' - k j f (cb k j i0) .snd .coh₁ x y z k j ) - k j f (cb k j i1) .snd .coh₁ x y z k j) - k j f (cb i0 j k) .snd .link x y j) - k j f (cb i1 j k) .snd .link x z j) - _ refl) - k j f (cb j i1 k) .snd .link y z j) - i - where - cb : I I I _ ∥₁ - cb i j k = squash₁ (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j) t k - - e : B Σ (A B) 3-Constant - e b .fst _ = b - e b .snd = record - { link = λ _ _ _ b - ; coh₁ = λ _ _ _ _ _ b - } - - eval : A Σ (A B) 3-Constant B - eval a₀ (g , _) = g a₀ - - module _ where - open 3-Constant - e-eval : ∀(a₀ : A) γ e (eval a₀ γ) γ - e-eval a₀ (g , 3kg) i .fst x = 3kg .link a₀ x i - e-eval a₀ (g , 3kg) i .snd .link x y = λ j 3kg .coh₁ a₀ x y j i - e-eval a₀ (g , 3kg) i .snd .coh₁ x y z - = Bgpd' - _ _ g a₀) - (3kg .coh₁ x y z) - k j 3kg .coh₁ a₀ x y j k) - k j 3kg .coh₁ a₀ x z j k) - _ refl) - k j 3kg .coh₁ a₀ y z j k) - i - - e-isEquiv : A isEquiv (e {A = A}) - e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ refl) - - preEquiv₂ : A ∥₁ B Σ (A B) 3-Constant - preEquiv₂ t = e , rec (isPropIsEquiv e) e-isEquiv t - - trunc→Gpd≃ : ( A ∥₁ B) Σ (A B) 3-Constant - trunc→Gpd≃ = compEquiv (equivΠCod preEquiv₂) preEquiv₁ - -open GpdElim using (rec→Gpd; trunc→Gpd≃) public - -squash₁ᵗ - : ∀(x y z : A) - Square (squash₁ x ∣₁ y ∣₁) (squash₁ x ∣₁ z ∣₁) refl (squash₁ y ∣₁ z ∣₁) -squash₁ᵗ x y z i = squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) - -elim→Gpd - : (P : A ∥₁ Type ) - (∀ t isGroupoid (P t)) - (f : (x : A) P x ∣₁) - (kf : x y PathP i P (squash₁ x ∣₁ y ∣₁ i)) (f x) (f y)) - (3kf : x y z - SquareP i j P (squash₁ᵗ x y z i j)) (kf x y) (kf x z) refl (kf y z)) - (t : A ∥₁) P t -elim→Gpd {A = A} P Pgpd f kf 3kf t = rec→Gpd (Pgpd t) g 3kg t - where - g : A P t - g x = transp i P (squash₁ x ∣₁ t i)) i0 (f x) - - open 3-Constant - - 3kg : 3-Constant g - 3kg .link x y i - = transp j P (squash₁ (squash₁ x ∣₁ y ∣₁ i) t j)) i0 (kf x y i) - 3kg .coh₁ x y z i j - = transp k P (squash₁ (squash₁ᵗ x y z i j) t k)) i0 (3kf x y z i j) - -RecHSet : (P : A TypeOfHLevel 2) 3-Constant P A ∥₁ TypeOfHLevel 2 -RecHSet P 3kP = rec→Gpd (isOfHLevelTypeOfHLevel 2) P 3kP - -∥∥-IdempotentL-⊎-≃ : A ∥₁ A′ ∥₁ A A′ ∥₁ -∥∥-IdempotentL-⊎-≃ = isoToEquiv ∥∥-IdempotentL-⊎-Iso - where ∥∥-IdempotentL-⊎-Iso : Iso ( A ∥₁ A′ ∥₁) ( A A′ ∥₁) - Iso.fun ∥∥-IdempotentL-⊎-Iso x = rec squash₁ lem x - where lem : A ∥₁ A′ A A′ ∥₁ - lem (inl x) = map a inl a) x - lem (inr x) = inr x ∣₁ - Iso.inv ∥∥-IdempotentL-⊎-Iso x = map lem x - where lem : A A′ A ∥₁ A′ - lem (inl x) = inl x ∣₁ - lem (inr x) = inr x - Iso.rightInv ∥∥-IdempotentL-⊎-Iso x = squash₁ (Iso.fun ∥∥-IdempotentL-⊎-Iso (Iso.inv ∥∥-IdempotentL-⊎-Iso x)) x - Iso.leftInv ∥∥-IdempotentL-⊎-Iso x = squash₁ (Iso.inv ∥∥-IdempotentL-⊎-Iso (Iso.fun ∥∥-IdempotentL-⊎-Iso x)) x - -∥∥-IdempotentL-⊎ : A ∥₁ A′ ∥₁ A A′ ∥₁ -∥∥-IdempotentL-⊎ = ua ∥∥-IdempotentL-⊎-≃ - -∥∥-IdempotentR-⊎-≃ : A A′ ∥₁ ∥₁ A A′ ∥₁ -∥∥-IdempotentR-⊎-≃ = isoToEquiv ∥∥-IdempotentR-⊎-Iso - where ∥∥-IdempotentR-⊎-Iso : Iso ( A A′ ∥₁ ∥₁) ( A A′ ∥₁) - Iso.fun ∥∥-IdempotentR-⊎-Iso x = rec squash₁ lem x - where lem : A A′ ∥₁ A A′ ∥₁ - lem (inl x) = inl x ∣₁ - lem (inr x) = map a inr a) x - Iso.inv ∥∥-IdempotentR-⊎-Iso x = map lem x - where lem : A A′ A A′ ∥₁ - lem (inl x) = inl x - lem (inr x) = inr x ∣₁ - Iso.rightInv ∥∥-IdempotentR-⊎-Iso x = squash₁ (Iso.fun ∥∥-IdempotentR-⊎-Iso (Iso.inv ∥∥-IdempotentR-⊎-Iso x)) x - Iso.leftInv ∥∥-IdempotentR-⊎-Iso x = squash₁ (Iso.inv ∥∥-IdempotentR-⊎-Iso (Iso.fun ∥∥-IdempotentR-⊎-Iso x)) x - -∥∥-IdempotentR-⊎ : A A′ ∥₁ ∥₁ A A′ ∥₁ -∥∥-IdempotentR-⊎ = ua ∥∥-IdempotentR-⊎-≃ - -∥∥-Idempotent-⊎ : {A : Type } {A′ : Type ℓ'} A ∥₁ A′ ∥₁ ∥₁ A A′ ∥₁ -∥∥-Idempotent-⊎ {A = A} {A′} = A ∥₁ A′ ∥₁ ∥₁ ≡⟨ ∥∥-IdempotentR-⊎ - A ∥₁ A′ ∥₁ ≡⟨ ∥∥-IdempotentL-⊎ - A A′ ∥₁ - -∥∥-IdempotentL-×-≃ : A ∥₁ × A′ ∥₁ A × A′ ∥₁ -∥∥-IdempotentL-×-≃ = isoToEquiv ∥∥-IdempotentL-×-Iso - where ∥∥-IdempotentL-×-Iso : Iso ( A ∥₁ × A′ ∥₁) ( A × A′ ∥₁) - Iso.fun ∥∥-IdempotentL-×-Iso x = rec squash₁ lem x - where lem : A ∥₁ × A′ A × A′ ∥₁ - lem (a , a′) = map2 a a′ a , a′) a a′ ∣₁ - Iso.inv ∥∥-IdempotentL-×-Iso x = map lem x - where lem : A × A′ A ∥₁ × A′ - lem (a , a′) = a ∣₁ , a′ - Iso.rightInv ∥∥-IdempotentL-×-Iso x = squash₁ (Iso.fun ∥∥-IdempotentL-×-Iso (Iso.inv ∥∥-IdempotentL-×-Iso x)) x - Iso.leftInv ∥∥-IdempotentL-×-Iso x = squash₁ (Iso.inv ∥∥-IdempotentL-×-Iso (Iso.fun ∥∥-IdempotentL-×-Iso x)) x - -∥∥-IdempotentL-× : A ∥₁ × A′ ∥₁ A × A′ ∥₁ -∥∥-IdempotentL-× = ua ∥∥-IdempotentL-×-≃ - -∥∥-IdempotentR-×-≃ : A × A′ ∥₁ ∥₁ A × A′ ∥₁ -∥∥-IdempotentR-×-≃ = isoToEquiv ∥∥-IdempotentR-×-Iso - where ∥∥-IdempotentR-×-Iso : Iso ( A × A′ ∥₁ ∥₁) ( A × A′ ∥₁) - Iso.fun ∥∥-IdempotentR-×-Iso x = rec squash₁ lem x - where lem : A × A′ ∥₁ A × A′ ∥₁ - lem (a , a′) = map2 a a′ a , a′) a ∣₁ a′ - Iso.inv ∥∥-IdempotentR-×-Iso x = map lem x - where lem : A × A′ A × A′ ∥₁ - lem (a , a′) = a , a′ ∣₁ - Iso.rightInv ∥∥-IdempotentR-×-Iso x = squash₁ (Iso.fun ∥∥-IdempotentR-×-Iso (Iso.inv ∥∥-IdempotentR-×-Iso x)) x - Iso.leftInv ∥∥-IdempotentR-×-Iso x = squash₁ (Iso.inv ∥∥-IdempotentR-×-Iso (Iso.fun ∥∥-IdempotentR-×-Iso x)) x - -∥∥-IdempotentR-× : A × A′ ∥₁ ∥₁ A × A′ ∥₁ -∥∥-IdempotentR-× = ua ∥∥-IdempotentR-×-≃ - -∥∥-Idempotent-× : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ ∥₁ A × A′ ∥₁ -∥∥-Idempotent-× {A = A} {A′} = A ∥₁ × A′ ∥₁ ∥₁ ≡⟨ ∥∥-IdempotentR-× - A ∥₁ × A′ ∥₁ ≡⟨ ∥∥-IdempotentL-× - A × A′ ∥₁ - -∥∥-Idempotent-×-≃ : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ ∥₁ A × A′ ∥₁ -∥∥-Idempotent-×-≃ {A = A} {A′} = compEquiv ∥∥-IdempotentR-×-≃ ∥∥-IdempotentL-×-≃ - -∥∥-×-≃ : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ A × A′ ∥₁ -∥∥-×-≃ {A = A} {A′} = compEquiv (invEquiv (propTruncIdempotent≃ (isProp× isPropPropTrunc isPropPropTrunc))) ∥∥-Idempotent-×-≃ - -∥∥-× : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ A × A′ ∥₁ -∥∥-× = ua ∥∥-×-≃ - --- using this we get a convenient recursor/eliminator for binary functions into sets -rec2→Set : {A B C : Type } (Cset : isSet C) - (f : A B C) - (∀ (a a' : A) (b b' : B) f a b f a' b') - A ∥₁ B ∥₁ C -rec2→Set {A = A} {B = B} {C = C} Cset f fconst = curry (g ∥∥-×-≃ .fst) - where - g : A × B ∥₁ C - g = rec→Set Cset (uncurry f) λ x y fconst (fst x) (fst y) (snd x) (snd y) +RecHProp P kP = rec→Set isSetHProp P kP + +squash₁ᵗ + : ∀(x y z : A) + Square (squash₁ x ∣₁ y ∣₁) (squash₁ x ∣₁ z ∣₁) refl (squash₁ y ∣₁ z ∣₁) +squash₁ᵗ x y z i = squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) + +module _ (B : A ∥₁ Type ) + (B-gpd : (a : _) isGroupoid (B a)) + (f : (a : A) B a ∣₁) + (f-coh : (x y : A) PathP i B (squash₁ x ∣₁ y ∣₁ i)) (f x) (f y)) + (f-coh-coh : (x y z : A) SquareP + i j B (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j)) + (f-coh x y) (f-coh x z) refl (f-coh y z)) + where + elim→Gpd : (t : A ∥₁) B t + private + pathHelper : (t u : A ∥₁) PathP i B (squash₁ t u i)) (elim→Gpd t) (elim→Gpd u) + triHelper₁ + : (t u v : A ∥₁) + SquareP i j B (squash₁ t (squash₁ u v i) j)) + (pathHelper t u) (pathHelper t v) + refl (pathHelper u v) + triHelper₂ + : (t u v : A ∥₁) + SquareP i j B (squash₁ (squash₁ t u i) v j)) + (pathHelper t v) (pathHelper u v) + (pathHelper t u) refl + triHelper₂Cube : (x y z : A ∥₁) + Cube j k squash₁ x z (k j)) + j k squash₁ y z j) + i k squash₁ x y i) + i k squash₁ x z (i k)) + i j squash₁ x (squash₁ y z j) i) + i j squash₁ (squash₁ x y i) z j) + + elim→Gpd x ∣₁ = f x + elim→Gpd (squash₁ t u i) = pathHelper t u i + triHelper₂Cube x y z = + isProp→PathP _ isOfHLevelPathP 1 (isOfHLevelPath 1 squash₁ _ _) _ _) _ _ + + pathHelper x ∣₁ y ∣₁ = f-coh x y + pathHelper (squash₁ t u j) v = triHelper₂ t u v j + pathHelper x ∣₁ (squash₁ u v j) = triHelper₁ x ∣₁ u v j + + triHelper₁ x ∣₁ y ∣₁ z ∣₁ = f-coh-coh x y z + triHelper₁ (squash₁ s t i) u v + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ s t i) (squash₁ u v i₁) j)) + (triHelper₁ s u v) (triHelper₁ t u v) + (triHelper₂ s t u) + (triHelper₂ s t v) + i j pathHelper s t i) + i j pathHelper u v j) + (B-gpd v) i + + triHelper₁ x ∣₁ (squash₁ t u i) v + = isGroupoid→CubeP i i₁ j B (squash₁ x ∣₁ (squash₁ (squash₁ t u i) v i₁) j)) + (triHelper₁ x ∣₁ t v) (triHelper₁ x ∣₁ u v) + (triHelper₁ x ∣₁ t u) + i j pathHelper x ∣₁ v j) + refl (triHelper₂ t u v) + (B-gpd v) i + triHelper₁ x ∣₁ y ∣₁ (squash₁ u v i) + = isGroupoid→CubeP i i₁ j B (squash₁ x ∣₁ (squash₁ y ∣₁ (squash₁ u v i) i₁) j)) + (triHelper₁ x ∣₁ y ∣₁ u) (triHelper₁ x ∣₁ y ∣₁ v) + i j f-coh x y j) (triHelper₁ x ∣₁ u v) + refl (triHelper₁ y ∣₁ u v) + (B-gpd v) i + triHelper₂ x ∣₁ y ∣₁ z ∣₁ i j = + comp k B (triHelper₂Cube x ∣₁ y ∣₁ z ∣₁ i j k)) + k λ {(i = i0) f-coh x z (k j) + ; (i = i1) f-coh y z j + ; (j = i0) f-coh x y i + ; (j = i1) f-coh x z (i k)}) + (f-coh-coh x y z j i) + triHelper₂ (squash₁ s t i) u v + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ (squash₁ s t i) u i₁) v j)) + (triHelper₂ s u v) (triHelper₂ t u v) + (triHelper₂ s t v) i j pathHelper u v j) + (triHelper₂ s t u) refl + (B-gpd v) i + triHelper₂ x ∣₁ (squash₁ t u i) v + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ x ∣₁ (squash₁ t u i) i₁) v j)) + (triHelper₂ x ∣₁ t v) (triHelper₂ x ∣₁ u v) + i j pathHelper x ∣₁ v j) (triHelper₂ t u v) + (triHelper₁ x ∣₁ t u) refl + (B-gpd v) i + triHelper₂ x ∣₁ y ∣₁ (squash₁ u v i) + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ x ∣₁ y ∣₁ i₁) (squash₁ u v i) j)) + (triHelper₂ x ∣₁ y ∣₁ u) (triHelper₂ x ∣₁ y ∣₁ v) + (triHelper₁ x ∣₁ u v) (triHelper₁ y ∣₁ u v) + refl i j pathHelper u v i) + (B-gpd v) i + + +module GpdElim (Bgpd : isGroupoid B) where + Bgpd' : isGroupoid' B + Bgpd' = isGroupoid→isGroupoid' Bgpd + + module _ (f : A B) (3kf : 3-Constant f) where + open 3-Constant 3kf + + rec→Gpd : A ∥₁ B + rec→Gpd = elim→Gpd _ B) _ Bgpd) f link coh₁ + + preEquiv₁ : ( A ∥₁ Σ (A B) 3-Constant) Σ (A B) 3-Constant + preEquiv₁ = isoToEquiv (iso fn const _ refl) retr) + where + open 3-Constant + + fn : ( A ∥₁ Σ (A B) 3-Constant) Σ (A B) 3-Constant + fn f .fst x = f x ∣₁ .fst x + fn f .snd .link x y i = f (squash₁ x ∣₁ y ∣₁ i) .snd .link x y i + fn f .snd .coh₁ x y z i j + = f (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j) .snd .coh₁ x y z i j + + retr : retract fn const + retr f i t .fst x = f (squash₁ x ∣₁ t i) .fst x + retr f i t .snd .link x y j + = f (squash₁ (squash₁ x ∣₁ y ∣₁ j) t i) .snd .link x y j + retr f i t .snd .coh₁ x y z + = Bgpd' + k j f (cb k j i0) .snd .coh₁ x y z k j ) + k j f (cb k j i1) .snd .coh₁ x y z k j) + k j f (cb i0 j k) .snd .link x y j) + k j f (cb i1 j k) .snd .link x z j) + _ refl) + k j f (cb j i1 k) .snd .link y z j) + i + where + cb : I I I _ ∥₁ + cb i j k = squash₁ (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j) t k + + e : B Σ (A B) 3-Constant + e b .fst _ = b + e b .snd = record + { link = λ _ _ _ b + ; coh₁ = λ _ _ _ _ _ b + } + + eval : A Σ (A B) 3-Constant B + eval a₀ (g , _) = g a₀ + + module _ where + open 3-Constant + e-eval : ∀(a₀ : A) γ e (eval a₀ γ) γ + e-eval a₀ (g , 3kg) i .fst x = 3kg .link a₀ x i + e-eval a₀ (g , 3kg) i .snd .link x y = λ j 3kg .coh₁ a₀ x y j i + e-eval a₀ (g , 3kg) i .snd .coh₁ x y z + = Bgpd' + _ _ g a₀) + (3kg .coh₁ x y z) + k j 3kg .coh₁ a₀ x y j k) + k j 3kg .coh₁ a₀ x z j k) + _ refl) + k j 3kg .coh₁ a₀ y z j k) + i + + e-isEquiv : A isEquiv (e {A = A}) + e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ refl) + + preEquiv₂ : A ∥₁ B Σ (A B) 3-Constant + preEquiv₂ t = e , rec (isPropIsEquiv e) e-isEquiv t + + trunc→Gpd≃ : ( A ∥₁ B) Σ (A B) 3-Constant + trunc→Gpd≃ = compEquiv (equivΠCod preEquiv₂) preEquiv₁ + +open GpdElim using (rec→Gpd; trunc→Gpd≃) public + +RecHSet : (P : A TypeOfHLevel 2) 3-Constant P A ∥₁ TypeOfHLevel 2 +RecHSet P 3kP = rec→Gpd (isOfHLevelTypeOfHLevel 2) P 3kP + +∥∥-IdempotentL-⊎-≃ : A ∥₁ A′ ∥₁ A A′ ∥₁ +∥∥-IdempotentL-⊎-≃ = isoToEquiv ∥∥-IdempotentL-⊎-Iso + where ∥∥-IdempotentL-⊎-Iso : Iso ( A ∥₁ A′ ∥₁) ( A A′ ∥₁) + Iso.fun ∥∥-IdempotentL-⊎-Iso x = rec squash₁ lem x + where lem : A ∥₁ A′ A A′ ∥₁ + lem (inl x) = map a inl a) x + lem (inr x) = inr x ∣₁ + Iso.inv ∥∥-IdempotentL-⊎-Iso x = map lem x + where lem : A A′ A ∥₁ A′ + lem (inl x) = inl x ∣₁ + lem (inr x) = inr x + Iso.rightInv ∥∥-IdempotentL-⊎-Iso x = squash₁ (Iso.fun ∥∥-IdempotentL-⊎-Iso (Iso.inv ∥∥-IdempotentL-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-⊎-Iso x = squash₁ (Iso.inv ∥∥-IdempotentL-⊎-Iso (Iso.fun ∥∥-IdempotentL-⊎-Iso x)) x + +∥∥-IdempotentL-⊎ : A ∥₁ A′ ∥₁ A A′ ∥₁ +∥∥-IdempotentL-⊎ = ua ∥∥-IdempotentL-⊎-≃ + +∥∥-IdempotentR-⊎-≃ : A A′ ∥₁ ∥₁ A A′ ∥₁ +∥∥-IdempotentR-⊎-≃ = isoToEquiv ∥∥-IdempotentR-⊎-Iso + where ∥∥-IdempotentR-⊎-Iso : Iso ( A A′ ∥₁ ∥₁) ( A A′ ∥₁) + Iso.fun ∥∥-IdempotentR-⊎-Iso x = rec squash₁ lem x + where lem : A A′ ∥₁ A A′ ∥₁ + lem (inl x) = inl x ∣₁ + lem (inr x) = map a inr a) x + Iso.inv ∥∥-IdempotentR-⊎-Iso x = map lem x + where lem : A A′ A A′ ∥₁ + lem (inl x) = inl x + lem (inr x) = inr x ∣₁ + Iso.rightInv ∥∥-IdempotentR-⊎-Iso x = squash₁ (Iso.fun ∥∥-IdempotentR-⊎-Iso (Iso.inv ∥∥-IdempotentR-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-⊎-Iso x = squash₁ (Iso.inv ∥∥-IdempotentR-⊎-Iso (Iso.fun ∥∥-IdempotentR-⊎-Iso x)) x + +∥∥-IdempotentR-⊎ : A A′ ∥₁ ∥₁ A A′ ∥₁ +∥∥-IdempotentR-⊎ = ua ∥∥-IdempotentR-⊎-≃ + +∥∥-Idempotent-⊎ : {A : Type } {A′ : Type ℓ'} A ∥₁ A′ ∥₁ ∥₁ A A′ ∥₁ +∥∥-Idempotent-⊎ {A = A} {A′} = A ∥₁ A′ ∥₁ ∥₁ ≡⟨ ∥∥-IdempotentR-⊎ + A ∥₁ A′ ∥₁ ≡⟨ ∥∥-IdempotentL-⊎ + A A′ ∥₁ + +∥∥-IdempotentL-×-≃ : A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-IdempotentL-×-≃ = isoToEquiv ∥∥-IdempotentL-×-Iso + where ∥∥-IdempotentL-×-Iso : Iso ( A ∥₁ × A′ ∥₁) ( A × A′ ∥₁) + Iso.fun ∥∥-IdempotentL-×-Iso x = rec squash₁ lem x + where lem : A ∥₁ × A′ A × A′ ∥₁ + lem (a , a′) = map2 a a′ a , a′) a a′ ∣₁ + Iso.inv ∥∥-IdempotentL-×-Iso x = map lem x + where lem : A × A′ A ∥₁ × A′ + lem (a , a′) = a ∣₁ , a′ + Iso.rightInv ∥∥-IdempotentL-×-Iso x = squash₁ (Iso.fun ∥∥-IdempotentL-×-Iso (Iso.inv ∥∥-IdempotentL-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-×-Iso x = squash₁ (Iso.inv ∥∥-IdempotentL-×-Iso (Iso.fun ∥∥-IdempotentL-×-Iso x)) x + +∥∥-IdempotentL-× : A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-IdempotentL-× = ua ∥∥-IdempotentL-×-≃ + +∥∥-IdempotentR-×-≃ : A × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-IdempotentR-×-≃ = isoToEquiv ∥∥-IdempotentR-×-Iso + where ∥∥-IdempotentR-×-Iso : Iso ( A × A′ ∥₁ ∥₁) ( A × A′ ∥₁) + Iso.fun ∥∥-IdempotentR-×-Iso x = rec squash₁ lem x + where lem : A × A′ ∥₁ A × A′ ∥₁ + lem (a , a′) = map2 a a′ a , a′) a ∣₁ a′ + Iso.inv ∥∥-IdempotentR-×-Iso x = map lem x + where lem : A × A′ A × A′ ∥₁ + lem (a , a′) = a , a′ ∣₁ + Iso.rightInv ∥∥-IdempotentR-×-Iso x = squash₁ (Iso.fun ∥∥-IdempotentR-×-Iso (Iso.inv ∥∥-IdempotentR-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-×-Iso x = squash₁ (Iso.inv ∥∥-IdempotentR-×-Iso (Iso.fun ∥∥-IdempotentR-×-Iso x)) x + +∥∥-IdempotentR-× : A × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-IdempotentR-× = ua ∥∥-IdempotentR-×-≃ + +∥∥-Idempotent-× : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-Idempotent-× {A = A} {A′} = A ∥₁ × A′ ∥₁ ∥₁ ≡⟨ ∥∥-IdempotentR-× + A ∥₁ × A′ ∥₁ ≡⟨ ∥∥-IdempotentL-× + A × A′ ∥₁ + +∥∥-Idempotent-×-≃ : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-Idempotent-×-≃ {A = A} {A′} = compEquiv ∥∥-IdempotentR-×-≃ ∥∥-IdempotentL-×-≃ + +∥∥-×-≃ : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-×-≃ {A = A} {A′} = compEquiv (invEquiv (propTruncIdempotent≃ (isProp× isPropPropTrunc isPropPropTrunc))) ∥∥-Idempotent-×-≃ + +∥∥-× : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-× = ua ∥∥-×-≃ + +-- using this we get a convenient recursor/eliminator for binary functions into sets +rec2→Set : {A B C : Type } (Cset : isSet C) + (f : A B C) + (∀ (a a' : A) (b b' : B) f a b f a' b') + A ∥₁ B ∥₁ C +rec2→Set {A = A} {B = B} {C = C} Cset f fconst = curry (g ∥∥-×-≃ .fst) + where + g : A × B ∥₁ C + g = rec→Set Cset (uncurry f) λ x y fconst (fst x) (fst y) (snd x) (snd y) \ No newline at end of file diff --git a/docs/Cubical.HITs.SetQuotients.Properties.html b/docs/Cubical.HITs.SetQuotients.Properties.html index 21f8edd..0ecf9d2 100644 --- a/docs/Cubical.HITs.SetQuotients.Properties.html +++ b/docs/Cubical.HITs.SetQuotients.Properties.html @@ -31,7 +31,7 @@ open import Cubical.HITs.PropositionalTruncation as PropTrunc using (∥_∥₁ ; ∣_∣₁ ; squash₁) renaming (rec to propRec) open import Cubical.HITs.SetTruncation as SetTrunc - using (∥_∥₂ ; ∣_∣₂ ; squash₂ ; isSetSetTrunc) + using (∥_∥₂ ; ∣_∣₂ ; squash₂ ; isSetSetTrunc) private @@ -41,20 +41,20 @@ R S T W : A A Type elimProp : {P : A / R Type } - (∀ x isProp (P x)) + (∀ x isProp (P x)) (∀ a P [ a ]) x P x elimProp prop f [ x ] = f x elimProp prop f (squash/ x y p q i j) = - isOfHLevel→isOfHLevelDep 2 x isProp→isSet (prop x)) + isOfHLevel→isOfHLevelDep 2 x isProp→isSet (prop x)) (g x) (g y) (cong g p) (cong g q) (squash/ x y p q) i j where g = elimProp prop f elimProp prop f (eq/ a b r i) = - isProp→PathP i prop (eq/ a b r i)) (f a) (f b) i + isProp→PathP i prop (eq/ a b r i)) (f a) (f b) i elimProp2 : {P : A / R B / S Type } - (∀ x y isProp (P x y)) + (∀ x y isProp (P x y)) (∀ a b P [ a ] [ b ]) x y P x y elimProp2 prop f = @@ -62,7 +62,7 @@ elimProp (prop [ a ]) (f a) elimProp3 : {P : A / R B / S C / T Type } - (∀ x y z isProp (P x y z)) + (∀ x y z isProp (P x y z)) (∀ a b c P [ a ] [ b ] [ c ]) x y z P x y z elimProp3 prop f = @@ -70,7 +70,7 @@ elimProp2 (prop [ a ]) (f a) elimProp4 : {P : A / R B / S C / T Q / W Type } - (∀ x y z t isProp (P x y z t)) + (∀ x y z t isProp (P x y z t)) (∀ a b c d P [ a ] [ b ] [ c ] [ d ]) x y z t P x y z t elimProp4 prop f = @@ -79,37 +79,37 @@ -- sometimes more convenient: elimContr : {P : A / R Type } - (∀ a isContr (P [ a ])) + (∀ a isContr (P [ a ])) x P x elimContr contr = - elimProp (elimProp _ isPropIsProp) λ _ isContr→isProp (contr _)) λ _ + elimProp (elimProp _ isPropIsProp) λ _ isContr→isProp (contr _)) λ _ contr _ .fst elimContr2 : {P : A / R B / S Type } - (∀ a b isContr (P [ a ] [ b ])) + (∀ a b isContr (P [ a ] [ b ])) x y P x y elimContr2 contr = elimContr λ _ - isOfHLevelΠ 0 (elimContr λ _ inhProp→isContr (contr _ _) isPropIsContr) + isOfHLevelΠ 0 (elimContr λ _ inhProp→isContr (contr _ _) isPropIsContr) -- lemma 6.10.2 in hott book []surjective : (x : A / R) ∃[ a A ] [ a ] x []surjective = elimProp x squash₁) a a , refl ∣₁) elim : {P : A / R Type } - (∀ x isSet (P x)) + (∀ x isSet (P x)) (f : (a : A) (P [ a ])) ((a b : A) (r : R a b) PathP i P (eq/ a b r i)) (f a) (f b)) x P x elim set f feq [ a ] = f a elim set f feq (eq/ a b r i) = feq a b r i elim set f feq (squash/ x y p q i j) = - isOfHLevel→isOfHLevelDep 2 set + isOfHLevel→isOfHLevelDep 2 set (g x) (g y) (cong g p) (cong g q) (squash/ x y p q) i j where g = elim set f feq -rec : isSet B +rec : isSet B (f : A B) ((a b : A) (r : R a b) f a f b) A / R B @@ -119,7 +119,7 @@ where g = rec set f feq -rec2 : isSet C +rec2 : isSet C (f : A B C) (∀ a b c R a b f a c f b c) (∀ a b c S b c f a b f a c) @@ -156,24 +156,24 @@ -- We start by proving that we can recover the set-quotient -- by set-truncating the (non-truncated type quotient) typeQuotSetTruncIso : Iso (A / R) A /ₜ R ∥₂ -Iso.fun typeQuotSetTruncIso = rec isSetSetTrunc a [ a ] ∣₂) +Iso.fun typeQuotSetTruncIso = rec isSetSetTrunc a [ a ] ∣₂) λ a b r cong ∣_∣₂ (eq/ a b r) -Iso.inv typeQuotSetTruncIso = SetTrunc.rec squash/ (TypeQuot.rec [_] eq/) -Iso.rightInv typeQuotSetTruncIso = SetTrunc.elim _ isProp→isSet (squash₂ _ _)) +Iso.inv typeQuotSetTruncIso = SetTrunc.rec squash/ (TypeQuot.rec [_] eq/) +Iso.rightInv typeQuotSetTruncIso = SetTrunc.elim _ isProp→isSet (squash₂ _ _)) (TypeQuot.elimProp _ squash₂ _ _) λ _ refl) Iso.leftInv typeQuotSetTruncIso = elimProp _ squash/ _ _) λ _ refl -module rec→Gpd {B : Type ℓ''} (Bgpd : isGroupoid B) +module rec→Gpd {B : Type ℓ''} (Bgpd : isGroupoid B) (f : A B) (feq : (a b : A) R a b f a f b) - (fprop : (a b : A) isProp (f a f b)) + (fprop : (a b : A) isProp (f a f b)) where fun : A / R B fun = f₁ f₂ where f₁ : A /ₜ R ∥₂ B - f₁ = SetTrunc.rec→Gpd.fun Bgpd f/ congF/Const + f₁ = SetTrunc.rec→Gpd.fun Bgpd f/ congF/Const where f/ : A /ₜ R B f/ = TypeQuot.rec f feq @@ -188,7 +188,7 @@ f₂ = Iso.fun typeQuotSetTruncIso -setQuotUniversalIso : isSet B +setQuotUniversalIso : isSet B Iso (A / R B) (Σ[ f (A B) ] ((a b : A) R a b f a f b)) Iso.fun (setQuotUniversalIso Bset) g = a g [ a ]) , λ a b r i g (eq/ a b r i) Iso.inv (setQuotUniversalIso Bset) h = rec Bset (fst h) (snd h) @@ -203,11 +203,11 @@ intro = Iso.fun (setQuotUniversalIso Bset) out = Iso.inv (setQuotUniversalIso Bset) -setQuotUniversal : isSet B +setQuotUniversal : isSet B (A / R B) (Σ[ f (A B) ] ((a b : A) R a b f a f b)) setQuotUniversal Bset = isoToEquiv (setQuotUniversalIso Bset) -open BinaryRelation +open BinaryRelation setQuotUnaryOp : (-_ : A A) (∀ a a' R a a' R (- a) (- a')) @@ -215,7 +215,7 @@ setQuotUnaryOp -_ h = rec squash/ a [ - a ]) a b x eq/ _ _ (h _ _ x)) -- characterisation of binary functions/operations on set-quotients -setQuotUniversal2Iso : isSet C isRefl R isRefl S +setQuotUniversal2Iso : isSet C isRefl R isRefl S Iso (A / R B / S C) (Σ[ _∗_ (A B C) ] (∀ a a' b b' R a a' S b b' a b a' b')) Iso.fun (setQuotUniversal2Iso {R = R} {S = S} Bset isReflR isReflS) _∗/_ = _∗_ , h @@ -237,7 +237,7 @@ Iso.leftInv (setQuotUniversal2Iso Bset isReflR isReflS) _∗/_ = funExt₂ (elimProp2 _ _ Bset _ _) λ _ _ refl) -setQuotUniversal2 : isSet C isRefl R isRefl S +setQuotUniversal2 : isSet C isRefl R isRefl S (A / R B / S C) (Σ[ _∗_ (A B C) ] (∀ a a' b b' R a a' S b b' a b a' b')) setQuotUniversal2 Bset isReflR isReflS = @@ -245,7 +245,7 @@ -- corollary for binary operations -- TODO: prove truncated inverse for effective relations -setQuotBinOp : isRefl R isRefl S +setQuotBinOp : isRefl R isRefl S (_∗_ : A B C) (∀ a a' b b' R a a' S b b' T (a b) (a' b')) (A / R B / S C / T) @@ -254,7 +254,7 @@ _ _ _ r eq/ _ _ (h _ _ _ _ r (isReflS _))) _ _ _ s eq/ _ _ (h _ _ _ _ (isReflR _) s)) -setQuotSymmBinOp : isRefl R isTrans R +setQuotSymmBinOp : isRefl R isTrans R (_∗_ : A A A) (∀ a b R (a b) (b a)) (∀ a a' b R a a' R (a b) (a' b)) @@ -268,17 +268,17 @@ (isTransR _ _ _ (∗Rsymm a' b) (isTransR _ _ _ (h b b' a' rb) (∗Rsymm b' a'))) -effective : (Rprop : isPropValued R) (Requiv : isEquivRel R) +effective : (Rprop : isPropValued R) (Requiv : isEquivRel R) (a b : A) [ a ] [ b ] R a b -effective {A = A} {R = R} Rprop (equivRel R/refl R/sym R/trans) a b p = +effective {A = A} {R = R} Rprop (equivRel R/refl R/sym R/trans) a b p = transport aa≡ab (R/refl _) where helper : A / R hProp _ helper = - rec isSetHProp + rec isSetHProp c (R a c , Rprop a c)) c d cd - Σ≡Prop _ isPropIsProp) + Σ≡Prop _ isPropIsProp) (hPropExt (Rprop a c) (Rprop a d) ac R/trans _ _ _ ac cd) ad R/trans _ _ _ ad (R/sym _ _ cd)))) @@ -286,14 +286,14 @@ aa≡ab : R a a R a b aa≡ab i = helper (p i) .fst -isEquivRel→effectiveIso : isPropValued R isEquivRel R +isEquivRel→effectiveIso : isPropValued R isEquivRel R (a b : A) Iso ([ a ] [ b ]) (R a b) Iso.fun (isEquivRel→effectiveIso {R = R} Rprop Req a b) = effective Rprop Req a b Iso.inv (isEquivRel→effectiveIso {R = R} Rprop Req a b) = eq/ a b Iso.rightInv (isEquivRel→effectiveIso {R = R} Rprop Req a b) _ = Rprop a b _ _ Iso.leftInv (isEquivRel→effectiveIso {R = R} Rprop Req a b) _ = squash/ _ _ _ _ -isEquivRel→isEffective : isPropValued R isEquivRel R isEffective R +isEquivRel→isEffective : isPropValued R isEquivRel R isEffective R isEquivRel→isEffective Rprop Req a b = isoToIsEquiv (invIso (isEquivRel→effectiveIso Rprop Req a b)) @@ -311,20 +311,20 @@ -- path-types for equivalence relations (not prop-valued) -- and their quotients -isEquivRel→TruncIso : isEquivRel R (a b : A) Iso ([ a ] [ b ]) R a b ∥₁ +isEquivRel→TruncIso : isEquivRel R (a b : A) Iso ([ a ] [ b ]) R a b ∥₁ isEquivRel→TruncIso {A = A} {R = R} Req a b = compIso (isProp→Iso (squash/ _ _) (squash/ _ _) (cong (Iso.fun truncRelIso)) (cong (Iso.inv truncRelIso))) (isEquivRel→effectiveIso _ _ PropTrunc.isPropPropTrunc) ∥R∥eq a b) where - open isEquivRel - ∥R∥eq : isEquivRel λ a b R a b ∥₁ - reflexive ∥R∥eq a = reflexive Req a ∣₁ - symmetric ∥R∥eq a b = PropTrunc.map (symmetric Req a b) - transitive ∥R∥eq a b c = PropTrunc.map2 (transitive Req a b c) + open isEquivRel + ∥R∥eq : isEquivRel λ a b R a b ∥₁ + reflexive ∥R∥eq a = reflexive Req a ∣₁ + symmetric ∥R∥eq a b = PropTrunc.map (symmetric Req a b) + transitive ∥R∥eq a b c = PropTrunc.map2 (transitive Req a b c) -discreteSetQuotients : isEquivRel R +discreteSetQuotients : isEquivRel R (∀ a₀ a₁ Dec (R a₀ a₁)) Discrete (A / R) discreteSetQuotients {A = A} {R = R} Req Rdec = @@ -341,7 +341,7 @@ Iso.rightInv (relBiimpl→TruncIso R→S S→R) = elimProp _ squash/ _ _) λ _ refl Iso.leftInv (relBiimpl→TruncIso R→S S→R) = elimProp _ squash/ _ _) λ _ refl -descendMapPath : {M : Type } (f g : A / R M) (isSetM : isSet M) +descendMapPath : {M : Type } (f g : A / R M) (isSetM : isSet M) ((x : A) f [ x ] g [ x ]) f g descendMapPath f g isSetM path i x = diff --git a/docs/Cubical.HITs.SetTruncation.Properties.html b/docs/Cubical.HITs.SetTruncation.Properties.html index 8092c6c..0f93aec 100644 --- a/docs/Cubical.HITs.SetTruncation.Properties.html +++ b/docs/Cubical.HITs.SetTruncation.Properties.html @@ -25,318 +25,322 @@ private variable - ℓ' ℓ'' : Level - A B C D : Type - -isSetPathImplicit : {x y : A ∥₂} isSet (x y) -isSetPathImplicit = isOfHLevelPath 2 squash₂ _ _ - -rec : isSet B (A B) A ∥₂ B -rec Bset f x ∣₂ = f x -rec Bset f (squash₂ x y p q i j) = - Bset _ _ (cong (rec Bset f) p) (cong (rec Bset f) q) i j - -rec2 : isSet C (A B C) A ∥₂ B ∥₂ C -rec2 Cset f x ∣₂ y ∣₂ = f x y -rec2 Cset f x ∣₂ (squash₂ y z p q i j) = - Cset _ _ (cong (rec2 Cset f x ∣₂) p) (cong (rec2 Cset f x ∣₂) q) i j -rec2 Cset f (squash₂ x y p q i j) z = - Cset _ _ (cong a rec2 Cset f a z) p) (cong a rec2 Cset f a z) q) i j - --- Old version: --- rec2 Cset f = rec (isSetΠ λ _ → Cset) λ x → rec Cset (f x) - --- lemma 6.9.1 in HoTT book -elim : {B : A ∥₂ Type } - (Bset : (x : A ∥₂) isSet (B x)) - (f : (a : A) B ( a ∣₂)) - (x : A ∥₂) B x -elim Bset f a ∣₂ = f a -elim Bset f (squash₂ x y p q i j) = - isOfHLevel→isOfHLevelDep 2 Bset _ _ - (cong (elim Bset f) p) (cong (elim Bset f) q) (squash₂ x y p q) i j - -elim2 : {C : A ∥₂ B ∥₂ Type } - (Cset : ((x : A ∥₂) (y : B ∥₂) isSet (C x y))) - (f : (a : A) (b : B) C a ∣₂ b ∣₂) - (x : A ∥₂) (y : B ∥₂) C x y -elim2 Cset f x ∣₂ y ∣₂ = f x y -elim2 Cset f x ∣₂ (squash₂ y z p q i j) = - isOfHLevel→isOfHLevelDep 2 a Cset x ∣₂ a) _ _ - (cong (elim2 Cset f x ∣₂) p) (cong (elim2 Cset f x ∣₂) q) (squash₂ y z p q) i j -elim2 Cset f (squash₂ x y p q i j) z = - isOfHLevel→isOfHLevelDep 2 a Cset a z) _ _ - (cong a elim2 Cset f a z) p) (cong a elim2 Cset f a z) q) (squash₂ x y p q) i j - --- Old version: --- elim2 Cset f = elim (λ _ → isSetΠ (λ _ → Cset _ _)) --- (λ a → elim (λ _ → Cset _ _) (f a)) - --- TODO: generalize -elim3 : {B : (x y z : A ∥₂) Type } - (Bset : ((x y z : A ∥₂) isSet (B x y z))) - (g : (a b c : A) B a ∣₂ b ∣₂ c ∣₂) - (x y z : A ∥₂) B x y z -elim3 Bset g = elim2 _ _ isSetΠ _ Bset _ _ _)) - a b elim _ Bset _ _ _) (g a b)) - -elim4 : {B : (w x y z : A ∥₂) Type } - (Bset : ((w x y z : A ∥₂) isSet (B w x y z))) - (g : (a b c d : A) B a ∣₂ b ∣₂ c ∣₂ d ∣₂) - (w x y z : A ∥₂) B w x y z -elim4 Bset g = elim3 _ _ _ isSetΠ λ _ Bset _ _ _ _) - λ a b c elim _ Bset _ _ _ _) (g a b c) - - --- the recursor for maps into groupoids following the "HIT proof" in: --- https://arxiv.org/abs/1507.01150 --- i.e. for any type A and groupoid B we can construct a map ∥ A ∥₂ → B --- from a map A → B satisfying the condition --- ∀ (a b : A) (p q : a ≡ b) → cong f p ≡ cong f q --- TODO: prove that this is an equivalence -module rec→Gpd {A : Type } {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A B) - (congFConst : (a b : A) (p q : a b) cong f p cong f q) where - - data H : Type where - η : A H - ε : (a b : A) a b ∥₁ η a η b -- prop. trunc. of a≡b - δ : (a b : A) (p : a b) ε a b p ∣₁ cong η p - gtrunc : isGroupoid H - - -- write elimination principle for H - module Helim {P : H Type ℓ''} (Pgpd : h isGroupoid (P h)) - (η* : (a : A) P (η a)) - (ε* : (a b : A) (∣p∣₁ : a b ∥₁) - PathP i P (ε a b ∣p∣₁ i)) (η* a) (η* b)) - (δ* : (a b : A) (p : a b) - PathP i PathP j P (δ a b p i j)) (η* a) (η* b)) - (ε* a b p ∣₁) (cong η* p)) where - - fun : (h : H) P h - fun (η a) = η* a - fun (ε a b ∣p∣₁ i) = ε* a b ∣p∣₁ i - fun (δ a b p i j) = δ* a b p i j - fun (gtrunc x y p q α β i j k) = isOfHLevel→isOfHLevelDep 3 Pgpd - (fun x) (fun y) - (cong fun p) (cong fun q) - (cong (cong fun) α) (cong (cong fun) β) - (gtrunc x y p q α β) i j k - - module Hrec {C : Type ℓ''} (Cgpd : isGroupoid C) - (η* : A C) - (ε* : (a b : A) a b ∥₁ η* a η* b) - (δ* : (a b : A) (p : a b) ε* a b p ∣₁ cong η* p) where - - fun : H C - fun (η a) = η* a - fun (ε a b ∣p∣₁ i) = ε* a b ∣p∣₁ i - fun (δ a b p i j) = δ* a b p i j - fun (gtrunc x y p q α β i j k) = Cgpd (fun x) (fun y) (cong fun p) (cong fun q) - (cong (cong fun) α) (cong (cong fun) β) i j k - - module HelimProp {P : H Type ℓ''} (Pprop : h isProp (P h)) - (η* : (a : A) P (η a)) where - - fun : h P h - fun = Helim.fun _ isSet→isGroupoid (isProp→isSet (Pprop _))) η* - a b ∣p∣₁ isOfHLevel→isOfHLevelDep 1 Pprop _ _ (ε a b ∣p∣₁)) - λ a b p isOfHLevel→isOfHLevelDep 1 - {B = λ p PathP i P (p i)) (η* a) (η* b)} - _ isOfHLevelPathP 1 (Pprop _) _ _) _ _ (δ a b p) - - -- The main trick: eliminating into hsets is easy - -- i.e. H has the universal property of set truncation... - module HelimSet {P : H Type ℓ''} (Pset : h isSet (P h)) - (η* : a P (η a)) where - - fun : (h : H) P h - fun = Helim.fun _ isSet→isGroupoid (Pset _)) η* ε* - λ a b p isOfHLevel→isOfHLevelDep 1 - {B = λ p PathP i P (p i)) (η* a) (η* b)} - _ isOfHLevelPathP' 1 (Pset _) _ _) _ _ (δ a b p) - where - ε* : (a b : A) (∣p∣₁ : a b ∥₁) PathP i P (ε a b ∣p∣₁ i)) (η* a) (η* b) - ε* a b = pElim _ isOfHLevelPathP' 1 (Pset _) (η* a) (η* b)) - λ p subst x PathP i P (x i)) (η* a) (η* b)) - (sym (δ a b p)) (cong η* p) - - - -- Now we need to prove that H is a set. - -- We start with a little lemma: - localHedbergLemma : {X : Type ℓ''} (P : X Type ℓ'') - (∀ x isProp (P x)) - ((x : X) P x (y : X) P y x y) - -------------------------------------------------- - (x : X) P x (y : X) isProp (x y) - localHedbergLemma {X = X} P Pprop P→≡ x px y = isPropRetract - p subst P p px) py sym (P→≡ x px x px) P→≡ x px y py) - isRetract (Pprop y) - where - isRetract : (p : x y) (sym (P→≡ x px x px)) P→≡ x px y (subst P p px) p - isRetract = J y' p' (sym (P→≡ x px x px)) P→≡ x px y' (subst P p' px) p') - (subst px' sym (P→≡ x px x px) P→≡ x px x px' refl) - (sym (substRefl {B = P} px)) (lCancel (P→≡ x px x px))) - - Hset : isSet H - Hset = HelimProp.fun _ isPropΠ λ _ isPropIsProp) baseCaseLeft - where - baseCaseLeft : (a₀ : A) (y : H) isProp (η a₀ y) - baseCaseLeft a₀ = localHedbergLemma x Q x .fst) x Q x .snd) Q→≡ _ refl ∣₁ - where - Q : H hProp - Q = HelimSet.fun _ isSetHProp) λ b a₀ b ∥₁ , isPropPropTrunc - -- Q (η b) = ∥ a ≡ b ∥₁ - - Q→≡ : (x : H) Q x .fst (y : H) Q y .fst x y - Q→≡ = HelimSet.fun _ isSetΠ3 λ _ _ _ gtrunc _ _) - λ a p HelimSet.fun _ isSetΠ λ _ gtrunc _ _) - λ b q sym (ε a₀ a p) ε a₀ b q - - -- our desired function will split through H, - -- i.e. we get a function ∥ A ∥₂ → H → B - fun : A ∥₂ B - fun = f₁ f₂ - where - f₁ : H B - f₁ = Hrec.fun Bgpd f εᶠ λ _ _ _ refl - where - εᶠ : (a b : A) a b ∥₁ f a f b - εᶠ a b = rec→Set (Bgpd _ _) (cong f) λ p q congFConst a b p q - -- this is the inductive step, - -- we use that maps ∥ A ∥₁ → B for an hset B - -- correspond to 2-Constant maps A → B (which cong f is by assumption) - f₂ : A ∥₂ H - f₂ = rec Hset η - - -map : (A B) A ∥₂ B ∥₂ -map f = rec squash₂ λ x f x ∣₂ - -map∙ : { ℓ' : Level} {A : Pointed } {B : Pointed ℓ'} - (f : A →∙ B) A ∥₂∙ →∙ B ∥₂∙ -fst (map∙ f) = map (fst f) -snd (map∙ f) = cong ∣_∣₂ (snd f) - -setTruncUniversal : isSet B ( A ∥₂ B) (A B) -setTruncUniversal {B = B} Bset = - isoToEquiv (iso h x h x ∣₂) (rec Bset) _ refl) rinv) - where - rinv : (f : A ∥₂ B) rec Bset x f x ∣₂) f - rinv f i x = - elim x isProp→isSet (Bset (rec Bset x f x ∣₂) x) (f x))) - _ refl) x i - -isSetSetTrunc : isSet A ∥₂ -isSetSetTrunc a b p q = squash₂ a b p q - -setTruncIdempotentIso : isSet A Iso A ∥₂ A -Iso.fun (setTruncIdempotentIso hA) = rec hA (idfun _) -Iso.inv (setTruncIdempotentIso hA) x = x ∣₂ -Iso.rightInv (setTruncIdempotentIso hA) _ = refl -Iso.leftInv (setTruncIdempotentIso hA) = elim _ isSet→isGroupoid isSetSetTrunc _ _) _ refl) - -setTruncIdempotent≃ : isSet A A ∥₂ A -setTruncIdempotent≃ {A = A} hA = isoToEquiv (setTruncIdempotentIso hA) - -setTruncIdempotent : isSet A A ∥₂ A -setTruncIdempotent hA = ua (setTruncIdempotent≃ hA) - -isContr→isContrSetTrunc : isContr A isContr ( A ∥₂) -isContr→isContrSetTrunc contr = fst contr ∣₂ - , elim _ isOfHLevelPath 2 (isSetSetTrunc) _ _) - λ a cong ∣_∣₂ (snd contr a) - - -setTruncIso : Iso A B Iso A ∥₂ B ∥₂ -Iso.fun (setTruncIso is) = rec isSetSetTrunc x Iso.fun is x ∣₂) -Iso.inv (setTruncIso is) = rec isSetSetTrunc x Iso.inv is x ∣₂) -Iso.rightInv (setTruncIso is) = - elim _ isOfHLevelPath 2 isSetSetTrunc _ _) - λ a cong ∣_∣₂ (Iso.rightInv is a) -Iso.leftInv (setTruncIso is) = - elim _ isOfHLevelPath 2 isSetSetTrunc _ _) - λ a cong ∣_∣₂ (Iso.leftInv is a) - -setSigmaIso : {B : A Type } Iso Σ A B ∥₂ Σ A x B x ∥₂) ∥₂ -setSigmaIso {A = A} {B = B} = iso fun funinv sect retr - where - {- writing it out explicitly to avoid yellow highlighting -} - fun : Σ A B ∥₂ Σ A x B x ∥₂) ∥₂ - fun = rec isSetSetTrunc λ {(a , p) a , p ∣₂ ∣₂} - funinv : Σ A x B x ∥₂) ∥₂ Σ A B ∥₂ - funinv = rec isSetSetTrunc {(a , p) rec isSetSetTrunc p a , p ∣₂) p}) - sect : section fun funinv - sect = elim _ isOfHLevelPath 2 isSetSetTrunc _ _) - λ { (a , p) elim {B = λ p fun (funinv a , p ∣₂) a , p ∣₂} - p isOfHLevelPath 2 isSetSetTrunc _ _) _ refl) p } - retr : retract fun funinv - retr = elim _ isOfHLevelPath 2 isSetSetTrunc _ _) - λ { _ refl } - -sigmaElim : {B : A ∥₂ Type } {C : Σ A ∥₂ B Type ℓ'} - (Bset : (x : Σ A ∥₂ B) isSet (C x)) - (g : (a : A) (b : B a ∣₂) C ( a ∣₂ , b)) - (x : Σ A ∥₂ B) C x -sigmaElim {B = B} {C = C} set g (x , y) = - elim {B = λ x (y : B x) C (x , y)} _ isSetΠ λ _ set _) g x y - -sigmaProdElim : {C : A ∥₂ × B ∥₂ Type } {D : Σ ( A ∥₂ × B ∥₂) C Type ℓ'} - (Bset : (x : Σ ( A ∥₂ × B ∥₂) C) isSet (D x)) - (g : (a : A) (b : B) (c : C ( a ∣₂ , b ∣₂)) D (( a ∣₂ , b ∣₂) , c)) - (x : Σ ( A ∥₂ × B ∥₂) C) D x -sigmaProdElim {B = B} {C = C} {D = D} set g ((x , y) , c) = - elim {B = λ x (y : B ∥₂) (c : C (x , y)) D ((x , y) , c)} - _ isSetΠ λ _ isSetΠ λ _ set _) - x elim _ isSetΠ λ _ set _) (g x)) - x y c - -prodElim : {C : A ∥₂ × B ∥₂ Type } - ((x : A ∥₂ × B ∥₂) isSet (C x)) - ((a : A) (b : B) C ( a ∣₂ , b ∣₂)) - (x : A ∥₂ × B ∥₂) C x -prodElim setC f (a , b) = elim2 x y setC (x , y)) f a b - -prodRec : {C : Type } isSet C (A B C) A ∥₂ × B ∥₂ C -prodRec setC f (a , b) = rec2 setC f a b - -prodElim2 : {E : ( A ∥₂ × B ∥₂) ( C ∥₂ × D ∥₂) Type } - ((x : A ∥₂ × B ∥₂) (y : C ∥₂ × D ∥₂) isSet (E x y)) - ((a : A) (b : B) (c : C) (d : D) E ( a ∣₂ , b ∣₂) ( c ∣₂ , d ∣₂)) - ((x : A ∥₂ × B ∥₂) (y : C ∥₂ × D ∥₂) (E x y)) -prodElim2 isset f = prodElim _ isSetΠ λ _ isset _ _) - λ a b prodElim _ isset _ _) - λ c d f a b c d - -setTruncOfProdIso : Iso A × B ∥₂ ( A ∥₂ × B ∥₂) -Iso.fun setTruncOfProdIso = rec (isSet× isSetSetTrunc isSetSetTrunc) λ { (a , b) a ∣₂ , b ∣₂ } -Iso.inv setTruncOfProdIso = prodRec isSetSetTrunc λ a b a , b ∣₂ -Iso.rightInv setTruncOfProdIso = - prodElim _ isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ refl -Iso.leftInv setTruncOfProdIso = - elim _ isOfHLevelPath 2 isSetSetTrunc _ _) λ {(a , b) refl} - -IsoSetTruncateSndΣ : {A : Type } {B : A Type ℓ'} Iso Σ A B ∥₂ Σ A x B x ∥₂) ∥₂ -Iso.fun IsoSetTruncateSndΣ = map λ a (fst a) , snd a ∣₂ -Iso.inv IsoSetTruncateSndΣ = rec isSetSetTrunc (uncurry λ x map λ b x , b) -Iso.rightInv IsoSetTruncateSndΣ = - elim _ isOfHLevelPath 2 isSetSetTrunc _ _) - (uncurry λ a elim _ isOfHLevelPath 2 isSetSetTrunc _ _) - λ _ refl) -Iso.leftInv IsoSetTruncateSndΣ = - elim _ isOfHLevelPath 2 isSetSetTrunc _ _) - λ _ refl - -PathIdTrunc₀Iso : {a b : A} Iso ( a ∣₂ b ∣₂) a b ∥₁ -Iso.fun (PathIdTrunc₀Iso {b = b}) p = - transport i rec {B = TypeOfHLevel _ 1} (isOfHLevelTypeOfHLevel 1) - a a b ∥₁ , squash₁) (p (~ i)) .fst) - refl ∣₁ -Iso.inv PathIdTrunc₀Iso = pRec (squash₂ _ _) (cong ∣_∣₂) -Iso.rightInv PathIdTrunc₀Iso _ = squash₁ _ _ -Iso.leftInv PathIdTrunc₀Iso _ = squash₂ _ _ _ _ - -mapFunctorial : {A B C : Type } (f : A B) (g : B C) - map g map f map (g f) -mapFunctorial f g = - funExt (elim x isSetPathImplicit) λ a refl) + ℓ' ℓ'' ℓa ℓb ℓc ℓd : Level + A : Type ℓa + B : Type ℓb + C : Type ℓc + D : Type ℓd + +isSetPathImplicit : {x y : A ∥₂} isSet (x y) +isSetPathImplicit = isOfHLevelPath 2 squash₂ _ _ + +rec : isSet B (A B) A ∥₂ B +rec Bset f x ∣₂ = f x +rec Bset f (squash₂ x y p q i j) = + Bset _ _ (cong (rec Bset f) p) (cong (rec Bset f) q) i j + +rec2 : isSet C (A B C) A ∥₂ B ∥₂ C +rec2 Cset f x ∣₂ y ∣₂ = f x y +rec2 Cset f x ∣₂ (squash₂ y z p q i j) = + Cset _ _ (cong (rec2 Cset f x ∣₂) p) (cong (rec2 Cset f x ∣₂) q) i j +rec2 Cset f (squash₂ x y p q i j) z = + Cset _ _ (cong a rec2 Cset f a z) p) (cong a rec2 Cset f a z) q) i j + +-- Old version: +-- rec2 Cset f = rec (isSetΠ λ _ → Cset) λ x → rec Cset (f x) + +-- lemma 6.9.1 in HoTT book +elim : {B : A ∥₂ Type } + (Bset : (x : A ∥₂) isSet (B x)) + (f : (a : A) B ( a ∣₂)) + (x : A ∥₂) B x +elim Bset f a ∣₂ = f a +elim Bset f (squash₂ x y p q i j) = + isOfHLevel→isOfHLevelDep 2 Bset _ _ + (cong (elim Bset f) p) (cong (elim Bset f) q) (squash₂ x y p q) i j + +elim2 : {C : A ∥₂ B ∥₂ Type } + (Cset : ((x : A ∥₂) (y : B ∥₂) isSet (C x y))) + (f : (a : A) (b : B) C a ∣₂ b ∣₂) + (x : A ∥₂) (y : B ∥₂) C x y +elim2 Cset f x ∣₂ y ∣₂ = f x y +elim2 Cset f x ∣₂ (squash₂ y z p q i j) = + isOfHLevel→isOfHLevelDep 2 a Cset x ∣₂ a) _ _ + (cong (elim2 Cset f x ∣₂) p) (cong (elim2 Cset f x ∣₂) q) (squash₂ y z p q) i j +elim2 Cset f (squash₂ x y p q i j) z = + isOfHLevel→isOfHLevelDep 2 a Cset a z) _ _ + (cong a elim2 Cset f a z) p) (cong a elim2 Cset f a z) q) (squash₂ x y p q) i j + +-- Old version: +-- elim2 Cset f = elim (λ _ → isSetΠ (λ _ → Cset _ _)) +-- (λ a → elim (λ _ → Cset _ _) (f a)) + +-- TODO: generalize +elim3 : {D : A ∥₂ B ∥₂ C ∥₂ Type } + (Dset : ((x : A ∥₂) (y : B ∥₂) (z : C ∥₂) isSet (D x y z))) + (g : (a : A) (b : B) (c : C) D a ∣₂ b ∣₂ c ∣₂) + (x : A ∥₂) (y : B ∥₂) (z : C ∥₂) D x y z +elim3 Dset g = elim2 _ _ isSetΠ _ Dset _ _ _)) + a b elim _ Dset _ _ _) (g a b)) + +elim4 : {E : A ∥₂ B ∥₂ C ∥₂ D ∥₂ Type } + (Eset : ((w : A ∥₂) (x : B ∥₂) (y : C ∥₂) (z : D ∥₂) + isSet (E w x y z))) + (g : (a : A) (b : B) (c : C) (d : D) E a ∣₂ b ∣₂ c ∣₂ d ∣₂) + (w : A ∥₂) (x : B ∥₂) (y : C ∥₂) (z : D ∥₂) E w x y z +elim4 Eset g = elim3 _ _ _ isSetΠ λ _ Eset _ _ _ _) + λ a b c elim _ Eset _ _ _ _) (g a b c) + + +-- the recursor for maps into groupoids following the "HIT proof" in: +-- https://arxiv.org/abs/1507.01150 +-- i.e. for any type A and groupoid B we can construct a map ∥ A ∥₂ → B +-- from a map A → B satisfying the condition +-- ∀ (a b : A) (p q : a ≡ b) → cong f p ≡ cong f q +-- TODO: prove that this is an equivalence +module rec→Gpd {A : Type } {B : Type ℓ'} (Bgpd : isGroupoid B) (f : A B) + (congFConst : (a b : A) (p q : a b) cong f p cong f q) where + + data H : Type where + η : A H + ε : (a b : A) a b ∥₁ η a η b -- prop. trunc. of a≡b + δ : (a b : A) (p : a b) ε a b p ∣₁ cong η p + gtrunc : isGroupoid H + + -- write elimination principle for H + module Helim {P : H Type ℓ''} (Pgpd : h isGroupoid (P h)) + (η* : (a : A) P (η a)) + (ε* : (a b : A) (∣p∣₁ : a b ∥₁) + PathP i P (ε a b ∣p∣₁ i)) (η* a) (η* b)) + (δ* : (a b : A) (p : a b) + PathP i PathP j P (δ a b p i j)) (η* a) (η* b)) + (ε* a b p ∣₁) (cong η* p)) where + + fun : (h : H) P h + fun (η a) = η* a + fun (ε a b ∣p∣₁ i) = ε* a b ∣p∣₁ i + fun (δ a b p i j) = δ* a b p i j + fun (gtrunc x y p q α β i j k) = isOfHLevel→isOfHLevelDep 3 Pgpd + (fun x) (fun y) + (cong fun p) (cong fun q) + (cong (cong fun) α) (cong (cong fun) β) + (gtrunc x y p q α β) i j k + + module Hrec {C : Type ℓ''} (Cgpd : isGroupoid C) + (η* : A C) + (ε* : (a b : A) a b ∥₁ η* a η* b) + (δ* : (a b : A) (p : a b) ε* a b p ∣₁ cong η* p) where + + fun : H C + fun (η a) = η* a + fun (ε a b ∣p∣₁ i) = ε* a b ∣p∣₁ i + fun (δ a b p i j) = δ* a b p i j + fun (gtrunc x y p q α β i j k) = Cgpd (fun x) (fun y) (cong fun p) (cong fun q) + (cong (cong fun) α) (cong (cong fun) β) i j k + + module HelimProp {P : H Type ℓ''} (Pprop : h isProp (P h)) + (η* : (a : A) P (η a)) where + + fun : h P h + fun = Helim.fun _ isSet→isGroupoid (isProp→isSet (Pprop _))) η* + a b ∣p∣₁ isOfHLevel→isOfHLevelDep 1 Pprop _ _ (ε a b ∣p∣₁)) + λ a b p isOfHLevel→isOfHLevelDep 1 + {B = λ p PathP i P (p i)) (η* a) (η* b)} + _ isOfHLevelPathP 1 (Pprop _) _ _) _ _ (δ a b p) + + -- The main trick: eliminating into hsets is easy + -- i.e. H has the universal property of set truncation... + module HelimSet {P : H Type ℓ''} (Pset : h isSet (P h)) + (η* : a P (η a)) where + + fun : (h : H) P h + fun = Helim.fun _ isSet→isGroupoid (Pset _)) η* ε* + λ a b p isOfHLevel→isOfHLevelDep 1 + {B = λ p PathP i P (p i)) (η* a) (η* b)} + _ isOfHLevelPathP' 1 (Pset _) _ _) _ _ (δ a b p) + where + ε* : (a b : A) (∣p∣₁ : a b ∥₁) PathP i P (ε a b ∣p∣₁ i)) (η* a) (η* b) + ε* a b = pElim _ isOfHLevelPathP' 1 (Pset _) (η* a) (η* b)) + λ p subst x PathP i P (x i)) (η* a) (η* b)) + (sym (δ a b p)) (cong η* p) + + + -- Now we need to prove that H is a set. + -- We start with a little lemma: + localHedbergLemma : {X : Type ℓ''} (P : X Type ℓ'') + (∀ x isProp (P x)) + ((x : X) P x (y : X) P y x y) + -------------------------------------------------- + (x : X) P x (y : X) isProp (x y) + localHedbergLemma {X = X} P Pprop P→≡ x px y = isPropRetract + p subst P p px) py sym (P→≡ x px x px) P→≡ x px y py) + isRetract (Pprop y) + where + isRetract : (p : x y) (sym (P→≡ x px x px)) P→≡ x px y (subst P p px) p + isRetract = J y' p' (sym (P→≡ x px x px)) P→≡ x px y' (subst P p' px) p') + (subst px' sym (P→≡ x px x px) P→≡ x px x px' refl) + (sym (substRefl {B = P} px)) (lCancel (P→≡ x px x px))) + + Hset : isSet H + Hset = HelimProp.fun _ isPropΠ λ _ isPropIsProp) baseCaseLeft + where + baseCaseLeft : (a₀ : A) (y : H) isProp (η a₀ y) + baseCaseLeft a₀ = localHedbergLemma x Q x .fst) x Q x .snd) Q→≡ _ refl ∣₁ + where + Q : H hProp + Q = HelimSet.fun _ isSetHProp) λ b a₀ b ∥₁ , isPropPropTrunc + -- Q (η b) = ∥ a ≡ b ∥₁ + + Q→≡ : (x : H) Q x .fst (y : H) Q y .fst x y + Q→≡ = HelimSet.fun _ isSetΠ3 λ _ _ _ gtrunc _ _) + λ a p HelimSet.fun _ isSetΠ λ _ gtrunc _ _) + λ b q sym (ε a₀ a p) ε a₀ b q + + -- our desired function will split through H, + -- i.e. we get a function ∥ A ∥₂ → H → B + fun : A ∥₂ B + fun = f₁ f₂ + where + f₁ : H B + f₁ = Hrec.fun Bgpd f εᶠ λ _ _ _ refl + where + εᶠ : (a b : A) a b ∥₁ f a f b + εᶠ a b = rec→Set (Bgpd _ _) (cong f) λ p q congFConst a b p q + -- this is the inductive step, + -- we use that maps ∥ A ∥₁ → B for an hset B + -- correspond to 2-Constant maps A → B (which cong f is by assumption) + f₂ : A ∥₂ H + f₂ = rec Hset η + + +map : (A B) A ∥₂ B ∥₂ +map f = rec squash₂ λ x f x ∣₂ + +map∙ : { ℓ' : Level} {A : Pointed } {B : Pointed ℓ'} + (f : A →∙ B) A ∥₂∙ →∙ B ∥₂∙ +fst (map∙ f) = map (fst f) +snd (map∙ f) = cong ∣_∣₂ (snd f) + +setTruncUniversal : isSet B ( A ∥₂ B) (A B) +setTruncUniversal {B = B} Bset = + isoToEquiv (iso h x h x ∣₂) (rec Bset) _ refl) rinv) + where + rinv : (f : A ∥₂ B) rec Bset x f x ∣₂) f + rinv f i x = + elim x isProp→isSet (Bset (rec Bset x f x ∣₂) x) (f x))) + _ refl) x i + +isSetSetTrunc : isSet A ∥₂ +isSetSetTrunc a b p q = squash₂ a b p q + +setTruncIdempotentIso : isSet A Iso A ∥₂ A +Iso.fun (setTruncIdempotentIso hA) = rec hA (idfun _) +Iso.inv (setTruncIdempotentIso hA) x = x ∣₂ +Iso.rightInv (setTruncIdempotentIso hA) _ = refl +Iso.leftInv (setTruncIdempotentIso hA) = elim _ isSet→isGroupoid isSetSetTrunc _ _) _ refl) + +setTruncIdempotent≃ : isSet A A ∥₂ A +setTruncIdempotent≃ {A = A} hA = isoToEquiv (setTruncIdempotentIso hA) + +setTruncIdempotent : isSet A A ∥₂ A +setTruncIdempotent hA = ua (setTruncIdempotent≃ hA) + +isContr→isContrSetTrunc : isContr A isContr ( A ∥₂) +isContr→isContrSetTrunc contr = fst contr ∣₂ + , elim _ isOfHLevelPath 2 (isSetSetTrunc) _ _) + λ a cong ∣_∣₂ (snd contr a) + + +setTruncIso : Iso A B Iso A ∥₂ B ∥₂ +Iso.fun (setTruncIso is) = rec isSetSetTrunc x Iso.fun is x ∣₂) +Iso.inv (setTruncIso is) = rec isSetSetTrunc x Iso.inv is x ∣₂) +Iso.rightInv (setTruncIso is) = + elim _ isOfHLevelPath 2 isSetSetTrunc _ _) + λ a cong ∣_∣₂ (Iso.rightInv is a) +Iso.leftInv (setTruncIso is) = + elim _ isOfHLevelPath 2 isSetSetTrunc _ _) + λ a cong ∣_∣₂ (Iso.leftInv is a) + +setSigmaIso : {B : A Type } Iso Σ A B ∥₂ Σ A x B x ∥₂) ∥₂ +setSigmaIso {A = A} {B = B} = iso fun funinv sect retr + where + {- writing it out explicitly to avoid yellow highlighting -} + fun : Σ A B ∥₂ Σ A x B x ∥₂) ∥₂ + fun = rec isSetSetTrunc λ {(a , p) a , p ∣₂ ∣₂} + funinv : Σ A x B x ∥₂) ∥₂ Σ A B ∥₂ + funinv = rec isSetSetTrunc {(a , p) rec isSetSetTrunc p a , p ∣₂) p}) + sect : section fun funinv + sect = elim _ isOfHLevelPath 2 isSetSetTrunc _ _) + λ { (a , p) elim {B = λ p fun (funinv a , p ∣₂) a , p ∣₂} + p isOfHLevelPath 2 isSetSetTrunc _ _) _ refl) p } + retr : retract fun funinv + retr = elim _ isOfHLevelPath 2 isSetSetTrunc _ _) + λ { _ refl } + +sigmaElim : {B : A ∥₂ Type } {C : Σ A ∥₂ B Type ℓ'} + (Bset : (x : Σ A ∥₂ B) isSet (C x)) + (g : (a : A) (b : B a ∣₂) C ( a ∣₂ , b)) + (x : Σ A ∥₂ B) C x +sigmaElim {B = B} {C = C} set g (x , y) = + elim {B = λ x (y : B x) C (x , y)} _ isSetΠ λ _ set _) g x y + +sigmaProdElim : {C : A ∥₂ × B ∥₂ Type } {D : Σ ( A ∥₂ × B ∥₂) C Type ℓ'} + (Bset : (x : Σ ( A ∥₂ × B ∥₂) C) isSet (D x)) + (g : (a : A) (b : B) (c : C ( a ∣₂ , b ∣₂)) D (( a ∣₂ , b ∣₂) , c)) + (x : Σ ( A ∥₂ × B ∥₂) C) D x +sigmaProdElim {B = B} {C = C} {D = D} set g ((x , y) , c) = + elim {B = λ x (y : B ∥₂) (c : C (x , y)) D ((x , y) , c)} + _ isSetΠ λ _ isSetΠ λ _ set _) + x elim _ isSetΠ λ _ set _) (g x)) + x y c + +prodElim : {C : A ∥₂ × B ∥₂ Type } + ((x : A ∥₂ × B ∥₂) isSet (C x)) + ((a : A) (b : B) C ( a ∣₂ , b ∣₂)) + (x : A ∥₂ × B ∥₂) C x +prodElim setC f (a , b) = elim2 x y setC (x , y)) f a b + +prodRec : {C : Type } isSet C (A B C) A ∥₂ × B ∥₂ C +prodRec setC f (a , b) = rec2 setC f a b + +prodElim2 : {E : ( A ∥₂ × B ∥₂) ( C ∥₂ × D ∥₂) Type } + ((x : A ∥₂ × B ∥₂) (y : C ∥₂ × D ∥₂) isSet (E x y)) + ((a : A) (b : B) (c : C) (d : D) E ( a ∣₂ , b ∣₂) ( c ∣₂ , d ∣₂)) + ((x : A ∥₂ × B ∥₂) (y : C ∥₂ × D ∥₂) (E x y)) +prodElim2 isset f = prodElim _ isSetΠ λ _ isset _ _) + λ a b prodElim _ isset _ _) + λ c d f a b c d + +setTruncOfProdIso : Iso A × B ∥₂ ( A ∥₂ × B ∥₂) +Iso.fun setTruncOfProdIso = rec (isSet× isSetSetTrunc isSetSetTrunc) λ { (a , b) a ∣₂ , b ∣₂ } +Iso.inv setTruncOfProdIso = prodRec isSetSetTrunc λ a b a , b ∣₂ +Iso.rightInv setTruncOfProdIso = + prodElim _ isOfHLevelPath 2 (isSet× isSetSetTrunc isSetSetTrunc) _ _) λ _ _ refl +Iso.leftInv setTruncOfProdIso = + elim _ isOfHLevelPath 2 isSetSetTrunc _ _) λ {(a , b) refl} + +IsoSetTruncateSndΣ : {A : Type } {B : A Type ℓ'} Iso Σ A B ∥₂ Σ A x B x ∥₂) ∥₂ +Iso.fun IsoSetTruncateSndΣ = map λ a (fst a) , snd a ∣₂ +Iso.inv IsoSetTruncateSndΣ = rec isSetSetTrunc (uncurry λ x map λ b x , b) +Iso.rightInv IsoSetTruncateSndΣ = + elim _ isOfHLevelPath 2 isSetSetTrunc _ _) + (uncurry λ a elim _ isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ refl) +Iso.leftInv IsoSetTruncateSndΣ = + elim _ isOfHLevelPath 2 isSetSetTrunc _ _) + λ _ refl + +PathIdTrunc₀Iso : {a b : A} Iso ( a ∣₂ b ∣₂) a b ∥₁ +Iso.fun (PathIdTrunc₀Iso {b = b}) p = + transport i rec {B = TypeOfHLevel _ 1} (isOfHLevelTypeOfHLevel 1) + a a b ∥₁ , squash₁) (p (~ i)) .fst) + refl ∣₁ +Iso.inv PathIdTrunc₀Iso = pRec (squash₂ _ _) (cong ∣_∣₂) +Iso.rightInv PathIdTrunc₀Iso _ = squash₁ _ _ +Iso.leftInv PathIdTrunc₀Iso _ = squash₂ _ _ _ _ + +mapFunctorial : {A B C : Type } (f : A B) (g : B C) + map g map f map (g f) +mapFunctorial f g = + funExt (elim x isSetPathImplicit) λ a refl) \ No newline at end of file diff --git a/docs/Cubical.HITs.TypeQuotients.Properties.html b/docs/Cubical.HITs.TypeQuotients.Properties.html index 94976bc..c3f69f9 100644 --- a/docs/Cubical.HITs.TypeQuotients.Properties.html +++ b/docs/Cubical.HITs.TypeQuotients.Properties.html @@ -40,14 +40,14 @@ rec f feq [ a ] = f a rec f feq (eq/ a b r i) = feq a b r i -elimProp : ((x : A /ₜ R ) isProp (B x)) +elimProp : ((x : A /ₜ R ) isProp (B x)) ((a : A) B ( [ a ])) --------------------------------- (x : A /ₜ R) B x elimProp Bprop f [ a ] = f a -elimProp Bprop f (eq/ a b r i) = isOfHLevel→isOfHLevelDep 1 Bprop (f a) (f b) (eq/ a b r) i +elimProp Bprop f (eq/ a b r i) = isOfHLevel→isOfHLevelDep 1 Bprop (f a) (f b) (eq/ a b r) i -elimProp2 : ((x y : A /ₜ R ) isProp (C x y)) +elimProp2 : ((x y : A /ₜ R ) isProp (C x y)) ((a b : A) C [ a ] [ b ]) -------------------------------------- (x y : A /ₜ R) C x y diff --git a/docs/Cubical.Induction.WellFounded.html b/docs/Cubical.Induction.WellFounded.html index eaa8701..c691557 100644 --- a/docs/Cubical.Induction.WellFounded.html +++ b/docs/Cubical.Induction.WellFounded.html @@ -5,45 +5,45 @@ open import Cubical.Foundations.Prelude -Rel : ∀{} Type ℓ' Type _ -Rel A = A A Type +module _ { ℓ'} {A : Type } (_<_ : A A Type ℓ') where + WFRec : ∀{ℓ''} (A Type ℓ'') A Type _ + WFRec P x = y y < x P y -module _ { ℓ'} {A : Type } (_<_ : A A Type ℓ') where - WFRec : ∀{ℓ''} (A Type ℓ'') A Type _ - WFRec P x = y y < x P y + data Acc (x : A) : Type (ℓ-max ℓ') where + acc : WFRec Acc x Acc x - data Acc (x : A) : Type (ℓ-max ℓ') where - acc : WFRec Acc x Acc x + WellFounded : Type _ + WellFounded = x Acc x - WellFounded : Type _ - WellFounded = x Acc x +module _ { ℓ'} {A : Type } {_<_ : A A Type ℓ'} where + isPropAcc : x isProp (Acc _<_ x) + isPropAcc x (acc p) (acc q) + = λ i acc y y<x isPropAcc y (p y y<x) (q y y<x) i) -module _ { ℓ'} {A : Type } {_<_ : A A Type ℓ'} where - isPropAcc : x isProp (Acc _<_ x) - isPropAcc x (acc p) (acc q) - = λ i acc y y<x isPropAcc y (p y y<x) (q y y<x) i) + isPropWellFounded : isProp (WellFounded _<_) + isPropWellFounded p q i a = isPropAcc a (p a) (q a) i - access : ∀{x} Acc _<_ x WFRec _<_ (Acc _<_) x - access (acc r) = r + access : ∀{x} Acc _<_ x WFRec _<_ (Acc _<_) x + access (acc r) = r - private - wfi : ∀{ℓ''} {P : A Type ℓ''} - x (wf : Acc _<_ x) - (∀ x (∀ y y < x P y) P x) - P x - wfi x (acc p) e = e x λ y y<x wfi y (p y y<x) e + private + wfi : ∀{ℓ''} {P : A Type ℓ''} + x (wf : Acc _<_ x) + (∀ x (∀ y y < x P y) P x) + P x + wfi x (acc p) e = e x λ y y<x wfi y (p y y<x) e - module WFI (wf : WellFounded _<_) where - module _ {ℓ''} {P : A Type ℓ''} (e : x (∀ y y < x P y) P x) where - private - wfi-compute : x ax wfi x ax e e x y _ wfi y (wf y) e) - wfi-compute x (acc p) - = λ i e x y y<x wfi y (isPropAcc y (p y y<x) (wf y) i) e) + module WFI (wf : WellFounded _<_) where + module _ {ℓ''} {P : A Type ℓ''} (e : x (∀ y y < x P y) P x) where + private + wfi-compute : x ax wfi x ax e e x y _ wfi y (wf y) e) + wfi-compute x (acc p) + = λ i e x y y<x wfi y (isPropAcc y (p y y<x) (wf y) i) e) - induction : x P x - induction x = wfi x (wf x) e + induction : x P x + induction x = wfi x (wf x) e - induction-compute : x induction x (e x λ y _ induction y) - induction-compute x = wfi-compute x (wf x) + induction-compute : x induction x (e x λ y _ induction y) + induction-compute x = wfi-compute x (wf x) \ No newline at end of file diff --git a/docs/Cubical.Relation.Binary.Base.html b/docs/Cubical.Relation.Binary.Base.html index 3a9046b..81e028b 100644 --- a/docs/Cubical.Relation.Binary.Base.html +++ b/docs/Cubical.Relation.Binary.Base.html @@ -20,234 +20,245 @@ open import Cubical.Relation.Nullary.Base -private - variable - ℓA ℓ≅A ℓA' ℓ≅A' : Level +open import Cubical.Induction.WellFounded -Rel : {} (A B : Type ) (ℓ' : Level) Type (ℓ-max (ℓ-suc ℓ')) -Rel A B ℓ' = A B Type ℓ' +private + variable + ℓA ℓ≅A ℓA' ℓ≅A' : Level -PropRel : {} (A B : Type ) (ℓ' : Level) Type (ℓ-max (ℓ-suc ℓ')) -PropRel A B ℓ' = Σ[ R Rel A B ℓ' ] a b isProp (R a b) +Rel : {ℓa ℓb} (A : Type ℓa) (B : Type ℓb) (ℓ' : Level) Type (ℓ-max (ℓ-max ℓa ℓb) (ℓ-suc ℓ')) +Rel A B ℓ' = A B Type ℓ' -idPropRel : {} (A : Type ) PropRel A A -idPropRel A .fst a a' = a a' ∥₁ -idPropRel A .snd _ _ = squash₁ +PropRel : {} (A B : Type ) (ℓ' : Level) Type (ℓ-max (ℓ-suc ℓ')) +PropRel A B ℓ' = Σ[ R Rel A B ℓ' ] a b isProp (R a b) -invPropRel : { ℓ'} {A B : Type } - PropRel A B ℓ' PropRel B A ℓ' -invPropRel R .fst b a = R .fst a b -invPropRel R .snd b a = R .snd a b +idPropRel : {} (A : Type ) PropRel A A +idPropRel A .fst a a' = a a' ∥₁ +idPropRel A .snd _ _ = squash₁ -compPropRel : { ℓ' ℓ''} {A B C : Type } - PropRel A B ℓ' PropRel B C ℓ'' PropRel A C (ℓ-max (ℓ-max ℓ' ℓ'')) -compPropRel R S .fst a c = Σ[ b _ ] (R .fst a b × S .fst b c) ∥₁ -compPropRel R S .snd _ _ = squash₁ +invPropRel : { ℓ'} {A B : Type } + PropRel A B ℓ' PropRel B A ℓ' +invPropRel R .fst b a = R .fst a b +invPropRel R .snd b a = R .snd a b -graphRel : {} {A B : Type } (A B) Rel A B -graphRel f a b = f a b +compPropRel : { ℓ' ℓ''} {A B C : Type } + PropRel A B ℓ' PropRel B C ℓ'' PropRel A C (ℓ-max (ℓ-max ℓ' ℓ'')) +compPropRel R S .fst a c = Σ[ b _ ] (R .fst a b × S .fst b c) ∥₁ +compPropRel R S .snd _ _ = squash₁ -module HeterogenousRelation { ℓ' : Level} {A B : Type } (R : Rel A B ℓ') where - isUniversalRel : Type (ℓ-max ℓ') - isUniversalRel = (a : A) (b : B) R a b +graphRel : {} {A B : Type } (A B) Rel A B +graphRel f a b = f a b -module BinaryRelation { ℓ' : Level} {A : Type } (R : Rel A A ℓ') where - isRefl : Type (ℓ-max ℓ') - isRefl = (a : A) R a a +module HeterogenousRelation { ℓ' : Level} {A B : Type } (R : Rel A B ℓ') where + isUniversalRel : Type (ℓ-max ℓ') + isUniversalRel = (a : A) (b : B) R a b - isIrrefl : Type (ℓ-max ℓ') - isIrrefl = (a : A) ¬ R a a +module BinaryRelation { ℓ' : Level} {A : Type } (R : Rel A A ℓ') where + isRefl : Type (ℓ-max ℓ') + isRefl = (a : A) R a a - isSym : Type (ℓ-max ℓ') - isSym = (a b : A) R a b R b a + isRefl' : Type (ℓ-max ℓ') + isRefl' = {a : A} R a a - isAsym : Type (ℓ-max ℓ') - isAsym = (a b : A) R a b ¬ R b a + isIrrefl : Type (ℓ-max ℓ') + isIrrefl = (a : A) ¬ R a a - isAntisym : Type (ℓ-max ℓ') - isAntisym = (a b : A) R a b R b a a b + isSym : Type (ℓ-max ℓ') + isSym = (a b : A) R a b R b a - isTrans : Type (ℓ-max ℓ') - isTrans = (a b c : A) R a b R b c R a c + isAsym : Type (ℓ-max ℓ') + isAsym = (a b : A) R a b ¬ R b a - -- Sum types don't play nicely with props, so we truncate - isCotrans : Type (ℓ-max ℓ') - isCotrans = (a b c : A) R a b (R a c ⊔′ R b c) + isAntisym : Type (ℓ-max ℓ') + isAntisym = (a b : A) R a b R b a a b - isWeaklyLinear : Type (ℓ-max ℓ') - isWeaklyLinear = (a b c : A) R a b R a c ⊔′ R c b + isTrans : Type (ℓ-max ℓ') + isTrans = (a b c : A) R a b R b c R a c - isConnected : Type (ℓ-max ℓ') - isConnected = (a b : A) ¬ (a b) R a b ⊔′ R b a + isTrans' : Type (ℓ-max ℓ') + isTrans' = {a b c : A} R a b R b c R a c - isStronglyConnected : Type (ℓ-max ℓ') - isStronglyConnected = (a b : A) R a b ⊔′ R b a + -- Sum types don't play nicely with props, so we truncate + isCotrans : Type (ℓ-max ℓ') + isCotrans = (a b c : A) R a b (R a c ⊔′ R b c) - isStronglyConnected→isConnected : isStronglyConnected isConnected - isStronglyConnected→isConnected strong a b _ = strong a b + isWeaklyLinear : Type (ℓ-max ℓ') + isWeaklyLinear = (a b c : A) R a b R a c ⊔′ R c b - isIrrefl×isTrans→isAsym : isIrrefl × isTrans isAsym - isIrrefl×isTrans→isAsym (irrefl , trans) a₀ a₁ Ra₀a₁ Ra₁a₀ - = irrefl a₀ (trans a₀ a₁ a₀ Ra₀a₁ Ra₁a₀) + isConnected : Type (ℓ-max ℓ') + isConnected = (a b : A) ¬ (a b) R a b ⊔′ R b a - IrreflKernel : Rel A A (ℓ-max ℓ') - IrreflKernel a b = R a b × (¬ a b) + isStronglyConnected : Type (ℓ-max ℓ') + isStronglyConnected = (a b : A) R a b ⊔′ R b a - ReflClosure : Rel A A (ℓ-max ℓ') - ReflClosure a b = R a b (a b) + isStronglyConnected→isConnected : isStronglyConnected isConnected + isStronglyConnected→isConnected strong a b _ = strong a b - SymKernel : Rel A A ℓ' - SymKernel a b = R a b × R b a + isIrrefl×isTrans→isAsym : isIrrefl × isTrans isAsym + isIrrefl×isTrans→isAsym (irrefl , trans) a₀ a₁ Ra₀a₁ Ra₁a₀ + = irrefl a₀ (trans a₀ a₁ a₀ Ra₀a₁ Ra₁a₀) - SymClosure : Rel A A ℓ' - SymClosure a b = R a b R b a + WellFounded→isIrrefl : WellFounded R isIrrefl + WellFounded→isIrrefl well = WFI.induction well λ a f Raa f a Raa Raa - AsymKernel : Rel A A ℓ' - AsymKernel a b = R a b × (¬ R b a) + IrreflKernel : Rel A A (ℓ-max ℓ') + IrreflKernel a b = R a b × (¬ a b) - NegationRel : Rel A A ℓ' - NegationRel a b = ¬ (R a b) + ReflClosure : Rel A A (ℓ-max ℓ') + ReflClosure a b = R a b (a b) - module _ - {ℓ'' : Level} - (P : Embedding A ℓ'') + SymKernel : Rel A A ℓ' + SymKernel a b = R a b × R b a - where + SymClosure : Rel A A ℓ' + SymClosure a b = R a b R b a - private - subtype : Type ℓ'' - subtype = (fst P) + AsymKernel : Rel A A ℓ' + AsymKernel a b = R a b × (¬ R b a) - toA : subtype A - toA = fst (snd P) + NegationRel : Rel A A ℓ' + NegationRel a b = ¬ (R a b) - InducedRelation : Rel subtype subtype ℓ' - InducedRelation a b = R (toA a) (toA b) + module _ + {ℓ'' : Level} + (P : Embedding A ℓ'') - record isEquivRel : Type (ℓ-max ℓ') where - constructor equivRel - field - reflexive : isRefl - symmetric : isSym - transitive : isTrans - - isUniversalRel→isEquivRel : HeterogenousRelation.isUniversalRel R isEquivRel - isUniversalRel→isEquivRel u .isEquivRel.reflexive a = u a a - isUniversalRel→isEquivRel u .isEquivRel.symmetric a b _ = u b a - isUniversalRel→isEquivRel u .isEquivRel.transitive a _ c _ _ = u a c - - isPropValued : Type (ℓ-max ℓ') - isPropValued = (a b : A) isProp (R a b) - - isSetValued : Type (ℓ-max ℓ') - isSetValued = (a b : A) isSet (R a b) - - isEffective : Type (ℓ-max ℓ') - isEffective = - (a b : A) isEquiv (eq/ {R = R} a b) - - - impliesIdentity : Type _ - impliesIdentity = {a a' : A} (R a a') (a a') - - inequalityImplies : Type _ - inequalityImplies = (a b : A) ¬ a b R a b - - -- the total space corresponding to the binary relation w.r.t. a - relSinglAt : (a : A) Type (ℓ-max ℓ') - relSinglAt a = Σ[ a' A ] (R a a') - - -- the statement that the total space is contractible at any a - contrRelSingl : Type (ℓ-max ℓ') - contrRelSingl = (a : A) isContr (relSinglAt a) - - isUnivalent : Type (ℓ-max ℓ') - isUnivalent = (a a' : A) (R a a') (a a') - - contrRelSingl→isUnivalent : isRefl contrRelSingl isUnivalent - contrRelSingl→isUnivalent ρ c a a' = isoToEquiv i - where - h : isProp (relSinglAt a) - h = isContr→isProp (c a) - aρa : relSinglAt a - aρa = a , ρ a - Q : (y : A) a y _ - Q y _ = R a y - i : Iso (R a a') (a a') - Iso.fun i r = cong fst (h aρa (a' , r)) - Iso.inv i = J Q (ρ a) - Iso.rightInv i = J y p cong fst (h aρa (y , J Q (ρ a) p)) p) - (J q _ cong fst (h aρa (a , q)) refl) - (J α _ cong fst α refl) refl - (isProp→isSet h _ _ refl (h _ _))) - (sym (JRefl Q (ρ a)))) - Iso.leftInv i r = J w β J Q (ρ a) (cong fst β) snd w) - (JRefl Q (ρ a)) (h aρa (a' , r)) - - isUnivalent→contrRelSingl : isUnivalent contrRelSingl - isUnivalent→contrRelSingl u a = q - where - abstract - f : (x : A) a x R a x - f x p = invEq (u a x) p - - t : singl a relSinglAt a - t (x , p) = x , f x p - - q : isContr (relSinglAt a) - q = isOfHLevelRespectEquiv 0 (t , totalEquiv _ _ f λ x invEquiv (u a x) .snd) - (isContrSingl a) - -EquivRel : {} (A : Type ) (ℓ' : Level) Type (ℓ-max (ℓ-suc ℓ')) -EquivRel A ℓ' = Σ[ R Rel A A ℓ' ] BinaryRelation.isEquivRel R - -EquivPropRel : {} (A : Type ) (ℓ' : Level) Type (ℓ-max (ℓ-suc ℓ')) -EquivPropRel A ℓ' = Σ[ R PropRel A A ℓ' ] BinaryRelation.isEquivRel (R .fst) - -record RelIso {A : Type ℓA} (_≅_ : Rel A A ℓ≅A) - {A' : Type ℓA'} (_≅'_ : Rel A' A' ℓ≅A') : Type (ℓ-max (ℓ-max ℓA ℓA') (ℓ-max ℓ≅A ℓ≅A')) where - constructor reliso - field - fun : A A' - inv : A' A - rightInv : (a' : A') fun (inv a') ≅' a' - leftInv : (a : A) inv (fun a) a - -open BinaryRelation - -RelIso→Iso : {A : Type ℓA} {A' : Type ℓA'} - (_≅_ : Rel A A ℓ≅A) (_≅'_ : Rel A' A' ℓ≅A') - (uni : impliesIdentity _≅_) (uni' : impliesIdentity _≅'_) - (f : RelIso _≅_ _≅'_) - Iso A A' -Iso.fun (RelIso→Iso _ _ _ _ f) = RelIso.fun f -Iso.inv (RelIso→Iso _ _ _ _ f) = RelIso.inv f -Iso.rightInv (RelIso→Iso _ _ uni uni' f) a' - = uni' (RelIso.rightInv f a') -Iso.leftInv (RelIso→Iso _ _ uni uni' f) a - = uni (RelIso.leftInv f a) - -isIrreflIrreflKernel : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isIrrefl (IrreflKernel R) -isIrreflIrreflKernel _ _ (_ , ¬a≡a) = ¬a≡a refl - -isReflReflClosure : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isRefl (ReflClosure R) -isReflReflClosure _ _ = inr refl - -isConnectedStronglyConnectedIrreflKernel : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') - isStronglyConnected R - isConnected (IrreflKernel R) -isConnectedStronglyConnectedIrreflKernel R strong a b ¬a≡b - = ∥₁.map x ⊎.rec Rab inl (Rab , ¬a≡b)) - Rba inr (Rba , b≡a ¬a≡b (sym b≡a)))) x) - (strong a b) - -isSymSymKernel : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isSym (SymKernel R) -isSymSymKernel _ _ _ (Rab , Rba) = Rba , Rab - -isSymSymClosure : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isSym (SymClosure R) -isSymSymClosure _ _ _ (inl Rab) = inr Rab -isSymSymClosure _ _ _ (inr Rba) = inl Rba - -isAsymAsymKernel : { ℓ'} {A : Type } (R : Rel A A ℓ') isAsym (AsymKernel R) -isAsymAsymKernel _ _ _ (Rab , _) (_ , ¬Rab) = ¬Rab Rab + where + + private + subtype : Type ℓ'' + subtype = (fst P) + + toA : subtype A + toA = fst (snd P) + + InducedRelation : Rel subtype subtype ℓ' + InducedRelation a b = R (toA a) (toA b) + + record isEquivRel : Type (ℓ-max ℓ') where + constructor equivRel + field + reflexive : isRefl + symmetric : isSym + transitive : isTrans + + isUniversalRel→isEquivRel : HeterogenousRelation.isUniversalRel R isEquivRel + isUniversalRel→isEquivRel u .isEquivRel.reflexive a = u a a + isUniversalRel→isEquivRel u .isEquivRel.symmetric a b _ = u b a + isUniversalRel→isEquivRel u .isEquivRel.transitive a _ c _ _ = u a c + + isPropValued : Type (ℓ-max ℓ') + isPropValued = (a b : A) isProp (R a b) + + isSetValued : Type (ℓ-max ℓ') + isSetValued = (a b : A) isSet (R a b) + + isEffective : Type (ℓ-max ℓ') + isEffective = + (a b : A) isEquiv (eq/ {R = R} a b) + + + impliesIdentity : Type _ + impliesIdentity = {a a' : A} (R a a') (a a') + + inequalityImplies : Type _ + inequalityImplies = (a b : A) ¬ a b R a b + + -- the total space corresponding to the binary relation w.r.t. a + relSinglAt : (a : A) Type (ℓ-max ℓ') + relSinglAt a = Σ[ a' A ] (R a a') + + -- the statement that the total space is contractible at any a + contrRelSingl : Type (ℓ-max ℓ') + contrRelSingl = (a : A) isContr (relSinglAt a) + + isUnivalent : Type (ℓ-max ℓ') + isUnivalent = (a a' : A) (R a a') (a a') + + contrRelSingl→isUnivalent : isRefl contrRelSingl isUnivalent + contrRelSingl→isUnivalent ρ c a a' = isoToEquiv i + where + h : isProp (relSinglAt a) + h = isContr→isProp (c a) + aρa : relSinglAt a + aρa = a , ρ a + Q : (y : A) a y _ + Q y _ = R a y + i : Iso (R a a') (a a') + Iso.fun i r = cong fst (h aρa (a' , r)) + Iso.inv i = J Q (ρ a) + Iso.rightInv i = J y p cong fst (h aρa (y , J Q (ρ a) p)) p) + (J q _ cong fst (h aρa (a , q)) refl) + (J α _ cong fst α refl) refl + (isProp→isSet h _ _ refl (h _ _))) + (sym (JRefl Q (ρ a)))) + Iso.leftInv i r = J w β J Q (ρ a) (cong fst β) snd w) + (JRefl Q (ρ a)) (h aρa (a' , r)) + + isUnivalent→contrRelSingl : isUnivalent contrRelSingl + isUnivalent→contrRelSingl u a = q + where + abstract + f : (x : A) a x R a x + f x p = invEq (u a x) p + + t : singl a relSinglAt a + t (x , p) = x , f x p + + q : isContr (relSinglAt a) + q = isOfHLevelRespectEquiv 0 (t , totalEquiv _ _ f λ x invEquiv (u a x) .snd) + (isContrSingl a) + +EquivRel : {} (A : Type ) (ℓ' : Level) Type (ℓ-max (ℓ-suc ℓ')) +EquivRel A ℓ' = Σ[ R Rel A A ℓ' ] BinaryRelation.isEquivRel R + +EquivPropRel : {} (A : Type ) (ℓ' : Level) Type (ℓ-max (ℓ-suc ℓ')) +EquivPropRel A ℓ' = Σ[ R PropRel A A ℓ' ] BinaryRelation.isEquivRel (R .fst) + +record RelIso {A : Type ℓA} (_≅_ : Rel A A ℓ≅A) + {A' : Type ℓA'} (_≅'_ : Rel A' A' ℓ≅A') : Type (ℓ-max (ℓ-max ℓA ℓA') (ℓ-max ℓ≅A ℓ≅A')) where + constructor reliso + field + fun : A A' + inv : A' A + rightInv : (a' : A') fun (inv a') ≅' a' + leftInv : (a : A) inv (fun a) a + +open BinaryRelation + +RelIso→Iso : {A : Type ℓA} {A' : Type ℓA'} + (_≅_ : Rel A A ℓ≅A) (_≅'_ : Rel A' A' ℓ≅A') + (uni : impliesIdentity _≅_) (uni' : impliesIdentity _≅'_) + (f : RelIso _≅_ _≅'_) + Iso A A' +Iso.fun (RelIso→Iso _ _ _ _ f) = RelIso.fun f +Iso.inv (RelIso→Iso _ _ _ _ f) = RelIso.inv f +Iso.rightInv (RelIso→Iso _ _ uni uni' f) a' + = uni' (RelIso.rightInv f a') +Iso.leftInv (RelIso→Iso _ _ uni uni' f) a + = uni (RelIso.leftInv f a) + +isIrreflIrreflKernel : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isIrrefl (IrreflKernel R) +isIrreflIrreflKernel _ _ (_ , ¬a≡a) = ¬a≡a refl + +isReflReflClosure : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isRefl (ReflClosure R) +isReflReflClosure _ _ = inr refl + +isConnectedStronglyConnectedIrreflKernel : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') + isStronglyConnected R + isConnected (IrreflKernel R) +isConnectedStronglyConnectedIrreflKernel R strong a b ¬a≡b + = ∥₁.map x ⊎.rec Rab inl (Rab , ¬a≡b)) + Rba inr (Rba , b≡a ¬a≡b (sym b≡a)))) x) + (strong a b) + +isSymSymKernel : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isSym (SymKernel R) +isSymSymKernel _ _ _ (Rab , Rba) = Rba , Rab + +isSymSymClosure : ∀{ ℓ'} {A : Type } (R : Rel A A ℓ') isSym (SymClosure R) +isSymSymClosure _ _ _ (inl Rab) = inr Rab +isSymSymClosure _ _ _ (inr Rba) = inl Rba + +isAsymAsymKernel : { ℓ'} {A : Type } (R : Rel A A ℓ') isAsym (AsymKernel R) +isAsymAsymKernel _ _ _ (Rab , _) (_ , ¬Rab) = ¬Rab Rab \ No newline at end of file diff --git a/docs/Cubical.Relation.Binary.Order.Preorder.Base.html b/docs/Cubical.Relation.Binary.Order.Preorder.Base.html index acb6986..7d75a3e 100644 --- a/docs/Cubical.Relation.Binary.Order.Preorder.Base.html +++ b/docs/Cubical.Relation.Binary.Order.Preorder.Base.html @@ -23,7 +23,7 @@ open import Cubical.Relation.Binary.Base open Iso -open BinaryRelation +open BinaryRelation private @@ -35,10 +35,10 @@ constructor ispreorder field - is-set : isSet A - is-prop-valued : isPropValued _≲_ - is-refl : isRefl _≲_ - is-trans : isTrans _≲_ + is-set : isSet A + is-prop-valued : isPropValued _≲_ + is-refl : isRefl _≲_ + is-trans : isTrans _≲_ unquoteDecl IsPreorderIsoΣ = declareRecordIsoΣ IsPreorderIsoΣ (quote IsPreorder) @@ -79,10 +79,10 @@ PreorderEquiv : (M : Preorder ℓ₀ ℓ₀') (M : Preorder ℓ₁ ℓ₁') Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') (ℓ-max ℓ₁ ℓ₁')) PreorderEquiv M N = Σ[ e M N ] IsPreorderEquiv (M .snd) e (N .snd) -isPropIsPreorder : {A : Type } (_≲_ : A A Type ℓ') isProp (IsPreorder _≲_) +isPropIsPreorder : {A : Type } (_≲_ : A A Type ℓ') isProp (IsPreorder _≲_) isPropIsPreorder _≲_ = isOfHLevelRetractFromIso 1 IsPreorderIsoΣ (isPropΣ isPropIsSet - λ isSetA isPropΣ (isPropΠ2 _ _ isPropIsProp)) + λ isSetA isPropΣ (isPropΠ2 _ _ isPropIsProp)) λ isPropValued≲ isProp× (isPropΠ _ isPropValued≲ _ _)) (isPropΠ4 λ _ _ _ _ isPropΠ λ _ isPropValued≲ _ _)) @@ -108,17 +108,17 @@ module S = PreorderStr (S .snd) module _ (isMon : x y x P.≲ y equivFun e x S.≲ equivFun e y) - (isMonInv : x y x S.≲ y invEq e x P.≲ invEq e y) where + (isMonInv : x y x S.≲ y invEq e x P.≲ invEq e y) where open IsPreorderEquiv open IsPreorder makeIsPreorderEquiv : IsPreorderEquiv (P .snd) e (S .snd) - pres≲ makeIsPreorderEquiv x y = propBiimpl→Equiv (P.isPreorder .is-prop-valued _ _) + pres≲ makeIsPreorderEquiv x y = propBiimpl→Equiv (P.isPreorder .is-prop-valued _ _) (S.isPreorder .is-prop-valued _ _) (isMon _ _) (isMonInv' _ _) where isMonInv' : x y equivFun e x S.≲ equivFun e y x P.≲ y - isMonInv' x y ex≲ey = transport i retEq e x i P.≲ retEq e y i) (isMonInv _ _ ex≲ey) + isMonInv' x y ex≲ey = transport i retEq e x i P.≲ retEq e y i) (isMonInv _ _ ex≲ey) module PreorderReasoning (P' : Preorder ℓ') where diff --git a/docs/Cubical.Relation.Binary.Order.Preorder.Properties.html b/docs/Cubical.Relation.Binary.Order.Preorder.Properties.html index 8f4d35a..d2ae906 100644 --- a/docs/Cubical.Relation.Binary.Order.Preorder.Properties.html +++ b/docs/Cubical.Relation.Binary.Order.Preorder.Properties.html @@ -21,21 +21,21 @@ module _ {A : Type } - {R : Rel A A ℓ'} + {R : Rel A A ℓ'} where - open BinaryRelation + open BinaryRelation - isPreorder→isEquivRelSymKernel : IsPreorder R isEquivRel (SymKernel R) + isPreorder→isEquivRelSymKernel : IsPreorder R isEquivRel (SymKernel R) isPreorder→isEquivRelSymKernel preorder - = equivRel a (IsPreorder.is-refl preorder a) + = equivRel a (IsPreorder.is-refl preorder a) , (IsPreorder.is-refl preorder a)) - (isSymSymKernel R) + (isSymSymKernel R) a b c (Rab , Rba) (Rbc , Rcb) IsPreorder.is-trans preorder a b c Rab Rbc , IsPreorder.is-trans preorder c b a Rcb Rba) - isPreorder→isStrictPosetAsymKernel : IsPreorder R IsStrictPoset (AsymKernel R) + isPreorder→isStrictPosetAsymKernel : IsPreorder R IsStrictPoset (AsymKernel R) isPreorder→isStrictPosetAsymKernel preorder = isstrictposet (IsPreorder.is-set preorder) a b isProp× (IsPreorder.is-prop-valued preorder a b) (isProp¬ (R b a))) @@ -43,9 +43,9 @@ a b c (Rab , ¬Rba) (Rbc , ¬Rcb) IsPreorder.is-trans preorder a b c Rab Rbc , λ Rca ¬Rcb (IsPreorder.is-trans preorder c a b Rca Rab)) - (isAsymAsymKernel R) + (isAsymAsymKernel R) - isPreorderInduced : IsPreorder R (B : Type ℓ'') (f : B A) IsPreorder (InducedRelation R (B , f)) + isPreorderInduced : IsPreorder R (B : Type ℓ'') (f : B A) IsPreorder (InducedRelation R (B , f)) isPreorderInduced pre B (f , emb) = ispreorder (Embedding-into-isSet→isSet (f , emb) (IsPreorder.is-set pre)) a b IsPreorder.is-prop-valued pre (f a) (f b)) @@ -54,6 +54,6 @@ Preorder→StrictPoset : Preorder ℓ' StrictPoset ℓ' Preorder→StrictPoset (_ , pre) - = _ , strictposetstr (BinaryRelation.AsymKernel (PreorderStr._≲_ pre)) + = _ , strictposetstr (BinaryRelation.AsymKernel (PreorderStr._≲_ pre)) (isPreorder→isStrictPosetAsymKernel (PreorderStr.isPreorder pre)) \ No newline at end of file diff --git a/docs/Cubical.Relation.Binary.Order.StrictPoset.Base.html b/docs/Cubical.Relation.Binary.Order.StrictPoset.Base.html index 253f67c..bc67a99 100644 --- a/docs/Cubical.Relation.Binary.Order.StrictPoset.Base.html +++ b/docs/Cubical.Relation.Binary.Order.StrictPoset.Base.html @@ -28,7 +28,7 @@ open import Cubical.Relation.Nullary.Properties open Iso -open BinaryRelation +open BinaryRelation private @@ -40,11 +40,11 @@ constructor isstrictposet field - is-set : isSet A - is-prop-valued : isPropValued _<_ - is-irrefl : isIrrefl _<_ - is-trans : isTrans _<_ - is-asym : isAsym _<_ + is-set : isSet A + is-prop-valued : isPropValued _<_ + is-irrefl : isIrrefl _<_ + is-trans : isTrans _<_ + is-asym : isAsym _<_ unquoteDecl IsStrictPosetIsoΣ = declareRecordIsoΣ IsStrictPosetIsoΣ (quote IsStrictPoset) @@ -85,10 +85,10 @@ StrictPosetEquiv : (M : StrictPoset ℓ₀ ℓ₀') (M : StrictPoset ℓ₁ ℓ₁') Type (ℓ-max (ℓ-max ℓ₀ ℓ₀') (ℓ-max ℓ₁ ℓ₁')) StrictPosetEquiv M N = Σ[ e M N ] IsStrictPosetEquiv (M .snd) e (N .snd) -isPropIsStrictPoset : {A : Type } (_<_ : A A Type ℓ') isProp (IsStrictPoset _<_) +isPropIsStrictPoset : {A : Type } (_<_ : A A Type ℓ') isProp (IsStrictPoset _<_) isPropIsStrictPoset _<_ = isOfHLevelRetractFromIso 1 IsStrictPosetIsoΣ (isPropΣ isPropIsSet - λ isSetA isPropΣ (isPropΠ2 _ _ isPropIsProp)) + λ isSetA isPropΣ (isPropΠ2 _ _ isPropIsProp)) λ isPropValued< isProp×2 (isPropΠ x isProp¬ (x < x))) (isPropΠ5 _ _ _ _ _ isPropValued< _ _)) (isPropΠ3 λ x y _ isProp¬ (y < x))) @@ -114,15 +114,15 @@ module S = StrictPosetStr (S .snd) module _ (isMon : x y x P.< y equivFun e x S.< equivFun e y) - (isMonInv : x y x S.< y invEq e x P.< invEq e y) where + (isMonInv : x y x S.< y invEq e x P.< invEq e y) where open IsStrictPosetEquiv open IsStrictPoset makeIsStrictPosetEquiv : IsStrictPosetEquiv (P .snd) e (S .snd) - pres< makeIsStrictPosetEquiv x y = propBiimpl→Equiv (P.isStrictPoset .is-prop-valued _ _) + pres< makeIsStrictPosetEquiv x y = propBiimpl→Equiv (P.isStrictPoset .is-prop-valued _ _) (S.isStrictPoset .is-prop-valued _ _) (isMon _ _) (isMonInv' _ _) where isMonInv' : x y equivFun e x S.< equivFun e y x P.< y - isMonInv' x y ex<ey = transport i retEq e x i P.< retEq e y i) (isMonInv _ _ ex<ey) + isMonInv' x y ex<ey = transport i retEq e x i P.< retEq e y i) (isMonInv _ _ ex<ey) \ No newline at end of file diff --git a/docs/Cubical.Relation.Binary.Properties.html b/docs/Cubical.Relation.Binary.Properties.html index 1a12a37..bae0642 100644 --- a/docs/Cubical.Relation.Binary.Properties.html +++ b/docs/Cubical.Relation.Binary.Properties.html @@ -3,41 +3,78 @@ module Cubical.Relation.Binary.Properties where open import Cubical.Foundations.Prelude -open import Cubical.Relation.Binary.Base +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Functions.FunExtEquiv -private - variable - : Level - A B : Type +open import Cubical.Data.Sigma +open import Cubical.Relation.Binary.Base --- Pulling back a relation along a function. --- This can for example be used when restricting an equivalence relation to a subset: --- _~'_ = on fst _~_ +private + variable + ℓ' : Level + A B : Type -module _ - (f : A B) - (R : Rel B B ) - where - open BinaryRelation +-- Pulling back a relation along a function. +-- This can for example be used when restricting an equivalence relation to a subset: +-- _~'_ = on fst _~_ - pulledbackRel : Rel A A - pulledbackRel x y = R (f x) (f y) +module _ + (f : A B) + (R : Rel B B ) + where - isReflPulledbackRel : isRefl R isRefl pulledbackRel - isReflPulledbackRel isReflR a = isReflR (f a) + open BinaryRelation - isSymPulledbackRel : isSym R isSym pulledbackRel - isSymPulledbackRel isSymR a a' = isSymR (f a) (f a') + pulledbackRel : Rel A A + pulledbackRel x y = R (f x) (f y) - isTransPulledbackRel : isTrans R isTrans pulledbackRel - isTransPulledbackRel isTransR a a' a'' = isTransR (f a) (f a') (f a'') + isReflPulledbackRel : isRefl R isRefl pulledbackRel + isReflPulledbackRel isReflR a = isReflR (f a) - open isEquivRel + isSymPulledbackRel : isSym R isSym pulledbackRel + isSymPulledbackRel isSymR a a' = isSymR (f a) (f a') - isEquivRelPulledbackRel : isEquivRel R isEquivRel pulledbackRel - reflexive (isEquivRelPulledbackRel isEquivRelR) = isReflPulledbackRel (reflexive isEquivRelR) - symmetric (isEquivRelPulledbackRel isEquivRelR) = isSymPulledbackRel (symmetric isEquivRelR) - transitive (isEquivRelPulledbackRel isEquivRelR) = isTransPulledbackRel (transitive isEquivRelR) + isTransPulledbackRel : isTrans R isTrans pulledbackRel + isTransPulledbackRel isTransR a a' a'' = isTransR (f a) (f a') (f a'') + + open isEquivRel + + isEquivRelPulledbackRel : isEquivRel R isEquivRel pulledbackRel + reflexive (isEquivRelPulledbackRel isEquivRelR) = isReflPulledbackRel (reflexive isEquivRelR) + symmetric (isEquivRelPulledbackRel isEquivRelR) = isSymPulledbackRel (symmetric isEquivRelR) + transitive (isEquivRelPulledbackRel isEquivRelR) = isTransPulledbackRel (transitive isEquivRelR) + +module _ {A B : Type } (e : A B) {_∼_ : Rel A A ℓ'} {_∻_ : Rel B B ℓ'} + (_h_ : x y (x y) ((fst e x) (fst e y))) where + + RelPathP : PathP i ua e i ua e i Type _) + _∼_ _∻_ + RelPathP i x y = Glue (ua-unglue e i x ua-unglue e i y) + λ { (i = i0) _ , x h y + ; (i = i1) _ , idEquiv _ } + + +module _ {ℓ''} {B : Type } {_∻_ : B B Type ℓ'} where + + JRelPathP-Goal : Type _ + JRelPathP-Goal = (A : Type ) (e : A B) (_~_ : A A Type ℓ') + (_h_ : x y x ~ y (fst e x fst e y)) Type ℓ'' + + + EquivJRel : (Goal : JRelPathP-Goal) + (Goal _ (idEquiv _) _∻_ λ _ _ idEquiv _ ) + {A} e _~_ _h_ Goal A e _~_ _h_ + EquivJRel Goal g {A} = EquivJ A e _~_ _h_ Goal A e _~_ _h_) + λ _~_ _h_ subst (uncurry (Goal B (idEquiv B))) + ((isPropRetract + (map-snd r funExt₂ λ x y sym (ua (r x y)))) + (map-snd r x y pathToEquiv λ i r (~ i) x y)) + (o , r) cong (o ,_) (funExt₂ λ x y equivEq + (funExt λ _ transportRefl _))) + (isPropSingl {a = _∻_})) _ _) g \ No newline at end of file diff --git a/docs/Cubical.Relation.Nullary.Base.html b/docs/Cubical.Relation.Nullary.Base.html index 8269c7f..c30a1e8 100644 --- a/docs/Cubical.Relation.Nullary.Base.html +++ b/docs/Cubical.Relation.Nullary.Base.html @@ -43,7 +43,7 @@ SplitSupport A = A ∥₁ A Collapsible : Type Type -Collapsible A = Σ[ f (A A) ] 2-Constant f +Collapsible A = Σ[ f (A A) ] 2-Constant f Populated ⟪_⟫ : Type Type Populated A = (f : Collapsible A) Fixpoint (f .fst) diff --git a/docs/Cubical.Relation.Nullary.Properties.html b/docs/Cubical.Relation.Nullary.Properties.html index 92d349f..8592ac0 100644 --- a/docs/Cubical.Relation.Nullary.Properties.html +++ b/docs/Cubical.Relation.Nullary.Properties.html @@ -41,9 +41,9 @@ EquivPresDiscrete : { ℓ'}{A : Type } {B : Type ℓ'} A B Discrete A Discrete B -EquivPresDiscrete e = isoPresDiscrete (equivToIso e) +EquivPresDiscrete e = isoPresDiscrete (equivToIso e) -isProp¬ : (A : Type ) isProp (¬ A) +isProp¬ : (A : Type ) isProp (¬ A) isProp¬ A p q i x = isProp⊥ (p x) (q x) i Stable¬ : Stable (¬ A) @@ -76,7 +76,7 @@ discreteDec Adis (no ¬p) (yes p) = ⊥.rec (¬p p) discreteDec {A = A} Adis (no ¬p) (no ¬p') = yes (cong no (isProp¬ A ¬p ¬p')) -isPropDec : (Aprop : isProp A) isProp (Dec A) +isPropDec : (Aprop : isProp A) isProp (Dec A) isPropDec Aprop (yes a) (yes a') = cong yes (Aprop a a') isPropDec Aprop (yes a) (no ¬a) = ⊥.rec (¬a a) isPropDec Aprop (no ¬a) (yes a) = ⊥.rec (¬a a) @@ -88,7 +88,7 @@ EquivPresDec : { ℓ'}{A : Type } {B : Type ℓ'} A B Dec A Dec B -EquivPresDec p = mapDec (p .fst) f f invEq p) +EquivPresDec p = mapDec (p .fst) f f invEq p) ¬→¬∥∥ : ¬ A ¬ A ∥₁ ¬→¬∥∥ ¬p a ∣₁ = ¬p a @@ -135,7 +135,7 @@ SplitSupport→Collapsible {A = A} hst = h , hIsConst where h : A A h p = hst p ∣₁ - hIsConst : 2-Constant h + hIsConst : 2-Constant h hIsConst p q i = hst (squash₁ p ∣₁ q ∣₁ i) Collapsible→SplitSupport : Collapsible A SplitSupport A @@ -149,7 +149,7 @@ -- stability of path space under truncation or path collapsability are necessary and sufficient properties -- for a a type to be a set. -Collapsible≡→isSet : Collapsible≡ A isSet A +Collapsible≡→isSet : Collapsible≡ A isSet A Collapsible≡→isSet {A = A} col a b p q = p≡q where f : (x : A) a x a x f x = col a x .fst @@ -163,10 +163,10 @@ ; (j = i0) rem p i k ; (j = i1) rem q i k }) a -HSeparated→isSet : HSeparated A isSet A +HSeparated→isSet : HSeparated A isSet A HSeparated→isSet = Collapsible≡→isSet HSeparated→Collapsible≡ -isSet→HSeparated : isSet A HSeparated A +isSet→HSeparated : isSet A HSeparated A isSet→HSeparated setA x y = extract where extract : x y ∥₁ x y extract p ∣₁ = p @@ -176,13 +176,13 @@ PStable≡→HSeparated : PStable≡ A HSeparated A PStable≡→HSeparated pst x y = PStable→SplitSupport (pst x y) -PStable≡→isSet : PStable≡ A isSet A +PStable≡→isSet : PStable≡ A isSet A PStable≡→isSet = HSeparated→isSet PStable≡→HSeparated Separated→PStable≡ : Separated A PStable≡ A Separated→PStable≡ st x y = Stable→PStable (st x y) -Separated→isSet : Separated A isSet A +Separated→isSet : Separated A isSet A Separated→isSet = PStable≡→isSet Separated→PStable≡ SeparatedΠ : (∀ x Separated (P x)) -> Separated ((x : A) -> P x) @@ -200,6 +200,6 @@ Discrete→Separated : Discrete A Separated A Discrete→Separated d x y = Dec→Stable (d x y) -Discrete→isSet : Discrete A isSet A +Discrete→isSet : Discrete A isSet A Discrete→isSet = Separated→isSet Discrete→Separated \ No newline at end of file diff --git a/docs/Cubical.Structures.Axioms.html b/docs/Cubical.Structures.Axioms.html index 2ec96dc..9b01e7c 100644 --- a/docs/Cubical.Structures.Axioms.html +++ b/docs/Cubical.Structures.Axioms.html @@ -36,18 +36,18 @@ axiomsUnivalentStr : {S : Type Type ℓ₁} (ι : (A B : TypeWithStr S) A .fst B .fst Type ℓ₁') {axioms : (X : Type ) S X Type ℓ₂} - (axioms-are-Props : (X : Type ) (s : S X) isProp (axioms X s)) + (axioms-are-Props : (X : Type ) (s : S X) isProp (axioms X s)) (θ : UnivalentStr S ι) UnivalentStr (AxiomsStructure S axioms) (AxiomsEquivStr ι axioms) axiomsUnivalentStr {S = S} ι {axioms = axioms} axioms-are-Props θ {X , s , a} {Y , t , b} e = ι (X , s) (Y , t) e - ≃⟨ θ e + ≃⟨ θ e PathP i S (ua e i)) s t - ≃⟨ invEquiv (Σ-contractSnd λ _ isOfHLevelPathP' 0 (axioms-are-Props _ _) _ _) + ≃⟨ invEquiv (Σ-contractSnd λ _ isOfHLevelPathP' 0 (axioms-are-Props _ _) _ _) Σ[ p PathP i S (ua e i)) s t ] PathP i axioms (ua e i) (p i)) a b - ≃⟨ ΣPath≃PathΣ + ≃⟨ ΣPath≃PathΣ PathP i AxiomsStructure S axioms (ua e i)) (s , a) (t , b) - + inducedStructure : {S : Type Type ℓ₁} {ι : (A B : TypeWithStr S) A .fst B .fst Type ℓ₁'} diff --git a/docs/Cubical.Structures.Pointed.html b/docs/Cubical.Structures.Pointed.html index e54549f..f1bf96d 100644 --- a/docs/Cubical.Structures.Pointed.html +++ b/docs/Cubical.Structures.Pointed.html @@ -28,7 +28,7 @@ PointedEquivStr A B f = equivFun f (pt A) pt B pointedUnivalentStr : UnivalentStr {} PointedStructure PointedEquivStr -pointedUnivalentStr f = invEquiv (ua-ungluePath-Equiv f) +pointedUnivalentStr f = invEquiv (ua-ungluePath-Equiv f) pointedSIP : (A B : Pointed ) A ≃[ PointedEquivStr ] B (A B) pointedSIP = SIP pointedUnivalentStr @@ -48,10 +48,10 @@ -} abstract pointed-sip⁻ : (A B : Pointed ) (A B) A ≃[ PointedEquivStr ] B - pointed-sip⁻ A B = invEq (pointedSIP A B) + pointed-sip⁻ A B = invEq (pointedSIP A B) pointed-sip⁻-refl : (A : Pointed ) pointed-sip⁻ A A refl idEquiv∙ A - pointed-sip⁻-refl A = sym (invEq (equivAdjointEquiv (pointedSIP A A)) (pointed-sip-idEquiv∙ A)) + pointed-sip⁻-refl A = sym (invEq (equivAdjointEquiv (pointedSIP A A)) (pointed-sip-idEquiv∙ A)) pointedEquivAction : EquivAction {} PointedStructure pointedEquivAction e = e diff --git a/docs/Cubical.Tactics.NatSolver.EvalHom.html b/docs/Cubical.Tactics.NatSolver.EvalHom.html new file mode 100644 index 0000000..3956463 --- /dev/null +++ b/docs/Cubical.Tactics.NatSolver.EvalHom.html @@ -0,0 +1,198 @@ + +Cubical.Tactics.NatSolver.EvalHom
{-# OPTIONS --safe #-}
+module Cubical.Tactics.NatSolver.EvalHom where
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+
+open import Cubical.Tactics.NatSolver.HornerForms
+
+private
+  variable
+     : Level
+
+module HomomorphismProperties where
+  open IteratedHornerOperations
+
+  evalHom+0 : {n : } (P : IteratedHornerForms n) (xs : Vec  n)
+       eval (0ₕ +ₕ P) xs  eval P xs
+  evalHom+0 (const x) [] = refl
+  evalHom+0 _ (x  xs) = refl
+
+  eval0H : {n : } (xs : Vec  n)
+          eval 0ₕ xs  0
+  eval0H [] = refl
+  eval0H (x  xs) = refl
+
+  eval1ₕ : {n : } (xs : Vec  n)
+          eval 1ₕ xs  1
+  eval1ₕ [] = refl
+  eval1ₕ (x  xs) =
+    eval 1ₕ (x  xs)                             ≡⟨ refl 
+    eval (0H ·X+ 1ₕ) (x  xs)                    ≡⟨ refl 
+    eval 0H (x  xs) · x + eval 1ₕ xs            ≡⟨ cong  u  u · x + eval 1ₕ xs)
+                                                        (eval0H (x  xs)) 
+    0 · x + eval 1ₕ xs                           ≡⟨ cong  u  0 · x + u)
+                                                        (eval1ₕ xs) 
+    1 
+
+
+  +ShufflePairs : (a b c d : )
+                 (a + b) + (c + d)  (a + c) + (b + d)
+  +ShufflePairs a b c d =
+    (a + b) + (c + d) ≡⟨ +-assoc (a + b) c d 
+    ((a + b) + c) + d ≡⟨ cong  u  u + d) (sym (+-assoc a b c)) 
+    (a + (b + c)) + d ≡⟨ cong  u  (a + u) + d) (+-comm b c) 
+    (a + (c + b)) + d ≡⟨ cong  u  u + d) (+-assoc a c b) 
+    ((a + c) + b) + d ≡⟨ sym (+-assoc (a + c) b d) 
+    (a + c) + (b + d) 
+
+  +Homeval :
+    {n : } (P Q : IteratedHornerForms n) (xs : Vec  n)
+     eval (P +ₕ Q) xs  (eval P xs) + (eval Q xs)
+  +Homeval (const x) (const y) [] = refl
+  +Homeval 0H Q xs =
+    eval (0H +ₕ Q) xs            ≡⟨ refl 
+    0 + eval Q xs               ≡⟨ cong  u  u + eval Q xs) (sym (eval0H xs)) 
+    eval 0H xs + eval Q xs 
+  +Homeval (P ·X+ Q) 0H xs =
+    eval ((P ·X+ Q) +ₕ 0H) xs       ≡⟨ refl 
+    eval (P ·X+ Q) xs              ≡⟨ sym (+-zero _) 
+    eval (P ·X+ Q) xs + 0          ≡⟨ cong  u  eval (P ·X+ Q) xs + u)
+                                           (sym (eval0H xs)) 
+    eval (P ·X+ Q) xs + eval 0H xs 
+  +Homeval (P ·X+ Q) (S ·X+ T) (x  xs) =
+    eval ((P ·X+ Q) +ₕ (S ·X+ T)) (x  xs)
+   ≡⟨ refl 
+    eval ((P +ₕ S) ·X+ (Q +ₕ T)) (x  xs)
+   ≡⟨ refl 
+    (eval (P +ₕ S) (x  xs)) · x + eval (Q +ₕ T) xs
+   ≡⟨ cong  u  (eval (P +ₕ S) (x  xs)) · x + u) (+Homeval Q T xs) 
+    (eval (P +ₕ S) (x  xs)) · x + (eval Q xs + eval T xs)
+   ≡⟨ cong  u  u · x + (eval Q xs + eval T xs)) (+Homeval P S (x  xs)) 
+    (eval P (x  xs) + eval S (x  xs)) · x
+    + (eval Q xs + eval T xs)
+   ≡⟨ cong  u  u + (eval Q xs + eval T xs))
+     (sym (·-distribʳ (eval P (x  xs)) (eval S (x  xs)) x)) 
+    (eval P (x  xs)) · x + (eval S (x  xs)) · x
+    + (eval Q xs + eval T xs)
+   ≡⟨ +ShufflePairs ((eval P (x  xs)) · x) ((eval S (x  xs)) · x) (eval Q xs) (eval T xs) 
+    ((eval P (x  xs)) · x + eval Q xs)
+    + ((eval S (x  xs)) · x + eval T xs)
+   
+
+  ⋆Homeval : {n : }
+             (r : IteratedHornerForms n)
+             (P : IteratedHornerForms (ℕ.suc n)) (x : ) (xs : Vec  n)
+            eval (r  P) (x  xs)  eval r xs · eval P (x  xs)
+
+
+  ⋆0LeftAnnihilates :
+    {n : } (P : IteratedHornerForms (ℕ.suc n)) (xs : Vec  (ℕ.suc n))
+     eval (0ₕ  P) xs  0
+
+  ·Homeval : {n : } (P Q : IteratedHornerForms n) (xs : Vec  n)
+     eval (P ·ₕ Q) xs  (eval P xs) · (eval Q xs)
+
+  ⋆0LeftAnnihilates 0H xs = eval0H xs
+  ⋆0LeftAnnihilates (P ·X+ Q) (x  xs) =
+      eval (0ₕ  (P ·X+ Q)) (x  xs)                    ≡⟨ refl 
+      eval ((0ₕ  P) ·X+ (0ₕ ·ₕ Q)) (x  xs)             ≡⟨ refl 
+      (eval (0ₕ  P) (x  xs)) · x + eval (0ₕ ·ₕ Q) xs
+    ≡⟨ cong  u  (u · x) + eval (0ₕ ·ₕ Q) _) (⋆0LeftAnnihilates P (x  xs)) 
+      0 · x + eval (0ₕ ·ₕ Q) xs
+    ≡⟨ ·Homeval 0ₕ Q _ 
+      eval 0ₕ xs · eval Q xs
+    ≡⟨ cong  u  u · eval Q xs) (eval0H xs) 
+      0 · eval Q xs 
+
+  ⋆Homeval r 0H x xs =
+    eval (r  0H) (x  xs)         ≡⟨ refl 
+    0                              ≡⟨ 0≡m·0 (eval r xs) 
+    eval r xs · 0                  ≡⟨ refl 
+    eval r xs · eval 0H (x  xs) 
+  ⋆Homeval r (P ·X+ Q) x xs =
+      eval (r  (P ·X+ Q)) (x  xs)                    ≡⟨ refl 
+      eval ((r  P) ·X+ (r ·ₕ Q)) (x  xs)              ≡⟨ refl 
+      (eval (r  P) (x  xs)) · x + eval (r ·ₕ Q) xs
+    ≡⟨ cong  u  u · x + eval (r ·ₕ Q) xs) (⋆Homeval r P x xs) 
+      (eval r xs · eval P (x  xs)) · x + eval (r ·ₕ Q) xs
+    ≡⟨ cong  u  (eval r xs · eval P (x  xs)) · x + u) (·Homeval r Q xs) 
+      (eval r xs · eval P (x  xs)) · x + eval r xs · eval Q xs
+    ≡⟨ cong  u  u  + eval r xs · eval Q xs) (sym (·-assoc (eval r xs) (eval P (x  xs)) x)) 
+      eval r xs · (eval P (x  xs) · x) + eval r xs · eval Q xs
+    ≡⟨ ·-distribˡ (eval r xs)  ((eval P (x  xs) · x)) (eval Q xs) 
+      eval r xs · ((eval P (x  xs) · x) + eval Q xs)
+    ≡⟨ refl 
+      eval r xs · eval (P ·X+ Q) (x  xs) 
+
+  combineCases :
+    {n : } (Q : IteratedHornerForms n) (P S : IteratedHornerForms (ℕ.suc n))
+    (xs : Vec  (ℕ.suc n))
+     eval ((P ·X+ Q) ·ₕ S) xs  eval (((P ·ₕ S) ·X+ 0ₕ) +ₕ (Q  S)) xs
+  combineCases Q P S (x  xs) with (P ·ₕ S)
+  ... | 0H =
+    eval (Q  S) (x  xs)                ≡⟨ refl 
+    0 + eval (Q  S) (x  xs)           ≡⟨ cong  u  u + eval (Q  S) (x  xs)) lemma 
+    eval (0H ·X+ 0ₕ) (x  xs)
+    + eval (Q  S) (x  xs)              ≡⟨ sym (+Homeval
+                                                  (0H ·X+ 0ₕ) (Q  S) (x  xs)) 
+    eval ((0H ·X+ 0ₕ) +ₕ (Q  S)) (x  xs) 
+    where lemma : 0  eval (0H ·X+ 0ₕ) (x  xs)
+          lemma = 0
+                ≡⟨ refl 
+                  0 + 0
+                ≡⟨ cong  u  u + 0) refl 
+                  0 · x + 0
+                ≡⟨ cong  u  0 · x + u) (sym (eval0H xs)) 
+                  0 · x + eval 0ₕ xs
+                ≡⟨ cong  u  u · x + eval 0ₕ xs) (sym (eval0H (x  xs))) 
+                  eval 0H (x  xs) · x + eval 0ₕ xs
+                ≡⟨ refl 
+                  eval (0H ·X+ 0ₕ) (x  xs) 
+  ... | (_ ·X+ _) = refl
+
+  ·Homeval (const x) (const y) [] = refl
+  ·Homeval 0H Q xs =
+    eval (0H ·ₕ Q) xs        ≡⟨ eval0H xs 
+    0                                 ≡⟨ refl 
+    0 · eval Q xs          ≡⟨ cong  u  u · eval Q xs) (sym (eval0H xs)) 
+    eval 0H xs · eval Q xs 
+  ·Homeval (P ·X+ Q) S (x  xs) =
+      eval ((P ·X+ Q) ·ₕ S) (x  xs)
+    ≡⟨ combineCases Q P S (x  xs) 
+      eval (((P ·ₕ S) ·X+ 0ₕ) +ₕ (Q  S)) (x  xs)
+    ≡⟨ +Homeval ((P ·ₕ S) ·X+ 0ₕ) (Q  S) (x  xs) 
+      eval ((P ·ₕ S) ·X+ 0ₕ) (x  xs) + eval (Q  S) (x  xs)
+    ≡⟨ refl 
+      (eval (P ·ₕ S) (x  xs) · x + eval 0ₕ xs)
+      + eval (Q  S) (x  xs)
+    ≡⟨ cong  u  u + eval (Q  S) (x  xs))
+          ((eval (P ·ₕ S) (x  xs) · x + eval 0ₕ xs)
+         ≡⟨ cong  u  eval (P ·ₕ S) (x  xs) · x + u) (eval0H xs) 
+           (eval (P ·ₕ S) (x  xs) · x + 0)
+         ≡⟨ +-zero _ 
+           (eval (P ·ₕ S) (x  xs) · x)
+         ≡⟨ cong  u  u · x) (·Homeval P S (x  xs)) 
+           ((eval P (x  xs) · eval S (x  xs)) · x)
+         ≡⟨ sym (·-assoc (eval P (x  xs)) (eval S (x  xs)) x) 
+           (eval P (x  xs) · (eval S (x  xs) · x))
+         ≡⟨ cong  u  eval P (x  xs) · u) (·-comm _ x) 
+           (eval P (x  xs) · (x · eval S (x  xs)))
+         ≡⟨ ·-assoc (eval P (x  xs)) x (eval S (x  xs)) 
+           (eval P (x  xs) · x) · eval S (x  xs)
+          ) 
+      (eval P (x  xs) · x) · eval S (x  xs)
+      + eval (Q  S) (x  xs)
+    ≡⟨ cong  u  (eval P (x  xs) · x) · eval S (x  xs) + u)
+            (⋆Homeval Q S x xs) 
+      (eval P (x  xs) · x) · eval S (x  xs)
+      + eval Q xs · eval S (x  xs)
+    ≡⟨ ·-distribʳ (eval P (x  xs) · x) (eval Q xs) (eval S (x  xs)) 
+      ((eval P (x  xs) · x) + eval Q xs) · eval S (x  xs)
+    ≡⟨ refl 
+      eval (P ·X+ Q) (x  xs) · eval S (x  xs) 
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.NatSolver.HornerForms.html b/docs/Cubical.Tactics.NatSolver.HornerForms.html new file mode 100644 index 0000000..6a6caa4 --- /dev/null +++ b/docs/Cubical.Tactics.NatSolver.HornerForms.html @@ -0,0 +1,102 @@ + +Cubical.Tactics.NatSolver.HornerForms
{-# OPTIONS --safe #-}
+module Cubical.Tactics.NatSolver.HornerForms where
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Data.Nat hiding (isZero)
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Bool using (Bool; true; false; if_then_else_)
+
+private
+  variable
+     : Level
+
+{-
+  This defines the type of multivariate Polynomials over ℕ.
+  The construction is based on the algebraic fact
+
+    ℕ[X₀][X₁]⋯[Xₙ] ≅ ℕ[X₀,⋯,Xₙ]
+
+  BUT: Contrary to algebraic convetions, we will give 'Xₙ' the lowest index
+  in the definition of 'Variable' below. So if 'Variable n k' is identified
+  with 'Xₖ', what we construct should rather be denoted with
+
+    ℕ[Xₙ][Xₙ₋₁]⋯[X₀]
+
+  or, to be precise about the evaluation order:
+
+    (⋯((ℕ[Xₙ])[Xₙ₋₁])⋯)[X₀]
+
+-}
+
+data IteratedHornerForms :   Type ℓ-zero where
+  const :   IteratedHornerForms ℕ.zero
+  0H : {n : }  IteratedHornerForms (ℕ.suc n)
+  _·X+_ : {n : }  IteratedHornerForms (ℕ.suc n)  IteratedHornerForms n
+                   IteratedHornerForms (ℕ.suc n)
+
+eval : {n : } (P : IteratedHornerForms n)
+        Vec  n  
+eval (const r) [] = r
+eval 0H (_  _) = 0
+eval (P ·X+ Q) (x  xs) =
+  (eval P (x  xs)) · x + eval Q xs
+
+module IteratedHornerOperations where
+
+  private
+    1H' : (n : )  IteratedHornerForms n
+    1H' ℕ.zero = const 1
+    1H' (ℕ.suc n) = 0H ·X+ 1H' n
+
+    0H' : (n : )  IteratedHornerForms n
+    0H' ℕ.zero = const 0
+    0H' (ℕ.suc n) = 0H
+
+  1ₕ : {n : }  IteratedHornerForms n
+  1ₕ {n = n} = 1H' n
+
+  0ₕ : {n : }  IteratedHornerForms n
+  0ₕ {n = n} = 0H' n
+
+  X : (n : ) (k : Fin n)  IteratedHornerForms n
+  X (ℕ.suc m) zero = 1ₕ ·X+ 0ₕ
+  X (ℕ.suc m) (suc k) = 0ₕ ·X+ X m k
+
+  _+ₕ_ : {n : }  IteratedHornerForms n  IteratedHornerForms n
+                IteratedHornerForms n
+  (const r) +ₕ (const s) = const (r + s)
+  0H +ₕ Q = Q
+  (P ·X+ r) +ₕ 0H = P ·X+ r
+  (P ·X+ r) +ₕ (Q ·X+ s) = (P +ₕ Q) ·X+ (r +ₕ s)
+
+  isZero : {n : }  IteratedHornerForms (ℕ.suc n)
+                    Bool
+  isZero 0H = true
+  isZero (P ·X+ P₁) = false
+
+  _⋆_ : {n : }  IteratedHornerForms n  IteratedHornerForms (ℕ.suc n)
+                 IteratedHornerForms (ℕ.suc n)
+  _·ₕ_ : {n : }  IteratedHornerForms n  IteratedHornerForms n
+                 IteratedHornerForms n
+  r  0H = 0H
+  r  (P ·X+ Q) = (r  P) ·X+ (r ·ₕ Q)
+
+  const x ·ₕ const y = const (x · y)
+  0H ·ₕ Q = 0H
+  (P ·X+ Q) ·ₕ S =
+     let
+        z = (P ·ₕ S)
+     in if (isZero z)
+        then (Q  S)
+        else (z ·X+ 0ₕ) +ₕ (Q  S)
+
+Variable : (n : ) (k : Fin n)  IteratedHornerForms n
+Variable n k = IteratedHornerOperations.X n k
+
+Constant : (n : ) (r : )  IteratedHornerForms n
+Constant ℕ.zero r = const r
+Constant (ℕ.suc n) r = IteratedHornerOperations.0ₕ ·X+ Constant n r
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.NatSolver.NatExpression.html b/docs/Cubical.Tactics.NatSolver.NatExpression.html new file mode 100644 index 0000000..f4e61f1 --- /dev/null +++ b/docs/Cubical.Tactics.NatSolver.NatExpression.html @@ -0,0 +1,30 @@ + +Cubical.Tactics.NatSolver.NatExpression
{-# OPTIONS --safe #-}
+module Cubical.Tactics.NatSolver.NatExpression where
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Data.FinData
+open import Cubical.Data.Nat
+open import Cubical.Data.Nat.Order using (zero-≤)
+open import Cubical.Data.Vec.Base
+
+infixl 6 _+'_
+infixl 7 _·'_
+
+-- Expression in a ring on A with n variables
+data Expr (n : ) : Type ℓ-zero where
+  K :   Expr n
+   : Fin n  Expr n
+  _+'_ : Expr n  Expr n  Expr n
+  _·'_ : Expr n  Expr n  Expr n
+
+module Eval where
+  open import Cubical.Data.Vec
+
+  ⟦_⟧ :  {n}  Expr n  Vec  n  
+   K r  v = r
+    k  v = lookup k v
+   x +' y  v =   x  v +  y  v
+   x ·' y  v =  x  v ·  y  v
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.NatSolver.Reflection.html b/docs/Cubical.Tactics.NatSolver.Reflection.html new file mode 100644 index 0000000..0ff39a4 --- /dev/null +++ b/docs/Cubical.Tactics.NatSolver.Reflection.html @@ -0,0 +1,147 @@ + +Cubical.Tactics.NatSolver.Reflection
{-# OPTIONS --safe #-}
+{-
+  This is inspired by/copied from:
+  https://github.com/agda/agda-stdlib/blob/master/src/Tactic/MonoidSolver.agda
+  and the 1lab.
+
+  Boilerplate code for calling the ring solver is constructed automatically
+  with agda's reflection features.
+-}
+module Cubical.Tactics.NatSolver.Reflection where
+
+open import Cubical.Foundations.Prelude hiding (Type)
+open import Cubical.Functions.Logic
+
+open import Agda.Builtin.Reflection hiding (Type)
+open import Agda.Builtin.String
+
+open import Cubical.Reflection.Base
+
+open import Cubical.Data.Maybe
+open import Cubical.Data.Sigma
+open import Cubical.Data.List
+open import Cubical.Data.Nat
+open import Cubical.Data.Bool
+open import Cubical.Data.Bool.SwitchStatement
+open import Cubical.Data.Vec using (Vec) renaming ([] to emptyVec; _∷_ to _∷vec_) public
+
+open import Cubical.Tactics.NatSolver.NatExpression
+open import Cubical.Tactics.NatSolver.Solver
+
+open import Cubical.Tactics.Reflection
+open import Cubical.Tactics.Reflection.Variables
+open import Cubical.Tactics.Reflection.Utilities
+
+open EqualityToNormalform renaming (solve to natSolve)
+private
+  variable
+     : Level
+
+  private
+    solverCallAsTerm : Arg Term  Term  Term  Term
+    solverCallAsTerm varList lhs rhs =
+      def
+         (quote natSolve)
+         (varg lhs  varg rhs
+            varList
+            varg (def (quote refl) [])  [])
+
+  solverCallWithVars :   Vars  Term  Term  Term
+  solverCallWithVars n vars lhs rhs =
+      solverCallAsTerm (variableList vars) lhs rhs
+      where
+        variableList : Vars  Arg Term
+        variableList [] = varg (con (quote emptyVec) [])
+        variableList (t  ts)
+          = varg (con (quote _∷vec_) (t v∷ (variableList ts)  []))
+
+module pr {n : } where
+  0' : Expr n
+  0' = K 0
+
+  1' : Expr n
+  1' = K 1
+
+module NatSolverReflection where
+  open pr
+
+  buildExpression : Term  TC (Template × Vars)
+
+  op2 : Name  Term  Term  TC (Template × Vars)
+  op2 op x y = do
+    r1  buildExpression x
+    r2  buildExpression y
+    returnTC ((λ ass  con op (fst r1 ass v∷ fst r2 ass v∷ [])) ,
+             appendWithoutRepetition (snd r1) (snd r2))
+
+  `_·_` : List (Arg Term)  TC (Template × Vars)
+  `_·_` (_ h∷ xs) = `_·_` xs
+  `_·_` (x v∷ y v∷ []) = op2 (quote _·'_) x y
+  `_·_` ts = errorOut ts
+
+  `_+_` : List (Arg Term)  TC (Template × Vars)
+  `_+_` (_ h∷ xs) = `_+_` xs
+  `_+_` (x v∷ y v∷ []) = op2 (quote _+'_) x y
+  `_+_` ts = errorOut ts
+
+  `1+_` : List (Arg Term)  TC (Template × Vars)
+  `1+_` (x v∷ []) = do
+    r1  buildExpression x
+    returnTC ((λ ass  con (quote _+'_) ((def (quote 1') []) v∷ fst r1 ass v∷ [])) ,
+              snd r1)
+  `1+_` ts = errorOut ts
+
+  K' : List (Arg Term)  TC (Template × Vars)
+  K' xs = returnTC ((λ _  con (quote K) xs) , [])
+
+  polynomialVariable : Maybe   Term
+  polynomialVariable (just n) = con (quote ) (finiteNumberAsTerm (just n) v∷ [])
+  polynomialVariable nothing = unknown
+
+  buildExpression v@(var _ _) =
+      returnTC ((λ ass  polynomialVariable (ass v)) ,
+           v  [])
+  buildExpression t@(lit n) = K' (t v∷ [])
+  buildExpression t@(def n xs) =
+    switch (n ==_) cases
+      case (quote _·_)  `_·_` xs   break
+      case (quote _+_)  `_+_` xs   break
+      default⇒ (K' xs)
+  buildExpression t@(con n xs) =
+    switch (n ==_) cases
+      case (quote suc)  `1+_` xs   break
+      default⇒ (K' xs)
+  buildExpression t = errorOut' t
+
+  toNatExpression : Term × Term  TC (Term × Term × Vars)
+  toNatExpression (lhs , rhs) = do
+      r1  buildExpression lhs
+      r2  buildExpression rhs
+      vars  returnTC (appendWithoutRepetition (snd r1) (snd r2))
+      returnTC (
+        let ass : VarAss
+            ass n = indexOf n vars
+        in (fst r1 ass , fst r2 ass , vars ))
+
+private
+
+  solve!-macro : Term  TC Unit
+  solve!-macro hole =
+    do
+      goal  inferType hole >>= normalise
+
+      just (lhs , rhs)  get-boundary goal
+        where
+          nothing
+             typeError(strErr "The NatSolver failed to parse the goal "
+                                termErr goal  [])
+
+      (lhs' , rhs' , vars)  NatSolverReflection.toNatExpression (lhs , rhs)
+      let solution = solverCallWithVars (length vars) vars lhs' rhs'
+      unify hole solution
+
+macro
+  solveℕ! : Term  TC _
+  solveℕ! = solve!-macro
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.NatSolver.Solver.html b/docs/Cubical.Tactics.NatSolver.Solver.html new file mode 100644 index 0000000..efddecf --- /dev/null +++ b/docs/Cubical.Tactics.NatSolver.Solver.html @@ -0,0 +1,125 @@ + +Cubical.Tactics.NatSolver.Solver
{-# OPTIONS --safe #-}
+module Cubical.Tactics.NatSolver.Solver where
+
+open import Cubical.Foundations.Prelude
+
+open import Cubical.Data.FinData
+open import Cubical.Data.Nat
+open import Cubical.Data.Nat.Order using (zero-≤)
+open import Cubical.Data.Vec.Base
+open import Cubical.Tactics.NatSolver.NatExpression
+open import Cubical.Tactics.NatSolver.HornerForms
+open import Cubical.Tactics.NatSolver.EvalHom
+
+private
+  variable
+     : Level
+
+module EqualityToNormalform where
+  open Eval
+  open IteratedHornerOperations
+  open HomomorphismProperties
+
+  normalize : {n : }  Expr n  IteratedHornerForms n
+  normalize {n = n} (K r) = Constant n r
+  normalize {n = n} ( k) = Variable n k
+  normalize (x +' y) =
+    (normalize x) +ₕ (normalize y)
+  normalize (x ·' y) =
+    (normalize x) ·ₕ (normalize y)
+
+  isEqualToNormalform :
+            {n : }
+            (e : Expr n) (xs : Vec  n)
+           eval (normalize e) xs   e  xs
+  isEqualToNormalform (K r) [] = refl
+  isEqualToNormalform {n = ℕ.suc n} (K r) (x  xs) =
+     eval (Constant (ℕ.suc n) r) (x  xs)           ≡⟨ refl 
+     eval (0ₕ ·X+ Constant n r) (x  xs)             ≡⟨ refl 
+     eval 0ₕ (x  xs) · x + eval (Constant n r) xs
+    ≡⟨ cong  u  u · x + eval (Constant n r) xs) (eval0H (x  xs)) 
+     0 · x + eval (Constant n r) xs
+    ≡⟨ refl 
+     eval (Constant n r) xs
+    ≡⟨ isEqualToNormalform (K r) xs 
+     r 
+
+  isEqualToNormalform ( zero) (x  xs) =
+    eval (1ₕ ·X+ 0ₕ) (x  xs)           ≡⟨ refl 
+    eval 1ₕ (x  xs) · x + eval 0ₕ xs   ≡⟨ cong  u  u · x + eval 0ₕ xs)
+                                               (eval1ₕ (x  xs)) 
+    1 · x + eval 0ₕ xs                  ≡⟨ cong  u  1 · x + u ) (eval0H xs) 
+    1 · x + 0                          ≡⟨ +-zero _ 
+    1 · x                               ≡⟨ ·-identityˡ _ 
+    x 
+  isEqualToNormalform {n = ℕ.suc n} ( (suc k)) (x  xs) =
+      eval (0ₕ ·X+ Variable n k) (x  xs)             ≡⟨ refl 
+      eval 0ₕ (x  xs) · x + eval (Variable n k) xs
+    ≡⟨ cong  u  u · x + eval (Variable n k) xs) (eval0H (x  xs)) 
+      0 · x + eval (Variable n k) xs
+    ≡⟨ refl 
+      eval (Variable n k) xs
+    ≡⟨ isEqualToNormalform ( k) xs 
+        (suc k)  (x  xs) 
+
+  isEqualToNormalform (e +' e₁) [] =
+        eval (normalize e +ₕ normalize e₁) []
+      ≡⟨ +Homeval (normalize e) _ [] 
+        eval (normalize e) []
+        + eval (normalize e₁) []
+      ≡⟨ cong  u  u + eval (normalize e₁) [])
+              (isEqualToNormalform e []) 
+         e  []
+        + eval (normalize e₁) []
+      ≡⟨ cong  u   e  [] + u) (isEqualToNormalform e₁ []) 
+         e  [] +  e₁  [] 
+  isEqualToNormalform (e +' e₁) (x  xs) =
+        eval (normalize e
+              +ₕ normalize e₁) (x  xs)
+      ≡⟨ +Homeval (normalize e) _ (x  xs) 
+        eval (normalize e) (x  xs)
+        + eval (normalize e₁) (x  xs)
+      ≡⟨ cong  u  u + eval (normalize e₁) (x  xs))
+              (isEqualToNormalform e (x  xs)) 
+         e  (x  xs)
+        + eval (normalize e₁) (x  xs)
+      ≡⟨ cong  u   e  (x  xs) + u)
+              (isEqualToNormalform e₁ (x  xs)) 
+         e  (x  xs) +  e₁  (x  xs) 
+
+  isEqualToNormalform (e ·' e₁) [] =
+        eval (normalize e ·ₕ normalize e₁) []
+      ≡⟨ ·Homeval (normalize e) _ [] 
+        eval (normalize e) []
+        · eval (normalize e₁) []
+      ≡⟨ cong  u  u · eval (normalize e₁) [])
+              (isEqualToNormalform e []) 
+         e  []
+        · eval (normalize e₁) []
+      ≡⟨ cong  u   e  [] · u) (isEqualToNormalform e₁ []) 
+         e  [] ·  e₁  [] 
+
+  isEqualToNormalform (e ·' e₁) (x  xs) =
+        eval (normalize e ·ₕ normalize e₁) (x  xs)
+      ≡⟨ ·Homeval (normalize e) _ (x  xs) 
+        eval (normalize e) (x  xs)
+        · eval (normalize e₁) (x  xs)
+      ≡⟨ cong  u  u · eval (normalize e₁) (x  xs))
+              (isEqualToNormalform e (x  xs)) 
+         e  (x  xs)
+        · eval (normalize e₁) (x  xs)
+      ≡⟨ cong  u   e  (x  xs) · u)
+              (isEqualToNormalform e₁ (x  xs)) 
+         e  (x  xs) ·  e₁  (x  xs) 
+
+  solve :
+    {n : } (e₁ e₂ : Expr n) (xs : Vec  n)
+    (p : eval (normalize e₁) xs  eval (normalize e₂) xs)
+      e₁  xs   e₂  xs
+  solve e₁ e₂ xs p =
+     e₁  xs                ≡⟨ sym (isEqualToNormalform e₁ xs) 
+    eval (normalize e₁) xs ≡⟨ p 
+    eval (normalize e₂) xs ≡⟨ isEqualToNormalform e₂ xs 
+     e₂  xs 
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.NatSolver.html b/docs/Cubical.Tactics.NatSolver.html new file mode 100644 index 0000000..8ae748c --- /dev/null +++ b/docs/Cubical.Tactics.NatSolver.html @@ -0,0 +1,14 @@ + +Cubical.Tactics.NatSolver
{-# OPTIONS --safe #-}
+{-
+  This is inspired by/copied from:
+  https://github.com/agda/agda-stdlib/blob/master/src/Tactic/MonoidSolver.agda
+  and the 1lab.
+
+  Boilerplate code for calling the ring solver is constructed automatically
+  with agda's reflection features.
+-}
+module Cubical.Tactics.NatSolver where
+
+open import Cubical.Tactics.NatSolver.Reflection public
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.Reflection.Utilities.html b/docs/Cubical.Tactics.Reflection.Utilities.html new file mode 100644 index 0000000..9fc2fa5 --- /dev/null +++ b/docs/Cubical.Tactics.Reflection.Utilities.html @@ -0,0 +1,37 @@ + +Cubical.Tactics.Reflection.Utilities
{-# OPTIONS --safe #-}
+module Cubical.Tactics.Reflection.Utilities where
+
+open import Cubical.Foundations.Prelude hiding (Type)
+
+open import Agda.Builtin.Reflection hiding (Type)
+open import Agda.Builtin.String
+open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_)
+
+open import Cubical.Reflection.Base
+open import Cubical.Data.List
+open import Cubical.Data.Maybe
+open import Cubical.Data.FinData using () renaming (zero to fzero; suc to fsuc)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Nat using ()
+
+open import Cubical.Tactics.Reflection
+open import Cubical.Tactics.Reflection.Variables
+
+
+finiteNumberAsTerm : Maybe   Term
+finiteNumberAsTerm (just ℕ.zero) = con (quote fzero) []
+finiteNumberAsTerm (just (ℕ.suc n)) = con (quote fsuc) (finiteNumberAsTerm (just n) v∷ [])
+finiteNumberAsTerm _ = unknown
+
+-- error message helper
+errorOut : List (Arg Term)  TC (Template × Vars)
+errorOut something = typeError (strErr "Don't know what to do with "  map  {(arg _ t)  termErr t}) something)
+
+errorOut' : Term  TC (Template × Vars)
+errorOut' something = typeError (strErr "Don't know what to do with "  termErr something  [])
+
+
+_==_ = primQNameEquality
+{-# INLINE _==_ #-}
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.Reflection.Variables.html b/docs/Cubical.Tactics.Reflection.Variables.html new file mode 100644 index 0000000..a1cc374 --- /dev/null +++ b/docs/Cubical.Tactics.Reflection.Variables.html @@ -0,0 +1,83 @@ + +Cubical.Tactics.Reflection.Variables
{-# OPTIONS --safe #-}
+{-
+  This code contains some helper functions for solvers.
+  Variables in the sense of this files are things that are treated like variables by a solver.
+  A ring solver might want to treat "f x" in an equation "f x + 0 ≡ f x" like a variable "y".
+  During the inspection of the lhs and rhs of an equation, terms like "f x" are found and saved
+  and later, indices are assigned to them. These indices will be the indices of the variables
+  in the normal forms the solver uses.
+-}
+module Cubical.Tactics.Reflection.Variables where
+
+open import Cubical.Foundations.Prelude hiding (Type)
+
+open import Agda.Builtin.Reflection hiding (Type)
+open import Agda.Builtin.String
+open import Agda.Builtin.Float
+open import Agda.Builtin.Word
+open import Agda.Builtin.Char
+open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_)
+
+open import Cubical.Reflection.Base
+open import Cubical.Data.Bool
+open import Cubical.Data.List
+open import Cubical.Data.Maybe
+open import Cubical.Data.Nat using ()
+
+open import Cubical.Tactics.Reflection
+
+private
+  variable
+     : Level
+
+
+Vars = List Term
+VarAss = Term  Maybe 
+Template = VarAss  Term
+
+private
+  _=N_ = primQNameEquality
+  _=M_ = primMetaEquality
+
+  _=L_ : Literal  Literal  Bool
+  nat n =L nat m = n =ℕ m
+  word64 n =L word64 m = primWord64ToNat n =ℕ primWord64ToNat m
+  float x =L float y = primFloatEquality x y
+  char c =L char c' = primCharEquality c c'
+  string s =L string s' = primStringEquality s s'
+  name x =L name y = x =N y
+  meta x =L meta y = x =M y
+  _ =L _ = false
+
+  compareVargs : List (Arg Term)  List (Arg Term)  Bool
+
+  _=T_ : Term  Term  Bool  -- this should be a TC, since it should error out sometimes
+  var index args =T var index' args' = (index =ℕ index') and compareVargs args args'
+  con c args =T con c' args'   = (c =N c') and compareVargs args args'
+  def f args =T def f' args'   = (f =N f') and compareVargs args args'
+  lit l =T lit l'              = l =L l'
+  meta x args =T meta x' args' = (x =M x') and compareVargs args args'
+  _ =T _                       = false  -- this should be fixed
+
+compareVargs [] [] = true
+compareVargs (x v∷ l) (x' v∷ l') = (x =T x') and compareVargs l l'
+compareVargs (_ h∷ l) (_ h∷ l') = compareVargs l l' -- ignore hargs, not sure this is good
+compareVargs _ _ = false
+
+addWithoutRepetition : Term  Vars  Vars
+addWithoutRepetition t l@(t'  l') = if (t =T t') then l else t'  addWithoutRepetition t l'
+addWithoutRepetition t []      = t  []
+
+appendWithoutRepetition : Vars  Vars  Vars
+appendWithoutRepetition (x  l) l' = appendWithoutRepetition l (addWithoutRepetition x l')
+appendWithoutRepetition [] l' = l'
+
+-- this can be used to get a map from variables to numbers 0,...,n
+indexOf : Term  Vars  Maybe 
+indexOf t (t'  l) =
+  if (t =T t')
+  then just 0
+  else map-Maybe  k  ℕ.suc k) (indexOf t l)
+indexOf t [] = nothing
+
\ No newline at end of file diff --git a/docs/Cubical.Tactics.Reflection.html b/docs/Cubical.Tactics.Reflection.html new file mode 100644 index 0000000..9e15672 --- /dev/null +++ b/docs/Cubical.Tactics.Reflection.html @@ -0,0 +1,116 @@ + +Cubical.Tactics.Reflection
-- SPDX-License-Identifier: BSD-3-Clause
+{-# OPTIONS --safe #-}
+module Cubical.Tactics.Reflection where
+
+{- Utilities common to different reflection solvers.
+
+  Most of these are copied/adapted from the 1Lab
+-}
+
+open import Cubical.Foundations.Prelude
+
+open import Agda.Builtin.Reflection hiding (Type)
+
+open import Cubical.Data.Bool
+open import Cubical.Data.List
+open import Cubical.Data.Maybe
+open import Cubical.Data.Sigma
+open import Cubical.Data.Unit
+
+open import Cubical.Reflection.Base
+
+private
+  variable
+     ℓ' : Level
+
+_<$>_ :  { ℓ'} {A : Type }{B : Type ℓ'}  (A  B)  TC A  TC B
+f <$> t = t >>= λ x  returnTC (f x)
+
+_<*>_ :  { ℓ'} {A : Type }{B : Type ℓ'}  TC (A  B)  TC A  TC B
+s <*> t = s >>= λ f  t >>= λ x  returnTC (f x)
+
+wait-for-args : List (Arg Term)  TC (List (Arg Term))
+wait-for-type : Term  TC Term
+
+wait-for-type (var x args) = var x <$> wait-for-args args
+wait-for-type (con c args) = con c <$> wait-for-args args
+wait-for-type (def f args) = def f <$> wait-for-args args
+wait-for-type (lam v (abs x t)) = returnTC (lam v (abs x t))
+wait-for-type (pat-lam cs args) = returnTC (pat-lam cs args)
+wait-for-type (pi (arg i a) (abs x b)) = do
+  a  wait-for-type a
+  b  wait-for-type b
+  returnTC (pi (arg i a) (abs x b))
+wait-for-type (agda-sort s) = returnTC (agda-sort s)
+wait-for-type (lit l) = returnTC (lit l)
+wait-for-type (meta x x₁) = blockOnMeta x
+wait-for-type unknown = returnTC unknown
+
+wait-for-args [] = returnTC []
+wait-for-args (arg i a  xs) =
+  (_∷_ <$> (arg i <$> wait-for-type a)) <*> wait-for-args xs
+
+unapply-path : Term  TC (Maybe (Term × Term × Term))
+unapply-path red@(def (quote PathP) (l h∷ T v∷ x v∷ y v∷ [])) = do
+  domain  newMeta (def (quote Type) (l v∷ []))
+  ty  returnTC (def (quote Path) (domain v∷ x v∷ y v∷ []))
+  debugPrint "tactic" 50
+    (strErr "(no reduction) unapply-path: got a "  termErr red
+     strErr " but I really want it to be "  termErr ty  [])
+  unify red ty
+  returnTC (just (domain , x , y))
+unapply-path tm = reduce tm >>= λ where
+  tm@(meta _ _)  do
+    dom  newMeta (def (quote Type) [])
+    l  newMeta dom
+    r  newMeta dom
+    unify tm (def (quote Type) (dom v∷ l v∷ r v∷ []))
+    wait-for-type l
+    wait-for-type r
+    returnTC (just (dom , l , r))
+  red@(def (quote PathP) (l h∷ T v∷ x v∷ y v∷ []))  do
+    domain  newMeta (def (quote Type) (l v∷ []))
+    ty  returnTC (def (quote Path) (domain v∷ x v∷ y v∷ []))
+    debugPrint "tactic" 50
+      (strErr "unapply-path: got a "  termErr red
+       strErr " but I really want it to be "  termErr ty  [])
+    unify red ty
+    returnTC (just (domain , x , y))
+  _  returnTC nothing
+
+{-
+  get-boundary maps a term 'x ≡ y' to the pair '(x,y)'
+-}
+get-boundary : Term  TC (Maybe (Term × Term))
+get-boundary tm = unapply-path tm >>= λ where
+  (just (_ , x , y))  returnTC (just (x , y))
+  nothing             returnTC nothing
+
+equation-solver : List Name  (Term -> Term -> TC Term)  Bool  Term  TC Unit
+equation-solver don't-Reduce mk-call debug hole =
+    withNormalisation false (
+    withReduceDefs (false , don't-Reduce) (
+    do
+      -- | First we normalize the goal
+      goal  inferType hole >>= reduce
+      -- | Then we parse the goal into an AST
+      just (lhs , rhs)  get-boundary goal
+        where
+          nothing
+             typeError(strErr "The functor solver failed to parse the goal"
+                                termErr goal  [])
+      -- | Then we invoke the solver
+      -- | And we unify the result of the solver with the original hole.
+      elhs  normalise lhs
+      erhs  normalise rhs
+      call  mk-call elhs erhs
+      let unify-with-goal = (unify hole call)
+      noConstraints (
+        if debug
+        then unify-with-goal
+        else (
+        unify-with-goal <|>
+        typeError ((strErr "Could not equate the following expressions:\n  "
+                   termErr elhs  strErr "\nAnd\n  "  termErr erhs  []))))))
+
\ No newline at end of file diff --git a/docs/Realizability.ApplicativeStructure.html b/docs/Realizability.ApplicativeStructure.html index ef9bf67..8499598 100644 --- a/docs/Realizability.ApplicativeStructure.html +++ b/docs/Realizability.ApplicativeStructure.html @@ -1,145 +1,173 @@ -Realizability.ApplicativeStructure
{-# OPTIONS --cubical --without-K --allow-unsolved-metas #-}
-open import Cubical.Core.Everything
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Relation.Nullary
-open import Cubical.Data.Nat
-open import Cubical.Data.Nat.Order
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Empty renaming (elim to ⊥elim)
-
-module Realizability.ApplicativeStructure where
-
-record ApplicativeStructure {} (A : Type ) : Type  where
-  infixl 20 _⨾_
-  field
-    isSetA : isSet A
-    _⨾_ : A  A  A
-
-module _ {} {A : Type } (as : ApplicativeStructure A) where
-  open ApplicativeStructure as
-  infix 23 `_
-  infixl 22 _̇_
-  data Term (n : ) : Type  where
-    # : Fin n  Term n
-    `_ : A  Term n
-    _̇_ : Term n  Term n  Term n
-
-  upgrade :  {n m}  n < m  Term n  Term m
-  upgrade _ (` a) = ` a
-  upgrade {n} {m} n<m (# k) = # (k .fst , <-trans (k .snd) n<m)
-  upgrade {n} {m} n<m (a ̇ b) = upgrade n<m a ̇ upgrade n<m b
-
-  substitute :  {n}  Term n  Vec A n  A
-  substitute (` a) _ = a
-  substitute {n} (# k) subs = lookup (Fin→FinData n k) subs
-  substitute (a ̇ b) subs = (substitute a subs)  (substitute b subs)
-
-  apply :  {n}  A  Vec A n  A
-  apply a [] = a
-  apply a (x  xs) = apply' (x  xs) a where
-                            apply' :  {n}  Vec A n  A  A
-                            apply' [] acc = acc
-                            apply' (x  xs) acc = apply' xs (acc  x)
-
-  applyWorks :  K a b  apply K (a  b  [])  K  a  b
-  applyWorks K a b = refl
-
-  record isInterpreted {n} (t : Term n) : Type  where
-    field
-      interpretation : A
-      naturality :  (subs : Vec A n)  apply interpretation subs  substitute t subs
-
-  isCombinatoriallyComplete : Type 
-  isCombinatoriallyComplete =  {n} (t : Term n)  isInterpreted t
-
-  record Feferman : Type  where
-    field
-      s : A
-      k : A
-      kab≡a :  a b  k  a  b  a
-      sabc≡ac_bc :  a b c  s  a  b  c  (a  c)  (b  c)
-
-  module _ (completeness : isCombinatoriallyComplete) where
-    open isInterpreted
-
-    preS : Term 3
-    preS = ((# 0) ̇ (# 2)) ̇ ((# 1) ̇ (# 2))
-
-    S : A
-    S = (completeness preS) .interpretation
-
-    preK : Term 2
-    preK = # 0
-
-    K : A
-    K = (completeness preK) .interpretation
-
-    Kab≡a :  a b  K  a  b  a
-    Kab≡a a b = (completeness preK) .naturality (a  b  [])
-
-    Sabc≡ac_bc :  a b c  S  a  b  c  (a  c)  (b  c)
-    Sabc≡ac_bc a b c = (completeness preS) .naturality (a  b  c  [])
-    open Feferman
-    completeness→feferman : Feferman
-    completeness→feferman .s = S
-    completeness→feferman .k = K
-    completeness→feferman .kab≡a = Kab≡a
-    completeness→feferman .sabc≡ac_bc = Sabc≡ac_bc
-
-  module _ (feferman : Feferman) where
-    open Feferman feferman
-    {-
-    This goofy definition is there to ensure that λ* computes.
-    For some reason the last branch of pattern-matching cannot definitionally equate y .fst and suc m
-    So we must postulate it.
-    But since we already know that y .fst = suc m we can use discreteℕ to get an actual proof and extract
-    it using fromYes. fromYes then gets a dummy proof
-    -}
-    λ* :  {n} (e : Term (suc n))  Term n
-    λ* (` a) = ` k ̇ ` a
-    λ* (a ̇ b) = ` s ̇ (λ* a) ̇ (λ* b)
-    λ* {n} (# y) with (discreteℕ (y .fst) zero)
-    ... | yes _ = ` s ̇ ` k ̇ ` k
-    ... | no ¬y≡zero with (y .fst)
-    ...     | zero = ⊥elim (¬y≡zero refl)
-    ...     | (suc m) = # (m , pred-≤-pred (subst  y'  suc y'  suc n) (fromYes fsty≡sucm (discreteℕ (y .fst) (suc m))) (y .snd))) where postulate fsty≡sucm : fst y  suc m
-
-    λ*-chainTerm :  n  Term n  Term zero
-    λ*-chainTerm zero t = t
-    λ*-chainTerm (suc n) t = λ*-chainTerm n (λ* t)
-
-    λ*-chain :  {n}  Term n  A
-    λ*-chain {n} t = substitute (λ*-chainTerm n t) []
-
-    ⟦_⟧ : Term zero  A
-     ` a  = a
-     a ̇ b  =  a    b 
-     # x  = ⊥elim {A = λ _  A} (¬Fin0 x)
-
-    λ*Computation :  (T : Term 1) (e : Term zero)   λ* T    e   substitute T ( e   [])
-    λ*Computation (# x) e = {!subst (λ x≡zero → ⟦ λ* (# x) ⟧ ⨾ ⟦ e ⟧ ≡ lookup (Fin→FinData 1 x) (⟦ e ⟧ ∷ [])) ? ?!}
-    λ*Computation (` x) e = kab≡a x  e 
-    λ*Computation (U ̇ V) e =
-      s   λ* U    λ* V    e 
-        ≡⟨ sabc≡ac_bc _ _ _ 
-       ( λ* U    e )  ( λ* V    e )
-        ≡⟨ cong  x  x  _) (λ*Computation U e) 
-        (substitute U ( e   []))  ( λ* V    e )
-        ≡⟨ cong  x  _  x) (λ*Computation V e) 
-        (substitute U ( e   []))  (substitute V ( e   []))
-        
-    
-
-    open isInterpreted
-
-    postulate λ*-naturality :  {n} (t : Term n) (subs : Vec A n)  apply (λ*-chain t) subs  substitute t subs
-    
-    feferman→completeness : isCombinatoriallyComplete
-    feferman→completeness t .interpretation = λ*-chain t
-    feferman→completeness t .naturality subs = λ*-naturality t subs
-    
-
+Realizability.ApplicativeStructure
open import Cubical.Core.Everything
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Relation.Nullary
+open import Cubical.Data.Nat
+open import Cubical.Data.Nat.Order
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Empty renaming (elim to ⊥elim)
+open import Cubical.Tactics.NatSolver
+
+module Realizability.ApplicativeStructure where
+
+private module _ {} {A : Type } where
+  -- Taken from Data.Vec.Base from agda-stdlib
+  foldlOp :  {ℓ'} (B :   Type ℓ')  Type _
+  foldlOp B =  {n}  B n  A  B (suc n)
+
+  opaque
+    foldl :  {ℓ'} {n : } (B :   Type ℓ')  foldlOp B  B zero  Vec A n  B n
+    foldl {ℓ'} {.zero} B op acc emptyVec = acc
+    foldl {ℓ'} {.(suc _)} B op acc (x ∷vec vec) = foldl  n  B (suc n)) op (op acc x) vec
+
+  opaque
+    reverse :  {n}  Vec A n  Vec A n
+    reverse vec = foldl  n  Vec A n)  acc curr  curr  acc) [] vec
+
+  opaque
+    chain :  {n}  Vec A (suc n)  (A  A  A)  A
+    chain {n} (x ∷vec vec) op = foldl  _  A)  acc curr  op acc curr) x vec
+
+record ApplicativeStructure {} (A : Type ) : Type  where
+  infixl 20 _⨾_
+  field
+    isSetA : isSet A
+    _⨾_ : A  A  A
+
+module _ {} {A : Type } (as : ApplicativeStructure A) where
+  open ApplicativeStructure as
+  infix 23 `_
+  infixl 22 _̇_
+  data Term (n : ) : Type  where
+    # : Fin n  Term n
+    `_ : A  Term n
+    _̇_ : Term n  Term n  Term n
+
+  ⟦_⟧ :  {n}  Term n  Vec A n  A
+  ⟦_⟧ (` a) _ = a
+  ⟦_⟧ {n} (# k) subs = lookup k subs
+  ⟦_⟧ (a ̇ b) subs = ( a  subs)  ( b  subs)
+
+  applicationChain :  {n m}  Vec (Term m) (suc n)  Term m
+  applicationChain {n} {m} vec = chain vec  x y  x ̇ y)
+
+  apply :  {n}  A  Vec A n  A
+  apply {n} a vec = chain (a  vec)  x y  x  y)
+  
+  private
+    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+      applyWorks :  K a b  apply K (a  b  [])  K  a  b
+      applyWorks K a b = refl
+
+  record Feferman : Type  where
+    field
+      s : A
+      k : A
+      kab≡a :  a b  k  a  b  a
+      sabc≡ac_bc :  a b c  s  a  b  c  (a  c)  (b  c)
+      
+  module ComputationRules (feferman : Feferman) where
+    open Feferman feferman
+
+    opaque
+      λ*abst :  {n}  (e : Term (suc n))  Term n
+      λ*abst {n} (# zero) = ` s ̇ ` k ̇ ` k
+      λ*abst {n} (# (suc x)) = ` k ̇ # x
+      λ*abst {n} (` x) = ` k ̇ ` x
+      λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁
+
+    -- Some shortcuts
+
+    λ* : Term one  A
+    λ* t =  λ*abst t  []
+
+    λ*2 : Term two  A
+    λ*2 t =  λ*abst (λ*abst t)  []
+
+    λ*3 : Term three  A
+    λ*3 t =  λ*abst (λ*abst (λ*abst t))  []
+
+    λ*4 : Term four  A
+    λ*4 t =  λ*abst (λ*abst (λ*abst (λ*abst t)))  []
+
+    opaque
+      unfolding λ*abst
+      βreduction :  {n}  (body : Term (suc n))  (prim : A)  (subs : Vec A n)   λ*abst body ̇ ` prim  subs   body  (prim  subs)
+      βreduction {n} (# zero) prim subs =
+        s  k  k  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+        k  prim  (k  prim)
+          ≡⟨ kab≡a _ _ 
+        prim
+          
+      βreduction {n} (# (suc x)) prim subs = kab≡a _ _
+      βreduction {n} (` x) prim subs = kab≡a _ _
+      βreduction {n} (rator ̇ rand) prim subs =
+        s   λ*abst rator  subs   λ*abst rand  subs  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+         λ*abst rator  subs  prim  ( λ*abst rand  subs  prim)
+          ≡⟨ cong₂  x y  x  y) (βreduction rator prim subs) (βreduction rand prim subs) 
+         rator  (prim  subs)   rand  (prim  subs)
+          
+
+    λ*chainTerm :  n  Term n  Term zero
+    λ*chainTerm zero t = t
+    λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t)
+
+    λ*chain :  {n}  Term n  A
+    λ*chain {n} t =  λ*chainTerm n t  []
+
+    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+
+      λ*ComputationRule :  (t : Term 1) (a : A)  λ* t  a   t  (a  [])
+      λ*ComputationRule t a =
+         λ*abst t  []  a
+          ≡⟨ βreduction t a [] 
+         t  (a  [])
+          
+
+      λ*2ComputationRule :  (t : Term 2) (a b : A)  λ*2 t  a  b   t  (b  a  [])
+      λ*2ComputationRule t a b =
+         λ*abst (λ*abst t)  []  a  b
+          ≡⟨ refl 
+         λ*abst (λ*abst t)  []   ` a  []   ` b  []
+          ≡⟨ refl 
+         λ*abst (λ*abst t) ̇ ` a  []   ` b  []
+          ≡⟨ cong  x  x  b) (βreduction (λ*abst t) a []) 
+         λ*abst t  (a  [])  b
+          ≡⟨ βreduction t b (a  []) 
+         t  (b  a  [])
+          
+          
+      λ*3ComputationRule :  (t : Term 3) (a b c : A)  λ*3 t  a  b  c   t  (c  b  a  [])
+      λ*3ComputationRule t a b c =
+         λ*abst (λ*abst (λ*abst t))  []   ` a  []   ` b  []   ` c  []
+          ≡⟨ cong  x  x  b  c) (βreduction (λ*abst (λ*abst t)) a []) 
+         λ*abst (λ*abst t)  (a  [])   ` b  (a  [])   ` c  (a  [])
+          ≡⟨ cong  x  x  c) (βreduction (λ*abst t) b (a  [])) 
+         λ*abst t  (b  a  [])   ` c  (b  a  [])
+          ≡⟨ βreduction t c (b  a  []) 
+         t  (c  b  a  [])
+          
+
+      λ*4ComputationRule :  (t : Term 4) (a b c d : A)  λ*4 t  a  b  c  d   t  (d  c  b  a  [])
+      λ*4ComputationRule t a b c d =
+         λ*abst (λ*abst (λ*abst (λ*abst t)))  []   ` a  []   ` b  []   ` c  []   ` d  []
+          ≡⟨ cong  x  x  b  c  d) (βreduction (λ*abst (λ*abst (λ*abst t))) a []) 
+         λ*abst (λ*abst (λ*abst t))  (a  [])   ` b  (a  [])   ` c  (a  [])   ` d  (a  [])
+          ≡⟨ cong  x  x  c  d) (βreduction (λ*abst (λ*abst t)) b (a  [])) 
+         λ*abst (λ*abst t)  (b  a  [])   ` c  (b  a  [])   ` d  (b  a  [])
+          ≡⟨ cong  x  x  d) (βreduction (λ*abst t) c (b  a  [])) 
+         λ*abst t  (c  b  a  [])   ` d  (c  b  a  [])
+          ≡⟨ βreduction t d (c  b  a  []) 
+         t  (d  c  b  a  [])
+          
 
\ No newline at end of file diff --git a/docs/Realizability.Choice.html b/docs/Realizability.Choice.html index 8d7739a..1980acf 100644 --- a/docs/Realizability.Choice.html +++ b/docs/Realizability.Choice.html @@ -10,5 +10,5 @@ module Realizability.Choice where Choice : ℓ' Type (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) -Choice ℓ' = (A : Type ) (B : Type ℓ') isSet A isSet B (f : A B) isSurjection f ∃[ f' (B A) ] section f f' +Choice ℓ' = (A : Type ) (B : Type ℓ') isSet A isSet B (f : A B) isSurjection f ∃[ f' (B A) ] section f f'
\ No newline at end of file diff --git a/docs/Realizability.CombinatoryAlgebra.html b/docs/Realizability.CombinatoryAlgebra.html index a8dc055..1cb4a61 100644 --- a/docs/Realizability.CombinatoryAlgebra.html +++ b/docs/Realizability.CombinatoryAlgebra.html @@ -6,192 +6,195 @@ open import Cubical.Data.Unit open import Cubical.Data.Nat -open import Realizability.ApplicativeStructure hiding (S;K) +open import Realizability.ApplicativeStructure -module Realizability.CombinatoryAlgebra where +module Realizability.CombinatoryAlgebra where -record CombinatoryAlgebra {} (A : Type ) : Type where - field - as : ApplicativeStructure A - completeness : isCombinatoriallyComplete as - fefermanStructure : Feferman as - fefermanStructure = completeness→feferman as completeness - open Feferman fefermanStructure public - open ApplicativeStructure as public +record CombinatoryAlgebra {} (A : Type ) : Type where + field + as : ApplicativeStructure A + fefermanStructure : Feferman as + open Feferman fefermanStructure public + open ApplicativeStructure as public + open ComputationRules as fefermanStructure public -module Combinators {} {A : Type } (ca : CombinatoryAlgebra A) where - open CombinatoryAlgebra ca +module Combinators {} {A : Type } (ca : CombinatoryAlgebra A) where + open CombinatoryAlgebra ca - i : A - i = s k k - - k' : A - k' = k i - - ia≡a : a i a a - ia≡a a = (cong x x a) refl) (sabc≡ac_bc k k a) (kab≡a a (k a)) - - k'ab≡b : a b k' a b b - k'ab≡b a b = k' a b - ≡⟨ refl - (k i a b) - ≡⟨ cong x x b) (kab≡a i a) - (i b) - ≡⟨ ia≡a b - b - - - true : A - true = k - - false : A - false = k' - - if_then_else_ : c t e A - if c then t else e = i c t e - - ifTrueThen : t e if true then t else e t - ifTrueThen t e = if true then t else e - ≡⟨ refl - i true t e - ≡⟨ cong x i x t e) refl - i k t e - ≡⟨ cong x x t e) (ia≡a k) - k t e - ≡⟨ kab≡a t e - t - - - ifFalseElse : t e if false then t else e e - ifFalseElse t e = if false then t else e - ≡⟨ refl - i false t e - ≡⟨ cong x i x t e) refl - i k' t e - ≡⟨ cong x x t e) (ia≡a k') - k' t e - ≡⟨ k'ab≡b t e - e - - - -- I used a Scheme script to generate this - pair : A - pair = s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) - (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) - (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) - (s (k (k)) (k (k))))) (s (k (k)) (k (k)))))) (s - (s (k (s)) (s (k (k)) (k (k)))) (s (k (k)) (s (k) (k))))))) - (s (s (k (s)) (s (k (k)) (k (k)))) (s (s (k (s)) (k (k))) (k (k)))) - - pr₁ : A - pr₁ = s (s k k) (k k) - - pr₂ : A - pr₂ = s (s k k) (k k') - - -- TODO : Prove computation rules - postulate pr₁pxy≡x : x y pr₁ (pair x y) x - postulate pr₂pxy≡y : x y pr₂ (pair x y) y - - -- Curry numbers - ℕ→curry : A - ℕ→curry zero = i - ℕ→curry (suc n) = pair k' (ℕ→curry n) - - Z : A - Z = pr₁ - - Zzero≡true : Z (ℕ→curry zero) true - Zzero≡true = Z (ℕ→curry zero) - ≡⟨ cong x Z x) refl - Z i - ≡⟨ cong x x i) refl - s (s k k) (k k) i - ≡⟨ sabc≡ac_bc (s k k) (k k) i - ((s k k) i) (k k i) - ≡⟨ cong x (x i) (k k i)) refl - (i i) (k k i) - ≡⟨ cong x x (k k i)) (ia≡a i) - i (k k i) - ≡⟨ cong x i x) (kab≡a k i) - i k - ≡⟨ ia≡a k - k - - - Zsuc≡false : n Z (ℕ→curry (suc n)) false - Zsuc≡false n = Z (ℕ→curry (suc n)) - ≡⟨ cong x Z x) refl - Z (pair k' (ℕ→curry n)) - ≡⟨ cong x x (pair k' (ℕ→curry n))) refl - pr₁ (pair k' (ℕ→curry n)) - ≡⟨ pr₁pxy≡x k' (ℕ→curry n) - false - - - S : A - S = pair k' - - Sn≡sucn : n S (ℕ→curry n) ℕ→curry (suc n) - Sn≡sucn n = S (ℕ→curry n) - ≡⟨ cong x x (ℕ→curry n)) refl - pair k' (ℕ→curry n) - - - P : A - P = s (s (s (k pr₁) i) (k (ℕ→curry zero))) (s (k (pr₂)) i) - - postulate Pzero≡zero : P (ℕ→curry zero) ℕ→curry zero - postulate Psucn≡n : n P (ℕ→curry (suc n)) ℕ→curry n - - B : g f A - B g f = s (k g) (s (k f) i) - - Ba≡gfa : g f a B g f a g (f a) - Ba≡gfa g f a = - s (k g) (s (k f) i) a - ≡⟨ sabc≡ac_bc (k g) (s (k f) i) a - (k g a) (s (k f) i a) - ≡⟨ cong x x (s (k f) i a)) (kab≡a g a) - g (s (k f) i a) - ≡⟨ cong x g x) (sabc≡ac_bc (k f) i a) - g ((k f a) (i a)) - ≡⟨ cong x g (x (i a))) (kab≡a f a) - g (f (i a)) - ≡⟨ cong x g (f x)) (ia≡a a) - g (f a) - - - -module Trivial {} {A : Type } (ca : CombinatoryAlgebra A) where - open CombinatoryAlgebra ca - open Combinators ca - module _ (isNonTrivial : s k ) where - - k≠k' : k k' - k≠k' k≡k' = isNonTrivial s≡k where - cond = if true then s else k - cond' = if false then s else k - condEq : cond cond' - condEq = cong x if x then s else k) k≡k' + i : A + i = s k k + + k' : A + k' = k i + + ia≡a : a i a a + ia≡a a = (cong x x a) refl) (sabc≡ac_bc k k a) (kab≡a a (k a)) + + k'ab≡b : a b k' a b b + k'ab≡b a b = k' a b + ≡⟨ refl + (k i a b) + ≡⟨ cong x x b) (kab≡a i a) + (i b) + ≡⟨ ia≡a b + b + + + true : A + true = k + + false : A + false = k' + + if_then_else_ : c t e A + if c then t else e = i c t e + + ifTrueThen : t e if true then t else e t + ifTrueThen t e = if true then t else e + ≡⟨ refl + i true t e + ≡⟨ cong x i x t e) refl + i k t e + ≡⟨ cong x x t e) (ia≡a k) + k t e + ≡⟨ kab≡a t e + t + + + ifFalseElse : t e if false then t else e e + ifFalseElse t e = if false then t else e + ≡⟨ refl + i false t e + ≡⟨ cong x i x t e) refl + i k' t e + ≡⟨ cong x x t e) (ia≡a k') + k' t e + ≡⟨ k'ab≡b t e + e + + + -- I used a Scheme script to generate this + opaque + pair : A + pair = s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) + (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) + (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) + (s (k (k)) (k (k))))) (s (k (k)) (k (k)))))) (s + (s (k (s)) (s (k (k)) (k (k)))) (s (k (k)) (s (k) (k))))))) + (s (s (k (s)) (s (k (k)) (k (k)))) (s (s (k (s)) (k (k))) (k (k)))) + + opaque + pr₁ : A + pr₁ = s (s k k) (k k) + + pr₂ : A + pr₂ = s (s k k) (k k') + + -- TODO : Prove computation rules + postulate pr₁pxy≡x : x y pr₁ (pair x y) x + postulate pr₂pxy≡y : x y pr₂ (pair x y) y + + -- Curry numbers + ℕ→curry : A + ℕ→curry zero = i + ℕ→curry (suc n) = pair k' (ℕ→curry n) + + Z : A + Z = pr₁ + + opaque + unfolding pr₁ + Zzero≡true : Z (ℕ→curry zero) true + Zzero≡true = Z (ℕ→curry zero) + ≡⟨ cong x Z x) refl + Z i + ≡⟨ cong x x i) refl + s (s k k) (k k) i + ≡⟨ sabc≡ac_bc (s k k) (k k) i + ((s k k) i) (k k i) + ≡⟨ cong x (x i) (k k i)) refl + (i i) (k k i) + ≡⟨ cong x x (k k i)) (ia≡a i) + i (k k i) + ≡⟨ cong x i x) (kab≡a k i) + i k + ≡⟨ ia≡a k + k + + + Zsuc≡false : n Z (ℕ→curry (suc n)) false + Zsuc≡false n = Z (ℕ→curry (suc n)) + ≡⟨ cong x Z x) refl + Z (pair k' (ℕ→curry n)) + ≡⟨ cong x x (pair k' (ℕ→curry n))) refl + pr₁ (pair k' (ℕ→curry n)) + ≡⟨ pr₁pxy≡x k' (ℕ→curry n) + false + + + S : A + S = pair k' + + Sn≡sucn : n S (ℕ→curry n) ℕ→curry (suc n) + Sn≡sucn n = S (ℕ→curry n) + ≡⟨ cong x x (ℕ→curry n)) refl + pair k' (ℕ→curry n) + + + P : A + P = s (s (s (k pr₁) i) (k (ℕ→curry zero))) (s (k (pr₂)) i) + + postulate Pzero≡zero : P (ℕ→curry zero) ℕ→curry zero + postulate Psucn≡n : n P (ℕ→curry (suc n)) ℕ→curry n + + B : g f A + B g f = s (k g) (s (k f) i) + + Ba≡gfa : g f a B g f a g (f a) + Ba≡gfa g f a = + s (k g) (s (k f) i) a + ≡⟨ sabc≡ac_bc (k g) (s (k f) i) a + (k g a) (s (k f) i a) + ≡⟨ cong x x (s (k f) i a)) (kab≡a g a) + g (s (k f) i a) + ≡⟨ cong x g x) (sabc≡ac_bc (k f) i a) + g ((k f a) (i a)) + ≡⟨ cong x g (x (i a))) (kab≡a f a) + g (f (i a)) + ≡⟨ cong x g (f x)) (ia≡a a) + g (f a) + + + +module Trivial {} {A : Type } (ca : CombinatoryAlgebra A) where + open CombinatoryAlgebra ca + open Combinators ca + module _ (isNonTrivial : s k ) where + + k≠k' : k k' + k≠k' k≡k' = isNonTrivial s≡k where + cond = if true then s else k + cond' = if false then s else k + condEq : cond cond' + condEq = cong x if x then s else k) k≡k' - cond≡s : cond s - cond≡s = ifTrueThen _ _ - - cond'≡k : cond' k - cond'≡k = ifFalseElse _ _ - - cond≡k : cond k - cond≡k = subst x x k) (sym condEq) cond'≡k - - s≡k : s k - s≡k = - s - ≡⟨ sym cond≡s - cond - ≡⟨ cond≡k - k - + cond≡s : cond s + cond≡s = ifTrueThen _ _ + + cond'≡k : cond' k + cond'≡k = ifFalseElse _ _ + + cond≡k : cond k + cond≡k = subst x x k) (sym condEq) cond'≡k + + s≡k : s k + s≡k = + s + ≡⟨ sym cond≡s + cond + ≡⟨ cond≡k + k + \ No newline at end of file diff --git a/docs/Realizability.PropResizing.html b/docs/Realizability.PropResizing.html new file mode 100644 index 0000000..5c3006b --- /dev/null +++ b/docs/Realizability.PropResizing.html @@ -0,0 +1,25 @@ + +Realizability.PropResizing
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Structure
+open import Cubical.Data.Sigma
+
+module Realizability.PropResizing where
+
+-- Formulation of propositional resizing inspired by the corresponding formulation
+-- in TypeTopology
+-- https://www.cs.bham.ac.uk/~mhe/TypeTopology/UF.Size.html
+
+copyOf :  {}  Type   (ℓ' : Level)  Type _
+copyOf {} X ℓ' = Σ[ copy  Type ℓ' ] X  copy
+
+copy = fst
+copyEquiv = snd
+
+-- We need the principle that TypeTopology calls Ω resizing
+-- that the universe of props in a universe 𝓤 has a copy in 𝓤
+-- This we call hPropResizing
+hPropResizing :    Type _
+hPropResizing  = copyOf (hProp ) 
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.BinProducts.html b/docs/Realizability.Topos.BinProducts.html new file mode 100644 index 0000000..aa2db2e --- /dev/null +++ b/docs/Realizability.Topos.BinProducts.html @@ -0,0 +1,602 @@ + +Realizability.Topos.BinProducts
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Unit
+open import Cubical.Data.Empty
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Limits.BinProduct
+
+module Realizability.Topos.BinProducts
+  { ℓ' ℓ''} {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
+
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+open FunctionalRelation
+open PartialEquivalenceRelation hiding (isSetX)
+module _
+  {X : Type ℓ'}
+  {Y : Type ℓ'}
+  (perX : PartialEquivalenceRelation X)
+  (perY : PartialEquivalenceRelation Y) where
+
+  opaque private
+    isSetX : isSet X
+    isSetX = PartialEquivalenceRelation.isSetX perX
+    isSetY : isSet Y
+    isSetY = PartialEquivalenceRelation.isSetX perY
+
+  opaque
+    {-# TERMINATING #-}
+    binProdObRT : PartialEquivalenceRelation (X × Y)
+    Predicate.isSetX (PartialEquivalenceRelation.equality binProdObRT) =
+      isSet× (isSet× isSetX isSetY) (isSet× isSetX isSetY)
+    Predicate.∣ PartialEquivalenceRelation.equality binProdObRT  ((x , y) , x' , y') r =
+      (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   perY .equality  (y , y')
+    Predicate.isPropValued (PartialEquivalenceRelation.equality binProdObRT) ((x , y) , x' , y') r =
+      isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _)
+    isPartialEquivalenceRelation.isSetX (PartialEquivalenceRelation.isPerEquality binProdObRT) = isSet× isSetX isSetY
+    isPartialEquivalenceRelation.isSymmetric (PartialEquivalenceRelation.isPerEquality binProdObRT) =
+      do
+        (sX , sX⊩isSymmetricX)  perX .isSymmetric
+        (sY , sY⊩isSymmetricY)  perY .isSymmetric
+        let
+          prover : ApplStrTerm as 1
+          prover = ` pair ̇ (` sX ̇ (` pr₁ ̇ # zero)) ̇ (` sY ̇ (` pr₂ ̇ # zero))
+        return
+          (λ* prover ,
+           { (x , y) (x' , y') r (pr₁r⊩x~x' , pr₂r⊩y~y') 
+            subst
+               r'  r'   perX .equality  (x' , x))
+              (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+              (sX⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x') ,
+            subst
+               r'  r'   perY .equality  (y' , y))
+              (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+              (sY⊩isSymmetricY y y' (pr₂  r) pr₂r⊩y~y') }))
+    isPartialEquivalenceRelation.isTransitive (PartialEquivalenceRelation.isPerEquality binProdObRT) =
+      do
+        (tX , tX⊩isTransitiveX)  perX .isTransitive
+        (tY , tY⊩isTransitiveY)  perY .isTransitive
+        let
+          prover : ApplStrTerm as 2
+          prover = ` pair ̇ (` tX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` tY ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
+        return
+          (λ*2 prover ,
+           { (x , y) (x' , y') (x'' , y'') a b (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x'~x'' , pr₂b⊩y'~y'') 
+            subst
+               r'  r'   perX .equality  (x , x''))
+              (sym (cong  x  pr₁  x) (λ*2ComputationRule prover a b)  pr₁pxy≡x _ _))
+              (tX⊩isTransitiveX x x' x'' (pr₁  a) (pr₁  b) pr₁a⊩x~x' pr₁b⊩x'~x'') ,
+            subst
+               r'  r'   perY .equality  (y , y''))
+              (sym (cong  x  pr₂  x) (λ*2ComputationRule prover a b)  pr₂pxy≡y _ _))
+              (tY⊩isTransitiveY y y' y'' (pr₂  a) (pr₂  b) pr₂a⊩y~y' pr₂b⊩y'~y'') }))
+
+  opaque
+    unfolding binProdObRT
+    unfolding idFuncRel
+    binProdPr₁FuncRel : FunctionalRelation binProdObRT perX
+    FunctionalRelation.relation binProdPr₁FuncRel =
+      record
+        { isSetX = isSet× (isSet× isSetX isSetY) isSetX
+        ; ∣_∣ = λ { ((x , y) , x') r  (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   perY .equality  (y , y) }
+        ; isPropValued =  { ((x , y) , x') r  isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) }) }
+    FunctionalRelation.isFuncRel binProdPr₁FuncRel =
+      record
+       { isStrictDomain =
+         do
+           (stD , stD⊩isStrictDomainEqX)  idFuncRel perX .isStrictDomain
+           let
+             prover : ApplStrTerm as 1
+             prover = ` pair ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (# zero))
+           return
+             (λ* prover ,
+              { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) 
+               subst
+                  r'  r'   perX .equality  (x , x))
+                 (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                 (stD⊩isStrictDomainEqX x x' (pr₁  r) pr₁r⊩x~x') ,
+               subst
+                  r'  r'   perY .equality  (y , y))
+                 (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                 pr₂r⊩y~y }))
+       ; isStrictCodomain =
+         do
+           (stC , stC⊩isStrictCodomainEqX)  idFuncRel perX .isStrictCodomain
+           let
+             prover : ApplStrTerm as 1
+             prover = ` stC ̇ (` pr₁ ̇ # zero)
+           return
+             (λ* prover ,
+              λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) 
+                subst
+                   r'  r'   perX .equality  (x' , x'))
+                  (sym (λ*ComputationRule prover r))
+                  (stC⊩isStrictCodomainEqX x x' (pr₁  r) pr₁r⊩x~x') })
+       ; isRelational =
+         do
+           (stC , stC⊩isStrictCodomainEqY)  idFuncRel perY .isStrictCodomain
+           (t , t⊩isTransitiveX)  perX .isTransitive
+           (s , s⊩isSymmetricX)  perX .isSymmetric
+           let
+             prover : ApplStrTerm as 3
+             prover = ` pair ̇ (` t ̇ (` s ̇ (` pr₁ ̇ # two)) ̇ (` t ̇ (` pr₁ ̇ # one) ̇ # zero)) ̇ (` stC ̇ (` pr₂ ̇ # two))
+           return
+             (λ*3 prover ,
+              ((λ { (x , y) (x' , y') x'' x''' a b c (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x~x'' , pr₂b⊩y~y) c⊩x''~x''' 
+                subst
+                   r'  r'   perX .equality  (x' , x'''))
+                  (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
+                  (t⊩isTransitiveX
+                    x' x x'''
+                    (s  (pr₁  a)) (t  (pr₁  b)  c)
+                    (s⊩isSymmetricX x x' (pr₁  a) pr₁a⊩x~x')
+                    (t⊩isTransitiveX x x'' x''' (pr₁  b) c pr₁b⊩x~x'' c⊩x''~x''')) ,
+                subst
+                   r'  r'   perY .equality  (y' , y'))
+                  (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _))
+                  (stC⊩isStrictCodomainEqY y y' (pr₂  a) pr₂a⊩y~y') })))
+       ; isSingleValued =
+         do
+           (t , t⊩isTransitive)  perX .isTransitive
+           (s , s⊩isSymmetric)  perX .isSymmetric
+           let
+             prover : ApplStrTerm as 2
+             prover = ` t ̇ (` s ̇ (` pr₁ ̇ # one)) ̇ (` pr₁ ̇ # zero)
+           return
+             (λ*2 prover ,
+               { (x , y) x' x'' r₁ r₂ (pr₁r₁⊩x~x' , pr₂r₁⊩y~y) (pr₁r₂⊩x~x'' , pr₂r₂⊩y~y) 
+                subst
+                   r'  r'   perX .equality  (x' , x''))
+                  (sym (λ*2ComputationRule prover r₁ r₂))
+                  (t⊩isTransitive x' x x'' (s  (pr₁  r₁)) (pr₁  r₂) (s⊩isSymmetric x x' (pr₁  r₁) pr₁r₁⊩x~x') pr₁r₂⊩x~x'')}))
+       ; isTotal =
+         do
+           return
+             (Id ,
+               { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) 
+                return
+                  (x ,
+                  ((subst  r'  r'   perX .equality  (x , x)) (cong  x  pr₁  x) (sym (Ida≡a _))) pr₁r⊩x~x) ,
+                   (subst  r'  r'   perY .equality  (y , y)) (cong  x  pr₂  x) (sym (Ida≡a _))) pr₂r⊩y~y))) }))
+       }
+
+  opaque
+    binProdPr₁RT : RTMorphism binProdObRT perX
+    binProdPr₁RT = [ binProdPr₁FuncRel ]
+
+  -- Code repetition to a degree
+  -- TODO : Refactor into a proper abstraction
+  opaque
+    unfolding binProdObRT
+    unfolding idFuncRel
+    binProdPr₂FuncRel : FunctionalRelation binProdObRT perY
+    FunctionalRelation.relation binProdPr₂FuncRel =
+      record
+        { isSetX = isSet× (isSet× isSetX isSetY) isSetY
+        ; ∣_∣ = λ { ((x , y) , y') r  (pr₁  r)   perY .equality  (y , y') × (pr₂  r)   perX .equality  (x , x) }
+        ; isPropValued = λ { ((x , y) , y') r  isProp× (perY .equality .isPropValued _ _) (perX .equality .isPropValued _ _) } }
+    FunctionalRelation.isFuncRel binProdPr₂FuncRel =
+      record
+       { isStrictDomain =
+         do
+           (stD , stD⊩isStrictDomainEqY)  idFuncRel perY .isStrictDomain
+           let
+             prover : ApplStrTerm as 1
+             prover = ` pair ̇ (` pr₂ ̇ (# zero)) ̇ (` stD ̇ (` pr₁ ̇ # zero))
+           return
+             (λ* prover ,
+              { (x , y) y' r (pr₁r⊩y~y' , pr₂r⊩x~x) 
+                subst
+                   r'  r'   perX .equality  (x , x))
+                  (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                  pr₂r⊩x~x ,
+                subst
+                   r'  r'   perY .equality  (y , y))
+                  (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                  (stD⊩isStrictDomainEqY y y' (pr₁  r) pr₁r⊩y~y') }))
+       ; isStrictCodomain =
+         do
+           (stC , stC⊩isStrictCodomainEqY)  idFuncRel perY .isStrictCodomain
+           let
+             prover : ApplStrTerm as 1
+             prover = ` stC ̇ (` pr₁ ̇ # zero)
+           return
+             (λ* prover ,
+              { (x , y) y' r (pr₁r⊩y~y' , pr₂r⊩x~x) 
+               subst
+                  r'  r'   perY .equality  (y' , y'))
+                 (sym (λ*ComputationRule prover r))
+                 (stC⊩isStrictCodomainEqY y y' (pr₁  r) pr₁r⊩y~y') }))
+       ; isRelational =
+         do
+           (stC , stC⊩isStrictCodomainEqX)  idFuncRel perX .isStrictCodomain
+           (relY , relY⊩isRelationalEqY)  idFuncRel perY .isRelational
+           let
+             prover : ApplStrTerm as 3
+             prover = ` pair ̇ (` relY ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇ (` stC ̇ (` pr₁ ̇ # two))
+           return
+             (λ*3 prover ,
+              { (x , y₁) (x' , y₂) y₃ y₄ a b c (pr₁a⊩x~x' , pr₂a⊩y₁~y₂) (pr₁b⊩y₁~y₃ , pr₂b⊩x~x) c⊩y₃~y₄ 
+               subst
+                  r'  r'   perY .equality  (y₂ , y₄))
+                 (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
+                 (relY⊩isRelationalEqY y₁ y₂ y₃ y₄ (pr₂  a) (pr₁  b) c pr₂a⊩y₁~y₂ pr₁b⊩y₁~y₃ c⊩y₃~y₄) ,
+               subst
+                  r'  r'   perX .equality  (x' , x'))
+                 (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _))
+                 (stC⊩isStrictCodomainEqX x x' (pr₁  a) pr₁a⊩x~x') }))
+       ; isSingleValued =
+         do
+           (svY , svY⊩isSingleValuedY)  idFuncRel perY .isSingleValued
+           let
+             prover : ApplStrTerm as 2
+             prover = ` svY ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
+           return
+             (λ*2 prover ,
+              { (x , y) y' y'' r₁ r₂ (pr₁r₁⊩y~y' , pr₂r₁⊩x~x) (pr₁r₂⊩y~y'' , pr₂r₂⊩) 
+               subst
+                  r'  r'   perY .equality  (y' , y''))
+                 (sym (λ*2ComputationRule prover r₁ r₂))
+                 (svY⊩isSingleValuedY y y' y'' (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩y~y' pr₁r₂⊩y~y'') }))
+       ; isTotal =
+         do
+           let
+             prover : ApplStrTerm as 1
+             prover = ` pair ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero)
+           return
+             (λ* prover ,
+              { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) 
+               return
+                 (y ,
+                   (subst
+                      r'  r'   perY .equality  (y , y))
+                     (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                     pr₂r⊩y~y ,
+                    subst
+                      r'  r'   perX .equality  (x , x))
+                     (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                     pr₁r⊩x~x)) }))
+       }
+
+  binProdPr₂RT : RTMorphism binProdObRT perY
+  binProdPr₂RT = [ binProdPr₂FuncRel ]
+
+  module UnivProp
+    {Z : Type ℓ'}
+    (perZ : PartialEquivalenceRelation Z)
+    (f : RTMorphism perZ perX)
+    (g : RTMorphism perZ perY) where
+
+    isSetZ = PartialEquivalenceRelation.isSetX perZ
+
+    opaque
+      unfolding binProdObRT
+      theFuncRel : (F : FunctionalRelation perZ perX)  (G : FunctionalRelation perZ perY)  FunctionalRelation perZ binProdObRT
+      theFuncRel F G =
+        record
+              { relation =
+                record
+                  { isSetX = isSet× isSetZ (isSet× isSetX isSetY)
+                  ; ∣_∣ = λ { (z , x , y) r  (pr₁  r)   F .relation  (z , x) × (pr₂  r)   G .relation  (z , y) }
+                ; isPropValued = λ { (z , x , y) r  isProp× (F .relation .isPropValued _ _) (G .relation .isPropValued _ _) } }
+              ; isFuncRel =
+                record
+                 { isStrictDomain =
+                   do
+                     (stFD , stFD⊩isStrictDomain)  F .isStrictDomain
+                     let
+                       prover : ApplStrTerm as 1
+                       prover = ` stFD ̇ (` pr₁ ̇ # zero)
+                     return
+                       (λ* prover ,
+                         { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
+                          subst
+                             r'  r'   perZ .equality  (z , z))
+                            (sym (λ*ComputationRule prover r))
+                            (stFD⊩isStrictDomain z x (pr₁  r) pr₁r⊩Fzx) }))
+                 ; isStrictCodomain =
+                   do
+                     (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
+                     (stGC , stGC⊩isStrictCodomainG)  G .isStrictCodomain
+                     let
+                       prover : ApplStrTerm as 1
+                       prover = ` pair ̇ (` stFC ̇ (` pr₁ ̇ # zero)) ̇ (` stGC ̇ (` pr₂ ̇ # zero))
+                     return
+                       (λ* prover ,
+                        { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
+                         subst
+                            r'  r'   perX .equality  (x , x))
+                           (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                           (stFC⊩isStrictCodomainF z x (pr₁  r) pr₁r⊩Fzx) ,
+                         subst
+                            r'  r'   perY .equality  (y , y))
+                           (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                           (stGC⊩isStrictCodomainG z y (pr₂  r) pr₂r⊩Gzy) }))
+                 ; isRelational =
+                   do
+                     (relF , relF⊩isRelationalF)  F .isRelational
+                     (relG , relG⊩isRelationalG)  G .isRelational
+                     let
+                       prover : ApplStrTerm as 3
+                       prover = ` pair ̇ (` relF ̇ # two ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` relG ̇ # two ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
+                     return
+                       (λ*3 prover ,
+                        { z z' (x , y) (x' , y') a b c a⊩z~z' (pr₁b⊩Fzx , pr₂b⊩Gzy) (pr₁c⊩x~x' , pr₂c⊩y~y') 
+                         (subst  r'  r'   F .relation  (z' , x')) (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _)) (relF⊩isRelationalF z z' x x' _ _ _ a⊩z~z' pr₁b⊩Fzx pr₁c⊩x~x')) ,
+                         subst  r'  r'   G .relation  (z' , y')) (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _)) (relG⊩isRelationalG z z' y y' _ _ _ a⊩z~z' pr₂b⊩Gzy pr₂c⊩y~y') }))
+                 ; isSingleValued =
+                   do
+                     (svF , svF⊩isSingleValuedF)  F .isSingleValued
+                     (svG , svG⊩isSingleValuedG)  G .isSingleValued
+                     let
+                       prover : ApplStrTerm as 2
+                       prover = ` pair ̇ (` svF ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` svG ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
+                     return
+                       (λ*2 prover ,
+                        { z (x , y) (x' , y') r₁ r₂ (pr₁r₁⊩Fzx , pr₂r₁⊩Gzy) (pr₁r₂⊩Fzx' , pr₂r₂⊩Gzy') 
+                         subst
+                            r'  r'   perX .equality  (x , x'))
+                           (sym (cong  x  pr₁  x) (λ*2ComputationRule prover r₁ r₂)  pr₁pxy≡x _ _))
+                           (svF⊩isSingleValuedF z x x' (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩Fzx pr₁r₂⊩Fzx') ,
+                         subst
+                            r'  r'   perY .equality  (y , y'))
+                           (sym (cong  x  pr₂  x) (λ*2ComputationRule prover r₁ r₂)  pr₂pxy≡y _ _))
+                           (svG⊩isSingleValuedG z y y' (pr₂  r₁) (pr₂  r₂) pr₂r₁⊩Gzy pr₂r₂⊩Gzy') }))
+                 ; isTotal =
+                   do
+                     (tlF , tlF⊩isTotalF)  F .isTotal
+                     (tlG , tlG⊩isTotalG)  G .isTotal
+                     let
+                       prover : ApplStrTerm as 1
+                       prover = ` pair ̇ (` tlF ̇ # zero) ̇ (` tlG ̇ # zero)
+                     return
+                       (λ* prover ,
+                        { z r r⊩z~z 
+                         do
+                           (x , tlFr⊩Fzx)  tlF⊩isTotalF z r r⊩z~z
+                           (y , tlGr⊩Gzy)  tlG⊩isTotalG z r r⊩z~z
+                           return
+                             ((x , y) ,
+                              (subst
+                                 r'  r'   F .relation  (z , x))
+                                (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                                tlFr⊩Fzx ,
+                               subst
+                                 r'  r'   G .relation  (z , y))
+                                (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                                tlGr⊩Gzy)) }))
+                 }}
+
+    opaque
+      unfolding binProdObRT
+      unfolding theFuncRel
+      theMap : RTMorphism perZ binProdObRT
+      theMap =
+        SQ.rec2
+          squash/
+           F G 
+            [ theFuncRel F G ])
+           { F F' G (F≤F' , F'≤F) 
+            let
+              answer =
+                do
+                  (s , s⊩F≤F')  F≤F'
+                  let
+                    prover : ApplStrTerm as 1
+                    prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+                  return
+                    (λ* prover ,
+                      { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
+                       subst
+                          r'  r'   F' .relation  (z , x))
+                         (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                         (s⊩F≤F' z x (pr₁  r) pr₁r⊩Fzx) ,
+                       subst
+                          r'  r'   G .relation  (z , y))
+                         (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                         pr₂r⊩Gzy }))
+            in
+            eq/ _ _ (answer , F≤G→G≤F perZ binProdObRT (theFuncRel F G) (theFuncRel F' G) answer) })
+           { F G G' (G≤G' , G'≤G) 
+            let
+              answer =
+                do
+                  (s , s⊩G≤G')  G≤G'
+                  let
+                    prover : ApplStrTerm as 1
+                    prover = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero))
+                  return
+                    (λ* prover ,
+                     { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
+                      (subst
+                         r'  r'   F .relation  (z , x))
+                        (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                        pr₁r⊩Fzx) ,
+                      (subst
+                         r'  r'   G' .relation  (z , y))
+                        (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                        (s⊩G≤G' z y (pr₂  r) pr₂r⊩Gzy)) }))
+            in eq/ _ _ (answer , (F≤G→G≤F perZ binProdObRT (theFuncRel F G) (theFuncRel F G') answer)) })
+          f g
+  opaque
+    unfolding UnivProp.theMap
+    unfolding UnivProp.theFuncRel
+    unfolding binProdPr₁RT
+    unfolding binProdPr₁FuncRel
+    unfolding composeRTMorphism
+    unfolding binProdPr₂FuncRel
+    binProductRT : BinProduct RT (X , perX) (Y , perY)
+    BinProduct.binProdOb binProductRT = X × Y , binProdObRT
+    BinProduct.binProdPr₁ binProductRT = binProdPr₁RT
+    BinProduct.binProdPr₂ binProductRT = binProdPr₂RT
+    BinProduct.univProp binProductRT {Z , perZ} f g =
+      uniqueExists
+        (UnivProp.theMap perZ f g)
+        -- There is probably a better less kluged version of this proof
+        -- But this is the best I could do
+        (SQ.elimProp3
+          {P = λ f g theMap'   (foo : theMap'  (UnivProp.theMap perZ f g))  composeRTMorphism _ _ _ theMap' binProdPr₁RT  f}
+           f g h  isPropΠ λ h≡  squash/ _ _)
+           F G theFuncRel' [theFuncRel']≡theMap 
+            let
+              answer =
+                do
+                  let
+                    (p , q) = (SQ.effective
+                         a b  isProp× isPropPropTrunc isPropPropTrunc)
+                        (isEquivRelBientailment perZ binProdObRT)
+                        theFuncRel'
+                        (UnivProp.theFuncRel perZ [ F ] [ G ] F G)
+                        [theFuncRel']≡theMap)
+                  (p , p⊩theFuncRel'≤theFuncRel)  p
+                  (q , q⊩theFuncRel≤theFuncRel')  q
+                  (relF , relF⊩isRelationalF)  F .isRelational
+                  (stD , stD⊩isStrictDomain)  theFuncRel' .isStrictDomain
+                  let
+                    prover : ApplStrTerm as 1
+                    prover = ` relF ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ (` p ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+                  return
+                    (λ* prover ,
+                    λ z x r r⊩∃ 
+                      transport
+                        (propTruncIdempotent (F .relation .isPropValued _ _))
+                        (do
+                          ((x' , y) , (pr₁r⊩theFuncRel'zx'y , (pr₁pr₂r⊩x~x' , pr₂pr₂r⊩y~y)))  r⊩∃
+                          return
+                            (subst
+                               r'  r'   F .relation  (z , x))
+                              (sym (λ*ComputationRule prover r))
+                              (relF⊩isRelationalF
+                                z z x' x
+                                (stD  (pr₁  r)) (pr₁  (p  (pr₁  r))) (pr₁  (pr₂  r))
+                                (stD⊩isStrictDomain z (x' , y) (pr₁  r) pr₁r⊩theFuncRel'zx'y )
+                                (p⊩theFuncRel'≤theFuncRel z (x' , y) (pr₁  r) pr₁r⊩theFuncRel'zx'y .fst)
+                                 pr₁pr₂r⊩x~x'))))
+            in
+            eq/ _ _ (answer , F≤G→G≤F perZ perX (composeFuncRel _ _ _ theFuncRel' binProdPr₁FuncRel) F answer))
+          f
+          g
+          (UnivProp.theMap perZ f g)
+          refl ,
+        SQ.elimProp3
+          {P = λ f g theMap'   (foo : theMap'  (UnivProp.theMap perZ f g))  composeRTMorphism _ _ _ theMap' binProdPr₂RT  g}
+           f g  h  isPropΠ λ h≡  squash/ _ _)
+           F G theFuncRel' [theFuncRel']≡theMap 
+            let
+              answer =
+                do
+                  let
+                    (p , q) = (SQ.effective
+                         a b  isProp× isPropPropTrunc isPropPropTrunc)
+                        (isEquivRelBientailment perZ binProdObRT)
+                        theFuncRel'
+                        (UnivProp.theFuncRel perZ [ F ] [ G ] F G)
+                        [theFuncRel']≡theMap)
+                  (p , p⊩theFuncRel'≤theFuncRel)  p
+                  (q , q⊩theFuncRel≤theFuncRel')  q
+                  (relG , relG⊩isRelationalG)  G .isRelational
+                  (st , st⊩isStrictDomainTheFuncRel')  theFuncRel' .isStrictDomain
+                  let
+                    prover : ApplStrTerm as 1
+                    prover = ` relG ̇ (` st ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (` p ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+                  return
+                    (λ* prover ,
+                     z y r r⊩∃ 
+                      transport
+                        (propTruncIdempotent (G .relation .isPropValued _ _))
+                        (do
+                          ((x , y') , (pr₁r⊩theFuncRel'zxy' , pr₁pr₂r⊩y'~y , pr₂pr₂r⊩x~x))  r⊩∃
+                          return
+                            (subst
+                               r'  r'   G .relation  (z , y))
+                              (sym (λ*ComputationRule prover r)) 
+                              (relG⊩isRelationalG
+                                z z y' y
+                                (st  (pr₁  r)) (pr₂  (p  (pr₁  r))) (pr₁  (pr₂  r))
+                                (st⊩isStrictDomainTheFuncRel' z (x , y') (pr₁  r) pr₁r⊩theFuncRel'zxy')
+                                (p⊩theFuncRel'≤theFuncRel z (x , y') (pr₁  r) pr₁r⊩theFuncRel'zxy' .snd)
+                                pr₁pr₂r⊩y'~y)))))
+            in
+            eq/ _ _ (answer , F≤G→G≤F perZ perY (composeFuncRel _ _ _ theFuncRel' binProdPr₂FuncRel) G answer))
+          f g
+          (UnivProp.theMap perZ f g)
+          refl)
+         !  isProp× (squash/ _ _) (squash/ _ _))
+        λ { !' (!'⋆π₁≡f , !'⋆π₂≡g) 
+          SQ.elimProp3
+            {P = λ !' f g   (foo : composeRTMorphism _ _ _ !' binProdPr₁RT  f) (bar : composeRTMorphism _ _ _ !' binProdPr₂RT  g)  UnivProp.theMap perZ f g  !'}
+             f g !'  isPropΠ λ _  isPropΠ λ _  squash/ _ _)
+             !' F G !'⋆π₁≡F !'⋆π₂≡G 
+              let
+                answer =
+                  do
+                    let
+                      (p , q)   = SQ.effective (isPropValuedBientailment perZ perX) (isEquivRelBientailment perZ perX) (composeFuncRel _ _ _ !' binProdPr₁FuncRel) F !'⋆π₁≡F
+                      (p' , q') = SQ.effective (isPropValuedBientailment perZ perY) (isEquivRelBientailment perZ perY) (composeFuncRel _ _ _ !' binProdPr₂FuncRel) G !'⋆π₂≡G
+                    (q , q⊩F≤!'⋆π₁)  q
+                    (q' , q'⊩G≤!'⋆π₂)  q'
+                    (rel!' , rel!'⊩isRelational!')  !' .isRelational
+                    (sv!' , sv!'⊩isSingleValued!')  !' .isSingleValued
+                    (tX , tX⊩isTransitiveX)  perX .isTransitive
+                    (sX , sX⊩isSymmetricX)  perX .isSymmetric
+                    (stD!' , stD!'⊩isStrictDomain!')  !' .isStrictDomain
+                    let
+                      realizer : ApplStrTerm as 1 -- cursed
+                      realizer =
+                        ` rel!' ̇ (` stD!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # zero)))) ̇
+                          (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # zero))) ̇
+                          (` pair ̇
+                           (` tX ̇
+                            (` sX ̇
+                             (` pr₁ ̇
+                              (` sv!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # zero)))))) ̇
+                            (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pr₁ ̇ # zero))))) ̇
+                           (` pr₁ ̇ (` pr₂ ̇ (` q' ̇ (` pr₂ ̇ # zero)))))
+                    return
+                      (λ* realizer ,
+                       { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
+                        transport
+                          (propTruncIdempotent (!' .relation .isPropValued _ _))
+                          (do
+                            ((x' , y') , ⊩!'zx'y' , ⊩x'~x , ⊩y'~y')  q⊩F≤!'⋆π₁ z x _ pr₁r⊩Fzx
+                            ((x'' , y'') , ⊩!'zx''y'' , ⊩y''~y , ⊩x''~x'')  q'⊩G≤!'⋆π₂ z y _ pr₂r⊩Gzy
+                            let
+                              (⊩x'~x'' , ⊩y'~y'') = sv!'⊩isSingleValued!' z (x' , y') (x'' , y'') _ _ ⊩!'zx'y' ⊩!'zx''y''
+                              ⊩x''~x = tX⊩isTransitiveX x'' x' x _ _ (sX⊩isSymmetricX x' x'' _ ⊩x'~x'') ⊩x'~x 
+                            return
+                              (subst
+                                 r'  r'   !' .relation  (z , x , y))
+                                (sym (λ*ComputationRule realizer r))
+                                (rel!'⊩isRelational!'
+                                  z z
+                                  (x'' , y'')
+                                  (x , y)
+                                  _ _ _
+                                  (stD!'⊩isStrictDomain!' z (x' , y') _ ⊩!'zx'y') ⊩!'zx''y'' ((subst  r'  r'   perX .equality  (x'' , x)) (sym (pr₁pxy≡x _ _)) ⊩x''~x) , (subst  r'  r'   perY .equality  (y'' , y)) (sym (pr₂pxy≡y _ _)) ⊩y''~y))))) }))
+              in
+              eq/ _ _ (answer , F≤G→G≤F perZ binProdObRT (UnivProp.theFuncRel perZ [ F ] [ G ] F G)
+                                 !' answer))
+            !' f g !'⋆π₁≡f !'⋆π₂≡g }
+
+binProductsRT : BinProducts RT
+binProductsRT (X , perX) (Y , perY) = binProductRT perX perY
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.Equalizer.html b/docs/Realizability.Topos.Equalizer.html new file mode 100644 index 0000000..916a1e2 --- /dev/null +++ b/docs/Realizability.Topos.Equalizer.html @@ -0,0 +1,619 @@ + +Realizability.Topos.Equalizer
{-
+
+EQUALISERS IN RT(A)
+────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────
+────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────
+
+Consider two parallel morphisms f g from (X, _=X_) to (Y , _=Y_)
+
+In order to construct their equaliser we need to first construct an auxillary object (X , _=E_)
+and construct the equaliser as an injection from (X , _=E_) to (X , _=X_)
+
+However, we cannot, in general show that RT has equalisers because the object (X , _=E_) that injects into (X , _=X_) depends
+on choice of representatives of f and g.
+
+We can, however, prove a weaker theorem. We can show that equalisers *merely* exist.
+
+More formally, we can show that ∃[ ob ∈ RTObject ] ∃[ eq ∈ RTMorphism ob (X , _=X_) ] (univPropEqualizer eq)
+
+To do this, we firstly show the universal property for the case when we have already been given the
+representatives.
+
+Since we are eliminating a set quotient into a proposition, we can choose any representatives.
+
+Thus we have shown that RT merely has equalisers.
+
+The idea of showing the mere existence of equalisers was suggested by Jon Sterling.
+
+See also : Remark 2.7 of "Tripos Theory" by JHP
+
+──────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────────
+
+An extra note worth adding is that the code is quite difficult to read and very ugly. This is mostly due to the fact that a lot
+of the things that are "implicit" in an informal setting need to be justified here. More so than usual.
+
+There is additional bureacracy because we have to deal with eliminators of set quotients. This makes things a little more complicated.
+
+-}
+open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Structure
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Unit
+open import Cubical.Data.Empty
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation as PT
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Limits.BinProduct
+
+module Realizability.Topos.Equalizer
+  { ℓ' ℓ''} {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
+
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+open FunctionalRelation
+open PartialEquivalenceRelation
+
+equalizerUnivProp :
+   {X Y : Type ℓ'}
+   (perX : PartialEquivalenceRelation X)
+   (perY : PartialEquivalenceRelation Y)
+   (f g : RTMorphism perX perY)
+   (equalizerOb : PartialEquivalenceRelation X)
+   (inc : RTMorphism equalizerOb perX)
+   Type _
+equalizerUnivProp {X} {Y} perX perY f g equalizerOb inc =
+    ((composeRTMorphism _ _ _ inc f)  (composeRTMorphism _ _ _ inc g)) ×
+     {Z : Type ℓ'} (perZ : PartialEquivalenceRelation Z) (inc' : RTMorphism perZ perX)
+     (composeRTMorphism _ _ _ inc' f)  (composeRTMorphism _ _ _ inc' g)
+    -----------------------------------------------------------------------------------
+     ∃![ !  RTMorphism perZ equalizerOb ] (composeRTMorphism _ _ _ ! inc  inc')
+
+module _
+  {X : Type ℓ'}
+  {Y : Type ℓ'}
+  (perX : PartialEquivalenceRelation X)
+  (perY : PartialEquivalenceRelation Y)
+  (f g : RTMorphism perX perY) where
+
+  opaque
+    equalizerPer :  (F G : FunctionalRelation perX perY)  PartialEquivalenceRelation X
+    equalizerPer F G =
+      record
+              { equality =
+                record
+                  { isSetX = isSet× (perX .isSetX) (perX .isSetX)
+                  ; ∣_∣ = λ { (x , x') r 
+                    ((pr₁  r)   perX .equality  (x , x')) ×
+                    (∃[ y  Y ] (pr₁  (pr₂  r))   F .relation  (x , y) × (pr₂  (pr₂  r))   G .relation  (x , y)) }
+                  ; isPropValued = λ { (x , x') r  isProp× (perX .equality .isPropValued _ _) isPropPropTrunc } }
+              ; isPerEquality =
+                record
+                  { isSetX = perX .isSetX
+                  ; isSymmetric =
+                    do
+                      (s , s⊩isSymmetricX)  perX .isSymmetric
+                      (relF , relF⊩isRelationalF)  F .isRelational
+                      (relG , relG⊩isRelationalG)  G .isRelational
+                      (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
+                      let
+                        prover : ApplStrTerm as 1
+                        prover =
+                          ` pair ̇
+                            (` s ̇ (` pr₁ ̇ # zero)) ̇
+                            (` pair ̇
+                              (` relF ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)))) ̇
+                              (` relG ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)))))
+                      return
+                        (λ* prover ,
+                         { x x' r (pr₁r⊩x~x' , pr₂r⊩∃) 
+                          subst
+                             r'  r'   perX .equality  (x' , x))
+                            (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                            (s⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x') ,
+                          do
+                            (y , pr₁pr₂r⊩Fxy , pr₂pr₂r⊩Gxy)  pr₂r⊩∃
+                            return
+                              (y ,
+                              subst
+                                 r'  r'   F .relation  (x' , y))
+                                (sym (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                                (relF⊩isRelationalF
+                                  x x' y y
+                                  (pr₁  r) (pr₁  (pr₂  r)) (stFC  (pr₁  (pr₂  r)))
+                                  pr₁r⊩x~x'
+                                  pr₁pr₂r⊩Fxy
+                                  (stFC⊩isStrictCodomainF x y (pr₁  (pr₂  r)) pr₁pr₂r⊩Fxy)) ,
+                              subst
+                                 r'  r'   G .relation  (x' , y))
+                                (sym (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                                (relG⊩isRelationalG
+                                  x x' y y
+                                  (pr₁  r) (pr₂  (pr₂  r)) (stFC  (pr₁  (pr₂  r)))
+                                  pr₁r⊩x~x'
+                                  pr₂pr₂r⊩Gxy
+                                  (stFC⊩isStrictCodomainF x y (pr₁  (pr₂  r)) pr₁pr₂r⊩Fxy))) }))
+                  ; isTransitive =
+                    do
+                      (t , t⊩isTransitiveX)  perX .isTransitive
+                      (relF , relF⊩isRelationalF)  F .isRelational
+                      (relG , relG⊩isRelationalG)  G .isRelational
+                      let
+                        prover : ApplStrTerm as 2
+                        prover = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ # one)) ̇ (` pr₂ ̇ (` pr₂ ̇ # one)))
+                      return
+                        (λ*2 prover ,
+                        λ { x₁ x₂ x₃ a b (pr₁a⊩x₁~x₂ , pr₂a⊩∃) (pr₁b⊩x₂~x₃ , pr₂b⊩∃) 
+                          subst
+                             r'  r'   perX .equality  (x₁ , x₃))
+                            (sym (cong  x  pr₁  x) (λ*2ComputationRule prover a b)  pr₁pxy≡x _ _))
+                            (t⊩isTransitiveX x₁ x₂ x₃ (pr₁  a) (pr₁  b) pr₁a⊩x₁~x₂ pr₁b⊩x₂~x₃) ,
+                          do
+                            (y , (pr₁pr₂a⊩Fx₁y , pr₂pr₂a⊩Gx₁y))  pr₂a⊩∃
+                            (y' , (pr₁pr₂a⊩Fx₂y' , pr₂pr₂a⊩Gx₂y'))  pr₂b⊩∃
+                            return
+                              (y ,
+                              subst
+                                 r'  r'   F .relation  (x₁ , y))
+                                (sym (cong  x  pr₁  (pr₂  x)) (λ*2ComputationRule prover a b)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                                pr₁pr₂a⊩Fx₁y ,
+                              subst
+                                 r'  r'   G .relation  (x₁ , y))
+                                (sym (cong  x  pr₂  (pr₂  x)) (λ*2ComputationRule prover a b)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                                pr₂pr₂a⊩Gx₁y) }) } }
+
+  opaque
+    unfolding idFuncRel
+    unfolding equalizerPer
+    equalizerFuncRel :  (F G : FunctionalRelation perX perY)  FunctionalRelation (equalizerPer F G) perX
+    equalizerFuncRel F G =
+      record
+        { relation = equalizerPer F G .equality
+        ; isFuncRel =
+          record
+           { isStrictDomain = idFuncRel (equalizerPer F G) .isStrictDomain
+           ; isStrictCodomain =
+             do
+               (stC , stC⊩isStrictCodomain)  idFuncRel perX .isStrictCodomain
+               let
+                 prover : ApplStrTerm as 1
+                 prover = ` stC ̇ (` pr₁ ̇ # zero)
+               return
+                 (λ* prover ,
+                  { x x' r (pr₁r⊩x~x' , pr₂r⊩∃) 
+                   subst  r'  r'   perX .equality  (x' , x')) (sym (λ*ComputationRule prover r)) (stC⊩isStrictCodomain x x' (pr₁  r) pr₁r⊩x~x') }))
+           ; isRelational =
+             do
+               (relEqX , relEqX⊩isRelationalEqX)  idFuncRel perX .isRelational
+               (relF , relF⊩isRelationalF)  F .isRelational
+               (relG , relG⊩isRelationalG)  G .isRelational
+               (svF , svF⊩isSingleValuedF)  F .isSingleValued
+               let
+                 prover : ApplStrTerm as 3
+                 prover =
+                   ` pair ̇
+                     (` relEqX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇
+                     (` pair ̇
+                       (` relF ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` pr₁ ̇ (` pr₂ ̇ # one)))) ̇
+                       (` relG ̇ (` pr₁ ̇ # two) ̇ (` pr₂ ̇ (` pr₂ ̇ # two)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` pr₁ ̇ (` pr₂ ̇ # one)))))
+               return
+                 (λ*3 prover ,
+                  x₁ x₂ x₃ x₄ a b c (pr₁a⊩x₁~x₂ , pr₂a⊩) (pr₁b⊩x₁~x₃ , pr₂b⊩) c⊩x₃~x₄ 
+                   subst
+                      r'  r'   perX .equality  (x₂ , x₄))
+                     (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
+                     (relEqX⊩isRelationalEqX x₁ x₂ x₃ x₄ (pr₁  a) (pr₁  b) c pr₁a⊩x₁~x₂ pr₁b⊩x₁~x₃ c⊩x₃~x₄) ,
+                   do
+                     (y , pr₁pr₂a⊩Fx₁y , pr₂pr₂a⊩Gx₁y)  pr₂a⊩
+                     (y' , pr₁pr₂b⊩Fx₁y' , pr₂pr₂b⊩Gx₁y')  pr₂b⊩
+                     let
+                       y~y' = svF⊩isSingleValuedF x₁ y y' (pr₁  (pr₂  a)) (pr₁  (pr₂  b)) pr₁pr₂a⊩Fx₁y pr₁pr₂b⊩Fx₁y'
+                     return
+                       (y' ,
+                       subst
+                          r'  r'   F .relation  (x₂ , y'))
+                         (sym (cong  x  pr₁  (pr₂  x)) (λ*3ComputationRule prover a b c)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                         (relF⊩isRelationalF
+                           x₁ x₂ y y'
+                           (pr₁  a) (pr₁  (pr₂  a)) (svF  (pr₁  (pr₂  a))  (pr₁  (pr₂  b)))
+                           pr₁a⊩x₁~x₂ pr₁pr₂a⊩Fx₁y y~y') ,
+                       subst
+                          r'  r'   G .relation  (x₂ , y'))
+                         (sym (cong  x  pr₂  (pr₂  x)) (λ*3ComputationRule prover a b c)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                         (relG⊩isRelationalG
+                           x₁ x₂ y y'
+                           (pr₁  a) (pr₂  (pr₂  a)) (svF  (pr₁  (pr₂  a))  (pr₁  (pr₂  b)))
+                           pr₁a⊩x₁~x₂ pr₂pr₂a⊩Gx₁y y~y'))))
+           ; isSingleValued =
+             do
+               (svEqX , svEqX⊩isSingleValuedEqX)  idFuncRel perX .isSingleValued
+               let
+                 prover : ApplStrTerm as 2
+                 prover = ` svEqX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
+               return
+                 (λ*2 prover ,
+                  x₁ x₂ x₃ r₁ r₂ (pr₁r₁⊩x₁~x₂ , pr₁r₁⊩) (pr₁r₂⊩x₁~x₃ , pr₂r₂⊩) 
+                   subst
+                      r'  r'   perX .equality  (x₂ , x₃))
+                     (sym (λ*2ComputationRule prover r₁ r₂))
+                     (svEqX⊩isSingleValuedEqX x₁ x₂ x₃ (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩x₁~x₂ pr₁r₂⊩x₁~x₃)))
+           ; isTotal = idFuncRel (equalizerPer F G) .isTotal
+           } }
+
+  opaque
+    equalizerMorphism :  (F G : FunctionalRelation perX perY)  RTMorphism (equalizerPer F G) perX
+    equalizerMorphism F G = [ equalizerFuncRel F G ]
+
+  opaque
+    unfolding equalizerMorphism
+    unfolding equalizerFuncRel
+    unfolding composeRTMorphism
+    inc⋆f≡inc⋆g :  (F G : FunctionalRelation perX perY)  composeRTMorphism _ _ _ (equalizerMorphism F G) [ F ]  composeRTMorphism _ _ _ (equalizerMorphism F G) [ G ]
+    inc⋆f≡inc⋆g F G =
+      let
+        answer =
+          do
+            (relG , relG⊩isRelationalG)  G .isRelational
+            (svF , svF⊩isSingleValuedF)  F .isSingleValued
+            (relF , relF⊩isRelationalF)  F .isRelational
+            (sX , sX⊩isSymmetricX)  perX .isSymmetric
+            (stCF , stCF⊩isStrictCodomainF)  F .isStrictCodomain
+            let
+              realizer : ApplStrTerm as 1
+              realizer =
+                ` pair ̇
+                  (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))))) ̇
+                  (` relG ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
+                     (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
+                     (` relF ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ # zero) ̇ (` stCF ̇ (` pr₂ ̇ # zero)))))
+            return
+              (λ* realizer ,
+              -- unfold everything and bring it back in together
+               x y r r⊩∃ 
+                do
+                  (x' , (⊩x~x' , ∃y) , ⊩Fx'y)  r⊩∃
+                  (y' , ⊩Fxy' , ⊩Gxy')  ∃y
+                  let
+                    y'~y =
+                      svF⊩isSingleValuedF x y' y _ _ ⊩Fxy' (relF⊩isRelationalF x' x y y _ _ _ (sX⊩isSymmetricX x x' _ ⊩x~x') ⊩Fx'y (stCF⊩isStrictCodomainF x' y _ ⊩Fx'y))
+                  return
+                    (x' ,
+                    (subst
+                       r'  r'   perX .equality  (x , x'))
+                      (sym (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _))
+                      ⊩x~x' ,
+                    do
+                      return
+                        (y' ,
+                        subst
+                           r'  r'   F .relation  (x , y'))
+                          (sym
+                            (cong  x  pr₁  (pr₂  (pr₁  x))) (λ*ComputationRule realizer r) 
+                             cong  x  pr₁  (pr₂  x)) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                             pr₁pxy≡x _ _))
+                          ⊩Fxy' ,
+                        subst
+                           r'  r'   G .relation  (x , y'))
+                          (sym
+                            (cong  x  pr₂  (pr₂  (pr₁  x))) (λ*ComputationRule realizer r) 
+                             cong  x  pr₂  (pr₂  x)) (pr₁pxy≡x _ _) 
+                             cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                             pr₂pxy≡y _ _))
+                          ⊩Gxy')) ,
+                    subst
+                       r'  r'   G .relation  (x' , y))
+                      (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                      (relG⊩isRelationalG x x' y' y _ _ _ ⊩x~x' ⊩Gxy' y'~y))))
+      in
+      eq/ _ _
+        (answer , F≤G→G≤F (equalizerPer F G) perY (composeFuncRel _ _ _ (equalizerFuncRel F G) F) (composeFuncRel _ _ _ (equalizerFuncRel F G) G) answer)
+
+  module UnivProp
+    (F G : FunctionalRelation perX perY)
+    {Z : Type ℓ'}
+    (perZ : PartialEquivalenceRelation Z)
+    (h : RTMorphism perZ perX)
+    (h⋆f≡h⋆g : composeRTMorphism _ _ _ h [ F ]  composeRTMorphism _ _ _ h [ G ]) where opaque
+    unfolding equalizerPer
+    unfolding composeRTMorphism
+    unfolding equalizerMorphism
+    unfolding equalizerFuncRel
+    
+    private
+      !funcRel :  (H : FunctionalRelation perZ perX) (H⋆F≡H⋆G : composeRTMorphism _ _ _ [ H ] [ F ]  composeRTMorphism _ _ _ [ H ] [ G ])  FunctionalRelation perZ (equalizerPer F G)
+      !funcRel H H⋆F≡H⋆G =
+        let
+          (p , q) =
+            SQ.effective
+              (isPropValuedBientailment perZ perY)
+              (isEquivRelBientailment perZ perY)
+              (composeFuncRel _ _ _ H F)
+              (composeFuncRel _ _ _ H G)
+              H⋆F≡H⋆G
+        in
+        record
+              { relation = H .relation
+              ; isFuncRel =
+                record
+                 { isStrictDomain = H .isStrictDomain
+                 ; isStrictCodomain =
+                   do
+                     (p , p⊩H⋆F≤H⋆G)  p
+                     (q , q⊩H⋆G≤H⋆F)  q
+                     (tlF , tlF⊩isTotalF)  F .isTotal
+                     (stCH , stCH⊩isStrictCodomainH)  H .isStrictCodomain
+                     (relG , relG⊩isRelationalG)  G .isRelational
+                     (svH , svH⊩isSingleValuedH)  H .isSingleValued
+                     (stCF , stCF⊩isStrictCodomainF)  F .isStrictCodomain
+                     let
+                       -- possibly the ugliest realizer out there
+                       prover : ApplStrTerm as 1
+                       prover =
+                         ` pair ̇
+                           (` stCH ̇ # zero) ̇
+                           (` pair ̇
+                             (` tlF ̇ (` stCH ̇ # zero)) ̇
+                             (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCH ̇ # zero))))) ̇ # zero) ̇
+                             (` pr₂ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCH ̇ # zero))))) ̇
+                              (` stCF ̇ (` tlF ̇ (` stCH ̇ # zero)))))
+                     return
+                       (λ* prover ,
+                        z x r r⊩Hzx 
+                         let
+                             x~x = stCH⊩isStrictCodomainH z x r r⊩Hzx
+                         in
+                         subst  r  r   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) x~x ,
+                         (do
+                           (y , ⊩Fxy)  tlF⊩isTotalF x (stCH  r) x~x
+                           let
+                             hope =
+                               p⊩H⋆F≤H⋆G
+                                 z y (pair  r  (tlF  (stCH  r)))
+                                 (return
+                                   (x ,
+                                    subst  r'  r'   H .relation  (z , x)) (sym (pr₁pxy≡x _ _)) r⊩Hzx ,
+                                    subst  r'  r'   F .relation  (x , y)) (sym (pr₂pxy≡y _ _)) ⊩Fxy))
+                           return
+                             (y ,
+                             subst
+                                r'  r'   F .relation  (x , y))
+                               (sym
+                                 (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r) 
+                                  cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                                  pr₁pxy≡x _ _))
+                               ⊩Fxy ,
+                             -- god I wish there was a better way to do this :(
+                             transport
+                               (propTruncIdempotent (G .relation .isPropValued _ _))
+                               (do
+                                 (x' , ⊩Hzx' , ⊩Gx'y)  hope
+                                 return
+                                   (subst
+                                      r'  r'   G .relation  (x , y))
+                                     (sym
+                                       (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r) 
+                                        cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                                        pr₂pxy≡y _ _))
+                                     (relG⊩isRelationalG x' x y y _ _ _ (svH⊩isSingleValuedH z x' x _ _ ⊩Hzx' r⊩Hzx) ⊩Gx'y (stCF⊩isStrictCodomainF x y _ ⊩Fxy))))))))
+                 ; isRelational =
+                   do
+                     (relH , relH⊩isRelationalH)  H .isRelational
+                     let
+                       prover : ApplStrTerm as 3
+                       prover = ` relH ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
+                     return
+                       (λ*3 prover ,
+                        z z' x x' a b c a⊩z~z' b⊩Hzx (pr₁c⊩x~x' , pr₂c⊩∃) 
+                         subst
+                            r'  r'   H .relation  (z' , x'))
+                           (sym (λ*3ComputationRule prover a b c))
+                           (relH⊩isRelationalH z z' x x' a b (pr₁  c) a⊩z~z' b⊩Hzx pr₁c⊩x~x')))
+                 ; isSingleValued =
+                   do
+                     (svH , svH⊩isSingleValuedH)  H .isSingleValued
+                     (tlF , tlF⊩isTotalF)  F .isTotal
+                     (tlG , tlG⊩isTotalG)  G .isTotal
+                     (stCH , stCH⊩isStrictCodomainH)  H .isStrictCodomain
+                     (stCF , stCF⊩isStrictCodomainF)  F .isStrictCodomain
+                     (relG , relG⊩isRelationalG)  G .isRelational
+                     (p , p⊩H⋆F≤H⋆G)  p
+                     let
+                       prover : ApplStrTerm as 2
+                       prover =
+                         ` pair ̇
+                           (` svH ̇ # one ̇ # zero) ̇
+                           (` pair ̇
+                             (` tlF ̇ (` stCH ̇ # one)) ̇
+                             (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # one ̇ (` tlF ̇ (` stCH ̇ # one))))) ̇ # one) ̇
+                               (` pr₂ ̇ (` p ̇ (` pair ̇ # one ̇ (` tlF ̇ (` stCH ̇ # one))))) ̇
+                               (` stCF ̇(` tlF ̇ (` stCH ̇ # one)))))
+                     return
+                       (λ*2 prover ,
+                        z x x' r₁ r₂ r₁⊩Hzx r₂⊩Hzx' 
+                         let
+                           x~x' = svH⊩isSingleValuedH z x x' r₁ r₂ r₁⊩Hzx r₂⊩Hzx'
+                           x~x = stCH⊩isStrictCodomainH z x r₁ r₁⊩Hzx
+                         in
+                         subst
+                            r'  r'   perX .equality  (x , x'))
+                           (sym (cong  x  pr₁  x) (λ*2ComputationRule prover r₁ r₂)  pr₁pxy≡x _ _))
+                           x~x' ,
+                         do
+                           (y , ⊩Fxy)  tlF⊩isTotalF x _ x~x
+                           let
+                             y~y = stCF⊩isStrictCodomainF x y _ ⊩Fxy
+                             hope =
+                               p⊩H⋆F≤H⋆G z y
+                               (pair  r₁  (tlF  (stCH  r₁)))
+                               (do
+                                 return
+                                   (x ,
+                                   (subst  r'  r'   H .relation  (z , x)) (sym (pr₁pxy≡x _ _)) r₁⊩Hzx) ,
+                                   (subst  r'  r'   F .relation  (x , y)) (sym (pr₂pxy≡y _ _)) ⊩Fxy)))
+                           (x'' , ⊩Hzx'' , ⊩Gx''y)  hope
+                           -- Can not use the fact that x ≐ x''
+                           let
+                             x''~x = svH⊩isSingleValuedH z x'' x _ _ ⊩Hzx'' r₁⊩Hzx
+                           return
+                             (y ,
+                             subst
+                                r'  r'   F .relation  (x , y))
+                               (sym
+                                 (cong  x  pr₁  (pr₂  x)) (λ*2ComputationRule prover r₁ r₂) 
+                                  cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                                  pr₁pxy≡x _ _)) ⊩Fxy ,
+                             subst
+                                r'  r'   G .relation  (x , y))
+                               (sym
+                                 (cong  x  pr₂  (pr₂  x)) (λ*2ComputationRule prover r₁ r₂) 
+                                  cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                                  pr₂pxy≡y _ _))
+                               (relG⊩isRelationalG x'' x y y _ _ _ x''~x ⊩Gx''y y~y))))
+                 ; isTotal = H .isTotal
+                 } }
+                 
+    mainMap :
+      Σ[ !  RTMorphism perZ (equalizerPer F G) ]
+        (composeRTMorphism _ _ _ ! (equalizerMorphism F G)  h) ×
+        (∀ (!' : RTMorphism perZ (equalizerPer F G)) (!'⋆inc≡h : composeRTMorphism _ _ _ !' (equalizerMorphism F G)  h)  !'  !)
+    mainMap =
+      SQ.elim
+        {P =
+          λ h 
+           (h⋆f≡h⋆g : composeRTMorphism _ _ _ h [ F ]  composeRTMorphism _ _ _ h [ G ])
+           Σ[ !  _ ] (composeRTMorphism _ _ _ ! (equalizerMorphism F G)  h) ×
+          (∀ (!' : RTMorphism perZ (equalizerPer F G)) (!'⋆inc≡h : composeRTMorphism _ _ _ !' (equalizerMorphism F G)  h)  !'  !)}
+         h  isSetΠ λ _  isSetΣ squash/ λ !  isSet× (isProp→isSet (squash/ _ _)) (isSetΠ λ !'  isSetΠ λ _  isProp→isSet (squash/ !' !)))
+         H H⋆F≡H⋆G 
+          [ !funcRel H H⋆F≡H⋆G ] ,
+          let    
+            answer =
+              do
+                (relH , relH⊩isRelationalH)  H .isRelational
+                (stDH , stDH⊩isStrictDomainH)  H .isStrictDomain
+                let
+                  prover : ApplStrTerm as 1
+                  prover = ` relH ̇ (` stDH ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+                return
+                  (λ* prover ,
+                   z x r r⊩∃x' 
+                    transport
+                      (propTruncIdempotent (H .relation .isPropValued _ _))
+                      (do
+                        (x' , pr₁r⊩Hzx' , (pr₁pr₂r⊩x'~x , _))  r⊩∃x'
+                        return
+                          (subst
+                             r'  r'   H .relation  (z , x))
+                            (sym (λ*ComputationRule prover r))
+                            (relH⊩isRelationalH z z x' x _ _ _ (stDH⊩isStrictDomainH z x' (pr₁  r) pr₁r⊩Hzx') pr₁r⊩Hzx' pr₁pr₂r⊩x'~x)))))
+            !funcRel⋆inc≡H = eq/ _ _ (answer , F≤G→G≤F _ _ (composeFuncRel _ _ _ (!funcRel H H⋆F≡H⋆G) (equalizerFuncRel F G)) H answer)
+          in !funcRel⋆inc≡H ,
+          λ !' !'⋆inc≡H 
+            SQ.elimProp
+              {P =
+                λ !' 
+                 (foo : composeRTMorphism _ _ _ !' (equalizerMorphism F G)  [ H ])
+                 !'  [ !funcRel H H⋆F≡H⋆G ]}
+               !'  isPropΠ λ _  squash/ _ _)
+               !'funcRel !'funcRel⋆inc≡H 
+                let
+                  (p , q) = SQ.effective (isPropValuedBientailment perZ perX) (isEquivRelBientailment perZ perX) (composeFuncRel _ _ _ !'funcRel (equalizerFuncRel F G)) H !'funcRel⋆inc≡H
+                  (p' , q') = SQ.effective (isPropValuedBientailment perZ perY) (isEquivRelBientailment perZ perY) (composeFuncRel _ _ _ H F) (composeFuncRel _ _ _ H G) H⋆F≡H⋆G
+                  answer =
+                    do
+                      (q , q⊩inc⋆!'≤H)  q
+                      (rel!' , rel!'⊩isRelational!'FuncRel)  !'funcRel .isRelational
+                      (stDH , stDH⊩isStrictDomainH)  H .isStrictDomain
+                      let
+                        prover : ApplStrTerm as 1
+                        prover = ` rel!' ̇ (` stDH ̇ # zero) ̇ (` pr₁ ̇ (` q ̇ # zero)) ̇ (` pr₂ ̇ (` q ̇ # zero))
+                      return
+                        (λ* prover ,
+                         z x r r⊩Hzx 
+                          transport
+                            (propTruncIdempotent (!'funcRel .relation .isPropValued _ _))
+                            (do
+                              (x' , pr₁qr⊩!'zx' , ⊩x~x' , foo)  q⊩inc⋆!'≤H z x r r⊩Hzx
+                              return
+                                (subst
+                                   r'  r'   !'funcRel .relation  (z , x))
+                                  (sym (λ*ComputationRule prover r))
+                                  (rel!'⊩isRelational!'FuncRel
+                                    z z x' x _ _ _
+                                    (stDH⊩isStrictDomainH z x r r⊩Hzx)
+                                    pr₁qr⊩!'zx'
+                                    (⊩x~x' , foo))))))
+                in
+                eq/ _ _ (F≤G→G≤F perZ (equalizerPer F G) (!funcRel H H⋆F≡H⋆G) !'funcRel answer , answer))
+              !'
+              !'⋆inc≡H)
+         { H H' (H≤H' , H'≤H) 
+          funExtDep
+            {A = λ i  composeRTMorphism _ _ _ (eq/ H H' (H≤H' , H'≤H) i) [ F ]  composeRTMorphism _ _ _ (eq/ H H' (H≤H' , H'≤H) i) [ G ]}
+            λ {H⋆F≡H⋆G} {H'⋆F≡H'⋆G} p 
+              ΣPathPProp
+                {A = λ i  RTMorphism perZ (equalizerPer F G)}
+                {B = λ i ! 
+                  ((composeRTMorphism _ _ _ ! (equalizerMorphism F G))  eq/ H H' (H≤H' , H'≤H) i) ×
+                  (∀ (!' : RTMorphism perZ (equalizerPer F G))  composeRTMorphism _ _ _ !' (equalizerMorphism F G)  eq/ H H' (H≤H' , H'≤H) i  !'  !)}
+                 !  isProp× (squash/ _ _) (isPropΠ λ !'  isPropΠ λ !'⋆inc≡H'  squash/ _ _))
+                let
+                  answer =
+                    (do
+                      (s , s⊩H≤H')  H≤H'
+                      return
+                        (s ,
+                         z x r r⊩Hzx 
+                          s⊩H≤H' z x r r⊩Hzx)))
+                in eq/ _ _ (answer , F≤G→G≤F perZ (equalizerPer F G) (!funcRel H H⋆F≡H⋆G) (!funcRel H' H'⋆F≡H'⋆G) answer) })
+        h
+        h⋆f≡h⋆g
+    
+  -- We have now done the major work and can simply eliminate f and g
+  opaque
+    unfolding idFuncRel
+    unfolding equalizerPer
+    equalizer :
+      ∃[ equalizerOb  PartialEquivalenceRelation X ]
+      ∃[ inc  RTMorphism equalizerOb perX ]
+      (equalizerUnivProp perX perY f g equalizerOb inc)
+    equalizer =
+      SQ.elimProp2
+        {P = λ f g  ∃[ equalizerOb  PartialEquivalenceRelation X ] ∃[ inc  RTMorphism equalizerOb perX ] (equalizerUnivProp perX perY f g equalizerOb inc)}
+         f g  isPropPropTrunc)
+         F G 
+            return
+              ((equalizerPer F G) ,
+              (return
+                  ((equalizerMorphism F G) ,
+                  ((inc⋆f≡inc⋆g F G) ,
+                   {Z} perZ inc' inc'⋆f≡inc'⋆g 
+                    let
+                      (! , !⋆inc≡inc' , unique!) = UnivProp.mainMap F G perZ inc' inc'⋆f≡inc'⋆g
+                    in
+                    uniqueExists
+                      !
+                      !⋆inc≡inc'
+                       !  squash/ _ _)
+                      λ !' !'⋆inc≡inc'  sym (unique! !' !'⋆inc≡inc')))))))
+        f g
+      
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.Everything.html b/docs/Realizability.Topos.Everything.html index 8a2ca6a..4ec7be1 100644 --- a/docs/Realizability.Topos.Everything.html +++ b/docs/Realizability.Topos.Everything.html @@ -3,4 +3,11 @@ open import Realizability.Topos.Object open import Realizability.Topos.FunctionalRelation +open import Realizability.Topos.TerminalObject +open import Realizability.Topos.BinProducts +open import Realizability.Topos.Equalizer +open import Realizability.Topos.MonicReprFuncRel +open import Realizability.Topos.StrictRelation +open import Realizability.Topos.ResizedPredicate +open import Realizability.Topos.SubobjectClassifier \ No newline at end of file diff --git a/docs/Realizability.Topos.FunctionalRelation.html b/docs/Realizability.Topos.FunctionalRelation.html index 0fca2f8..dcc7b73 100644 --- a/docs/Realizability.Topos.FunctionalRelation.html +++ b/docs/Realizability.Topos.FunctionalRelation.html @@ -1,377 +1,611 @@ -Realizability.Topos.FunctionalRelation
{-# OPTIONS --allow-unsolved-metas #-}
-open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
-open import Realizability.CombinatoryAlgebra
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Structure
-open import Cubical.Foundations.HLevels
-open import Cubical.Data.Vec
-open import Cubical.Data.Nat
-open import Cubical.Data.FinData
-open import Cubical.Data.Fin hiding (Fin; _/_)
-open import Cubical.Data.Sigma
-open import Cubical.Data.Empty
-open import Cubical.Data.Unit
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.HITs.SetQuotients renaming (rec to setQuotRec; rec2 to setQuotRec2; elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2; elimProp3 to setQuotElimProp3)
-open import Cubical.Categories.Category
-
-module Realizability.Topos.FunctionalRelation
-  { ℓ' ℓ''}
-  {A : Type }
-  (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
-  where
-
-open import Realizability.Tripos.Logic.Syntax { = ℓ'}
-open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate')
-open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
-open import Realizability.Tripos.Prealgebra.Meets.Identity ca
-open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
-
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Predicate' renaming (isSetX to isSetPredicateBase)
-open PredicateProperties
-open Morphism
-
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
-
-private
-  Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''}
-
-open PartialEquivalenceRelation
-
-record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
-  equalityX = perX .equality
-  equalityY = perY .equality
-
-  field
-    relation : Predicate (X × Y)
-    isStrict :
-      ∃[ st  A ]
-      (∀ x y r
-       r   relation  (x , y)
-      ------------------------------------------------------------------------------------
-       (pr₁  (st  r))   equalityX  (x , x) × ((pr₂  (st  r))   equalityY  (y , y)))
-    isRelational :
-      ∃[ rl  A ]
-      (∀ x x' y y' a b c
-       a   equalityX  (x , x')
-       b   relation  (x , y)
-       c   equalityY  (y , y')
-      ------------------------------------------
-       (rl  (pair  a  (pair  b  c)))   relation  (x' , y'))
-    isSingleValued :
-      ∃[ sv  A ]
-      (∀ x y y' r₁ r₂
-       r₁   relation  (x , y)
-       r₂   relation  (x , y')
-      -----------------------------------
-       (sv  (pair  r₁  r₂))   equalityY  (y , y'))
-    isTotal :
-      ∃[ tl  A ]
-      (∀ x r  r   equalityX  (x , x)  ∃[ y  Y ] (tl  r)   relation  (x , y))
-
-open FunctionalRelation
-
-pointwiseEntailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  (F G : FunctionalRelation perX perY)  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
-pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe  A ] (∀ x y r  r   F .relation  (x , y)  (pe  r)   G .relation  (x , y))
-
--- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong
--- Lemma 4.3.5
-F≤G→G≤F :
-   {X Y : Type ℓ'}
-   (perX : PartialEquivalenceRelation X)
-   (perY : PartialEquivalenceRelation Y)
-   (F G : FunctionalRelation perX perY)
-   pointwiseEntailment perX perY F G
-   pointwiseEntailment perX perY G F
-F≤G→G≤F {X} {Y} perX perY F G F≤G =
-  do
-    (r , r⊩F≤G)  F≤G
-    (stG , stG⊩isStrictG)  G .isStrict
-    (tlF , tlF⊩isTotalF)  F .isTotal
-    (svG , svG⊩isSingleValuedG)  G .isSingleValued
-    (rlF , rlF⊩isRelational)  F .isRelational
-    let
-      prover : ApplStrTerm as 1
-      prover = ` rlF ̇ (` pair ̇ (` pr₁ ̇ (` stG ̇ # fzero)) ̇ (` pair ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero))) ̇ (` svG ̇ (` pair ̇ (` r ̇ (` tlF ̇ (` pr₁ ̇ (` stG ̇ # fzero)))) ̇ # fzero))))
-    return
-      (λ* prover ,
-       x y r' r'⊩Gxy 
-        subst
-           r''  r''   F .relation  (x , y))
-          (sym (λ*ComputationRule prover (r'  [])))
-          (transport
-            (propTruncIdempotent (F .relation .isPropValued (x , y) _))
-            (do
-              (y' , Fxy')  tlF⊩isTotalF x (pr₁  (stG  r')) (stG⊩isStrictG x y r' r'⊩Gxy .fst)
-              return
-                (rlF⊩isRelational
-                  x x y' y
-                  (pr₁  (stG  r'))
-                  (tlF  (pr₁  (stG  r')))
-                  (svG  (pair  (r  (tlF  (pr₁  (stG  r'))))  r'))
-                  (stG⊩isStrictG x y r' r'⊩Gxy .fst)
-                  Fxy'
-                  (svG⊩isSingleValuedG x y' y (r  (tlF  (pr₁  (stG  r')))) r' (r⊩F≤G x y' (tlF  (pr₁  (stG  r'))) Fxy') r'⊩Gxy))))))
-
-equalityFuncRel :  {X : Type ℓ'}  (perX : PartialEquivalenceRelation X)  FunctionalRelation perX perX
-relation (equalityFuncRel {X} perX) = perX .equality
-isStrict (equalityFuncRel {X} perX) =
-  do
-    (s , s⊩isSymmetric)  perX .isSymmetric
-    (t , t⊩isTransitive)  perX .isTransitive
-    let
-      prover : ApplStrTerm as 1
-      prover = ` pair ̇ (` t ̇ (` pair ̇ # fzero ̇ (` s ̇ # fzero))) ̇ (` t ̇ (` pair ̇ (` s ̇ # fzero) ̇ # fzero))
-    return
-      (λ* prover ,
-      λ x x' r r⊩x~x' 
-        let
-          pr₁proofEq : pr₁  (λ* prover  r)  t  (pair  r  (s  r))
-          pr₁proofEq =
-            pr₁  (λ* prover  r)
-              ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (r  [])) 
-            pr₁  (pair  (t  (pair  r  (s  r)))  (t  (pair  (s  r)  r)))
-              ≡⟨ pr₁pxy≡x _ _ 
-            t  (pair  r  (s  r))
-              
-
-          pr₂proofEq : pr₂  (λ* prover  r)  t  (pair  (s  r)  r)
-          pr₂proofEq =
-            pr₂  (λ* prover  r)
-              ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (r  [])) 
-            pr₂  (pair  (t  (pair  r  (s  r)))  (t  (pair  (s  r)  r)))
-              ≡⟨ pr₂pxy≡y _ _ 
-            t  (pair  (s  r)  r)
-              
-        in
-        subst  r'  r'   perX .equality  (x , x)) (sym pr₁proofEq) (t⊩isTransitive x x' x r (s  r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')) ,
-        subst  r'  r'   perX .equality  (x' , x')) (sym pr₂proofEq) (t⊩isTransitive x' x x' (s  r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x'))
-isRelational (equalityFuncRel {X} perX) =
-  do
-    (s , s⊩isSymmetric)  perX .isSymmetric
-    (t , t⊩isTransitive)  perX .isTransitive
-    let
-      prover : ApplStrTerm as 1
-      prover = ` t ̇ (` pair ̇ (` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))
-
-    return
-      (λ* prover ,
-       x x' y y' a b c a⊩x~x' b⊩x~y c⊩y~y' 
-        let
-          proofNormalForm : (λ* prover  (pair  a  (pair  b  c)))  (t  (pair  (t  (pair  (s  a)  b))  c))
-          proofNormalForm =
-            λ* prover  (pair  a  (pair  b  c))
-              ≡⟨ λ*ComputationRule prover ((pair  a  (pair  b  c))  []) 
-            (t 
-            (pair 
-             (t 
-              (pair  (s  (pr₁  (pair  a  (pair  b  c)))) 
-               (pr₁  (pr₂  (pair  a  (pair  b  c))))))
-              (pr₂  (pr₂  (pair  a  (pair  b  c))))))
-             ≡⟨
-               cong₂
-                  x y  (t  (pair  (t  (pair  (s  x)  y))  (pr₂  (pr₂  (pair  a  (pair  b  c)))))))
-                 (pr₁pxy≡x _ _)
-                 (pr₁  (pr₂  (pair  a  (pair  b  c)))
-                   ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
-                 pr₁  (pair  b  c)
-                   ≡⟨ pr₁pxy≡x _ _ 
-                 b
-                   )
-              
-            t  (pair  (t  (pair  (s  a)  b))  (pr₂  (pr₂  (pair  a  (pair  b  c)))))
-              ≡⟨
-                cong
-                   x  t  (pair  (t  (pair  (s  a)  b))  x))
-                  (pr₂  (pr₂  (pair  a  (pair  b  c)))
-                    ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
-                  pr₂  (pair  b  c)
-                    ≡⟨ pr₂pxy≡y _ _ 
-                  c
-                    )
-               
-            t  (pair  (t  (pair  (s  a)  b))  c)
-              
-        in
-        subst
-           r  r   perX .equality  (x' , y'))
-          (sym proofNormalForm)
-          (t⊩isTransitive x' y y' (t  (pair  (s  a)  b)) c (t⊩isTransitive x' x y (s  a) b (s⊩isSymmetric x x' a a⊩x~x') b⊩x~y) c⊩y~y')))
-isSingleValued (equalityFuncRel {X} perX) =
-  do
-    (s , s⊩isSymmetric)  perX .isSymmetric
-    (t , t⊩isTransitive)  perX .isTransitive
-    let
-      prover : ApplStrTerm as 1
-      prover = ` t ̇ (` pair ̇ (` s ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero))
-    return
-      (λ* prover ,
-       x y y' r₁ r₂ r₁⊩x~y r₂⊩x~y' 
-        let
-          proofEq : λ* prover  (pair  r₁  r₂)  t  (pair  (s  r₁)  r₂)
-          proofEq = {!!}
-        in
-        subst
-           r  r   perX .equality  (y , y'))
-          (sym proofEq)
-          (t⊩isTransitive y x y' (s  r₁) r₂ (s⊩isSymmetric x y r₁ r₁⊩x~y) r₂⊩x~y')))
-isTotal (equalityFuncRel {X} perX) =
-  do
-    (s , s⊩isSymmetric)  perX .isSymmetric
-    (t , t⊩isTransitive)  perX .isTransitive
-    return (Id ,  x r r⊩x~x   x , (subst  r  r   perX .equality  (x , x)) (sym (Ida≡a r)) r⊩x~x) ∣₁))
-
-composeFuncRel :
-   {X Y Z : Type ℓ'}
-   (perX : PartialEquivalenceRelation X)
-   (perY : PartialEquivalenceRelation Y)
-   (perZ : PartialEquivalenceRelation Z)
-   (F : FunctionalRelation perX perY)
-   (G : FunctionalRelation perY perZ)
-   FunctionalRelation perX perZ
-
-(relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
-  record
-    { isSetX = isSet× (perX .isSetX) (perZ .isSetX)
-    ; ∣_∣ = λ { (x , z) r  ∃[ y  Y ] (pr₁  r)   F .relation  (x , y) × (pr₂  r)   G .relation  (y , z) }
-    ; isPropValued = λ x a  isPropPropTrunc }
-isStrict (composeFuncRel {X} {Y} {Z} perX perY perZ F G) =
-  do
-    (stF , stF⊩isStrictF)  F .isStrict
-    (stG , stG⊩isStrictG)  G .isStrict
-    let
-      prover : ApplStrTerm as 1
-      prover = ` pair ̇ (` pr₁ ̇ (` stF ̇ (` pr₁ ̇ # fzero))) ̇ (` pr₂ ̇ (` stG ̇ (` pr₂ ̇ # fzero)))
-    return
-      (λ* prover ,
-       x z r r⊩∃y 
-        transport (propTruncIdempotent {!!})
-        (do
-          (y , pr₁r⊩Fxy , pr₂r⊩Gxy)  r⊩∃y
-          let
-            pr₁proofEq : pr₁  (λ* prover  r)  pr₁  (stF  (pr₁  r))
-            pr₁proofEq =
-              pr₁  (λ* prover  r)
-                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (r  [])) 
-              pr₁  (pair  (pr₁  (stF  (pr₁  r)))  (pr₂  (stG  (pr₂  r))))
-                ≡⟨ pr₁pxy≡x _ _ 
-              pr₁  (stF  (pr₁  r))
-                
-
-            pr₂proofEq : pr₂  (λ* prover  r)  pr₂  (stG  (pr₂  r))
-            pr₂proofEq = {!!}
-          return
-            ((subst  r'  r'   perX .equality  (x , x)) (sym pr₁proofEq) (stF⊩isStrictF x y (pr₁  r) pr₁r⊩Fxy .fst)) ,
-              subst  r'  r'   perZ .equality  (z , z)) (sym pr₂proofEq) (stG⊩isStrictG y z (pr₂  r) pr₂r⊩Gxy .snd)))))
-isRelational (composeFuncRel {X} {Y} {Z} perX perY perZ F G) =
-  do
-    (rlF , rlF⊩isRelationalF)  F .isRelational
-    (stF , stF⊩isStrictF)  F .isStrict
-    (rlG , rlG⊩isRelationalG)  G .isRelational
-    let
-      prover : ApplStrTerm as 1
-      prover =
-        ` pair ̇
-          (` rlF ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero)))))))) ̇
-          (` rlG ̇ (` pair ̇ (` pr₂ ̇ (` stF ̇ (` pr₁ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))))) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))))
-    return
-      (λ* prover ,
-       x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' 
-        do
-          (y , pr₁b⊩Fxy , pr₂b⊩Gyz)  b⊩∃y
-          let
-            proofEq : λ* prover  (pair  a  (pair  b  c))  pair  (rlF  (pair  a  (pair  (pr₁  b)  (pr₂  (stF  (pr₁  b))))))  (rlG  (pair  (pr₂  (stF  (pr₁  b)))  (pair  (pr₂  b)  c)))
-            proofEq =
-              λ* prover  (pair  a  (pair  b  c))
-                ≡⟨ λ*ComputationRule prover ((pair  a  (pair  b  c))  []) 
-              {!!}
-          return
-            (y ,
-            (subst  r  r   F .relation  (x' , y)) (sym {!!}) (rlF⊩isRelationalF x x' y y a (pr₁  b) (pr₂  (stF  (pr₁  b))) a⊩x~x' pr₁b⊩Fxy (stF⊩isStrictF x y (pr₁  b) pr₁b⊩Fxy .snd)) ,
-            (subst  r  r   G .relation  (y , z')) (sym {!!}) (rlG⊩isRelationalG y y z z' (pr₂  (stF  (pr₁  b))) (pr₂  b) c (stF⊩isStrictF x y (pr₁  b) pr₁b⊩Fxy .snd) pr₂b⊩Gyz c⊩z~z'))))))
-isSingleValued (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!}
-isTotal (composeFuncRel {X} {Y} {Z} perX perY perZ F G) = {!!}
-
-RTMorphism :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  Type _
-RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / λ F G  pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F
-
-idRTMorphism :  {X : Type ℓ'}  (perX : PartialEquivalenceRelation X)  RTMorphism perX perX
-idRTMorphism {X} perX = [ equalityFuncRel perX ]
-
-composeRTMorphism :
-   {X Y Z : Type ℓ'}
-   (perX : PartialEquivalenceRelation X)
-   (perY : PartialEquivalenceRelation Y)
-   (perZ : PartialEquivalenceRelation Z)
-   (f : RTMorphism perX perY)
-   (g : RTMorphism perY perZ)
-  ----------------------------------------
-   RTMorphism perX perZ
-composeRTMorphism {X} {Y} {Z} perX perY perZ f g =
-  setQuotRec2
-    squash/
-     F G  [ composeFuncRel perX perY perZ F G ])
-     { F F' G (F≤F' , F'≤F)  eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) ({!!} , {!!}) })
-     { F G G' (G≤G' , G'≤G)  eq/ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') ({!!} , {!!}) })
-    f g
-
-idLRTMorphism :
-   {X Y : Type ℓ'}
-   (perX : PartialEquivalenceRelation X)
-   (perY : PartialEquivalenceRelation Y)
-   (f : RTMorphism perX perY)
-   composeRTMorphism perX perX perY (idRTMorphism perX) f  f
-idLRTMorphism {X} {Y} perX perY f =
-  setQuotElimProp
-     f  squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f)
-     f  eq/ _ _ ({!!} , {!!}))
-    f
-
-idRRTMorphism :
-   {X Y : Type ℓ'}
-   (perX : PartialEquivalenceRelation X)
-   (perY : PartialEquivalenceRelation Y)
-   (f : RTMorphism perX perY)
-   composeRTMorphism perX perY perY f (idRTMorphism perY)  f
-idRRTMorphism {X} {Y} perX perY f =
-  setQuotElimProp
-     f  squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f)
-     f  eq/ _ _ ({!!} , {!!}))
-    f
-
-assocRTMorphism :
-   {X Y Z W : Type ℓ'}
-   (perX : PartialEquivalenceRelation X)
-   (perY : PartialEquivalenceRelation Y)
-   (perZ : PartialEquivalenceRelation Z)
-   (perW : PartialEquivalenceRelation W)
-   (f : RTMorphism perX perY)
-   (g : RTMorphism perY perZ)
-   (h : RTMorphism perZ perW)
-   composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h  composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h)
-assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h =
-  setQuotElimProp3
-     f g h 
-      squash/
-        (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h)
-        (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h)))
-     f g h  eq/ _ _ ({!!} , {!!}))
-    f g h
-
-RT : Category (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ'')))
-Category.ob RT = Σ[ X  Type ℓ' ] PartialEquivalenceRelation X
-Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism perX perY
-Category.id RT {X , perX} = idRTMorphism perX
-Category._⋆_ RT {X , perX} {y , perY} {Z , perZ} f g = composeRTMorphism perX perY perZ f g
-Category.⋆IdL RT {X , perX} {Y , perY} f = idLRTMorphism perX perY f
-Category.⋆IdR RT {X , perX} {Y , perY} f = idRRTMorphism perX perY f
-Category.⋆Assoc RT {X , perX} {Y , perY} {Z , perZ} {W , perW} f g h = assocRTMorphism perX perY perZ perW f g h
-Category.isSetHom RT = squash/
+Realizability.Topos.FunctionalRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData
+open import Cubical.Data.Fin hiding (Fin; _/_)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Relation.Binary
+
+module Realizability.Topos.FunctionalRelation
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
+
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+open PartialEquivalenceRelation
+
+module _
+  {X Y : Type ℓ'}
+  (perX : PartialEquivalenceRelation X)
+  (perY : PartialEquivalenceRelation Y)
+  (relation : Predicate (X × Y)) where
+  equalityX = perX .equality
+  equalityY = perY .equality
+  
+  realizesStrictDomain : A  Type _
+  realizesStrictDomain stD = (∀ x y r  r   relation  (x , y)  (stD  r)   equalityX  (x , x))
+
+  realizesStrictCodomain : A  Type _
+  realizesStrictCodomain stC = (∀ x y r  r   relation  (x , y)  (stC  r)   equalityY  (y , y))
+
+  realizesRelational : A  Type _
+  realizesRelational rel =
+        (∀ x x' y y' a b c
+         a   equalityX  (x , x')
+         b   relation  (x , y)
+         c   equalityY  (y , y')
+        ------------------------------------------
+         (rel  a  b  c)   relation  (x' , y'))
+
+  realizesSingleValued : A  Type _
+  realizesSingleValued sv =
+        (∀ x y y' r₁ r₂
+         r₁   relation  (x , y)
+         r₂   relation  (x , y')
+        -----------------------------------
+         (sv  r₁  r₂)   equalityY  (y , y'))
+
+  realizesTotal : A  Type _
+  realizesTotal tl =
+        (∀ x r  r   equalityX  (x , x)  ∃[ y  Y ] (tl  r)   relation  (x , y))
+    
+  record isFunctionalRelation : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+    constructor makeIsFunctionalRelation
+    field
+      isStrictDomain : ∃[ stD  A ] (realizesStrictDomain stD)
+      isStrictCodomain : ∃[ stC  A ] (realizesStrictCodomain stC)
+      isRelational : ∃[ rl  A ] (realizesRelational rl)
+      isSingleValued : ∃[ sv  A ] (realizesSingleValued sv)
+      isTotal : ∃[ tl  A ] (realizesTotal tl)
+
+record FunctionalRelation {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  constructor makeFunctionalRelation
+  field
+    relation : Predicate (X × Y)
+    isFuncRel : isFunctionalRelation perX perY relation
+  open isFunctionalRelation isFuncRel public
+  
+open FunctionalRelation
+
+pointwiseEntailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  (F G : FunctionalRelation perX perY)  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
+pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe  A ] (∀ x y r  r   F .relation  (x , y)  (pe  r)   G .relation  (x , y))
+
+-- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong
+-- Lemma 4.3.5
+opaque
+  F≤G→G≤F :
+     {X Y : Type ℓ'}
+     (perX : PartialEquivalenceRelation X)
+     (perY : PartialEquivalenceRelation Y)
+     (F G : FunctionalRelation perX perY)
+     pointwiseEntailment perX perY F G
+     pointwiseEntailment perX perY G F
+  F≤G→G≤F {X} {Y} perX perY F G F≤G =
+    do
+      (r , r⊩F≤G)  F≤G
+      (tlF , tlF⊩isTotalF)  F .isTotal
+      (svG , svG⊩isSingleValuedG)  G .isSingleValued
+      (rlF , rlF⊩isRelationalF)  F .isRelational
+      (stGD , stGD⊩isStrictDomainG)  G .isStrictDomain
+      let
+        prover : ApplStrTerm as 1
+        prover = ` rlF ̇ (` stGD ̇ # zero) ̇ (` tlF ̇ (` stGD ̇ # zero)) ̇ (` svG ̇ (` r ̇ (` tlF ̇ (` stGD ̇ # zero))) ̇ # zero)
+      return
+        (λ* prover ,
+         x y s s⊩Gxy 
+          subst
+             r'  r'   F .relation  (x , y))
+            (sym (λ*ComputationRule prover s))
+            (transport
+              (propTruncIdempotent (F .relation .isPropValued _ _))
+              (do
+                (y' , tlF⨾stGDs⊩Fxy')  tlF⊩isTotalF x (stGD  s) (stGD⊩isStrictDomainG x y s s⊩Gxy)
+                return
+                  (rlF⊩isRelationalF
+                    x x y' y
+                    (stGD  s) (tlF  (stGD  s)) (svG  (r  (tlF  (stGD  s)))  s)
+                    (stGD⊩isStrictDomainG x y s s⊩Gxy)
+                    tlF⨾stGDs⊩Fxy'
+                    (svG⊩isSingleValuedG x y' y (r  (tlF  (stGD  s))) s (r⊩F≤G x y' (tlF  (stGD  s)) tlF⨾stGDs⊩Fxy') s⊩Gxy))))))
+
+bientailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  FunctionalRelation perX perY  FunctionalRelation perX perY  Type _
+bientailment {X} {Y} perX perY F G = pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F
+
+isPropValuedBientailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  (F G : FunctionalRelation perX perY)  isProp (bientailment perX perY F G)
+isPropValuedBientailment {X} {Y} perX perY F G = isProp× isPropPropTrunc isPropPropTrunc
+
+RTMorphism :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  Type _
+RTMorphism {X} {Y} perX perY = FunctionalRelation perX perY / bientailment perX perY
+
+isEquivRelBientailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  BinaryRelation.isEquivRel (bientailment perX perY)
+BinaryRelation.isEquivRel.reflexive (isEquivRelBientailment {X} {Y} perX perY) =
+  λ A 
+   Id ,  x y r r⊩Axy  subst  r'  r'   A .relation  (x , y)) (sym (Ida≡a _)) r⊩Axy) ∣₁ ,
+   Id ,  x y r r⊩Axy  subst  r'  r'   A .relation  (x , y)) (sym (Ida≡a _)) r⊩Axy) ∣₁
+BinaryRelation.isEquivRel.symmetric (isEquivRelBientailment {X} {Y} perX perY) F G (F≤G , G≤F) = G≤F , F≤G
+BinaryRelation.isEquivRel.transitive (isEquivRelBientailment {X} {Y} perX perY) F G H (F≤G , G≤F) (G≤H , H≤G) =
+  let
+    answer =
+      do
+        (s , s⊩F≤G)  F≤G
+        (p , p⊩G≤H)  G≤H
+        let
+          prover : ApplStrTerm as 1
+          prover = ` p ̇ (` s ̇ # zero)
+        return
+          (λ* prover ,
+           x y r r⊩Fxy  subst  r'  r'   H .relation  (x , y)) (sym (λ*ComputationRule prover r)) (p⊩G≤H x y (s  r) (s⊩F≤G x y r r⊩Fxy))))
+  in
+  answer , F≤G→G≤F perX perY F H answer
+
+opaque
+  idFuncRel :  {X : Type ℓ'}  (perX : PartialEquivalenceRelation X)  FunctionalRelation perX perX
+  relation (idFuncRel {X} perX) = perX .equality
+  isFunctionalRelation.isStrictDomain (isFuncRel (idFuncRel {X} perX)) =
+    do
+      (s , s⊩isSymmetric)  perX .isSymmetric
+      (t , t⊩isTransitive)  perX .isTransitive
+      let
+        prover : ApplStrTerm as 1
+        prover = ` t ̇ # zero ̇ (` s ̇ # zero)
+      return
+        (λ* prover ,
+         λ x x' r r⊩x~x' 
+           subst
+              r'  r'   perX .equality  (x , x))
+             (sym (λ*ComputationRule prover r))
+             (t⊩isTransitive x x' x r (s  r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')))
+  isFunctionalRelation.isStrictCodomain (isFuncRel (idFuncRel {X} perX)) =
+    do
+      (s , s⊩isSymmetric)  perX .isSymmetric
+      (t , t⊩isTransitive)  perX .isTransitive
+      let
+        prover : ApplStrTerm as 1
+        prover = ` t ̇ (` s ̇ # zero) ̇ # zero
+      return
+        (λ* prover ,
+         x x' r r⊩x~x' 
+          subst
+             r'  r'   perX .equality  (x' , x'))
+            (sym (λ*ComputationRule prover r))
+            (t⊩isTransitive x' x x' (s  r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x')))
+  isFunctionalRelation.isRelational (isFuncRel (idFuncRel {X} perX)) =
+    do
+      (s , s⊩isSymmetric)  perX .isSymmetric
+      (t , t⊩isTransitive)  perX .isTransitive
+      let
+        prover : ApplStrTerm as 3
+        prover = ` t ̇ (` t ̇ (` s ̇ # two) ̇ # one) ̇ # zero
+      return
+        (λ*3 prover ,
+         x₁ x₂ x₃ x₄ a b c a⊩x₁~x₂ b⊩x₁~x₃ c⊩x₃~x₄ 
+          subst
+             r'  r'   perX .equality  (x₂ , x₄))
+            (sym (λ*3ComputationRule prover a b c))
+            (t⊩isTransitive x₂ x₃ x₄ (t  (s  a)  b) c (t⊩isTransitive x₂ x₁ x₃ (s  a) b (s⊩isSymmetric x₁ x₂ a a⊩x₁~x₂) b⊩x₁~x₃) c⊩x₃~x₄)))
+  isFunctionalRelation.isSingleValued (isFuncRel (idFuncRel {X} perX)) =
+    do
+      (s , s⊩isSymmetric)  perX .isSymmetric
+      (t , t⊩isTransitive)  perX .isTransitive
+      let
+        prover : ApplStrTerm as 2
+        prover = ` t ̇ (` s ̇ # one) ̇ # zero
+      return
+        (λ*2 prover ,
+         x₁ x₂ x₃ r₁ r₂ r₁⊩x₁~x₂ r₂⊩x₁~x₃ 
+          subst
+             r'  r'   perX .equality  (x₂ , x₃))
+            (sym (λ*2ComputationRule prover r₁ r₂))
+            (t⊩isTransitive x₂ x₁ x₃ (s  r₁) r₂ (s⊩isSymmetric x₁ x₂ r₁ r₁⊩x₁~x₂) r₂⊩x₁~x₃)))
+  isFunctionalRelation.isTotal (isFuncRel (idFuncRel {X} perX)) =
+    do
+      (s , s⊩isSymmetric)  perX .isSymmetric
+      (t , t⊩isTransitive)  perX .isTransitive
+      return
+        (Id ,
+         x r r⊩x~x   x , subst  r'  r'   perX .equality  (x , x)) (sym (Ida≡a _)) r⊩x~x ∣₁))
+
+idRTMorphism :  {X : Type ℓ'}  (perX : PartialEquivalenceRelation X)  RTMorphism perX perX
+idRTMorphism {X} perX = [ idFuncRel perX ]
+
+opaque
+  {-# TERMINATING #-} -- bye bye, type-checking with --safe 😔💔
+  composeFuncRel :
+     {X Y Z : Type ℓ'}
+     (perX : PartialEquivalenceRelation X)
+     (perY : PartialEquivalenceRelation Y)
+     (perZ : PartialEquivalenceRelation Z)
+     FunctionalRelation perX perY
+     FunctionalRelation perY perZ
+     FunctionalRelation perX perZ
+  isSetPredicateBase (relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = isSet× (perX .isSetX) (perZ .isSetX)
+   relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)  (x , z) r =
+    ∃[ y  Y ] (pr₁  r)   F .relation  (x , y) × (pr₂  r)   G .relation  (y , z)
+  isPropValued (relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) (x , z) r = isPropPropTrunc
+  isFunctionalRelation.isStrictDomain (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
+    do
+      (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
+      let
+        prover : ApplStrTerm as 1
+        prover = ` stFD ̇ (` pr₁ ̇ # zero)
+      return
+        (λ* prover ,
+         x z r r⊩∃y 
+          subst
+             r'  r'   perX .equality  (x , x))
+            (sym (λ*ComputationRule prover r))
+            (transport
+              (propTruncIdempotent (perX .equality .isPropValued _ _))
+              (do
+                (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
+                return (stFD⊩isStrictDomainF x y (pr₁  r) pr₁r⊩Fxy)))))
+  isFunctionalRelation.isStrictCodomain (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
+    do
+      (stGC , stGC⊩isStrictCodomainG)  G .isStrictCodomain
+      let
+        prover : ApplStrTerm as 1
+        prover = ` stGC ̇ (` pr₂ ̇ # zero)
+      return
+        (λ* prover ,
+         λ x z r r⊩∃y 
+           subst
+              r'  r'   perZ .equality  (z , z))
+             (sym (λ*ComputationRule prover r))
+             (transport
+               (propTruncIdempotent (perZ .equality .isPropValued _ _))
+               (do
+                 (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
+                 return (stGC⊩isStrictCodomainG y z (pr₂  r) pr₂r⊩Gyz))))
+  isFunctionalRelation.isRelational (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
+    do
+      (rlF , rlF⊩isRelationalF)  F .isRelational
+      (rlG , rlG⊩isRelationalG)  G .isRelational
+      (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
+      let
+        prover : ApplStrTerm as 3
+        prover = ` pair ̇ (` rlF ̇ # two ̇ (` pr₁ ̇ # one) ̇ (` stFC ̇ (` pr₁ ̇ # one))) ̇ (` rlG ̇ (` stFC ̇ (` pr₁ ̇ # one)) ̇ (` pr₂ ̇ # one) ̇ # zero)
+      return
+        (λ*3 prover ,
+         x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' 
+          do
+            (y , pr₁b⊩Fxy , pr₂b⊩Gyz)  b⊩∃y
+            let
+              pr₁proofEq : pr₁  (λ*3 prover  a  b  c)  rlF  a  (pr₁  b)  (stFC  (pr₁  b))
+              pr₁proofEq = cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _
+
+              pr₂proofEq : pr₂  (λ*3 prover  a  b  c)  rlG  (stFC  (pr₁  b))  (pr₂  b)  c
+              pr₂proofEq = cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _
+            return
+              (y ,
+               subst
+                  r'  r'   F .relation  (x' , y))
+                 (sym pr₁proofEq)
+                 (rlF⊩isRelationalF x x' y y a (pr₁  b) (stFC  (pr₁  b)) a⊩x~x' pr₁b⊩Fxy (stFC⊩isStrictCodomainF x y (pr₁  b) pr₁b⊩Fxy)) ,
+               subst
+                  r'  r'   G .relation  (y , z'))
+                 (sym pr₂proofEq)
+                 (rlG⊩isRelationalG y y z z' (stFC  (pr₁  b)) (pr₂  b) c (stFC⊩isStrictCodomainF x y (pr₁  b) pr₁b⊩Fxy) pr₂b⊩Gyz c⊩z~z'))))
+  isFunctionalRelation.isSingleValued (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
+    do
+      (svF , svF⊩isSingleValuedF)  F .isSingleValued
+      (svG , svG⊩isSingleValuedG)  G .isSingleValued
+      (relG , relG⊩isRelationalG)  G .isRelational
+      (stGC , stGC⊩isStrictCodomainG)  G .isStrictCodomain
+      let
+        prover : ApplStrTerm as 2
+        prover = ` svG ̇ (` pr₂ ̇ # one) ̇ (` relG ̇ (` svF ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ # one)) ̇ (` pr₂ ̇ # zero) ̇ (` stGC ̇ (` pr₂ ̇ # zero)))
+      return
+        (λ*2 prover ,
+         x z z' r₁ r₂ r₁⊩∃y r₂⊩∃y 
+          transport
+            (propTruncIdempotent (perZ .equality .isPropValued _ _))
+            (do
+              (y , pr₁r₁⊩Fxy , pr₂r₁⊩Gyz)  r₁⊩∃y
+              (y' , pr₁r₂⊩Fxy' , pr₂r₂⊩Gy'z')  r₂⊩∃y
+              return
+                (subst
+                   r'  r'   perZ .equality  (z , z'))
+                  (sym (λ*2ComputationRule prover r₁ r₂))
+                  (svG⊩isSingleValuedG
+                    y z z'
+                    (pr₂  r₁)
+                    (relG  (svF  (pr₁  r₂)  (pr₁  r₁))  (pr₂  r₂)  (stGC  (pr₂  r₂)))
+                    pr₂r₁⊩Gyz
+                    (relG⊩isRelationalG
+                      y' y z' z'
+                      (svF  (pr₁  r₂)  (pr₁  r₁))
+                      (pr₂  r₂)
+                      (stGC  (pr₂  r₂))
+                      (svF⊩isSingleValuedF x y' y (pr₁  r₂) (pr₁  r₁) pr₁r₂⊩Fxy' pr₁r₁⊩Fxy)
+                      pr₂r₂⊩Gy'z'
+                      (stGC⊩isStrictCodomainG y' z' (pr₂  r₂) pr₂r₂⊩Gy'z')))))))
+  isFunctionalRelation.isTotal (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
+    do
+      (tlF , tlF⊩isTotalF)  F .isTotal
+      (tlG , tlG⊩isTotalG)  G .isTotal
+      (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pair ̇ (` tlF ̇ # zero) ̇ (` tlG ̇ (` stFC ̇ (` tlF ̇ # zero)))
+      return
+        (λ* prover ,
+         x r r⊩x~x 
+          do
+            (y , ⊩Fxy)  tlF⊩isTotalF x r r⊩x~x
+            (z , ⊩Gyz)  tlG⊩isTotalG y (stFC  (tlF  r)) (stFC⊩isStrictCodomainF x y (tlF  r) ⊩Fxy)
+            return
+              (z ,
+              return
+                (y ,
+                ((subst  r'  r'   F .relation  (x , y)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) ⊩Fxy) ,
+                 (subst  r'  r'   G .relation  (y , z)) (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _)) ⊩Gyz))))))
+
+opaque
+  unfolding composeFuncRel
+  composeRTMorphism :
+     {X Y Z : Type ℓ'}
+     (perX : PartialEquivalenceRelation X)
+     (perY : PartialEquivalenceRelation Y)
+     (perZ : PartialEquivalenceRelation Z)
+     (f : RTMorphism perX perY)
+     (g : RTMorphism perY perZ)
+    ----------------------------------------
+     RTMorphism perX perZ
+  composeRTMorphism {X} {Y} {Z} perX perY perZ f g =
+    SQ.rec2
+      squash/
+       F G  [ composeFuncRel perX perY perZ F G ])
+       { F F' G (F≤F' , F'≤F) 
+        eq/ _ _
+          let answer = (do
+              (s , s⊩F≤F')  F≤F'
+              let
+                prover : ApplStrTerm as 1
+                prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+              return
+                (λ* prover ,
+                 x z r r⊩∃y 
+                  do
+                    (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
+                    return
+                      (y ,
+                       subst
+                          r'  r'   F' .relation  (x , y))
+                         (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                         (s⊩F≤F' x y (pr₁  r) pr₁r⊩Fxy) ,
+                       subst
+                          r'  r'   G .relation  (y , z))
+                         (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                         pr₂r⊩Gyz))))
+          in
+        (answer , F≤G→G≤F perX perZ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) answer) })
+       { F G G' (G≤G' , G'≤G) 
+        eq/ _ _
+          let answer = (do
+            (s , s⊩G≤G')  G≤G'
+            let
+              prover : ApplStrTerm as 1
+              prover = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero))
+            return
+              (λ* prover ,
+               x z r r⊩∃y 
+                 do
+                   (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
+
+                   return
+                     (y ,
+                      subst  r'  r'   F .relation  (x , y)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) pr₁r⊩Fxy ,
+                      subst  r'  r'   G' .relation  (y , z)) (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _)) (s⊩G≤G' y z (pr₂  r) pr₂r⊩Gyz)))))
+          in
+        (answer , F≤G→G≤F perX perZ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') answer) })
+      f g
+
+opaque
+  unfolding composeRTMorphism
+  unfolding idFuncRel
+  idLRTMorphism :
+     {X Y : Type ℓ'}
+     (perX : PartialEquivalenceRelation X)
+     (perY : PartialEquivalenceRelation Y)
+     (f : RTMorphism perX perY)
+     composeRTMorphism perX perX perY (idRTMorphism perX) f  f
+  idLRTMorphism {X} {Y} perX perY f =
+    SQ.elimProp
+       f  squash/ (composeRTMorphism perX perX perY (idRTMorphism perX) f) f)
+       F 
+        let
+          answer : pointwiseEntailment perX perY (composeFuncRel perX perX perY (idFuncRel perX) F) F
+          answer =
+            do
+              (relF , relF⊩isRelationalF)  F .isRelational
+              (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
+              (sX , sX⊩isSymmetricX)  perX .isSymmetric
+              let
+                prover : ApplStrTerm as 1
+                prover = ` relF ̇ (` sX ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero) ̇ (` stFC ̇ (` pr₂ ̇ # zero))
+              return
+                (λ* prover ,
+                  x y r r⊩∃x' 
+                   transport
+                     (propTruncIdempotent (F .relation .isPropValued _ _))
+                     (do
+                       (x' , pr₁r⊩x~x' , pr₂r⊩Fx'y)  r⊩∃x'
+                       return
+                         (subst
+                            r'  r'   F .relation  (x , y))
+                           (sym (λ*ComputationRule prover r))
+                           (relF⊩isRelationalF
+                             x' x y y
+                             (sX  (pr₁  r)) (pr₂  r) (stFC  (pr₂  r))
+                             (sX⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x')
+                             pr₂r⊩Fx'y
+                             (stFC⊩isStrictCodomainF x' y (pr₂  r) pr₂r⊩Fx'y))))))
+        in
+        eq/ _ _ (answer , F≤G→G≤F perX perY (composeFuncRel perX perX perY (idFuncRel perX) F) F answer))
+      f
+
+opaque
+  unfolding composeRTMorphism
+  unfolding idFuncRel
+  idRRTMorphism :
+     {X Y : Type ℓ'}
+     (perX : PartialEquivalenceRelation X)
+     (perY : PartialEquivalenceRelation Y)
+     (f : RTMorphism perX perY)
+     composeRTMorphism perX perY perY f (idRTMorphism perY)  f
+  idRRTMorphism {X} {Y} perX perY f =
+    SQ.elimProp
+       f  squash/ (composeRTMorphism perX perY perY f (idRTMorphism perY)) f)
+       F 
+        let
+          answer : pointwiseEntailment perX perY (composeFuncRel perX perY perY F (idFuncRel perY)) F
+          answer =
+            do
+              (relF , relF⊩isRelationalF)  F .isRelational
+              (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
+              let
+                prover : ApplStrTerm as 1
+                prover = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ # zero)
+              return
+                (λ* prover ,
+                 x y r r⊩∃y' 
+                  transport
+                    (propTruncIdempotent (F .relation .isPropValued _ _))
+                    (do
+                      (y' , pr₁r⊩Fxy' , pr₂r⊩y'~y)  r⊩∃y'
+                      return
+                        (subst
+                           r'  r'   F .relation  (x , y))
+                          (sym (λ*ComputationRule prover r))
+                          (relF⊩isRelationalF x x y' y (stFD  (pr₁  r)) (pr₁  r) (pr₂  r) (stFD⊩isStrictDomainF x y' (pr₁  r) pr₁r⊩Fxy') pr₁r⊩Fxy' pr₂r⊩y'~y)))))
+        in
+        eq/ _ _ (answer , F≤G→G≤F perX perY (composeFuncRel perX perY perY F (idFuncRel perY)) F answer))
+      f
+
+opaque
+  unfolding composeRTMorphism
+  assocRTMorphism :
+     {X Y Z W : Type ℓ'}
+     (perX : PartialEquivalenceRelation X)
+     (perY : PartialEquivalenceRelation Y)
+     (perZ : PartialEquivalenceRelation Z)
+     (perW : PartialEquivalenceRelation W)
+     (f : RTMorphism perX perY)
+     (g : RTMorphism perY perZ)
+     (h : RTMorphism perZ perW)
+     composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h  composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h)
+  assocRTMorphism {X} {Y} {Z} {W} perX perY perZ perW f g h =
+    SQ.elimProp3
+       f g h 
+        squash/
+          (composeRTMorphism perX perZ perW (composeRTMorphism perX perY perZ f g) h)
+          (composeRTMorphism perX perY perW f (composeRTMorphism perY perZ perW g h)))
+       F G H 
+        let
+          answer =
+            do
+              let
+                prover : ApplStrTerm as 1
+                prover = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero))
+              return
+                (λ* prover ,
+                 x w r r⊩∃z 
+                  transport
+                    (propTruncIdempotent isPropPropTrunc)
+                    (do
+                      (z , pr₁r⊩∃y , pr₂r⊩Hzw)  r⊩∃z
+                      (y , pr₁pr₁r⊩Fxy , pr₂pr₁r⊩Gyz)  pr₁r⊩∃y
+                      return
+                        (return
+                          (y ,
+                            (subst
+                               r'  r'   F .relation  (x , y))
+                              (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                              pr₁pr₁r⊩Fxy ,
+                            return
+                              (z ,
+                                ((subst
+                                   r'  r'   G .relation  (y , z))
+                                  (sym
+                                    (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r) 
+                                     cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                                  pr₂pr₁r⊩Gyz) ,
+                                 (subst
+                                   r'  r'   H .relation  (z , w))
+                                  (sym
+                                    (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r) 
+                                     cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                                  pr₂r⊩Hzw)))))))))
+        in
+        eq/ _ _
+          (answer ,
+           F≤G→G≤F
+             perX perW
+             (composeFuncRel perX perZ perW (composeFuncRel perX perY perZ F G) H)
+             (composeFuncRel perX perY perW F (composeFuncRel perY perZ perW G H))
+             answer))
+      f g h
+
+-- Very useful helper functions to prevent type-checking time from exploding
+opaque
+  [F]≡[G]→F≤G :  {X Y : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y}  (F G : FunctionalRelation perX perY)  [_] {R = bientailment perX perY} F  [_] {R = bientailment perX perY} G  pointwiseEntailment perX perY F G
+  [F]≡[G]→F≤G {X} {Y} {perX} {perY} F G iso = SQ.effective (isPropValuedBientailment perX perY) (isEquivRelBientailment perX perY) F G iso .fst
+
+  [F]≡[G]→G≤F :  {X Y : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y}  (F G : FunctionalRelation perX perY)  [_] {R = bientailment perX perY} F  [_] {R = bientailment perX perY} G  pointwiseEntailment perX perY G F
+  [F]≡[G]→G≤F {X} {Y} {perX} {perY} F G iso = SQ.effective (isPropValuedBientailment perX perY) (isEquivRelBientailment perX perY) F G iso .snd
+
+opaque
+  unfolding composeRTMorphism
+  [F]⋆[G]≡[H]⋆[I]→F⋆G≤H⋆I :  {X Y Z W : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z} {perW : PartialEquivalenceRelation W}  (F : FunctionalRelation perX perY) (G : FunctionalRelation perY perZ) (H : FunctionalRelation perX perW) (I : FunctionalRelation perW perZ)  composeRTMorphism _ _ _ [ F ] [ G ]  composeRTMorphism _ _ _ [ H ] [ I ]  pointwiseEntailment _ _ (composeFuncRel _ _ _ F G) (composeFuncRel _ _ _ H I)
+  [F]⋆[G]≡[H]⋆[I]→F⋆G≤H⋆I {X} {Y} {Z} {W} {perX} {perY} {perZ} {perW} F G H I iso =
+    SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) (composeFuncRel _ _ _ F G) (composeFuncRel _ _ _ H I) iso .fst
+
+opaque
+  unfolding composeRTMorphism
+  [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G :  {X Y Z W : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z} {perW : PartialEquivalenceRelation W}  (F : FunctionalRelation perX perY) (G : FunctionalRelation perY perZ) (H : FunctionalRelation perX perW) (I : FunctionalRelation perW perZ)  composeRTMorphism _ _ _ [ F ] [ G ]  composeRTMorphism _ _ _ [ H ] [ I ]  pointwiseEntailment _ _ (composeFuncRel _ _ _ H I) (composeFuncRel _ _ _ F G)
+  [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G {X} {Y} {Z} {W} {perX} {perY} {perZ} {perW} F G H I iso =
+    SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) (composeFuncRel _ _ _ F G) (composeFuncRel _ _ _ H I) iso .snd
+
+opaque
+  unfolding composeRTMorphism
+  [F]≡[G]⋆[H]→F≤G⋆H :  {X Y Z : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z}  (F : FunctionalRelation perX perZ) (G : FunctionalRelation perX perY) (H : FunctionalRelation perY perZ)  [ F ]  composeRTMorphism _ _ _ [ G ] [ H ]  pointwiseEntailment _ _ F (composeFuncRel _ _ _ G H)
+  [F]≡[G]⋆[H]→F≤G⋆H {X} {Y} {Z} {perX} {perY} {perZ} F G H iso =
+    SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) F (composeFuncRel _ _ _ G H) iso .fst
+
+opaque
+  unfolding composeRTMorphism
+  [F]≡[G]⋆[H]→G⋆H≤F :  {X Y Z : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z}  (F : FunctionalRelation perX perZ) (G : FunctionalRelation perX perY) (H : FunctionalRelation perY perZ)  [ F ]  composeRTMorphism _ _ _ [ G ] [ H ]  pointwiseEntailment _ _ (composeFuncRel _ _ _ G H) F
+  [F]≡[G]⋆[H]→G⋆H≤F {X} {Y} {Z} {perX} {perY} {perZ} F G H iso = SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) F (composeFuncRel _ _ _ G H) iso .snd
+
+RT : Category (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ'')))
+Category.ob RT = Σ[ X  Type ℓ' ] PartialEquivalenceRelation X
+Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism perX perY
+Category.id RT {X , perX} = idRTMorphism perX
+Category._⋆_ RT {X , perX} {y , perY} {Z , perZ} f g = composeRTMorphism perX perY perZ f g
+Category.⋆IdL RT {X , perX} {Y , perY} f = idLRTMorphism perX perY f
+Category.⋆IdR RT {X , perX} {Y , perY} f = idRRTMorphism perX perY f
+Category.⋆Assoc RT {X , perX} {Y , perY} {Z , perZ} {W , perW} f g h = assocRTMorphism perX perY perZ perW f g h
+Category.isSetHom RT = squash/
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.MonicReprFuncRel.html b/docs/Realizability.Topos.MonicReprFuncRel.html new file mode 100644 index 0000000..5c0622c --- /dev/null +++ b/docs/Realizability.Topos.MonicReprFuncRel.html @@ -0,0 +1,280 @@ + +Realizability.Topos.MonicReprFuncRel
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Foundations.HLevels
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData
+open import Cubical.Data.Fin hiding (Fin; _/_)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Morphism
+open import Cubical.Relation.Binary
+
+module Realizability.Topos.MonicReprFuncRel
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
+
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+open import Realizability.Topos.Equalizer {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+open import Realizability.Topos.BinProducts {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+open PartialEquivalenceRelation
+open FunctionalRelation
+open Category RT
+
+-- Monics in RT
+module _ {X Y : Type ℓ'} (perX : PartialEquivalenceRelation X) (perY : PartialEquivalenceRelation Y) (F : FunctionalRelation perX perY) where
+
+  opaque
+    isInjectiveFuncRel : Type (ℓ-max  (ℓ-max ℓ' ℓ''))
+    isInjectiveFuncRel =
+      ∃[ inj  A ] (∀ x x' y r₁ r₂  r₁   F .relation  (x , y)  r₂   F .relation  (x' , y)  (inj  r₁  r₂)   perX .equality  (x , x'))
+
+  opaque
+    unfolding isInjectiveFuncRel
+    isPropIsInjectiveFuncRel : isProp isInjectiveFuncRel
+    isPropIsInjectiveFuncRel = isPropPropTrunc
+
+  -- This is the easier part
+  -- Essentially just a giant realizer that uses the injectivity
+  opaque
+    unfolding composeRTMorphism
+    unfolding composeFuncRel
+    unfolding isInjectiveFuncRel
+    isInjectiveFuncRel→isMonic : isInjectiveFuncRel  isMonic RT [ F ]
+    isInjectiveFuncRel→isMonic isInjectiveF {Z , perZ} {a} {b} a⋆[F]≡b⋆[F] =
+      SQ.elimProp2
+        {P = λ a b  composeRTMorphism _ _ _ a [ F ]  composeRTMorphism _ _ _ b [ F ]  a  b}
+         a b  isPropΠ λ _  squash/ a b)
+         A B A⋆F≡B⋆F 
+          let
+            (p , q) = SQ.effective (isPropValuedBientailment perZ perY) (isEquivRelBientailment perZ perY) _ _ A⋆F≡B⋆F
+            answer =
+              do
+                (p , p⊩A⋆F≤B⋆F)  p
+                (stCA , stCA⊩isStrictCodomainA)  A .isStrictCodomain
+                (stDA , stDA⊩isStrictDomainA)  A .isStrictDomain
+                (tlF , tlF⊩isTotalF)  F .isTotal
+                (relB , relB⊩isRelationalB)  B .isRelational
+                (injF , injF⊩isInjectiveF)  isInjectiveF
+                let
+                  realizer : ApplStrTerm as 1
+                  realizer =
+                    ` relB ̇ (` stDA ̇ # zero) ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCA ̇ # zero))))) ̇
+                      (` injF ̇ (` pr₂ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCA ̇ # zero))))) ̇
+                      (` tlF ̇ (` stCA ̇ # zero)))
+                return
+                  (λ* realizer ,
+                   z x r r⊩Azx 
+                    transport
+                      (propTruncIdempotent (B .relation .isPropValued _ _))
+                      (do
+                        let
+                          x~x = stCA⊩isStrictCodomainA z x r r⊩Azx
+                          z~z = stDA⊩isStrictDomainA z x r r⊩Azx
+                        (y , ⊩Fxy)  tlF⊩isTotalF x _ x~x
+                        (x' , ⊩Bzx' , ⊩Fx'y)  
+                          p⊩A⋆F≤B⋆F
+                            z y
+                            (pair  r  (tlF  (stCA  r)))
+                            (return
+                              (x ,
+                              subst  r'  r'   A .relation  (z , x)) (sym (pr₁pxy≡x _ _)) r⊩Azx ,
+                              subst  r'  r'   F .relation  (x , y)) (sym (pr₂pxy≡y _ _)) ⊩Fxy))
+                        let
+                          x'~x = injF⊩isInjectiveF x' x y _ _ ⊩Fx'y ⊩Fxy -- this is the only place where we actually need the injectivity
+                        return
+                          (subst
+                             r'  r'   B .relation  (z , x))
+                            (sym (λ*ComputationRule realizer r))
+                            (relB⊩isRelationalB z z x' x _ _ _ z~z ⊩Bzx' x'~x)))))
+          in
+          eq/ A B (answer , F≤G→G≤F perZ perX A B answer))
+        a b
+        a⋆[F]≡b⋆[F]
+
+  opaque
+    unfolding binProdPr₁RT
+    unfolding binProdPr₁FuncRel
+    unfolding binProdPr₂FuncRel
+    unfolding equalizerMorphism
+    unfolding composeRTMorphism
+
+    π₁ : FunctionalRelation (binProdObRT perX perX) perX
+    π₁ = binProdPr₁FuncRel perX perX
+
+    π₂ : FunctionalRelation (binProdObRT perX perX) perX
+    π₂ = binProdPr₂FuncRel perX perX
+
+    kernelPairEqualizerFuncRel :
+      FunctionalRelation
+        (equalizerPer -- hehe
+          (binProdObRT perX perX) perY
+          ([ π₁ ]  [ F ])
+          ([ π₂ ]  [ F ])
+          (composeFuncRel _ _ _ π₁ F)
+          (composeFuncRel _ _ _ π₂ F))
+        (binProdObRT perX perX)
+    kernelPairEqualizerFuncRel =
+      equalizerFuncRel _ _
+        ((binProdPr₁RT perX perX)  [ F ])
+        ((binProdPr₂RT perX perX)  [ F ])
+        (composeFuncRel _ _ _ (binProdPr₁FuncRel perX perX) F)
+        (composeFuncRel _ _ _ (binProdPr₂FuncRel perX perX) F)
+
+    kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ :
+      composeRTMorphism _ _ _ [ kernelPairEqualizerFuncRel ] (composeRTMorphism _ _ _ [ π₁ ] [ F ])  composeRTMorphism _ _ _ [ kernelPairEqualizerFuncRel ] (composeRTMorphism _ _ _ [ π₂ ] [ F ])
+    kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ =
+      inc⋆f≡inc⋆g
+        (binProdObRT perX perX) perY
+        (composeRTMorphism _ _ _ [ π₁ ] [ F ])
+        (composeRTMorphism _ _ _ [ π₂ ] [ F ])
+        (composeFuncRel _ _ _ π₁ F)
+        (composeFuncRel _ _ _ π₂ F)
+
+    mainKernelPairEquation : ([ kernelPairEqualizerFuncRel ]  [ π₁ ])  [ F ]  ([ kernelPairEqualizerFuncRel ]  [ π₂ ])  [ F ]
+    mainKernelPairEquation =
+      ([ kernelPairEqualizerFuncRel ]  [ π₁ ])  [ F ]
+        ≡⟨ ⋆Assoc [ kernelPairEqualizerFuncRel ] [ π₁ ] [ F ]   -- Agda cannot solve these as constraints 😔
+      [ kernelPairEqualizerFuncRel ]  ([ π₁ ]  [ F ])
+        ≡⟨ kernelPairEqualizer⋆π₁≡kernelPairEqualizer⋆π₂ 
+      [ kernelPairEqualizerFuncRel ]  ([ π₂ ]  [ F ])
+        ≡⟨ sym (⋆Assoc [ kernelPairEqualizerFuncRel ] [ π₂ ] [ F ]) 
+      ([ kernelPairEqualizerFuncRel ]  [ π₂ ])  [ F ]
+        
+
+  opaque
+    unfolding isInjectiveFuncRel
+    unfolding composeRTMorphism
+    unfolding kernelPairEqualizerFuncRel
+    unfolding equalizerFuncRel
+    unfolding equalizerPer
+    unfolding binProdPr₁RT
+    unfolding binProdPr₂FuncRel
+    isMonic→isInjectiveFuncRel : isMonic RT [ F ]  isInjectiveFuncRel
+    isMonic→isInjectiveFuncRel isMonicF =
+      do
+        let
+          equation = isMonicF {a = [ kernelPairEqualizerFuncRel ]  [ π₁ ]} {a' = [ kernelPairEqualizerFuncRel ]  [ π₂ ]} mainKernelPairEquation
+          (p , q) = SQ.effective (isPropValuedBientailment _ _) (isEquivRelBientailment _ _) _ _ equation
+        (p , p⊩kπ₁≤kπ₂)  p
+        (q , q⊩kπ₂≤kπ₁)  q
+        (stCF , stCF⊩isStrictCodomainF)  F .isStrictCodomain
+        (stDF , stDF⊩isStrictDomainF)  F .isStrictDomain
+        (s , s⊩isSymmetricEqX)  perX .isSymmetric
+        (t , t⊩isTransitiveEqX)  perX .isTransitive
+        let
+          cursed : ApplStrTerm as 2
+          cursed =
+             (` pair ̇
+                  (` pair ̇
+                    (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)) ̇
+                    (` pair ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)) ̇ # one) ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # zero) ̇ (` stDF ̇ # one)) ̇ # zero))) ̇
+                  (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)))
+          realizer : ApplStrTerm as 2
+          realizer = ` t ̇ (` s ̇ (` pr₁ ̇ (` pr₂ ̇ (` p ̇ cursed)))) ̇ (` s ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₁ ̇ (` p ̇ cursed)))))
+        return
+          (λ*2 realizer ,
+           x₁ x₂ y r₁ r₂ r₁⊩Fx₁y r₂⊩Fx₂y 
+            let
+              x₁~x₁ = stDF⊩isStrictDomainF x₁ y r₁ r₁⊩Fx₁y
+              x₂~x₂ = stDF⊩isStrictDomainF x₂ y r₂ r₂⊩Fx₂y
+              foo =
+                p⊩kπ₁≤kπ₂ (x₁ , x₂) x₁
+                (pair 
+                  (pair 
+                    (pair  (stDF  r₁)  (stDF  r₂)) 
+                    (pair  (pair  (pair  (stDF  r₁)  (stDF  r₂))  r₁)  (pair  (pair  (stDF  r₂)  (stDF  r₁))  r₂))) 
+                  (pair  (stDF  r₁)  (stDF  r₂)))
+                (return
+                  (((x₁ , x₂)) ,
+                  ((subst  r'  r'   perX .equality  (x₁ , x₁)) (sym (cong  x  pr₁  (pr₁  x)) (pr₁pxy≡x _ _)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _)) x₁~x₁ ,
+                    subst  r'  r'   perX .equality  (x₂ , x₂)) (sym (cong  x  pr₂  (pr₁  x)) (pr₁pxy≡x _ _)  cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _)) x₂~x₂) ,
+                 return
+                  (y ,
+                    return
+                      (x₁ ,
+                        (subst  r'  r'   perX .equality  (x₁ , x₁))
+                          (sym
+                            (cong  x  pr₁  (pr₁  (pr₁  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  (pr₁  (pr₁  x))) (pr₂pxy≡y _ _) 
+                             cong  x  pr₁  (pr₁  x)) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  x) (pr₁pxy≡x _ _) 
+                             pr₁pxy≡x _ _))
+                          x₁~x₁ ,
+                         subst  r'  r'   perX .equality  (x₂ , x₂))
+                           (sym
+                             (cong  x  pr₂  (pr₁  (pr₁  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₁  (pr₁  x))) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  (pr₁  x)) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                              pr₂pxy≡y _ _))
+                           x₂~x₂) ,
+                         subst  r'  r'   F .relation  (x₁ , y))
+                           (sym
+                             (cong  x  pr₂  (pr₁  (pr₂  x))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₁  x)) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                              pr₂pxy≡y _ _))
+                           r₁⊩Fx₁y) ,
+                    return
+                      (x₂ ,
+                        (subst  r'  r'   perX .equality  (x₂ , x₂))
+                          (sym
+                            (cong  x  pr₁  (pr₁  (pr₂  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  (pr₁  (pr₂  x))) (pr₂pxy≡y _ _) 
+                             cong  x  pr₁  (pr₁  x)) (pr₂pxy≡y _ _) 
+                             cong  x  pr₁  x) (pr₁pxy≡x _ _) 
+                             pr₁pxy≡x _ _))
+                          x₂~x₂ ,
+                         subst  r'  r'   perX .equality  (x₁ , x₁))
+                           (sym
+                             (cong  x  pr₂  (pr₁  (pr₂  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₁  (pr₂  x))) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  (pr₁  x)) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                              pr₂pxy≡y _ _))
+                           x₁~x₁) ,
+                         subst  r'  r'   F .relation  (x₂ , y))
+                           (sym
+                             (cong  x  pr₂  (pr₂  (pr₂  x))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₂  x)) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                              pr₂pxy≡y _ _)) r₂⊩Fx₂y))) ,
+                         subst  r'  r'   perX .equality  (x₁ , x₁)) (sym (cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _)) x₁~x₁ ,
+                         subst  r'  r'   perX .equality  (x₂ , x₂)) (sym (cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _)) x₂~x₂))
+            in
+            transport
+              (propTruncIdempotent (perX .equality .isPropValued _ _))
+              (do
+                ((x₁' , x₂') , ((x₁~x₁' , x₂~x₂') , kp2) , p2)  foo
+                (y' , bar1 , bar2)  kp2
+                (x₁'' , (x₁~x₁'' , x₂~'x₂) , ⊩Fx₁''y')  bar1
+                (x₂'' , (x₂~x₂'' , x₁~'x₁) , ⊩Fx₂''y')  bar2
+                let
+                  (x₂'~x₁ , foo') = p2
+                return
+                  (subst
+                     r'  r'   perX .equality  (x₁ , x₂))
+                    (sym (λ*2ComputationRule realizer r₁ r₂))
+                    (t⊩isTransitiveEqX x₁ x₂' x₂ _ _ (s⊩isSymmetricEqX x₂' x₁ _ x₂'~x₁) (s⊩isSymmetricEqX x₂ x₂' _ x₂~x₂'))))))
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.Object.html b/docs/Realizability.Topos.Object.html index b87ceb6..1737fdc 100644 --- a/docs/Realizability.Topos.Object.html +++ b/docs/Realizability.Topos.Object.html @@ -1,41 +1,80 @@ -Realizability.Topos.Object
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+Realizability.Topos.Object
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.Structure
-open import Cubical.Data.Vec
-open import Cubical.Data.Nat
-open import Cubical.Data.FinData renaming (zero to fzero)
-open import Cubical.Data.Sigma
-open import Cubical.Data.Empty
-
-module Realizability.Topos.Object
-  { ℓ' ℓ''}
-  {A : Type }
-  (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
-  where
-
-open import Realizability.Tripos.Logic.Syntax { = ℓ'}
-open import Realizability.Tripos.Logic.Semantics {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate')
-open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Predicate' renaming (isSetX to isSetPredicateBase)
-open PredicateProperties
-open Morphism
-
-private
-  Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''}
-
-record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
-  field
-    isSetX : isSet X
-    equality : Predicate (X × X)
-    isSymmetric : ∃[ s  A ] (∀ x y r  r   equality  (x , y)  (s  r)   equality  (y , x))
-  open PredicateProperties {ℓ'' = ℓ''} (X × X)
-  field
-    isTransitive : ∃[ t  A ] (∀ x y z a b  a   equality  (x , y)  b   equality  (y , z)  (t  (pair  a  b))   equality  (x , z))
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Equiv
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.Reflection.RecordEquiv
+
+module Realizability.Topos.Object
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
   
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate renaming (isSetX to isSetPredicateBase)
+open PredicateProperties
+open Morphism
+
+record isPartialEquivalenceRelation (X : Type ℓ') (equality : Predicate (X × X)) : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  field
+    isSetX : isSet X
+    isSymmetric : ∃[ s  A ] (∀ x y r  r   equality  (x , y)  (s  r)   equality  (y , x))
+    isTransitive : ∃[ t  A ] (∀ x y z a b  a   equality  (x , y)  b   equality  (y , z)  (t  a  b)   equality  (x , z))
+
+open isPartialEquivalenceRelation
+isPropIsPartialEquivalenceRelation :  {X : Type ℓ'}  (equality : Predicate (X × X))  isProp (isPartialEquivalenceRelation X equality)
+isPropIsPartialEquivalenceRelation {X} equality x y i =
+  record { isSetX = isProp→PathP  i  isPropIsSet) (x .isSetX) (y .isSetX) i ; isSymmetric = squash₁ (x .isSymmetric) (y .isSymmetric) i ; isTransitive = squash₁ (x .isTransitive) (y .isTransitive) i }
+
+record PartialEquivalenceRelation (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  field
+    equality : Predicate (X × X)
+    isPerEquality : isPartialEquivalenceRelation X equality
+  open isPartialEquivalenceRelation isPerEquality public
+
+-- Directly from previous commit
+unquoteDecl PartialEquivalenceRelationIsoΣ = declareRecordIsoΣ PartialEquivalenceRelationIsoΣ (quote PartialEquivalenceRelation)
+
+PartialEquivalenceRelationΣ : (X : Type ℓ')  Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ''))
+PartialEquivalenceRelationΣ X = Σ[ equality  Predicate (X × X) ] isPartialEquivalenceRelation X equality
+
+open PartialEquivalenceRelation
+module _ (X : Type ℓ') where opaque
+  open Iso
+  PartialEquivalenceRelationΣ≡ : (perA perB : PartialEquivalenceRelationΣ X)  perA .fst  perB .fst  perA  perB
+  PartialEquivalenceRelationΣ≡ perA perB predicateEq = Σ≡Prop  x  isPropIsPartialEquivalenceRelation x) predicateEq 
+
+  PartialEquivalenceRelationΣ≃ : (perA perB : PartialEquivalenceRelationΣ X)  (perA .fst  perB .fst)  (perA  perB)
+  PartialEquivalenceRelationΣ≃ perA perB = Σ≡PropEquiv λ x  isPropIsPartialEquivalenceRelation x
+
+  PartialEquivalenceRelationIso : (perA perB : PartialEquivalenceRelation X)  Iso (Iso.fun PartialEquivalenceRelationIsoΣ perA  Iso.fun PartialEquivalenceRelationIsoΣ perB) (perA  perB)
+  Iso.fun (PartialEquivalenceRelationIso perA perB) p i = Iso.inv PartialEquivalenceRelationIsoΣ (p i)
+  inv (PartialEquivalenceRelationIso perA perB) = cong  x  Iso.fun PartialEquivalenceRelationIsoΣ x)
+  rightInv (PartialEquivalenceRelationIso perA perB) b = refl
+  leftInv (PartialEquivalenceRelationIso perA perB) a = refl
+
+  -- Main SIP
+  PartialEquivalenceRelation≃ : (perA perB : PartialEquivalenceRelation X)  (perA .equality  perB .equality)  (perA  perB)
+  PartialEquivalenceRelation≃ perA perB =
+    perA .equality  perB .equality
+      ≃⟨ idEquiv (perA .equality  perB .equality) 
+    Iso.fun PartialEquivalenceRelationIsoΣ perA .fst  Iso.fun PartialEquivalenceRelationIsoΣ perB .fst
+      ≃⟨ PartialEquivalenceRelationΣ≃ (Iso.fun PartialEquivalenceRelationIsoΣ perA) (Iso.fun PartialEquivalenceRelationIsoΣ perB) 
+    Iso.fun PartialEquivalenceRelationIsoΣ perA  Iso.fun PartialEquivalenceRelationIsoΣ perB
+      ≃⟨ isoToEquiv (PartialEquivalenceRelationIso perA perB) 
+    perA  perB
+      
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.ResizedPredicate.html b/docs/Realizability.Topos.ResizedPredicate.html new file mode 100644 index 0000000..0a56df2 --- /dev/null +++ b/docs/Realizability.Topos.ResizedPredicate.html @@ -0,0 +1,94 @@ + +Realizability.Topos.ResizedPredicate
-- Before we can talk about power objects in RT
+-- we need to use propositional resizing to get
+-- a copy of A-valued predicates in Type ℓ'
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Equiv
+open import Cubical.Data.Empty
+open import Cubical.Data.Sigma
+open import Realizability.PropResizing
+open import Realizability.CombinatoryAlgebra
+
+module Realizability.Topos.ResizedPredicate
+  {}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (resizing : hPropResizing )
+  where
+
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = } {ℓ'' = } ca
+open import Realizability.Topos.Object {ℓ' = } {ℓ'' = } ca isNonTrivial
+
+open CombinatoryAlgebra ca
+open Predicate renaming (isSetX to isSetPredicateBase)
+
+smallHProp = resizing .fst
+smallHProp≃hProp = resizing .snd
+
+ResizedPredicate : Type   Type 
+ResizedPredicate X = Σ[ rel  (X  A  smallHProp) ] isSet X
+
+PredicateΣ≃ResizedPredicate :  X  PredicateΣ X  ResizedPredicate X
+PredicateΣ≃ResizedPredicate X =
+  Σ-cong-equiv-prop
+    (equivΠ
+      (idEquiv X)
+       x 
+        equivΠ
+          (idEquiv A)
+          λ a 
+            smallHProp≃hProp))
+     _  isPropIsSet)
+     _  isPropIsSet)
+     _ answer  answer)
+     _ answer  answer)
+
+Predicate≃ResizedPredicate :  X  Predicate X  ResizedPredicate X
+Predicate≃ResizedPredicate X = compEquiv (Predicate≃PredicateΣ X) (PredicateΣ≃ResizedPredicate X)
+
+isSetResizedPredicate :  {X}  isSet (ResizedPredicate X)
+isSetResizedPredicate {X} = isOfHLevelRespectEquiv 2 (Predicate≃ResizedPredicate X) (isSetPredicate X)
+
+ResizedPredicate≃Predicate :  X  ResizedPredicate X  Predicate X
+ResizedPredicate≃Predicate X = invEquiv (Predicate≃ResizedPredicate X)
+
+toPredicate :  {X}  ResizedPredicate X  Predicate X
+toPredicate {X} ϕ = equivFun (ResizedPredicate≃Predicate X) ϕ
+
+fromPredicate :  {X}  Predicate X  ResizedPredicate X
+fromPredicate {X} ϕ = equivFun (Predicate≃ResizedPredicate X) ϕ
+
+compIsIdEquiv :  X  compEquiv (Predicate≃ResizedPredicate X) (ResizedPredicate≃Predicate X)  idEquiv (Predicate X)
+compIsIdEquiv X = invEquiv-is-rinv (Predicate≃ResizedPredicate X)
+
+compIsIdFunc :  {X}  (p : Predicate X)  toPredicate (fromPredicate p)  p
+compIsIdFunc {X} p i = equivFun (compIsIdEquiv X i) p
+
+module ResizedPredicateProps {X} (perX : PartialEquivalenceRelation X) where
+  open PartialEquivalenceRelation
+
+  entailmentResizedPredicate :  (ϕ ψ : ResizedPredicate X)  A  Type 
+  entailmentResizedPredicate ϕ ψ r =  (x : X) (a : A) (⊩ϕx : a   toPredicate ϕ  x)  (r  a)   toPredicate ψ  x
+
+  isPropEntailmentResizedPredicate :  ϕ ψ a  isProp (entailmentResizedPredicate ϕ ψ a)
+  isPropEntailmentResizedPredicate ϕ ψ a =
+    isPropΠ λ x  isPropΠ λ b  isPropΠ λ _  (toPredicate ψ) .isPropValued _ _
+
+  isStrictResizedPredicate :  (ϕ : ResizedPredicate X)  A  Type 
+  isStrictResizedPredicate ϕ r =  (x : X) (a : A) (⊩ϕx : a   toPredicate ϕ  x)  (r  a)   perX .equality  (x , x)
+
+  isPropIsStrictResizedPredicate :  ϕ r  isProp (isStrictResizedPredicate ϕ r)
+  isPropIsStrictResizedPredicate ϕ r =
+    isPropΠ λ x  isPropΠ λ a  isPropΠ λ _  perX .equality .isPropValued _ _
+
+  isRelationalResizedPredicate :  (ϕ : ResizedPredicate X)  A  Type 
+  isRelationalResizedPredicate ϕ r =
+     (x x' : X) (a b : A) (⊩x~x' : a   perX .equality  (x , x')) (⊩ϕx : b   toPredicate ϕ  x)  (r  a  b)   toPredicate ϕ  x'
+
+  isPropIsRelationalResizedPredicate :  ϕ r  isProp (isRelationalResizedPredicate ϕ r)
+  isPropIsRelationalResizedPredicate ϕ r =
+    isPropΠ λ x  isPropΠ λ x'  isPropΠ λ a  isPropΠ λ b  isPropΠ λ _  isPropΠ λ _  toPredicate ϕ .isPropValued _ _
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.StrictRelation.html b/docs/Realizability.Topos.StrictRelation.html new file mode 100644 index 0000000..5594652 --- /dev/null +++ b/docs/Realizability.Topos.StrictRelation.html @@ -0,0 +1,634 @@ + +Realizability.Topos.StrictRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Equiv
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData
+open import Cubical.Data.Fin hiding (Fin; _/_)
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Morphism
+open import Cubical.Categories.Constructions.SubObject
+open import Cubical.Categories.Constructions.Slice
+open import Cubical.Relation.Binary
+
+module Realizability.Topos.StrictRelation
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  where
+
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
+open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+open import Realizability.Topos.Equalizer {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+open import Realizability.Topos.BinProducts {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+open import Realizability.Topos.MonicReprFuncRel {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate renaming (isSetX to isSetPredicateBase)
+open Morphism
+open PartialEquivalenceRelation
+open FunctionalRelation
+open Category RT
+
+record isStrictRelation {X : Type ℓ'} (perX : PartialEquivalenceRelation X) (ϕ : Predicate X) : Type (ℓ-max  (ℓ-max ℓ' ℓ'')) where
+  field
+    isStrict : ∃[ st  A ] (∀ x r  r   ϕ  x  (st  r)   perX .equality  (x , x))
+    isRelational : ∃[ rel  A ] (∀ x x' r s  r   ϕ  x  s   perX .equality  (x , x')  (rel  r  s)   ϕ  x')
+
+record StrictRelation {X : Type ℓ'} (perX : PartialEquivalenceRelation X) : Type (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) where
+  field
+    predicate : Predicate X
+    isStrictRelationPredicate : isStrictRelation perX predicate
+  open isStrictRelation isStrictRelationPredicate public
+
+-- Every strict relation induces a subobject
+module InducedSubobject {X : Type ℓ'} (perX : PartialEquivalenceRelation X) (ϕ : StrictRelation perX) where
+  open StrictRelation
+  -- the subobject induced by ϕ
+  {-# TERMINATING #-}
+  subPer : PartialEquivalenceRelation X
+  Predicate.isSetX (equality subPer) = isSet× (perX .isSetX) (perX .isSetX)
+   equality subPer  (x , x') r = (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   ϕ .predicate  x
+  isPropValued (equality subPer) (x , x') r = isProp× (perX .equality .isPropValued _ _) (ϕ .predicate .isPropValued _ _)
+  isPartialEquivalenceRelation.isSetX (isPerEquality subPer) = perX .isSetX
+  isPartialEquivalenceRelation.isSymmetric (isPerEquality subPer) =
+    do
+      -- Trivial : use symmetry of ~X and relationality of ϕ
+      (s , s⊩isSymmetricX)  perX .isSymmetric
+      (relϕ , relϕ⊩isRelationalϕ)  ϕ .isRelational
+      let
+        realizer : ApplStrTerm as 1
+        realizer = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` relϕ ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero))
+      return
+        (λ* realizer ,
+         { x x' r (pr₁r⊩x~x' , pr₂r⊩ϕx) 
+          subst
+             r'  r'   perX .equality  (x' , x))
+            (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+            (s⊩isSymmetricX x x' _ pr₁r⊩x~x') ,
+          subst
+             r'  r'   ϕ .predicate  x')
+            (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+            (relϕ⊩isRelationalϕ x x' _ _ pr₂r⊩ϕx pr₁r⊩x~x') }))
+  isPartialEquivalenceRelation.isTransitive (isPerEquality subPer) =
+    do
+      (t , t⊩isTransitiveX)  perX .isTransitive
+      (relϕ , relϕ⊩isRelationalϕ)  ϕ .isRelational
+      let
+        realizer : ApplStrTerm as 2
+        realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # one)
+      return
+        (λ*2 realizer ,
+         { x₁ x₂ x₃ a b (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₂~x₃ , ⊩ϕx₂) 
+          subst
+             r'  r'   perX .equality  (x₁ , x₃))
+            (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer a b)  pr₁pxy≡x _ _))
+            (t⊩isTransitiveX x₁ x₂ x₃ _ _ ⊩x₁~x₂ ⊩x₂~x₃) ,
+          subst
+             r'  r'   ϕ .predicate  x₁)
+            (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer a b)  pr₂pxy≡y _ _))
+            ⊩ϕx₁ }))
+
+  opaque
+    unfolding idFuncRel
+    {-# TERMINATING #-}
+    incFuncRel : FunctionalRelation subPer perX
+    Predicate.isSetX (relation incFuncRel) = isSet× (perX .isSetX) (perX .isSetX)
+    Predicate.∣ relation incFuncRel  (x , x') r = (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   ϕ .predicate  x
+    Predicate.isPropValued (relation incFuncRel) (x , x') r = isProp× (perX .equality .isPropValued _ _) (ϕ .predicate .isPropValued _ _)
+    isFunctionalRelation.isStrictDomain (isFuncRel incFuncRel) =
+      do
+        (stD , stD⊩isStrictDomain)  idFuncRel perX .isStrictDomain
+        let
+          realizer : ApplStrTerm as 1
+          realizer = ` pair ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+        return
+          (λ* realizer ,
+           { x x' r (⊩x~x' , ⊩ϕx) 
+            (subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (stD⊩isStrictDomain x x' _ ⊩x~x')) ,
+            (subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx) }))
+    isFunctionalRelation.isStrictCodomain (isFuncRel incFuncRel) =
+      do
+        (stC , stC⊩isStrictCodomain)  idFuncRel perX .isStrictCodomain
+        let
+          realizer : ApplStrTerm as 1
+          realizer = ` stC ̇ (` pr₁ ̇ # zero)
+        return
+          (λ* realizer ,
+           { x x' r (⊩x~x' , ⊩ϕx)  subst  r'  r'   perX .equality  (x' , x')) (sym (λ*ComputationRule realizer r)) (stC⊩isStrictCodomain x x' _ ⊩x~x')}))
+    isFunctionalRelation.isRelational (isFuncRel incFuncRel) =
+      do
+        (relX , relX⊩isRelationalX)  idFuncRel perX .isRelational
+        (relϕ , relϕ⊩isRelationalϕ)  ϕ .isRelational
+        let
+          realizer : ApplStrTerm as 3
+          realizer = ` pair ̇ (` relX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇ (` relϕ ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # two))
+        return
+          (λ*3 realizer ,
+           { x₁ x₂ x₃ x₄ a b c (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₁~x₃ , ⊩ϕx₁') c⊩x₃~x₄ 
+            subst
+               r'  r'   perX .equality  (x₂ , x₄))
+              (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
+              (relX⊩isRelationalX x₁ x₂ x₃ x₄ _ _ _ ⊩x₁~x₂ ⊩x₁~x₃ c⊩x₃~x₄) ,
+            subst
+               r'  r'   ϕ .predicate  x₂)
+              (sym (cong  x  pr₂  x) (λ*3ComputationRule realizer a b c)  pr₂pxy≡y _ _))
+              (relϕ⊩isRelationalϕ x₁ x₂ _ _ ⊩ϕx₁ ⊩x₁~x₂) }))
+    isFunctionalRelation.isSingleValued (isFuncRel incFuncRel) =
+      do
+        (sv , sv⊩isSingleValuedX)  idFuncRel perX .isSingleValued
+        let
+          realizer : ApplStrTerm as 2
+          realizer = ` sv ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
+        return
+          (λ*2 realizer ,
+           { x x' x'' r₁ r₂ (⊩x~x' , ⊩ϕx) (⊩x~x'' , ⊩ϕx') 
+            subst  r'  r'   perX .equality  (x' , x'')) (sym (λ*2ComputationRule realizer r₁ r₂)) (sv⊩isSingleValuedX x x' x'' _ _ ⊩x~x' ⊩x~x'') }))
+    isFunctionalRelation.isTotal (isFuncRel incFuncRel) =
+      do
+        return
+          (Id ,
+           { x r (pr₁r⊩x~x , pr₂r⊩ϕx) 
+            return
+              (x ,
+              subst  r'  r'   perX .equality  (x , x)) (cong  x  pr₁  x) (sym (Ida≡a _))) pr₁r⊩x~x ,
+              subst  r'  r'   ϕ .predicate  x) (cong  x  pr₂  x) (sym (Ida≡a _))) pr₂r⊩ϕx) }))
+
+  opaque
+    unfolding isInjectiveFuncRel
+    unfolding incFuncRel
+    isInjectiveIncFuncRel : isInjectiveFuncRel subPer perX incFuncRel
+    isInjectiveIncFuncRel =
+      do
+        (t , t⊩isTransitiveX)  perX .isTransitive
+        (s , s⊩isSymmetricX)  perX .isSymmetric
+        let
+          realizer : ApplStrTerm as 2
+          realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` s ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ # one)
+        return
+          (λ*2 realizer ,
+           x₁ x₂ x₃ r₁ r₂ (⊩x₁~x₃ , ⊩ϕx₁) (⊩x₂~x₃ , ⊩ϕx₂) 
+            subst
+               r'  r'   perX .equality  (x₁ , x₂))
+              (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
+              (t⊩isTransitiveX x₁ x₃ x₂ _ _ ⊩x₁~x₃ (s⊩isSymmetricX x₂ x₃ _ ⊩x₂~x₃)) ,
+            subst
+               r'  r'   ϕ .predicate  x₁)
+              (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁  r₂)  pr₂pxy≡y _ _))
+              ⊩ϕx₁))
+
+  isMonicInc : isMonic RT [ incFuncRel ]
+  isMonicInc = isInjectiveFuncRel→isMonic subPer perX incFuncRel isInjectiveIncFuncRel
+
+-- Every subobject representing functional relation is isomorphic (as a subobject) to a subobject induced by a strict relation
+module SubobjectIsoMonicFuncRel
+  {X Y : Type ℓ'}
+  (perX : PartialEquivalenceRelation X)
+  (perY : PartialEquivalenceRelation Y)
+  (F : FunctionalRelation perY perX)
+  (isMonicF : isMonic RT [ F ]) where
+
+  {-# TERMINATING #-}
+  ψ : StrictRelation perX
+  Predicate.isSetX (StrictRelation.predicate ψ) = perX .isSetX
+  Predicate.∣ StrictRelation.predicate ψ  x r = ∃[ y  Y ] r   F .relation  (y , x)
+  Predicate.isPropValued (StrictRelation.predicate ψ) x r = isPropPropTrunc
+  isStrictRelation.isStrict (StrictRelation.isStrictRelationPredicate ψ) =
+    do
+      (stCF , stCF⊩isStrictCodomainF)  F .isStrictCodomain
+      return
+        (stCF ,
+         x r r⊩∃y 
+          transport
+            (propTruncIdempotent (perX .equality .isPropValued _ _))
+            (do
+              (y , ⊩Fyx)  r⊩∃y
+              return (stCF⊩isStrictCodomainF y x _ ⊩Fyx))))
+  isStrictRelation.isRelational (StrictRelation.isStrictRelationPredicate ψ) =
+    do
+      (relF , relF⊩isRelationalF)  F .isRelational
+      (stDF , stDF⊩isStrictDomainF)  F .isStrictDomain
+      let
+        realizer : ApplStrTerm as 2
+        realizer = ` relF ̇ (` stDF ̇ # one) ̇ # one ̇ # zero
+      return
+        (λ*2 realizer ,
+         x x' r s r⊩∃y s⊩x~x' 
+          do
+            (y , ⊩Fyx)  r⊩∃y
+            return
+              (y ,
+              subst
+                 r'  r'   F .relation  (y , x'))
+                (sym (λ*2ComputationRule realizer r s))
+                (relF⊩isRelationalF y y x x' _ _ _ (stDF⊩isStrictDomainF y x _ ⊩Fyx) ⊩Fyx s⊩x~x'))))
+
+  perψ : PartialEquivalenceRelation X
+  perψ = InducedSubobject.subPer perX ψ
+
+  -- ≤ as subobjects
+  -- TODO : formalise the preorder category of subobjects
+  {-# TERMINATING #-}
+  perY≤perψFuncRel : FunctionalRelation perY perψ
+  Predicate.isSetX (relation perY≤perψFuncRel) = isSet× (perY .isSetX) (perX .isSetX)
+  Predicate.∣ relation perY≤perψFuncRel  =  F .relation 
+  Predicate.isPropValued (relation perY≤perψFuncRel) = F .relation .isPropValued
+  isFunctionalRelation.isStrictDomain (isFuncRel perY≤perψFuncRel) =
+    isFunctionalRelation.isStrictDomain (F .isFuncRel)
+  isFunctionalRelation.isStrictCodomain (isFuncRel perY≤perψFuncRel) =
+    do
+      (stCF , stCF⊩isStrictCodomain)  F .isStrictCodomain
+      let
+        realizer : ApplStrTerm as 1
+        realizer = ` pair ̇ (` stCF ̇ # zero) ̇ # zero
+      return
+        (λ* realizer ,
+         y x r ⊩Fyx 
+          subst
+             r'  r'   perX .equality  (x , x))
+            (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+            (stCF⊩isStrictCodomain y x _ ⊩Fyx) ,
+           y ,
+            subst
+               r'  r'   F .relation  (y , x))
+              (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+              ⊩Fyx ∣₁))
+  isFunctionalRelation.isRelational (isFuncRel perY≤perψFuncRel) =
+    do
+      (relF , relF⊩isRelationalF)  F .isRelational
+      let
+        realizer : ApplStrTerm as 3
+        realizer = ` relF ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
+      return
+        (λ*3 realizer ,
+         { y y' x x' a b c ⊩y~y' ⊩Fyx (⊩x~x' , ⊩Fy''x) 
+          subst  r'  r'   F .relation  (y' , x')) (sym (λ*3ComputationRule realizer a b c)) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') }))
+  isFunctionalRelation.isSingleValued (isFuncRel perY≤perψFuncRel) =
+    do
+      (svF , svF⊩isSingleValuedF)  F .isSingleValued
+      let
+        realizer : ApplStrTerm as 2
+        realizer = ` pair ̇ (` svF ̇ # one ̇ # zero) ̇ # one
+      return
+        (λ*2 realizer ,
+         y x x' r₁ r₂ ⊩Fyx ⊩Fyx' 
+          subst
+             r'  r'   perX .equality  (x , x'))
+            (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
+            (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') ,
+           y ,
+            (subst
+               r'  r'   F .relation  (y , x))
+              (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _))
+              ⊩Fyx) ∣₁))
+  isFunctionalRelation.isTotal (isFuncRel perY≤perψFuncRel) =
+    do
+      (tlF , tlF⊩isTotalF)  F .isTotal
+      return
+        (tlF ,
+         y r ⊩y~y 
+          do
+            (x , ⊩Fyx)  tlF⊩isTotalF y _ ⊩y~y
+            return (x , ⊩Fyx)))
+
+  -- perY truly is ≤ perψ
+  opaque
+    unfolding composeRTMorphism
+    unfolding InducedSubobject.incFuncRel
+    perY≤perψCommutes : [ perY≤perψFuncRel ]  [ InducedSubobject.incFuncRel perX ψ ]  [ F ]
+    perY≤perψCommutes =
+      let
+        answer =
+          do
+            (stDF , stDF⊩isStrictDomainF)  F .isStrictDomain
+            (relF , relF⊩isRelationalF)  F .isRelational
+            let
+              realizer : ApplStrTerm as 1
+              realizer = ` relF ̇ (` stDF ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+            return
+              (λ* realizer ,
+               y x r r⊩∃x' 
+                transport
+                  (propTruncIdempotent (F .relation .isPropValued _ _))
+                  (do
+                    (x' , ⊩Fyx' , ⊩x'~x , ⊩ψx')  r⊩∃x'
+                    return
+                      (subst
+                         r  r   F .relation  (y , x))
+                        (sym (λ*ComputationRule realizer r))
+                        (relF⊩isRelationalF y y x' x _ _ _ (stDF⊩isStrictDomainF y x' _ ⊩Fyx') ⊩Fyx' ⊩x'~x)))))
+      in
+      eq/ _ _ (answer , F≤G→G≤F perY perX (composeFuncRel _ _ _ perY≤perψFuncRel (InducedSubobject.incFuncRel perX ψ)) F answer)
+
+  opaque
+    unfolding isInjectiveFuncRel
+    {-# TERMINATING #-}
+    perψ≤perYFuncRel : FunctionalRelation perψ perY
+    Predicate.isSetX (relation perψ≤perYFuncRel) = isSet× (perX .isSetX) (perY .isSetX)
+    Predicate.∣ relation perψ≤perYFuncRel  (x , y) r = r   F .relation  (y , x)
+    Predicate.isPropValued (relation perψ≤perYFuncRel) (x , y) r = F .relation .isPropValued _ _
+    isFunctionalRelation.isStrictDomain (isFuncRel perψ≤perYFuncRel) =
+      do
+        (stCF , stCF⊩isStrictCodomainF)  F .isStrictCodomain
+        let
+          realizer : ApplStrTerm as 1
+          realizer = ` pair ̇ (` stCF ̇ # zero) ̇ # zero
+        return
+          (λ* realizer ,
+           x y r ⊩Fyx 
+            (subst
+               r'  r'   perX .equality  (x , x))
+              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+              (stCF⊩isStrictCodomainF y x _ ⊩Fyx)) ,
+            (return
+              (y ,
+              (subst
+                 r'  r'   F .relation  (y , x))
+                (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                ⊩Fyx)))))
+    isFunctionalRelation.isStrictCodomain (isFuncRel perψ≤perYFuncRel) =
+      do
+        (stDF , stDF⊩isStrictDomainF)  F .isStrictDomain
+        return
+          (stDF ,
+           x y r ⊩Fyx  stDF⊩isStrictDomainF y x _ ⊩Fyx))
+    isFunctionalRelation.isRelational (isFuncRel perψ≤perYFuncRel) =
+      do
+        (relF , relF⊩isRelationalF)  F .isRelational
+        let
+          realizer : ApplStrTerm as 3
+          realizer = ` relF ̇ # zero ̇ # one ̇ (` pr₁ ̇ # two)
+        return
+          (λ*3 realizer ,
+           { x x' y y' a b c (⊩x~x' , ⊩ψx) ⊩Fyx ⊩y~y' 
+            subst  r'  r'   F .relation  (y' , x')) (sym (λ*3ComputationRule realizer a b c)) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') }))
+    isFunctionalRelation.isSingleValued (isFuncRel perψ≤perYFuncRel) =
+      let
+        isInjectiveFuncRelF = isMonic→isInjectiveFuncRel perY perX F isMonicF
+      in
+      do
+        (injF , injF⊩isInjectiveF)  isInjectiveFuncRelF
+        return
+          (injF ,
+           x y y' r₁ r₂ ⊩Fyx ⊩Fy'x 
+            injF⊩isInjectiveF y y' x _ _ ⊩Fyx ⊩Fy'x))
+    isFunctionalRelation.isTotal (isFuncRel perψ≤perYFuncRel) =
+        return
+          (pr₂ ,
+           { x r (⊩x~x , ⊩ψx)  ⊩ψx }))
+
+  opaque
+    unfolding composeRTMorphism
+    unfolding composeFuncRel
+    unfolding InducedSubobject.incFuncRel
+    unfolding perψ≤perYFuncRel
+    perψ≤perYCommutes : [ perψ≤perYFuncRel ]  [ F ]  [ InducedSubobject.incFuncRel perX ψ ]
+    perψ≤perYCommutes =
+      let
+        answer =
+          do
+            (svF , svF⊩isSingleValuedF)  F .isSingleValued
+            let
+              realizer : ApplStrTerm as 1
+              realizer = ` pair ̇ (` svF ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero)
+            return
+              (λ* realizer ,
+               x x' r r⊩∃y 
+                transport
+                  (propTruncIdempotent (isProp× (perX .equality .isPropValued _ _) isPropPropTrunc))
+                  (do
+                    (y , ⊩Fyx , ⊩Fyx')  r⊩∃y
+                    return
+                      (subst
+                         r'  r'   perX .equality  (x , x'))
+                        (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                        (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') ,
+                      return
+                        (y ,
+                        (subst
+                           r'  r'   F .relation  (y , x))
+                          (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                          ⊩Fyx))))))
+      in
+      eq/ _ _ (answer , F≤G→G≤F perψ perX (composeFuncRel _ _ _ perψ≤perYFuncRel F) (InducedSubobject.incFuncRel perX ψ) answer)
+
+-- For strict relations, subobject inclusion is identified with pointwise entailment
+module InclusionEntailment
+  {X : Type ℓ'}
+  (perX : PartialEquivalenceRelation X)
+  (ϕ ψ : StrictRelation perX) where
+  open StrictRelation
+  open PredicateProperties X
+  SubObjX = SubObjCat RT (X , perX)
+  SubObjHom = Category.Hom[_,_] SubObjX
+
+  perϕ = InducedSubobject.subPer perX ϕ
+  perψ = InducedSubobject.subPer perX ψ
+
+  incϕ = InducedSubobject.incFuncRel perX ϕ
+  incψ = InducedSubobject.incFuncRel perX ψ
+
+  ϕsubObject : Category.ob SubObjX
+  ϕsubObject = sliceob [ InducedSubobject.incFuncRel perX ϕ ] , InducedSubobject.isMonicInc perX ϕ
+
+  ψsubObject : Category.ob SubObjX
+  ψsubObject = sliceob [ InducedSubobject.incFuncRel perX ψ ] , InducedSubobject.isMonicInc perX ψ
+
+  opaque
+    unfolding composeRTMorphism
+    unfolding composeFuncRel
+    unfolding InducedSubobject.incFuncRel
+    SubObjHom→ϕ≤ψ : SubObjHom ϕsubObject ψsubObject  (ϕ .predicate  ψ .predicate)
+    SubObjHom→ϕ≤ψ (slicehom f f⋆incψ≡incϕ) =
+      SQ.elimProp
+        {P = λ f  (f  [ incψ ]  [ incϕ ])  ϕ .predicate  ψ .predicate}
+         f  isPropΠ λ f⋆incψ≡incϕ  isProp≤ (ϕ .predicate) (ψ .predicate))
+         F F⋆incψ≡incϕ 
+          let
+            (p , q) =
+              SQ.effective
+                (isPropValuedBientailment (InducedSubobject.subPer perX ϕ) perX)
+                (isEquivRelBientailment (InducedSubobject.subPer perX ϕ) perX)
+                (composeFuncRel _ _ _ F incψ)
+                incϕ
+                F⋆incψ≡incϕ
+          in
+          do
+            (stϕ , stϕ⊩isStrictϕ)  ϕ .isStrict
+            (relψ , relψ⊩isRelationalψ)  ψ .isRelational
+            (q , q⊩incϕ≤F⋆incψ)  q
+            let
+              realizer : ApplStrTerm as 1
+              realizer = ` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # zero) ̇ # zero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # zero) ̇ # zero))))
+            return
+              (λ* realizer ,
+               x a a⊩ϕx 
+                transport
+                  (propTruncIdempotent (ψ .predicate .isPropValued _ _))
+                  (do
+                    (x' , ⊩Fxx' , ⊩x'~x , ⊩ψx') 
+                      q⊩incϕ≤F⋆incψ
+                        x x
+                        (pair  (stϕ  a)  a)
+                        ((subst  r'  r'   perX .equality  (x , x)) (sym (pr₁pxy≡x _ _)) (stϕ⊩isStrictϕ x a a⊩ϕx)) ,
+                         (subst  r'  r'   ϕ .predicate  x) (sym (pr₂pxy≡y _ _)) a⊩ϕx))
+                    return (subst  r'  r'   ψ .predicate  x) (sym (λ*ComputationRule realizer a)) (relψ⊩isRelationalψ x' x _ _ ⊩ψx' ⊩x'~x))))))
+        f
+        f⋆incψ≡incϕ
+
+  module _ (ϕ≤ψ : ϕ .predicate  ψ .predicate) where opaque
+    unfolding idFuncRel
+    unfolding composeRTMorphism
+    unfolding composeFuncRel
+    unfolding InducedSubobject.incFuncRel
+
+    {-# TERMINATING #-}
+    funcRel : FunctionalRelation perϕ perψ
+    Predicate.isSetX (relation funcRel) = isSet× (perX .isSetX) (perX .isSetX)
+    Predicate.∣ relation funcRel  (x , x') r = (pr₁  r)   perX .equality  (x , x') × ((pr₁  (pr₂  r))   ϕ .predicate  x) × ((pr₂  (pr₂  r))   ψ .predicate  x)
+    Predicate.isPropValued (relation funcRel) (x , x') r = isProp× (perX .equality .isPropValued _ _) (isProp× (ϕ .predicate .isPropValued _ _) (ψ .predicate .isPropValued _ _))
+    isFunctionalRelation.isStrictDomain (isFuncRel funcRel) =
+      do
+        (stϕ , stϕ⊩isStrictϕ)  ϕ .isStrict
+        let
+          realizer : ApplStrTerm as 1
+          realizer = ` pair ̇ (` stϕ ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+        return
+          (λ* realizer ,
+           { x x' r (⊩x~x' , ⊩ϕx , ⊩ψx) 
+            subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (stϕ⊩isStrictϕ x _ ⊩ϕx) ,
+            subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx}))
+    isFunctionalRelation.isStrictCodomain (isFuncRel funcRel) =
+      do
+        (stCX , stCX⊩isStrictCodomainX)  idFuncRel perX .isStrictCodomain
+        (relψ , relψ⊩isRelationalψ)  ψ .isRelational
+        let
+          realizer : ApplStrTerm as 1
+          realizer = ` pair ̇ (` stCX ̇ (` pr₁ ̇ # zero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero))
+        return
+          (λ* realizer ,
+           { x x' r (⊩x~x' , ⊩ϕx , ⊩ψx) 
+            subst
+               r'  r'   perX .equality  (x' , x'))
+              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+              (stCX⊩isStrictCodomainX x x' _ ⊩x~x') ,
+            subst
+               r'  r'   ψ .predicate  x')
+              (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+              (relψ⊩isRelationalψ x x' _ _ ⊩ψx ⊩x~x')}))
+    isFunctionalRelation.isRelational (isFuncRel funcRel) =
+      do
+        (relX , relX⊩isRelationalX)  idFuncRel perX .isRelational
+        (relϕ , relϕ⊩isRelationalϕ)  ϕ .isRelational
+        (relψ , relψ⊩isRelationalψ)  ψ .isRelational
+        let
+          realizer : ApplStrTerm as 3
+          realizer =
+            ` pair ̇ (` relX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` relϕ ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # two)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` pr₁ ̇ # two)))
+        return
+          (λ*3 realizer ,
+          λ { x₁ x₂ x₃ x₄ a b c (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₁~x₃ , ⊩'ϕx₁ , ⊩ψx₁) (⊩x₃~x₄ , ⊩ψx₃) 
+            subst
+               r'  r'   perX .equality  (x₂ , x₄))
+              (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
+              (relX⊩isRelationalX x₁ x₂ x₃ x₄ _ _ _ ⊩x₁~x₂ ⊩x₁~x₃ ⊩x₃~x₄) ,
+            subst
+               r'  r'   ϕ .predicate  x₂)
+              (sym (cong  x  pr₁  (pr₂  x)) (λ*3ComputationRule realizer a b c)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+              (relϕ⊩isRelationalϕ x₁ x₂ _ _ ⊩ϕx₁ ⊩x₁~x₂) ,
+            subst
+               r'  r'   ψ .predicate  x₂)
+              (sym (cong  x  pr₂  (pr₂  x)) (λ*3ComputationRule realizer a b c)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+              (relψ⊩isRelationalψ x₁ x₂ _ _ ⊩ψx₁ ⊩x₁~x₂)})
+    isFunctionalRelation.isSingleValued (isFuncRel funcRel) =
+      do
+        (svX , svX⊩isSingleValuedX)  idFuncRel perX .isSingleValued
+        (relψ , relψ⊩isRelationalψ)  ψ .isRelational
+        let
+          realizer : ApplStrTerm as 2
+          realizer = ` pair ̇ (` svX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` pr₁ ̇ # one))
+        return
+          (λ*2 realizer ,
+           { x₁ x₂ x₃ r₁ r₂ (⊩x₁~x₂ , ⊩ϕx , ⊩ψx) (⊩x₁~x₃ , ⊩'ϕx , ⊩'ψx) 
+            (subst  r'  r'   perX .equality  (x₂ , x₃)) (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _)) (svX⊩isSingleValuedX x₁ x₂ x₃ _ _ ⊩x₁~x₂ ⊩x₁~x₃)) ,
+             subst  r'  r'   ψ .predicate  x₂) (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _)) (relψ⊩isRelationalψ x₁ x₂ _ _ ⊩ψx ⊩x₁~x₂)}))
+    isFunctionalRelation.isTotal (isFuncRel funcRel) =
+      do
+        (tl , tl⊩isTotalIncψ)  incψ .isTotal
+        (s , s⊩ϕ≤ψ)  ϕ≤ψ
+        let
+          realizer : ApplStrTerm as 1
+          realizer = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` pair ̇ (` pr₂ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero)))
+        return
+          (λ* realizer ,
+           { x r (⊩x~x , ⊩ϕx) 
+            return
+              (x ,
+              subst
+                 r'  r'   perX .equality  (x , x))
+                (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                ⊩x~x ,
+              subst
+                 r'  r'   ϕ .predicate  x)
+                (sym (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                ⊩ϕx ,
+              subst
+                 r'  r'   ψ .predicate  x)
+                (sym (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule realizer r)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                (s⊩ϕ≤ψ x _ ⊩ϕx))}))
+    
+    funcRel⋆incψ≡incϕ : [ funcRel ]  [ incψ ]  [ incϕ ]
+    funcRel⋆incψ≡incϕ =
+      let
+        answer =
+          do
+            (t , t⊩isTransitiveX)  perX .isTransitive
+            let
+              realizer : ApplStrTerm as 1
+              realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)))
+            return
+              (λ* realizer ,
+               { x x' r ⊩∃x'' 
+                transport
+                  (propTruncIdempotent (isPropΣ (perX .equality .isPropValued _ _) λ _  ϕ .predicate .isPropValued _ _))
+                  (do
+                    (x'' , (⊩x~x'' , ⊩ϕx , ⊩ψx) , (⊩x''~x' , ⊩'ψx))  ⊩∃x''
+                    return
+                      ((subst
+                         r'  r'   perX .equality  (x , x'))
+                        (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                        (t⊩isTransitiveX x x'' x' _ _ ⊩x~x'' ⊩x''~x')) ,
+                       (subst
+                          r'  r'   ϕ .predicate  x)
+                         (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                         ⊩ϕx)))}))
+      in
+      eq/ _ _ (answer , F≤G→G≤F perϕ perX (composeFuncRel _ _ _ funcRel incψ) incϕ answer)
+
+    ϕ≤ψ→SubObjHom : SubObjHom ϕsubObject ψsubObject
+    ϕ≤ψ→SubObjHom =
+      slicehom [ funcRel ] funcRel⋆incψ≡incϕ
+
+  SubObjHom≃ϕ≤ψ : SubObjHom ϕsubObject ψsubObject  (ϕ .predicate  ψ .predicate)
+  SubObjHom≃ϕ≤ψ =
+    propBiimpl→Equiv
+      (isPropSubObjMor RT (X , perX) ϕsubObject ψsubObject)
+      (isProp≤ (ϕ .predicate) (ψ .predicate))
+      SubObjHom→ϕ≤ψ
+      ϕ≤ψ→SubObjHom
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.SubobjectClassifier.html b/docs/Realizability.Topos.SubobjectClassifier.html new file mode 100644 index 0000000..4d86d16 --- /dev/null +++ b/docs/Realizability.Topos.SubobjectClassifier.html @@ -0,0 +1,940 @@ + +Realizability.Topos.SubobjectClassifier
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; ⟦_⟧ to pre⟦_⟧)
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Equiv
+open import Cubical.Data.Empty
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Morphism
+open import Realizability.PropResizing
+open import Realizability.CombinatoryAlgebra
+
+module Realizability.Topos.SubobjectClassifier
+  {}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (resizing : hPropResizing )
+  where
+  
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = } {ℓ'' = } ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = } {ℓ'' = } ca
+open import Realizability.Topos.Object {ℓ' = } {ℓ'' = } ca isNonTrivial 
+open import Realizability.Topos.FunctionalRelation {ℓ' = } {ℓ'' = } ca isNonTrivial
+open import Realizability.Topos.Equalizer {ℓ' = } {ℓ'' = } ca isNonTrivial
+open import Realizability.Topos.BinProducts {ℓ' = } {ℓ'' = } ca isNonTrivial
+open import Realizability.Topos.MonicReprFuncRel {ℓ' = } {ℓ'' = } ca isNonTrivial
+open import Realizability.Topos.ResizedPredicate ca isNonTrivial resizing
+open import Realizability.Topos.TerminalObject {ℓ' = } {ℓ'' = } ca isNonTrivial
+open import Realizability.Topos.StrictRelation {ℓ' = } {ℓ'' = } ca isNonTrivial
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate renaming (isSetX to isSetPredicateBase)
+open Morphism
+open PartialEquivalenceRelation
+open FunctionalRelation
+open Category RT
+
+⟦_⟧ = pre⟦_⟧ as
+
+Ωper : PartialEquivalenceRelation (ResizedPredicate Unit*)
+Predicate.isSetX (equality Ωper) = isSet× isSetResizedPredicate isSetResizedPredicate
+Predicate.∣ equality Ωper  (α , β) r =
+  (∀ (a : A) (⊩α : a   toPredicate α  tt*)  ((pr₁  r)  a)   toPredicate β  tt*) ×
+  (∀ (a : A) (⊩β : a   toPredicate β  tt*)  ((pr₂  r)  a)   toPredicate α  tt*)
+Predicate.isPropValued (equality Ωper) (α , β) r =
+  isProp×
+    (isPropΠ  _  isPropΠ λ _  (toPredicate β) .isPropValued _ _))
+    (isPropΠ  _  isPropΠ λ _  (toPredicate α) .isPropValued _ _))
+isPartialEquivalenceRelation.isSetX (isPerEquality Ωper) = isSetResizedPredicate
+isPartialEquivalenceRelation.isSymmetric (isPerEquality Ωper) =
+  do
+    let
+      ent₁ : ApplStrTerm as 2
+      ent₁ = ` pr₂ ̇ # one ̇ # zero
+
+      ent₂ : ApplStrTerm as 2
+      ent₂ = ` pr₁ ̇ # one ̇ # zero
+
+      realizer : ApplStrTerm as 1
+      realizer = ` pair ̇ (λ*abst ent₁) ̇ (λ*abst ent₂)
+    return
+      (λ* realizer ,
+      λ { α β r (pr₁r⊩α≤β , pr₂r⊩β≤α) 
+         a a⊩β 
+          let
+            eq : pr₁  (λ* realizer  r)  a  pr₂  r  a
+            eq =
+              pr₁  (λ* realizer  r)  a
+                ≡⟨ cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
+              pr₁  (pair  _  _)  a
+                ≡⟨ cong  x  x  a) (pr₁pxy≡x _ _) 
+               (λ*abst ent₁)  (r  [])  a
+                ≡⟨ βreduction ent₁ a (r  []) 
+              pr₂  r  a
+                
+          in
+          subst  r'  r'   toPredicate α  tt*) (sym eq) (pr₂r⊩β≤α a a⊩β)) ,
+         a a⊩α 
+          let
+            eq : pr₂  (λ* realizer  r)  a  pr₁  r  a
+            eq =
+              pr₂  (λ* realizer  r)  a
+                ≡⟨ cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+              pr₂  (pair  _  _)  a
+                ≡⟨ cong  x  x  a) (pr₂pxy≡y _ _) 
+               λ*abst ent₂  (r  [])  a
+                ≡⟨ βreduction ent₂ a (r  []) 
+              pr₁  r  a
+                
+          in
+          subst  r'  r'   toPredicate β  tt*) (sym eq) (pr₁r⊩α≤β a a⊩α)) })
+isPartialEquivalenceRelation.isTransitive (isPerEquality Ωper) =
+  do
+    let
+      closure1 : ApplStrTerm as 3
+      closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ # two ̇ # zero)
+
+      closure2 : ApplStrTerm as 3
+      closure2 = ` pr₂ ̇ # two ̇ (` pr₂ ̇ # one ̇ # zero)
+
+      realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2)
+    return
+      (λ*2 realizer ,
+       { x y z a b (⊩x≤y , ⊩y≤x) (⊩y≤z , ⊩z≤y) 
+         r r⊩x 
+          subst
+             r'  r'   toPredicate z  tt*)
+            (sym
+              (cong  x  pr₁  x  r) (λ*2ComputationRule realizer a b) 
+               cong  x  x  r) (pr₁pxy≡x _ _) 
+               βreduction closure1 r (b  a  [])))
+            (⊩y≤z _ (⊩x≤y r r⊩x))) ,
+         r r⊩z 
+          subst
+             r'  r'   toPredicate x  tt*)
+            (sym
+              (cong  x  pr₂  x  r) (λ*2ComputationRule realizer a b) 
+               cong  x  x  r) (pr₂pxy≡y _ _) 
+               βreduction closure2 r (b  a  [])))
+            (⊩y≤x _ (⊩z≤y r r⊩z))) }))
+
+opaque
+  unfolding terminalPer
+  trueFuncRel : FunctionalRelation terminalPer Ωper
+  Predicate.isSetX (relation trueFuncRel) = isSet× isSetUnit* isSetResizedPredicate
+  Predicate.∣ relation trueFuncRel  (tt* , p) r =  (a : A)  (r  a)   toPredicate p  tt*
+  Predicate.isPropValued (relation trueFuncRel) (tt* , p) r = isPropΠ λ a  (toPredicate p) .isPropValued _ _
+  isFunctionalRelation.isStrictDomain (isFuncRel trueFuncRel) =
+    do
+      return
+        (k ,
+         { tt* y r r⊩⊤≤y  tt*}))
+  isFunctionalRelation.isStrictCodomain (isFuncRel trueFuncRel) =
+    do
+      let
+        idClosure : ApplStrTerm as 2
+        idClosure = # zero
+        realizer : ApplStrTerm as 1
+        realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure)
+      return
+        (λ* realizer ,
+         { tt* y r r⊩⊤≤y 
+           a a⊩y 
+            subst
+               r'  r'   toPredicate y  tt*)
+              (sym
+                (cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₁pxy≡x _ _) 
+                 βreduction idClosure a (r  [])))
+              a⊩y) ,
+           a a⊩y 
+            subst
+               r'  r'   toPredicate y  tt*)
+              (sym
+                (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₂pxy≡y _ _) 
+                 βreduction idClosure a (r  [])))
+              a⊩y)}))
+  isFunctionalRelation.isRelational (isFuncRel trueFuncRel) =
+    do
+      let
+        realizer : ApplStrTerm as 4
+        realizer = ` pr₁ ̇ # one ̇ (# two  ̇ ` k)
+      return
+        (λ*4 realizer ,
+         { tt* tt* x y a b c tt* b⊩⊤≤x (pr₁c⊩x≤y , pr₂c⊩y≤x) r 
+          subst
+             r'  r'   toPredicate y  tt*)
+            (sym (λ*4ComputationRule realizer a b c r))
+            (pr₁c⊩x≤y (b  k) (b⊩⊤≤x k))}))
+  isFunctionalRelation.isSingleValued (isFuncRel trueFuncRel) =
+    do
+      let
+        closure1 : ApplStrTerm as 3
+        closure1 = # one ̇ ` k
+
+        closure2 : ApplStrTerm as 3
+        closure2 = # two ̇ ` k
+        
+        realizer : ApplStrTerm as 2
+        realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2)
+      return
+        (λ*2 realizer ,
+         { tt* x y r₁ r₂ r₁⊩⊤≤x r₂⊩⊤≤y 
+           a a⊩x 
+            subst
+               r'  r'   toPredicate y  tt*)
+              (sym
+                (cong  x  pr₁  x  a) (λ*2ComputationRule realizer r₁ r₂) 
+                 cong  x  x  a) (pr₁pxy≡x _ _) 
+                 βreduction closure1 a (r₂  r₁  [])))
+              (r₂⊩⊤≤y k)) ,
+           a a⊩y 
+            subst
+               r'  r'   toPredicate x  tt*)
+              (sym
+                (cong  x  pr₂  x  a) (λ*2ComputationRule realizer r₁ r₂) 
+                 cong  x  x  a) (pr₂pxy≡y _ _) 
+                 βreduction closure2 a (r₂  r₁  [])))
+              (r₁⊩⊤≤x k))}))
+  isFunctionalRelation.isTotal (isFuncRel trueFuncRel) =
+    do
+      return
+        (k ,
+         { tt* r tt* 
+          let
+             = pre1 Unit* isSetUnit* isNonTrivial
+          in
+          
+            fromPredicate  ,
+             a 
+              subst  p  (k  r  a)   p  tt*) (sym (compIsIdFunc )) tt*)
+          ∣₁ }))
+
+opaque
+  unfolding isInjectiveFuncRel
+  unfolding terminalPer
+  isInjectiveTrueFuncRel : isInjectiveFuncRel _ _ trueFuncRel
+  isInjectiveTrueFuncRel =
+    do
+      return
+        (k ,
+         { tt* tt* p r₁ r₂ r₁⊩⊤≤p r₂⊩⊤≤p  tt* }))
+
+truePredicate : Predicate Unit*
+Predicate.isSetX truePredicate = isSetUnit*
+Predicate.∣ truePredicate  tt* r = Unit*
+Predicate.isPropValued truePredicate tt* r = isPropUnit*
+
+ = fromPredicate truePredicate
+
+-- The subobject classifier classifies subobjects represented by strict relations
+module ClassifiesStrictRelations
+  (X : Type )
+  (perX : PartialEquivalenceRelation X)
+  (ϕ : StrictRelation perX) where
+
+  open InducedSubobject perX ϕ
+  open StrictRelation
+  resizedϕ = fromPredicate (ϕ .predicate)
+
+  -- the functional relation that represents the unique indicator map
+  {-# TERMINATING #-}
+  charFuncRel : FunctionalRelation perX Ωper
+  Predicate.isSetX (relation charFuncRel) = isSet× (perX .isSetX) isSetResizedPredicate
+  Predicate.∣ relation charFuncRel  (x , p) r =
+    (pr₁  r)   perX .equality  (x , x) ×
+    (∀ (b : A) (b⊩ϕx : b   ϕ .predicate  x)  (pr₁  (pr₂  r)  b)   toPredicate p  tt*) ×
+    (∀ (b : A) (b⊩px : b   toPredicate p  tt*)  (pr₂  (pr₂  r)  b)   ϕ .predicate  x)
+  Predicate.isPropValued (relation charFuncRel) (x , p) r =
+    isProp×
+      (perX .equality .isPropValued _ _)
+      (isProp×
+        (isPropΠ  _  isPropΠ λ _  (toPredicate p) .isPropValued _ _))
+        (isPropΠ λ _  isPropΠ λ _  ϕ .predicate .isPropValued _ _))
+  isFunctionalRelation.isStrictDomain (isFuncRel charFuncRel) =
+    do
+      return
+        (pr₁ ,
+         { x p r (pr₁r⊩x~x , ⊩ϕx≤p , ⊩p≤ϕx)  pr₁r⊩x~x}))
+  isFunctionalRelation.isStrictCodomain (isFuncRel charFuncRel) =
+    do
+      let
+        idClosure : ApplStrTerm as 2
+        idClosure = # zero
+        realizer : ApplStrTerm as 1
+        realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure)
+      return
+        (λ* realizer ,
+         x y r x₁ 
+           a a⊩y 
+            subst
+               r'  r'   toPredicate y  tt*)
+              (sym
+                (cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₁pxy≡x _ _) 
+                 βreduction idClosure a (r  [])))
+              a⊩y) ,
+           a a⊩y 
+            subst
+               r'  r'   toPredicate y  tt*)
+              (sym
+                (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₂pxy≡y _ _) 
+                 βreduction idClosure a (r  [])))
+              a⊩y)))
+  isFunctionalRelation.isRelational (isFuncRel charFuncRel) =
+    do
+      (sX , sX⊩isSymmetricX)  perX .isSymmetric
+      (tX , tX⊩isTransitiveX)  perX .isTransitive
+      (relϕ , relϕ⊩isRelationalϕ)  isStrictRelation.isRelational (ϕ .isStrictRelationPredicate)
+      let
+        closure1 : ApplStrTerm as 4
+        closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` relϕ ̇ # zero ̇ (` sX ̇ # three)))
+
+        closure2 : ApplStrTerm as 4
+        closure2 = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ # one ̇ # zero)) ̇ # three
+
+        realizer : ApplStrTerm as 3
+        realizer = ` pair ̇ (` tX ̇ (` sX ̇ # two) ̇ # two) ̇ (` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2))
+      return
+        (λ*3 realizer ,
+         { x x' p p' a b c a⊩x~x' (⊩x~x , ⊩ϕx≤p , ⊩p≤ϕx) (⊩p≤p' , ⊩p'≤p) 
+          let
+            ⊩x'~x = sX⊩isSymmetricX x x' a a⊩x~x'
+            ⊩x'~x' = tX⊩isTransitiveX x' x x' _ _ ⊩x'~x a⊩x~x'
+          in
+          subst
+             r'  r'   perX .equality  (x' , x'))
+            (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
+            ⊩x'~x' ,
+           r r⊩ϕx' 
+            subst
+               r'  r'   toPredicate p'  tt*)
+              (sym
+                (cong  x  pr₁  (pr₂  x)  r) (λ*3ComputationRule realizer a b c) 
+                 cong  x  pr₁  x  r) (pr₂pxy≡y _ _) 
+                 cong  x  x  r) (pr₁pxy≡x _ _) 
+                 βreduction closure1 r (c  b  a  [])))
+              (⊩p≤p' _ (⊩ϕx≤p _ (relϕ⊩isRelationalϕ x' x _ _ r⊩ϕx' ⊩x'~x)))) ,
+          λ r r⊩p' 
+            subst
+               r'  r'   ϕ .predicate  x')
+              (sym
+                (cong  x  pr₂  (pr₂  x)  r) (λ*3ComputationRule realizer a b c) 
+                 cong  x  pr₂  x  r) (pr₂pxy≡y _ _) 
+                 cong  x  x  r) (pr₂pxy≡y _ _) 
+                 βreduction closure2 r (c  b  a  [])))
+              (relϕ⊩isRelationalϕ x x' _ _ (⊩p≤ϕx _ (⊩p'≤p r r⊩p')) a⊩x~x') }))
+  isFunctionalRelation.isSingleValued (isFuncRel charFuncRel) =
+    do
+      let
+        closure1 : ApplStrTerm as 3
+        closure1 = ` pr₁ ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ # zero)
+
+        closure2 : ApplStrTerm as 3
+        closure2 = ` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ (` pr₂ ̇ # one) ̇ # zero)
+
+        realizer : ApplStrTerm as 2
+        realizer = ` pair ̇ λ*abst closure1 ̇ λ*abst closure2
+      return
+        (λ*2 realizer ,
+         { x y y' r₁ r₂ (⊩x~x , ⊩ϕx≤y , ⊩y≤ϕx) (⊩'x~x , ⊩ϕx≤y' , ⊩y'≤ϕx) 
+           a a⊩y 
+            subst
+               r'  r'   toPredicate y'  tt*)
+              (sym (cong  x  pr₁  x  a) (λ*2ComputationRule realizer r₁ r₂)  cong  x  x  a) (pr₁pxy≡x _ _)  βreduction closure1 a (r₂  r₁  [])))
+              (⊩ϕx≤y' _ (⊩y≤ϕx a a⊩y))) ,
+           a a⊩y' 
+            subst
+               r'  r'   toPredicate y  tt*)
+              (sym (cong  x  pr₂  x  a) (λ*2ComputationRule realizer r₁ r₂)  cong  x  x  a) (pr₂pxy≡y _ _)  βreduction closure2 a (r₂  r₁  [])))
+              (⊩ϕx≤y _ (⊩y'≤ϕx a a⊩y'))) }))
+  isFunctionalRelation.isTotal (isFuncRel charFuncRel) =
+    do
+      let
+        idClosure : ApplStrTerm as 2
+        idClosure = # zero
+
+        realizer : ApplStrTerm as 1
+        realizer = ` pair ̇ # zero ̇ (` pair ̇ λ*abst idClosure ̇ λ*abst idClosure)
+      return
+        (λ* realizer ,
+         x r r⊩x~x 
+          let
+            resultPredicate : Predicate Unit*
+            resultPredicate =
+              makePredicate
+                isSetUnit*
+                 { tt* s  s   ϕ .predicate  x })
+                 { tt* s  ϕ .predicate .isPropValued _ _ })
+          in
+          return
+            (fromPredicate resultPredicate ,
+            subst
+               r'  r'   perX .equality  (x , x))
+              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+              r⊩x~x ,
+             b b⊩ϕx 
+              subst
+                 r  r   toPredicate (fromPredicate resultPredicate)  tt*)
+                (sym
+                  (cong  x  pr₁  (pr₂  x)  b) (λ*ComputationRule realizer r) 
+                   cong  x  pr₁  x  b) (pr₂pxy≡y _ _) 
+                   cong  x  x  b) (pr₁pxy≡x _ _) 
+                   βreduction idClosure b (r  [])))
+                (subst  p  b   p  tt*) (sym (compIsIdFunc resultPredicate)) b⊩ϕx)) ,
+             b b⊩'ϕx 
+              subst
+                 r  r   ϕ .predicate  x)
+                (sym
+                  (cong  x  pr₂  (pr₂  x)  b) (λ*ComputationRule realizer r) 
+                   cong  x  pr₂  x  b) (pr₂pxy≡y _ _) 
+                   cong  x  x  b) (pr₂pxy≡y _ _) 
+                   βreduction idClosure b (r  [])))
+                let foo = subst  p  b   p  tt*) (compIsIdFunc resultPredicate) b⊩'ϕx in foo))))
+
+  subobjectCospan :  char  Cospan RT
+  Cospan.l (subobjectCospan char) = X , perX
+  Cospan.m (subobjectCospan char) = ResizedPredicate Unit* , Ωper
+  Cospan.r (subobjectCospan char) = Unit* , terminalPer
+  Cospan.s₁ (subobjectCospan char) = char
+  Cospan.s₂ (subobjectCospan char) = [ trueFuncRel ]
+
+  opaque
+    unfolding composeRTMorphism
+    unfolding composeFuncRel
+    unfolding terminalFuncRel
+    unfolding trueFuncRel
+    unfolding incFuncRel
+    subobjectSquareCommutes : [ incFuncRel ]  [ charFuncRel ]  [ terminalFuncRel subPer ]  [ trueFuncRel ]
+    subobjectSquareCommutes =
+      let
+        answer =
+          do
+            (stX , stX⊩isStrictDomainX)  idFuncRel perX .isStrictDomain
+            (relϕ , relϕ⊩isRelationalϕ)  StrictRelation.isRelational ϕ
+            let
+              closure : ApplStrTerm as 2
+              closure = (` pr₁ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` relϕ ̇ (` pr₂ ̇ (` pr₁ ̇ # one)) ̇ (` pr₁ ̇ (` pr₁ ̇ # one))))
+              realizer : ApplStrTerm as 1
+              realizer =
+                ` pair ̇
+                  (` pair ̇ (` stX ̇ (` pr₁ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
+                  λ*abst closure
+            return
+              (λ* realizer ,
+               { x p r r⊩∃x' 
+                do
+                  (x' , (⊩x~x' , ⊩ϕx) , ⊩x'~x' , ⊩ϕx'≤p , ⊩p≤ϕx')  r⊩∃x'
+                  return
+                    (tt* ,
+                    ((subst
+                       r'  r'   perX .equality  (x , x))
+                      (sym (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _))
+                      (stX⊩isStrictDomainX x x' _ ⊩x~x')) ,
+                     (subst
+                        r'  r'   ϕ .predicate  x)
+                       (sym (cong  x  pr₂  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _))
+                       ⊩ϕx)) ,
+                    λ r' 
+                      let
+                        eq : pr₂  (λ* realizer  r)  r'  pr₁  (pr₂  (pr₂  r))  (relϕ  (pr₂  (pr₁  r))  (pr₁  (pr₁  r)))
+                        eq =
+                          cong  x  pr₂  x  r') (λ*ComputationRule realizer r) 
+                          cong  x  x  r') (pr₂pxy≡y _ _) 
+                          βreduction closure r' (r  [])
+                      in
+                      subst
+                         r'  r'   toPredicate p  tt*)
+                        (sym eq)
+                        (⊩ϕx'≤p _ (relϕ⊩isRelationalϕ x x' _ _ ⊩ϕx ⊩x~x'))) }))
+      in
+      eq/ _ _ (answer , F≤G→G≤F subPer Ωper (composeFuncRel _ _ _ incFuncRel charFuncRel) (composeFuncRel _ _ _ (terminalFuncRel subPer) trueFuncRel) answer)
+
+  module
+    UnivPropWithRepr
+    {Y : Type }
+    (perY : PartialEquivalenceRelation Y)
+    (F : FunctionalRelation perY perX)
+    (G : FunctionalRelation perY terminalPer)
+    (entailment : pointwiseEntailment perY Ωper (composeFuncRel _ _ _ G trueFuncRel) (composeFuncRel _ _ _ F charFuncRel)) where
+
+    opaque
+      unfolding terminalFuncRel
+      G≤idY : pointwiseEntailment perY terminalPer G (terminalFuncRel perY)
+      G≤idY =
+        do
+          (stDG , stDG⊩isStrictDomainG)  G .isStrictDomain
+          return
+            (stDG ,
+             { y tt* r r⊩Gy  stDG⊩isStrictDomainG y tt* r r⊩Gy }))
+
+    opaque
+      idY≤G : pointwiseEntailment perY terminalPer (terminalFuncRel perY) G
+      idY≤G = F≤G→G≤F perY terminalPer G (terminalFuncRel perY) G≤idY
+
+    opaque
+      unfolding trueFuncRel
+      trueFuncRelTruePredicate :  a  (a   trueFuncRel .relation  (tt* , fromPredicate truePredicate))
+      trueFuncRelTruePredicate a = λ b  subst  p  (a  b)   p  tt*) (sym (compIsIdFunc truePredicate)) tt*
+
+    opaque
+      unfolding composeFuncRel
+      unfolding terminalFuncRel
+      {-# TERMINATING #-}
+      H : FunctionalRelation perY subPer
+      Predicate.isSetX (relation H) = isSet× (perY .isSetX) (perX .isSetX)
+      Predicate.∣ relation H  (y , x) r = r   F .relation  (y , x)
+      Predicate.isPropValued (relation H) (y , x) r = F .relation .isPropValued _ _
+      isFunctionalRelation.isStrictDomain (isFuncRel H) =
+        do
+          (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
+          return
+            (stFD ,
+              y x r r⊩Hyx  stFD⊩isStrictDomainF y x r r⊩Hyx))
+      isFunctionalRelation.isStrictCodomain (isFuncRel H) =
+        do
+          (ent , ent⊩entailment)  entailment
+          (a , a⊩idY≤G)  idY≤G
+          (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
+          (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
+          (svF , svF⊩isSingleValuedF)  F .isSingleValued
+          (relϕ , relϕ⊩isRelationalϕ)  StrictRelation.isRelational ϕ
+          let
+            realizer : ApplStrTerm as 1
+            realizer =
+              ` pair ̇
+                (` stFC ̇ # zero) ̇
+                (` relϕ ̇
+                  (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k)))) ̇ ` k) ̇
+                  (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k))) ̇ # zero))
+          return
+            (λ* realizer ,
+              y x r r⊩Hyx 
+               subst
+                  r'  r'   perX .equality  (x , x))
+                 (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                 (stFC⊩isStrictCodomainF y x _ r⊩Hyx) ,
+               (equivFun
+                 (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _))
+                 (do
+                   (x' , ⊩Fyx' , ⊩x'~x' , ⊩ϕx'≤⊤ , ⊩⊤≤ϕx') 
+                     ent⊩entailment
+                     y
+                     (fromPredicate truePredicate)
+                     (pair  (a  (stFD  r))  k)
+                     (return
+                       (tt* ,
+                        subst
+                           r  r   G .relation  (y , tt*))
+                          (sym (pr₁pxy≡x _ _))
+                          (a⊩idY≤G y tt* (stFD  r) (stFD⊩isStrictDomainF y x _ r⊩Hyx))  ,
+                        trueFuncRelTruePredicate _))
+                   let
+                     ⊩x'~x = svF⊩isSingleValuedF y x' x _ _ ⊩Fyx' r⊩Hyx
+                     ⊩ϕx = relϕ⊩isRelationalϕ x' x _ _ (⊩⊤≤ϕx' k (subst  p  k   p  tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x'~x
+                   return (subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx)))))
+      isFunctionalRelation.isRelational (isFuncRel H) =
+        do
+          (relF , relF⊩isRelationalF)  isFunctionalRelation.isRelational (F .isFuncRel)
+          let
+            realizer : ApplStrTerm as 3
+            realizer = ` relF ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
+          return
+            (λ*3 realizer ,
+              y y' x x' a b c ⊩y~y' ⊩Fyx (⊩x~x' , ⊩ϕx) 
+               subst
+                  r'  r'   F .relation  (y' , x'))
+                 (sym (λ*3ComputationRule realizer a b c))
+                 (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x')))
+      isFunctionalRelation.isSingleValued (isFuncRel H) =
+        do
+          (ent , ent⊩entailment)  entailment
+          (a , a⊩idY≤G)  idY≤G
+          (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
+          (svF , svF⊩isSingleValuedF)  F .isSingleValued
+          (relϕ , relϕ⊩isRelationalϕ)  StrictRelation.isRelational ϕ
+          let
+            realizer : ApplStrTerm as 2
+            realizer =
+              ` pair ̇
+                (` svF ̇ # one ̇ # zero) ̇
+                (` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD  ̇ # one)) ̇ ` k)))) ̇ ` k) ̇ (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # one)) ̇ ` k))) ̇ # one))
+          return
+            (λ*2 realizer ,
+              y x x' r₁ r₂ ⊩Fyx ⊩Fyx' 
+               subst
+                  r'  r'   perX .equality  (x , x'))
+                 (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
+                 (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') ,
+               (equivFun
+                 (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _))
+                 (do
+                   (x'' , ⊩Fyx'' , ⊩x''~x'' , ⊩ϕx''≤⊤ , ⊩⊤≤ϕx'') 
+                     ent⊩entailment
+                     y
+                     (fromPredicate truePredicate)
+                     (pair  (a  (stFD  r₁))  k)
+                     (return
+                       (tt* ,
+                        subst  r  r   G .relation  (y , tt*)) (sym (pr₁pxy≡x _ _)) (a⊩idY≤G y tt* _ (stFD⊩isStrictDomainF y x _ ⊩Fyx))  ,
+                        trueFuncRelTruePredicate _))
+                   let
+                     ⊩x''~x = svF⊩isSingleValuedF y x'' x _ _ ⊩Fyx'' ⊩Fyx
+                     ⊩ϕx = relϕ⊩isRelationalϕ x'' x _ _ (⊩⊤≤ϕx'' k (subst  p  k   p  tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x''~x
+                   return
+                     (subst
+                        r'  r'   ϕ .predicate  x)
+                       (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _))
+                       ⊩ϕx)))))
+      isFunctionalRelation.isTotal (isFuncRel H) =
+        do
+          (ent , ent⊩entailment)  entailment
+          (a , a⊩idY≤G)  idY≤G
+          let
+            realizer : ApplStrTerm as 1
+            realizer = ` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ # zero) ̇ ` k))
+          return
+            (λ* realizer ,
+             { y r r⊩y~y 
+              do
+                (x , ⊩Fyx , ⊩x~x , ⊩ϕx≤⊤ , ⊩⊤≤ϕx) 
+                  ent⊩entailment
+                    y
+                    (fromPredicate truePredicate)
+                    (pair  (a  r)  k)
+                    (return
+                      (tt* ,
+                       subst  r  r   G .relation  (y , tt*)) (sym (pr₁pxy≡x _ _)) (a⊩idY≤G y tt* r r⊩y~y)  ,
+                       trueFuncRelTruePredicate _))
+                return (x , subst  r'  r'   F .relation  (y , x)) (sym (λ*ComputationRule realizer r)) ⊩Fyx) }))
+
+    opaque
+      unfolding composeRTMorphism
+      unfolding incFuncRel
+      unfolding H
+      F≡H⋆inc : [ F ]  [ H ]  [ incFuncRel ]
+      F≡H⋆inc =
+        let
+          answer =
+            do
+              (relF , relF⊩isRelationalF)  isFunctionalRelation.isRelational (F .isFuncRel)
+              (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
+              let
+                realizer : ApplStrTerm as 1
+                realizer = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+              return
+                 (λ* realizer ,
+                  y x r ⊩∃x' 
+                   equivFun
+                     (propTruncIdempotent≃ (F .relation .isPropValued _ _))
+                     (do
+                       (x' , ⊩Hyx' , ⊩x'~x , ⊩ϕx')  ⊩∃x'
+                       return
+                         (subst
+                            r'  r'   F .relation  (y , x))
+                           (sym (λ*ComputationRule realizer r))
+                           (relF⊩isRelationalF y y x' x _ _ _ (stFD⊩isStrictDomainF y x' _ ⊩Hyx') ⊩Hyx' ⊩x'~x)))))
+        in eq/ _ _ (F≤G→G≤F perY perX (composeFuncRel _ _ _ H incFuncRel) F answer , answer)
+
+    opaque
+      unfolding composeRTMorphism
+      unfolding terminalFuncRel
+      G≡H⋆terminal : [ G ]  [ H ]  [ terminalFuncRel subPer ]
+      G≡H⋆terminal =
+        let
+          answer =
+            do
+              (stHD , stHD⊩isStrictDomainH)  H .isStrictDomain
+              (a , a⊩idY≤G)  idY≤G
+              let
+                realizer : ApplStrTerm as 1
+                realizer = ` a ̇ (` stHD ̇ (` pr₁ ̇ # zero))
+              return
+                (λ* realizer ,
+                  { y tt* r r⊩∃x 
+                   equivFun
+                     (propTruncIdempotent≃ (G .relation .isPropValued _ _))
+                     (do
+                       (x , ⊩Hyx , ⊩x~x , ⊩ϕx)  r⊩∃x
+                       return (subst  r'  r'   G .relation  (y , tt*)) (sym (λ*ComputationRule realizer r)) (a⊩idY≤G y tt* _ (stHD⊩isStrictDomainH y x _ ⊩Hyx)))) }))
+        in eq/ _ _ (F≤G→G≤F perY terminalPer (composeFuncRel _ _ _ H (terminalFuncRel subPer)) G answer , answer)
+
+    opaque
+      unfolding composeRTMorphism
+      unfolding H
+      unfolding incFuncRel
+      isUniqueH :  (H' : FunctionalRelation perY subPer)  [ F ]  [ H' ]  [ incFuncRel ]  [ G ]  [ H' ]  [ terminalFuncRel subPer ]  [_] {R = bientailment perY subPer} H  [ H' ]
+      isUniqueH H' F≡H'⋆inc G≡H'⋆term =
+        let
+          F≤H'⋆inc = [F]≡[G]→F≤G F (composeFuncRel _ _ _ H' incFuncRel) F≡H'⋆inc
+          answer : pointwiseEntailment _ _ H H'
+          answer =
+            do
+              (a , a⊩F≤H'⋆inc)  F≤H'⋆inc
+              (relH' , relH'⊩isRelationalH)  isFunctionalRelation.isRelational (H' .isFuncRel)
+              (stDH , stDH⊩isStrictDomainH)  H .isStrictDomain
+              let
+                realizer : ApplStrTerm as 1
+                realizer = ` relH' ̇ (` stDH ̇ # zero) ̇ (` pr₁ ̇ (` a ̇ # zero)) ̇ (` pr₂ ̇ (` a ̇ # zero))
+              return
+                (λ* realizer ,
+                  y x r r⊩Hyx 
+                   equivFun
+                     (propTruncIdempotent≃ (H' .relation .isPropValued _ _))
+                     (do
+                       (x' , ⊩H'yx' , ⊩x'~x , ⊩ϕx')  a⊩F≤H'⋆inc y x r r⊩Hyx
+                       return
+                         (subst
+                            r'  r'   H' .relation  (y , x))
+                           (sym (λ*ComputationRule realizer r))
+                           (relH'⊩isRelationalH y y x' x _ _ _ (stDH⊩isStrictDomainH y x r r⊩Hyx) ⊩H'yx' (⊩x'~x , ⊩ϕx'))))))
+        in
+        eq/ _ _ (answer , (F≤G→G≤F _ _ H H' answer))
+
+  opaque
+    classifies : isPullback RT (subobjectCospan [ charFuncRel ]) [ incFuncRel ] [ terminalFuncRel subPer ] subobjectSquareCommutes
+    classifies {Y , perY} f g f⋆char≡g⋆true =
+      SQ.elimProp2
+        {P = λ f g   (commutes : f  [ charFuncRel ]  g  [ trueFuncRel ])  ∃![ hk  RTMorphism perY subPer ] (f  hk  [ incFuncRel ]) × (g  hk  [ terminalFuncRel subPer ])}
+         f g  isPropΠ λ _  isPropIsContr)
+         F G F⋆char≡G⋆true 
+           let
+             entailment = [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G F charFuncRel G trueFuncRel F⋆char≡G⋆true
+           in
+           uniqueExists
+             [ UnivPropWithRepr.H perY F G entailment ]
+             (UnivPropWithRepr.F≡H⋆inc perY F G entailment ,
+             UnivPropWithRepr.G≡H⋆terminal perY F G entailment)
+              hk'  isProp× (squash/ _ _) (squash/ _ _))
+             -- nested eliminator 🤮
+             λ { h' (f≡h'⋆inc , g≡h'⋆term) 
+               SQ.elimProp
+                 {P = λ h'   (comm1 : [ F ]  h'  [ incFuncRel ]) (comm2 : [ G ]  h'  [ terminalFuncRel subPer ])  [ UnivPropWithRepr.H perY F G entailment ]  h'}
+                  h'  isPropΠ λ _  isPropΠ λ _  squash/ _ _)
+                  H' F≡H'⋆inc G≡H'⋆term 
+                   UnivPropWithRepr.isUniqueH perY F G entailment H' F≡H'⋆inc G≡H'⋆term)
+                 h'
+                 f≡h'⋆inc
+                 g≡h'⋆term })
+        f g f⋆char≡g⋆true
+
+  module
+    PullbackHelper
+    (C : FunctionalRelation perX Ωper)
+    (commutes : [ incFuncRel ]  [ C ]  [ terminalFuncRel subPer ]  [ trueFuncRel ])
+    (classifies : isPullback RT (subobjectCospan [ C ]) [ incFuncRel ] [ terminalFuncRel subPer ] commutes) where
+
+    {-# TERMINATING #-}
+    ψ : StrictRelation perX
+    Predicate.isSetX (predicate ψ) = perX .isSetX
+    Predicate.∣ predicate ψ  x r = r   C .relation  (x , )
+    Predicate.isPropValued (predicate ψ) x r = C .relation .isPropValued _ _
+    isStrictRelation.isStrict (isStrictRelationPredicate ψ) =
+      do
+        (stDC , stDC⊩isStrictDomainC)  C .isStrictDomain
+        return
+          (stDC ,
+           λ x r r⊩Cx⊤  stDC⊩isStrictDomainC x (fromPredicate truePredicate) r r⊩Cx⊤)
+    isStrictRelation.isRelational (isStrictRelationPredicate ψ) =
+      do
+        (relC , relC⊩isRelationalC)  isFunctionalRelation.isRelational (C .isFuncRel)
+        (stCC , stCC⊩isStrictCodomainC)  C .isStrictCodomain
+        let
+          realizer : ApplStrTerm as 2
+          realizer = ` relC ̇ # zero ̇ # one ̇ (` stCC ̇ # one)
+        return
+          (λ*2 realizer ,
+           λ x x' a b a⊩Cx⊤ b⊩x~x' 
+             subst  r'  r'   C .relation  (x' , )) (sym (λ*2ComputationRule realizer a b)) (relC⊩isRelationalC x x'   _ _ _ b⊩x~x' a⊩Cx⊤ (stCC⊩isStrictCodomainC x  a a⊩Cx⊤)))
+
+    perψ = InducedSubobject.subPer perX ψ
+    incFuncRelψ = InducedSubobject.incFuncRel perX ψ
+
+    opaque
+      unfolding composeRTMorphism
+      unfolding InducedSubobject.incFuncRel
+      unfolding terminalFuncRel
+      unfolding trueFuncRel
+      pbSqCommutes : [ incFuncRelψ ]  [ C ]  [ terminalFuncRel perψ ]  [ trueFuncRel ]
+      pbSqCommutes =
+        let
+          answer =
+            do
+              (stDC , stDC⊩isStrictDomainC)  C .isStrictDomain
+              (stCC , stCC⊩isStrictCodomainC)  C .isStrictCodomain
+              (svC , svC⊩isSingleValuedC)  C .isSingleValued
+              (relC , relC⊩isRelationalC)  isFunctionalRelation.isRelational (C .isFuncRel)
+              (sX , sX⊩isSymmetricX)  perX .isSymmetric
+              let
+                closure : ApplStrTerm as 2
+                closure = ` pr₁ ̇ (` svC ̇ (` pr₂ ̇ (` pr₁ ̇ # one)) ̇ (` relC ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # one))) ̇ (` pr₂ ̇ # one) ̇ (` stCC ̇ (` pr₂ ̇ # one)))) ̇ ` k
+
+                realizer : ApplStrTerm as 1
+                realizer = ` pair ̇ (` pair ̇ (` stDC ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (λ*abst closure)
+              return
+                (λ* realizer ,
+                 λ { x p r r⊩∃x' 
+                   do
+                     (x' , (⊩x~x' , ⊩Cx⊤) , ⊩Cx'p)  r⊩∃x'
+                     let
+                       ⊩Cxp = relC⊩isRelationalC x' x p p _ _ _ (sX⊩isSymmetricX x x' _ ⊩x~x') ⊩Cx'p (stCC⊩isStrictCodomainC x' p _ ⊩Cx'p) 
+                       (⊩⊤≤p , p≤⊤) = svC⊩isSingleValuedC x  p _ _ ⊩Cx⊤ ⊩Cxp
+                     return
+                       (tt* ,
+                       (subst
+                          r'  r'   perX .equality  (x , x))
+                         (sym
+                           (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r) 
+                            cong  x  pr₁  x) (pr₁pxy≡x _ _) 
+                            pr₁pxy≡x _ _ ))
+                         (stDC⊩isStrictDomainC x  _ ⊩Cx⊤) ,
+                        subst
+                           r'  r'   C .relation  (x , ))
+                          (sym
+                            (cong  x  pr₂  (pr₁  x)) (λ*ComputationRule realizer r) 
+                             cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                             pr₂pxy≡y _ _))
+                          ⊩Cx⊤) ,
+                        λ a 
+                          subst
+                             r'  r'   toPredicate p  tt*)
+                            (sym
+                              (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+                               cong  x  x  a) (pr₂pxy≡y _ _) 
+                               βreduction closure a (r  [])))
+                            (⊩⊤≤p k (subst  q  k   q  tt*) (sym (compIsIdFunc truePredicate)) tt*))) })
+        in eq/ _ _ (answer , F≤G→G≤F _ _ (composeFuncRel _ _ _ incFuncRelψ C) (composeFuncRel _ _ _ (terminalFuncRel perψ) trueFuncRel) answer)
+
+    opaque
+      unfolding InducedSubobject.incFuncRel
+      unfolding composeFuncRel
+      ⊩Cx⊤≤ϕx : ∃[ ent  A ] (∀ (x : X) (r : A)  r   C .relation  (x , )  (ent  r)   ϕ .predicate  x)
+      ⊩Cx⊤≤ϕx =
+        let
+          ((h , incψ≡h⋆incϕ , termψ≡h⋆termϕ) , isUniqueH) = classifies [ incFuncRelψ ] [ terminalFuncRel perψ ] pbSqCommutes
+        in
+        SQ.elimProp
+          {P = λ h   (incψ≡h⋆incϕ : [ incFuncRelψ ]  h  [ incFuncRel ])  ∃[ ent  A ] (∀ (x : X) (r : A)  r   C .relation  (x , )  (ent  r)   ϕ .predicate  x)}
+           h  isPropΠ λ _  isPropPropTrunc)
+           H incψ≡H⋆incϕ 
+            do
+              (a , a⊩incψ≤H⋆incϕ)  [F]≡[G]⋆[H]→F≤G⋆H incFuncRelψ H incFuncRel incψ≡H⋆incϕ
+              (stDC , stDC⊩isStrictDomainC)  C .isStrictDomain
+              (relϕ , relϕ⊩isRelationalϕ)  isStrictRelation.isRelational (ϕ .isStrictRelationPredicate)
+              let
+                realizer = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) 
+              return
+                (λ* realizer ,
+                  x r r⊩Cx⊤ 
+                   equivFun
+                     (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _))
+                     (do
+                       (x' , ⊩Hxx' , ⊩x'~x , ⊩ϕx') 
+                           a⊩incψ≤H⋆incϕ
+                           x x
+                           (pair  (stDC  r)  r)
+                           (subst  r'  r'   perX .equality  (x , x)) (sym (pr₁pxy≡x _ _)) (stDC⊩isStrictDomainC x  r r⊩Cx⊤) ,
+                            subst  r'  r'   C .relation  (x , )) (sym (pr₂pxy≡y _ _)) r⊩Cx⊤)
+                       return
+                         (subst  r'  r'   ϕ .predicate  x) (sym (λ*ComputationRule realizer r)) (relϕ⊩isRelationalϕ x' x _ _ ⊩ϕx' ⊩x'~x))))))
+          h
+          incψ≡h⋆incϕ
+
+  opaque
+    unfolding trueFuncRel
+    unfolding composeFuncRel
+    unfolding incFuncRel
+    unfolding terminalFuncRel
+    isUniqueCharMorphism :
+       (char : RTMorphism perX Ωper)
+       (commutes : [ incFuncRel ]  char  [ terminalFuncRel subPer ]  [ trueFuncRel ])
+       (classifies : isPullback RT (subobjectCospan char) [ incFuncRel ] [ terminalFuncRel subPer ] commutes)
+       char  [ charFuncRel ]
+    isUniqueCharMorphism char commutes classifies =
+      SQ.elimProp
+        {P =
+          λ char 
+           (commutes : [ incFuncRel ]  char  [ terminalFuncRel subPer ]  [ trueFuncRel ])
+            (classifies : isPullback RT (subobjectCospan char) [ incFuncRel ] [ terminalFuncRel subPer ] commutes)
+           char  [ charFuncRel ]}
+         char  isPropΠ λ commutes  isPropΠ λ classifies  squash/ _ _)
+         charFuncRel' commutes classifies 
+          let
+            answer =
+              do
+                (stDX' , stDX'⊩isStrictDomainX')  charFuncRel' .isStrictDomain
+                (relX' , relX'⊩isRelationalX')  isFunctionalRelation.isRelational (charFuncRel' .isFuncRel)
+                (a , a⊩inc⋆X'≤term⋆true)  [F]⋆[G]≡[H]⋆[I]→F⋆G≤H⋆I incFuncRel charFuncRel' (terminalFuncRel subPer) trueFuncRel commutes
+                (b , b⊩term⋆true≤inc⋆X')  [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G incFuncRel charFuncRel' (terminalFuncRel subPer) trueFuncRel commutes
+                (d , d⊩X'x⊤≤ϕx)  PullbackHelper.⊩Cx⊤≤ϕx charFuncRel' commutes classifies
+                let
+                  closure1 : ApplStrTerm as 2
+                  closure1 = ` pr₂ ̇ (` a ̇ (` pair ̇ (` pair ̇ (` stDX'  ̇ # one) ̇ # zero) ̇ # one)) ̇ ` k
+                  closure2 : ApplStrTerm as 2
+                  closure2 = ` d ̇ (` relX' ̇ (` stDX' ̇ # one) ̇ # one ̇ (` pair ̇ ` k ̇ (` k ̇ # zero)))
+                  realizer : ApplStrTerm as 1
+                  realizer = ` pair ̇ (` stDX' ̇ # zero) ̇ (` pair ̇ λ*abst closure1 ̇ λ*abst closure2)
+                return
+                  (λ* realizer ,
+                    { x p r r⊩X'xp 
+                     let
+                       ⊩x~x = stDX'⊩isStrictDomainX' x p r r⊩X'xp
+                     in
+                     subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x  _ _)) ⊩x~x ,
+                      b b⊩ϕx 
+                       let
+                         goal =
+                           a⊩inc⋆X'≤term⋆true
+                             x p (pair  (pair  (stDX'  r)  b)  r)
+                             (return
+                               (x , (subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _)) ⊩x~x ,
+                               subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _)) b⊩ϕx) ,
+                               subst  r'  r'   charFuncRel' .relation  (x , p)) (sym (pr₂pxy≡y _ _)) r⊩X'xp))
+
+                         eq : pr₁  (pr₂  (λ* realizer  r))  b  pr₂  (a  (pair  (pair  (stDX'  r)  b)  r))  k
+                         eq =
+                           cong  x  pr₁  (pr₂  x)  b) (λ*ComputationRule realizer r)  cong  x  pr₁  x  b) (pr₂pxy≡y _ _)  cong  x  x  b) (pr₁pxy≡x _ _)  βreduction closure1 b (r  [])
+                       in
+                       equivFun
+                         (propTruncIdempotent≃ (toPredicate p .isPropValued _ _))
+                         (do
+                           (tt* , ⊩'x~x , ⊩⊤≤p)  goal
+                           return (subst  r'  r'   toPredicate p  tt*) (sym eq) (⊩⊤≤p k)))) ,
+                      c c⊩p 
+                       let
+                         ⊩X'x⊤ =
+                           relX'⊩isRelationalX'
+                             x x p  _ _
+                             (pair  k  (k  c))
+                             ⊩x~x r⊩X'xp
+                             ((λ b b⊩p  subst  q  (pr₁  (pair  k  (k  c)))   q  tt*) (sym (compIsIdFunc truePredicate)) tt*) ,
+                               b b⊩⊤  subst  r'  r'   toPredicate p  tt*) (sym (cong  x  x  b) (pr₂pxy≡y _ _)  kab≡a _ _)) c⊩p))
+
+                         eq : pr₂  (pr₂  (λ* realizer  r))  c  d  (relX'  (stDX'  r)  r  (pair  k  (k  c)))
+                         eq =
+                           cong  x  pr₂  (pr₂  x)  c) (λ*ComputationRule realizer r) 
+                           cong  x  pr₂  x  c) (pr₂pxy≡y _ _) 
+                           cong  x  x  c) (pr₂pxy≡y _ _) 
+                           βreduction closure2 c (r  [])
+                       in
+                       subst
+                          r'  r'   ϕ .predicate  x)
+                         (sym eq)
+                         (d⊩X'x⊤≤ϕx x _ ⊩X'x⊤)) }))
+          in eq/ _ _ (answer , F≤G→G≤F _ _ charFuncRel' charFuncRel answer))
+        char
+        commutes
+        classifies
+
\ No newline at end of file diff --git a/docs/Realizability.Topos.TerminalObject.html b/docs/Realizability.Topos.TerminalObject.html new file mode 100644 index 0000000..0eb2b9f --- /dev/null +++ b/docs/Realizability.Topos.TerminalObject.html @@ -0,0 +1,112 @@ + +Realizability.Topos.TerminalObject
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Unit
+open import Cubical.Data.Empty
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Limits.Terminal
+
+module Realizability.Topos.TerminalObject
+  { ℓ' ℓ''}
+  {A : Type }
+  (ca : CombinatoryAlgebra A)
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
+
+open CombinatoryAlgebra ca
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Topos.Object {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
+
+open Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open PartialEquivalenceRelation
+open Predicate renaming (isSetX to isSetPredicateBase)
+
+opaque
+  terminalPer : PartialEquivalenceRelation Unit*
+  isSetPredicateBase (equality terminalPer) = isSet× isSetUnit* isSetUnit*
+   equality terminalPer  (tt* , tt*) _ = Unit*
+  isPropValued (equality terminalPer) _ _ = isPropUnit*
+  isPartialEquivalenceRelation.isSetX (isPerEquality terminalPer) = isSetUnit*
+  isPartialEquivalenceRelation.isSymmetric (isPerEquality terminalPer) =
+    return (k ,  { tt* tt* r tt*  tt* }))
+  isPartialEquivalenceRelation.isTransitive (isPerEquality terminalPer) =
+    return (k ,  { tt* tt* tt* _ _ tt* tt*  tt* }))
+
+open FunctionalRelation
+
+opaque
+  unfolding terminalPer
+  terminalFuncRel :  {Y : Type ℓ'}  (perY : PartialEquivalenceRelation Y)  FunctionalRelation perY terminalPer
+  terminalFuncRel {Y} perY =
+    record
+      { relation =
+        record
+          { isSetX = isSet× (perY .isSetX) isSetUnit*
+          ; ∣_∣ = λ { (y , tt*) r  r   perY .equality  (y , y) }
+          ; isPropValued = λ { (y , tt*) r  perY .equality .isPropValued _ _ } }
+      ; isFuncRel =
+        record
+          { isStrictDomain = return (Id ,  { y tt* r r⊩y~y  subst  r'  r'   perY .equality  (y , y)) (sym (Ida≡a _)) r⊩y~y }))
+          ; isStrictCodomain = return (k ,  { y tt* r r⊩y~y  tt* }))
+          ; isRelational =
+            (do
+            (t , t⊩isTransitive)  perY .isTransitive
+            (s , s⊩isSymmetric)  perY .isSymmetric
+            let
+              prover : ApplStrTerm as 3
+              prover = ` t ̇ (` s ̇ # two) ̇ # two
+            return
+              (λ*3 prover ,
+               { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt* 
+                subst
+                   r'  r'   perY .equality  (y' , y'))
+                  (sym (λ*3ComputationRule prover a b c))
+                  (t⊩isTransitive y' y y' (s  a) a (s⊩isSymmetric y y' a a⊩y~y') a⊩y~y') })))
+          ; isSingleValued = (return (k ,  { y tt* tt* r₁ r₂ r₁⊩y~y r₂⊩y~y  tt* })))
+          ; isTotal = return
+                      (Id ,
+                       y r r⊩y~y 
+                        return (tt* , (subst  r'  r'   perY .equality  (y , y)) (sym (Ida≡a _)) r⊩y~y))))
+                                    } }
+opaque
+  unfolding terminalPer
+  isTerminalTerminalPer :  {Y : Type ℓ'}  (perY : PartialEquivalenceRelation Y)  isContr (RTMorphism perY terminalPer)
+  isTerminalTerminalPer {Y} perY =
+    inhProp→isContr
+      [ terminalFuncRel perY ]
+      λ f g 
+        SQ.elimProp2
+           f g  squash/ f g)
+           F G 
+            let
+              answer : pointwiseEntailment perY terminalPer F G
+              answer =
+                do
+                  (tlG , tlG⊩isTotalG)  G .isTotal
+                  (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
+                  let
+                    prover : ApplStrTerm as 1
+                    prover = ` tlG ̇ (` stFD ̇ # zero)
+                  return
+                    (λ* prover ,
+                     { y tt* r r⊩Fy 
+                      transport
+                        (propTruncIdempotent (G .relation .isPropValued _ _))
+                        (do
+                          (tt* , tlGstFD⊩Gy)  tlG⊩isTotalG y (stFD  r) (stFD⊩isStrictDomainF y tt* r r⊩Fy)
+                          return (subst  r'  r'   G .relation  (y , tt*)) (sym (λ*ComputationRule prover r)) tlGstFD⊩Gy)) }))
+            in
+            eq/ _ _ (answer , F≤G→G≤F perY terminalPer F G answer))
+          f g
+
+TerminalRT : Terminal RT
+TerminalRT =
+  (Unit* , terminalPer) ,  { (Y , perY)  isTerminalTerminalPer perY})
+
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html b/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html index 2f2b484..f96528d 100644 --- a/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html +++ b/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html @@ -1,79 +1,76 @@ Realizability.Tripos.Prealgebra.Meets.Commutativity
open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Equiv
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum renaming (rec to sumRec)
-open import Cubical.Relation.Binary.Order.Preorder
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
+open import Realizability.ApplicativeStructure
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Equiv
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum renaming (rec to sumRec)
+open import Cubical.Relation.Binary.Order.Preorder
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
 
-module Realizability.Tripos.Prealgebra.Meets.Commutativity {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Tripos.Prealgebra.Meets.Commutativity { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open import Realizability.Tripos.Prealgebra.Predicate ca
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
 
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
-
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where
+module _ (X : Type ℓ') (isSetX' : isSet X) where
   
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  open PreorderReasoning preorder≤
+  private PredicateX = Predicate  X
+  open Predicate
+  open PredicateProperties  X
+  open PreorderReasoning preorder≤
 
-  x⊓y≤y⊓x :  x y  x  y  y  x
-  x⊓y≤y⊓x x y =
-    do
-      let
-        proof : Term as 1
-        proof = ` pair ̇ (` pr₂ ̇ # fzero) ̇ (` pr₁ ̇ # fzero)
-      return
-        (λ* proof ,
-           x' a a⊩x⊓y 
-            subst
-               r  r   y  x  x')
-              (sym (λ*ComputationRule proof (a  []) ))
-              ((subst  r  r   y  x') (sym (pr₁pxy≡x _ _)) (a⊩x⊓y .snd)) ,
-               (subst  r  r   x  x') (sym (pr₂pxy≡y _ _)) (a⊩x⊓y .fst)))))
+  x⊓y≤y⊓x :  x y  x  y  y  x
+  x⊓y≤y⊓x x y =
+    do
+      let
+        proof : Term as 1
+        proof = ` pair ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero)
+      return
+        (λ* proof ,
+           x' a a⊩x⊓y 
+            subst
+               r  r   y  x  x')
+              (sym (λ*ComputationRule proof a))
+              ((subst  r  r   y  x') (sym (pr₁pxy≡x _ _)) (a⊩x⊓y .snd)) ,
+               (subst  r  r   x  x') (sym (pr₂pxy≡y _ _)) (a⊩x⊓y .fst)))))
 
-  antiSym→x⊓z≤y⊓z :  x y z  x  y  y  x  x  z  y  z
-  antiSym→x⊓z≤y⊓z x y z x≤y y≤x =
-    do
-      (f , f⊩x≤y)  x≤y
-      (g , g⊩y≤x)  y≤x
-      let
-        proof : Term as 1
-        proof = ` pair ̇ (` f ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero)
-      return
-        ((λ* proof) ,
-           x' a a⊩x⊓z 
-            subst
-               r  r   y  z  x')
-              (sym (λ*ComputationRule proof (a  [])))
-              ((subst  r  r   y  x') (sym (pr₁pxy≡x _ _)) (f⊩x≤y x' (pr₁  a) (a⊩x⊓z .fst))) ,
-               (subst  r  r   z  x') (sym (pr₂pxy≡y _ _)) (a⊩x⊓z .snd)))))
+  antiSym→x⊓z≤y⊓z :  x y z  x  y  y  x  x  z  y  z
+  antiSym→x⊓z≤y⊓z x y z x≤y y≤x =
+    do
+      (f , f⊩x≤y)  x≤y
+      (g , g⊩y≤x)  y≤x
+      let
+        proof : Term as 1
+        proof = ` pair ̇ (` f ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+      return
+        ((λ* proof) ,
+           x' a a⊩x⊓z 
+            subst
+               r  r   y  z  x')
+              (sym (λ*ComputationRule proof a))
+              ((subst  r  r   y  x') (sym (pr₁pxy≡x _ _)) (f⊩x≤y x' (pr₁  a) (a⊩x⊓z .fst))) ,
+               (subst  r  r   z  x') (sym (pr₂pxy≡y _ _)) (a⊩x⊓z .snd)))))
 
 
-  antiSym→x⊓y≤x⊓z :  x y z  y  z  z  y  x  y  x  z
-  antiSym→x⊓y≤x⊓z x y z y≤z z≤y =
-    do
-      (f , f⊩y≤z)  y≤z
-      (g , g⊩z≤y)  z≤y
-      let
-        proof : Term as 1
-        proof = ` pair ̇ (`  pr₁ ̇ # fzero) ̇ (` f ̇ (` pr₂ ̇ # fzero))
-      return
-        ((λ* proof) ,
-           { x' a (pr₁a⊩x , pr₂a⊩y) 
-            subst
-               r  r   x  z  x')
-              (sym (λ*ComputationRule proof (a  [])))
-              ((subst  r  r   x  x') (sym (pr₁pxy≡x _ _)) pr₁a⊩x) ,
-               (subst  r  r   z  x') (sym (pr₂pxy≡y _ _)) (f⊩y≤z x' (pr₂  a) pr₂a⊩y))) }))
+  antiSym→x⊓y≤x⊓z :  x y z  y  z  z  y  x  y  x  z
+  antiSym→x⊓y≤x⊓z x y z y≤z z≤y =
+    do
+      (f , f⊩y≤z)  y≤z
+      (g , g⊩z≤y)  z≤y
+      let
+        proof : Term as 1
+        proof = ` pair ̇ (`  pr₁ ̇ # zero) ̇ (` f ̇ (` pr₂ ̇ # zero))
+      return
+        ((λ* proof) ,
+           { x' a (pr₁a⊩x , pr₂a⊩y) 
+            subst
+               r  r   x  z  x')
+              (sym (λ*ComputationRule proof a))
+              ((subst  r  r   x  x') (sym (pr₁pxy≡x _ _)) pr₁a⊩x) ,
+               (subst  r  r   z  x') (sym (pr₂pxy≡y _ _)) (f⊩y≤z x' (pr₂  a) pr₂a⊩y))) }))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html b/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html index 62ba9bd..90b9965 100644 --- a/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html +++ b/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html @@ -1,75 +1,72 @@ Realizability.Tripos.Prealgebra.Meets.Identity
open import Cubical.Foundations.Prelude
 open import Cubical.Data.Unit
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum
-open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum
+open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure 
 
-module Realizability.Tripos.Prealgebra.Meets.Identity {} {A : Type } (ca : CombinatoryAlgebra A) where
-open import Realizability.Tripos.Prealgebra.Predicate ca
-open import Realizability.Tripos.Prealgebra.Meets.Commutativity ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module Realizability.Tripos.Prealgebra.Meets.Identity { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Meets.Commutativity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
+module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
+  private PredicateX = Predicate X
+  open Predicate
+  open PredicateProperties X
+  open PreorderReasoning preorder≤
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  open PreorderReasoning preorder≤
+  pre1 : PredicateX
+  Predicate.isSetX pre1 = isSetX'
+  Predicate.∣ pre1  = λ x a  Unit*
+  Predicate.isPropValued pre1 = λ x a  isPropUnit*
 
-  pre1 : PredicateX
-  Predicate.isSetX pre1 = isSetX'
-  Predicate.∣ pre1  = λ x a  Unit*
-  Predicate.isPropValued pre1 = λ x a  isPropUnit*
+  x⊓1≤x :  x  x  pre1  x
+  x⊓1≤x x =  pr₁ ,  x' a a⊩x⊓1  a⊩x⊓1 .fst) ∣₁
 
-  x⊓1≤x :  x  x  pre1  x
-  x⊓1≤x x =  pr₁ ,  x' a a⊩x⊓1  a⊩x⊓1 .fst) ∣₁
+  x≤x⊓1 :  x  x  x  pre1
+  x≤x⊓1 x =
+    do
+      let
+        proof : Term as 1
+        proof = ` pair ̇ # zero ̇ ` true
+      return
+        ((λ* proof) ,
+           x' a a⊩x 
+            subst
+               r   x  pre1  x' r)
+              (sym (λ*ComputationRule proof a))
+              (subst
+                 r  r   x  x')
+                (sym (pr₁pxy≡x _ _))
+                a⊩x , tt*)))
 
-  x≤x⊓1 :  x  x  x  pre1
-  x≤x⊓1 x =
-    do
-      let
-        proof : Term as 1
-        proof = ` pair ̇ # fzero ̇ ` true
-      return
-        ((λ* proof) ,
-           x' a a⊩x 
-            subst
-               r   x  pre1  x' r)
-              (sym (λ*ComputationRule proof (a  [])))
-              (subst
-                 r  r   x  x')
-                (sym (pr₁pxy≡x _ _))
-                a⊩x , tt*)))
+  1⊓x≤x :  x  pre1  x  x
+  1⊓x≤x x =  pr₂ ,  x' a a⊩1⊓x  a⊩1⊓x .snd) ∣₁
 
-  1⊓x≤x :  x  pre1  x  x
-  1⊓x≤x x =  pr₂ ,  x' a a⊩1⊓x  a⊩1⊓x .snd) ∣₁
-
-  x≤1⊓x :  x  x  pre1  x
-  x≤1⊓x x =
-    do
-      let
-        proof : Term as 1
-        proof = ` pair ̇ ` false ̇ # fzero
-      return
-        ((λ* proof) ,
-           x' a a⊩x 
-            subst
-               r  r   pre1  x  x')
-              (sym (λ*ComputationRule proof (a  [])))
-              (tt* ,
-              (subst
-                 r  r   x  x')
-                (sym (pr₂pxy≡y _ _))
-                a⊩x))))
+  x≤1⊓x :  x  x  pre1  x
+  x≤1⊓x x =
+    do
+      let
+        proof : Term as 1
+        proof = ` pair ̇ ` false ̇ # zero
+      return
+        ((λ* proof) ,
+           x' a a⊩x 
+            subst
+               r  r   pre1  x  x')
+              (sym (λ*ComputationRule proof a))
+              (tt* ,
+              (subst
+                 r  r   x  x')
+                (sym (pr₂pxy≡y _ _))
+                a⊩x))))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Predicate.Base.html b/docs/Realizability.Tripos.Prealgebra.Predicate.Base.html index b629916..6ff832b 100644 --- a/docs/Realizability.Tripos.Prealgebra.Predicate.Base.html +++ b/docs/Realizability.Tripos.Prealgebra.Predicate.Base.html @@ -1,72 +1,76 @@ Realizability.Tripos.Prealgebra.Predicate.Base
open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (⟦_⟧ to `⟦_⟧; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Univalence
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Data.Sigma
-open import Cubical.Functions.FunExtEquiv
+open import Realizability.ApplicativeStructure
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Equiv
+open import Cubical.Data.Sigma
+open import Cubical.Functions.FunExtEquiv
 
-module Realizability.Tripos.Prealgebra.Predicate.Base {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Tripos.Prealgebra.Predicate.Base { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-record Predicate {ℓ' ℓ''} (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
-  field
-    isSetX : isSet X
-    ∣_∣ : X  A  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
-    isPropValued :  x a  isProp (∣_∣ x a)
-  infix 25 ∣_∣
+record Predicate (X : Type ℓ') : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
+  constructor makePredicate
+  field
+    isSetX : isSet X
+    ∣_∣ : X  A  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
+    isPropValued :  x a  isProp (∣_∣ x a)
+  infix 25 ∣_∣
 
-open Predicate
-infix 26 _⊩_
-_⊩_ :  {ℓ'}  A  (A  Type ℓ')  Type ℓ'
-a  ϕ = ϕ a
+open Predicate
+infix 26 _⊩_
+_⊩_ :  {ℓ'}  A  (A  Type ℓ')  Type ℓ'
+a  ϕ = ϕ a
 
-PredicateΣ :  {ℓ' ℓ''}  (X : Type ℓ')  Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ''))
-PredicateΣ {ℓ'} {ℓ''} X = Σ[ rel  (X  A  hProp (ℓ-max (ℓ-max  ℓ') ℓ'')) ] (isSet X)
+PredicateΣ :  (X : Type ℓ')  Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ''))
+PredicateΣ X = Σ[ rel  (X  A  hProp (ℓ-max (ℓ-max  ℓ') ℓ'')) ] (isSet X)
 
-isSetPredicateΣ :  {ℓ' ℓ''} (X : Type ℓ')  isSet (PredicateΣ {ℓ'' = ℓ''} X)
-isSetPredicateΣ X = isSetΣ (isSetΠ  x  isSetΠ λ a  isSetHProp)) λ _  isProp→isSet isPropIsSet
+isSetPredicateΣ :  (X : Type ℓ')  isSet (PredicateΣ X)
+isSetPredicateΣ X = isSetΣ (isSetΠ  x  isSetΠ λ a  isSetHProp)) λ _  isProp→isSet isPropIsSet
 
-PredicateIsoΣ :  {ℓ' ℓ''} (X : Type ℓ')  Iso (Predicate {ℓ'' = ℓ''} X) (PredicateΣ {ℓ'' = ℓ''} X)
-PredicateIsoΣ {ℓ'} {ℓ''} X =
-  iso
-     p   x a  (a   p  x) , p .isPropValued x a) , p .isSetX)
-     p  record { isSetX = p .snd ; ∣_∣ = λ x a  p .fst x a .fst ; isPropValued = λ x a  p .fst x a .snd })
-     b  refl)
-    λ a  refl
+PredicateIsoΣ :  (X : Type ℓ')  Iso (Predicate  X) (PredicateΣ  X)
+PredicateIsoΣ X =
+  iso
+     p   x a  (a   p  x) , p .isPropValued x a) , p .isSetX)
+     p  record { isSetX = p .snd ; ∣_∣ = λ x a  p .fst x a .fst ; isPropValued = λ x a  p .fst x a .snd })
+     b  refl)
+    λ a  refl
 
-Predicate≡PredicateΣ :  {ℓ' ℓ''} (X : Type ℓ')  Predicate {ℓ'' = ℓ''} X  PredicateΣ {ℓ'' = ℓ''} X
-Predicate≡PredicateΣ {ℓ'} {ℓ''} X = isoToPath (PredicateIsoΣ X)
+Predicate≃PredicateΣ :  (X : Type ℓ')  Predicate X  PredicateΣ X
+Predicate≃PredicateΣ X = isoToEquiv (PredicateIsoΣ X)
 
-isSetPredicate :  {ℓ' ℓ''} (X : Type ℓ')  isSet (Predicate {ℓ'' = ℓ''} X)
-isSetPredicate {ℓ'} {ℓ''} X = subst  predicateType  isSet predicateType) (sym (Predicate≡PredicateΣ X)) (isSetPredicateΣ {ℓ'' = ℓ''} X)
+Predicate≡PredicateΣ :  (X : Type ℓ')  Predicate  X  PredicateΣ  X
+Predicate≡PredicateΣ X = isoToPath (PredicateIsoΣ X)
 
-PredicateΣ≡ :  {ℓ' ℓ''} (X : Type ℓ')  (P Q : PredicateΣ {ℓ'' = ℓ''} X)  (P .fst  Q .fst)  P  Q
-PredicateΣ≡ X P Q ∣P∣≡∣Q∣ = Σ≡Prop  _  isPropIsSet) ∣P∣≡∣Q∣
+isSetPredicate :   (X : Type ℓ')  isSet (Predicate  X)
+isSetPredicate X = subst  predicateType  isSet predicateType) (sym (Predicate≡PredicateΣ X)) (isSetPredicateΣ  X)
 
-Predicate≡ :
-   {ℓ' ℓ''} (X : Type ℓ')
-   (P Q : Predicate {ℓ'' = ℓ''} X)
-   (∀ x a  a   P  x  a   Q  x)
-   (∀ x a  a   Q  x  a   P  x)
-  -----------------------------------
-   P  Q
-Predicate≡ {ℓ'} {ℓ''} X P Q P→Q Q→P i =
-  PredicateIsoΣ X .inv
-    (PredicateΣ≡
-      {ℓ'' = ℓ''}
-      X
-      (PredicateIsoΣ X .fun P)
-      (PredicateIsoΣ X .fun Q)
-      (funExt₂
-         x a  Σ≡Prop  A  isPropIsProp)
-        (hPropExt
-          (P .isPropValued x a)
-          (Q .isPropValued x a)
-          (P→Q x a)
-          (Q→P x a)))) i) where open Iso
+PredicateΣ≡ :   (X : Type ℓ')  (P Q : PredicateΣ  X)  (P .fst  Q .fst)  P  Q
+PredicateΣ≡ X P Q ∣P∣≡∣Q∣ = Σ≡Prop  _  isPropIsSet) ∣P∣≡∣Q∣
+
+Predicate≡ :
+   (X : Type ℓ')
+   (P Q : Predicate  X)
+   (∀ x a  a   P  x  a   Q  x)
+   (∀ x a  a   Q  x  a   P  x)
+  -----------------------------------
+   P  Q
+Predicate≡ X P Q P→Q Q→P i =
+  PredicateIsoΣ X .inv
+    (PredicateΣ≡
+      X
+      (PredicateIsoΣ X .fun P)
+      (PredicateIsoΣ X .fun Q)
+      (funExt₂
+         x a  Σ≡Prop  A  isPropIsProp)
+        (hPropExt
+          (P .isPropValued x a)
+          (Q .isPropValued x a)
+          (P→Q x a)
+          (Q→P x a)))) i) where open Iso
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html b/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html index 8131a20..50a1a61 100644 --- a/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html +++ b/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html @@ -1,321 +1,362 @@ Realizability.Tripos.Prealgebra.Predicate.Properties
open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Univalence
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Function
-open import Cubical.Functions.FunExtEquiv
-open import Cubical.Data.Sigma
-open import Cubical.Data.Empty
-open import Cubical.Data.Unit
-open import Cubical.Data.Sum
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-
-module
-  Realizability.Tripos.Prealgebra.Predicate.Properties
-  {} {A : Type } (ca : CombinatoryAlgebra A) where
-
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca
-
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Predicate
-module PredicateProperties {ℓ' ℓ''} (X : Type ℓ') where
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  _≤_ : Predicate {ℓ'' = ℓ''} X  Predicate {ℓ'' = ℓ''} X  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
-  ϕ  ψ = ∃[ b  A ] (∀ (x : X) (a : A)  a  ( ϕ  x)  (b  a)   ψ  x)
-
-  isProp≤ :  ϕ ψ  isProp (ϕ  ψ)
-  isProp≤ ϕ ψ = isPropPropTrunc
-
-  isRefl≤ :  ϕ  ϕ  ϕ
-  isRefl≤ ϕ =  Id ,  x a a⊩ϕx  subst  r  r   ϕ  x) (sym (Ida≡a a)) a⊩ϕx) ∣₁
-
-  isTrans≤ :  ϕ ψ ξ  ϕ  ψ  ψ  ξ  ϕ  ξ
-  isTrans≤ ϕ ψ ξ ϕ≤ψ ψ≤ξ = do
-                           (a , ϕ≤[a]ψ)  ϕ≤ψ
-                           (b , ψ≤[b]ξ)  ψ≤ξ
-                           return
-                             ((B b a) ,
-                              x a' a'⊩ϕx 
-                               subst
-                                  r  r   ξ  x)
-                                 (sym (Ba≡gfa b a a'))
-                                 (ψ≤[b]ξ x (a  a')
-                                 (ϕ≤[a]ψ x a' a'⊩ϕx))))
+open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Cubical.Foundations.Prelude as P
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty
+open import Cubical.Data.Unit
+open import Cubical.Data.Sum
+open import Cubical.Data.Vec
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+
+module
+  Realizability.Tripos.Prealgebra.Predicate.Properties
+  { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Tripos.Prealgebra.Predicate.Base { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate
+
+module PredicateProperties (X : Type ℓ') where
+  private PredicateX = Predicate X
+  open Predicate
+  _≤_ : Predicate  X  Predicate  X  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
+  ϕ  ψ = ∃[ b  A ] (∀ (x : X) (a : A)  a  ( ϕ  x)  (b  a)   ψ  x)
+
+  isProp≤ :  ϕ ψ  isProp (ϕ  ψ)
+  isProp≤ ϕ ψ = isPropPropTrunc
+
+  isRefl≤ :  ϕ  ϕ  ϕ
+  isRefl≤ ϕ =  Id ,  x a a⊩ϕx  subst  r  r   ϕ  x) (sym (Ida≡a a)) a⊩ϕx) ∣₁
+
+  isTrans≤ :  ϕ ψ ξ  ϕ  ψ  ψ  ξ  ϕ  ξ
+  isTrans≤ ϕ ψ ξ ϕ≤ψ ψ≤ξ = do
+                           (a , ϕ≤[a]ψ)  ϕ≤ψ
+                           (b , ψ≤[b]ξ)  ψ≤ξ
+                           return
+                             ((B b a) ,
+                              x a' a'⊩ϕx 
+                               subst
+                                  r  r   ξ  x)
+                                 (sym (Ba≡gfa b a a'))
+                                 (ψ≤[b]ξ x (a  a')
+                                 (ϕ≤[a]ψ x a' a'⊩ϕx))))
     
 
-  open IsPreorder renaming
-    (is-set to isSet
-    ;is-prop-valued to isPropValued
-    ;is-refl to isRefl
-    ;is-trans to isTrans)
+  open IsPreorder renaming
+    (is-set to isSet
+    ;is-prop-valued to isPropValued
+    ;is-refl to isRefl
+    ;is-trans to isTrans)
 
-  preorder≤ : _
-  preorder≤ = preorder (Predicate X) _≤_ (ispreorder (isSetPredicate X) isProp≤ isRefl≤ isTrans≤)
+  preorder≤ : _
+  preorder≤ = preorder (Predicate X) _≤_ (ispreorder (isSetPredicate X) isProp≤ isRefl≤ isTrans≤)
 
-  {-
+  {-
   It is not necessary to truncate the underlying predicate but it is very convenient.
   We can prove that the underlying type is a proposition if the combinatory algebra
   is non-trivial. This would require some effort to do in Agda, so I have deferred it
   for later.
   -}
-  infix 25 _⊔_
-  _⊔_ : PredicateX  PredicateX  PredicateX
-  (ϕ  ψ) .isSetX = ϕ .isSetX
-   ϕ  ψ  x a =  ((pr₁  a  k) × ((pr₂  a)   ϕ  x))  ((pr₁  a  k') × ((pr₂  a)   ψ  x)) ∥₁
-  (ϕ  ψ) .isPropValued x a = isPropPropTrunc
-
-  infix 25 _⊓_
-  _⊓_ : PredicateX  PredicateX  PredicateX
-  (ϕ  ψ) .isSetX = ϕ .isSetX
-   ϕ  ψ  x a = ((pr₁  a)   ϕ  x) × ((pr₂  a)   ψ  x)
-  (ϕ  ψ) .isPropValued x a = isProp× (ϕ .isPropValued x (pr₁  a)) (ψ .isPropValued x (pr₂  a))
-
-  infix 25 _⇒_
-  _⇒_ : PredicateX  PredicateX  PredicateX
-  (ϕ  ψ) .isSetX = ϕ .isSetX
-   ϕ  ψ  x a =  b  (b   ϕ  x)  (a  b)   ψ  x
-  (ϕ  ψ) .isPropValued x a = isPropΠ λ a  isPropΠ λ a⊩ϕx  ψ .isPropValued _ _
-
-
-module Morphism {ℓ' ℓ''} {X Y : Type ℓ'} (isSetX : isSet X) (isSetY : isSet Y)  where
-  PredicateX = Predicate {ℓ'' = ℓ''} X
-  PredicateY = Predicate {ℓ'' = ℓ''} Y
-  module PredicatePropertiesX = PredicateProperties {ℓ'' = ℓ''} X
-  module PredicatePropertiesY = PredicateProperties {ℓ'' = ℓ''} Y
-  open PredicatePropertiesX renaming (_≤_ to _≤X_ ; isProp≤ to isProp≤X)
-  open PredicatePropertiesY renaming (_≤_ to _≤Y_ ; isProp≤ to isProp≤Y)
-  open Predicate hiding (isSetX)
-
-  ⋆_ : (X  Y)  (PredicateY  PredicateX)
-   f =
-    λ ϕ 
-      record
-        { isSetX = isSetX
-        ; ∣_∣ = λ x a  a   ϕ  (f x)
-        ; isPropValued = λ x a  ϕ .isPropValued (f x) a }
-
-  `∀[_] : (X  Y)  (PredicateX  PredicateY)
-  `∀[ f ] =
-    λ ϕ 
-      record
-        { isSetX = isSetY
-        ; ∣_∣ = λ y a  (∀ b x  f x  y  (a  b)   ϕ  x)
-        ; isPropValued = λ y a  isPropΠ λ a'  isPropΠ λ x  isPropΠ λ fx≡y  ϕ .isPropValued x (a  a') }
-
-  `∃[_] : (X  Y)  (PredicateX  PredicateY)
-  `∃[ f ] =
-    λ ϕ 
-      record
-        { isSetX = isSetY
-        ; ∣_∣ = λ y a  ∃[ x  X ] (f x  y) × (a   ϕ  x)
-        ; isPropValued = λ y a  isPropPropTrunc }
-
-  -- Adjunction proofs
-
-  `∃isLeftAdjoint→ :  ϕ ψ f  `∃[ f ] ϕ ≤Y ψ  ϕ ≤X ( f) ψ
-  `∃isLeftAdjoint→ ϕ ψ f p =
-    do
-      (a~ , a~proves)  p
-      return (a~ ,  x a a⊩ϕx  a~proves (f x) a  x , refl , a⊩ϕx ∣₁))
-
-
-  `∃isLeftAdjoint← :  ϕ ψ f  ϕ ≤X ( f) ψ  `∃[ f ] ϕ ≤Y ψ
-  `∃isLeftAdjoint← ϕ ψ f p =
-    do
-      (a~ , a~proves)  p
-      return
-        (a~ ,
-         y b b⊩∃fϕ 
-          equivFun
-            (propTruncIdempotent≃
-              (ψ .isPropValued y (a~  b)))
-              (do
-                (x , fx≡y , b⊩ϕx)  b⊩∃fϕ
-                return (subst  y'  (a~  b)   ψ  y') fx≡y (a~proves x b b⊩ϕx)))))
-
-  `∃isLeftAdjoint :  ϕ ψ f  `∃[ f ] ϕ ≤Y ψ  ϕ ≤X ( f) ψ
-  `∃isLeftAdjoint ϕ ψ f =
-    hPropExt
-      (isProp≤Y (`∃[ f ] ϕ) ψ)
-      (isProp≤X ϕ (( f) ψ))
-      (`∃isLeftAdjoint→ ϕ ψ f)
-      (`∃isLeftAdjoint← ϕ ψ f)
-
-  `∀isRightAdjoint→ :  ϕ ψ f  ψ ≤Y `∀[ f ] ϕ  ( f) ψ ≤X ϕ
-  `∀isRightAdjoint→ ϕ ψ f p =
-    do
-      (a~ , a~proves)  p
-      let realizer = (s  (s  (k  a~)  Id)  (k  k))
-      return
-        (realizer ,
-         x a a⊩ψfx 
-          equivFun
-            (propTruncIdempotent≃
-              (ϕ .isPropValued x (realizer  a) ))
-              (do
-                let ∀prover = a~proves (f x) a a⊩ψfx
-                return
-                  (subst
-                     ϕ~  ϕ~   ϕ  x)
-                    (sym
-                      (realizer  a
-                        ≡⟨ refl 
-                       s  (s  (k  a~)  Id)  (k  k)  a
-                        ≡⟨ sabc≡ac_bc _ _ _ 
-                       s  (k  a~)  Id  a  (k  k  a)
-                        ≡⟨ cong  x  x  (k  k  a)) (sabc≡ac_bc _ _ _) 
-                       k  a~  a  (Id  a)  (k  k  a)
-                        ≡⟨ cong  x  k  a~  a  x  (k  k  a)) (Ida≡a a) 
-                       k  a~  a  a  (k  k  a)
-                        ≡⟨ cong  x  k  a~  a  a  x) (kab≡a _ _) 
-                       (k  a~  a)  a  k
-                        ≡⟨ cong  x  x  a  k) (kab≡a _ _) 
-                       a~  a  k
-                         ))
-                    (∀prover k x refl)))))
-
-  `∀isRightAdjoint← :  ϕ ψ f  ( f) ψ ≤X ϕ  ψ ≤Y `∀[ f ] ϕ
-  `∀isRightAdjoint← ϕ ψ f p =
-    do
-      (a~ , a~proves)  p
-      let realizer = (s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id))
-      return
-        (realizer ,
-         y b b⊩ψy a x fx≡y 
-          subst
-             r  r   ϕ  x)
-            (sym
-              (realizer  b  a
-                 ≡⟨ refl 
-               s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id)  b  a
-                 ≡⟨ cong  x  x  a) (sabc≡ac_bc _ _ _) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  (s  (k  k)  Id  b)  a
-                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  x  a) (sabc≡ac_bc (k  k) Id b) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  ((k  k  b)  (Id  b))  a
-                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (x  (Id  b))  a) (kab≡a _ _) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  (Id  b))  a
-                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (k  x)  a) (Ida≡a b) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  b)  a
-                 ≡⟨ cong  x  x  (k  b)  a) (sabc≡ac_bc _ _ _) 
-               k  s  b  (s  (k  k)  (k  a~)  b)  (k  b)  a
-                 ≡⟨ cong  x  x  (s  (k  k)  (k  a~)  b)  (k  b)  a) (kab≡a _ _) 
-               s  (s  (k  k)  (k  a~)  b)  (k  b)  a
-                 ≡⟨ sabc≡ac_bc _ _ _ 
-               s  (k  k)  (k  a~)  b  a  (k  b  a)
-                 ≡⟨ cong  x  s  (k  k)  (k  a~)  b  a  x) (kab≡a b a) 
-               s  (k  k)  (k  a~)  b  a  b
-                 ≡⟨ cong  x  x  a  b) (sabc≡ac_bc (k  k) (k  a~) b) 
-               k  k  b  (k  a~  b)  a  b
-                 ≡⟨ cong  x  x  (k  a~  b)  a  b) (kab≡a _ _) 
-               k  (k  a~  b)  a  b
-                 ≡⟨ cong  x  k  x  a  b) (kab≡a _ _) 
-               k  a~  a  b
-                 ≡⟨ cong  x  x  b) (kab≡a _ _) 
-               a~  b
-                   ))
-            (a~proves x b (subst  x'  b   ψ  x') (sym fx≡y) b⊩ψy))))
-
-  `∀isRightAdjoint :  ϕ ψ f  ( f) ψ ≤X ϕ  ψ ≤Y `∀[ f ] ϕ
-  `∀isRightAdjoint ϕ ψ f =
-    hPropExt
-      (isProp≤X (( f) ψ) ϕ)
-      (isProp≤Y ψ (`∀[ f ] ϕ))
-      (`∀isRightAdjoint← ϕ ψ f)
-      (`∀isRightAdjoint→ ϕ ψ f)
-
--- The proof is trivial but I am the reader it was left to as an exercise
-module BeckChevalley
-    {ℓ' ℓ'' : Level}
-    (I J K : Type ℓ')
-    (isSetI : isSet I)
-    (isSetJ : isSet J)
-    (isSetK : isSet K)
-    (f : J  I)
-    (g : K  I) where
-
-    module Morphism' = Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''}
-    open Morphism'
+  infix 25 _⊔_
+  _⊔_ : PredicateX  PredicateX  PredicateX
+  (ϕ  ψ) .isSetX = ϕ .isSetX
+   ϕ  ψ  x a =  ((pr₁  a  k) × ((pr₂  a)   ϕ  x))  ((pr₁  a  k') × ((pr₂  a)   ψ  x)) ∥₁
+  (ϕ  ψ) .isPropValued x a = isPropPropTrunc
+
+  infix 25 _⊓_
+  _⊓_ : PredicateX  PredicateX  PredicateX
+  (ϕ  ψ) .isSetX = ϕ .isSetX
+   ϕ  ψ  x a = ((pr₁  a)   ϕ  x) × ((pr₂  a)   ψ  x)
+  (ϕ  ψ) .isPropValued x a = isProp× (ϕ .isPropValued x (pr₁  a)) (ψ .isPropValued x (pr₂  a))
+
+  infix 25 _⇒_
+  _⇒_ : PredicateX  PredicateX  PredicateX
+  (ϕ  ψ) .isSetX = ϕ .isSetX
+   ϕ  ψ  x a =  b  (b   ϕ  x)  (a  b)   ψ  x
+  (ϕ  ψ) .isPropValued x a = isPropΠ λ a  isPropΠ λ a⊩ϕx  ψ .isPropValued _ _
+
+module _ where
+  open PredicateProperties Unit*
+  private
+    Predicate' = Predicate 
+  module NotAntiSym (antiSym :  (a b : Predicate' Unit*)  (a≤b : a  b)  (b≤a : b  a)  a  b) where
+    Lift' = Lift {i = } {j = (ℓ-max ℓ' ℓ'')}
+
+    kRealized : Predicate' Unit*
+    kRealized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a  Lift' (a  k) ; isPropValued = λ x a  isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k) }
+
+    k'Realized : Predicate' Unit*
+    k'Realized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a  Lift' (a  k') ; isPropValued = λ x a  isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k') }
+
+    kRealized≤k'Realized : kRealized  k'Realized
+    kRealized≤k'Realized =
+      do
+        let
+          prover : ApplStrTerm as 1
+          prover = ` k'
+        return (λ* prover , λ { x a (lift a≡k)  lift (λ*ComputationRule prover a) })
+
+    k'Realized≤kRealized : k'Realized  kRealized
+    k'Realized≤kRealized =
+      do
+        let
+          prover : ApplStrTerm as 1
+          prover = ` k
+        return (λ* prover , λ { x a (lift a≡k')  lift (λ*ComputationRule prover a) })
+
+    kRealized≡k'Realized : kRealized  k'Realized
+    kRealized≡k'Realized = antiSym kRealized k'Realized kRealized≤k'Realized k'Realized≤kRealized
+
+    Lift≡ : Lift' (k  k)  Lift' (k  k')
+    Lift≡ i =  kRealized≡k'Realized i  tt* k
+
+    Liftk≡k' : Lift' (k  k')
+    Liftk≡k' = transport Lift≡ (lift refl)
+
+    k≡k' : k  k'
+    k≡k' = Liftk≡k' .lower
+
+module Morphism {X Y : Type ℓ'} (isSetX : isSet X) (isSetY : isSet Y)  where
+  PredicateX = Predicate  X
+  PredicateY = Predicate  Y
+  module PredicatePropertiesX = PredicateProperties X
+  module PredicatePropertiesY = PredicateProperties Y
+  open PredicatePropertiesX renaming (_≤_ to _≤X_ ; isProp≤ to isProp≤X)
+  open PredicatePropertiesY renaming (_≤_ to _≤Y_ ; isProp≤ to isProp≤Y)
+  open Predicate hiding (isSetX)
+
+  ⋆_ : (X  Y)  (PredicateY  PredicateX)
+   f =
+    λ ϕ 
+      record
+        { isSetX = isSetX
+        ; ∣_∣ = λ x a  a   ϕ  (f x)
+        ; isPropValued = λ x a  ϕ .isPropValued (f x) a }
+
+  `∀[_] : (X  Y)  (PredicateX  PredicateY)
+  `∀[ f ] =
+    λ ϕ 
+      record
+        { isSetX = isSetY
+        ; ∣_∣ = λ y a  (∀ b x  f x  y  (a  b)   ϕ  x)
+        ; isPropValued = λ y a  isPropΠ λ a'  isPropΠ λ x  isPropΠ λ fx≡y  ϕ .isPropValued x (a  a') }
+
+  `∃[_] : (X  Y)  (PredicateX  PredicateY)
+  `∃[ f ] =
+    λ ϕ 
+      record
+        { isSetX = isSetY
+        ; ∣_∣ = λ y a  ∃[ x  X ] (f x  y) × (a   ϕ  x)
+        ; isPropValued = λ y a  isPropPropTrunc }
+
+  -- Adjunction proofs
+
+  `∃isLeftAdjoint→ :  ϕ ψ f  `∃[ f ] ϕ ≤Y ψ  ϕ ≤X ( f) ψ
+  `∃isLeftAdjoint→ ϕ ψ f p =
+    do
+      (a~ , a~proves)  p
+      return (a~ ,  x a a⊩ϕx  a~proves (f x) a  x , refl , a⊩ϕx ∣₁))
+
+
+  `∃isLeftAdjoint← :  ϕ ψ f  ϕ ≤X ( f) ψ  `∃[ f ] ϕ ≤Y ψ
+  `∃isLeftAdjoint← ϕ ψ f p =
+    do
+      (a~ , a~proves)  p
+      return
+        (a~ ,
+         y b b⊩∃fϕ 
+          equivFun
+            (propTruncIdempotent≃
+              (ψ .isPropValued y (a~  b)))
+              (do
+                (x , fx≡y , b⊩ϕx)  b⊩∃fϕ
+                return (subst  y'  (a~  b)   ψ  y') fx≡y (a~proves x b b⊩ϕx)))))
+
+  `∃isLeftAdjoint :  ϕ ψ f  `∃[ f ] ϕ ≤Y ψ  ϕ ≤X ( f) ψ
+  `∃isLeftAdjoint ϕ ψ f =
+    hPropExt
+      (isProp≤Y (`∃[ f ] ϕ) ψ)
+      (isProp≤X ϕ (( f) ψ))
+      (`∃isLeftAdjoint→ ϕ ψ f)
+      (`∃isLeftAdjoint← ϕ ψ f)
+
+  `∀isRightAdjoint→ :  ϕ ψ f  ψ ≤Y `∀[ f ] ϕ  ( f) ψ ≤X ϕ
+  `∀isRightAdjoint→ ϕ ψ f p =
+    do
+      (a~ , a~proves)  p
+      let realizer = (s  (s  (k  a~)  Id)  (k  k))
+      return
+        (realizer ,
+         x a a⊩ψfx 
+          equivFun
+            (propTruncIdempotent≃
+              (ϕ .isPropValued x (realizer  a) ))
+              (do
+                let ∀prover = a~proves (f x) a a⊩ψfx
+                return
+                  (subst
+                     ϕ~  ϕ~   ϕ  x)
+                    (sym
+                      (realizer  a
+                        ≡⟨ refl 
+                       s  (s  (k  a~)  Id)  (k  k)  a
+                        ≡⟨ sabc≡ac_bc _ _ _ 
+                       s  (k  a~)  Id  a  (k  k  a)
+                        ≡⟨ cong  x  x  (k  k  a)) (sabc≡ac_bc _ _ _) 
+                       k  a~  a  (Id  a)  (k  k  a)
+                        ≡⟨ cong  x  k  a~  a  x  (k  k  a)) (Ida≡a a) 
+                       k  a~  a  a  (k  k  a)
+                        ≡⟨ cong  x  k  a~  a  a  x) (kab≡a _ _) 
+                       (k  a~  a)  a  k
+                        ≡⟨ cong  x  x  a  k) (kab≡a _ _) 
+                       a~  a  k
+                         ))
+                    (∀prover k x refl)))))
+
+  `∀isRightAdjoint← :  ϕ ψ f  ( f) ψ ≤X ϕ  ψ ≤Y `∀[ f ] ϕ
+  `∀isRightAdjoint← ϕ ψ f p =
+    do
+      (a~ , a~proves)  p
+      let realizer = (s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id))
+      return
+        (realizer ,
+         y b b⊩ψy a x fx≡y 
+          subst
+             r  r   ϕ  x)
+            (sym
+              (realizer  b  a
+                 ≡⟨ refl 
+               s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id)  b  a
+                 ≡⟨ cong  x  x  a) (sabc≡ac_bc _ _ _) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  (s  (k  k)  Id  b)  a
+                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  x  a) (sabc≡ac_bc (k  k) Id b) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  ((k  k  b)  (Id  b))  a
+                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (x  (Id  b))  a) (kab≡a _ _) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  (Id  b))  a
+                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (k  x)  a) (Ida≡a b) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  b)  a
+                 ≡⟨ cong  x  x  (k  b)  a) (sabc≡ac_bc _ _ _) 
+               k  s  b  (s  (k  k)  (k  a~)  b)  (k  b)  a
+                 ≡⟨ cong  x  x  (s  (k  k)  (k  a~)  b)  (k  b)  a) (kab≡a _ _) 
+               s  (s  (k  k)  (k  a~)  b)  (k  b)  a
+                 ≡⟨ sabc≡ac_bc _ _ _ 
+               s  (k  k)  (k  a~)  b  a  (k  b  a)
+                 ≡⟨ cong  x  s  (k  k)  (k  a~)  b  a  x) (kab≡a b a) 
+               s  (k  k)  (k  a~)  b  a  b
+                 ≡⟨ cong  x  x  a  b) (sabc≡ac_bc (k  k) (k  a~) b) 
+               k  k  b  (k  a~  b)  a  b
+                 ≡⟨ cong  x  x  (k  a~  b)  a  b) (kab≡a _ _) 
+               k  (k  a~  b)  a  b
+                 ≡⟨ cong  x  k  x  a  b) (kab≡a _ _) 
+               k  a~  a  b
+                 ≡⟨ cong  x  x  b) (kab≡a _ _) 
+               a~  b
+                   ))
+            (a~proves x b (subst  x'  b   ψ  x') (sym fx≡y) b⊩ψy))))
+
+  `∀isRightAdjoint :  ϕ ψ f  ( f) ψ ≤X ϕ  ψ ≤Y `∀[ f ] ϕ
+  `∀isRightAdjoint ϕ ψ f =
+    hPropExt
+      (isProp≤X (( f) ψ) ϕ)
+      (isProp≤Y ψ (`∀[ f ] ϕ))
+      (`∀isRightAdjoint← ϕ ψ f)
+      (`∀isRightAdjoint→ ϕ ψ f)
+
+-- The proof is trivial but I am the reader it was left to as an exercise
+module BeckChevalley
+    (I J K : Type ℓ')
+    (isSetI : isSet I)
+    (isSetJ : isSet J)
+    (isSetK : isSet K)
+    (f : J  I)
+    (g : K  I) where
+
+    module Morphism' = Morphism
+    open Morphism'
     
-    L = Σ[ k  K ] Σ[ j  J ] (g k  f j)
-
-    isSetL : isSet L
-    isSetL = isSetΣ isSetK λ k  isSetΣ isSetJ λ j  isProp→isSet (isSetI _ _)
-
-    p : L  K
-    p (k , _ , _) = k
-
-    q : L  J
-    q (_ , l , _) = l
-
-    q* = ⋆_ isSetL isSetJ q
-    g* = ⋆_ isSetK isSetI g
-
-    module `f = Morphism' isSetJ isSetI
-    open `f renaming (`∃[_] to `∃[J→I][_]; `∀[_] to `∀[J→I][_])
-
-    module `q = Morphism' isSetL isSetK
-    open `q renaming (`∃[_] to `∃[L→K][_]; `∀[_] to `∀[L→K][_])
-
-    `∃BeckChevalley→ :  ϕ k a  a   g* (`∃[J→I][ f ] ϕ)  k  a   `∃[L→K][ p ] (q* ϕ)  k
-    `∃BeckChevalley→ ϕ k a p =
-      do
-        (j , fj≡gk , a⊩ϕj)  p
-        return ((k , (j , (sym fj≡gk))) , (refl , a⊩ϕj))
-
-    `∃BeckChevalley← :  ϕ k a  a   `∃[L→K][ p ] (q* ϕ)  k  a   g* (`∃[J→I][ f ] ϕ)  k
-    `∃BeckChevalley← ϕ k a p =
-      do
-        (x@(k' , j , gk'≡fj) , k'≡k , a⊩ϕqj)  p
-        return (j , (subst  k  f j  g k) k'≡k (sym gk'≡fj)) , a⊩ϕqj)
-
-    open Iso
-    `∃BeckChevalley : g*  `∃[J→I][ f ]  `∃[L→K][ p ]  q*
-    `∃BeckChevalley =
-      funExt λ ϕ i 
-        PredicateIsoΣ K .inv
-          (PredicateΣ≡ {ℓ'' = ℓ''} K
-            ((λ k a  ( (g*  `∃[J→I][ f ]) ϕ  k a) , ((g*  `∃[J→I][ f ]) ϕ .isPropValued k a)) , isSetK)
-            ((λ k a  ( (`∃[L→K][ p ]  q*) ϕ  k a) , ((`∃[L→K][ p ]  q*) ϕ .isPropValued k a)) , isSetK)
-            (funExt₂
-               k a 
-                Σ≡Prop
-                   x  isPropIsProp {A = x})
-                  (hPropExt
-                    (g* (`∃[J→I][ f ] ϕ) .isPropValued k a)
-                    (`∃[L→K][ p ] (q* ϕ) .isPropValued k a)
-                    (`∃BeckChevalley→ ϕ k a)
-                    (`∃BeckChevalley← ϕ k a))))
-           i)
-
-    `∀BeckChevalley→ :  ϕ k a  a   g* (`∀[J→I][ f ] ϕ)  k  a   `∀[L→K][ p ] (q* ϕ)  k
-    `∀BeckChevalley→ ϕ k a p b (k' , j , gk'≡fj) k'≡k = p b j (sym (subst  k''  g k''  f j) k'≡k gk'≡fj))
-
-    `∀BeckChevalley← :  ϕ k a  a   `∀[L→K][ p ] (q* ϕ)  k  a   g* (`∀[J→I][ f ] ϕ)  k
-    `∀BeckChevalley← ϕ k a p b j fj≡gk = p b (k , j , sym fj≡gk) refl
-
-    `∀BeckChevalley : g*  `∀[J→I][ f ]  `∀[L→K][ p ]  q*
-    `∀BeckChevalley =
-      funExt λ ϕ i 
-        PredicateIsoΣ K .inv
-          (PredicateΣ≡ {ℓ'' = ℓ''} K
-            ((λ k a  (a   g* (`∀[J→I][ f ] ϕ)  k) , (g* (`∀[J→I][ f ] ϕ) .isPropValued k a)) , isSetK)
-            ((λ k a  (a   `∀[L→K][ p ] (q* ϕ)  k) , (`∀[L→K][ p ] (q* ϕ) .isPropValued k a)) , isSetK)
-            (funExt₂
-               k a 
-                Σ≡Prop
-                   x  isPropIsProp {A = x})
-                  (hPropExt
-                    (g* (`∀[J→I][ f ] ϕ) .isPropValued k a)
-                    (`∀[L→K][ p ] (q* ϕ) .isPropValued k a)
-                    (`∀BeckChevalley→ ϕ k a)
-                    (`∀BeckChevalley← ϕ k a))))
-          i)
+    L = Σ[ k  K ] Σ[ j  J ] (g k  f j)
+
+    isSetL : isSet L
+    isSetL = isSetΣ isSetK λ k  isSetΣ isSetJ λ j  isProp→isSet (isSetI _ _)
+
+    p : L  K
+    p (k , _ , _) = k
+
+    q : L  J
+    q (_ , l , _) = l
+
+    q* = ⋆_ isSetL isSetJ q
+    g* = ⋆_ isSetK isSetI g
+
+    module `f = Morphism' isSetJ isSetI
+    open `f renaming (`∃[_] to `∃[J→I][_]; `∀[_] to `∀[J→I][_])
+
+    module `q = Morphism' isSetL isSetK
+    open `q renaming (`∃[_] to `∃[L→K][_]; `∀[_] to `∀[L→K][_])
+
+    `∃BeckChevalley→ :  ϕ k a  a   g* (`∃[J→I][ f ] ϕ)  k  a   `∃[L→K][ p ] (q* ϕ)  k
+    `∃BeckChevalley→ ϕ k a p =
+      do
+        (j , fj≡gk , a⊩ϕj)  p
+        return ((k , (j , (sym fj≡gk))) , (refl , a⊩ϕj))
+
+    `∃BeckChevalley← :  ϕ k a  a   `∃[L→K][ p ] (q* ϕ)  k  a   g* (`∃[J→I][ f ] ϕ)  k
+    `∃BeckChevalley← ϕ k a p =
+      do
+        (x@(k' , j , gk'≡fj) , k'≡k , a⊩ϕqj)  p
+        return (j , (subst  k  f j  g k) k'≡k (sym gk'≡fj)) , a⊩ϕqj)
+
+    open Iso
+    `∃BeckChevalley : g*  `∃[J→I][ f ]  `∃[L→K][ p ]  q*
+    `∃BeckChevalley =
+      funExt λ ϕ i 
+        PredicateIsoΣ K .inv
+          (PredicateΣ≡  K
+            ((λ k a  ( (g*  `∃[J→I][ f ]) ϕ  k a) , ((g*  `∃[J→I][ f ]) ϕ .isPropValued k a)) , isSetK)
+            ((λ k a  ( (`∃[L→K][ p ]  q*) ϕ  k a) , ((`∃[L→K][ p ]  q*) ϕ .isPropValued k a)) , isSetK)
+            (funExt₂
+               k a 
+                Σ≡Prop
+                   x  isPropIsProp {A = x})
+                  (hPropExt
+                    (g* (`∃[J→I][ f ] ϕ) .isPropValued k a)
+                    (`∃[L→K][ p ] (q* ϕ) .isPropValued k a)
+                    (`∃BeckChevalley→ ϕ k a)
+                    (`∃BeckChevalley← ϕ k a))))
+           i)
+
+    `∀BeckChevalley→ :  ϕ k a  a   g* (`∀[J→I][ f ] ϕ)  k  a   `∀[L→K][ p ] (q* ϕ)  k
+    `∀BeckChevalley→ ϕ k a p b (k' , j , gk'≡fj) k'≡k = p b j (sym (subst  k''  g k''  f j) k'≡k gk'≡fj))
+
+    `∀BeckChevalley← :  ϕ k a  a   `∀[L→K][ p ] (q* ϕ)  k  a   g* (`∀[J→I][ f ] ϕ)  k
+    `∀BeckChevalley← ϕ k a p b j fj≡gk = p b (k , j , sym fj≡gk) refl
+
+    `∀BeckChevalley : g*  `∀[J→I][ f ]  `∀[L→K][ p ]  q*
+    `∀BeckChevalley =
+      funExt λ ϕ i 
+        PredicateIsoΣ K .inv
+          (PredicateΣ≡ K
+            ((λ k a  (a   g* (`∀[J→I][ f ] ϕ)  k) , (g* (`∀[J→I][ f ] ϕ) .isPropValued k a)) , isSetK)
+            ((λ k a  (a   `∀[L→K][ p ] (q* ϕ)  k) , (`∀[L→K][ p ] (q* ϕ) .isPropValued k a)) , isSetK)
+            (funExt₂
+               k a 
+                Σ≡Prop
+                   x  isPropIsProp {A = x})
+                  (hPropExt
+                    (g* (`∀[J→I][ f ] ϕ) .isPropValued k a)
+                    (`∀[L→K][ p ] (q* ϕ) .isPropValued k a)
+                    (`∀BeckChevalley→ ϕ k a)
+                    (`∀BeckChevalley← ϕ k a))))
+          i)
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Predicate.html b/docs/Realizability.Tripos.Prealgebra.Predicate.html index 1f6df21..88eab61 100644 --- a/docs/Realizability.Tripos.Prealgebra.Predicate.html +++ b/docs/Realizability.Tripos.Prealgebra.Predicate.html @@ -2,8 +2,8 @@ Realizability.Tripos.Prealgebra.Predicate
open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 
-module Realizability.Tripos.Prealgebra.Predicate {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Tripos.Prealgebra.Predicate { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca public
-open import Realizability.Tripos.Prealgebra.Predicate.Properties ca public
+open import Realizability.Tripos.Prealgebra.Predicate.Base { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca public
+open import Realizability.Tripos.Prealgebra.Predicate.Properties {ℓ' = ℓ'} {ℓ'' = ℓ''} ca public
 
\ No newline at end of file diff --git a/docs/index.html b/docs/index.html index 414742e..b05e2a9 100644 --- a/docs/index.html +++ b/docs/index.html @@ -4,11 +4,6 @@ open import Realizability.CombinatoryAlgebra open import Realizability.ApplicativeStructure -open import Realizability.Assembly.Everything -open import Realizability.Tripos.Everything -open import Realizability.Topos.Everything -open import Realizability.Choice -open import Tripoi.Tripos -open import Tripoi.HeytingAlgebra -open import Tripoi.PosetReflection +open import Realizability.Topos.Everything +open import Realizability.Choice
\ No newline at end of file diff --git a/src/Realizability/ApplicativeStructure.agda b/src/Realizability/ApplicativeStructure.agda index afaa19c..d44bcbc 100644 --- a/src/Realizability/ApplicativeStructure.agda +++ b/src/Realizability/ApplicativeStructure.agda @@ -91,6 +91,9 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where λ*3 : Term three → A λ*3 t = ⟦ λ*abst (λ*abst (λ*abst t)) ⟧ [] + λ*4 : Term four → A + λ*4 t = ⟦ λ*abst (λ*abst (λ*abst (λ*abst t))) ⟧ [] + opaque unfolding λ*abst βreduction : ∀ {n} → (body : Term (suc n)) → (prim : A) → (subs : Vec A n) → ⟦ λ*abst body ̇ ` prim ⟧ subs ≡ ⟦ body ⟧ (prim ∷ subs) @@ -153,3 +156,16 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where ≡⟨ βreduction t c (b ∷ a ∷ []) ⟩ ⟦ t ⟧ (c ∷ b ∷ a ∷ []) ∎ + + λ*4ComputationRule : ∀ (t : Term 4) (a b c d : A) → λ*4 t ⨾ a ⨾ b ⨾ c ⨾ d ≡ ⟦ t ⟧ (d ∷ c ∷ b ∷ a ∷ []) + λ*4ComputationRule t a b c d = + ⟦ λ*abst (λ*abst (λ*abst (λ*abst t))) ⟧ [] ⨾ ⟦ ` a ⟧ [] ⨾ ⟦ ` b ⟧ [] ⨾ ⟦ ` c ⟧ [] ⨾ ⟦ ` d ⟧ [] + ≡⟨ cong (λ x → x ⨾ b ⨾ c ⨾ d) (βreduction (λ*abst (λ*abst (λ*abst t))) a []) ⟩ + ⟦ λ*abst (λ*abst (λ*abst t)) ⟧ (a ∷ []) ⨾ ⟦ ` b ⟧ (a ∷ []) ⨾ ⟦ ` c ⟧ (a ∷ []) ⨾ ⟦ ` d ⟧ (a ∷ []) + ≡⟨ cong (λ x → x ⨾ c ⨾ d) (βreduction (λ*abst (λ*abst t)) b (a ∷ [])) ⟩ + ⟦ λ*abst (λ*abst t) ⟧ (b ∷ a ∷ []) ⨾ ⟦ ` c ⟧ (b ∷ a ∷ []) ⨾ ⟦ ` d ⟧ (b ∷ a ∷ []) + ≡⟨ cong (λ x → x ⨾ d) (βreduction (λ*abst t) c (b ∷ a ∷ [])) ⟩ + ⟦ λ*abst t ⟧ (c ∷ b ∷ a ∷ []) ⨾ ⟦ ` d ⟧ (c ∷ b ∷ a ∷ []) + ≡⟨ βreduction t d (c ∷ b ∷ a ∷ []) ⟩ + ⟦ t ⟧ (d ∷ c ∷ b ∷ a ∷ []) + ∎ diff --git a/src/Realizability/Topos/FunctionalRelation.agda b/src/Realizability/Topos/FunctionalRelation.agda index aed4695..b4c151e 100644 --- a/src/Realizability/Topos/FunctionalRelation.agda +++ b/src/Realizability/Topos/FunctionalRelation.agda @@ -567,6 +567,37 @@ opaque answer)) f g h +-- Very useful helper functions to prevent type-checking time from exploding +opaque + [F]≡[G]→F≤G : ∀ {X Y : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} → (F G : FunctionalRelation perX perY) → [_] {R = bientailment perX perY} F ≡ [_] {R = bientailment perX perY} G → pointwiseEntailment perX perY F G + [F]≡[G]→F≤G {X} {Y} {perX} {perY} F G iso = SQ.effective (isPropValuedBientailment perX perY) (isEquivRelBientailment perX perY) F G iso .fst + + [F]≡[G]→G≤F : ∀ {X Y : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} → (F G : FunctionalRelation perX perY) → [_] {R = bientailment perX perY} F ≡ [_] {R = bientailment perX perY} G → pointwiseEntailment perX perY G F + [F]≡[G]→G≤F {X} {Y} {perX} {perY} F G iso = SQ.effective (isPropValuedBientailment perX perY) (isEquivRelBientailment perX perY) F G iso .snd + +opaque + unfolding composeRTMorphism + [F]⋆[G]≡[H]⋆[I]→F⋆G≤H⋆I : ∀ {X Y Z W : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z} {perW : PartialEquivalenceRelation W} → (F : FunctionalRelation perX perY) (G : FunctionalRelation perY perZ) (H : FunctionalRelation perX perW) (I : FunctionalRelation perW perZ) → composeRTMorphism _ _ _ [ F ] [ G ] ≡ composeRTMorphism _ _ _ [ H ] [ I ] → pointwiseEntailment _ _ (composeFuncRel _ _ _ F G) (composeFuncRel _ _ _ H I) + [F]⋆[G]≡[H]⋆[I]→F⋆G≤H⋆I {X} {Y} {Z} {W} {perX} {perY} {perZ} {perW} F G H I iso = + SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) (composeFuncRel _ _ _ F G) (composeFuncRel _ _ _ H I) iso .fst + +opaque + unfolding composeRTMorphism + [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G : ∀ {X Y Z W : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z} {perW : PartialEquivalenceRelation W} → (F : FunctionalRelation perX perY) (G : FunctionalRelation perY perZ) (H : FunctionalRelation perX perW) (I : FunctionalRelation perW perZ) → composeRTMorphism _ _ _ [ F ] [ G ] ≡ composeRTMorphism _ _ _ [ H ] [ I ] → pointwiseEntailment _ _ (composeFuncRel _ _ _ H I) (composeFuncRel _ _ _ F G) + [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G {X} {Y} {Z} {W} {perX} {perY} {perZ} {perW} F G H I iso = + SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) (composeFuncRel _ _ _ F G) (composeFuncRel _ _ _ H I) iso .snd + +opaque + unfolding composeRTMorphism + [F]≡[G]⋆[H]→F≤G⋆H : ∀ {X Y Z : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z} → (F : FunctionalRelation perX perZ) (G : FunctionalRelation perX perY) (H : FunctionalRelation perY perZ) → [ F ] ≡ composeRTMorphism _ _ _ [ G ] [ H ] → pointwiseEntailment _ _ F (composeFuncRel _ _ _ G H) + [F]≡[G]⋆[H]→F≤G⋆H {X} {Y} {Z} {perX} {perY} {perZ} F G H iso = + SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) F (composeFuncRel _ _ _ G H) iso .fst + +opaque + unfolding composeRTMorphism + [F]≡[G]⋆[H]→G⋆H≤F : ∀ {X Y Z : Type ℓ'} {perX : PartialEquivalenceRelation X} {perY : PartialEquivalenceRelation Y} {perZ : PartialEquivalenceRelation Z} → (F : FunctionalRelation perX perZ) (G : FunctionalRelation perX perY) (H : FunctionalRelation perY perZ) → [ F ] ≡ composeRTMorphism _ _ _ [ G ] [ H ] → pointwiseEntailment _ _ (composeFuncRel _ _ _ G H) F + [F]≡[G]⋆[H]→G⋆H≤F {X} {Y} {Z} {perX} {perY} {perZ} F G H iso = SQ.effective (isPropValuedBientailment perX perZ) (isEquivRelBientailment perX perZ) F (composeFuncRel _ _ _ G H) iso .snd + RT : Category (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) (ℓ-max (ℓ-suc ℓ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) Category.ob RT = Σ[ X ∈ Type ℓ' ] PartialEquivalenceRelation X Category.Hom[_,_] RT (X , perX) (Y , perY) = RTMorphism perX perY diff --git a/src/Realizability/Topos/PowerObject.agda b/src/Realizability/Topos/PowerObject.agda index 228682c..dfbacd7 100644 --- a/src/Realizability/Topos/PowerObject.agda +++ b/src/Realizability/Topos/PowerObject.agda @@ -8,7 +8,9 @@ open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category +open import Cubical.Categories.Limits.Pullback open import Realizability.PropResizing open import Realizability.CombinatoryAlgebra @@ -193,3 +195,45 @@ module _ (X : Type ℓ) (perX : PartialEquivalenceRelation X) where isFunctionalRelation.isRelational (isFuncRel topArrowFuncRel) = {!!} isFunctionalRelation.isSingleValued (isFuncRel topArrowFuncRel) = {!!} isFunctionalRelation.isTotal (isFuncRel topArrowFuncRel) = {!!} + + powerObjectCospan : RTMorphism (binProdObRT perX perY) (binProdObRT perX 𝓟) → Cospan RT + Cospan.l (powerObjectCospan f) = X × Y , binProdObRT perX perY + Cospan.m (powerObjectCospan f) = X × ResizedPredicate X , binProdObRT perX 𝓟 + Cospan.r (powerObjectCospan f) = X × ResizedPredicate X , ∈subPer + Cospan.s₁ (powerObjectCospan f) = f + Cospan.s₂ (powerObjectCospan f) = [ ∈incFuncRel ] + + F : FunctionalRelation (binProdObRT perX perY) (binProdObRT perX 𝓟) + Predicate.isSetX (relation F) = isSet× (isSet× (perX .isSetX) (perY .isSetX)) (isSet× (perX .isSetX) isSetResizedPredicate) + Predicate.∣ relation F ∣ ((x , y) , (x' , p)) r = (pr₁ ⨾ (pr₁ ⨾ r)) ⊩ ∣ perY .equality ∣ (y , y) × (pr₂ ⨾ (pr₁ ⨾ r)) ⊩ ∣ 𝓟 .equality ∣ (p , p) × {!∀ !} × {!!} + Predicate.isPropValued (relation F) = {!!} + isFunctionalRelation.isStrictDomain (isFuncRel F) = {!!} + isFunctionalRelation.isStrictCodomain (isFuncRel F) = {!!} + isFunctionalRelation.isRelational (isFuncRel F) = {!!} + isFunctionalRelation.isSingleValued (isFuncRel F) = {!!} + isFunctionalRelation.isTotal (isFuncRel F) = {!!} + + opaque + unfolding composeRTMorphism + unfolding composeFuncRel + pullbackSquareCommutes : [ ϕincFuncRel ] ⋆ [ F ] ≡ [ topArrowFuncRel ] ⋆ [ ∈incFuncRel ] + pullbackSquareCommutes = + eq/ _ _ {!!} + + isPowerObjectUnivProp : Type _ + isPowerObjectUnivProp = + ∃![ f ∈ RTMorphism (binProdObRT perX perY) (binProdObRT perX 𝓟) ] + Σ[ commutes ∈ [ ϕincFuncRel ] ⋆ f ≡ [ topArrowFuncRel ] ⋆ [ ∈incFuncRel ] ] + (isPullback RT (powerObjectCospan f) {c = X × Y , ϕsubPer} [ ϕincFuncRel ] [ topArrowFuncRel ] commutes) + + isPropIsPowerObjectUnivProp : isProp isPowerObjectUnivProp + isPropIsPowerObjectUnivProp = isPropIsContr + + isPowerObject : isPowerObjectUnivProp + isPowerObject = + uniqueExists + [ F ] + (pullbackSquareCommutes , {!!}) + (λ F' → isPropΣ (squash/ _ _) λ commutes → isPropIsPullback RT (powerObjectCospan F') [ ϕincFuncRel ] [ topArrowFuncRel ] commutes) + (λ { f' (commutes , isPullback) → + {!!} }) diff --git a/src/Realizability/Topos/SubobjectClassifier.agda b/src/Realizability/Topos/SubobjectClassifier.agda index 14844b4..d1d9a56 100644 --- a/src/Realizability/Topos/SubobjectClassifier.agda +++ b/src/Realizability/Topos/SubobjectClassifier.agda @@ -12,6 +12,7 @@ open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Morphism open import Realizability.PropResizing open import Realizability.CombinatoryAlgebra @@ -98,11 +99,33 @@ isPartialEquivalenceRelation.isSymmetric (isPerEquality Ωper) = subst (λ r' → r' ⊩ ∣ toPredicate β ∣ tt*) (sym eq) (pr₁r⊩α≤β a a⊩α)) }) isPartialEquivalenceRelation.isTransitive (isPerEquality Ωper) = do + let + closure1 : ApplStrTerm as 3 + closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ # two ̇ # zero) + + closure2 : ApplStrTerm as 3 + closure2 = ` pr₂ ̇ # two ̇ (` pr₂ ̇ # one ̇ # zero) + + realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2) return - ({!!} , + (λ*2 realizer , (λ { x y z a b (⊩x≤y , ⊩y≤x) (⊩y≤z , ⊩z≤y) → - (λ r r⊩x → subst (λ r' → r' ⊩ ∣ toPredicate z ∣ tt*) {!!} (⊩y≤z _ (⊩x≤y r r⊩x))) , - (λ r r⊩z → subst (λ r' → r' ⊩ ∣ toPredicate x ∣ tt*) {!!} (⊩y≤x _ (⊩z≤y r r⊩z))) })) + (λ r r⊩x → + subst + (λ r' → r' ⊩ ∣ toPredicate z ∣ tt*) + (sym + (cong (λ x → pr₁ ⨾ x ⨾ r) (λ*2ComputationRule realizer a b) ∙ + cong (λ x → x ⨾ r) (pr₁pxy≡x _ _) ∙ + βreduction closure1 r (b ∷ a ∷ []))) + (⊩y≤z _ (⊩x≤y r r⊩x))) , + (λ r r⊩z → + subst + (λ r' → r' ⊩ ∣ toPredicate x ∣ tt*) + (sym + (cong (λ x → pr₂ ⨾ x ⨾ r) (λ*2ComputationRule realizer a b) ∙ + cong (λ x → x ⨾ r) (pr₂pxy≡y _ _) ∙ + βreduction closure2 r (b ∷ a ∷ []))) + (⊩y≤x _ (⊩z≤y r r⊩z))) })) opaque unfolding terminalPer @@ -117,24 +140,72 @@ opaque (λ { tt* y r r⊩⊤≤y → tt*})) isFunctionalRelation.isStrictCodomain (isFuncRel trueFuncRel) = do + let + idClosure : ApplStrTerm as 2 + idClosure = # zero + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure) return - ({!!} , + (λ* realizer , (λ { tt* y r r⊩⊤≤y → - (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y) , - (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y)})) + (λ a a⊩y → + subst + (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) + (sym + (cong (λ x → pr₁ ⨾ x ⨾ a) (λ*ComputationRule realizer r) ∙ + cong (λ x → x ⨾ a) (pr₁pxy≡x _ _) ∙ + βreduction idClosure a (r ∷ []))) + a⊩y) , + (λ a a⊩y → + subst + (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) + (sym + (cong (λ x → pr₂ ⨾ x ⨾ a) (λ*ComputationRule realizer r) ∙ + cong (λ x → x ⨾ a) (pr₂pxy≡y _ _) ∙ + βreduction idClosure a (r ∷ []))) + a⊩y)})) isFunctionalRelation.isRelational (isFuncRel trueFuncRel) = do + let + realizer : ApplStrTerm as 4 + realizer = ` pr₁ ̇ # one ̇ (# two ̇ ` k) return - ({!!} , + (λ*4 realizer , (λ { tt* tt* x y a b c tt* b⊩⊤≤x (pr₁c⊩x≤y , pr₂c⊩y≤x) r → - subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} (pr₁c⊩x≤y (b ⨾ k) (b⊩⊤≤x k))})) + subst + (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) + (sym (λ*4ComputationRule realizer a b c r)) + (pr₁c⊩x≤y (b ⨾ k) (b⊩⊤≤x k))})) isFunctionalRelation.isSingleValued (isFuncRel trueFuncRel) = do + let + closure1 : ApplStrTerm as 3 + closure1 = # one ̇ ` k + + closure2 : ApplStrTerm as 3 + closure2 = # two ̇ ` k + + realizer : ApplStrTerm as 2 + realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2) return - ({!!} , + (λ*2 realizer , (λ { tt* x y r₁ r₂ r₁⊩⊤≤x r₂⊩⊤≤y → - (λ a a⊩x → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} (r₂⊩⊤≤y k)) , - (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate x ∣ tt*) {!!} (r₁⊩⊤≤x k))})) + (λ a a⊩x → + subst + (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) + (sym + (cong (λ x → pr₁ ⨾ x ⨾ a) (λ*2ComputationRule realizer r₁ r₂) ∙ + cong (λ x → x ⨾ a) (pr₁pxy≡x _ _) ∙ + βreduction closure1 a (r₂ ∷ r₁ ∷ []))) + (r₂⊩⊤≤y k)) , + (λ a a⊩y → + subst + (λ r' → r' ⊩ ∣ toPredicate x ∣ tt*) + (sym + (cong (λ x → pr₂ ⨾ x ⨾ a) (λ*2ComputationRule realizer r₁ r₂) ∙ + cong (λ x → x ⨾ a) (pr₂pxy≡y _ _) ∙ + βreduction closure2 a (r₂ ∷ r₁ ∷ []))) + (r₁⊩⊤≤x k))})) isFunctionalRelation.isTotal (isFuncRel trueFuncRel) = do return @@ -159,8 +230,15 @@ opaque (k , (λ { tt* tt* p r₁ r₂ r₁⊩⊤≤p r₂⊩⊤≤p → tt* })) --- The subobject classifier indeed classifies subobjects -module _ +truePredicate : Predicate Unit* +Predicate.isSetX truePredicate = isSetUnit* +Predicate.∣ truePredicate ∣ tt* r = Unit* +Predicate.isPropValued truePredicate tt* r = isPropUnit* + +⊤ = fromPredicate truePredicate + +-- The subobject classifier classifies subobjects represented by strict relations +module ClassifiesStrictRelations (X : Type ℓ) (perX : PartialEquivalenceRelation X) (ϕ : StrictRelation perX) where @@ -170,6 +248,7 @@ module _ resizedϕ = fromPredicate (ϕ .predicate) -- the functional relation that represents the unique indicator map + {-# TERMINATING #-} charFuncRel : FunctionalRelation perX Ωper Predicate.isSetX (relation charFuncRel) = isSet× (perX .isSetX) isSetResizedPredicate Predicate.∣ relation charFuncRel ∣ (x , p) r = @@ -189,29 +268,107 @@ module _ (λ { x p r (pr₁r⊩x~x , ⊩ϕx≤p , ⊩p≤ϕx) → pr₁r⊩x~x})) isFunctionalRelation.isStrictCodomain (isFuncRel charFuncRel) = do + let + idClosure : ApplStrTerm as 2 + idClosure = # zero + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure) return - ({!!} , + (λ* realizer , (λ x y r x₁ → - (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y) , - (λ a a⊩y → subst (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) {!!} a⊩y))) + (λ a a⊩y → + subst + (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) + (sym + (cong (λ x → pr₁ ⨾ x ⨾ a) (λ*ComputationRule realizer r) ∙ + cong (λ x → x ⨾ a) (pr₁pxy≡x _ _) ∙ + βreduction idClosure a (r ∷ []))) + a⊩y) , + (λ a a⊩y → + subst + (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) + (sym + (cong (λ x → pr₂ ⨾ x ⨾ a) (λ*ComputationRule realizer r) ∙ + cong (λ x → x ⨾ a) (pr₂pxy≡y _ _) ∙ + βreduction idClosure a (r ∷ []))) + a⊩y))) isFunctionalRelation.isRelational (isFuncRel charFuncRel) = do + (sX , sX⊩isSymmetricX) ← perX .isSymmetric + (tX , tX⊩isTransitiveX) ← perX .isTransitive + (relϕ , relϕ⊩isRelationalϕ) ← isStrictRelation.isRelational (ϕ .isStrictRelationPredicate) + let + closure1 : ApplStrTerm as 4 + closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` relϕ ̇ # zero ̇ (` sX ̇ # three))) + + closure2 : ApplStrTerm as 4 + closure2 = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ # one ̇ # zero)) ̇ # three + + realizer : ApplStrTerm as 3 + realizer = ` pair ̇ (` tX ̇ (` sX ̇ # two) ̇ # two) ̇ (` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2)) return - ({!!} , - (λ { x x' p p' a b c a⊩x~x' (pr₁b⊩x~x , pr₁pr₂b⊩ϕx≤y , pr₂pr₂b⊩y≤ϕx) (pr₁c⊩y≤y' , pr₂c⊩y'≤y) → - {!!} , - (λ r r⊩ϕx' → {!!}) , - λ r r⊩p' → {!!} })) + (λ*3 realizer , + (λ { x x' p p' a b c a⊩x~x' (⊩x~x , ⊩ϕx≤p , ⊩p≤ϕx) (⊩p≤p' , ⊩p'≤p) → + let + ⊩x'~x = sX⊩isSymmetricX x x' a a⊩x~x' + ⊩x'~x' = tX⊩isTransitiveX x' x x' _ _ ⊩x'~x a⊩x~x' + in + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x' , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*3ComputationRule realizer a b c) ∙ pr₁pxy≡x _ _)) + ⊩x'~x' , + (λ r r⊩ϕx' → + subst + (λ r' → r' ⊩ ∣ toPredicate p' ∣ tt*) + (sym + (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x) ⨾ r) (λ*3ComputationRule realizer a b c) ∙ + cong (λ x → pr₁ ⨾ x ⨾ r) (pr₂pxy≡y _ _) ∙ + cong (λ x → x ⨾ r) (pr₁pxy≡x _ _) ∙ + βreduction closure1 r (c ∷ b ∷ a ∷ []))) + (⊩p≤p' _ (⊩ϕx≤p _ (relϕ⊩isRelationalϕ x' x _ _ r⊩ϕx' ⊩x'~x)))) , + λ r r⊩p' → + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x') + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x) ⨾ r) (λ*3ComputationRule realizer a b c) ∙ + cong (λ x → pr₂ ⨾ x ⨾ r) (pr₂pxy≡y _ _) ∙ + cong (λ x → x ⨾ r) (pr₂pxy≡y _ _) ∙ + βreduction closure2 r (c ∷ b ∷ a ∷ []))) + (relϕ⊩isRelationalϕ x x' _ _ (⊩p≤ϕx _ (⊩p'≤p r r⊩p')) a⊩x~x') })) isFunctionalRelation.isSingleValued (isFuncRel charFuncRel) = do + let + closure1 : ApplStrTerm as 3 + closure1 = ` pr₁ ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ # zero) + + closure2 : ApplStrTerm as 3 + closure2 = ` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ (` pr₂ ̇ # one) ̇ # zero) + + realizer : ApplStrTerm as 2 + realizer = ` pair ̇ λ*abst closure1 ̇ λ*abst closure2 return - ({!!} , - (λ { x y y' r₁ r₂ (pr₁r₁⊩x~x , pr₁pr₂r₁⊩ϕx≤y , pr₂pr₂r₁⊩y≤ϕx) (pr₂r₂⊩x~x , pr₁pr₂r₂⊩ϕx≤y' , pr₂pr₂r₂⊩y'≤ϕx) → - {!!} })) + (λ*2 realizer , + (λ { x y y' r₁ r₂ (⊩x~x , ⊩ϕx≤y , ⊩y≤ϕx) (⊩'x~x , ⊩ϕx≤y' , ⊩y'≤ϕx) → + (λ a a⊩y → + subst + (λ r' → r' ⊩ ∣ toPredicate y' ∣ tt*) + (sym (cong (λ x → pr₁ ⨾ x ⨾ a) (λ*2ComputationRule realizer r₁ r₂) ∙ cong (λ x → x ⨾ a) (pr₁pxy≡x _ _) ∙ βreduction closure1 a (r₂ ∷ r₁ ∷ []))) + (⊩ϕx≤y' _ (⊩y≤ϕx a a⊩y))) , + (λ a a⊩y' → + subst + (λ r' → r' ⊩ ∣ toPredicate y ∣ tt*) + (sym (cong (λ x → pr₂ ⨾ x ⨾ a) (λ*2ComputationRule realizer r₁ r₂) ∙ cong (λ x → x ⨾ a) (pr₂pxy≡y _ _) ∙ βreduction closure2 a (r₂ ∷ r₁ ∷ []))) + (⊩ϕx≤y _ (⊩y'≤ϕx a a⊩y'))) })) isFunctionalRelation.isTotal (isFuncRel charFuncRel) = do + let + idClosure : ApplStrTerm as 2 + idClosure = # zero + + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ # zero ̇ (` pair ̇ λ*abst idClosure ̇ λ*abst idClosure) return - ({!!} , + (λ* realizer , (λ x r r⊩x~x → let resultPredicate : Predicate Unit* @@ -223,17 +380,35 @@ module _ in return (fromPredicate resultPredicate , - subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym {!!}) r⊩x~x , - (λ b b⊩ϕx → {!compIsIdFunc!}) , - λ b b⊩'ϕx → {!!}))) - + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₁pxy≡x _ _)) + r⊩x~x , + (λ b b⊩ϕx → + subst + (λ r → r ⊩ ∣ toPredicate (fromPredicate resultPredicate) ∣ tt*) + (sym + (cong (λ x → pr₁ ⨾ (pr₂ ⨾ x) ⨾ b) (λ*ComputationRule realizer r) ∙ + cong (λ x → pr₁ ⨾ x ⨾ b) (pr₂pxy≡y _ _) ∙ + cong (λ x → x ⨾ b) (pr₁pxy≡x _ _) ∙ + βreduction idClosure b (r ∷ []))) + (subst (λ p → b ⊩ ∣ p ∣ tt*) (sym (compIsIdFunc resultPredicate)) b⊩ϕx)) , + (λ b b⊩'ϕx → + subst + (λ r → r ⊩ ∣ ϕ .predicate ∣ x) + (sym + (cong (λ x → pr₂ ⨾ (pr₂ ⨾ x) ⨾ b) (λ*ComputationRule realizer r) ∙ + cong (λ x → pr₂ ⨾ x ⨾ b) (pr₂pxy≡y _ _) ∙ + cong (λ x → x ⨾ b) (pr₂pxy≡y _ _) ∙ + βreduction idClosure b (r ∷ []))) + let foo = subst (λ p → b ⊩ ∣ p ∣ tt*) (compIsIdFunc resultPredicate) b⊩'ϕx in foo)))) - subobjectCospan : Cospan RT - Cospan.l subobjectCospan = X , perX - Cospan.m subobjectCospan = ResizedPredicate Unit* , Ωper - Cospan.r subobjectCospan = Unit* , terminalPer - Cospan.s₁ subobjectCospan = [ charFuncRel ] - Cospan.s₂ subobjectCospan = [ trueFuncRel ] + subobjectCospan : ∀ char → Cospan RT + Cospan.l (subobjectCospan char) = X , perX + Cospan.m (subobjectCospan char) = ResizedPredicate Unit* , Ωper + Cospan.r (subobjectCospan char) = Unit* , terminalPer + Cospan.s₁ (subobjectCospan char) = char + Cospan.s₂ (subobjectCospan char) = [ trueFuncRel ] opaque unfolding composeRTMorphism @@ -286,49 +461,478 @@ module _ in eq/ _ _ (answer , F≤G→G≤F subPer Ωper (composeFuncRel _ _ _ incFuncRel charFuncRel) (composeFuncRel _ _ _ (terminalFuncRel subPer) trueFuncRel) answer) - classifies : isPullback RT subobjectCospan [ incFuncRel ] [ terminalFuncRel subPer ] subobjectSquareCommutes - classifies {Y , perY} f g f⋆char≡g⋆true = + module + UnivPropWithRepr + {Y : Type ℓ} + (perY : PartialEquivalenceRelation Y) + (F : FunctionalRelation perY perX) + (G : FunctionalRelation perY terminalPer) + (entailment : pointwiseEntailment perY Ωper (composeFuncRel _ _ _ G trueFuncRel) (composeFuncRel _ _ _ F charFuncRel)) where + + opaque + unfolding terminalFuncRel + G≤idY : pointwiseEntailment perY terminalPer G (terminalFuncRel perY) + G≤idY = + do + (stDG , stDG⊩isStrictDomainG) ← G .isStrictDomain + return + (stDG , + (λ { y tt* r r⊩Gy → stDG⊩isStrictDomainG y tt* r r⊩Gy })) + + opaque + idY≤G : pointwiseEntailment perY terminalPer (terminalFuncRel perY) G + idY≤G = F≤G→G≤F perY terminalPer G (terminalFuncRel perY) G≤idY + + opaque + unfolding trueFuncRel + trueFuncRelTruePredicate : ∀ a → (a ⊩ ∣ trueFuncRel .relation ∣ (tt* , fromPredicate truePredicate)) + trueFuncRelTruePredicate a = λ b → subst (λ p → (a ⨾ b) ⊩ ∣ p ∣ tt*) (sym (compIsIdFunc truePredicate)) tt* + + opaque + unfolding composeFuncRel + unfolding terminalFuncRel + {-# TERMINATING #-} + H : FunctionalRelation perY subPer + Predicate.isSetX (relation H) = isSet× (perY .isSetX) (perX .isSetX) + Predicate.∣ relation H ∣ (y , x) r = r ⊩ ∣ F .relation ∣ (y , x) + Predicate.isPropValued (relation H) (y , x) r = F .relation .isPropValued _ _ + isFunctionalRelation.isStrictDomain (isFuncRel H) = + do + (stFD , stFD⊩isStrictDomainF) ← F .isStrictDomain + return + (stFD , + (λ y x r r⊩Hyx → stFD⊩isStrictDomainF y x r r⊩Hyx)) + isFunctionalRelation.isStrictCodomain (isFuncRel H) = + do + (ent , ent⊩entailment) ← entailment + (a , a⊩idY≤G) ← idY≤G + (stFD , stFD⊩isStrictDomainF) ← F .isStrictDomain + (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + (relϕ , relϕ⊩isRelationalϕ) ← StrictRelation.isRelational ϕ + let + realizer : ApplStrTerm as 1 + realizer = + ` pair ̇ + (` stFC ̇ # zero) ̇ + (` relϕ ̇ + (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k)))) ̇ ` k) ̇ + (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k))) ̇ # zero)) + return + (λ* realizer , + (λ y x r r⊩Hyx → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₁pxy≡x _ _)) + (stFC⊩isStrictCodomainF y x _ r⊩Hyx) , + (equivFun + (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _)) + (do + (x' , ⊩Fyx' , ⊩x'~x' , ⊩ϕx'≤⊤ , ⊩⊤≤ϕx') ← + ent⊩entailment + y + (fromPredicate truePredicate) + (pair ⨾ (a ⨾ (stFD ⨾ r)) ⨾ k) + (return + (tt* , + subst + (λ r → r ⊩ ∣ G .relation ∣ (y , tt*)) + (sym (pr₁pxy≡x _ _)) + (a⊩idY≤G y tt* (stFD ⨾ r) (stFD⊩isStrictDomainF y x _ r⊩Hyx)) , + trueFuncRelTruePredicate _)) + let + ⊩x'~x = svF⊩isSingleValuedF y x' x _ _ ⊩Fyx' r⊩Hyx + ⊩ϕx = relϕ⊩isRelationalϕ x' x _ _ (⊩⊤≤ϕx' k (subst (λ p → k ⊩ ∣ p ∣ tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x'~x + return (subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₂pxy≡y _ _)) ⊩ϕx))))) + isFunctionalRelation.isRelational (isFuncRel H) = + do + (relF , relF⊩isRelationalF) ← isFunctionalRelation.isRelational (F .isFuncRel) + let + realizer : ApplStrTerm as 3 + realizer = ` relF ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero) + return + (λ*3 realizer , + (λ y y' x x' a b c ⊩y~y' ⊩Fyx (⊩x~x' , ⊩ϕx) → + subst + (λ r' → r' ⊩ ∣ F .relation ∣ (y' , x')) + (sym (λ*3ComputationRule realizer a b c)) + (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x'))) + isFunctionalRelation.isSingleValued (isFuncRel H) = + do + (ent , ent⊩entailment) ← entailment + (a , a⊩idY≤G) ← idY≤G + (stFD , stFD⊩isStrictDomainF) ← F .isStrictDomain + (svF , svF⊩isSingleValuedF) ← F .isSingleValued + (relϕ , relϕ⊩isRelationalϕ) ← StrictRelation.isRelational ϕ + let + realizer : ApplStrTerm as 2 + realizer = + ` pair ̇ + (` svF ̇ # one ̇ # zero) ̇ + (` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # one)) ̇ ` k)))) ̇ ` k) ̇ (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # one)) ̇ ` k))) ̇ # one)) + return + (λ*2 realizer , + (λ y x x' r₁ r₂ ⊩Fyx ⊩Fyx' → + subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x')) + (sym (cong (λ x → pr₁ ⨾ x) (λ*2ComputationRule realizer r₁ r₂) ∙ pr₁pxy≡x _ _)) + (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') , + (equivFun + (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _)) + (do + (x'' , ⊩Fyx'' , ⊩x''~x'' , ⊩ϕx''≤⊤ , ⊩⊤≤ϕx'') ← + ent⊩entailment + y + (fromPredicate truePredicate) + (pair ⨾ (a ⨾ (stFD ⨾ r₁)) ⨾ k) + (return + (tt* , + subst (λ r → r ⊩ ∣ G .relation ∣ (y , tt*)) (sym (pr₁pxy≡x _ _)) (a⊩idY≤G y tt* _ (stFD⊩isStrictDomainF y x _ ⊩Fyx)) , + trueFuncRelTruePredicate _)) + let + ⊩x''~x = svF⊩isSingleValuedF y x'' x _ _ ⊩Fyx'' ⊩Fyx + ⊩ϕx = relϕ⊩isRelationalϕ x'' x _ _ (⊩⊤≤ϕx'' k (subst (λ p → k ⊩ ∣ p ∣ tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x''~x + return + (subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) + (sym (cong (λ x → pr₂ ⨾ x) (λ*2ComputationRule realizer r₁ r₂) ∙ pr₂pxy≡y _ _)) + ⊩ϕx))))) + isFunctionalRelation.isTotal (isFuncRel H) = + do + (ent , ent⊩entailment) ← entailment + (a , a⊩idY≤G) ← idY≤G + let + realizer : ApplStrTerm as 1 + realizer = ` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ # zero) ̇ ` k)) + return + (λ* realizer , + (λ { y r r⊩y~y → + do + (x , ⊩Fyx , ⊩x~x , ⊩ϕx≤⊤ , ⊩⊤≤ϕx) ← + ent⊩entailment + y + (fromPredicate truePredicate) + (pair ⨾ (a ⨾ r) ⨾ k) + (return + (tt* , + subst (λ r → r ⊩ ∣ G .relation ∣ (y , tt*)) (sym (pr₁pxy≡x _ _)) (a⊩idY≤G y tt* r r⊩y~y) , + trueFuncRelTruePredicate _)) + return (x , subst (λ r' → r' ⊩ ∣ F .relation ∣ (y , x)) (sym (λ*ComputationRule realizer r)) ⊩Fyx) })) + + opaque + unfolding composeRTMorphism + unfolding incFuncRel + unfolding H + F≡H⋆inc : [ F ] ≡ [ H ] ⋆ [ incFuncRel ] + F≡H⋆inc = + let + answer = + do + (relF , relF⊩isRelationalF) ← isFunctionalRelation.isRelational (F .isFuncRel) + (stFD , stFD⊩isStrictDomainF) ← F .isStrictDomain + let + realizer : ApplStrTerm as 1 + realizer = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)) + return + (λ* realizer , + (λ y x r ⊩∃x' → + equivFun + (propTruncIdempotent≃ (F .relation .isPropValued _ _)) + (do + (x' , ⊩Hyx' , ⊩x'~x , ⊩ϕx') ← ⊩∃x' + return + (subst + (λ r' → r' ⊩ ∣ F .relation ∣ (y , x)) + (sym (λ*ComputationRule realizer r)) + (relF⊩isRelationalF y y x' x _ _ _ (stFD⊩isStrictDomainF y x' _ ⊩Hyx') ⊩Hyx' ⊩x'~x))))) + in eq/ _ _ (F≤G→G≤F perY perX (composeFuncRel _ _ _ H incFuncRel) F answer , answer) + + opaque + unfolding composeRTMorphism + unfolding terminalFuncRel + G≡H⋆terminal : [ G ] ≡ [ H ] ⋆ [ terminalFuncRel subPer ] + G≡H⋆terminal = + let + answer = + do + (stHD , stHD⊩isStrictDomainH) ← H .isStrictDomain + (a , a⊩idY≤G) ← idY≤G + let + realizer : ApplStrTerm as 1 + realizer = ` a ̇ (` stHD ̇ (` pr₁ ̇ # zero)) + return + (λ* realizer , + (λ { y tt* r r⊩∃x → + equivFun + (propTruncIdempotent≃ (G .relation .isPropValued _ _)) + (do + (x , ⊩Hyx , ⊩x~x , ⊩ϕx) ← r⊩∃x + return (subst (λ r' → r' ⊩ ∣ G .relation ∣ (y , tt*)) (sym (λ*ComputationRule realizer r)) (a⊩idY≤G y tt* _ (stHD⊩isStrictDomainH y x _ ⊩Hyx)))) })) + in eq/ _ _ (F≤G→G≤F perY terminalPer (composeFuncRel _ _ _ H (terminalFuncRel subPer)) G answer , answer) + + opaque + unfolding composeRTMorphism + unfolding H + unfolding incFuncRel + isUniqueH : ∀ (H' : FunctionalRelation perY subPer) → [ F ] ≡ [ H' ] ⋆ [ incFuncRel ] → [ G ] ≡ [ H' ] ⋆ [ terminalFuncRel subPer ] → [_] {R = bientailment perY subPer} H ≡ [ H' ] + isUniqueH H' F≡H'⋆inc G≡H'⋆term = + let + F≤H'⋆inc = [F]≡[G]→F≤G F (composeFuncRel _ _ _ H' incFuncRel) F≡H'⋆inc + answer : pointwiseEntailment _ _ H H' + answer = + do + (a , a⊩F≤H'⋆inc) ← F≤H'⋆inc + (relH' , relH'⊩isRelationalH) ← isFunctionalRelation.isRelational (H' .isFuncRel) + (stDH , stDH⊩isStrictDomainH) ← H .isStrictDomain + let + realizer : ApplStrTerm as 1 + realizer = ` relH' ̇ (` stDH ̇ # zero) ̇ (` pr₁ ̇ (` a ̇ # zero)) ̇ (` pr₂ ̇ (` a ̇ # zero)) + return + (λ* realizer , + (λ y x r r⊩Hyx → + equivFun + (propTruncIdempotent≃ (H' .relation .isPropValued _ _)) + (do + (x' , ⊩H'yx' , ⊩x'~x , ⊩ϕx') ← a⊩F≤H'⋆inc y x r r⊩Hyx + return + (subst + (λ r' → r' ⊩ ∣ H' .relation ∣ (y , x)) + (sym (λ*ComputationRule realizer r)) + (relH'⊩isRelationalH y y x' x _ _ _ (stDH⊩isStrictDomainH y x r r⊩Hyx) ⊩H'yx' (⊩x'~x , ⊩ϕx')))))) + in + eq/ _ _ (answer , (F≤G→G≤F _ _ H H' answer)) + + opaque + classifies : isPullback RT (subobjectCospan [ charFuncRel ]) [ incFuncRel ] [ terminalFuncRel subPer ] subobjectSquareCommutes + classifies {Y , perY} f g f⋆char≡g⋆true = SQ.elimProp2 {P = λ f g → ∀ (commutes : f ⋆ [ charFuncRel ] ≡ g ⋆ [ trueFuncRel ]) → ∃![ hk ∈ RTMorphism perY subPer ] (f ≡ hk ⋆ [ incFuncRel ]) × (g ≡ hk ⋆ [ terminalFuncRel subPer ])} (λ f g → isPropΠ λ _ → isPropIsContr) (λ F G F⋆char≡G⋆true → - let - H : FunctionalRelation perY subPer - H = - makeFunctionalRelation - (makePredicate - (isSet× (perY .isSetX) (perX .isSetX)) - (λ { (y , x) r → (pr₁ ⨾ r) ⊩ ∣ F .relation ∣ (y , x) × (pr₂ ⨾ r) ⊩ ∣ ϕ .predicate ∣ x }) - λ { (y , x) r → isProp× (F .relation .isPropValued _ _) (ϕ .predicate .isPropValued _ _) }) - (makeIsFunctionalRelation - (do - (stF , stF⊩isStrictDomainF) ← F .isStrictDomain - return - ({!!} , - (λ { y x r (pr₁r⊩Fyx , pr₂r⊩ϕx) → - subst (λ r' → r' ⊩ ∣ perY .equality ∣ (y , y)) (sym {!!}) (stF⊩isStrictDomainF y x (pr₁ ⨾ r) pr₁r⊩Fyx)}))) - (do - (stFC , stFC⊩isStrictCodomainF) ← F .isStrictCodomain - return - ({!!} , - (λ { y x r (pr₁r⊩Fyx , pr₂r⊩ϕx) → - {!stFC⊩isStrictCodomainF y x (pr₁ ⨾ r) pr₁r⊩Fyx!} , - {!pr₂r⊩ϕx!} }))) - (do - return {!!}) - (do - return {!!}) - (do - return - ({!!} , - (λ { y r r⊩y~y → - return - ({!!} , {!!} , {!!})})))) - in - uniqueExists - [ H ] - ({!!} , - {!!}) - (λ hk' → isProp× (squash/ _ _) (squash/ _ _)) - {!!}) + let + entailment = [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G F charFuncRel G trueFuncRel F⋆char≡G⋆true + in + uniqueExists + [ UnivPropWithRepr.H perY F G entailment ] + (UnivPropWithRepr.F≡H⋆inc perY F G entailment , + UnivPropWithRepr.G≡H⋆terminal perY F G entailment) + (λ hk' → isProp× (squash/ _ _) (squash/ _ _)) + -- nested eliminator 🤮 + λ { h' (f≡h'⋆inc , g≡h'⋆term) → + SQ.elimProp + {P = λ h' → ∀ (comm1 : [ F ] ≡ h' ⋆ [ incFuncRel ]) (comm2 : [ G ] ≡ h' ⋆ [ terminalFuncRel subPer ]) → [ UnivPropWithRepr.H perY F G entailment ] ≡ h'} + (λ h' → isPropΠ λ _ → isPropΠ λ _ → squash/ _ _) + (λ H' F≡H'⋆inc G≡H'⋆term → + UnivPropWithRepr.isUniqueH perY F G entailment H' F≡H'⋆inc G≡H'⋆term) + h' + f≡h'⋆inc + g≡h'⋆term }) f g f⋆char≡g⋆true + + module + PullbackHelper + (C : FunctionalRelation perX Ωper) + (commutes : [ incFuncRel ] ⋆ [ C ] ≡ [ terminalFuncRel subPer ] ⋆ [ trueFuncRel ]) + (classifies : isPullback RT (subobjectCospan [ C ]) [ incFuncRel ] [ terminalFuncRel subPer ] commutes) where + + {-# TERMINATING #-} + ψ : StrictRelation perX + Predicate.isSetX (predicate ψ) = perX .isSetX + Predicate.∣ predicate ψ ∣ x r = r ⊩ ∣ C .relation ∣ (x , ⊤) + Predicate.isPropValued (predicate ψ) x r = C .relation .isPropValued _ _ + isStrictRelation.isStrict (isStrictRelationPredicate ψ) = + do + (stDC , stDC⊩isStrictDomainC) ← C .isStrictDomain + return + (stDC , + λ x r r⊩Cx⊤ → stDC⊩isStrictDomainC x (fromPredicate truePredicate) r r⊩Cx⊤) + isStrictRelation.isRelational (isStrictRelationPredicate ψ) = + do + (relC , relC⊩isRelationalC) ← isFunctionalRelation.isRelational (C .isFuncRel) + (stCC , stCC⊩isStrictCodomainC) ← C .isStrictCodomain + let + realizer : ApplStrTerm as 2 + realizer = ` relC ̇ # zero ̇ # one ̇ (` stCC ̇ # one) + return + (λ*2 realizer , + λ x x' a b a⊩Cx⊤ b⊩x~x' → + subst (λ r' → r' ⊩ ∣ C .relation ∣ (x' , ⊤)) (sym (λ*2ComputationRule realizer a b)) (relC⊩isRelationalC x x' ⊤ ⊤ _ _ _ b⊩x~x' a⊩Cx⊤ (stCC⊩isStrictCodomainC x ⊤ a a⊩Cx⊤))) + + perψ = InducedSubobject.subPer perX ψ + incFuncRelψ = InducedSubobject.incFuncRel perX ψ + + opaque + unfolding composeRTMorphism + unfolding InducedSubobject.incFuncRel + unfolding terminalFuncRel + unfolding trueFuncRel + pbSqCommutes : [ incFuncRelψ ] ⋆ [ C ] ≡ [ terminalFuncRel perψ ] ⋆ [ trueFuncRel ] + pbSqCommutes = + let + answer = + do + (stDC , stDC⊩isStrictDomainC) ← C .isStrictDomain + (stCC , stCC⊩isStrictCodomainC) ← C .isStrictCodomain + (svC , svC⊩isSingleValuedC) ← C .isSingleValued + (relC , relC⊩isRelationalC) ← isFunctionalRelation.isRelational (C .isFuncRel) + (sX , sX⊩isSymmetricX) ← perX .isSymmetric + let + closure : ApplStrTerm as 2 + closure = ` pr₁ ̇ (` svC ̇ (` pr₂ ̇ (` pr₁ ̇ # one)) ̇ (` relC ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # one))) ̇ (` pr₂ ̇ # one) ̇ (` stCC ̇ (` pr₂ ̇ # one)))) ̇ ` k + + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` pair ̇ (` stDC ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (λ*abst closure) + return + (λ* realizer , + λ { x p r r⊩∃x' → + do + (x' , (⊩x~x' , ⊩Cx⊤) , ⊩Cx'p) ← r⊩∃x' + let + ⊩Cxp = relC⊩isRelationalC x' x p p _ _ _ (sX⊩isSymmetricX x x' _ ⊩x~x') ⊩Cx'p (stCC⊩isStrictCodomainC x' p _ ⊩Cx'p) + (⊩⊤≤p , p≤⊤) = svC⊩isSingleValuedC x ⊤ p _ _ ⊩Cx⊤ ⊩Cxp + return + (tt* , + (subst + (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) + (sym + (cong (λ x → pr₁ ⨾ (pr₁ ⨾ x)) (λ*ComputationRule realizer r) ∙ + cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₁pxy≡x _ _ )) + (stDC⊩isStrictDomainC x ⊤ _ ⊩Cx⊤) , + subst + (λ r' → r' ⊩ ∣ C .relation ∣ (x , ⊤)) + (sym + (cong (λ x → pr₂ ⨾ (pr₁ ⨾ x)) (λ*ComputationRule realizer r) ∙ + cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ + pr₂pxy≡y _ _)) + ⊩Cx⊤) , + λ a → + subst + (λ r' → r' ⊩ ∣ toPredicate p ∣ tt*) + (sym + (cong (λ x → pr₂ ⨾ x ⨾ a) (λ*ComputationRule realizer r) ∙ + cong (λ x → x ⨾ a) (pr₂pxy≡y _ _) ∙ + βreduction closure a (r ∷ []))) + (⊩⊤≤p k (subst (λ q → k ⊩ ∣ q ∣ tt*) (sym (compIsIdFunc truePredicate)) tt*))) }) + in eq/ _ _ (answer , F≤G→G≤F _ _ (composeFuncRel _ _ _ incFuncRelψ C) (composeFuncRel _ _ _ (terminalFuncRel perψ) trueFuncRel) answer) + + opaque + unfolding InducedSubobject.incFuncRel + unfolding composeFuncRel + ⊩Cx⊤≤ϕx : ∃[ ent ∈ A ] (∀ (x : X) (r : A) → r ⊩ ∣ C .relation ∣ (x , ⊤) → (ent ⨾ r) ⊩ ∣ ϕ .predicate ∣ x) + ⊩Cx⊤≤ϕx = + let + ((h , incψ≡h⋆incϕ , termψ≡h⋆termϕ) , isUniqueH) = classifies [ incFuncRelψ ] [ terminalFuncRel perψ ] pbSqCommutes + in + SQ.elimProp + {P = λ h → ∀ (incψ≡h⋆incϕ : [ incFuncRelψ ] ≡ h ⋆ [ incFuncRel ]) → ∃[ ent ∈ A ] (∀ (x : X) (r : A) → r ⊩ ∣ C .relation ∣ (x , ⊤) → (ent ⨾ r) ⊩ ∣ ϕ .predicate ∣ x)} + (λ h → isPropΠ λ _ → isPropPropTrunc) + (λ H incψ≡H⋆incϕ → + do + (a , a⊩incψ≤H⋆incϕ) ← [F]≡[G]⋆[H]→F≤G⋆H incFuncRelψ H incFuncRel incψ≡H⋆incϕ + (stDC , stDC⊩isStrictDomainC) ← C .isStrictDomain + (relϕ , relϕ⊩isRelationalϕ) ← isStrictRelation.isRelational (ϕ .isStrictRelationPredicate) + let + realizer = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) + return + (λ* realizer , + (λ x r r⊩Cx⊤ → + equivFun + (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _)) + (do + (x' , ⊩Hxx' , ⊩x'~x , ⊩ϕx') ← + a⊩incψ≤H⋆incϕ + x x + (pair ⨾ (stDC ⨾ r) ⨾ r) + (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (pr₁pxy≡x _ _)) (stDC⊩isStrictDomainC x ⊤ r r⊩Cx⊤) , + subst (λ r' → r' ⊩ ∣ C .relation ∣ (x , ⊤)) (sym (pr₂pxy≡y _ _)) r⊩Cx⊤) + return + (subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (sym (λ*ComputationRule realizer r)) (relϕ⊩isRelationalϕ x' x _ _ ⊩ϕx' ⊩x'~x)))))) + h + incψ≡h⋆incϕ + + opaque + unfolding trueFuncRel + unfolding composeFuncRel + unfolding incFuncRel + unfolding terminalFuncRel + isUniqueCharMorphism : + ∀ (char : RTMorphism perX Ωper) + → (commutes : [ incFuncRel ] ⋆ char ≡ [ terminalFuncRel subPer ] ⋆ [ trueFuncRel ]) + → (classifies : isPullback RT (subobjectCospan char) [ incFuncRel ] [ terminalFuncRel subPer ] commutes) + → char ≡ [ charFuncRel ] + isUniqueCharMorphism char commutes classifies = + SQ.elimProp + {P = + λ char → + ∀ (commutes : [ incFuncRel ] ⋆ char ≡ [ terminalFuncRel subPer ] ⋆ [ trueFuncRel ]) + (classifies : isPullback RT (subobjectCospan char) [ incFuncRel ] [ terminalFuncRel subPer ] commutes) + → char ≡ [ charFuncRel ]} + (λ char → isPropΠ λ commutes → isPropΠ λ classifies → squash/ _ _) + (λ charFuncRel' commutes classifies → + let + answer = + do + (stDX' , stDX'⊩isStrictDomainX') ← charFuncRel' .isStrictDomain + (relX' , relX'⊩isRelationalX') ← isFunctionalRelation.isRelational (charFuncRel' .isFuncRel) + (a , a⊩inc⋆X'≤term⋆true) ← [F]⋆[G]≡[H]⋆[I]→F⋆G≤H⋆I incFuncRel charFuncRel' (terminalFuncRel subPer) trueFuncRel commutes + (b , b⊩term⋆true≤inc⋆X') ← [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G incFuncRel charFuncRel' (terminalFuncRel subPer) trueFuncRel commutes + (d , d⊩X'x⊤≤ϕx) ← PullbackHelper.⊩Cx⊤≤ϕx charFuncRel' commutes classifies + let + closure1 : ApplStrTerm as 2 + closure1 = ` pr₂ ̇ (` a ̇ (` pair ̇ (` pair ̇ (` stDX' ̇ # one) ̇ # zero) ̇ # one)) ̇ ` k + closure2 : ApplStrTerm as 2 + closure2 = ` d ̇ (` relX' ̇ (` stDX' ̇ # one) ̇ # one ̇ (` pair ̇ ` k ̇ (` k ̇ # zero))) + realizer : ApplStrTerm as 1 + realizer = ` pair ̇ (` stDX' ̇ # zero) ̇ (` pair ̇ λ*abst closure1 ̇ λ*abst closure2) + return + (λ* realizer , + (λ { x p r r⊩X'xp → + let + ⊩x~x = stDX'⊩isStrictDomainX' x p r r⊩X'xp + in + subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₁pxy≡x _ _)) ⊩x~x , + (λ b b⊩ϕx → + let + goal = + a⊩inc⋆X'≤term⋆true + x p (pair ⨾ (pair ⨾ (stDX' ⨾ r) ⨾ b) ⨾ r) + (return + (x , (subst (λ r' → r' ⊩ ∣ perX .equality ∣ (x , x)) (sym (cong (λ x → pr₁ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₁pxy≡x _ _)) ⊩x~x , + subst (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) (sym (cong (λ x → pr₂ ⨾ x) (pr₁pxy≡x _ _) ∙ pr₂pxy≡y _ _)) b⊩ϕx) , + subst (λ r' → r' ⊩ ∣ charFuncRel' .relation ∣ (x , p)) (sym (pr₂pxy≡y _ _)) r⊩X'xp)) + + eq : pr₁ ⨾ (pr₂ ⨾ (λ* realizer ⨾ r)) ⨾ b ≡ pr₂ ⨾ (a ⨾ (pair ⨾ (pair ⨾ (stDX' ⨾ r) ⨾ b) ⨾ r)) ⨾ k + eq = + cong (λ x → pr₁ ⨾ (pr₂ ⨾ x) ⨾ b) (λ*ComputationRule realizer r) ∙ cong (λ x → pr₁ ⨾ x ⨾ b) (pr₂pxy≡y _ _) ∙ cong (λ x → x ⨾ b) (pr₁pxy≡x _ _) ∙ βreduction closure1 b (r ∷ []) + in + equivFun + (propTruncIdempotent≃ (toPredicate p .isPropValued _ _)) + (do + (tt* , ⊩'x~x , ⊩⊤≤p) ← goal + return (subst (λ r' → r' ⊩ ∣ toPredicate p ∣ tt*) (sym eq) (⊩⊤≤p k)))) , + (λ c c⊩p → + let + ⊩X'x⊤ = + relX'⊩isRelationalX' + x x p ⊤ _ _ + (pair ⨾ k ⨾ (k ⨾ c)) + ⊩x~x r⊩X'xp + ((λ b b⊩p → subst (λ q → (pr₁ ⨾ (pair ⨾ k ⨾ (k ⨾ c))) ⊩ ∣ q ∣ tt*) (sym (compIsIdFunc truePredicate)) tt*) , + (λ b b⊩⊤ → subst (λ r' → r' ⊩ ∣ toPredicate p ∣ tt*) (sym (cong (λ x → x ⨾ b) (pr₂pxy≡y _ _) ∙ kab≡a _ _)) c⊩p)) + + eq : pr₂ ⨾ (pr₂ ⨾ (λ* realizer ⨾ r)) ⨾ c ≡ d ⨾ (relX' ⨾ (stDX' ⨾ r) ⨾ r ⨾ (pair ⨾ k ⨾ (k ⨾ c))) + eq = + cong (λ x → pr₂ ⨾ (pr₂ ⨾ x) ⨾ c) (λ*ComputationRule realizer r) ∙ + cong (λ x → pr₂ ⨾ x ⨾ c) (pr₂pxy≡y _ _) ∙ + cong (λ x → x ⨾ c) (pr₂pxy≡y _ _) ∙ + βreduction closure2 c (r ∷ []) + in + subst + (λ r' → r' ⊩ ∣ ϕ .predicate ∣ x) + (sym eq) + (d⊩X'x⊤≤ϕx x _ ⊩X'x⊤)) })) + in eq/ _ _ (answer , F≤G→G≤F _ _ charFuncRel' charFuncRel answer)) + char + commutes + classifies From 05772cc2384158c451f535bfc4a43220668501a4 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 16 Apr 2024 11:10:38 +0530 Subject: [PATCH 41/61] Start working on general case for subobject classifiers --- .../Topos/SubobjectClassifier.agda | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/src/Realizability/Topos/SubobjectClassifier.agda b/src/Realizability/Topos/SubobjectClassifier.agda index d1d9a56..5cc8c95 100644 --- a/src/Realizability/Topos/SubobjectClassifier.agda +++ b/src/Realizability/Topos/SubobjectClassifier.agda @@ -936,3 +936,33 @@ module ClassifiesStrictRelations char commutes classifies + +module ClassifiesSubobjects where + + subobjectClassifierUnivProp : Type _ + subobjectClassifierUnivProp = + ∀ {X Y : Type ℓ} + {perX : PartialEquivalenceRelation X} + {perY : PartialEquivalenceRelation Y} + → (f : RTMorphism perX perY) + → isMonic RT f + → ∃![ char ∈ RTMorphism perY Ωper ] + Σ[ commutes ∈ f ⋆ char ≡ [ terminalFuncRel perX ] ⋆ [ trueFuncRel ] ] + isPullback RT (cospan (Y , perY) ((ResizedPredicate Unit*) , Ωper) (Unit* , terminalPer) char [ trueFuncRel ]) f [ terminalFuncRel perX ] commutes + + isSubobjectClassifier : subobjectClassifierUnivProp + isSubobjectClassifier {X} {Y} {perX} {perY} f isMonicF = + SQ.elimProp + {P = λ f → ∀ (isMonic : isMonic RT f) → ∃![ char ∈ RTMorphism perY Ωper ] (Σ[ commutes ∈ f ⋆ char ≡ [ terminalFuncRel perX ] ⋆ [ trueFuncRel ] ] isPullback RT (cospan (Y , perY) ((ResizedPredicate Unit*) , Ωper) (Unit* , terminalPer) char [ trueFuncRel ]) f [ terminalFuncRel perX ] commutes) } + (λ f → isPropΠ λ isMonicF → isPropIsContr) + (λ F isMonicF → + let + ϕ = SubobjectIsoMonicFuncRel.ψ perY perX F isMonicF + in + uniqueExists + [ ClassifiesStrictRelations.charFuncRel Y perY ϕ ] + ({!ClassifiesStrictRelations.subobjectSquareCommutes Y perY ϕ!} , {!!}) + {!!} + {!!}) + f + isMonicF From 4e956d4251aa72a23476139a1aa20c0b021d5ad8 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 2 May 2024 22:14:37 +0530 Subject: [PATCH 42/61] Define PERs --- .gitignore | 1 + Dockerfile | 62 -- ...ure.agda => ApplicativeStructure.lagda.md} | 71 +- src/Realizability/PERs/PER.agda | 82 ++ src/Realizability/PERs/SystemF.agda | 40 + src/Realizability/Topos/Counterexample.agda | 63 ++ src/Realizability/Topos/Everything.agda | 1 + src/docs/Agda.Builtin.Bool.html | 15 + src/docs/Agda.Builtin.Char.html | 18 + src/docs/Agda.Builtin.Cubical.Equiv.html | 107 +++ src/docs/Agda.Builtin.Cubical.Glue.html | 18 + src/docs/Agda.Builtin.Cubical.HCompU.html | 78 ++ src/docs/Agda.Builtin.Cubical.Path.html | 15 + src/docs/Agda.Builtin.Cubical.Sub.html | 18 + src/docs/Agda.Builtin.Float.html | 209 +++++ src/docs/Agda.Builtin.FromNat.html | 16 + src/docs/Agda.Builtin.FromNeg.html | 16 + src/docs/Agda.Builtin.Int.html | 18 + src/docs/Agda.Builtin.List.html | 16 + src/docs/Agda.Builtin.Maybe.html | 9 + src/docs/Agda.Builtin.Nat.html | 134 +++ src/docs/Agda.Builtin.Reflection.html | 470 ++++++++++ src/docs/Agda.Builtin.Sigma.html | 17 + src/docs/Agda.Builtin.String.html | 36 + src/docs/Agda.Builtin.Unit.html | 10 + src/docs/Agda.Builtin.Word.html | 13 + src/docs/Agda.Primitive.Cubical.html | 78 ++ src/docs/Agda.Primitive.html | 41 + src/docs/Agda.css | 41 + src/docs/Cubical.Core.Everything.html | 8 + src/docs/Cubical.Core.Glue.html | 139 +++ src/docs/Cubical.Core.Primitives.html | 209 +++++ src/docs/Cubical.Data.Bool.Base.html | 101 +++ src/docs/Cubical.Data.Bool.Properties.html | 431 +++++++++ .../Cubical.Data.Bool.SwitchStatement.html | 42 + src/docs/Cubical.Data.Bool.html | 6 + src/docs/Cubical.Data.Empty.Base.html | 25 + src/docs/Cubical.Data.Empty.Properties.html | 37 + src/docs/Cubical.Data.Empty.html | 5 + src/docs/Cubical.Data.FinData.Base.html | 86 ++ src/docs/Cubical.Data.FinData.Properties.html | 373 ++++++++ src/docs/Cubical.Data.FinData.html | 5 + src/docs/Cubical.Data.List.Base.html | 52 ++ src/docs/Cubical.Data.List.Properties.html | 181 ++++ src/docs/Cubical.Data.List.html | 5 + src/docs/Cubical.Data.Maybe.Base.html | 30 + src/docs/Cubical.Data.Maybe.Properties.html | 181 ++++ src/docs/Cubical.Data.Maybe.html | 5 + src/docs/Cubical.Data.Nat.Base.html | 90 ++ src/docs/Cubical.Data.Nat.Literals.html | 22 + src/docs/Cubical.Data.Nat.Order.html | 529 +++++++++++ src/docs/Cubical.Data.Nat.Properties.html | 342 +++++++ src/docs/Cubical.Data.Nat.html | 5 + src/docs/Cubical.Data.Prod.Base.html | 60 ++ src/docs/Cubical.Data.Sigma.Base.html | 52 ++ src/docs/Cubical.Data.Sigma.Properties.html | 486 ++++++++++ src/docs/Cubical.Data.Sigma.html | 6 + src/docs/Cubical.Data.Sum.Base.html | 37 + src/docs/Cubical.Data.Sum.Properties.html | 294 ++++++ src/docs/Cubical.Data.Sum.html | 5 + src/docs/Cubical.Data.Unit.Base.html | 25 + src/docs/Cubical.Data.Unit.Properties.html | 121 +++ src/docs/Cubical.Data.Unit.html | 5 + src/docs/Cubical.Data.Vec.Base.html | 67 ++ src/docs/Cubical.Data.Vec.NAry.html | 55 ++ src/docs/Cubical.Data.Vec.Properties.html | 141 +++ src/docs/Cubical.Data.Vec.html | 5 + .../Cubical.Foundations.CartesianKanOps.html | 180 ++++ src/docs/Cubical.Foundations.Equiv.Base.html | 65 ++ .../Cubical.Foundations.Equiv.Fiberwise.html | 99 +++ ...Cubical.Foundations.Equiv.HalfAdjoint.html | 138 +++ .../Cubical.Foundations.Equiv.Properties.html | 260 ++++++ src/docs/Cubical.Foundations.Equiv.html | 326 +++++++ src/docs/Cubical.Foundations.Function.html | 163 ++++ .../Cubical.Foundations.GroupoidLaws.html | 532 +++++++++++ src/docs/Cubical.Foundations.HLevels.html | 841 ++++++++++++++++++ src/docs/Cubical.Foundations.Isomorphism.html | 225 +++++ src/docs/Cubical.Foundations.Path.html | 439 +++++++++ .../Cubical.Foundations.Pointed.Base.html | 151 ++++ .../Cubical.Foundations.Pointed.FunExt.html | 48 + ...bical.Foundations.Pointed.Homogeneous.html | 216 +++++ .../Cubical.Foundations.Pointed.Homotopy.html | 119 +++ ...ubical.Foundations.Pointed.Properties.html | 246 +++++ src/docs/Cubical.Foundations.Pointed.html | 9 + src/docs/Cubical.Foundations.Powerset.html | 65 ++ src/docs/Cubical.Foundations.Prelude.html | 616 +++++++++++++ src/docs/Cubical.Foundations.SIP.html | 124 +++ src/docs/Cubical.Foundations.Structure.html | 47 + src/docs/Cubical.Foundations.Transport.html | 221 +++++ src/docs/Cubical.Foundations.Univalence.html | 343 +++++++ src/docs/Cubical.Functions.Embedding.html | 476 ++++++++++ src/docs/Cubical.Functions.Fibration.html | 111 +++ src/docs/Cubical.Functions.Fixpoint.html | 43 + src/docs/Cubical.Functions.FunExtEquiv.html | 193 ++++ src/docs/Cubical.Functions.Involution.html | 42 + src/docs/Cubical.Functions.Logic.html | 290 ++++++ ...cal.HITs.PropositionalTruncation.Base.html | 17 + ...Ts.PropositionalTruncation.MagicTrick.html | 88 ++ ...Ts.PropositionalTruncation.Properties.html | 564 ++++++++++++ .../Cubical.HITs.PropositionalTruncation.html | 7 + src/docs/Cubical.Homotopy.Base.html | 19 + src/docs/Cubical.Induction.WellFounded.html | 47 + src/docs/Cubical.Reflection.Base.html | 50 ++ src/docs/Cubical.Reflection.StrictEquiv.html | 81 ++ src/docs/Cubical.Relation.Nullary.Base.html | 70 ++ .../Cubical.Relation.Nullary.Properties.html | 203 +++++ src/docs/Cubical.Relation.Nullary.html | 5 + src/docs/Cubical.Structures.Axioms.html | 69 ++ src/docs/Cubical.Structures.Pointed.html | 59 ++ .../Cubical.Tactics.NatSolver.EvalHom.html | 196 ++++ ...Cubical.Tactics.NatSolver.HornerForms.html | 100 +++ ...bical.Tactics.NatSolver.NatExpression.html | 28 + .../Cubical.Tactics.NatSolver.Reflection.html | 145 +++ .../Cubical.Tactics.NatSolver.Solver.html | 123 +++ src/docs/Cubical.Tactics.NatSolver.html | 12 + .../Cubical.Tactics.Reflection.Utilities.html | 35 + .../Cubical.Tactics.Reflection.Variables.html | 81 ++ src/docs/Cubical.Tactics.Reflection.html | 114 +++ .../Realizability.ApplicativeStructure.html | 231 +++++ .../Realizability.ApplicativeStructure.md | 215 +++++ 120 files changed, 15103 insertions(+), 65 deletions(-) delete mode 100644 Dockerfile rename src/Realizability/{ApplicativeStructure.agda => ApplicativeStructure.lagda.md} (70%) create mode 100644 src/Realizability/PERs/PER.agda create mode 100644 src/Realizability/PERs/SystemF.agda create mode 100644 src/Realizability/Topos/Counterexample.agda create mode 100644 src/docs/Agda.Builtin.Bool.html create mode 100644 src/docs/Agda.Builtin.Char.html create mode 100644 src/docs/Agda.Builtin.Cubical.Equiv.html create mode 100644 src/docs/Agda.Builtin.Cubical.Glue.html create mode 100644 src/docs/Agda.Builtin.Cubical.HCompU.html create mode 100644 src/docs/Agda.Builtin.Cubical.Path.html create mode 100644 src/docs/Agda.Builtin.Cubical.Sub.html create mode 100644 src/docs/Agda.Builtin.Float.html create mode 100644 src/docs/Agda.Builtin.FromNat.html create mode 100644 src/docs/Agda.Builtin.FromNeg.html create mode 100644 src/docs/Agda.Builtin.Int.html create mode 100644 src/docs/Agda.Builtin.List.html create mode 100644 src/docs/Agda.Builtin.Maybe.html create mode 100644 src/docs/Agda.Builtin.Nat.html create mode 100644 src/docs/Agda.Builtin.Reflection.html create mode 100644 src/docs/Agda.Builtin.Sigma.html create mode 100644 src/docs/Agda.Builtin.String.html create mode 100644 src/docs/Agda.Builtin.Unit.html create mode 100644 src/docs/Agda.Builtin.Word.html create mode 100644 src/docs/Agda.Primitive.Cubical.html create mode 100644 src/docs/Agda.Primitive.html create mode 100644 src/docs/Agda.css create mode 100644 src/docs/Cubical.Core.Everything.html create mode 100644 src/docs/Cubical.Core.Glue.html create mode 100644 src/docs/Cubical.Core.Primitives.html create mode 100644 src/docs/Cubical.Data.Bool.Base.html create mode 100644 src/docs/Cubical.Data.Bool.Properties.html create mode 100644 src/docs/Cubical.Data.Bool.SwitchStatement.html create mode 100644 src/docs/Cubical.Data.Bool.html create mode 100644 src/docs/Cubical.Data.Empty.Base.html create mode 100644 src/docs/Cubical.Data.Empty.Properties.html create mode 100644 src/docs/Cubical.Data.Empty.html create mode 100644 src/docs/Cubical.Data.FinData.Base.html create mode 100644 src/docs/Cubical.Data.FinData.Properties.html create mode 100644 src/docs/Cubical.Data.FinData.html create mode 100644 src/docs/Cubical.Data.List.Base.html create mode 100644 src/docs/Cubical.Data.List.Properties.html create mode 100644 src/docs/Cubical.Data.List.html create mode 100644 src/docs/Cubical.Data.Maybe.Base.html create mode 100644 src/docs/Cubical.Data.Maybe.Properties.html create mode 100644 src/docs/Cubical.Data.Maybe.html create mode 100644 src/docs/Cubical.Data.Nat.Base.html create mode 100644 src/docs/Cubical.Data.Nat.Literals.html create mode 100644 src/docs/Cubical.Data.Nat.Order.html create mode 100644 src/docs/Cubical.Data.Nat.Properties.html create mode 100644 src/docs/Cubical.Data.Nat.html create mode 100644 src/docs/Cubical.Data.Prod.Base.html create mode 100644 src/docs/Cubical.Data.Sigma.Base.html create mode 100644 src/docs/Cubical.Data.Sigma.Properties.html create mode 100644 src/docs/Cubical.Data.Sigma.html create mode 100644 src/docs/Cubical.Data.Sum.Base.html create mode 100644 src/docs/Cubical.Data.Sum.Properties.html create mode 100644 src/docs/Cubical.Data.Sum.html create mode 100644 src/docs/Cubical.Data.Unit.Base.html create mode 100644 src/docs/Cubical.Data.Unit.Properties.html create mode 100644 src/docs/Cubical.Data.Unit.html create mode 100644 src/docs/Cubical.Data.Vec.Base.html create mode 100644 src/docs/Cubical.Data.Vec.NAry.html create mode 100644 src/docs/Cubical.Data.Vec.Properties.html create mode 100644 src/docs/Cubical.Data.Vec.html create mode 100644 src/docs/Cubical.Foundations.CartesianKanOps.html create mode 100644 src/docs/Cubical.Foundations.Equiv.Base.html create mode 100644 src/docs/Cubical.Foundations.Equiv.Fiberwise.html create mode 100644 src/docs/Cubical.Foundations.Equiv.HalfAdjoint.html create mode 100644 src/docs/Cubical.Foundations.Equiv.Properties.html create mode 100644 src/docs/Cubical.Foundations.Equiv.html create mode 100644 src/docs/Cubical.Foundations.Function.html create mode 100644 src/docs/Cubical.Foundations.GroupoidLaws.html create mode 100644 src/docs/Cubical.Foundations.HLevels.html create mode 100644 src/docs/Cubical.Foundations.Isomorphism.html create mode 100644 src/docs/Cubical.Foundations.Path.html create mode 100644 src/docs/Cubical.Foundations.Pointed.Base.html create mode 100644 src/docs/Cubical.Foundations.Pointed.FunExt.html create mode 100644 src/docs/Cubical.Foundations.Pointed.Homogeneous.html create mode 100644 src/docs/Cubical.Foundations.Pointed.Homotopy.html create mode 100644 src/docs/Cubical.Foundations.Pointed.Properties.html create mode 100644 src/docs/Cubical.Foundations.Pointed.html create mode 100644 src/docs/Cubical.Foundations.Powerset.html create mode 100644 src/docs/Cubical.Foundations.Prelude.html create mode 100644 src/docs/Cubical.Foundations.SIP.html create mode 100644 src/docs/Cubical.Foundations.Structure.html create mode 100644 src/docs/Cubical.Foundations.Transport.html create mode 100644 src/docs/Cubical.Foundations.Univalence.html create mode 100644 src/docs/Cubical.Functions.Embedding.html create mode 100644 src/docs/Cubical.Functions.Fibration.html create mode 100644 src/docs/Cubical.Functions.Fixpoint.html create mode 100644 src/docs/Cubical.Functions.FunExtEquiv.html create mode 100644 src/docs/Cubical.Functions.Involution.html create mode 100644 src/docs/Cubical.Functions.Logic.html create mode 100644 src/docs/Cubical.HITs.PropositionalTruncation.Base.html create mode 100644 src/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html create mode 100644 src/docs/Cubical.HITs.PropositionalTruncation.Properties.html create mode 100644 src/docs/Cubical.HITs.PropositionalTruncation.html create mode 100644 src/docs/Cubical.Homotopy.Base.html create mode 100644 src/docs/Cubical.Induction.WellFounded.html create mode 100644 src/docs/Cubical.Reflection.Base.html create mode 100644 src/docs/Cubical.Reflection.StrictEquiv.html create mode 100644 src/docs/Cubical.Relation.Nullary.Base.html create mode 100644 src/docs/Cubical.Relation.Nullary.Properties.html create mode 100644 src/docs/Cubical.Relation.Nullary.html create mode 100644 src/docs/Cubical.Structures.Axioms.html create mode 100644 src/docs/Cubical.Structures.Pointed.html create mode 100644 src/docs/Cubical.Tactics.NatSolver.EvalHom.html create mode 100644 src/docs/Cubical.Tactics.NatSolver.HornerForms.html create mode 100644 src/docs/Cubical.Tactics.NatSolver.NatExpression.html create mode 100644 src/docs/Cubical.Tactics.NatSolver.Reflection.html create mode 100644 src/docs/Cubical.Tactics.NatSolver.Solver.html create mode 100644 src/docs/Cubical.Tactics.NatSolver.html create mode 100644 src/docs/Cubical.Tactics.Reflection.Utilities.html create mode 100644 src/docs/Cubical.Tactics.Reflection.Variables.html create mode 100644 src/docs/Cubical.Tactics.Reflection.html create mode 100644 src/docs/Realizability.ApplicativeStructure.html create mode 100644 src/docs/Realizability.ApplicativeStructure.md diff --git a/.gitignore b/.gitignore index 1fbe28c..c02bd83 100644 --- a/.gitignore +++ b/.gitignore @@ -1,3 +1,4 @@ *.agdai *.agda~ *.DS_Store +*.lagda.md~ \ No newline at end of file diff --git a/Dockerfile b/Dockerfile deleted file mode 100644 index 926e7d3..0000000 --- a/Dockerfile +++ /dev/null @@ -1,62 +0,0 @@ -# syntax=docker/dockerfile:1 - -# Comments are provided throughout this file to help you get started. -# If you need more help, visit the Dockerfile reference guide at -# https://docs.docker.com/go/dockerfile-reference/ - -################################################################################ -# Pick a base image to serve as the foundation for the other build stages in -# this file. -# -# For illustrative purposes, the following FROM command -# is using the alpine image (see https://hub.docker.com/_/alpine). -# By specifying the "latest" tag, it will also use whatever happens to be the -# most recent version of that image when you build your Dockerfile. -# If reproducability is important, consider using a versioned tag -# (e.g., alpine:3.17.2) or SHA (e.g., alpine@sha256:c41ab5c992deb4fe7e5da09f67a8804a46bd0592bfdf0b1847dde0e0889d2bff). -FROM alpine:latest as base - -################################################################################ -# Create a stage for building/compiling the application. -# -# The following commands will leverage the "base" stage above to generate -# a "hello world" script and make it executable, but for a real application, you -# would issue a RUN command for your application's build process to generate the -# executable. For language-specific examples, take a look at the Dockerfiles in -# the Awesome Compose repository: https://github.com/docker/awesome-compose -FROM base as build -COPY < +```agda open import Cubical.Core.Everything open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels @@ -28,13 +40,21 @@ private module _ {ℓ} {A : Type ℓ} where opaque chain : ∀ {n} → Vec A (suc n) → (A → A → A) → A chain {n} (x ∷vec vec) op = foldl (λ _ → A) (λ acc curr → op acc curr) x vec +``` + +```agda record ApplicativeStructure {ℓ} (A : Type ℓ) : Type ℓ where infixl 20 _⨾_ field isSetA : isSet A _⨾_ : A → A → A +``` +Since being a set is a property - we will generally not have to think about it too much. Also, since `A` is a set - we can drop the relevance of paths and simply talk about "equality". + +We can define the notion of a term over an applicative structure. +```agda module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where open ApplicativeStructure as infix 23 `_ @@ -43,7 +63,11 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where # : Fin n → Term n `_ : A → Term n _̇_ : Term n → Term n → Term n +``` + +These terms can be evaluated into $A$ if we can give the values of the free variables. +```agda ⟦_⟧ : ∀ {n} → Term n → Vec A n → A ⟦_⟧ (` a) _ = a ⟦_⟧ {n} (# k) subs = lookup k subs @@ -54,7 +78,10 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where apply : ∀ {n} → A → Vec A n → A apply {n} a vec = chain (a ∷ vec) (λ x y → x ⨾ y) - +``` + +
+```agda private opaque unfolding reverse @@ -62,14 +89,30 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where unfolding chain applyWorks : ∀ K a b → apply K (a ∷ b ∷ []) ≡ K ⨾ a ⨾ b applyWorks K a b = refl +``` +
+ +On an applicative structure we can define Feferman structure (or SK structure). We call an applicative structure endowed with Feferman structure a **combinatory algebra**. +```agda record Feferman : Type ℓ where field s : A k : A kab≡a : ∀ a b → k ⨾ a ⨾ b ≡ a sabc≡ac_bc : ∀ a b c → s ⨾ a ⨾ b ⨾ c ≡ (a ⨾ c) ⨾ (b ⨾ c) - +``` + +Feferman structure allows us to construct **combinatorial completeness** structure. + +Imagine we have a term `t : Term n` (for some `n : ℕ`). We can ask if `A` has a "copy" of `t` so that application would correspond to subsitution. That is, we may ask if we can find an `a : A` such that +`a ⨾ < a¹ ⨾ a² ⨾ .... ⨾ aⁿ >` (here the `< ... >` notation represents a chain of applications) would be equal to `t [a¹ / # 0 , a² / # 1 , .... , aⁿ / # (pred n) ]`. If the applicative structure additionally can be endowed with Feferman structure - then the answer is yes. + +To get to such a term, we first need to define a function that takes `Term (suc n)` to `Term n` by "abstracting" the free variable represented by the index `# 0`. + +We will call this `λ*abst` and this will turn out to behave very similar to λ abstraction - and we will also show that it validates a kind of β reduction rule. + +```agda module ComputationRules (feferman : Feferman) where open Feferman feferman @@ -79,9 +122,15 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where λ*abst {n} (# (suc x)) = ` k ̇ # x λ*abst {n} (` x) = ` k ̇ ` x λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁ +``` + +**Remark** : It is important to note that in general, realizability is developed using **partial combinatory algebras** and **partial applicative structures**. In this case, `λ*abst` is not particularly well-behaved. The β reduction-esque rule we derive also does not behave *completely* like β reduction. See Jonh Longley's PhD thesis "Realizability Toposes and Language Semantics" Theorem 1.1.9. - -- Some shortcuts +**Remark** : We declare the definition as `opaque` - this is important. If we let Agda unfold this definition all the way we ocassionally end up with unreadable goals containing a mess of `s` and `k`. +We define meta-syntactic sugar for some of the more common cases : + +```agda λ* : Term one → A λ* t = ⟦ λ*abst t ⟧ [] @@ -93,7 +142,14 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where λ*4 : Term four → A λ*4 t = ⟦ λ*abst (λ*abst (λ*abst (λ*abst t))) ⟧ [] +``` + +We now show that we have a β-reduction-esque operation. We proceed by induction on the structure of the term and the number of free variables. +For the particular combinatory algebra Λ/β (terms of the untyped λ calculus quotiented by β equality) - this β reduction actually coincides with the "actual" β reduction. +TODO : Prove this. + +```agda opaque unfolding λ*abst βreduction : ∀ {n} → (body : Term (suc n)) → (prim : A) → (subs : Vec A n) → ⟦ λ*abst body ̇ ` prim ⟧ subs ≡ ⟦ body ⟧ (prim ∷ subs) @@ -113,14 +169,22 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where ≡⟨ cong₂ (λ x y → x ⨾ y) (βreduction rator prim subs) (βreduction rand prim subs) ⟩ ⟦ rator ⟧ (prim ∷ subs) ⨾ ⟦ rand ⟧ (prim ∷ subs) ∎ +``` +
+```agda λ*chainTerm : ∀ n → Term n → Term zero λ*chainTerm zero t = t λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t) λ*chain : ∀ {n} → Term n → A λ*chain {n} t = ⟦ λ*chainTerm n t ⟧ [] +``` +
+ +We provide useful reasoning combinators that are useful and frequent. +```agda opaque unfolding reverse unfolding foldl @@ -169,3 +233,4 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where ≡⟨ βreduction t d (c ∷ b ∷ a ∷ []) ⟩ ⟦ t ⟧ (d ∷ c ∷ b ∷ a ∷ []) ∎ +``` diff --git a/src/Realizability/PERs/PER.agda b/src/Realizability/PERs/PER.agda new file mode 100644 index 0000000..a686f24 --- /dev/null +++ b/src/Realizability/PERs/PER.agda @@ -0,0 +1,82 @@ +open import Realizability.ApplicativeStructure +open import Realizability.CombinatoryAlgebra +open import Realizability.PropResizing +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Binary +open import Cubical.Data.Sigma + +module Realizability.PERs.PER + {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +isPartialEquivalenceRelation : PropRel A A ℓ → Type _ +isPartialEquivalenceRelation (rel , isPropValuedRel) = BinaryRelation.isSym rel × BinaryRelation.isTrans rel + +isPropIsPartialEquivalenceRelation : ∀ r → isProp (isPartialEquivalenceRelation r) +isPropIsPartialEquivalenceRelation (rel , isPropValuedRel) = + isProp× + (isPropΠ (λ x → isPropΠ λ y → isProp→ (isPropValuedRel y x))) + (isPropΠ λ x → isPropΠ λ y → isPropΠ λ z → isProp→ (isProp→ (isPropValuedRel x z))) + +record PER : Type (ℓ-suc ℓ) where + no-eta-equality + constructor makePER + field + relation : A → A → Type ℓ + isPropValuedRelation : ∀ x y → isProp (relation x y) + isPER : isPartialEquivalenceRelation (relation , isPropValuedRelation) + isSymmetric = isPER .fst + isTransitive = isPER .snd + +open PER + +PERΣ : Type (ℓ-suc ℓ) +PERΣ = Σ[ relation ∈ (A → A → hProp ℓ) ] isPartialEquivalenceRelation ((λ a b → ⟨ relation a b ⟩) , λ a b → str (relation a b)) + +IsoPERΣPER : Iso PERΣ PER +PER.relation (Iso.fun IsoPERΣPER (relation , isPER)) x y = ⟨ relation x y ⟩ +PER.isPropValuedRelation (Iso.fun IsoPERΣPER (relation , isPER)) x y = str (relation x y) +PER.isPER (Iso.fun IsoPERΣPER (relation , isPER)) = isPER +Iso.inv IsoPERΣPER per = (λ x y → per .relation x y , per .isPropValuedRelation x y) , (isSymmetric per) , (isTransitive per) +-- refl does not work because of no-eta-equality maybe? +relation (Iso.rightInv IsoPERΣPER per i) = λ a b → per .relation a b +isPropValuedRelation (Iso.rightInv IsoPERΣPER per i) = λ a b → per .isPropValuedRelation a b +isPER (Iso.rightInv IsoPERΣPER per i) = (isSymmetric per) , (isTransitive per) +Iso.leftInv IsoPERΣPER perΣ = + Σ≡Prop + (λ x → isPropIsPartialEquivalenceRelation ((λ a b → ⟨ x a b ⟩) , (λ a b → str (x a b)))) + (funExt₂ λ a b → Σ≡Prop (λ x → isPropIsProp) refl) + +PERΣ≃PER : PERΣ ≃ PER +PERΣ≃PER = isoToEquiv IsoPERΣPER + +isSetPERΣ : isSet PERΣ +isSetPERΣ = isSetΣ (isSet→ (isSet→ isSetHProp)) (λ rel → isProp→isSet (isPropIsPartialEquivalenceRelation ((λ a b → ⟨ rel a b ⟩) , (λ a b → str (rel a b))))) + +isSetPER : isSet PER +isSetPER = isOfHLevelRespectEquiv 2 PERΣ≃PER isSetPERΣ + +module ResizedPER (resizing : hPropResizing ℓ) where + private + smallHProp = resizing .fst + hProp≃smallHProp = resizing .snd + smallHProp≃hProp = invEquiv hProp≃smallHProp + + ResizedPER : Type ℓ + ResizedPER = Σ[ relation ∈ (A → A → smallHProp) ] isPartialEquivalenceRelation ((λ a b → ⟨ equivFun (smallHProp≃hProp) (relation a b) ⟩) , λ a b → str (equivFun (smallHProp≃hProp) (relation a b))) + + PERΣ≃ResizedPER : PERΣ ≃ ResizedPER + PERΣ≃ResizedPER = + Σ-cong-equiv-prop + (equiv→ (idEquiv A) (equiv→ (idEquiv A) hProp≃smallHProp)) + (λ relation → isPropIsPartialEquivalenceRelation ((λ a b → ⟨ relation a b ⟩) , (λ a b → str (relation a b)))) + (λ resizedRelation → isPropIsPartialEquivalenceRelation ((λ a b → ⟨ equivFun (smallHProp≃hProp) (resizedRelation a b) ⟩) , λ a b → str (equivFun smallHProp≃hProp (resizedRelation a b)))) + (λ relation isPERRelation → (λ a b aRb → {!!}) , λ a b c aRb bRc → {!!}) + λ relation isPERRelation → {!!} + + PER≃ResizedPER : PER ≃ ResizedPER + PER≃ResizedPER = compEquiv (invEquiv PERΣ≃PER) PERΣ≃ResizedPER diff --git a/src/Realizability/PERs/SystemF.agda b/src/Realizability/PERs/SystemF.agda new file mode 100644 index 0000000..70a186b --- /dev/null +++ b/src/Realizability/PERs/SystemF.agda @@ -0,0 +1,40 @@ +open import Cubical.Foundations.Prelude + +module Realizability.PERs.SystemF where + +module Syntax where + -- Only one kind for now + -- System Fω has a simply typed lambda calculus at the type level + -- Inspired heavily by + -- "System F in Agda for Fun and Profit" + data Kind : Type where + tp : Kind + + data TypeCtxt : Type where + [] : TypeCtxt + _,_ : TypeCtxt → Kind → TypeCtxt + + data _∈*_ : Kind → TypeCtxt → Type where + here : ∀ {k Γ} → k ∈* (Γ , k) + there : ∀ {k Γ k'} → k ∈* Γ → k ∈* (Γ , k') + + data Tp : TypeCtxt → Kind → Type where + var : ∀ {Γ k} → k ∈* Γ → Tp Γ k + funcTp : ∀ {Γ k} → Tp Γ k → Tp Γ k → Tp Γ k + prodTp : ∀ {Γ k} → Tp Γ k → Tp Γ k → Tp Γ k + forallTp : ∀ {Γ k} → Tp (Γ , k) tp → Tp Γ tp + + data TermCtxt : TypeCtxt → Type where + [] : TermCtxt [] + _,*_ : ∀ {Γ k} → TermCtxt Γ → k ∈* Γ → TermCtxt (Γ , k) + _,_ : ∀ {Γ} → TermCtxt Γ → Tp Γ tp → TermCtxt Γ + + -- This is a better notion of renaming than as an inductive type? + Ren* : TypeCtxt → TypeCtxt → Type + Ren* Γ Δ = ∀ {K} → K ∈* Γ → K ∈* Δ + + data _∈_ : ∀ {Γ} → Tp Γ tp → TermCtxt Γ → Type where + here : ∀ {Γ} {A : Tp Γ tp} {Θ : TermCtxt Γ} → A ∈ (Θ , A) + thereType : ∀ {Γ} {A B : Tp Γ tp} {Θ : TermCtxt Γ} → A ∈ Θ → A ∈ (Θ , B) + + diff --git a/src/Realizability/Topos/Counterexample.agda b/src/Realizability/Topos/Counterexample.agda new file mode 100644 index 0000000..45a11b0 --- /dev/null +++ b/src/Realizability/Topos/Counterexample.agda @@ -0,0 +1,63 @@ +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) +open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Foundations.HLevels +open import Cubical.Data.Vec +open import Cubical.Data.Nat +open import Cubical.Data.FinData +open import Cubical.Data.Fin hiding (Fin; _/_) +open import Cubical.Data.Sigma +open import Cubical.Data.Empty +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Relation.Binary + +module Realizability.Topos.Counterexample + {ℓ} + {A : Type ℓ} + (ca : CombinatoryAlgebra A) + (isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca → ⊥) + where + +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ} {ℓ'' = ℓ} ca +open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ} {ℓ'' = ℓ} ca +open import Realizability.Topos.Object {ℓ = ℓ} {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial +open import Realizability.Topos.TerminalObject {ℓ' = ℓ} {ℓ'' = ℓ} ca isNonTrivial + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Predicate renaming (isSetX to isSetPredicateBase) +open PredicateProperties +open Morphism + +module + ChoiceCounterexample + (choice : + ∀ {X Y : Type ℓ} + (perX : PartialEquivalenceRelation X) + (perY : PartialEquivalenceRelation Y) + (f : RTMorphism perX perY) + → Σ[ F ∈ FunctionalRelation perX perY ] ([ F ] ≡ f)) where + + module _ (a : A) where + opaque + unfolding terminalPer + singleRealizerIdFuncRel : FunctionalRelation terminalPer terminalPer + Predicate.isSetX (FunctionalRelation.relation singleRealizerIdFuncRel) = isSet× isSetUnit* isSetUnit* + Predicate.∣ FunctionalRelation.relation singleRealizerIdFuncRel ∣ (tt* , tt*) r = r ≡ a + Predicate.isPropValued (FunctionalRelation.relation singleRealizerIdFuncRel) (tt* , tt*) r = isSetA _ _ + isFunctionalRelation.isStrictDomain (FunctionalRelation.isFuncRel singleRealizerIdFuncRel) = return (k , (λ { tt* tt* r r≡a → tt* })) + isFunctionalRelation.isStrictCodomain (FunctionalRelation.isFuncRel singleRealizerIdFuncRel) = return (k , (λ { tt* tt* r r≡a → tt* })) + isFunctionalRelation.isRelational (FunctionalRelation.isFuncRel singleRealizerIdFuncRel) = + let + realizer : ApplStrTerm as 3 + realizer = # one + in + return (λ*3 realizer , (λ { tt* tt* tt* tt* a b c tt* b≡a tt* → λ*3ComputationRule realizer a b c ∙ b≡a})) + isFunctionalRelation.isSingleValued (FunctionalRelation.isFuncRel singleRealizerIdFuncRel) = return (k , (λ { x y y' r₁ r₂ x₁ x₂ → tt* })) + isFunctionalRelation.isTotal (FunctionalRelation.isFuncRel singleRealizerIdFuncRel) = return (k ⨾ a , (λ { tt* r tt* → ∣ tt* , (kab≡a a r) ∣₁})) diff --git a/src/Realizability/Topos/Everything.agda b/src/Realizability/Topos/Everything.agda index bd8d4bd..7783f54 100644 --- a/src/Realizability/Topos/Everything.agda +++ b/src/Realizability/Topos/Everything.agda @@ -1,3 +1,4 @@ +{-# OPTIONS --allow-unsolved-metas #-} module Realizability.Topos.Everything where open import Realizability.Topos.Object diff --git a/src/docs/Agda.Builtin.Bool.html b/src/docs/Agda.Builtin.Bool.html new file mode 100644 index 0000000..848c014 --- /dev/null +++ b/src/docs/Agda.Builtin.Bool.html @@ -0,0 +1,15 @@ +{-# OPTIONS --cubical-compatible --safe --no-universe-polymorphism + --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Bool where + +data Bool : Set where + false true : Bool + +{-# BUILTIN BOOL Bool #-} +{-# BUILTIN FALSE false #-} +{-# BUILTIN TRUE true #-} + +{-# COMPILE JS Bool = function (x,v) { return ((x)? v["true"]() : v["false"]()); } #-} +{-# COMPILE JS false = false #-} +{-# COMPILE JS true = true #-} diff --git a/src/docs/Agda.Builtin.Char.html b/src/docs/Agda.Builtin.Char.html new file mode 100644 index 0000000..6a79684 --- /dev/null +++ b/src/docs/Agda.Builtin.Char.html @@ -0,0 +1,18 @@ +{-# OPTIONS --cubical-compatible --safe --no-universe-polymorphism + --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Char where + +open import Agda.Builtin.Nat +open import Agda.Builtin.Bool + +postulate Char : Set +{-# BUILTIN CHAR Char #-} + +primitive + primIsLower primIsDigit primIsAlpha primIsSpace primIsAscii + primIsLatin1 primIsPrint primIsHexDigit : Char Bool + primToUpper primToLower : Char Char + primCharToNat : Char Nat + primNatToChar : Nat Char + primCharEquality : Char Char Bool diff --git a/src/docs/Agda.Builtin.Cubical.Equiv.html b/src/docs/Agda.Builtin.Cubical.Equiv.html new file mode 100644 index 0000000..d589055 --- /dev/null +++ b/src/docs/Agda.Builtin.Cubical.Equiv.html @@ -0,0 +1,107 @@ +{-# OPTIONS --erased-cubical --safe --no-sized-types --no-guardedness #-} + +module Agda.Builtin.Cubical.Equiv where + +open import Agda.Primitive +open import Agda.Builtin.Sigma +open import Agda.Primitive.Cubical renaming (primINeg to ~_; primIMax to _∨_; primIMin to _∧_; + primHComp to hcomp; primTransp to transp; primComp to comp; + itIsOne to 1=1) +open import Agda.Builtin.Cubical.Path +open import Agda.Builtin.Cubical.Sub renaming (Sub to _[_↦_]) +import Agda.Builtin.Cubical.HCompU as HCompU + +module Helpers = HCompU.Helpers + +open Helpers + + +-- We make this a record so that isEquiv can be proved using +-- copatterns. This is good because copatterns don't get unfolded +-- unless a projection is applied so it should be more efficient. +record isEquiv { ℓ'} {A : Set } {B : Set ℓ'} (f : A B) : Set ( ℓ') where + no-eta-equality + field + equiv-proof : (y : B) isContr (fiber f y) + +open isEquiv public + +infix 4 _≃_ + +_≃_ : { ℓ'} (A : Set ) (B : Set ℓ') Set ( ℓ') +A B = Σ (A B) \ f (isEquiv f) + +equivFun : { ℓ'} {A : Set } {B : Set ℓ'} A B A B +equivFun e = fst e + +-- Improved version of equivProof compared to Lemma 5 in CCHM. We put +-- the (φ = i0) face in contr' making it be definitionally c in this +-- case. This makes the computational behavior better, in particular +-- for transp in Glue. +equivProof : {la lt} (T : Set la) (A : Set lt) (w : T A) (a : A) + ψ (f : Partial ψ (fiber (w .fst) a)) fiber (w .fst) a [ ψ f ] +equivProof A B w a ψ fb = + inS (contr' {A = fiber (w .fst) a} (w .snd .equiv-proof a) ψ fb) + where + contr' : {} {A : Set } isContr A (φ : I) (u : Partial φ A) A + contr' {A = A} (c , p) φ u = hcomp i λ { (φ = i1) p (u 1=1) i + ; (φ = i0) c }) c + + +{-# BUILTIN EQUIV _≃_ #-} +{-# BUILTIN EQUIVFUN equivFun #-} +{-# BUILTIN EQUIVPROOF equivProof #-} + +module _ { : I Level} (P : (i : I) Set ( i)) where + private + E : (i : I) Set ( i) + E = λ i P i + ~E : (i : I) Set ( (~ i)) + ~E = λ i P (~ i) + + A = P i0 + B = P i1 + + f : A B + f x = transp E i0 x + + g : B A + g y = transp ~E i0 y + + u : i A E i + u i x = transp j E (i j)) (~ i) x + + v : i B E i + v i y = transp j ~E ( ~ i j)) i y + + fiberPath : (y : B) (xβ0 xβ1 : fiber f y) xβ0 xβ1 + fiberPath y (x0 , β0) (x1 , β1) k = ω , λ j δ (~ j) where + module _ (j : I) where + private + sys : A i PartialP (~ j j) _ E (~ i)) + sys x i (j = i0) = v (~ i) y + sys x i (j = i1) = u (~ i) x + ω0 = comp ~E (sys x0) ((β0 (~ j))) + ω1 = comp ~E (sys x1) ((β1 (~ j))) + θ0 = fill ~E (sys x0) (inS (β0 (~ j))) + θ1 = fill ~E (sys x1) (inS (β1 (~ j))) + sys = λ {j (k = i0) ω0 j ; j (k = i1) ω1 j} + ω = hcomp sys (g y) + θ = hfill sys (inS (g y)) + δ = λ (j : I) comp E + i λ { (j = i0) v i y ; (k = i0) θ0 j (~ i) + ; (j = i1) u i ω ; (k = i1) θ1 j (~ i) }) + (θ j) + + γ : (y : B) y f (g y) + γ y j = comp E i λ { (j = i0) v i y + ; (j = i1) u i (g y) }) (g y) + + pathToisEquiv : isEquiv f + pathToisEquiv .equiv-proof y .fst .fst = g y + pathToisEquiv .equiv-proof y .fst .snd = sym (γ y) + pathToisEquiv .equiv-proof y .snd = fiberPath y _ + + pathToEquiv : A B + pathToEquiv .fst = f + pathToEquiv .snd = pathToisEquiv diff --git a/src/docs/Agda.Builtin.Cubical.Glue.html b/src/docs/Agda.Builtin.Cubical.Glue.html new file mode 100644 index 0000000..0ec02a6 --- /dev/null +++ b/src/docs/Agda.Builtin.Cubical.Glue.html @@ -0,0 +1,18 @@ +{-# OPTIONS --cubical --safe --no-sized-types --no-guardedness #-} + +module Agda.Builtin.Cubical.Glue where + +open import Agda.Primitive +open import Agda.Primitive.Cubical +open import Agda.Builtin.Cubical.Equiv public + +primitive + primGlue : { ℓ'} (A : Set ) {φ : I} + (T : Partial φ (Set ℓ')) (e : PartialP φ o T o A)) + Set ℓ' + prim^glue : { ℓ'} {A : Set } {φ : I} + {T : Partial φ (Set ℓ')} {e : PartialP φ o T o A)} + (t : PartialP φ T) (a : A) primGlue A T e + prim^unglue : { ℓ'} {A : Set } {φ : I} + {T : Partial φ (Set ℓ')} {e : PartialP φ o T o A)} + primGlue A T e A diff --git a/src/docs/Agda.Builtin.Cubical.HCompU.html b/src/docs/Agda.Builtin.Cubical.HCompU.html new file mode 100644 index 0000000..2cb5781 --- /dev/null +++ b/src/docs/Agda.Builtin.Cubical.HCompU.html @@ -0,0 +1,78 @@ +{-# OPTIONS --erased-cubical --safe --no-sized-types --no-guardedness #-} + +module Agda.Builtin.Cubical.HCompU where + +open import Agda.Primitive +open import Agda.Builtin.Sigma +open import Agda.Primitive.Cubical renaming (primINeg to ~_; primIMax to _∨_; primIMin to _∧_; + primHComp to hcomp; primTransp to transp; primComp to comp; + itIsOne to 1=1) +open import Agda.Builtin.Cubical.Path +open import Agda.Builtin.Cubical.Sub renaming (Sub to _[_↦_]; primSubOut to outS) + +module Helpers where + -- Homogeneous filling + hfill : {} {A : Set } {φ : I} + (u : i Partial φ A) + (u0 : A [ φ u i0 ]) (i : I) A + hfill {φ = φ} u u0 i = + hcomp j \ { (φ = i1) u (i j) 1=1 + ; (i = i0) outS u0 }) + (outS u0) + + -- Heterogeneous filling defined using comp + fill : { : I Level} (A : i Set ( i)) {φ : I} + (u : i Partial φ (A i)) + (u0 : A i0 [ φ u i0 ]) + i A i + fill A {φ = φ} u u0 i = + comp j A (i j)) + j \ { (φ = i1) u (i j) 1=1 + ; (i = i0) outS u0 }) + (outS {φ = φ} u0) + + module _ {} {A : Set } where + refl : {x : A} x x + refl {x = x} = λ _ x + + sym : {x y : A} x y y x + sym p = λ i p (~ i) + + cong : {ℓ'} {B : A Set ℓ'} {x y : A} + (f : (a : A) B a) (p : x y) + PathP i B (p i)) (f x) (f y) + cong f p = λ i f (p i) + + isContr : {} Set Set + isContr A = Σ A \ x (∀ y x y) + + fiber : { ℓ'} {A : Set } {B : Set ℓ'} (f : A B) (y : B) Set ( ℓ') + fiber {A = A} f y = Σ A \ x f x y + +open Helpers + + +primitive + prim^glueU : {la : Level} {φ : I} {T : I Partial φ (Set la)} + {A : Set la [ φ T i0 ]} + PartialP φ (T i1) outS A hcomp T (outS A) + prim^unglueU : {la : Level} {φ : I} {T : I Partial φ (Set la)} + {A : Set la [ φ T i0 ]} + hcomp T (outS A) outS A + -- Needed for transp. + primFaceForall : (I I) I + +transpProof : {l} (e : I Set l) (φ : I) (a : Partial φ (e i0)) (b : e i1 [ φ (\ o transp (\ i e i) i0 (a o)) ] ) fiber (transp (\ i e i) i0) (outS b) +transpProof e φ a b = f , \ j comp (\ i e i) (\ i + \ { (φ = i1) transp (\ j e (j i)) (~ i) (a 1=1) + ; (j = i0) transp (\ j e (j i)) (~ i) f + ; (j = i1) g (~ i) }) + f + where + b' = outS {u = (\ o transp (\ i e i) i0 (a o))} b + g : (k : I) e (~ k) + g k = fill (\ i e (~ i)) (\ i \ { (φ = i1) transp (\ j e (j ~ i)) i (a 1=1) + ; (φ = i0) transp (\ j e (~ j ~ i)) (~ i) b' }) (inS b') k + f = comp (\ i e (~ i)) (\ i \ { (φ = i1) transp (\ j e (j ~ i)) i (a 1=1); (φ = i0) transp (\ j e (~ j ~ i)) (~ i) b' }) b' + +{-# BUILTIN TRANSPPROOF transpProof #-} diff --git a/src/docs/Agda.Builtin.Cubical.Path.html b/src/docs/Agda.Builtin.Cubical.Path.html new file mode 100644 index 0000000..c795c3d --- /dev/null +++ b/src/docs/Agda.Builtin.Cubical.Path.html @@ -0,0 +1,15 @@ +{-# OPTIONS --erased-cubical --safe --no-sized-types --no-guardedness #-} + +module Agda.Builtin.Cubical.Path where + + open import Agda.Primitive.Cubical using (PathP) public + + + infix 4 _≡_ + + -- We have a variable name in `(λ i → A)` as a hint for case + -- splitting. + _≡_ : {} {A : Set } A A Set + _≡_ {A = A} = PathP i A) + + {-# BUILTIN PATH _≡_ #-} diff --git a/src/docs/Agda.Builtin.Cubical.Sub.html b/src/docs/Agda.Builtin.Cubical.Sub.html new file mode 100644 index 0000000..5e0e5e1 --- /dev/null +++ b/src/docs/Agda.Builtin.Cubical.Sub.html @@ -0,0 +1,18 @@ +{-# OPTIONS --erased-cubical --safe --no-sized-types --no-guardedness #-} + +module Agda.Builtin.Cubical.Sub where + + open import Agda.Primitive.Cubical + + {-# BUILTIN SUB Sub #-} + + postulate + inS : {} {A : Set } {φ} (x : A) Sub A φ _ x) + + {-# BUILTIN SUBIN inS #-} + + -- Sub A φ u is treated as A. + {-# COMPILE JS inS = _ => _ => _ => x => x #-} + + primitive + primSubOut : {} {A : Set } {φ : I} {u : Partial φ A} Sub _ φ u A diff --git a/src/docs/Agda.Builtin.Float.html b/src/docs/Agda.Builtin.Float.html new file mode 100644 index 0000000..80e0afb --- /dev/null +++ b/src/docs/Agda.Builtin.Float.html @@ -0,0 +1,209 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Float where + +open import Agda.Builtin.Bool +open import Agda.Builtin.Int +open import Agda.Builtin.Maybe +open import Agda.Builtin.Nat +open import Agda.Builtin.Sigma +open import Agda.Builtin.String +open import Agda.Builtin.Word + +postulate Float : Set +{-# BUILTIN FLOAT Float #-} + +primitive + -- Relations + primFloatInequality : Float Float Bool + primFloatEquality : Float Float Bool + primFloatLess : Float Float Bool + primFloatIsInfinite : Float Bool + primFloatIsNaN : Float Bool + primFloatIsNegativeZero : Float Bool + primFloatIsSafeInteger : Float Bool + -- Conversions + primFloatToWord64 : Float Maybe Word64 + primNatToFloat : Nat Float + primIntToFloat : Int Float + primFloatRound : Float Maybe Int + primFloatFloor : Float Maybe Int + primFloatCeiling : Float Maybe Int + primFloatToRatio : Float (Σ Int λ _ Int) + primRatioToFloat : Int Int Float + primFloatDecode : Float Maybe (Σ Int λ _ Int) + primFloatEncode : Int Int Maybe Float + primShowFloat : Float String + -- Operations + primFloatPlus : Float Float Float + primFloatMinus : Float Float Float + primFloatTimes : Float Float Float + primFloatDiv : Float Float Float + primFloatPow : Float Float Float + primFloatNegate : Float Float + primFloatSqrt : Float Float + primFloatExp : Float Float + primFloatLog : Float Float + primFloatSin : Float Float + primFloatCos : Float Float + primFloatTan : Float Float + primFloatASin : Float Float + primFloatACos : Float Float + primFloatATan : Float Float + primFloatATan2 : Float Float Float + primFloatSinh : Float Float + primFloatCosh : Float Float + primFloatTanh : Float Float + primFloatASinh : Float Float + primFloatACosh : Float Float + primFloatATanh : Float Float + +{-# COMPILE JS + primFloatRound = function(x) { + x = agdaRTS._primFloatRound(x); + if (x === null) { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["nothing"]; + } + else { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["just"](x); + } + }; +#-} +{-# COMPILE JS + primFloatFloor = function(x) { + x = agdaRTS._primFloatFloor(x); + if (x === null) { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["nothing"]; + } + else { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["just"](x); + } + }; +#-} +{-# COMPILE JS + primFloatCeiling = function(x) { + x = agdaRTS._primFloatCeiling(x); + if (x === null) { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["nothing"]; + } + else { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["just"](x); + } + }; +#-} +{-# COMPILE JS + primFloatToRatio = function(x) { + x = agdaRTS._primFloatToRatio(x); + return z_jAgda_Agda_Builtin_Sigma["_,_"](x.numerator)(x.denominator); + }; +#-} +{-# COMPILE JS + primFloatDecode = function(x) { + x = agdaRTS._primFloatDecode(x); + if (x === null) { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["nothing"]; + } + else { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["just"]( + z_jAgda_Agda_Builtin_Sigma["_,_"](x.mantissa)(x.exponent)); + } + }; +#-} +{-# COMPILE JS + primFloatEncode = function(x) { + return function (y) { + x = agdaRTS.uprimFloatEncode(x, y); + if (x === null) { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["nothing"]; + } + else { + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["just"](x); + } + } + }; +#-} + +primFloatNumericalEquality = primFloatEquality +{-# WARNING_ON_USAGE primFloatNumericalEquality +"Warning: primFloatNumericalEquality was deprecated in Agda v2.6.2. +Please use primFloatEquality instead." +#-} + +primFloatNumericalLess = primFloatLess +{-# WARNING_ON_USAGE primFloatNumericalLess +"Warning: primFloatNumericalLess was deprecated in Agda v2.6.2. +Please use primFloatLess instead." +#-} + +primRound = primFloatRound +{-# WARNING_ON_USAGE primRound +"Warning: primRound was deprecated in Agda v2.6.2. +Please use primFloatRound instead." +#-} + +primFloor = primFloatFloor +{-# WARNING_ON_USAGE primFloor +"Warning: primFloor was deprecated in Agda v2.6.2. +Please use primFloatFloor instead." +#-} + +primCeiling = primFloatCeiling +{-# WARNING_ON_USAGE primCeiling +"Warning: primCeiling was deprecated in Agda v2.6.2. +Please use primFloatCeiling instead." +#-} + +primExp = primFloatExp +{-# WARNING_ON_USAGE primExp +"Warning: primExp was deprecated in Agda v2.6.2. +Please use primFloatExp instead." +#-} + +primLog = primFloatLog +{-# WARNING_ON_USAGE primLog +"Warning: primLog was deprecated in Agda v2.6.2. +Please use primFloatLog instead." +#-} + +primSin = primFloatSin +{-# WARNING_ON_USAGE primSin +"Warning: primSin was deprecated in Agda v2.6.2. +Please use primFloatSin instead." +#-} + +primCos = primFloatCos +{-# WARNING_ON_USAGE primCos +"Warning: primCos was deprecated in Agda v2.6.2. +Please use primFloatCos instead." +#-} + +primTan = primFloatTan +{-# WARNING_ON_USAGE primTan +"Warning: primTan was deprecated in Agda v2.6.2. +Please use primFloatTan instead." +#-} + +primASin = primFloatASin +{-# WARNING_ON_USAGE primASin +"Warning: primASin was deprecated in Agda v2.6.2. +Please use primFloatASin instead." +#-} + + +primACos = primFloatACos +{-# WARNING_ON_USAGE primACos +"Warning: primACos was deprecated in Agda v2.6.2. +Please use primFloatACos instead." +#-} + +primATan = primFloatATan +{-# WARNING_ON_USAGE primATan +"Warning: primATan was deprecated in Agda v2.6.2. +Please use primFloatATan instead." +#-} + +primATan2 = primFloatATan2 +{-# WARNING_ON_USAGE primATan2 +"Warning: primATan2 was deprecated in Agda v2.6.2. +Please use primFloatATan2 instead." +#-} diff --git a/src/docs/Agda.Builtin.FromNat.html b/src/docs/Agda.Builtin.FromNat.html new file mode 100644 index 0000000..080b2ab --- /dev/null +++ b/src/docs/Agda.Builtin.FromNat.html @@ -0,0 +1,16 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.FromNat where + +open import Agda.Primitive +open import Agda.Builtin.Nat + +record Number {a} (A : Set a) : Set (lsuc a) where + field + Constraint : Nat Set a + fromNat : n {{_ : Constraint n}} A + +open Number {{...}} public using (fromNat) + +{-# BUILTIN FROMNAT fromNat #-} +{-# DISPLAY Number.fromNat _ n = fromNat n #-} diff --git a/src/docs/Agda.Builtin.FromNeg.html b/src/docs/Agda.Builtin.FromNeg.html new file mode 100644 index 0000000..db2ba5a --- /dev/null +++ b/src/docs/Agda.Builtin.FromNeg.html @@ -0,0 +1,16 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.FromNeg where + +open import Agda.Primitive +open import Agda.Builtin.Nat + +record Negative {a} (A : Set a) : Set (lsuc a) where + field + Constraint : Nat Set a + fromNeg : n {{_ : Constraint n}} A + +open Negative {{...}} public using (fromNeg) + +{-# BUILTIN FROMNEG fromNeg #-} +{-# DISPLAY Negative.fromNeg _ n = fromNeg n #-} diff --git a/src/docs/Agda.Builtin.Int.html b/src/docs/Agda.Builtin.Int.html new file mode 100644 index 0000000..29018d7 --- /dev/null +++ b/src/docs/Agda.Builtin.Int.html @@ -0,0 +1,18 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Int where + +open import Agda.Builtin.Nat +open import Agda.Builtin.String + +infix 8 pos -- Standard library uses this as +_ + +data Int : Set where + pos : (n : Nat) Int + negsuc : (n : Nat) Int + +{-# BUILTIN INTEGER Int #-} +{-# BUILTIN INTEGERPOS pos #-} +{-# BUILTIN INTEGERNEGSUC negsuc #-} + +primitive primShowInteger : Int String diff --git a/src/docs/Agda.Builtin.List.html b/src/docs/Agda.Builtin.List.html new file mode 100644 index 0000000..b32b1cd --- /dev/null +++ b/src/docs/Agda.Builtin.List.html @@ -0,0 +1,16 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.List where + +infixr 5 _∷_ +data List {a} (A : Set a) : Set a where + [] : List A + _∷_ : (x : A) (xs : List A) List A + +{-# BUILTIN LIST List #-} + +{-# COMPILE JS List = function(x,v) { + if (x.length < 1) { return v["[]"](); } else { return v["_∷_"](x[0], x.slice(1)); } +} #-} +{-# COMPILE JS [] = Array() #-} +{-# COMPILE JS _∷_ = function (x) { return function(y) { return Array(x).concat(y); }; } #-} diff --git a/src/docs/Agda.Builtin.Maybe.html b/src/docs/Agda.Builtin.Maybe.html new file mode 100644 index 0000000..1894080 --- /dev/null +++ b/src/docs/Agda.Builtin.Maybe.html @@ -0,0 +1,9 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Maybe where + +data Maybe {a} (A : Set a) : Set a where + just : A Maybe A + nothing : Maybe A + +{-# BUILTIN MAYBE Maybe #-} diff --git a/src/docs/Agda.Builtin.Nat.html b/src/docs/Agda.Builtin.Nat.html new file mode 100644 index 0000000..4b1bb67 --- /dev/null +++ b/src/docs/Agda.Builtin.Nat.html @@ -0,0 +1,134 @@ +{-# OPTIONS --cubical-compatible --safe --no-universe-polymorphism + --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Nat where + +open import Agda.Builtin.Bool + +data Nat : Set where + zero : Nat + suc : (n : Nat) Nat + +{-# BUILTIN NATURAL Nat #-} + +infix 4 _==_ _<_ +infixl 6 _+_ _-_ +infixl 7 _*_ + +_+_ : Nat Nat Nat +zero + m = m +suc n + m = suc (n + m) + +{-# BUILTIN NATPLUS _+_ #-} + +_-_ : Nat Nat Nat +n - zero = n +zero - suc m = zero +suc n - suc m = n - m + +{-# BUILTIN NATMINUS _-_ #-} + +_*_ : Nat Nat Nat +zero * m = zero +suc n * m = m + n * m + +{-# BUILTIN NATTIMES _*_ #-} + +_==_ : Nat Nat Bool +zero == zero = true +suc n == suc m = n == m +_ == _ = false + +{-# BUILTIN NATEQUALS _==_ #-} + +_<_ : Nat Nat Bool +_ < zero = false +zero < suc _ = true +suc n < suc m = n < m + +{-# BUILTIN NATLESS _<_ #-} + +-- Helper function div-helper for Euclidean division. +--------------------------------------------------------------------------- +-- +-- div-helper computes n / 1+m via iteration on n. +-- +-- n div (suc m) = div-helper 0 m n m +-- +-- The state of the iterator has two accumulator variables: +-- +-- k: The quotient, returned once n=0. Initialized to 0. +-- +-- j: A counter, initialized to the divisor m, decreased on each iteration step. +-- Once it reaches 0, the quotient k is increased and j reset to m, +-- starting the next countdown. +-- +-- Under the precondition j ≤ m, the invariant is +-- +-- div-helper k m n j = k + (n + m - j) div (1 + m) + +div-helper : (k m n j : Nat) Nat +div-helper k m zero j = k +div-helper k m (suc n) zero = div-helper (suc k) m n m +div-helper k m (suc n) (suc j) = div-helper k m n j + +{-# BUILTIN NATDIVSUCAUX div-helper #-} + +-- Proof of the invariant by induction on n. +-- +-- clause 1: div-helper k m 0 j +-- = k by definition +-- = k + (0 + m - j) div (1 + m) since m - j < 1 + m +-- +-- clause 2: div-helper k m (1 + n) 0 +-- = div-helper (1 + k) m n m by definition +-- = 1 + k + (n + m - m) div (1 + m) by induction hypothesis +-- = 1 + k + n div (1 + m) by simplification +-- = k + (n + (1 + m)) div (1 + m) by expansion +-- = k + (1 + n + m - 0) div (1 + m) by expansion +-- +-- clause 3: div-helper k m (1 + n) (1 + j) +-- = div-helper k m n j by definition +-- = k + (n + m - j) div (1 + m) by induction hypothesis +-- = k + ((1 + n) + m - (1 + j)) div (1 + m) by expansion +-- +-- Q.e.d. + +-- Helper function mod-helper for the remainder computation. +--------------------------------------------------------------------------- +-- +-- (Analogous to div-helper.) +-- +-- mod-helper computes n % 1+m via iteration on n. +-- +-- n mod (suc m) = mod-helper 0 m n m +-- +-- The invariant is: +-- +-- m = k + j ==> mod-helper k m n j = (n + k) mod (1 + m). + +mod-helper : (k m n j : Nat) Nat +mod-helper k m zero j = k +mod-helper k m (suc n) zero = mod-helper 0 m n m +mod-helper k m (suc n) (suc j) = mod-helper (suc k) m n j + +{-# BUILTIN NATMODSUCAUX mod-helper #-} + +-- Proof of the invariant by induction on n. +-- +-- clause 1: mod-helper k m 0 j +-- = k by definition +-- = (0 + k) mod (1 + m) since m = k + j, thus k < m +-- +-- clause 2: mod-helper k m (1 + n) 0 +-- = mod-helper 0 m n m by definition +-- = (n + 0) mod (1 + m) by induction hypothesis +-- = (n + (1 + m)) mod (1 + m) by expansion +-- = (1 + n) + k) mod (1 + m) since k = m (as l = 0) +-- +-- clause 3: mod-helper k m (1 + n) (1 + j) +-- = mod-helper (1 + k) m n j by definition +-- = (n + (1 + k)) mod (1 + m) by induction hypothesis +-- = ((1 + n) + k) mod (1 + m) by commutativity +-- +-- Q.e.d. diff --git a/src/docs/Agda.Builtin.Reflection.html b/src/docs/Agda.Builtin.Reflection.html new file mode 100644 index 0000000..d38270b --- /dev/null +++ b/src/docs/Agda.Builtin.Reflection.html @@ -0,0 +1,470 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Reflection where + +open import Agda.Builtin.Unit +open import Agda.Builtin.Bool +open import Agda.Builtin.Nat +open import Agda.Builtin.Word +open import Agda.Builtin.List +open import Agda.Builtin.String +open import Agda.Builtin.Char +open import Agda.Builtin.Float +open import Agda.Builtin.Int +open import Agda.Builtin.Sigma +open import Agda.Primitive + +-- Names -- + +postulate Name : Set +{-# BUILTIN QNAME Name #-} + +primitive + primQNameEquality : Name Name Bool + primQNameLess : Name Name Bool + primShowQName : Name String + +-- Fixity -- + +data Associativity : Set where + left-assoc : Associativity + right-assoc : Associativity + non-assoc : Associativity + +data Precedence : Set where + related : Float Precedence + unrelated : Precedence + +data Fixity : Set where + fixity : Associativity Precedence Fixity + +{-# BUILTIN ASSOC Associativity #-} +{-# BUILTIN ASSOCLEFT left-assoc #-} +{-# BUILTIN ASSOCRIGHT right-assoc #-} +{-# BUILTIN ASSOCNON non-assoc #-} + +{-# BUILTIN PRECEDENCE Precedence #-} +{-# BUILTIN PRECRELATED related #-} +{-# BUILTIN PRECUNRELATED unrelated #-} + +{-# BUILTIN FIXITY Fixity #-} +{-# BUILTIN FIXITYFIXITY fixity #-} + +{-# COMPILE GHC Associativity = data MAlonzo.RTE.Assoc (MAlonzo.RTE.LeftAssoc | MAlonzo.RTE.RightAssoc | MAlonzo.RTE.NonAssoc) #-} +{-# COMPILE GHC Precedence = data MAlonzo.RTE.Precedence (MAlonzo.RTE.Related | MAlonzo.RTE.Unrelated) #-} +{-# COMPILE GHC Fixity = data MAlonzo.RTE.Fixity (MAlonzo.RTE.Fixity) #-} + +{-# COMPILE JS Associativity = function (x,v) { return v[x](); } #-} +{-# COMPILE JS left-assoc = "left-assoc" #-} +{-# COMPILE JS right-assoc = "right-assoc" #-} +{-# COMPILE JS non-assoc = "non-assoc" #-} + +{-# COMPILE JS Precedence = + function (x,v) { + if (x === "unrelated") { return v[x](); } else { return v["related"](x); }} #-} +{-# COMPILE JS related = function(x) { return x; } #-} +{-# COMPILE JS unrelated = "unrelated" #-} + +{-# COMPILE JS Fixity = function (x,v) { return v["fixity"](x["assoc"], x["prec"]); } #-} +{-# COMPILE JS fixity = function (x) { return function (y) { return { "assoc": x, "prec": y}; }; } #-} + +primitive + primQNameFixity : Name Fixity + primQNameToWord64s : Name Σ Word64 _ Word64) + +-- Metavariables -- + +postulate Meta : Set +{-# BUILTIN AGDAMETA Meta #-} + +primitive + primMetaEquality : Meta Meta Bool + primMetaLess : Meta Meta Bool + primShowMeta : Meta String + primMetaToNat : Meta Nat + +-- Arguments -- + +-- Arguments can be (visible), {hidden}, or {{instance}}. +data Visibility : Set where + visible hidden instance′ : Visibility + +{-# BUILTIN HIDING Visibility #-} +{-# BUILTIN VISIBLE visible #-} +{-# BUILTIN HIDDEN hidden #-} +{-# BUILTIN INSTANCE instance′ #-} + +-- Arguments can be relevant or irrelevant. +data Relevance : Set where + relevant irrelevant : Relevance + +{-# BUILTIN RELEVANCE Relevance #-} +{-# BUILTIN RELEVANT relevant #-} +{-# BUILTIN IRRELEVANT irrelevant #-} + +-- Arguments also have a quantity. +data Quantity : Set where + quantity-0 quantity-ω : Quantity + +{-# BUILTIN QUANTITY Quantity #-} +{-# BUILTIN QUANTITY-0 quantity-0 #-} +{-# BUILTIN QUANTITY-ω quantity-ω #-} + +-- Relevance and quantity are combined into a modality. +data Modality : Set where + modality : (r : Relevance) (q : Quantity) Modality + +{-# BUILTIN MODALITY Modality #-} +{-# BUILTIN MODALITY-CONSTRUCTOR modality #-} + +data ArgInfo : Set where + arg-info : (v : Visibility) (m : Modality) ArgInfo + +data Arg {a} (A : Set a) : Set a where + arg : (i : ArgInfo) (x : A) Arg A + +{-# BUILTIN ARGINFO ArgInfo #-} +{-# BUILTIN ARGARGINFO arg-info #-} +{-# BUILTIN ARG Arg #-} +{-# BUILTIN ARGARG arg #-} + +data Blocker : Set where + blockerAny : List Blocker Blocker + blockerAll : List Blocker Blocker + blockerMeta : Meta Blocker + +{-# BUILTIN AGDABLOCKER Blocker #-} +{-# BUILTIN AGDABLOCKERANY blockerAny #-} +{-# BUILTIN AGDABLOCKERALL blockerAll #-} +{-# BUILTIN AGDABLOCKERMETA blockerMeta #-} + +-- Name abstraction -- + +data Abs {a} (A : Set a) : Set a where + abs : (s : String) (x : A) Abs A + +{-# BUILTIN ABS Abs #-} +{-# BUILTIN ABSABS abs #-} + +-- Literals -- + +data Literal : Set where + nat : (n : Nat) Literal + word64 : (n : Word64) Literal + float : (x : Float) Literal + char : (c : Char) Literal + string : (s : String) Literal + name : (x : Name) Literal + meta : (x : Meta) Literal + +{-# BUILTIN AGDALITERAL Literal #-} +{-# BUILTIN AGDALITNAT nat #-} +{-# BUILTIN AGDALITWORD64 word64 #-} +{-# BUILTIN AGDALITFLOAT float #-} +{-# BUILTIN AGDALITCHAR char #-} +{-# BUILTIN AGDALITSTRING string #-} +{-# BUILTIN AGDALITQNAME name #-} +{-# BUILTIN AGDALITMETA meta #-} + + +-- Terms and patterns -- + +data Term : Set +data Sort : Set +data Pattern : Set +data Clause : Set +Type = Term +Telescope = List (Σ String λ _ Arg Type) + +data Term where + var : (x : Nat) (args : List (Arg Term)) Term + con : (c : Name) (args : List (Arg Term)) Term + def : (f : Name) (args : List (Arg Term)) Term + lam : (v : Visibility) (t : Abs Term) Term + pat-lam : (cs : List Clause) (args : List (Arg Term)) Term + pi : (a : Arg Type) (b : Abs Type) Term + agda-sort : (s : Sort) Term + lit : (l : Literal) Term + meta : (x : Meta) List (Arg Term) Term + unknown : Term + +data Sort where + set : (t : Term) Sort + lit : (n : Nat) Sort + prop : (t : Term) Sort + propLit : (n : Nat) Sort + inf : (n : Nat) Sort + unknown : Sort + +data Pattern where + con : (c : Name) (ps : List (Arg Pattern)) Pattern + dot : (t : Term) Pattern + var : (x : Nat) Pattern + lit : (l : Literal) Pattern + proj : (f : Name) Pattern + absurd : (x : Nat) Pattern -- absurd patterns counts as variables + +data Clause where + clause : (tel : Telescope) (ps : List (Arg Pattern)) (t : Term) Clause + absurd-clause : (tel : Telescope) (ps : List (Arg Pattern)) Clause + +{-# BUILTIN AGDATERM Term #-} +{-# BUILTIN AGDASORT Sort #-} +{-# BUILTIN AGDAPATTERN Pattern #-} +{-# BUILTIN AGDACLAUSE Clause #-} + +{-# BUILTIN AGDATERMVAR var #-} +{-# BUILTIN AGDATERMCON con #-} +{-# BUILTIN AGDATERMDEF def #-} +{-# BUILTIN AGDATERMMETA meta #-} +{-# BUILTIN AGDATERMLAM lam #-} +{-# BUILTIN AGDATERMEXTLAM pat-lam #-} +{-# BUILTIN AGDATERMPI pi #-} +{-# BUILTIN AGDATERMSORT agda-sort #-} +{-# BUILTIN AGDATERMLIT lit #-} +{-# BUILTIN AGDATERMUNSUPPORTED unknown #-} + +{-# BUILTIN AGDASORTSET set #-} +{-# BUILTIN AGDASORTLIT lit #-} +{-# BUILTIN AGDASORTPROP prop #-} +{-# BUILTIN AGDASORTPROPLIT propLit #-} +{-# BUILTIN AGDASORTINF inf #-} +{-# BUILTIN AGDASORTUNSUPPORTED unknown #-} + +{-# BUILTIN AGDAPATCON con #-} +{-# BUILTIN AGDAPATDOT dot #-} +{-# BUILTIN AGDAPATVAR var #-} +{-# BUILTIN AGDAPATLIT lit #-} +{-# BUILTIN AGDAPATPROJ proj #-} +{-# BUILTIN AGDAPATABSURD absurd #-} + +{-# BUILTIN AGDACLAUSECLAUSE clause #-} +{-# BUILTIN AGDACLAUSEABSURD absurd-clause #-} + +-- Definitions -- + +data Definition : Set where + function : (cs : List Clause) Definition + data-type : (pars : Nat) (cs : List Name) Definition + record-type : (c : Name) (fs : List (Arg Name)) Definition + data-cons : (d : Name) Definition + axiom : Definition + prim-fun : Definition + +{-# BUILTIN AGDADEFINITION Definition #-} +{-# BUILTIN AGDADEFINITIONFUNDEF function #-} +{-# BUILTIN AGDADEFINITIONDATADEF data-type #-} +{-# BUILTIN AGDADEFINITIONRECORDDEF record-type #-} +{-# BUILTIN AGDADEFINITIONDATACONSTRUCTOR data-cons #-} +{-# BUILTIN AGDADEFINITIONPOSTULATE axiom #-} +{-# BUILTIN AGDADEFINITIONPRIMITIVE prim-fun #-} + +-- Errors -- + +data ErrorPart : Set where + strErr : String ErrorPart + termErr : Term ErrorPart + pattErr : Pattern ErrorPart + nameErr : Name ErrorPart + +{-# BUILTIN AGDAERRORPART ErrorPart #-} +{-# BUILTIN AGDAERRORPARTSTRING strErr #-} +{-# BUILTIN AGDAERRORPARTTERM termErr #-} +{-# BUILTIN AGDAERRORPARTPATT pattErr #-} +{-# BUILTIN AGDAERRORPARTNAME nameErr #-} + +-- TC monad -- + +postulate + TC : {a} Set a Set a + returnTC : {a} {A : Set a} A TC A + bindTC : {a b} {A : Set a} {B : Set b} TC A (A TC B) TC B + unify : Term Term TC + typeError : {a} {A : Set a} List ErrorPart TC A + inferType : Term TC Type + checkType : Term Type TC Term + normalise : Term TC Term + reduce : Term TC Term + catchTC : {a} {A : Set a} TC A TC A TC A + quoteTC : {a} {A : Set a} A TC Term + unquoteTC : {a} {A : Set a} Term TC A + quoteωTC : {A : Setω} A TC Term + getContext : TC Telescope + extendContext : {a} {A : Set a} String Arg Type TC A TC A + inContext : {a} {A : Set a} Telescope TC A TC A + freshName : String TC Name + declareDef : Arg Name Type TC + declarePostulate : Arg Name Type TC + declareData : Name Nat Type TC + defineData : Name List (Σ Name _ Type)) TC + defineFun : Name List Clause TC + getType : Name TC Type + getDefinition : Name TC Definition + blockTC : {a} {A : Set a} Blocker TC A + commitTC : TC + isMacro : Name TC Bool + pragmaForeign : String String TC + pragmaCompile : String Name String TC + + -- If 'true', makes the following primitives also normalise + -- their results: inferType, checkType, quoteTC, getType, and getContext + withNormalisation : {a} {A : Set a} Bool TC A TC A + askNormalisation : TC Bool + + -- If 'true', makes the following primitives to reconstruct hidden arguments: + -- getDefinition, normalise, reduce, inferType, checkType and getContext + withReconstructed : {a} {A : Set a} Bool TC A TC A + askReconstructed : TC Bool + + -- Whether implicit arguments at the end should be turned into metavariables + withExpandLast : {a} {A : Set a} Bool TC A TC A + askExpandLast : TC Bool + + -- White/blacklist specific definitions for reduction while executing the TC computation + -- 'true' for whitelist, 'false' for blacklist + withReduceDefs : {a} {A : Set a} (Σ Bool λ _ List Name) TC A TC A + askReduceDefs : TC (Σ Bool λ _ List Name) + + formatErrorParts : List ErrorPart TC String + -- Prints the third argument if the corresponding verbosity level is turned + -- on (with the -v flag to Agda). + debugPrint : String Nat List ErrorPart TC + + -- Fail if the given computation gives rise to new, unsolved + -- "blocking" constraints. + noConstraints : {a} {A : Set a} TC A TC A + + -- Run the given TC action and return the first component. Resets to + -- the old TC state if the second component is 'false', or keep the + -- new TC state if it is 'true'. + runSpeculative : {a} {A : Set a} TC (Σ A λ _ Bool) TC A + + -- Get a list of all possible instance candidates for the given meta + -- variable (it does not have to be an instance meta). + getInstances : Meta TC (List Term) + +{-# BUILTIN AGDATCM TC #-} +{-# BUILTIN AGDATCMRETURN returnTC #-} +{-# BUILTIN AGDATCMBIND bindTC #-} +{-# BUILTIN AGDATCMUNIFY unify #-} +{-# BUILTIN AGDATCMTYPEERROR typeError #-} +{-# BUILTIN AGDATCMINFERTYPE inferType #-} +{-# BUILTIN AGDATCMCHECKTYPE checkType #-} +{-# BUILTIN AGDATCMNORMALISE normalise #-} +{-# BUILTIN AGDATCMREDUCE reduce #-} +{-# BUILTIN AGDATCMCATCHERROR catchTC #-} +{-# BUILTIN AGDATCMQUOTETERM quoteTC #-} +{-# BUILTIN AGDATCMUNQUOTETERM unquoteTC #-} +{-# BUILTIN AGDATCMQUOTEOMEGATERM quoteωTC #-} +{-# BUILTIN AGDATCMGETCONTEXT getContext #-} +{-# BUILTIN AGDATCMEXTENDCONTEXT extendContext #-} +{-# BUILTIN AGDATCMINCONTEXT inContext #-} +{-# BUILTIN AGDATCMFRESHNAME freshName #-} +{-# BUILTIN AGDATCMDECLAREDEF declareDef #-} +{-# BUILTIN AGDATCMDECLAREPOSTULATE declarePostulate #-} +{-# BUILTIN AGDATCMDECLAREDATA declareData #-} +{-# BUILTIN AGDATCMDEFINEDATA defineData #-} +{-# BUILTIN AGDATCMDEFINEFUN defineFun #-} +{-# BUILTIN AGDATCMGETTYPE getType #-} +{-# BUILTIN AGDATCMGETDEFINITION getDefinition #-} +{-# BUILTIN AGDATCMBLOCK blockTC #-} +{-# BUILTIN AGDATCMCOMMIT commitTC #-} +{-# BUILTIN AGDATCMISMACRO isMacro #-} +{-# BUILTIN AGDATCMPRAGMAFOREIGN pragmaForeign #-} +{-# BUILTIN AGDATCMPRAGMACOMPILE pragmaCompile #-} +{-# BUILTIN AGDATCMWITHNORMALISATION withNormalisation #-} +{-# BUILTIN AGDATCMWITHRECONSTRUCTED withReconstructed #-} +{-# BUILTIN AGDATCMWITHEXPANDLAST withExpandLast #-} +{-# BUILTIN AGDATCMWITHREDUCEDEFS withReduceDefs #-} +{-# BUILTIN AGDATCMASKNORMALISATION askNormalisation #-} +{-# BUILTIN AGDATCMASKRECONSTRUCTED askReconstructed #-} +{-# BUILTIN AGDATCMASKEXPANDLAST askExpandLast #-} +{-# BUILTIN AGDATCMASKREDUCEDEFS askReduceDefs #-} +{-# BUILTIN AGDATCMFORMATERRORPARTS formatErrorParts #-} +{-# BUILTIN AGDATCMDEBUGPRINT debugPrint #-} +{-# BUILTIN AGDATCMNOCONSTRAINTS noConstraints #-} +{-# BUILTIN AGDATCMRUNSPECULATIVE runSpeculative #-} +{-# BUILTIN AGDATCMGETINSTANCES getInstances #-} + +-- All the TC primitives are compiled to functions that return +-- undefined, rather than just undefined, in an attempt to make sure +-- that code will run properly. +{-# COMPILE JS returnTC = _ => _ => _ => undefined #-} +{-# COMPILE JS bindTC = _ => _ => _ => _ => + _ => _ => undefined #-} +{-# COMPILE JS unify = _ => _ => undefined #-} +{-# COMPILE JS typeError = _ => _ => _ => undefined #-} +{-# COMPILE JS inferType = _ => undefined #-} +{-# COMPILE JS checkType = _ => _ => undefined #-} +{-# COMPILE JS normalise = _ => undefined #-} +{-# COMPILE JS reduce = _ => undefined #-} +{-# COMPILE JS catchTC = _ => _ => _ => _ => undefined #-} +{-# COMPILE JS quoteTC = _ => _ => _ => undefined #-} +{-# COMPILE JS unquoteTC = _ => _ => _ => undefined #-} +{-# COMPILE JS quoteωTC = _ => _ => undefined #-} +{-# COMPILE JS getContext = undefined #-} +{-# COMPILE JS extendContext = _ => _ => _ => _ => _ => undefined #-} +{-# COMPILE JS inContext = _ => _ => _ => _ => undefined #-} +{-# COMPILE JS freshName = _ => undefined #-} +{-# COMPILE JS declareDef = _ => _ => undefined #-} +{-# COMPILE JS declarePostulate = _ => _ => undefined #-} +{-# COMPILE JS declareData = _ => _ => _ => undefined #-} +{-# COMPILE JS defineData = _ => _ => undefined #-} +{-# COMPILE JS defineFun = _ => _ => undefined #-} +{-# COMPILE JS getType = _ => undefined #-} +{-# COMPILE JS getDefinition = _ => undefined #-} +{-# COMPILE JS blockTC = _ => _ => undefined #-} +{-# COMPILE JS commitTC = undefined #-} +{-# COMPILE JS isMacro = _ => undefined #-} +{-# COMPILE JS pragmaForeign = _ => _ => undefined #-} +{-# COMPILE JS pragmaCompile = _ => _ => _ => undefined #-} +{-# COMPILE JS withNormalisation = _ => _ => _ => _ => undefined #-} +{-# COMPILE JS withReconstructed = _ => _ => _ => _ => undefined #-} +{-# COMPILE JS withExpandLast = _ => _ => _ => _ => undefined #-} +{-# COMPILE JS withReduceDefs = _ => _ => _ => _ => undefined #-} +{-# COMPILE JS askNormalisation = undefined #-} +{-# COMPILE JS askReconstructed = undefined #-} +{-# COMPILE JS askExpandLast = undefined #-} +{-# COMPILE JS askReduceDefs = undefined #-} +{-# COMPILE JS debugPrint = _ => _ => _ => undefined #-} +{-# COMPILE JS noConstraints = _ => _ => _ => undefined #-} +{-# COMPILE JS runSpeculative = _ => _ => _ => undefined #-} +{-# COMPILE JS getInstances = _ => undefined #-} + +private + filter : (Name Bool) List Name List Name + filter p [] = [] + filter p (x xs) with p x + ... | true = x filter p xs + ... | false = filter p xs + + _∈_ : Name List Name Bool + n [] = false + n (n' l) with primQNameEquality n n' + ... | true = true + ... | false = n l + + _∉_ : Name List Name Bool + n l with n l + ... | true = false + ... | false = true + + _++_ : List Name List Name List Name + [] ++ l = l + (x xs) ++ l = x (xs ++ l) + + combineReduceDefs : (Σ Bool λ _ List Name) (Σ Bool λ _ List Name) (Σ Bool λ _ List Name) + combineReduceDefs (true , defs₁) (true , defs₂) = (true , filter (_∈ defs₁) defs₂) + combineReduceDefs (false , defs₁) (true , defs₂) = (true , filter (_∉ defs₁) defs₂) + combineReduceDefs (true , defs₁) (false , defs₂) = (true , filter (_∉ defs₂) defs₁) + combineReduceDefs (false , defs₁) (false , defs₂) = (false , defs₁ ++ defs₂) + +onlyReduceDefs dontReduceDefs : {a} {A : Set a} List Name TC A TC A +onlyReduceDefs defs x = bindTC askReduceDefs exDefs withReduceDefs (combineReduceDefs (true , defs) exDefs) x) +dontReduceDefs defs x = bindTC askReduceDefs exDefs withReduceDefs (combineReduceDefs (false , defs) exDefs) x) + +blockOnMeta : {a} {A : Set a} Meta TC A +blockOnMeta m = blockTC (blockerMeta m) + +{-# WARNING_ON_USAGE onlyReduceDefs "DEPRECATED: Use `withReduceDefs` instead of `onlyReduceDefs`" #-} +{-# WARNING_ON_USAGE dontReduceDefs "DEPRECATED: Use `withReduceDefs` instead of `dontReduceDefs`" #-} diff --git a/src/docs/Agda.Builtin.Sigma.html b/src/docs/Agda.Builtin.Sigma.html new file mode 100644 index 0000000..69098bb --- /dev/null +++ b/src/docs/Agda.Builtin.Sigma.html @@ -0,0 +1,17 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Sigma where + +open import Agda.Primitive + +record Σ {a b} (A : Set a) (B : A Set b) : Set (a b) where + constructor _,_ + field + fst : A + snd : B fst + +open Σ public + +infixr 4 _,_ + +{-# BUILTIN SIGMA Σ #-} diff --git a/src/docs/Agda.Builtin.String.html b/src/docs/Agda.Builtin.String.html new file mode 100644 index 0000000..d61d218 --- /dev/null +++ b/src/docs/Agda.Builtin.String.html @@ -0,0 +1,36 @@ +{-# OPTIONS --cubical-compatible --safe --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.String where + +open import Agda.Builtin.Bool +open import Agda.Builtin.Char +open import Agda.Builtin.List +open import Agda.Builtin.Maybe +open import Agda.Builtin.Nat using (Nat) +open import Agda.Builtin.Sigma + +postulate String : Set +{-# BUILTIN STRING String #-} + +primitive + primStringUncons : String Maybe (Σ Char _ String)) + primStringToList : String List Char + primStringFromList : List Char String + primStringAppend : String String String + primStringEquality : String String Bool + primShowChar : Char String + primShowString : String String + primShowNat : Nat String + +{-# COMPILE JS primStringUncons = function(x) { + if (x === "") { return z_jAgda_Agda_Builtin_Maybe["Maybe"]["nothing"]; }; + return z_jAgda_Agda_Builtin_Maybe["Maybe"]["just"](z_jAgda_Agda_Builtin_Sigma["_,_"](x.charAt(0))(x.slice(1))); + } + #-} +{-# COMPILE JS primStringToList = function(x) { return x.split(""); } #-} +{-# COMPILE JS primStringFromList = function(x) { return x.join(""); } #-} +{-# COMPILE JS primStringAppend = function(x) { return function(y) { return x+y; }; } #-} +{-# COMPILE JS primStringEquality = function(x) { return function(y) { return x===y; }; } #-} +{-# COMPILE JS primShowChar = function(x) { return JSON.stringify(x); } #-} +{-# COMPILE JS primShowString = function(x) { return JSON.stringify(x); } #-} +{-# COMPILE JS primShowNat = function(x) { return JSON.stringify(x); } #-} diff --git a/src/docs/Agda.Builtin.Unit.html b/src/docs/Agda.Builtin.Unit.html new file mode 100644 index 0000000..37e0b7d --- /dev/null +++ b/src/docs/Agda.Builtin.Unit.html @@ -0,0 +1,10 @@ +{-# OPTIONS --cubical-compatible --safe --no-universe-polymorphism + --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Unit where + +record : Set where + instance constructor tt + +{-# BUILTIN UNIT #-} +{-# COMPILE GHC = data () (()) #-} diff --git a/src/docs/Agda.Builtin.Word.html b/src/docs/Agda.Builtin.Word.html new file mode 100644 index 0000000..9151f5b --- /dev/null +++ b/src/docs/Agda.Builtin.Word.html @@ -0,0 +1,13 @@ +{-# OPTIONS --cubical-compatible --safe --no-universe-polymorphism + --no-sized-types --no-guardedness --level-universe #-} + +module Agda.Builtin.Word where + +open import Agda.Builtin.Nat + +postulate Word64 : Set +{-# BUILTIN WORD64 Word64 #-} + +primitive + primWord64ToNat : Word64 Nat + primWord64FromNat : Nat Word64 diff --git a/src/docs/Agda.Primitive.Cubical.html b/src/docs/Agda.Primitive.Cubical.html new file mode 100644 index 0000000..5f462f1 --- /dev/null +++ b/src/docs/Agda.Primitive.Cubical.html @@ -0,0 +1,78 @@ +{-# OPTIONS --erased-cubical #-} + +module Agda.Primitive.Cubical where + +{-# BUILTIN CUBEINTERVALUNIV IUniv #-} -- IUniv : SSet₁ +{-# BUILTIN INTERVAL I #-} -- I : IUniv + +{-# BUILTIN IZERO i0 #-} +{-# BUILTIN IONE i1 #-} + +-- I is treated as the type of booleans. +{-# COMPILE JS i0 = false #-} +{-# COMPILE JS i1 = true #-} + +infix 30 primINeg +infixr 20 primIMin primIMax + +primitive + primIMin : I I I + primIMax : I I I + primINeg : I I + +{-# BUILTIN ISONE IsOne #-} -- IsOne : I → Setω + +postulate + itIsOne : IsOne i1 + IsOne1 : i j IsOne i IsOne (primIMax i j) + IsOne2 : i j IsOne j IsOne (primIMax i j) + +{-# BUILTIN ITISONE itIsOne #-} +{-# BUILTIN ISONE1 IsOne1 #-} +{-# BUILTIN ISONE2 IsOne2 #-} + +-- IsOne i is treated as the unit type. +{-# COMPILE JS itIsOne = { "tt" : a => a["tt"]() } #-} +{-# COMPILE JS IsOne1 = + _ => _ => _ => { return { "tt" : a => a["tt"]() } } + #-} +{-# COMPILE JS IsOne2 = + _ => _ => _ => { return { "tt" : a => a["tt"]() } } + #-} + +-- Partial : ∀{ℓ} (i : I) (A : Set ℓ) → Set ℓ +-- Partial i A = IsOne i → A + +{-# BUILTIN PARTIAL Partial #-} +{-# BUILTIN PARTIALP PartialP #-} + +postulate + isOneEmpty : {} {A : Partial i0 (Set )} PartialP i0 A + +{-# BUILTIN ISONEEMPTY isOneEmpty #-} + +-- Partial i A and PartialP i A are treated as IsOne i → A. +{-# COMPILE JS isOneEmpty = + _ => x => _ => x({ "tt" : a => a["tt"]() }) + #-} + +primitive + primPOr : {} (i j : I) {A : Partial (primIMax i j) (Set )} + (u : PartialP i z A (IsOne1 i j z))) + (v : PartialP j z A (IsOne2 i j z))) + PartialP (primIMax i j) A + + -- Computes in terms of primHComp and primTransp + primComp : {} (A : (i : I) Set ( i)) {φ : I} (u : i Partial φ (A i)) (a : A i0) A i1 + +syntax primPOr p q u t = [ p u , q t ] + +primitive + primTransp : {} (A : (i : I) Set ( i)) (φ : I) (a : A i0) A i1 + primHComp : {} {A : Set } {φ : I} (u : i Partial φ A) (a : A) A + + +postulate + PathP : {} (A : I Set ) A i0 A i1 Set + +{-# BUILTIN PATHP PathP #-} diff --git a/src/docs/Agda.Primitive.html b/src/docs/Agda.Primitive.html new file mode 100644 index 0000000..5ade3e1 --- /dev/null +++ b/src/docs/Agda.Primitive.html @@ -0,0 +1,41 @@ +-- The Agda primitives (preloaded). + +{-# OPTIONS --cubical-compatible --no-import-sorts --level-universe #-} + +module Agda.Primitive where + +------------------------------------------------------------------------ +-- Universe levels +------------------------------------------------------------------------ + +infixl 6 _⊔_ + +{-# BUILTIN PROP Prop #-} +{-# BUILTIN TYPE Set #-} +{-# BUILTIN STRICTSET SSet #-} + +{-# BUILTIN PROPOMEGA Propω #-} +{-# BUILTIN SETOMEGA Setω #-} +{-# BUILTIN STRICTSETOMEGA SSetω #-} + +{-# BUILTIN LEVELUNIV LevelUniv #-} + +-- Level is the first thing we need to define. +-- The other postulates can only be checked if built-in Level is known. + +postulate + Level : LevelUniv + +-- MAlonzo compiles Level to (). This should be safe, because it is +-- not possible to pattern match on levels. + +{-# BUILTIN LEVEL Level #-} + +postulate + lzero : Level + lsuc : ( : Level) Level + _⊔_ : (ℓ₁ ℓ₂ : Level) Level + +{-# BUILTIN LEVELZERO lzero #-} +{-# BUILTIN LEVELSUC lsuc #-} +{-# BUILTIN LEVELMAX _⊔_ #-} diff --git a/src/docs/Agda.css b/src/docs/Agda.css new file mode 100644 index 0000000..86813a5 --- /dev/null +++ b/src/docs/Agda.css @@ -0,0 +1,41 @@ +/* Aspects. */ +.Agda .Comment { color: #B22222 } +.Agda .Background {} +.Agda .Markup { color: #000000 } +.Agda .Keyword { color: #CD6600 } +.Agda .String { color: #B22222 } +.Agda .Number { color: #A020F0 } +.Agda .Symbol { color: #404040 } +.Agda .PrimitiveType { color: #0000CD } +.Agda .Pragma { color: black } +.Agda .Operator {} +.Agda .Hole { background: #B4EEB4 } + +/* NameKinds. */ +.Agda .Bound { color: black } +.Agda .Generalizable { color: black } +.Agda .InductiveConstructor { color: #008B00 } +.Agda .CoinductiveConstructor { color: #8B7500 } +.Agda .Datatype { color: #0000CD } +.Agda .Field { color: #EE1289 } +.Agda .Function { color: #0000CD } +.Agda .Module { color: #A020F0 } +.Agda .Postulate { color: #0000CD } +.Agda .Primitive { color: #0000CD } +.Agda .Record { color: #0000CD } + +/* OtherAspects. */ +.Agda .DottedPattern {} +.Agda .UnsolvedMeta { color: black; background: yellow } +.Agda .UnsolvedConstraint { color: black; background: yellow } +.Agda .TerminationProblem { color: black; background: #FFA07A } +.Agda .IncompletePattern { color: black; background: #F5DEB3 } +.Agda .Error { color: red; text-decoration: underline } +.Agda .TypeChecks { color: black; background: #ADD8E6 } +.Agda .Deadcode { color: black; background: #808080 } +.Agda .ShadowingInTelescope { color: black; background: #808080 } + +/* Standard attributes. */ +.Agda a { text-decoration: none } +.Agda a[href]:hover { background-color: #B4EEB4 } +.Agda [href].hover-highlight { background-color: #B4EEB4; } diff --git a/src/docs/Cubical.Core.Everything.html b/src/docs/Cubical.Core.Everything.html new file mode 100644 index 0000000..6c633b2 --- /dev/null +++ b/src/docs/Cubical.Core.Everything.html @@ -0,0 +1,8 @@ +{-# OPTIONS --safe #-} +module Cubical.Core.Everything where + +-- Basic primitives (some are from Agda.Primitive) +open import Cubical.Core.Primitives public + +-- Definition of equivalences and Glue types +open import Cubical.Core.Glue public diff --git a/src/docs/Cubical.Core.Glue.html b/src/docs/Cubical.Core.Glue.html new file mode 100644 index 0000000..404a546 --- /dev/null +++ b/src/docs/Cubical.Core.Glue.html @@ -0,0 +1,139 @@ +{- + +This file contains: + +- Definitions equivalences + +- Glue types + +-} +{-# OPTIONS --safe #-} +module Cubical.Core.Glue where + +open import Cubical.Core.Primitives + +open import Agda.Builtin.Cubical.Glue public + using ( isEquiv -- ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} (f : A → B) → Type (ℓ ⊔ ℓ') + + ; equiv-proof -- ∀ (y : B) → isContr (fiber f y) + + ; _≃_ -- ∀ {ℓ ℓ'} (A : Type ℓ) (B : Type ℓ') → Type (ℓ ⊔ ℓ') + + ; equivFun -- ∀ {ℓ ℓ'} {A : Type ℓ} {B : Type ℓ'} → A ≃ B → A → B + + ; equivProof -- ∀ {ℓ ℓ'} (T : Type ℓ) (A : Type ℓ') (w : T ≃ A) (a : A) φ → + -- Partial φ (fiber (equivFun w) a) → fiber (equivFun w) a + + ; primGlue -- ∀ {ℓ ℓ'} (A : Type ℓ) {φ : I} (T : Partial φ (Type ℓ')) + -- → (e : PartialP φ (λ o → T o ≃ A)) → Type ℓ' + + ; prim^unglue -- ∀ {ℓ ℓ'} {A : Type ℓ} {φ : I} {T : Partial φ (Type ℓ')} + -- → {e : PartialP φ (λ o → T o ≃ A)} → primGlue A T e → A + + -- The ∀ operation on I. This is commented out as it is not currently used for anything + -- ; primFaceForall -- (I → I) → I + ) + renaming ( prim^glue to glue -- ∀ {ℓ ℓ'} {A : Type ℓ} {φ : I} {T : Partial φ (Type ℓ')} + -- → {e : PartialP φ (λ o → T o ≃ A)} + -- → PartialP φ T → A → primGlue A T e + ) + +private + variable + ℓ' : Level + +-- Uncurry Glue to make it more pleasant to use +Glue : (A : Type ) {φ : I} + (Te : Partial φ (Σ[ T Type ℓ' ] T A)) + Type ℓ' +Glue A Te = primGlue A x Te x .fst) x Te x .snd) + +-- Make the φ argument of prim^unglue explicit +unglue : {A : Type } (φ : I) {T : Partial φ (Type ℓ')} + {e : PartialP φ o T o A)} primGlue A T e A +unglue φ = prim^unglue {φ = φ} + +-- People unfamiliar with [Glue], [glue] and [uglue] can find the types below more +-- informative as they demonstrate the computational behavior. +-- +-- Full inference rules can be found in Section 6 of CCHM: +-- https://arxiv.org/pdf/1611.02108.pdf +-- Cubical Type Theory: a constructive interpretation of the univalence axiom +-- Cyril Cohen, Thierry Coquand, Simon Huber, Anders Mörtberg +private + + module GluePrims (A : Type ) {φ : I} (Te : Partial φ (Σ[ T Type ℓ' ] T A)) where + T : Partial φ (Type ℓ') + T φ1 = Te φ1 .fst + e : PartialP φ φ T φ A) + e φ1 = Te φ1 .snd + + -- Glue can be seen as a subtype of Type that, at φ, is definitionally equal to the left type + -- of the given equivalences. + Glue-S : Type ℓ' [ φ T ] + Glue-S = inS (Glue A Te) + + -- Which means partial elements of T are partial elements of Glue + coeT→G : + (t : PartialP φ T) + Partial φ (Glue A Te) + coeT→G t (φ = i1) = t 1=1 + + -- ... and elements of Glue can be seen as partial elements of T + coeG→T : + (g : Glue A Te) + PartialP φ T + coeG→T g (φ = i1) = g + + -- What about elements that are applied to the equivalences? + trans-e : + (t : PartialP φ T) + Partial φ A + trans-e t ϕ1 = equivFun (e ϕ1) (t ϕ1) + + -- glue gives a partial element of Glue given an element of A. Note that it "undoes" + -- the application of the equivalences! + glue-S : + (t : PartialP φ T) + A [ φ trans-e t ] + Glue A Te [ φ coeT→G t ] + glue-S t s = inS (glue t (outS s)) + + -- typechecking glue enforces this, e.g. you can not simply write + -- glue-bad : (t : PartialP φ T) → A → Glue A Te + -- glue-bad t s = glue t s + + -- unglue does the inverse: + unglue-S : + (b : Glue A Te) + A [ φ trans-e (coeG→T b) ] + unglue-S b = inS (unglue φ b) + + module GlueTransp (A : I Type ) (Te : (i : I) Partial (i ~ i) (Σ[ T Type ℓ' ] T A i)) where + A0 A1 : Type + A0 = A i0 + A1 = A i1 + T : (i : I) Partial (i ~ i) (Type ℓ') + T i φ = Te i φ .fst + e : (i : I) PartialP (i ~ i) φ T i φ A i) + e i φ = Te i φ .snd + T0 T1 : Type ℓ' + T0 = T i0 1=1 + T1 = T i1 1=1 + e0 : T0 A0 + e0 = e i0 1=1 + e1 : T1 A1 + e1 = e i1 1=1 + + open import Cubical.Foundations.Prelude using (transport) + transportA : A0 A1 + transportA = transport i A i) + + -- copied over from Foundations/Equiv for readability, can't directly import due to cyclic dependency + invEq : {X : Type ℓ'} {ℓ''} {Y : Type ℓ''} (w : X Y) Y X + invEq w y = w .snd .equiv-proof y .fst .fst + + -- transport in Glue reduces to transport in A + the application of the equivalences in forward and backward + -- direction. + transp-S : (t0 : T0) T1 [ i1 _ invEq e1 (transportA (equivFun e0 t0))) ] + transp-S t0 = inS (transport i Glue (A i) (Te i)) t0) diff --git a/src/docs/Cubical.Core.Primitives.html b/src/docs/Cubical.Core.Primitives.html new file mode 100644 index 0000000..afa41a7 --- /dev/null +++ b/src/docs/Cubical.Core.Primitives.html @@ -0,0 +1,209 @@ +{- +This file document and export the main primitives of Cubical Agda. It +also defines some basic derived operations (composition and filling). + +-} +{-# OPTIONS --safe #-} +module Cubical.Core.Primitives where + +open import Agda.Builtin.Cubical.Path public +open import Agda.Builtin.Cubical.Sub public + renaming (primSubOut to outS) +open import Agda.Primitive.Cubical public + renaming ( primIMin to _∧_ -- I → I → I + ; primIMax to _∨_ -- I → I → I + ; primINeg to ~_ -- I → I + ; isOneEmpty to empty + ; primComp to comp + ; primHComp to hcomp + ; primTransp to transp + ; itIsOne to 1=1 ) + +-- These two are to make sure all the primitives are loaded and ready +-- to compute hcomp/transp for the universe. +import Agda.Builtin.Cubical.Glue +-- HCompU is already imported from Glue, and older Agda versions do +-- not have it. So we comment it out for now. +-- import Agda.Builtin.Cubical.HCompU + +open import Agda.Primitive public + using ( Level + ; SSet ) + renaming ( lzero to ℓ-zero + ; lsuc to ℓ-suc + ; _⊔_ to ℓ-max + ; Set to Type + ; Setω to Typeω ) +open import Agda.Builtin.Sigma public + +-- This file document the Cubical Agda primitives. The primitives +-- themselves are bound by the Agda files imported above. + +-- * The Interval +-- I : IUniv + +-- Endpoints, Connections, Reversal +-- i0 i1 : I +-- _∧_ _∨_ : I → I → I +-- ~_ : I → I + + +-- * Dependent path type. (Path over Path) + +-- Introduced with lambda abstraction and eliminated with application, +-- just like function types. + +-- PathP : ∀ {ℓ} (A : I → Type ℓ) → A i0 → A i1 → Type ℓ + +infix 4 _[_≡_] + +_[_≡_] : {} (A : I Type ) A i0 A i1 Type +_[_≡_] = PathP + + +-- Non dependent path types + +Path : {} (A : Type ) A A Type +Path A a b = PathP _ A) a b + +-- PathP (λ i → A) x y gets printed as x ≡ y when A does not mention i. +-- _≡_ : ∀ {ℓ} {A : Type ℓ} → A → A → Type ℓ +-- _≡_ {A = A} = PathP (λ _ → A) + + +-- * @IsOne r@ represents the constraint "r = i1". +-- Often we will use "φ" for elements of I, when we intend to use them +-- with IsOne (or Partial[P]). +-- IsOne : I → SSet ℓ-zero + +-- i1 is indeed equal to i1. +-- 1=1 : IsOne i1 + + +-- * Types of partial elements, and their dependent version. + +-- "Partial φ A" is a special version of "IsOne φ → A" with a more +-- extensional judgmental equality. +-- "PartialP φ A" allows "A" to be defined only on "φ". + +-- Partial : ∀ {ℓ} → I → Type ℓ → SSet ℓ +-- PartialP : ∀ {ℓ} → (φ : I) → Partial φ (Type ℓ) → SSet ℓ + +-- Partial elements are introduced by pattern matching with (r = i0) +-- or (r = i1) constraints, like so: + +private + sys : i Partial (i ~ i) Type₁ + sys i (i = i0) = Type₀ + sys i (i = i1) = Type₀ Type₀ + + -- It also works with pattern matching lambdas: + -- http://wiki.portal.chalmers.se/agda/pmwiki.php?n=ReferenceManual.PatternMatchingLambdas + sys' : i Partial (i ~ i) Type₁ + sys' i = λ { (i = i0) Type₀ + ; (i = i1) Type₀ Type₀ + } + + -- When the cases overlap they must agree. + sys2 : i j Partial (i (i j)) Type₁ + sys2 i j = λ { (i = i1) Type₀ + ; (i = i1) (j = i1) Type₀ + } + + -- (i0 = i1) is actually absurd. + sys3 : Partial i0 Type₁ + sys3 = λ { () } + + +-- * There are cubical subtypes as in CCHM. Note that these are not +-- fibrant (hence in SSet ℓ): + +_[_↦_] : {} (A : Type ) (φ : I) (u : Partial φ A) SSet +A [ φ u ] = Sub A φ u + +infix 4 _[_↦_] + +-- Any element u : A can be seen as an element of A [ φ ↦ u ] which +-- agrees with u on φ: + +-- inS : ∀ {ℓ} {A : Type ℓ} {φ} (u : A) → A [ φ ↦ (λ _ → u) ] + +-- One can also forget that an element agrees with u on φ: + +-- outS : ∀ {ℓ} {A : Type ℓ} {φ : I} {u : Partial φ A} → A [ φ ↦ u ] → A + + +-- * Composition operation according to [CCHM 18]. +-- When calling "comp A φ u a" Agda makes sure that "a" agrees with "u i0" on "φ". +-- compCCHM : ∀ {ℓ} (A : (i : I) → Type ℓ) (φ : I) (u : ∀ i → Partial φ (A i)) (a : A i0) → A i1 + +-- Note: this is not recommended to use, instead use the CHM +-- primitives! The reason is that these work with HITs and produce +-- fewer empty systems. + + +-- * Generalized transport and homogeneous composition [CHM 18]. + +-- When calling "transp A φ a" Agda makes sure that "A" is constant on "φ" (see below). +-- transp : ∀ {ℓ} (A : I → Type ℓ) (φ : I) (a : A i0) → A i1 + +-- "A" being constant on "φ" means that "A" should be a constant function whenever the +-- constraint "φ = i1" is satisfied. For example: +-- - If "φ" is "i0" then "A" can be anything, since this condition is vacuously true. +-- - If "φ" is "i1" then "A" must be a constant function. +-- - If "φ" is some in-scope variable "i" then "A" only needs to be a constant function +-- when substituting "i1" for "i". + +-- When calling "hcomp A φ u a" Agda makes sure that "a" agrees with "u i0" on "φ". +-- hcomp : ∀ {ℓ} {A : Type ℓ} {φ : I} (u : I → Partial φ A) (a : A) → A + +private + variable + : Level + ℓ' : I Level + +-- Homogeneous filling +hfill : {A : Type } + {φ : I} + (u : i Partial φ A) + (u0 : A [ φ u i0 ]) + ----------------------- + (i : I) A +hfill {φ = φ} u u0 i = + hcomp j λ { (φ = i1) u (i j) 1=1 + ; (i = i0) outS u0 }) + (outS u0) + +-- Heterogeneous composition can defined as in CHM, however we use the +-- builtin one as it doesn't require u0 to be a cubical subtype. This +-- reduces the number of inS's a lot. +-- comp : (A : ∀ i → Type (ℓ' i)) +-- {φ : I} +-- (u : ∀ i → Partial φ (A i)) +-- (u0 : A i0 [ φ ↦ u i0 ]) +-- → --------------------------- +-- A i1 +-- comp A {φ = φ} u u0 = +-- hcomp (λ i → λ { (φ = i1) → transp (λ j → A (i ∨ j)) i (u _ 1=1) }) +-- (transp A i0 (outS u0)) + +-- Heterogeneous filling defined using comp +fill : (A : i Type (ℓ' i)) + {φ : I} + (u : i Partial φ (A i)) + (u0 : A i0 [ φ u i0 ]) + --------------------------- + (i : I) A i +fill A {φ = φ} u u0 i = + comp j A (i j)) + j λ { (φ = i1) u (i j) 1=1 + ; (i = i0) outS u0 }) + (outS u0) + +-- Σ-types +infix 2 Σ-syntax + +Σ-syntax : { ℓ'} (A : Type ) (B : A Type ℓ') Type (ℓ-max ℓ') +Σ-syntax = Σ + +syntax Σ-syntax A x B) = Σ[ x A ] B diff --git a/src/docs/Cubical.Data.Bool.Base.html b/src/docs/Cubical.Data.Bool.Base.html new file mode 100644 index 0000000..e898f51 --- /dev/null +++ b/src/docs/Cubical.Data.Bool.Base.html @@ -0,0 +1,101 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool.Base where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Empty +open import Cubical.Data.Sum.Base +open import Cubical.Data.Unit.Base + +open import Cubical.Relation.Nullary.Base + +-- Obtain the booleans +open import Agda.Builtin.Bool public + +private + variable + : Level + A : Type + +infixr 6 _and_ +infixr 5 _or_ +infix 0 if_then_else_ + +not : Bool Bool +not true = false +not false = true + +_or_ : Bool Bool Bool +false or x = x +true or _ = true + +_and_ : Bool Bool Bool +false and _ = false +true and x = x + +-- xor / mod-2 addition +_⊕_ : Bool Bool Bool +false x = x +true x = not x + +if_then_else_ : Bool A A A +if true then x else y = x +if false then x else y = y + +_≟_ : Discrete Bool +false false = yes refl +false true = no λ p subst b if b then else Bool) p true +true false = no λ p subst b if b then Bool else ) p true +true true = yes refl + +Dec→Bool : Dec A Bool +Dec→Bool (yes p) = true +Dec→Bool (no ¬p) = false + +-- Helpers for automatic proof +Bool→Type : Bool Type₀ +Bool→Type true = Unit +Bool→Type false = + +Bool→Type* : Bool Type +Bool→Type* true = Unit* +Bool→Type* false = ⊥* + +True : Dec A Type₀ +True Q = Bool→Type (Dec→Bool Q) + +False : Dec A Type₀ +False Q = Bool→Type (not (Dec→Bool Q)) + +toWitness : {Q : Dec A} True Q A +toWitness {Q = yes p} _ = p + +toWitnessFalse : {Q : Dec A} False Q ¬ A +toWitnessFalse {Q = no ¬p} _ = ¬p + +dichotomyBool : (x : Bool) (x true) (x false) +dichotomyBool true = inl refl +dichotomyBool false = inr refl + +dichotomyBoolSym : (x : Bool) (x false) (x true) +dichotomyBoolSym false = inl refl +dichotomyBoolSym true = inr refl + +-- TODO: this should be uncommented and implemented using instance arguments +-- _==_ : {dA : Discrete A} → A → A → Bool +-- _==_ {dA = dA} x y = Dec→Bool (dA x y) + +-- Universe lifted booleans +Bool* : {} Type +Bool* = Lift Bool + +true* : {} Bool* {} +true* = lift true + +false* : {} Bool* {} +false* = lift false + +-- Pointed version +Bool*∙ : {} Σ[ X Type ] X +fst Bool*∙ = Bool* +snd Bool*∙ = true* diff --git a/src/docs/Cubical.Data.Bool.Properties.html b/src/docs/Cubical.Data.Bool.Properties.html new file mode 100644 index 0000000..858a2e0 --- /dev/null +++ b/src/docs/Cubical.Data.Bool.Properties.html @@ -0,0 +1,431 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Functions.Involution + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed + +open import Cubical.Data.Sum hiding (elim) +open import Cubical.Data.Bool.Base +open import Cubical.Data.Empty hiding (elim) +open import Cubical.Data.Empty as Empty hiding (elim) +open import Cubical.Data.Sigma +open import Cubical.Data.Unit using (Unit; isPropUnit) + +open import Cubical.HITs.PropositionalTruncation hiding (elim; rec) + +open import Cubical.Relation.Nullary + +private + variable + : Level + A : Type + +elim : {} {A : Bool Type } + A true + A false + (b : Bool) A b +elim t f false = f +elim t f true = t + +notnot : x not (not x) x +notnot true = refl +notnot false = refl + +notIso : Iso Bool Bool +Iso.fun notIso = not +Iso.inv notIso = not +Iso.rightInv notIso = notnot +Iso.leftInv notIso = notnot + +notIsEquiv : isEquiv not +notIsEquiv = involIsEquiv {f = not} notnot + +notEquiv : Bool Bool +notEquiv = involEquiv {f = not} notnot + +notEq : Bool Bool +notEq = involPath {f = not} notnot + +private + -- This computes to false as expected + nfalse : Bool + nfalse = transp i notEq i) i0 true + + -- Sanity check + nfalsepath : nfalse false + nfalsepath = refl + +K-Bool + : (P : {b : Bool} b b Type ) + (∀{b} P {b} refl) + ∀{b} (q : b b) P q +K-Bool P Pr {false} = J (λ{ false q P q ; true _ Lift }) Pr +K-Bool P Pr {true} = J (λ{ true q P q ; false _ Lift }) Pr + +isSetBool : isSet Bool +isSetBool a b = J _ p q p q) (K-Bool (refl ≡_) refl) + +true≢false : ¬ true false +true≢false p = subst b if b then Bool else ) p true + +false≢true : ¬ false true +false≢true p = subst b if b then else Bool) p true + +¬true→false : (x : Bool) ¬ x true x false +¬true→false false _ = refl +¬true→false true p = Empty.rec (p refl) + +¬false→true : (x : Bool) ¬ x false x true +¬false→true false p = Empty.rec (p refl) +¬false→true true _ = refl + +not≢const : x ¬ not x x +not≢const false = true≢false +not≢const true = false≢true + +-- or properties +or-zeroˡ : x true or x true +or-zeroˡ _ = refl + +or-zeroʳ : x x or true true +or-zeroʳ false = refl +or-zeroʳ true = refl + +or-identityˡ : x false or x x +or-identityˡ _ = refl + +or-identityʳ : x x or false x +or-identityʳ false = refl +or-identityʳ true = refl + +or-comm : x y x or y y or x +or-comm false false = refl +or-comm false true = refl +or-comm true false = refl +or-comm true true = refl + +or-assoc : x y z x or (y or z) (x or y) or z +or-assoc false y z = refl +or-assoc true y z = refl + +or-idem : x x or x x +or-idem false = refl +or-idem true = refl + +-- and properties +and-zeroˡ : x false and x false +and-zeroˡ _ = refl + +and-zeroʳ : x x and false false +and-zeroʳ false = refl +and-zeroʳ true = refl + +and-identityˡ : x true and x x +and-identityˡ _ = refl + +and-identityʳ : x x and true x +and-identityʳ false = refl +and-identityʳ true = refl + +and-comm : x y x and y y and x +and-comm false false = refl +and-comm false true = refl +and-comm true false = refl +and-comm true true = refl + +and-assoc : x y z x and (y and z) (x and y) and z +and-assoc false y z = refl +and-assoc true y z = refl + +and-idem : x x and x x +and-idem false = refl +and-idem true = refl + +-- xor properties +⊕-identityʳ : x x false x +⊕-identityʳ false = refl +⊕-identityʳ true = refl + +⊕-comm : x y x y y x +⊕-comm false false = refl +⊕-comm false true = refl +⊕-comm true false = refl +⊕-comm true true = refl + +⊕-assoc : x y z x (y z) (x y) z +⊕-assoc false y z = refl +⊕-assoc true false z = refl +⊕-assoc true true z = notnot z + +not-⊕ˡ : x y not (x y) not x y +not-⊕ˡ false y = refl +not-⊕ˡ true y = notnot y + +⊕-invol : x y x (x y) y +⊕-invol false x = refl +⊕-invol true x = notnot x + +isEquiv-⊕ : x isEquiv (x ⊕_) +isEquiv-⊕ x = involIsEquiv (⊕-invol x) + +⊕-Path : x Bool Bool +⊕-Path x = involPath {f = x ⊕_} (⊕-invol x) + +⊕-Path-refl : ⊕-Path false refl +⊕-Path-refl = isInjectiveTransport refl + +¬transportNot : ∀(P : Bool Bool) b ¬ (transport P (not b) transport P b) +¬transportNot P b eq = not≢const b sub + where + sub : not b b + sub = subst {A = Bool Bool} f f (not b) f b) + i c transport⁻Transport P c i) (cong (transport⁻ P) eq) + +module BoolReflection where + data Table (A : Type₀) (P : Bool A) : Type₀ where + inspect : (b c : A) + transport P false b + transport P true c + Table A P + + table : P Table Bool P + table = J Table (inspect false true refl refl) + + reflLemma : (P : Bool Bool) + transport P false false + transport P true true + transport P transport (⊕-Path false) + reflLemma P ff tt i false = ff i + reflLemma P ff tt i true = tt i + + notLemma : (P : Bool Bool) + transport P false true + transport P true false + transport P transport (⊕-Path true) + notLemma P ft tf i false = ft i + notLemma P ft tf i true = tf i + + categorize : P transport P transport (⊕-Path (transport P false)) + categorize P with table P + categorize P | inspect false true p q + = subst b transport P transport (⊕-Path b)) (sym p) (reflLemma P p q) + categorize P | inspect true false p q + = subst b transport P transport (⊕-Path b)) (sym p) (notLemma P p q) + categorize P | inspect false false p q + = Empty.rec (¬transportNot P false (q sym p)) + categorize P | inspect true true p q + = Empty.rec (¬transportNot P false (q sym p)) + + ⊕-complete : P P ⊕-Path (transport P false) + ⊕-complete P = isInjectiveTransport (categorize P) + + ⊕-comp : p q ⊕-Path p ⊕-Path q ⊕-Path (q p) + ⊕-comp p q = isInjectiveTransport i x ⊕-assoc q p x i) + + open Iso + + reflectIso : Iso Bool (Bool Bool) + reflectIso .fun = ⊕-Path + reflectIso .inv P = transport P false + reflectIso .leftInv = ⊕-identityʳ + reflectIso .rightInv P = sym (⊕-complete P) + + reflectEquiv : Bool (Bool Bool) + reflectEquiv = isoToEquiv reflectIso + +_≤_ : Bool Bool Type +true false = +_ _ = Unit + +_≥_ : Bool Bool Type +false true = +_ _ = Unit + +isProp≤ : b c isProp (b c) +isProp≤ true false = isProp⊥ +isProp≤ true true = isPropUnit +isProp≤ false false = isPropUnit +isProp≤ false true = isPropUnit + +isProp≥ : b c isProp (b c) +isProp≥ false true = isProp⊥ +isProp≥ true true = isPropUnit +isProp≥ false false = isPropUnit +isProp≥ true false = isPropUnit + +isProp-Bool→Type : b isProp (Bool→Type b) +isProp-Bool→Type false = isProp⊥ +isProp-Bool→Type true = isPropUnit + +isPropDep-Bool→Type : isPropDep Bool→Type +isPropDep-Bool→Type = isOfHLevel→isOfHLevelDep 1 isProp-Bool→Type + +IsoBool→∙ : {} {A : Pointed } Iso ((Bool , true) →∙ A) (typ A) +Iso.fun IsoBool→∙ f = fst f false +fst (Iso.inv IsoBool→∙ a) false = a +fst (Iso.inv (IsoBool→∙ {A = A}) a) true = pt A +snd (Iso.inv IsoBool→∙ a) = refl +Iso.rightInv IsoBool→∙ a = refl +Iso.leftInv IsoBool→∙ (f , p) = + ΣPathP ((funExt { false refl ; true sym p})) + , λ i j p (~ i j)) + +-- import here to avoid conflicts +open import Cubical.Data.Unit + +-- relation to hProp + +BoolProp≃BoolProp* : {a : Bool} Bool→Type a Bool→Type* {} a +BoolProp≃BoolProp* {a = true} = Unit≃Unit* +BoolProp≃BoolProp* {a = false} = uninhabEquiv Empty.rec Empty.rec* + +Bool→TypeInj : (a b : Bool) Bool→Type a Bool→Type b a b +Bool→TypeInj true true _ = refl +Bool→TypeInj true false p = Empty.rec (p .fst tt) +Bool→TypeInj false true p = Empty.rec (invEq p tt) +Bool→TypeInj false false _ = refl + +Bool→TypeInj* : (a b : Bool) Bool→Type* {} a Bool→Type* {} b a b +Bool→TypeInj* true true _ = refl +Bool→TypeInj* true false p = Empty.rec* (p .fst tt*) +Bool→TypeInj* false true p = Empty.rec* (invEq p tt*) +Bool→TypeInj* false false _ = refl + +isPropBool→Type : {a : Bool} isProp (Bool→Type a) +isPropBool→Type {a = true} = isPropUnit +isPropBool→Type {a = false} = isProp⊥ + +isPropBool→Type* : {a : Bool} isProp (Bool→Type* {} a) +isPropBool→Type* {a = true} = isPropUnit* +isPropBool→Type* {a = false} = isProp⊥* + +DecBool→Type : {a : Bool} Dec (Bool→Type a) +DecBool→Type {a = true} = yes tt +DecBool→Type {a = false} = no x x) + +DecBool→Type* : {a : Bool} Dec (Bool→Type* {} a) +DecBool→Type* {a = true} = yes tt* +DecBool→Type* {a = false} = no (lift x) x) + +Dec→DecBool : {P : Type } (dec : Dec P) P Bool→Type (Dec→Bool dec) +Dec→DecBool (yes p) _ = tt +Dec→DecBool (no ¬p) q = Empty.rec (¬p q) + +DecBool→Dec : {P : Type } (dec : Dec P) Bool→Type (Dec→Bool dec) P +DecBool→Dec (yes p) _ = p + +Dec≃DecBool : {P : Type } (h : isProp P)(dec : Dec P) P Bool→Type (Dec→Bool dec) +Dec≃DecBool h dec = propBiimpl→Equiv h isPropBool→Type (Dec→DecBool dec) (DecBool→Dec dec) + +Bool≡BoolDec : {a : Bool} a Dec→Bool (DecBool→Type {a = a}) +Bool≡BoolDec {a = true} = refl +Bool≡BoolDec {a = false} = refl + +Dec→DecBool* : {P : Type } (dec : Dec P) P Bool→Type* {} (Dec→Bool dec) +Dec→DecBool* (yes p) _ = tt* +Dec→DecBool* (no ¬p) q = Empty.rec (¬p q) + +DecBool→Dec* : {P : Type } (dec : Dec P) Bool→Type* {} (Dec→Bool dec) P +DecBool→Dec* (yes p) _ = p + +Dec≃DecBool* : {P : Type } (h : isProp P)(dec : Dec P) P Bool→Type* {} (Dec→Bool dec) +Dec≃DecBool* h dec = propBiimpl→Equiv h isPropBool→Type* (Dec→DecBool* dec) (DecBool→Dec* dec) + +Bool≡BoolDec* : {a : Bool} a Dec→Bool (DecBool→Type* {} {a = a}) +Bool≡BoolDec* {a = true} = refl +Bool≡BoolDec* {a = false} = refl + +Bool→Type× : (a b : Bool) Bool→Type (a and b) Bool→Type a × Bool→Type b +Bool→Type× true true _ = tt , tt +Bool→Type× true false p = Empty.rec p +Bool→Type× false true p = Empty.rec p +Bool→Type× false false p = Empty.rec p + +Bool→Type×' : (a b : Bool) Bool→Type a × Bool→Type b Bool→Type (a and b) +Bool→Type×' true true _ = tt +Bool→Type×' true false (_ , p) = Empty.rec p +Bool→Type×' false true (p , _) = Empty.rec p +Bool→Type×' false false (p , _) = Empty.rec p + +Bool→Type×≃ : (a b : Bool) Bool→Type a × Bool→Type b Bool→Type (a and b) +Bool→Type×≃ a b = + propBiimpl→Equiv (isProp× isPropBool→Type isPropBool→Type) isPropBool→Type + (Bool→Type×' a b) (Bool→Type× a b) + +Bool→Type⊎ : (a b : Bool) Bool→Type (a or b) Bool→Type a Bool→Type b +Bool→Type⊎ true true _ = inl tt +Bool→Type⊎ true false _ = inl tt +Bool→Type⊎ false true _ = inr tt +Bool→Type⊎ false false p = Empty.rec p + +Bool→Type⊎' : (a b : Bool) Bool→Type a Bool→Type b Bool→Type (a or b) +Bool→Type⊎' true true _ = tt +Bool→Type⊎' true false _ = tt +Bool→Type⊎' false true _ = tt +Bool→Type⊎' false false (inl p) = Empty.rec p +Bool→Type⊎' false false (inr p) = Empty.rec p + +PropBoolP→P : (dec : Dec A) Bool→Type (Dec→Bool dec) A +PropBoolP→P (yes p) _ = p + +P→PropBoolP : (dec : Dec A) A Bool→Type (Dec→Bool dec) +P→PropBoolP (yes p) _ = tt +P→PropBoolP (no ¬p) = ¬p + +Bool≡ : Bool Bool Bool +Bool≡ true true = true +Bool≡ true false = false +Bool≡ false true = false +Bool≡ false false = true + +Bool≡≃ : (a b : Bool) (a b) Bool→Type (Bool≡ a b) +Bool≡≃ true true = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) +Bool≡≃ true false = uninhabEquiv true≢false Empty.rec +Bool≡≃ false true = uninhabEquiv false≢true Empty.rec +Bool≡≃ false false = isContr→≃Unit (inhProp→isContr refl (isSetBool _ _)) +open Iso + +-- Bool is equivalent to bi-point type + +Iso-⊤⊎⊤-Bool : Iso (Unit Unit) Bool +Iso-⊤⊎⊤-Bool .fun (inl tt) = true +Iso-⊤⊎⊤-Bool .fun (inr tt) = false +Iso-⊤⊎⊤-Bool .inv true = inl tt +Iso-⊤⊎⊤-Bool .inv false = inr tt +Iso-⊤⊎⊤-Bool .leftInv (inl tt) = refl +Iso-⊤⊎⊤-Bool .leftInv (inr tt) = refl +Iso-⊤⊎⊤-Bool .rightInv true = refl +Iso-⊤⊎⊤-Bool .rightInv false = refl + +separatedBool : Separated Bool +separatedBool = Discrete→Separated _≟_ + + +Bool→Bool→∙Bool : Bool (Bool , true) →∙ (Bool , true) +Bool→Bool→∙Bool false = idfun∙ _ +Bool→Bool→∙Bool true = const∙ _ _ + +Iso-Bool→∙Bool-Bool : Iso ((Bool , true) →∙ (Bool , true)) Bool +Iso.fun Iso-Bool→∙Bool-Bool f = fst f false +Iso.inv Iso-Bool→∙Bool-Bool = Bool→Bool→∙Bool +Iso.rightInv Iso-Bool→∙Bool-Bool false = refl +Iso.rightInv Iso-Bool→∙Bool-Bool true = refl +Iso.leftInv Iso-Bool→∙Bool-Bool f = Σ≡Prop _ isSetBool _ _) (help _ refl) + where + help : (x : Bool) fst f false x + Bool→Bool→∙Bool (fst f false) .fst f .fst + help false p = funExt + λ { false j Bool→Bool→∙Bool (p j) .fst false) sym p + ; true j Bool→Bool→∙Bool (p j) .fst true) sym (snd f)} + help true p = j Bool→Bool→∙Bool (p j) .fst) + funExt λ { false sym p ; true sym (snd f)} diff --git a/src/docs/Cubical.Data.Bool.SwitchStatement.html b/src/docs/Cubical.Data.Bool.SwitchStatement.html new file mode 100644 index 0000000..89c5c7c --- /dev/null +++ b/src/docs/Cubical.Data.Bool.SwitchStatement.html @@ -0,0 +1,42 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool.SwitchStatement where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Bool.Base +open import Cubical.Data.Nat + +{- + Switch-case: + + _==_ : A → A → Bool + + _ : B + _ = switch (λ x → x == fixedValue) cases + case value1 ⇒ result1 break + case value2 ⇒ result2 break + ... + case valueN ⇒ resultN break + default⇒ defaultResult +-} + + +private + variable + ℓ' : Level + + +infixr 6 default⇒_ +infixr 5 case_⇒_break_ +infixr 4 switch_cases_ + +switch_cases_ : {A : Type } {B : Type ℓ'} (A Bool) ((A Bool) B) B +switch caseIndicator cases caseData = caseData caseIndicator + +case_⇒_break_ : {A : Type } {B : Type ℓ'} A B (otherCases : (A Bool) B) (A Bool) B +case forValue result break otherCases = λ caseIndicator if (caseIndicator forValue) then result else (otherCases caseIndicator) + +default⇒_ : {A : Type } {B : Type ℓ'} B (A Bool) B +default⇒_ value caseIndicator = value diff --git a/src/docs/Cubical.Data.Bool.html b/src/docs/Cubical.Data.Bool.html new file mode 100644 index 0000000..f12edfc --- /dev/null +++ b/src/docs/Cubical.Data.Bool.html @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Bool where + +open import Cubical.Data.Bool.Base public + +open import Cubical.Data.Bool.Properties public diff --git a/src/docs/Cubical.Data.Empty.Base.html b/src/docs/Cubical.Data.Empty.Base.html new file mode 100644 index 0000000..924b91e --- /dev/null +++ b/src/docs/Cubical.Data.Empty.Base.html @@ -0,0 +1,25 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Empty.Base where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ' : Level + +data : Type₀ where + +⊥* : Type +⊥* = Lift + +rec : {A : Type } A +rec () + +rec* : {A : Type } ⊥* { = ℓ'} A +rec* () + +elim : {A : Type } (x : ) A x +elim () + +elim* : {A : ⊥* {ℓ'} Type } (x : ⊥* {ℓ'}) A x +elim* () diff --git a/src/docs/Cubical.Data.Empty.Properties.html b/src/docs/Cubical.Data.Empty.Properties.html new file mode 100644 index 0000000..b982325 --- /dev/null +++ b/src/docs/Cubical.Data.Empty.Properties.html @@ -0,0 +1,37 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Empty.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Empty.Base + +isProp⊥ : isProp +isProp⊥ () + +isProp⊥* : {} isProp {} ⊥* +isProp⊥* _ () + +isContr⊥→A : {} {A : Type } isContr ( A) +fst isContr⊥→A () +snd isContr⊥→A f i () + +isContrΠ⊥ : {} {A : Type } isContr ((x : ) A x) +fst isContrΠ⊥ () +snd isContrΠ⊥ f i () + +isContrΠ⊥* : { ℓ'} {A : ⊥* {} Type ℓ'} isContr ((x : ⊥*) A x) +fst isContrΠ⊥* () +snd isContrΠ⊥* f i () + +uninhabEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + (A ) (B ) A B +uninhabEquiv ¬a ¬b = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun a = rec (¬a a) + isom .inv b = rec (¬b b) + isom .rightInv b = rec (¬b b) + isom .leftInv a = rec (¬a a) diff --git a/src/docs/Cubical.Data.Empty.html b/src/docs/Cubical.Data.Empty.html new file mode 100644 index 0000000..c10f7e0 --- /dev/null +++ b/src/docs/Cubical.Data.Empty.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Empty where + +open import Cubical.Data.Empty.Base public +open import Cubical.Data.Empty.Properties public diff --git a/src/docs/Cubical.Data.FinData.Base.html b/src/docs/Cubical.Data.FinData.Base.html new file mode 100644 index 0000000..45a0258 --- /dev/null +++ b/src/docs/Cubical.Data.FinData.Base.html @@ -0,0 +1,86 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.FinData.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +import Cubical.Data.Empty as +open import Cubical.Data.Nat using (; zero; suc; _+_; _·_; +-assoc) +open import Cubical.Data.Bool.Base +open import Cubical.Relation.Nullary + +private + variable + : Level + A B : Type + +data Fin : Type₀ where + zero : {n : } Fin (suc n) + suc : {n : } (i : Fin n) Fin (suc n) + +-- useful patterns +pattern one = suc zero +pattern two = suc one +pattern three = suc two +pattern four = suc three +pattern five = suc four + +toℕ : {n} Fin n +toℕ zero = 0 +toℕ (suc i) = suc (toℕ i) + +fromℕ : (n : ) Fin (suc n) +fromℕ zero = zero +fromℕ (suc n) = suc (fromℕ n) + +toFromId : (n : ) toℕ (fromℕ n) n +toFromId zero = refl +toFromId (suc n) = cong suc (toFromId n) + +¬Fin0 : ¬ Fin 0 +¬Fin0 () + +_==_ : {n} Fin n Fin n Bool +zero == zero = true +zero == suc _ = false +suc _ == zero = false +suc m == suc n = m == n + +weakenFin : {n : } Fin n Fin (suc n) +weakenFin zero = zero +weakenFin (suc i) = suc (weakenFin i) + +predFin : {n : } Fin (suc (suc n)) Fin (suc n) +predFin zero = zero +predFin (suc x) = x + +foldrFin : {n} (A B B) B (Fin n A) B +foldrFin {n = zero} _ b _ = b +foldrFin {n = suc n} f b l = f (l zero) (foldrFin f b (l suc)) + +elim + : ∀(P : ∀{k} Fin k Type ) + (∀{k} P {suc k} zero) + (∀{k} {fn : Fin k} P fn P (suc fn)) + {k : } (fn : Fin k) P fn + +elim P fz fs {zero} = ⊥.rec ¬Fin0 +elim P fz fs {suc k} zero = fz +elim P fz fs {suc k} (suc fj) = fs (elim P fz fs fj) + + +rec : ∀{k} (a0 aS : A) Fin k A +rec a0 aS zero = a0 +rec a0 aS (suc x) = aS + +FinVec : (A : Type ) (n : ) Type +FinVec A n = Fin n A + +replicateFinVec : (n : ) A FinVec A n +replicateFinVec _ a _ = a + + +_++Fin_ : {n m : } FinVec A n FinVec A m FinVec A (n + m) +_++Fin_ {n = zero} _ W i = W i +_++Fin_ {n = suc n} V _ zero = V zero +_++Fin_ {n = suc n} V W (suc i) = ((V suc) ++Fin W) i diff --git a/src/docs/Cubical.Data.FinData.Properties.html b/src/docs/Cubical.Data.FinData.Properties.html new file mode 100644 index 0000000..f02ebeb --- /dev/null +++ b/src/docs/Cubical.Data.FinData.Properties.html @@ -0,0 +1,373 @@ + +{-# OPTIONS --safe #-} +module Cubical.Data.FinData.Properties where + +-- WARNING : fromℕ' is in triple ! => to clean ! +-- sort file + mix with Fin folder + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.Sum +open import Cubical.Data.Sigma +open import Cubical.Data.FinData.Base as Fin +open import Cubical.Data.Nat renaming (zero to ℕzero ; suc to ℕsuc + ;znots to ℕznots ; snotz to ℕsnotz) +open import Cubical.Data.Nat.Order +open import Cubical.Data.Empty as +open import Cubical.Data.Maybe + +open import Cubical.Relation.Nullary + +open import Cubical.Structures.Pointed + +private + variable + ℓ' : Level + A : Type + m n k : + +toℕ<n : {n} (i : Fin n) toℕ i < n +toℕ<n {n = ℕsuc n} zero = n , +-comm n 1 +toℕ<n {n = ℕsuc n} (suc i) = toℕ<n i .fst , +-suc _ _ cong ℕsuc (toℕ<n i .snd) + +znots : ∀{k} {m : Fin k} ¬ (zero (suc m)) +znots {k} {m} x = subst (Fin.rec (Fin k) ) x m + +znotsP : {k0 k1 : } {k : k0 k1} {m1 : Fin k1} + ¬ PathP i Fin (ℕsuc (k i))) zero (suc m1) +znotsP p = ℕznots (congP i toℕ) p) + +snotz : ∀{k} {m : Fin k} ¬ ((suc m) zero) +snotz {k} {m} x = subst (Fin.rec (Fin k)) x m + +snotzP : {k0 k1 : } {k : k0 k1} {m0 : Fin k0} + ¬ PathP i Fin (ℕsuc (k i))) (suc m0) zero +snotzP p = ℕsnotz (congP i toℕ) p) + +-- alternative from +fromℕ' : (n : ) (k : ) (k < n) Fin n +fromℕ' ℕzero k infkn = ⊥.rec (¬-<-zero infkn) +fromℕ' (ℕsuc n) ℕzero infkn = zero +fromℕ' (ℕsuc n) (ℕsuc k) infkn = suc (fromℕ' n k (pred-≤-pred infkn)) + +toFromId' : (n : ) (k : ) (infkn : k < n) toℕ (fromℕ' n k infkn) k +toFromId' ℕzero k infkn = ⊥.rec (¬-<-zero infkn) +toFromId' (ℕsuc n) ℕzero infkn = refl +toFromId' (ℕsuc n) (ℕsuc k) infkn = cong ℕsuc (toFromId' n k (pred-≤-pred infkn)) + +fromToId' : (n : ) (k : Fin n ) (r : toℕ k < n) fromℕ' n (toℕ k) r k +fromToId' (ℕsuc n) zero r = refl +fromToId' (ℕsuc n) (suc k) r = cong suc (fromToId' n k (pred-≤-pred r)) + +inj-toℕ : {n : } {k l : Fin n} (toℕ k toℕ l) k l +inj-toℕ {ℕsuc n} {zero} {zero} x = refl +inj-toℕ {ℕsuc n} {zero} {suc l} x = ⊥.rec (ℕznots x) +inj-toℕ {ℕsuc n} {suc k} {zero} x = ⊥.rec (ℕsnotz x) +inj-toℕ {ℕsuc n} {suc k} {suc l} x = cong suc (inj-toℕ (injSuc x)) + +inj-cong : {n : } {k l : Fin n} (p : toℕ k toℕ l) cong toℕ (inj-toℕ p) p +inj-cong p = isSetℕ _ _ _ _ + +isPropFin0 : isProp (Fin 0) +isPropFin0 = ⊥.rec ¬Fin0 + +isContrFin1 : isContr (Fin 1) +isContrFin1 .fst = zero +isContrFin1 .snd zero = refl + +injSucFin : {n} {p q : Fin n} suc p suc q p q +injSucFin {ℕsuc ℕzero} {zero} {zero} pf = refl +injSucFin {ℕsuc (ℕsuc n)} pf = cong predFin pf + +injSucFinP : {n0 n1 : } {pn : n0 n1} {p0 : Fin n0} {p1 : Fin n1} + PathP i Fin (ℕsuc (pn i))) (suc p0) (suc p1) + PathP i Fin (pn i)) p0 p1 +injSucFinP {one} {one} {pn} {zero} {zero} sucp = + transport j PathP i Fin (eqn j i)) zero zero) refl + where eqn : refl pn + eqn = isSetℕ one one refl pn +injSucFinP {one} {ℕsuc (ℕsuc n1)} {pn} {p0} {p1} sucp = ⊥.rec (ℕznots (injSuc pn)) +injSucFinP {ℕsuc (ℕsuc n0)} {one} {pn} {p0} {p1} sucp = ⊥.rec (ℕsnotz (injSuc pn)) +injSucFinP {ℕsuc (ℕsuc n0)} {ℕsuc (ℕsuc n1)} {pn} {p0} {p1} sucp = + transport j PathP i Fin (eqn j i)) p0 p1) ( + congP i predFin) ( + transport j PathP i Fin (ℕsuc (eqn (~ j) i))) (suc p0) (suc p1)) sucp + ) + ) + where pn' : 2 + n0 2 + n1 + pn' = cong ℕsuc (injSuc pn) + eqn : pn' pn + eqn = isSetℕ (2 + n0) (2 + n1) pn' pn + +discreteFin : ∀{k} Discrete (Fin k) +discreteFin zero zero = yes refl +discreteFin zero (suc y) = no znots +discreteFin (suc x) zero = no snotz +discreteFin (suc x) (suc y) with discreteFin x y +... | yes p = yes (cong suc p) +... | no ¬p = no q ¬p (injSucFin q)) + +isSetFin : ∀{k} isSet (Fin k) +isSetFin = Discrete→isSet discreteFin + +isWeaken? : {n} (p : Fin (ℕsuc n)) Dec (Σ[ q Fin n ] p weakenFin q) +isWeaken? {ℕzero} zero = no λ (q , eqn) case q of λ () +isWeaken? {ℕsuc n} zero = yes (zero , refl) +isWeaken? {ℕsuc n} (suc p) with isWeaken? {n} p +... | yes (q , p≡wq) = yes (suc q , cong suc p≡wq) +... | no p≢wq = no λ + { (zero , sp≡wq) snotz sp≡wq + ; (suc q , sp≡wq) p≢wq (q , cong predFin sp≡wq) + } + +data biEq {n : } (i j : Fin n) : Type where + eq : i j biEq i j + ¬eq : ¬ i j biEq i j + +data triEq {n : } (i j a : Fin n) : Type where + leq : a i triEq i j a + req : a j triEq i j a + ¬eq : (¬ a i) × (¬ a j) triEq i j a + +biEq? : (i j : Fin n) biEq i j +biEq? i j = case (discreteFin i j) return _ biEq i j) + of λ { (yes p) eq p ; (no ¬p) ¬eq ¬p } + +triEq? : (i j a : Fin n) triEq i j a +triEq? i j a = + case (discreteFin a i) return _ triEq i j a) of + λ { (yes p) leq p + ; (no ¬p) + case (discreteFin a j) return _ triEq i j a) of + λ { (yes q) req q + ; (no ¬q) ¬eq (¬p , ¬q) }} + + +weakenRespToℕ : {n} (i : Fin n) toℕ (weakenFin i) toℕ i +weakenRespToℕ zero = refl +weakenRespToℕ (suc i) = cong ℕsuc (weakenRespToℕ i) + +toFin : {n : } (m : ) m < n Fin n +toFin {n = ℕzero} _ m<0 = ⊥.rec (¬-<-zero m<0) +toFin {n = ℕsuc n} _ (ℕzero , _) = fromℕ n --in this case we have m≡n +toFin {n = ℕsuc n} m (ℕsuc k , p) = weakenFin (toFin m (k , cong predℕ p)) + +toFin0≡0 : {n : } (p : 0 < ℕsuc n) toFin 0 p zero +toFin0≡0 (ℕzero , p) = subst x fromℕ x zero) (cong predℕ p) refl +toFin0≡0 {ℕzero} (ℕsuc k , p) = ⊥.rec (ℕsnotz (+-comm 1 k (cong predℕ p))) +toFin0≡0 {ℕsuc n} (ℕsuc k , p) = + subst x weakenFin x zero) (sym (toFin0≡0 (k , cong predℕ p))) refl + +genδ-FinVec : (n k : ) (a b : A) FinVec A n +genδ-FinVec (ℕsuc n) ℕzero a b zero = a +genδ-FinVec (ℕsuc n) ℕzero a b (suc x) = b +genδ-FinVec (ℕsuc n) (ℕsuc k) a b zero = b +genδ-FinVec (ℕsuc n) (ℕsuc k) a b (suc x) = genδ-FinVec n k a b x + +δℕ-FinVec : (n k : ) FinVec n +δℕ-FinVec n k = genδ-FinVec n k 1 0 + +-- WARNING : harder to prove things about +genδ-FinVec' : (n k : ) (a b : A) FinVec A n +genδ-FinVec' n k a b x with discreteℕ (toℕ x) k +... | yes p = a +... | no ¬p = b + +-- doing induction on toFin is awkward, so the following alternative +enum : (m : ) m < n Fin n +enum {n = ℕzero} _ m<0 = ⊥.rec (¬-<-zero m<0) +enum {n = ℕsuc n} 0 _ = zero +enum {n = ℕsuc n} (ℕsuc m) p = suc (enum m (pred-≤-pred p)) + +enum∘toℕ : (i : Fin n)(p : toℕ i < n) enum (toℕ i) p i +enum∘toℕ {n = ℕsuc n} zero _ = refl +enum∘toℕ {n = ℕsuc n} (suc i) p t = suc (enum∘toℕ i (pred-≤-pred p) t) + +toℕ∘enum : (m : )(p : m < n) toℕ (enum m p) m +toℕ∘enum {n = ℕzero} _ m<0 = ⊥.rec (¬-<-zero m<0) +toℕ∘enum {n = ℕsuc n} 0 _ = refl +toℕ∘enum {n = ℕsuc n} (ℕsuc m) p i = ℕsuc (toℕ∘enum m (pred-≤-pred p) i) + +enumExt : {m m' : }(p : m < n)(p' : m' < n) m m' enum m p enum m' p' +enumExt p p' q i = enum (q i) (isProp→PathP i isProp≤ {m = ℕsuc (q i)}) p p' i) + +enumInj : (p : m < k) (q : n < k) enum m p enum n q m n +enumInj p q path = sym (toℕ∘enum _ p) cong toℕ path toℕ∘enum _ q + +enumIndStep : + (P : Fin n Type ) + (k : )(p : ℕsuc k < n) + ((m : )(q : m < n)(q' : m k ) P (enum m q)) + P (enum (ℕsuc k) p) + ((m : )(q : m < n)(q' : m ℕsuc k) P (enum m q)) +enumIndStep P k p f x m q q' = + case (≤-split q') return _ P (enum m q)) of + λ { (inl r') f m q (pred-≤-pred r') + ; (inr r') subst P (enumExt p q (sym r')) x } + +enumElim : + (P : Fin n Type ) + (k : )(p : k < n)(h : ℕsuc k n) + ((m : )(q : m < n)(q' : m k) P (enum m q)) + (i : Fin n) P i +enumElim P k p h f i = + subst P (enum∘toℕ i (toℕ<n i)) (f (toℕ i) (toℕ<n i) + (pred-≤-pred (subst a toℕ i < a) (sym h) (toℕ<n i)))) + + +++FinAssoc : {n m k : } (U : FinVec A n) (V : FinVec A m) (W : FinVec A k) + PathP i FinVec A (+-assoc n m k i)) (U ++Fin (V ++Fin W)) ((U ++Fin V) ++Fin W) +++FinAssoc {n = ℕzero} _ _ _ = refl +++FinAssoc {n = ℕsuc n} U V W i zero = U zero +++FinAssoc {n = ℕsuc n} U V W i (suc ind) = ++FinAssoc (U suc) V W i ind + +++FinRid : {n : } (U : FinVec A n) (V : FinVec A 0) + PathP i FinVec A (+-zero n i)) (U ++Fin V) U +++FinRid {n = ℕzero} U V = funExt λ i ⊥.rec (¬Fin0 i) +++FinRid {n = ℕsuc n} U V i zero = U zero +++FinRid {n = ℕsuc n} U V i (suc ind) = ++FinRid (U suc) V i ind + +++FinElim : {P : A Type ℓ'} {n m : } (U : FinVec A n) (V : FinVec A m) + (∀ i P (U i)) (∀ i P (V i)) i P ((U ++Fin V) i) +++FinElim {n = ℕzero} _ _ _ PVHyp i = PVHyp i +++FinElim {n = ℕsuc n} _ _ PUHyp _ zero = PUHyp zero +++FinElim {P = P} {n = ℕsuc n} U V PUHyp PVHyp (suc i) = + ++FinElim {P = P} (U suc) V i PUHyp (suc i)) PVHyp i + +++FinPres∈ : {n m : } {α : FinVec A n} {β : FinVec A m} (S : A) + (∀ i α i S) (∀ i β i S) i (α ++Fin β) i S +++FinPres∈ {n = ℕzero} S i = i +++FinPres∈ {n = ℕsuc n} S zero = zero +++FinPres∈ {n = ℕsuc n} S (suc i) = ++FinPres∈ S ( suc) i + +-- sends i to n+i if toℕ i < m and to i∸n otherwise +-- then +Shuffle²≡id and over the induced path (i.e. in PathP (ua +ShuffleEquiv)) +-- ++Fin is commutative, but how to go from there? ++Shuffle : (m n : ) Fin (m + n) Fin (n + m) ++Shuffle m n i with <Dec (toℕ i) m +... | yes i<m = toFin (n + (toℕ i)) (<-k+ i<m) +... | no ¬i<m = toFin (toℕ i m) + (subst x toℕ i m < x) (+-comm m n) (≤<-trans (∸-≤ (toℕ i) m) (toℕ<n i))) + + +finSucMaybeIso : Iso (Fin (ℕ.suc n)) (Maybe (Fin n)) +Iso.fun finSucMaybeIso zero = nothing +Iso.fun finSucMaybeIso (suc i) = just i +Iso.inv finSucMaybeIso nothing = zero +Iso.inv finSucMaybeIso (just i) = suc i +Iso.rightInv finSucMaybeIso nothing = refl +Iso.rightInv finSucMaybeIso (just i) = refl +Iso.leftInv finSucMaybeIso zero = refl +Iso.leftInv finSucMaybeIso (suc i) = refl + +finSuc≡Maybe : Fin (ℕ.suc n) Maybe (Fin n) +finSuc≡Maybe = isoToPath finSucMaybeIso + +finSuc≡Maybe∙ : (Fin (ℕ.suc n) , zero) Maybe∙ (Fin n) +finSuc≡Maybe∙ = pointed-sip _ _ ((isoToEquiv finSucMaybeIso) , refl) + +-- Proof that Fin n ⊎ Fin m ≃ Fin (n+m) +module FinSumChar where + + fun : (n m : ) Fin n Fin m Fin (n + m) + fun ℕzero m (inr i) = i + fun (ℕsuc n) m (inl zero) = zero + fun (ℕsuc n) m (inl (suc i)) = suc (fun n m (inl i)) + fun (ℕsuc n) m (inr i) = suc (fun n m (inr i)) + + invSucAux : (n m : ) Fin n Fin m Fin (ℕsuc n) Fin m + invSucAux n m (inl i) = inl (suc i) + invSucAux n m (inr i) = inr i + + inv : (n m : ) Fin (n + m) Fin n Fin m + inv ℕzero m i = inr i + inv (ℕsuc n) m zero = inl zero + inv (ℕsuc n) m (suc i) = invSucAux n m (inv n m i) + + ret : (n m : ) (i : Fin n Fin m) inv n m (fun n m i) i + ret ℕzero m (inr i) = refl + ret (ℕsuc n) m (inl zero) = refl + ret (ℕsuc n) m (inl (suc i)) = subst x invSucAux n m x inl (suc i)) + (sym (ret n m (inl i))) refl + ret (ℕsuc n) m (inr i) = subst x invSucAux n m x inr i) (sym (ret n m (inr i))) refl + + sec : (n m : ) (i : Fin (n + m)) fun n m (inv n m i) i + sec ℕzero m i = refl + sec (ℕsuc n) m zero = refl + sec (ℕsuc n) m (suc i) = helperPath (inv n m i) cong suc (sec n m i) + where + helperPath : x fun (ℕsuc n) m (invSucAux n m x) suc (fun n m x) + helperPath (inl _) = refl + helperPath (inr _) = refl + + Equiv : (n m : ) Fin n Fin m Fin (n + m) + Equiv n m = isoToEquiv (iso (fun n m) (inv n m) (sec n m) (ret n m)) + + ++FinInl : (n m : ) (U : FinVec A n) (W : FinVec A m) (i : Fin n) + U i (U ++Fin W) (fun n m (inl i)) + ++FinInl (ℕsuc n) m U W zero = refl + ++FinInl (ℕsuc n) m U W (suc i) = ++FinInl n m (U suc) W i + + ++FinInr : (n m : ) (U : FinVec A n) (W : FinVec A m) (i : Fin m) + W i (U ++Fin W) (fun n m (inr i)) + ++FinInr ℕzero (ℕsuc m) U W i = refl + ++FinInr (ℕsuc n) m U W i = ++FinInr n m (U suc) W i + +-- Proof that Fin n × Fin m ≃ Fin nm +module FinProdChar where + + open Iso + sucProdToSumIso : (n m : ) Iso (Fin (ℕsuc n) × Fin m) (Fin m (Fin n × Fin m)) + fun (sucProdToSumIso n m) (zero , j) = inl j + fun (sucProdToSumIso n m) (suc i , j) = inr (i , j) + inv (sucProdToSumIso n m) (inl j) = zero , j + inv (sucProdToSumIso n m) (inr (i , j)) = suc i , j + rightInv (sucProdToSumIso n m) (inl j) = refl + rightInv (sucProdToSumIso n m) (inr (i , j)) = refl + leftInv (sucProdToSumIso n m) (zero , j) = refl + leftInv (sucProdToSumIso n m) (suc i , j) = refl + + Equiv : (n m : ) (Fin n × Fin m) Fin (n · m) + Equiv ℕzero m = uninhabEquiv x ¬Fin0 (fst x)) ¬Fin0 + Equiv (ℕsuc n) m = Fin (ℕsuc n) × Fin m ≃⟨ isoToEquiv (sucProdToSumIso n m) + Fin m (Fin n × Fin m) ≃⟨ isoToEquiv (⊎Iso idIso (equivToIso (Equiv n m))) + Fin m Fin (n · m) ≃⟨ FinSumChar.Equiv m (n · m) + Fin (m + n · m) + +-- Exhaustion of decidable predicate + +∀Dec : + (P : Fin m Type ) + (dec : (i : Fin m) Dec (P i)) + ((i : Fin m) P i) (Σ[ i Fin m ] ¬ P i) +∀Dec {m = 0} _ _ = inl λ () +∀Dec {m = ℕsuc m} P dec = helper (dec zero) (∀Dec _ (dec suc)) + where + helper : + Dec (P zero) + ((i : Fin m) P (suc i)) (Σ[ i Fin m ] ¬ P (suc i)) + ((i : Fin (ℕsuc m)) P i) (Σ[ i Fin (ℕsuc m) ] ¬ P i) + helper (yes p) (inl q) = inl λ { zero p ; (suc i) q i } + helper (yes _) (inr q) = inr (suc (q .fst) , q .snd) + helper (no ¬p) _ = inr (zero , ¬p) + +∀Dec2 : + (P : Fin m Fin n Type ) + (dec : (i : Fin m)(j : Fin n) Dec (P i j)) + ((i : Fin m)(j : Fin n) P i j) (Σ[ i Fin m ] Σ[ j Fin n ] ¬ P i j) +∀Dec2 {m = 0} {n = n} _ _ = inl λ () +∀Dec2 {m = ℕsuc m} {n = n} P dec = helper (∀Dec (P zero) (dec zero)) (∀Dec2 (P suc) (dec suc)) + where + helper : + ((j : Fin n) P zero j) (Σ[ j Fin n ] ¬ P zero j) + ((i : Fin m)(j : Fin n) P (suc i) j) (Σ[ i Fin m ] Σ[ j Fin n ] ¬ P (suc i) j) + ((i : Fin (ℕsuc m))(j : Fin n) P i j) (Σ[ i Fin (ℕsuc m) ] Σ[ j Fin n ] ¬ P i j) + helper (inl p) (inl q) = inl λ { zero j p j ; (suc i) j q i j } + helper (inl _) (inr q) = inr (suc (q .fst) , q .snd .fst , q .snd .snd) + helper (inr p) _ = inr (zero , p) diff --git a/src/docs/Cubical.Data.FinData.html b/src/docs/Cubical.Data.FinData.html new file mode 100644 index 0000000..27b0da8 --- /dev/null +++ b/src/docs/Cubical.Data.FinData.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.FinData where + +open import Cubical.Data.FinData.Base public +open import Cubical.Data.FinData.Properties public diff --git a/src/docs/Cubical.Data.List.Base.html b/src/docs/Cubical.Data.List.Base.html new file mode 100644 index 0000000..3ede1c5 --- /dev/null +++ b/src/docs/Cubical.Data.List.Base.html @@ -0,0 +1,52 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.List.Base where + +open import Agda.Builtin.List public +open import Cubical.Core.Everything +open import Cubical.Data.Maybe.Base as Maybe +open import Cubical.Data.Nat.Base + +module _ {} {A : Type } where + + infixr 5 _++_ + infixl 5 _∷ʳ_ + + [_] : A List A + [ a ] = a [] + + _++_ : List A List A List A + [] ++ ys = ys + (x xs) ++ ys = x xs ++ ys + + rev : List A List A + rev [] = [] + rev (x xs) = rev xs ++ [ x ] + + _∷ʳ_ : List A A List A + xs ∷ʳ x = xs ++ x [] + + length : List A + length [] = 0 + length (x l) = 1 + length l + + map : {ℓ'} {B : Type ℓ'} (A B) List A List B + map f [] = [] + map f (x xs) = f x map f xs + + map2 : {ℓ' ℓ''} {B : Type ℓ'} {C : Type ℓ''} + (A B C) List A List B List C + map2 f [] _ = [] + map2 f _ [] = [] + map2 f (x xs) (y ys) = f x y map2 f xs ys + + filterMap : {ℓ'} {B : Type ℓ'} (A Maybe B) List A List B + filterMap f [] = [] + filterMap f (x xs) = Maybe.rec (filterMap f xs) (_∷ filterMap f xs) (f x) + + foldr : {ℓ'} {B : Type ℓ'} (A B B) B List A B + foldr f b [] = b + foldr f b (x xs) = f x (foldr f b xs) + + foldl : {ℓ'} {B : Type ℓ'} (B A B) B List A B + foldl f b [] = b + foldl f b (x xs) = foldl f (f b x) xs diff --git a/src/docs/Cubical.Data.List.Properties.html b/src/docs/Cubical.Data.List.Properties.html new file mode 100644 index 0000000..cd224c4 --- /dev/null +++ b/src/docs/Cubical.Data.List.Properties.html @@ -0,0 +1,181 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.List.Properties where + +open import Agda.Builtin.List +open import Cubical.Core.Everything +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Prelude +open import Cubical.Data.Empty as +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Unit +open import Cubical.Relation.Nullary + +open import Cubical.Data.List.Base + +module _ {} {A : Type } where + + ++-unit-r : (xs : List A) xs ++ [] xs + ++-unit-r [] = refl + ++-unit-r (x xs) = cong (_∷_ x) (++-unit-r xs) + + ++-assoc : (xs ys zs : List A) (xs ++ ys) ++ zs xs ++ ys ++ zs + ++-assoc [] ys zs = refl + ++-assoc (x xs) ys zs = cong (_∷_ x) (++-assoc xs ys zs) + + rev-snoc : (xs : List A) (y : A) rev (xs ++ [ y ]) y rev xs + rev-snoc [] y = refl + rev-snoc (x xs) y = cong (_++ [ x ]) (rev-snoc xs y) + + rev-++ : (xs ys : List A) rev (xs ++ ys) rev ys ++ rev xs + rev-++ [] ys = sym (++-unit-r (rev ys)) + rev-++ (x xs) ys = + cong zs zs ++ [ x ]) (rev-++ xs ys) + ++-assoc (rev ys) (rev xs) [ x ] + + rev-rev : (xs : List A) rev (rev xs) xs + rev-rev [] = refl + rev-rev (x xs) = rev-snoc (rev xs) x cong (_∷_ x) (rev-rev xs) + + rev-rev-snoc : (xs : List A) (y : A) + Square (rev-rev (xs ++ [ y ])) (cong (_++ [ y ]) (rev-rev xs)) (cong rev (rev-snoc xs y)) refl + rev-rev-snoc [] y = sym (lUnit refl) + rev-rev-snoc (x xs) y i j = + hcomp + k λ + { (i = i1) compPath-filler (rev-snoc (rev xs) x) (cong (x ∷_) (rev-rev xs)) k j ++ [ y ] + ; (j = i0) rev (rev-snoc xs y i ++ [ x ]) + ; (j = i1) x rev-rev-snoc xs y i k + }) + (rev-snoc (rev-snoc xs y i) x j) + + data SnocView : List A Type where + nil : SnocView [] + snoc : (x : A) (xs : List A) (sx : SnocView xs) SnocView (xs ∷ʳ x) + + snocView : (xs : List A) SnocView xs + snocView xs = helper nil xs + where + helper : {l : List A} -> SnocView l -> (r : List A) -> SnocView (l ++ r) + helper {l} sl [] = subst SnocView (sym (++-unit-r l)) sl + helper {l} sl (x r) = subst SnocView (++-assoc l (x []) r) (helper (snoc x l sl) r) + +-- Path space of list type +module ListPath {} {A : Type } where + + Cover : List A List A Type + Cover [] [] = Lift Unit + Cover [] (_ _) = Lift + Cover (_ _) [] = Lift + Cover (x xs) (y ys) = (x y) × Cover xs ys + + reflCode : xs Cover xs xs + reflCode [] = lift tt + reflCode (_ xs) = refl , reflCode xs + + encode : xs ys (p : xs ys) Cover xs ys + encode xs _ = J ys _ Cover xs ys) (reflCode xs) + + encodeRefl : xs encode xs xs refl reflCode xs + encodeRefl xs = JRefl ys _ Cover xs ys) (reflCode xs) + + decode : xs ys Cover xs ys xs ys + decode [] [] _ = refl + decode [] (_ _) (lift ()) + decode (x xs) [] (lift ()) + decode (x xs) (y ys) (p , c) = cong₂ _∷_ p (decode xs ys c) + + decodeRefl : xs decode xs xs (reflCode xs) refl + decodeRefl [] = refl + decodeRefl (x xs) = cong (cong₂ _∷_ refl) (decodeRefl xs) + + decodeEncode : xs ys (p : xs ys) decode xs ys (encode xs ys p) p + decodeEncode xs _ = + J ys p decode xs ys (encode xs ys p) p) + (cong (decode xs xs) (encodeRefl xs) decodeRefl xs) + + isOfHLevelCover : (n : HLevel) (p : isOfHLevel (suc (suc n)) A) + (xs ys : List A) isOfHLevel (suc n) (Cover xs ys) + isOfHLevelCover n p [] [] = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isPropUnit) + isOfHLevelCover n p [] (y ys) = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (x xs) [] = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (x xs) (y ys) = + isOfHLevelΣ (suc n) (p x y) (\ _ isOfHLevelCover n p xs ys) + +isOfHLevelList : {} (n : HLevel) {A : Type } + isOfHLevel (suc (suc n)) A isOfHLevel (suc (suc n)) (List A) +isOfHLevelList n ofLevel xs ys = + isOfHLevelRetract (suc n) + (ListPath.encode xs ys) + (ListPath.decode xs ys) + (ListPath.decodeEncode xs ys) + (ListPath.isOfHLevelCover n ofLevel xs ys) + +private + variable + : Level + A : Type + + caseList : { ℓ'} {A : Type } {B : Type ℓ'} (n c : B) List A B + caseList n _ [] = n + caseList _ c (_ _) = c + + safe-head : A List A A + safe-head x [] = x + safe-head _ (x _) = x + + safe-tail : List A List A + safe-tail [] = [] + safe-tail (_ xs) = xs + +cons-inj₁ : {x y : A} {xs ys} x xs y ys x y +cons-inj₁ {x = x} p = cong (safe-head x) p + +cons-inj₂ : {x y : A} {xs ys} x xs y ys xs ys +cons-inj₂ = cong safe-tail + +¬cons≡nil : {x : A} {xs} ¬ (x xs []) +¬cons≡nil {A = A} p = lower (subst (caseList (Lift ) (List A)) p []) + +¬nil≡cons : {x : A} {xs} ¬ ([] x xs) +¬nil≡cons {A = A} p = lower (subst (caseList (List A) (Lift )) p []) + +¬snoc≡nil : {x : A} {xs} ¬ (xs ∷ʳ x []) +¬snoc≡nil {xs = []} contra = ¬cons≡nil contra +¬snoc≡nil {xs = x xs} contra = ¬cons≡nil contra + +¬nil≡snoc : {x : A} {xs} ¬ ([] xs ∷ʳ x) +¬nil≡snoc contra = ¬snoc≡nil (sym contra) + +cons≡rev-snoc : (x : A) (xs : List A) x rev xs rev (xs ∷ʳ x) +cons≡rev-snoc _ [] = refl +cons≡rev-snoc x (y ys) = λ i cons≡rev-snoc x ys i ++ y [] + +isContr[]≡[] : isContr (Path (List A) [] []) +isContr[]≡[] = refl , ListPath.decodeEncode [] [] + +isPropXs≡[] : {xs : List A} isProp (xs []) +isPropXs≡[] {xs = []} = isOfHLevelSuc 0 isContr[]≡[] +isPropXs≡[] {xs = x xs} = λ p _ ⊥.rec (¬cons≡nil p) + +discreteList : Discrete A Discrete (List A) +discreteList eqA [] [] = yes refl +discreteList eqA [] (y ys) = no ¬nil≡cons +discreteList eqA (x xs) [] = no ¬cons≡nil +discreteList eqA (x xs) (y ys) with eqA x y | discreteList eqA xs ys +... | yes p | yes q = yes i p i q i) +... | yes _ | no ¬q = no p ¬q (cons-inj₂ p)) +... | no ¬p | _ = no q ¬p (cons-inj₁ q)) + +foldrCons : (xs : List A) foldr _∷_ [] xs xs +foldrCons [] = refl +foldrCons (x xs) = cong (x ∷_) (foldrCons xs) + +length-map : {ℓA ℓB} {A : Type ℓA} {B : Type ℓB} (f : A B) (as : List A) + length (map f as) length as +length-map f [] = refl +length-map f (a as) = cong suc (length-map f as) diff --git a/src/docs/Cubical.Data.List.html b/src/docs/Cubical.Data.List.html new file mode 100644 index 0000000..408383f --- /dev/null +++ b/src/docs/Cubical.Data.List.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.List where + +open import Cubical.Data.List.Base public +open import Cubical.Data.List.Properties public diff --git a/src/docs/Cubical.Data.Maybe.Base.html b/src/docs/Cubical.Data.Maybe.Base.html new file mode 100644 index 0000000..449d0ed --- /dev/null +++ b/src/docs/Cubical.Data.Maybe.Base.html @@ -0,0 +1,30 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Maybe.Base where + +open import Cubical.Core.Everything + +private + variable + ℓA ℓB : Level + A : Type ℓA + B : Type ℓB + +data Maybe (A : Type ) : Type where + nothing : Maybe A + just : A Maybe A + +caseMaybe : (n j : B) Maybe A B +caseMaybe n _ nothing = n +caseMaybe _ j (just _) = j + +map-Maybe : (A B) Maybe A Maybe B +map-Maybe _ nothing = nothing +map-Maybe f (just x) = just (f x) + +rec : B (A B) Maybe A B +rec n j nothing = n +rec n j (just a) = j a + +elim : {A : Type ℓA} (B : Maybe A Type ℓB) B nothing ((x : A) B (just x)) (mx : Maybe A) B mx +elim B n j nothing = n +elim B n j (just a) = j a diff --git a/src/docs/Cubical.Data.Maybe.Properties.html b/src/docs/Cubical.Data.Maybe.Properties.html new file mode 100644 index 0000000..f2a5091 --- /dev/null +++ b/src/docs/Cubical.Data.Maybe.Properties.html @@ -0,0 +1,181 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Maybe.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function using (_∘_; idfun) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Pointed.Base using (Pointed; _→∙_; pt) +open import Cubical.Foundations.Structure using (⟨_⟩) + +open import Cubical.Functions.Embedding using (isEmbedding) + +open import Cubical.Data.Empty as using (; isProp⊥) +open import Cubical.Data.Unit +open import Cubical.Data.Nat using (suc) +open import Cubical.Data.Sum using (_⊎_; inl; inr) +open import Cubical.Data.Sigma using (ΣPathP) + +open import Cubical.Relation.Nullary using (¬_; Discrete; yes; no) + +open import Cubical.Data.Maybe.Base as Maybe + +Maybe∙ : {} (A : Type ) Pointed +Maybe∙ A .fst = Maybe A +Maybe∙ A .snd = nothing + +-- Maybe∙ is the "free pointing" functor, that is, left adjoint to the +-- forgetful functor forgetting the base point. +module _ {} (A : Type ) {ℓ'} (B : Pointed ℓ') where + + freelyPointedIso : Iso (Maybe∙ A →∙ B) (A B ) + Iso.fun freelyPointedIso f∙ = fst f∙ just + Iso.inv freelyPointedIso f = Maybe.rec (pt B) f , refl + Iso.rightInv freelyPointedIso f = refl + Iso.leftInv freelyPointedIso f∙ = + ΣPathP + ( funExt (Maybe.elim _ (sym (snd f∙)) a refl)) + , λ i j snd f∙ (~ i j)) + +map-Maybe-id : {} {A : Type } m map-Maybe (idfun A) m m +map-Maybe-id nothing = refl +map-Maybe-id (just _) = refl + +-- Path space of Maybe type +module MaybePath {} {A : Type } where + Cover : Maybe A Maybe A Type + Cover nothing nothing = Lift Unit + Cover nothing (just _) = Lift + Cover (just _) nothing = Lift + Cover (just a) (just a') = a a' + + reflCode : (c : Maybe A) Cover c c + reflCode nothing = lift tt + reflCode (just b) = refl + + encode : c c' c c' Cover c c' + encode c _ = J c' _ Cover c c') (reflCode c) + + encodeRefl : c encode c c refl reflCode c + encodeRefl c = JRefl c' _ Cover c c') (reflCode c) + + decode : c c' Cover c c' c c' + decode nothing nothing _ = refl + decode (just _) (just _) p = cong just p + + decodeRefl : c decode c c (reflCode c) refl + decodeRefl nothing = refl + decodeRefl (just _) = refl + + decodeEncode : c c' (p : c c') decode c c' (encode c c' p) p + decodeEncode c _ = + J c' p decode c c' (encode c c' p) p) + (cong (decode c c) (encodeRefl c) decodeRefl c) + + encodeDecode : c c' (d : Cover c c') encode c c' (decode c c' d) d + encodeDecode nothing nothing _ = refl + encodeDecode (just a) (just a') = + J a' p encode (just a) (just a') (cong just p) p) (encodeRefl (just a)) + + Cover≃Path : c c' Cover c c' (c c') + Cover≃Path c c' = isoToEquiv + (iso (decode c c') (encode c c') (decodeEncode c c') (encodeDecode c c')) + + Cover≡Path : c c' Cover c c' (c c') + Cover≡Path c c' = isoToPath + (iso (decode c c') (encode c c') (decodeEncode c c') (encodeDecode c c')) + + isOfHLevelCover : (n : HLevel) + isOfHLevel (suc (suc n)) A + c c' isOfHLevel (suc n) (Cover c c') + isOfHLevelCover n p nothing nothing = isOfHLevelLift (suc n) (isOfHLevelUnit (suc n)) + isOfHLevelCover n p nothing (just a') = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (just a) nothing = isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p (just a) (just a') = p a a' + +isOfHLevelMaybe : {} (n : HLevel) {A : Type } + isOfHLevel (suc (suc n)) A + isOfHLevel (suc (suc n)) (Maybe A) +isOfHLevelMaybe n lA c c' = + isOfHLevelRetract (suc n) + (MaybePath.encode c c') + (MaybePath.decode c c') + (MaybePath.decodeEncode c c') + (MaybePath.isOfHLevelCover n lA c c') + +private + variable + : Level + A : Type + +fromJust-def : A Maybe A A +fromJust-def a nothing = a +fromJust-def _ (just a) = a + +just-inj : (x y : A) just x just y x y +just-inj x _ eq = cong (fromJust-def x) eq + +isEmbedding-just : isEmbedding (just {A = A}) +isEmbedding-just w z = MaybePath.Cover≃Path (just w) (just z) .snd + +¬nothing≡just : {x : A} ¬ (nothing just x) +¬nothing≡just {A = A} {x = x} p = lower (subst (caseMaybe (Maybe A) (Lift )) p (just x)) + +¬just≡nothing : {x : A} ¬ (just x nothing) +¬just≡nothing {A = A} {x = x} p = lower (subst (caseMaybe (Lift ) (Maybe A)) p (just x)) + +isProp-x≡nothing : (x : Maybe A) isProp (x nothing) +isProp-x≡nothing nothing x w = + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w +isProp-x≡nothing (just _) p _ = ⊥.rec (¬just≡nothing p) + +isProp-nothing≡x : (x : Maybe A) isProp (nothing x) +isProp-nothing≡x nothing x w = + subst isProp (MaybePath.Cover≡Path nothing nothing) (isOfHLevelLift 1 isPropUnit) x w +isProp-nothing≡x (just _) p _ = ⊥.rec (¬nothing≡just p) + +isContr-nothing≡nothing : isContr (nothing {A = A} nothing) +isContr-nothing≡nothing = inhProp→isContr refl (isProp-x≡nothing _) + +discreteMaybe : Discrete A Discrete (Maybe A) +discreteMaybe eqA nothing nothing = yes refl +discreteMaybe eqA nothing (just a') = no ¬nothing≡just +discreteMaybe eqA (just a) nothing = no ¬just≡nothing +discreteMaybe eqA (just a) (just a') with eqA a a' +... | yes p = yes (cong just p) +... | no ¬p = no p ¬p (just-inj _ _ p)) + +module SumUnit where + Maybe→SumUnit : Maybe A Unit A + Maybe→SumUnit nothing = inl tt + Maybe→SumUnit (just a) = inr a + + SumUnit→Maybe : Unit A Maybe A + SumUnit→Maybe (inl _) = nothing + SumUnit→Maybe (inr a) = just a + + Maybe→SumUnit→Maybe : (x : Maybe A) SumUnit→Maybe (Maybe→SumUnit x) x + Maybe→SumUnit→Maybe nothing = refl + Maybe→SumUnit→Maybe (just _) = refl + + SumUnit→Maybe→SumUnit : (x : Unit A) Maybe→SumUnit (SumUnit→Maybe x) x + SumUnit→Maybe→SumUnit (inl _) = refl + SumUnit→Maybe→SumUnit (inr _) = refl + +Maybe≡SumUnit : Maybe A Unit A +Maybe≡SumUnit = isoToPath (iso Maybe→SumUnit SumUnit→Maybe SumUnit→Maybe→SumUnit Maybe→SumUnit→Maybe) + where open SumUnit + +congMaybeEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + A B Maybe A Maybe B +congMaybeEquiv e = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = map-Maybe (equivFun e) + isom .inv = map-Maybe (invEq e) + isom .rightInv nothing = refl + isom .rightInv (just b) = cong just (secEq e b) + isom .leftInv nothing = refl + isom .leftInv (just a) = cong just (retEq e a) diff --git a/src/docs/Cubical.Data.Maybe.html b/src/docs/Cubical.Data.Maybe.html new file mode 100644 index 0000000..c12907f --- /dev/null +++ b/src/docs/Cubical.Data.Maybe.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Maybe where + +open import Cubical.Data.Maybe.Base public +open import Cubical.Data.Maybe.Properties public diff --git a/src/docs/Cubical.Data.Nat.Base.html b/src/docs/Cubical.Data.Nat.Base.html new file mode 100644 index 0000000..fafebe9 --- /dev/null +++ b/src/docs/Cubical.Data.Nat.Base.html @@ -0,0 +1,90 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Data.Nat.Base where + +open import Cubical.Core.Primitives + +open import Agda.Builtin.Nat public + using (zero; suc; _+_) + renaming (Nat to ; _-_ to _∸_; _*_ to _·_) + +open import Cubical.Data.Nat.Literals public +open import Cubical.Data.Bool.Base +open import Cubical.Data.Sum.Base hiding (elim) +open import Cubical.Data.Empty.Base hiding (elim) +open import Cubical.Data.Unit.Base +open import Cubical.Data.Sigma.Base + +predℕ : +predℕ zero = zero +predℕ (suc n) = n + +caseNat : {} {A : Type } (a0 aS : A) A +caseNat a0 aS zero = a0 +caseNat a0 aS (suc n) = aS + +doubleℕ : +doubleℕ zero = zero +doubleℕ (suc x) = suc (suc (doubleℕ x)) + +-- doublesℕ n m = 2^n · m +doublesℕ : +doublesℕ zero m = m +doublesℕ (suc n) m = doublesℕ n (doubleℕ m) + +-- iterate +iter : {} {A : Type } (A A) A A +iter zero f z = z +iter (suc n) f z = f (iter n f z) + +elim : {} {A : Type } + A zero + ((n : ) A n A (suc n)) + (n : ) A n +elim a₀ _ zero = a₀ +elim a₀ f (suc n) = f n (elim a₀ f n) + +elim+2 : {} {A : Type } A 0 A 1 + ((n : ) (A (suc n) A (suc (suc n)))) + (n : ) A n +elim+2 a0 a1 ind zero = a0 +elim+2 a0 a1 ind (suc zero) = a1 +elim+2 {A = A} a0 a1 ind (suc (suc n)) = + ind n (elim+2 {A = A} a0 a1 ind (suc n)) + +isEven isOdd : Bool +isEven zero = true +isEven (suc n) = isOdd n +isOdd zero = false +isOdd (suc n) = isEven n + +--Typed version +private + toType : Bool Type + toType false = + toType true = Unit + +isEvenT : Type +isEvenT n = toType (isEven n) + +isOddT : Type +isOddT n = isEvenT (suc n) + +isZero : Bool +isZero zero = true +isZero (suc n) = false + +-- exponential + +_^_ : +m ^ 0 = 1 +m ^ (suc n) = m · m ^ n + + +-- Iterated product +_ˣ_ : {} (A : Type ) (n : ) Type +A ˣ zero = A zero +A ˣ suc n = (A ˣ n) × A (suc n) + + : {} (A : Type ) (0A : (n : ) A n) (n : ) A ˣ n + A 0A zero = 0A zero + A 0A (suc n) = ( A 0A n) , (0A (suc n)) diff --git a/src/docs/Cubical.Data.Nat.Literals.html b/src/docs/Cubical.Data.Nat.Literals.html new file mode 100644 index 0000000..cccd141 --- /dev/null +++ b/src/docs/Cubical.Data.Nat.Literals.html @@ -0,0 +1,22 @@ +{- + + Importing and re-exporting this module allows for (constrained) natural number + and negative integer literals for any type (e.g. Int, ℕ₋₁, ℕ₋₂, ℕ₊₁). + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Data.Nat.Literals where + +open import Agda.Builtin.FromNat public + renaming (Number to HasFromNat) +open import Agda.Builtin.FromNeg public + renaming (Negative to HasFromNeg) +open import Cubical.Data.Unit.Base public + +-- Natural number literals for ℕ + +open import Agda.Builtin.Nat renaming (Nat to ) + +instance + fromNatℕ : HasFromNat + fromNatℕ = record { Constraint = λ _ Unit ; fromNat = λ n n } diff --git a/src/docs/Cubical.Data.Nat.Order.html b/src/docs/Cubical.Data.Nat.Order.html new file mode 100644 index 0000000..e31ad47 --- /dev/null +++ b/src/docs/Cubical.Data.Nat.Order.html @@ -0,0 +1,529 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Data.Nat.Order where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels + + +open import Cubical.Data.Empty as +open import Cubical.Data.Sigma +open import Cubical.Data.Sum as + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Nat.Properties + +open import Cubical.Induction.WellFounded + +open import Cubical.Relation.Nullary + +private + variable + : Level + +infix 4 _≤_ _<_ _≥_ _>_ + +_≤_ : Type₀ +m n = Σ[ k ] k + m n + +_<_ : Type₀ +m < n = suc m n + +_≥_ : Type₀ +m n = n m + +_>_ : Type₀ +m > n = n < m + +data Trichotomy (m n : ) : Type₀ where + lt : m < n Trichotomy m n + eq : m n Trichotomy m n + gt : n < m Trichotomy m n + +private + variable + k l m n : + +private + witness-prop : j isProp (j + m n) + witness-prop {m} {n} j = isSetℕ (j + m) n + +isProp≤ : isProp (m n) +isProp≤ {m} {n} (k , p) (l , q) + = Σ≡Prop witness-prop lemma + where + lemma : k l + lemma = inj-+m (p (sym q)) + +zero-≤ : 0 n +zero-≤ {n} = n , +-zero n + +suc-≤-suc : m n suc m suc n +suc-≤-suc (k , p) = k , (+-suc k _) (cong suc p) + +≤-+k : m n m + k n + k +≤-+k {m} {k = k} (i , p) + = i , +-assoc i m k cong (_+ k) p + +≤SumRight : n k + n +≤SumRight {n} {k} = ≤-+k zero-≤ + +≤-k+ : m n k + m k + n +≤-k+ {m} {n} {k} + = subst (_≤ k + n) (+-comm m k) + subst (m + k ≤_) (+-comm n k) + ≤-+k + +≤SumLeft : n n + k +≤SumLeft {n} {k} = subst (n ≤_) (+-comm k n) (≤-+k zero-≤) + +pred-≤-pred : suc m suc n m n +pred-≤-pred (k , p) = k , injSuc ((sym (+-suc k _)) p) + +≤-refl : m m +≤-refl = 0 , refl + +≤-suc : m n m suc n +≤-suc (k , p) = suc k , cong suc p + +suc-< : suc m < n m < n +suc-< p = pred-≤-pred (≤-suc p) + +≤-sucℕ : n suc n +≤-sucℕ = ≤-suc ≤-refl + +≤-predℕ : predℕ n n +≤-predℕ {zero} = ≤-refl +≤-predℕ {suc n} = ≤-suc ≤-refl + +≤-trans : k m m n k n +≤-trans {k} {m} {n} (i , p) (j , q) = i + j , l2 (l1 q) + where + l1 : j + i + k j + m + l1 = (sym (+-assoc j i k)) (cong (j +_) p) + l2 : i + j + k j + i + k + l2 = cong (_+ k) (+-comm i j) + +≤-antisym : m n n m m n +≤-antisym {m} (i , p) (j , q) = (cong (_+ m) l3) p + where + l1 : j + i + m m + l1 = (sym (+-assoc j i m)) ((cong (j +_) p) q) + l2 : j + i 0 + l2 = m+n≡n→m≡0 l1 + l3 : 0 i + l3 = sym (snd (m+n≡0→m≡0×n≡0 l2)) + +≤-+-≤ : m n l k m + l n + k +≤-+-≤ p q = ≤-trans (≤-+k p) (≤-k+ q) + +≤-k+-cancel : k + m k + n m n +≤-k+-cancel {k} {m} (l , p) = l , inj-m+ (sub k m p) + where + sub : k m k + (l + m) l + (k + m) + sub k m = +-assoc k l m cong (_+ m) (+-comm k l) sym (+-assoc l k m) + +≤-+k-cancel : m + k n + k m n +≤-+k-cancel {m} {k} {n} (l , p) = l , cancelled + where + cancelled : l + m n + cancelled = inj-+m (sym (+-assoc l m k) p) + +≤-+k-trans : (m + k n) m n +≤-+k-trans {m} {k} {n} p = ≤-trans (k , +-comm k m) p + +≤-k+-trans : (k + m n) m n +≤-k+-trans {m} {k} {n} p = ≤-trans (m , refl) p + +≤-·k : m n m · k n · k +≤-·k {m} {n} {k} (d , r) = d · k , reason where + reason : d · k + m · k n · k + reason = d · k + m · k ≡⟨ ·-distribʳ d m k + (d + m) · k ≡⟨ cong ( k) r + n · k + +<-k+-cancel : k + m < k + n m < n +<-k+-cancel {k} {m} {n} = ≤-k+-cancel subst (_≤ k + n) (sym (+-suc k m)) + +¬-<-zero : ¬ m < 0 +¬-<-zero (k , p) = snotz ((sym (+-suc k _)) p) + +¬m<m : ¬ m < m +¬m<m {m} = ¬-<-zero ≤-+k-cancel {k = m} + +≤0→≡0 : n 0 n 0 +≤0→≡0 {zero} ineq = refl +≤0→≡0 {suc n} ineq = ⊥.rec (¬-<-zero ineq) + +predℕ-≤-predℕ : m n (predℕ m) (predℕ n) +predℕ-≤-predℕ {zero} {zero} ineq = ≤-refl +predℕ-≤-predℕ {zero} {suc n} ineq = zero-≤ +predℕ-≤-predℕ {suc m} {zero} ineq = ⊥.rec (¬-<-zero ineq) +predℕ-≤-predℕ {suc m} {suc n} ineq = pred-≤-pred ineq + +¬m+n<m : ¬ m + n < m +¬m+n<m {m} {n} = ¬-<-zero <-k+-cancel subst (m + n <_) (sym (+-zero m)) + +<-weaken : m < n m n +<-weaken (k , p) = suc k , sym (+-suc k _) p + +≤<-trans : l m m < n l < n +≤<-trans p = ≤-trans (suc-≤-suc p) + +<≤-trans : l < m m n l < n +<≤-trans = ≤-trans + +<-trans : l < m m < n l < n +<-trans p = ≤<-trans (<-weaken p) + +<-asym : m < n ¬ n m +<-asym m<n = ¬m<m <≤-trans m<n + +<-+k : m < n m + k < n + k +<-+k p = ≤-+k p + +<-k+ : m < n k + m < k + n +<-k+ {m} {n} {k} p = subst km km k + n) (+-suc k m) (≤-k+ p) + +<-+k-trans : (m + k < n) m < n +<-+k-trans {m} {k} {n} p = ≤<-trans (k , +-comm k m) p + +<-k+-trans : (k + m < n) m < n +<-k+-trans {m} {k} {n} p = ≤<-trans (m , refl) p + +<-+-< : m < n k < l m + k < n + l +<-+-< m<n k<l = <-trans (<-+k m<n) (<-k+ k<l) + +<-+-≤ : m < n k l m + k < n + l +<-+-≤ p q = <≤-trans (<-+k p) (≤-k+ q) + +<-·sk : m < n m · suc k < n · suc k +<-·sk {m} {n} {k} (d , r) = (d · suc k + k) , reason where + reason : (d · suc k + k) + suc (m · suc k) n · suc k + reason = (d · suc k + k) + suc (m · suc k) ≡⟨ sym (+-assoc (d · suc k) k _) + d · suc k + (k + suc (m · suc k)) ≡[ i ]⟨ d · suc k + +-suc k (m · suc k) i + d · suc k + suc m · suc k ≡⟨ ·-distribʳ d (suc m) (suc k) + (d + suc m) · suc k ≡⟨ cong ( suc k) r + n · suc k + +∸-≤ : m n m n m +∸-≤ m zero = ≤-refl +∸-≤ zero (suc n) = ≤-refl +∸-≤ (suc m) (suc n) = ≤-trans (∸-≤ m n) (1 , refl) + +≤-∸-+-cancel : m n (n m) + m n +≤-∸-+-cancel {zero} {n} _ = +-zero _ +≤-∸-+-cancel {suc m} {zero} m≤n = ⊥.rec (¬-<-zero m≤n) +≤-∸-+-cancel {suc m} {suc n} m+1≤n+1 = +-suc _ _ cong suc (≤-∸-+-cancel (pred-≤-pred m+1≤n+1)) + +≤-∸-suc : m n suc (n m) suc n m +≤-∸-suc {zero} {n} m≤n = refl +≤-∸-suc {suc m} {zero} m≤n = ⊥.rec (¬-<-zero m≤n) +≤-∸-suc {suc m} {suc n} m+1≤n+1 = ≤-∸-suc (pred-≤-pred m+1≤n+1) + +≤-∸-k : m n k + (n m) (k + n) m +≤-∸-k {m} {n} {zero} r = refl +≤-∸-k {m} {n} {suc k} r = cong suc (≤-∸-k r) ≤-∸-suc (≤-trans r (k , refl)) + +left-≤-max : m max m n +left-≤-max {zero} {n} = zero-≤ +left-≤-max {suc m} {zero} = ≤-refl +left-≤-max {suc m} {suc n} = suc-≤-suc left-≤-max + +right-≤-max : n max m n +right-≤-max {zero} {m} = zero-≤ +right-≤-max {suc n} {zero} = ≤-refl +right-≤-max {suc n} {suc m} = suc-≤-suc right-≤-max + +min-≤-left : min m n m +min-≤-left {zero} {n} = ≤-refl +min-≤-left {suc m} {zero} = zero-≤ +min-≤-left {suc m} {suc n} = suc-≤-suc min-≤-left + +min-≤-right : min m n n +min-≤-right {zero} {n} = zero-≤ +min-≤-right {suc m} {zero} = ≤-refl +min-≤-right {suc m} {suc n} = suc-≤-suc min-≤-right + +≤Dec : m n Dec (m n) +≤Dec zero n = yes (n , +-zero _) +≤Dec (suc m) zero = no ¬-<-zero +≤Dec (suc m) (suc n) with ≤Dec m n +... | yes m≤n = yes (suc-≤-suc m≤n) +... | no m≰n = no λ m+1≤n+1 m≰n (pred-≤-pred m+1≤n+1 ) + +≤Stable : m n Stable (m n) +≤Stable m n = Dec→Stable (≤Dec m n) + +<Dec : m n Dec (m < n) +<Dec m n = ≤Dec (suc m) n + +<Stable : m n Stable (m < n) +<Stable m n = Dec→Stable (<Dec m n) + +Trichotomy-suc : Trichotomy m n Trichotomy (suc m) (suc n) +Trichotomy-suc (lt m<n) = lt (suc-≤-suc m<n) +Trichotomy-suc (eq m=n) = eq (cong suc m=n) +Trichotomy-suc (gt n<m) = gt (suc-≤-suc n<m) + +_≟_ : m n Trichotomy m n +zero zero = eq refl +zero suc n = lt (n , +-comm n 1) +suc m zero = gt (m , +-comm m 1) +suc m suc n = Trichotomy-suc (m n) + +Dichotomyℕ : (n m : ) (n m) (n > m) +Dichotomyℕ n m with (n m) +... | lt x = inl (<-weaken x) +... | Trichotomy.eq x = inl (0 , x) +... | gt x = inr x + +splitℕ-≤ : (m n : ) (m n) (n < m) +splitℕ-≤ m n with m n +... | lt x = inl (<-weaken x) +... | eq x = inl (0 , x) +... | gt x = inr x + +splitℕ-< : (m n : ) (m < n) (n m) +splitℕ-< m n with m n +... | lt x = inl x +... | eq x = inr (0 , (sym x)) +... | gt x = inr (<-weaken x) + +≤CaseInduction : {P : Type } {n m : } + (n m P n m) (m n P n m) + P n m +≤CaseInduction {n = n} {m = m} p q with n m +... | lt x = p (<-weaken x) +... | eq x = p (subst (n ≤_) x ≤-refl) +... | gt x = q (<-weaken x) + +<-split : m < suc n (m < n) (m n) +<-split {n = zero} = inr snd m+n≡0→m≡0×n≡0 snd pred-≤-pred +<-split {zero} {suc n} = λ _ inl (suc-≤-suc zero-≤) +<-split {suc m} {suc n} = ⊎.map suc-≤-suc (cong suc) <-split pred-≤-pred + +≤-split : m n (m < n) (m n) +≤-split p = <-split (suc-≤-suc p) + +≤→< : m n ¬ m n m < n +≤→< p q = + case (≤-split p) of + λ { (inl r) r + ; (inr r) ⊥.rec (q r) } + +≤-suc-≢ : m suc n (m suc n ) m n +≤-suc-≢ p ¬q = pred-≤-pred (≤→< p ¬q) + +≤-+-split : n m k k n + m (n k) (m (n + m) k) +≤-+-split n m k k≤n+m with n k +... | eq p = inl (0 , p) +... | lt n<k = inl (<-weaken n<k) +... | gt k<n with m ((n + m) k) +... | eq p = inr (0 , p) +... | lt m<n+m∸k = inr (<-weaken m<n+m∸k) +... | gt n+m∸k<m = + ⊥.rec (¬m<m (transport i ≤-∸-+-cancel k≤n+m i < +-comm m n i) (<-+-< n+m∸k<m k<n))) + +<-asym'-case : Trichotomy m n ¬ m < n n m +<-asym'-case (lt p) q = ⊥.rec (q p) +<-asym'-case (eq p) _ = _ , sym p +<-asym'-case (gt p) _ = <-weaken p + +<-asym' : ¬ m < n n m +<-asym' = <-asym'-case (_≟_ _ _) + +private + acc-suc : Acc _<_ n Acc _<_ (suc n) + acc-suc a + = acc λ y y<sn + case <-split y<sn of λ + { (inl y<n) access a y y<n + ; (inr y≡n) subst _ (sym y≡n) a + } + +<-wellfounded : WellFounded _<_ +<-wellfounded zero = acc λ _ ⊥.rec ¬-<-zero +<-wellfounded (suc n) = acc-suc (<-wellfounded n) + +<→≢ : n < m ¬ n m +<→≢ {n} {m} p q = ¬m<m (subst (_< m) q p) + +module _ + (b₀ : ) + (P : Type₀) + (base : n n < suc b₀ P n) + (step : n P n P (suc b₀ + n)) + where + open WFI (<-wellfounded) + + private + dichotomy : b n (n < b) (Σ[ m ] n b + m) + dichotomy b n + = case n b return _ (n < b) (Σ[ m ] n b + m)) of λ + { (lt o) inl o + ; (eq p) inr (0 , p sym (+-zero b)) + ; (gt (m , p)) inr (suc m , sym p +-suc m b +-comm (suc m) b) + } + + dichotomy<≡ : b n (n<b : n < b) dichotomy b n inl n<b + dichotomy<≡ b n n<b + = case dichotomy b n return d d inl n<b) of λ + { (inl x) cong inl (isProp≤ x n<b) + ; (inr (m , p)) ⊥.rec (<-asym n<b (m , sym (p +-comm b m))) + } + + dichotomy+≡ : b m n (p : n b + m) dichotomy b n inr (m , p) + dichotomy+≡ b m n p + = case dichotomy b n return d d inr (m , p)) of λ + { (inl n<b) ⊥.rec (<-asym n<b (m , +-comm m b sym p)) + ; (inr (m' , q)) + cong inr (Σ≡Prop x isSetℕ n (b + x)) (inj-m+ {m = b} (sym q p))) + } + + b = suc b₀ + + lemma₁ : ∀{x y z} x suc z + y y < x + lemma₁ {y = y} {z} p = z , +-suc z y sym p + + subStep : (n : ) (∀ m m < n P m) (n < b) (Σ[ m ] n b + m) P n + subStep n _ (inl l) = base n l + subStep n rec (inr (m , p)) + = transport (cong P (sym p)) (step m (rec m (lemma₁ p))) + + wfStep : (n : ) (∀ m m < n P m) P n + wfStep n rec = subStep n rec (dichotomy b n) + + wfStepLemma₀ : n (n<b : n < b) rec wfStep n rec base n n<b + wfStepLemma₀ n n<b rec = cong (subStep n rec) (dichotomy<≡ b n n<b) + + wfStepLemma₁ : n rec wfStep (b + n) rec step n (rec n (lemma₁ refl)) + wfStepLemma₁ n rec + = cong (subStep (b + n) rec) (dichotomy+≡ b n (b + n) refl) + transportRefl _ + + +induction : n P n + +induction = induction wfStep + + +inductionBase : n (l : n < b) +induction n base n l + +inductionBase n l = induction-compute wfStep n wfStepLemma₀ n l _ + + +inductionStep : n +induction (b + n) step n (+induction n) + +inductionStep n = induction-compute wfStep (b + n) wfStepLemma₁ n _ + +module <-Reasoning where + -- TODO: would it be better to mirror the way it is done in the agda-stdlib? + infixr 2 _<⟨_⟩_ _≤<⟨_⟩_ _≤⟨_⟩_ _<≤⟨_⟩_ _≡<⟨_⟩_ _≡≤⟨_⟩_ _<≡⟨_⟩_ _≤≡⟨_⟩_ + _<⟨_⟩_ : k k < n n < m k < m + _ <⟨ p q = <-trans p q + + _≤<⟨_⟩_ : k k n n < m k < m + _ ≤<⟨ p q = ≤<-trans p q + + _≤⟨_⟩_ : k k n n m k m + _ ≤⟨ p q = ≤-trans p q + + _<≤⟨_⟩_ : k k < n n m k < m + _ <≤⟨ p q = <≤-trans p q + + _≡≤⟨_⟩_ : k k l l m k m + _ ≡≤⟨ p q = subst k k _) (sym p) q + + _≡<⟨_⟩_ : k k l l < m k < m + _ ≡<⟨ p q = _ ≡≤⟨ cong suc p q + + _≤≡⟨_⟩_ : k k l l m k m + _ ≤≡⟨ p q = subst l _ l) q p + + _<≡⟨_⟩_ : k k < l l m k < m + _ <≡⟨ p q = _ ≤≡⟨ p q + + +-- Some lemmas about ∸ +suc∸-fst : (n m : ) m < n suc (n m) (suc n) m +suc∸-fst zero zero p = refl +suc∸-fst zero (suc m) p = ⊥.rec (¬-<-zero p) +suc∸-fst (suc n) zero p = refl +suc∸-fst (suc n) (suc m) p = (suc∸-fst n m (pred-≤-pred p)) + +n∸m≡0 : (n m : ) n m (n m) 0 +n∸m≡0 zero zero p = refl +n∸m≡0 (suc n) zero p = ⊥.rec (¬-<-zero p) +n∸m≡0 zero (suc m) p = refl +n∸m≡0 (suc n) (suc m) p = n∸m≡0 n m (pred-≤-pred p) + +n∸n≡0 : (n : ) n n 0 +n∸n≡0 zero = refl +n∸n≡0 (suc n) = n∸n≡0 n + +n∸l>0 : (n l : ) (l < n) 0 < (n l) +n∸l>0 zero zero r = ⊥.rec (¬-<-zero r) +n∸l>0 zero (suc l) r = ⊥.rec (¬-<-zero r) +n∸l>0 (suc n) zero r = suc-≤-suc zero-≤ +n∸l>0 (suc n) (suc l) r = n∸l>0 n l (pred-≤-pred r) + +-- automation + +≤-solver-type : (m n : ) Trichotomy m n Type +≤-solver-type m n (lt p) = m n +≤-solver-type m n (eq p) = m n +≤-solver-type m n (gt p) = n < m + +≤-solver-case : (m n : ) (p : Trichotomy m n) ≤-solver-type m n p +≤-solver-case m n (lt p) = <-weaken p +≤-solver-case m n (eq p) = _ , p +≤-solver-case m n (gt p) = p + +≤-solver : (m n : ) ≤-solver-type m n (m n) +≤-solver m n = ≤-solver-case m n (m n) + + + +-- inductive order relation taken from agda-stdlib +data _≤'_ : Type where + z≤ : {n} zero ≤' n + s≤s : {m n} m ≤' n suc m ≤' suc n + +_<'_ : Type +m <' n = suc m ≤' n + +-- Smart constructors of _<_ +pattern z<s {n} = s≤s (z≤ {n}) +pattern s<s {m} {n} m<n = s≤s {m} {n} m<n + +¬-<'-zero : ¬ m <' 0 +¬-<'-zero {zero} () +¬-<'-zero {suc m} () + +≤'Dec : m n Dec (m ≤' n) +≤'Dec zero n = yes z≤ +≤'Dec (suc m) zero = no ¬-<'-zero +≤'Dec (suc m) (suc n) with ≤'Dec m n +... | yes m≤'n = yes (s≤s m≤'n) +... | no m≰'n = no λ { (s≤s m≤'n) m≰'n m≤'n } + +≤'IsPropValued : m n isProp (m ≤' n) +≤'IsPropValued zero n z≤ z≤ = refl +≤'IsPropValued (suc m) zero () +≤'IsPropValued (suc m) (suc n) (s≤s x) (s≤s y) = cong s≤s (≤'IsPropValued m n x y) + +≤-∸-≤ : m n l m n m l n l +≤-∸-≤ m n zero r = r +≤-∸-≤ zero zero (suc l) r = ≤-refl +≤-∸-≤ zero (suc n) (suc l) r = (n l) , (+-zero _) +≤-∸-≤ (suc m) zero (suc l) r = ⊥.rec (¬-<-zero r) +≤-∸-≤ (suc m) (suc n) (suc l) r = ≤-∸-≤ m n l (pred-≤-pred r) + +<-∸-< : m n l m < n l < n m l < n l +<-∸-< m n zero r q = r +<-∸-< zero zero (suc l) r q = ⊥.rec (¬-<-zero r) +<-∸-< zero (suc n) (suc l) r q = n∸l>0 (suc n) (suc l) q +<-∸-< (suc m) zero (suc l) r q = ⊥.rec (¬-<-zero r) +<-∸-< (suc m) (suc n) (suc l) r q = <-∸-< m n l (pred-≤-pred r) (pred-≤-pred q) + +≤-∸-≥ : n l k l k n k n l +≤-∸-≥ n zero zero r = ≤-refl +≤-∸-≥ n zero (suc k) r = ∸-≤ n (suc k) +≤-∸-≥ n (suc l) zero r = ⊥.rec (¬-<-zero r) +≤-∸-≥ zero (suc l) (suc k) r = ≤-refl +≤-∸-≥ (suc n) (suc l) (suc k) r = ≤-∸-≥ n l k (pred-≤-pred r) diff --git a/src/docs/Cubical.Data.Nat.Properties.html b/src/docs/Cubical.Data.Nat.Properties.html new file mode 100644 index 0000000..6083be9 --- /dev/null +++ b/src/docs/Cubical.Data.Nat.Properties.html @@ -0,0 +1,342 @@ +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Data.Nat.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Transport + +open import Cubical.Data.Nat.Base +open import Cubical.Data.Empty as +open import Cubical.Data.Sigma +open import Cubical.Data.Sum.Base + +open import Cubical.Relation.Nullary + +private + variable + l m n : + +min : +min zero m = zero +min (suc n) zero = zero +min (suc n) (suc m) = suc (min n m) + +minComm : (n m : ) min n m min m n +minComm zero zero = refl +minComm zero (suc m) = refl +minComm (suc n) zero = refl +minComm (suc n) (suc m) = cong suc (minComm n m) + +max : +max zero m = m +max (suc n) zero = suc n +max (suc n) (suc m) = suc (max n m) + +maxComm : (n m : ) max n m max m n +maxComm zero zero = refl +maxComm zero (suc m) = refl +maxComm (suc n) zero = refl +maxComm (suc n) (suc m) = cong suc (maxComm n m) + +znots : ¬ (0 suc n) +znots eq = subst (caseNat ) eq 0 + +snotz : ¬ (suc n 0) +snotz eq = subst (caseNat ) eq 0 + +injSuc : suc m suc n m n +injSuc p = cong predℕ p + +-- encode decode caracterisation of equality +codeℕ : (n m : ) Type ℓ-zero +codeℕ zero zero = Unit +codeℕ zero (suc m) = +codeℕ (suc n) zero = +codeℕ (suc n) (suc m) = codeℕ n m + +encodeℕ : (n m : ) (n m) codeℕ n m +encodeℕ n m p = subst m codeℕ n m) p (reflCode n) + where + reflCode : (n : ) codeℕ n n + reflCode zero = tt + reflCode (suc n) = reflCode n + +compute-eqℕ : (n m : ) (n m ) codeℕ n m +compute-eqℕ zero zero p = tt +compute-eqℕ zero (suc m) p = znots p +compute-eqℕ (suc n) zero p = snotz p +compute-eqℕ (suc n) (suc m) p = compute-eqℕ n m (injSuc p) + +decodeℕ : (n m : ) codeℕ n m (n m) +decodeℕ zero zero = λ _ refl +decodeℕ zero (suc m) = ⊥.rec +decodeℕ (suc n) zero = ⊥.rec +decodeℕ (suc n) (suc m) = λ r cong suc (decodeℕ n m r) + +≡ℕ≃Codeℕ : (n m : ) (n m) codeℕ n m +≡ℕ≃Codeℕ n m = isoToEquiv is + where + is : Iso (n m) (codeℕ n m) + Iso.fun is = encodeℕ n m + Iso.inv is = decodeℕ n m + Iso.rightInv is = sect n m + where + sect : (n m : ) (r : codeℕ n m) (encodeℕ n m (decodeℕ n m r) r) + sect zero zero tt = refl + sect zero (suc m) r = ⊥.rec r + sect (suc n) zero r = ⊥.rec r + sect (suc n) (suc m) r = sect n m r + Iso.leftInv is = retr n m + where + reflRetr : (n : ) decodeℕ n n (encodeℕ n n refl) refl + reflRetr zero = refl + reflRetr (suc n) i = cong suc (reflRetr n i) + + retr : (n m : ) (p : n m) (decodeℕ n m (encodeℕ n m p) p) + retr n m p = J m p decodeℕ n m (encodeℕ n m p) p) (reflRetr n) p + + +≡ℕ≃Codeℕ' : (n m : ) (n m) codeℕ n m +≡ℕ≃Codeℕ' n m = isoToEquiv is + where + is : Iso (n m) (codeℕ n m) + Iso.fun is = compute-eqℕ n m + Iso.inv is = decodeℕ n m + Iso.rightInv is = sect n m + where + sect : (n m : ) (r : codeℕ n m) compute-eqℕ n m (decodeℕ n m r) r + sect zero zero tt = refl + sect (suc n) (suc m) r = sect n m r + Iso.leftInv is = retr n m + where + reflRetr : (n : ) decodeℕ n n (compute-eqℕ n n refl) refl + reflRetr zero = refl + reflRetr (suc n) i = cong suc (reflRetr n i) + + retr : (n m : ) (p : n m) decodeℕ n m (compute-eqℕ n m p) p + retr n m p = J m p decodeℕ n m (compute-eqℕ n m p) p) (reflRetr n) p + + +discreteℕ : Discrete +discreteℕ zero zero = yes refl +discreteℕ zero (suc n) = no znots +discreteℕ (suc m) zero = no snotz +discreteℕ (suc m) (suc n) with discreteℕ m n +... | yes p = yes (cong suc p) +... | no p = no x p (injSuc x)) + +separatedℕ : Separated +separatedℕ = Discrete→Separated discreteℕ + +isSetℕ : isSet +isSetℕ = Discrete→isSet discreteℕ + +-- Arithmetic facts about predℕ + +suc-predℕ : n ¬ n 0 n suc (predℕ n) +suc-predℕ zero p = ⊥.rec (p refl) +suc-predℕ (suc n) p = refl + +-- Arithmetic facts about + + ++-zero : m m + 0 m ++-zero zero = refl ++-zero (suc m) = cong suc (+-zero m) + ++-suc : m n m + suc n suc (m + n) ++-suc zero n = refl ++-suc (suc m) n = cong suc (+-suc m n) + ++-comm : m n m + n n + m ++-comm m zero = +-zero m ++-comm m (suc n) = (+-suc m n) (cong suc (+-comm m n)) + +-- Addition is associative ++-assoc : m n o m + (n + o) (m + n) + o ++-assoc zero _ _ = refl ++-assoc (suc m) n o = cong suc (+-assoc m n o) + +inj-m+ : m + l m + n l n +inj-m+ {zero} p = p +inj-m+ {suc m} p = inj-m+ (injSuc p) + +inj-+m : l + m n + m l n +inj-+m {l} {m} {n} p = inj-m+ ((+-comm m l) (p (+-comm n m))) + +m+n≡n→m≡0 : m + n n m 0 +m+n≡n→m≡0 {n = zero} = λ p (sym (+-zero _)) p +m+n≡n→m≡0 {n = suc n} p = m+n≡n→m≡0 (injSuc ((sym (+-suc _ n)) p)) + +m+n≡0→m≡0×n≡0 : m + n 0 (m 0) × (n 0) +m+n≡0→m≡0×n≡0 {zero} = refl ,_ +m+n≡0→m≡0×n≡0 {suc m} p = ⊥.rec (snotz p) + +-- Arithmetic facts about · + +0≡m·0 : m 0 m · 0 +0≡m·0 zero = refl +0≡m·0 (suc m) = 0≡m·0 m + +·-suc : m n m · suc n m + m · n +·-suc zero n = refl +·-suc (suc m) n + = cong suc + ( n + m · suc n ≡⟨ cong (n +_) (·-suc m n) + n + (m + m · n) ≡⟨ +-assoc n m (m · n) + (n + m) + m · n ≡⟨ cong (_+ m · n) (+-comm n m) + (m + n) + m · n ≡⟨ sym (+-assoc m n (m · n)) + m + (n + m · n) + ) + +·-comm : m n m · n n · m +·-comm zero n = 0≡m·0 n +·-comm (suc m) n = cong (n +_) (·-comm m n) sym (·-suc n m) + +·-distribʳ : m n o (m · o) + (n · o) (m + n) · o +·-distribʳ zero _ _ = refl +·-distribʳ (suc m) n o = sym (+-assoc o (m · o) (n · o)) cong (o +_) (·-distribʳ m n o) + +·-distribˡ : o m n (o · m) + (o · n) o · (m + n) +·-distribˡ o m n = i ·-comm o m i + ·-comm o n i) ·-distribʳ m n o ·-comm (m + n) o + +·-assoc : m n o m · (n · o) (m · n) · o +·-assoc zero _ _ = refl +·-assoc (suc m) n o = cong (n · o +_) (·-assoc m n o) ·-distribʳ n (m · n) o + +·-identityˡ : m 1 · m m +·-identityˡ m = +-zero m + +·-identityʳ : m m · 1 m +·-identityʳ zero = refl +·-identityʳ (suc m) = cong suc (·-identityʳ m) + +0≡n·sm→0≡n : 0 n · suc m 0 n +0≡n·sm→0≡n {n = zero} p = refl +0≡n·sm→0≡n {n = suc n} p = ⊥.rec (znots p) + +inj-·sm : l · suc m n · suc m l n +inj-·sm {zero} {m} {n} p = 0≡n·sm→0≡n p +inj-·sm {l} {m} {zero} p = sym (0≡n·sm→0≡n (sym p)) +inj-·sm {suc l} {m} {suc n} p = cong suc (inj-·sm (inj-m+ {m = suc m} p)) + +inj-sm· : suc m · l suc m · n l n +inj-sm· {m} {l} {n} p = inj-·sm (·-comm l (suc m) p ·-comm (suc m) n) + +integral-domain-· : {k l : } (k 0 ) (l 0 ) (k · l 0 ) +integral-domain-· {zero} {l} ¬p ¬q r = ¬p refl +integral-domain-· {suc k} {zero} ¬p ¬q r = ¬q refl +integral-domain-· {suc k} {suc l} ¬p ¬q r = snotz r + +-- Arithmetic facts about ∸ + +zero∸ : n zero n zero +zero∸ zero = refl +zero∸ (suc _) = refl + + +n∸n : (n : ) n n 0 +n∸n zero = refl +n∸n (suc n) = n∸n n + +∸-cancelˡ : k m n (k + m) (k + n) m n +∸-cancelˡ zero = λ _ _ refl +∸-cancelˡ (suc k) = ∸-cancelˡ k + ++∸ : k n (k + n) n k ++∸ zero n = n∸n n ++∸ (suc k) zero = cong suc (+-comm k zero) ++∸ (suc k) (suc n) = cong (_∸ n) (+-suc k n) +∸ (suc k) n + +∸+ : k n (n + k) n k +∸+ k n = cong X X n) (+-comm n k) +∸ k n + +∸-cancelʳ : m n k (m + k) (n + k) m n +∸-cancelʳ m n k = i +-comm m k i +-comm n k i) ∸-cancelˡ k m n + +∸-distribʳ : m n k (m n) · k m · k n · k +∸-distribʳ m zero k = refl +∸-distribʳ zero (suc n) k = sym (zero∸ (k + n · k)) +∸-distribʳ (suc m) (suc n) k = ∸-distribʳ m n k sym (∸-cancelˡ k (m · k) (n · k)) + + + +-- factorial: +_! : +zero ! = 1 +suc n ! = (suc n) · (n !) + +--binomial coefficient: +_choose_ : +n choose zero = 1 +zero choose suc k = 0 +suc n choose suc k = n choose (suc k) + n choose k + +evenOrOdd : (n : ) isEvenT n isOddT n +evenOrOdd zero = inl tt +evenOrOdd (suc zero) = inr tt +evenOrOdd (suc (suc n)) = evenOrOdd n + +¬evenAndOdd : (n : ) ¬ isEvenT n × isOddT n +¬evenAndOdd zero (p , ()) +¬evenAndOdd (suc zero) () +¬evenAndOdd (suc (suc n)) = ¬evenAndOdd n + +isPropIsEvenT : (n : ) isProp (isEvenT n) +isPropIsEvenT zero x y = refl +isPropIsEvenT (suc zero) = isProp⊥ +isPropIsEvenT (suc (suc n)) = isPropIsEvenT n + +isPropIsOddT : (n : ) isProp (isOddT n) +isPropIsOddT n = isPropIsEvenT (suc n) + +isPropEvenOrOdd : (n : ) isProp (isEvenT n isOddT n) +isPropEvenOrOdd n (inl x) (inl x₁) = cong inl (isPropIsEvenT n x x₁) +isPropEvenOrOdd n (inl x) (inr x₁) = ⊥.rec (¬evenAndOdd n (x , x₁)) +isPropEvenOrOdd n (inr x) (inl x₁) = ⊥.rec (¬evenAndOdd (suc n) (x , x₁)) +isPropEvenOrOdd n (inr x) (inr x₁) = cong inr (isPropIsEvenT (suc n) x x₁) + +module PlusBis where + + _+'_ : + zero +' b = b + suc a +' zero = suc a + suc a +' suc b = 2 + (a + b) + + +'≡+ : (n m : ) n +' m n + m + +'≡+ zero m = refl + +'≡+ (suc n) zero = cong suc (sym (+-comm n zero)) + +'≡+ (suc n) (suc m) = cong suc (sym (+-suc n m)) + + +'-comm : (n m : ) n +' m m +' n + +'-comm n m = +'≡+ n m ∙∙ +-comm n m ∙∙ sym (+'≡+ m n) + + +'-assoc : (n m l : ) (n +' (m +' l)) ((n +' m) +' l) + +'-assoc n m l = + i +'≡+ n (+'≡+ m l i) i) + ∙∙ +-assoc n m l + ∙∙ i +'≡+ (+'≡+ n m (~ i)) l (~ i)) + + +'-rid : (n : ) n +' 0 n + +'-rid zero = refl + +'-rid (suc n) = refl + + +'-lid : (n : ) 0 +' n n + +'-lid n = refl + + +'-suc : (n m : ) suc (n +' m) suc n +' m + +'-suc zero zero = refl + +'-suc zero (suc m) = refl + +'-suc (suc n) zero = refl + +'-suc (suc n) (suc m) = refl + +-- Neat transport lemma for ℕ +compSubstℕ : {} {A : Type } {n m l : } + (p : n m) (q : m l) (r : n l) + {x : _} + subst A q (subst A p x) + subst A r x +compSubstℕ {A = A} p q r {x = x} = + sym (substComposite A p q x) + λ i subst A (isSetℕ _ _ (p q) r i) x diff --git a/src/docs/Cubical.Data.Nat.html b/src/docs/Cubical.Data.Nat.html new file mode 100644 index 0000000..d06d8d0 --- /dev/null +++ b/src/docs/Cubical.Data.Nat.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Nat where + +open import Cubical.Data.Nat.Base public +open import Cubical.Data.Nat.Properties public diff --git a/src/docs/Cubical.Data.Prod.Base.html b/src/docs/Cubical.Data.Prod.Base.html new file mode 100644 index 0000000..07476e6 --- /dev/null +++ b/src/docs/Cubical.Data.Prod.Base.html @@ -0,0 +1,60 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Prod.Base where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function + +-- Here we define an inductive version of the product type, see below +-- for its uses. + +-- See `Cubical.Data.Sigma` for `_×_` defined as a special case of +-- sigma types, which is the generally preferred one. + +-- If × is defined using Σ then transp/hcomp will be compute +-- "negatively", that is, they won't reduce unless we project out the +-- first of second component. This is not always what we want so this +-- implementation is done using a datatype which computes positively. + + +private + variable + ℓ' : Level + +data _×_ (A : Type ) (B : Type ℓ') : Type (ℓ-max ℓ') where + _,_ : A B A × B + +infixr 5 _×_ + +proj₁ : {A : Type } {B : Type ℓ'} A × B A +proj₁ (x , _) = x + +proj₂ : {A : Type } {B : Type ℓ'} A × B B +proj₂ (_ , x) = x + + +private + variable + A : Type + B C : A Type + +intro : (∀ a B a) (∀ a C a) a B a × C a +intro f g a = f a , g a + +map : {B : Type } {D : B Type ℓ'} + (∀ a C a) (∀ b D b) (x : A × B) C (proj₁ x) × D (proj₂ x) +map f g = intro (f proj₁) (g proj₂) + + +×-η : {A : Type } {B : Type ℓ'} (x : A × B) x ((proj₁ x) , (proj₂ x)) +×-η (x , x₁) = refl + + +-- The product type with one parameter in Typeω + +record _×ω_ {a} (A : Type a) (B : Typeω) : Typeω where + constructor _,_ + field + fst : A + snd : B diff --git a/src/docs/Cubical.Data.Sigma.Base.html b/src/docs/Cubical.Data.Sigma.Base.html new file mode 100644 index 0000000..cb51e76 --- /dev/null +++ b/src/docs/Cubical.Data.Sigma.Base.html @@ -0,0 +1,52 @@ +{- Basic definitions using Σ-types + +Σ-types are defined in Core/Primitives as they are needed for Glue types. + +The file contains: + +- Non-dependent pair types: A × B +- Mere existence: ∃[x ∈ A] B +- Unique existence: ∃![x ∈ A] B + +-} +{-# OPTIONS --safe #-} +module Cubical.Data.Sigma.Base where + +open import Cubical.Core.Primitives public + +open import Cubical.Foundations.Prelude +open import Cubical.HITs.PropositionalTruncation.Base + + +-- Non-dependent pair types + +_×_ : { ℓ'} (A : Type ) (B : Type ℓ') Type (ℓ-max ℓ') +A × B = Σ A _ B) + +infixr 5 _×_ + + +-- Mere existence + + : { ℓ'} (A : Type ) (B : A Type ℓ') Type (ℓ-max ℓ') + A B = Σ A B ∥₁ + +infix 2 ∃-syntax + +∃-syntax : { ℓ'} (A : Type ) (B : A Type ℓ') Type (ℓ-max ℓ') +∃-syntax = + +syntax ∃-syntax A x B) = ∃[ x A ] B + + +-- Unique existence + +∃! : { ℓ'} (A : Type ) (B : A Type ℓ') Type (ℓ-max ℓ') +∃! A B = isContr (Σ A B) + +infix 2 ∃!-syntax + +∃!-syntax : { ℓ'} (A : Type ) (B : A Type ℓ') Type (ℓ-max ℓ') +∃!-syntax = ∃! + +syntax ∃!-syntax A x B) = ∃![ x A ] B diff --git a/src/docs/Cubical.Data.Sigma.Properties.html b/src/docs/Cubical.Data.Sigma.Properties.html new file mode 100644 index 0000000..c1bca62 --- /dev/null +++ b/src/docs/Cubical.Data.Sigma.Properties.html @@ -0,0 +1,486 @@ +{- + +Basic properties about Σ-types + +- Action of Σ on functions ([map-fst], [map-snd]) +- Characterization of equality in Σ-types using dependent paths ([ΣPath{Iso,≃,≡}PathΣ], [Σ≡Prop]) +- Proof that discrete types are closed under Σ ([discreteΣ]) +- Commutativity and associativity ([Σ-swap-*, Σ-assoc-*]) +- Distributivity of Π over Σ ([Σ-Π-*]) +- Action of Σ on isomorphisms, equivalences, and paths ([Σ-cong-fst], [Σ-cong-snd], ...) +- Characterization of equality in Σ-types using transport ([ΣPathTransport{≃,≡}PathΣ]) +- Σ with a contractible base is its fiber ([Σ-contractFst, ΣUnit]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Data.Sigma.Properties where + +open import Cubical.Data.Sigma.Base + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Relation.Nullary +open import Cubical.Data.Unit.Base +open import Cubical.Data.Empty.Base + +open import Cubical.Reflection.StrictEquiv + +open Iso + +private + variable + ℓ' ℓ'' : Level + A A' : Type + B B' : (a : A) Type + C : (a : A) (b : B a) Type + +map-fst : {B : Type } (f : A A') A × B A' × B +map-fst f (a , b) = (f a , b) + +map-snd : (∀ {a} B a B' a) Σ A B Σ A B' +map-snd f (a , b) = (a , f b) + +map-× : {B : Type } {B' : Type ℓ'} (A A') (B B') A × B A' × B' +map-× f g (a , b) = (f a , g b) + +≡-× : {A : Type } {B : Type ℓ'} {x y : A × B} fst x fst y snd x snd y x y +≡-× p q i = (p i) , (q i) + + +-- Characterization of paths in Σ using dependent paths + +module _ {A : I Type } {B : (i : I) A i Type ℓ'} + {x : Σ (A i0) (B i0)} {y : Σ (A i1) (B i1)} + where + + ΣPathP : Σ[ p PathP A (fst x) (fst y) ] PathP i B i (p i)) (snd x) (snd y) + PathP i Σ (A i) (B i)) x y + ΣPathP eq i = fst eq i , snd eq i + + PathPΣ : PathP i Σ (A i) (B i)) x y + Σ[ p PathP A (fst x) (fst y) ] PathP i B i (p i)) (snd x) (snd y) + PathPΣ eq = i fst (eq i)) , i snd (eq i)) + + -- allows one to write + -- open PathPΣ somePathInΣAB renaming (fst ... ) + module PathPΣ (p : PathP i Σ (A i) (B i)) x y) where + open Σ (PathPΣ p) public + + ΣPathIsoPathΣ : Iso (Σ[ p PathP A (fst x) (fst y) ] (PathP i B i (p i)) (snd x) (snd y))) + (PathP i Σ (A i) (B i)) x y) + fun ΣPathIsoPathΣ = ΣPathP + inv ΣPathIsoPathΣ = PathPΣ + rightInv ΣPathIsoPathΣ _ = refl + leftInv ΣPathIsoPathΣ _ = refl + + unquoteDecl ΣPath≃PathΣ = declStrictIsoToEquiv ΣPath≃PathΣ ΣPathIsoPathΣ + + ΣPath≡PathΣ : (Σ[ p PathP A (fst x) (fst y) ] (PathP i B i (p i)) (snd x) (snd y))) + (PathP i Σ (A i) (B i)) x y) + ΣPath≡PathΣ = ua ΣPath≃PathΣ + +×≡Prop : isProp A' {u v : A × A'} u .fst v .fst u v +×≡Prop pB {u} {v} p i = (p i) , (pB (u .snd) (v .snd) i) + +-- Useful lemma to prove unique existence +uniqueExists : (a : A) (b : B a) (h : (a' : A) isProp (B a')) (H : (a' : A) B a' a a') ∃![ a A ] B a +fst (uniqueExists a b h H) = (a , b) +snd (uniqueExists a b h H) (a' , b') = ΣPathP (H a' b' , isProp→PathP i h (H a' b' i)) b b') + + +-- Characterization of dependent paths in Σ + +module _ {A : I Type } {B : (i : I) (a : A i) Type ℓ'} + {x : Σ (A i0) (B i0)} {y : Σ (A i1) (B i1)} + where + + ΣPathPIsoPathPΣ : + Iso (Σ[ p PathP A (x .fst) (y .fst) ] PathP i B i (p i)) (x .snd) (y .snd)) + (PathP i Σ (A i) (B i)) x y) + ΣPathPIsoPathPΣ .fun (p , q) i = p i , q i + ΣPathPIsoPathPΣ .inv pq .fst i = pq i .fst + ΣPathPIsoPathPΣ .inv pq .snd i = pq i .snd + ΣPathPIsoPathPΣ .rightInv _ = refl + ΣPathPIsoPathPΣ .leftInv _ = refl + + unquoteDecl ΣPathP≃PathPΣ = declStrictIsoToEquiv ΣPathP≃PathPΣ ΣPathPIsoPathPΣ + + ΣPathP≡PathPΣ = ua ΣPathP≃PathPΣ + +-- Σ of discrete types + +discreteΣ : Discrete A ((a : A) Discrete (B a)) Discrete (Σ A B) +discreteΣ {B = B} Adis Bdis (a0 , b0) (a1 , b1) = discreteΣ' (Adis a0 a1) + where + discreteΣ' : Dec (a0 a1) Dec ((a0 , b0) (a1 , b1)) + discreteΣ' (yes p) = J a1 p b1 Dec ((a0 , b0) (a1 , b1))) (discreteΣ'') p b1 + where + discreteΣ'' : (b1 : B a0) Dec ((a0 , b0) (a0 , b1)) + discreteΣ'' b1 with Bdis a0 b0 b1 + ... | (yes q) = yes (transport ΣPath≡PathΣ (refl , q)) + ... | (no ¬q) = no r ¬q (subst X PathP i B (X i)) b0 b1) (Discrete→isSet Adis a0 a0 (cong fst r) refl) (cong snd r))) + discreteΣ' (no ¬p) = no r ¬p (cong fst r)) + +lUnit×Iso : Iso (Unit × A) A +fun lUnit×Iso = snd +inv lUnit×Iso = tt ,_ +rightInv lUnit×Iso _ = refl +leftInv lUnit×Iso _ = refl + +lUnit*×Iso : ∀{} Iso (Unit* {} × A) A +fun lUnit*×Iso = snd +inv lUnit*×Iso = tt* ,_ +rightInv lUnit*×Iso _ = refl +leftInv lUnit*×Iso _ = refl + +rUnit×Iso : Iso (A × Unit) A +fun rUnit×Iso = fst +inv rUnit×Iso = _, tt +rightInv rUnit×Iso _ = refl +leftInv rUnit×Iso _ = refl + +rUnit*×Iso : ∀{} Iso (A × Unit* {}) A +fun rUnit*×Iso = fst +inv rUnit*×Iso = _, tt* +rightInv rUnit*×Iso _ = refl +leftInv rUnit*×Iso _ = refl + +module _ {A : Type } {A' : Type ℓ'} where + Σ-swap-Iso : Iso (A × A') (A' × A) + fun Σ-swap-Iso (x , y) = (y , x) + inv Σ-swap-Iso (x , y) = (y , x) + rightInv Σ-swap-Iso _ = refl + leftInv Σ-swap-Iso _ = refl + + unquoteDecl Σ-swap-≃ = declStrictIsoToEquiv Σ-swap-≃ Σ-swap-Iso + +module _ {A : Type } {B : A Type ℓ'} {C : a B a Type ℓ''} where + Σ-assoc-Iso : Iso (Σ[ a Σ A B ] C (fst a) (snd a)) (Σ[ a A ] Σ[ b B a ] C a b) + fun Σ-assoc-Iso ((x , y) , z) = (x , (y , z)) + inv Σ-assoc-Iso (x , (y , z)) = ((x , y) , z) + rightInv Σ-assoc-Iso _ = refl + leftInv Σ-assoc-Iso _ = refl + + unquoteDecl Σ-assoc-≃ = declStrictIsoToEquiv Σ-assoc-≃ Σ-assoc-Iso + + Σ-Π-Iso : Iso ((a : A) Σ[ b B a ] C a b) (Σ[ f ((a : A) B a) ] a C a (f a)) + fun Σ-Π-Iso f = (fst f , snd f) + inv Σ-Π-Iso (f , g) x = (f x , g x) + rightInv Σ-Π-Iso _ = refl + leftInv Σ-Π-Iso _ = refl + + unquoteDecl Σ-Π-≃ = declStrictIsoToEquiv Σ-Π-≃ Σ-Π-Iso + +module _ {A : Type } {B : A Type ℓ'} {B' : a Type ℓ''} where + Σ-assoc-swap-Iso : Iso (Σ[ a Σ A B ] B' (fst a)) (Σ[ a Σ A B' ] B (fst a)) + fun Σ-assoc-swap-Iso ((x , y) , z) = ((x , z) , y) + inv Σ-assoc-swap-Iso ((x , z) , y) = ((x , y) , z) + rightInv Σ-assoc-swap-Iso _ = refl + leftInv Σ-assoc-swap-Iso _ = refl + + unquoteDecl Σ-assoc-swap-≃ = declStrictIsoToEquiv Σ-assoc-swap-≃ Σ-assoc-swap-Iso + +Σ-cong-iso-fst : (isom : Iso A A') Iso (Σ A (B fun isom)) (Σ A' B) +fun (Σ-cong-iso-fst isom) x = fun isom (x .fst) , x .snd +inv (Σ-cong-iso-fst {B = B} isom) x = inv isom (x .fst) , subst B (sym (ε (x .fst))) (x .snd) + where + ε = isHAEquiv.rinv (snd (iso→HAEquiv isom)) +rightInv (Σ-cong-iso-fst {B = B} isom) (x , y) = ΣPathP (ε x , toPathP goal) + where + ε = isHAEquiv.rinv (snd (iso→HAEquiv isom)) + goal : subst B (ε x) (subst B (sym (ε x)) y) y + goal = sym (substComposite B (sym (ε x)) (ε x) y) + ∙∙ cong x subst B x y) (lCancel (ε x)) + ∙∙ substRefl {B = B} y +leftInv (Σ-cong-iso-fst {A = A} {B = B} isom) (x , y) = ΣPathP (leftInv isom x , toPathP goal) + where + ε = isHAEquiv.rinv (snd (iso→HAEquiv isom)) + γ = isHAEquiv.com (snd (iso→HAEquiv isom)) + + lem : (x : A) sym (ε (fun isom x)) cong (fun isom) (leftInv isom x) refl + lem x = cong a sym (ε (fun isom x)) a) (γ x) lCancel (ε (fun isom x)) + + goal : subst B (cong (fun isom) (leftInv isom x)) (subst B (sym (ε (fun isom x))) y) y + goal = sym (substComposite B (sym (ε (fun isom x))) (cong (fun isom) (leftInv isom x)) y) + ∙∙ cong a subst B a y) (lem x) + ∙∙ substRefl {B = B} y + +Σ-cong-equiv-fst : (e : A A') Σ A (B equivFun e) Σ A' B +-- we could just do this: +-- Σ-cong-equiv-fst e = isoToEquiv (Σ-cong-iso-fst (equivToIso e)) +-- but the following reduces slightly better +Σ-cong-equiv-fst {A = A} {A' = A'} {B = B} e = intro , isEqIntro + where + intro : Σ A (B equivFun e) Σ A' B + intro (a , b) = equivFun e a , b + isEqIntro : isEquiv intro + isEqIntro .equiv-proof x = ctr , isCtr where + PB : {x y} x y B x B y Type _ + PB p = PathP i B (p i)) + + open Σ x renaming (fst to a'; snd to b) + open Σ (equivCtr e a') renaming (fst to ctrA; snd to α) + ctrB : B (equivFun e ctrA) + ctrB = subst B (sym α) b + ctrP : PB α ctrB b + ctrP = symP (transport-filler i B (sym α i)) b) + ctr : fiber intro x + ctr = (ctrA , ctrB) , ΣPathP (α , ctrP) + + isCtr : y ctr y + isCtr ((r , s) , p) = λ i (a≡r i , b!≡s i) , ΣPathP (α≡ρ i , coh i) where + open PathPΣ p renaming (fst to ρ; snd to σ) + open PathPΣ (equivCtrPath e a' (r , ρ)) renaming (fst to a≡r; snd to α≡ρ) + + b!≡s : PB (cong (equivFun e) a≡r) ctrB s + b!≡s i = comp k B (α≡ρ i (~ k))) k + { (i = i0) ctrP (~ k) + ; (i = i1) σ (~ k) + })) b + + coh : PathP i PB (α≡ρ i) (b!≡s i) b) ctrP σ + coh i j = fill k B (α≡ρ i (~ k))) k + { (i = i0) ctrP (~ k) + ; (i = i1) σ (~ k) + })) (inS b) (~ j) + +Σ-cong-fst : (p : A A') Σ A (B transport p) Σ A' B +Σ-cong-fst {B = B} p i = Σ (p i) (B transp j p (i j)) i) + +Σ-cong-iso-snd : ((x : A) Iso (B x) (B' x)) Iso (Σ A B) (Σ A B') +fun (Σ-cong-iso-snd isom) (x , y) = x , fun (isom x) y +inv (Σ-cong-iso-snd isom) (x , y') = x , inv (isom x) y' +rightInv (Σ-cong-iso-snd isom) (x , y) = ΣPathP (refl , rightInv (isom x) y) +leftInv (Σ-cong-iso-snd isom) (x , y') = ΣPathP (refl , leftInv (isom x) y') + +Σ-cong-equiv-snd : (∀ a B a B' a) Σ A B Σ A B' +Σ-cong-equiv-snd h = isoToEquiv (Σ-cong-iso-snd (equivToIso h)) + +Σ-cong-snd : ((x : A) B x B' x) Σ A B Σ A B' +Σ-cong-snd {A = A} p i = Σ[ x A ] (p x i) + +Σ-cong-iso : (isom : Iso A A') + ((x : A) Iso (B x) (B' (fun isom x))) + Iso (Σ A B) (Σ A' B') +Σ-cong-iso isom isom' = compIso (Σ-cong-iso-snd isom') (Σ-cong-iso-fst isom) + +Σ-cong-equiv : (e : A A') + ((x : A) B x B' (equivFun e x)) + Σ A B Σ A' B' +Σ-cong-equiv e e' = isoToEquiv (Σ-cong-iso (equivToIso e) (equivToIso e')) + +Σ-cong' : (p : A A') PathP i p i Type ℓ') B B' Σ A B Σ A' B' +Σ-cong' p p' = cong₂ (A : Type _) (B : A Type _) Σ A B) p p' + +Σ-cong-equiv-prop : + (e : A A') + ((x : A ) isProp (B x)) + ((x : A') isProp (B' x)) + ((x : A) B x B' (equivFun e x)) + ((x : A) B' (equivFun e x) B x) + Σ A B Σ A' B' +Σ-cong-equiv-prop e prop prop' prop→ prop← = + Σ-cong-equiv e x propBiimpl→Equiv (prop x) (prop' (equivFun e x)) (prop→ x) (prop← x)) + +-- Alternative version for path in Σ-types, as in the HoTT book + +ΣPathTransport : (a b : Σ A B) Type _ +ΣPathTransport {B = B} a b = Σ[ p (fst a fst b) ] transport i B (p i)) (snd a) snd b + +IsoΣPathTransportPathΣ : (a b : Σ A B) Iso (ΣPathTransport a b) (a b) +IsoΣPathTransportPathΣ {B = B} a b = + compIso (Σ-cong-iso-snd p invIso (PathPIsoPath i B (p i)) _ _))) + ΣPathIsoPathΣ + +ΣPathTransport≃PathΣ : (a b : Σ A B) ΣPathTransport a b (a b) +ΣPathTransport≃PathΣ {B = B} a b = isoToEquiv (IsoΣPathTransportPathΣ a b) + +ΣPathTransport→PathΣ : (a b : Σ A B) ΣPathTransport a b (a b) +ΣPathTransport→PathΣ a b = Iso.fun (IsoΣPathTransportPathΣ a b) + +PathΣ→ΣPathTransport : (a b : Σ A B) (a b) ΣPathTransport a b +PathΣ→ΣPathTransport a b = Iso.inv (IsoΣPathTransportPathΣ a b) + +ΣPathTransport≡PathΣ : (a b : Σ A B) ΣPathTransport a b (a b) +ΣPathTransport≡PathΣ a b = ua (ΣPathTransport≃PathΣ a b) + +Σ-contractFstIso : (c : isContr A) Iso (Σ A B) (B (c .fst)) +fun (Σ-contractFstIso {B = B} c) p = subst B (sym (c .snd (fst p))) (snd p) +inv (Σ-contractFstIso {B = B} c) b = _ , b +rightInv (Σ-contractFstIso {B = B} c) b = + cong p subst B p b) (isProp→isSet (isContr→isProp c) _ _ _ _) transportRefl _ +fst (leftInv (Σ-contractFstIso {B = B} c) p j) = c .snd (fst p) j +snd (leftInv (Σ-contractFstIso {B = B} c) p j) = + transp i B (c .snd (fst p) (~ i j))) j (snd p) + +Σ-contractFst : (c : isContr A) Σ A B B (c .fst) +Σ-contractFst {B = B} c = isoToEquiv (Σ-contractFstIso c) + +-- a special case of the above +module _ (A : Unit Type ) where + ΣUnit : Σ Unit A A tt + unquoteDef ΣUnit = defStrictEquiv ΣUnit snd { x (tt , x) }) + +Σ-contractSnd : ((a : A) isContr (B a)) Σ A B A +Σ-contractSnd c = isoToEquiv isom + where + isom : Iso _ _ + isom .fun = fst + isom .inv a = a , c a .fst + isom .rightInv _ = refl + isom .leftInv (a , b) = cong (a ,_) (c a .snd b) + +isEmbeddingFstΣProp : ((x : A) isProp (B x)) + {u v : Σ A B} + isEquiv (p : u v) cong fst p) +isEmbeddingFstΣProp {B = B} pB {u = u} {v = v} .equiv-proof x = ctr , isCtr + where + ctrP : u v + ctrP = ΣPathP (x , isProp→PathP _ pB _) _ _) + ctr : fiber (p : u v) cong fst p) x + ctr = ctrP , refl + + isCtr : z ctr z + isCtr (z , p) = ΣPathP (ctrP≡ , cong (sym snd) fzsingl) where + fzsingl : Path (singl x) (x , refl) (cong fst z , sym p) + fzsingl = isContrSingl x .snd (cong fst z , sym p) + ctrSnd : SquareP i j B (fzsingl i .fst j)) (cong snd ctrP) (cong snd z) _ _ + ctrSnd = isProp→SquareP _ _ pB _) _ _ _ _ + ctrP≡ : ctrP z + ctrP≡ i = ΣPathP (fzsingl i .fst , ctrSnd i) + +Σ≡PropEquiv : ((x : A) isProp (B x)) {u v : Σ A B} + (u .fst v .fst) (u v) +Σ≡PropEquiv pB = invEquiv (_ , isEmbeddingFstΣProp pB) + +Σ≡Prop : ((x : A) isProp (B x)) {u v : Σ A B} + (p : u .fst v .fst) u v +Σ≡Prop pB p = equivFun (Σ≡PropEquiv pB) p + +-- dependent version +ΣPathPProp : { ℓ'} {A : I Type } {B : (i : I) A i Type ℓ'} + {u : Σ (A i0) (B i0)} {v : Σ (A i1) (B i1)} + ((a : A (i1)) isProp (B i1 a)) + PathP i A i) (fst u) (fst v) + PathP i Σ (A i) (B i)) u v +fst (ΣPathPProp {u = u} {v = v} pB p i) = p i +snd (ΣPathPProp {B = B} {u = u} {v = v} pB p i) = lem i + where + lem : PathP i B i (p i)) (snd u) (snd v) + lem = toPathP (pB _ _ _) + +≃-× : {ℓ'' ℓ'''} {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} A C B D A × B C × D +≃-× eq1 eq2 = + map-× (fst eq1) (fst eq2) + , record + { equiv-proof + = λ {(c , d) ((eq1⁻ c .fst .fst + , eq2⁻ d .fst .fst) + , ≡-× (eq1⁻ c .fst .snd) + (eq2⁻ d .fst .snd)) + , λ {((a , b) , p) ΣPathP (≡-× (cong fst (eq1⁻ c .snd (a , cong fst p))) + (cong fst (eq2⁻ d .snd (b , cong snd p))) + , λ i ≡-× (snd ((eq1⁻ c .snd (a , cong fst p)) i)) + (snd ((eq2⁻ d .snd (b , cong snd p)) i)))}}} + where + eq1⁻ = equiv-proof (eq1 .snd) + eq2⁻ = equiv-proof (eq2 .snd) + +{- Some simple ismorphisms -} + +prodIso : { ℓ' ℓ'' ℓ'''} {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + Iso A C + Iso B D + Iso (A × B) (C × D) +Iso.fun (prodIso iAC iBD) (a , b) = (Iso.fun iAC a) , Iso.fun iBD b +Iso.inv (prodIso iAC iBD) (c , d) = (Iso.inv iAC c) , Iso.inv iBD d +Iso.rightInv (prodIso iAC iBD) (c , d) = ΣPathP ((Iso.rightInv iAC c) , (Iso.rightInv iBD d)) +Iso.leftInv (prodIso iAC iBD) (a , b) = ΣPathP ((Iso.leftInv iAC a) , (Iso.leftInv iBD b)) + +prodEquivToIso : {ℓ'' ℓ'''} {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + (e : A C)(e' : B D) + prodIso (equivToIso e) (equivToIso e') equivToIso (≃-× e e') +Iso.fun (prodEquivToIso e e' i) = Iso.fun (equivToIso (≃-× e e')) +Iso.inv (prodEquivToIso e e' i) = Iso.inv (equivToIso (≃-× e e')) +Iso.rightInv (prodEquivToIso e e' i) = Iso.rightInv (equivToIso (≃-× e e')) +Iso.leftInv (prodEquivToIso e e' i) = Iso.leftInv (equivToIso (≃-× e e')) + +toProdIso : {B C : A Type } + Iso ((a : A) B a × C a) (((a : A) B a) × ((a : A) C a)) +Iso.fun toProdIso = λ f a fst (f a)) , a snd (f a)) +Iso.inv toProdIso (f , g) = λ a (f a) , (g a) +Iso.rightInv toProdIso (f , g) = refl +Iso.leftInv toProdIso b = refl + +module _ {A : Type } {B : A Type ℓ'} {C : a B a Type ℓ''} where + curryIso : Iso (((a , b) : Σ A B) C a b) ((a : A) (b : B a) C a b) + Iso.fun curryIso f a b = f (a , b) + Iso.inv curryIso f a = f (fst a) (snd a) + Iso.rightInv curryIso a = refl + Iso.leftInv curryIso f = refl + + unquoteDecl curryEquiv = declStrictIsoToEquiv curryEquiv curryIso + +-- Sigma type with empty base + +module _ (A : Type ) where + + open Iso + + ΣEmptyIso : Iso (Σ A) + fun ΣEmptyIso (* , _) = * + + ΣEmpty : Σ A + ΣEmpty = isoToEquiv ΣEmptyIso + +module _ { : Level} (A : ⊥* {} Type ) where + + open Iso + + ΣEmpty*Iso : Iso (Σ ⊥* A) ⊥* + fun ΣEmpty*Iso (* , _) = * + +-- fiber of projection map + +module _ + (A : Type ) + (B : A Type ℓ') where + + private + proj : Σ A B A + proj (a , b) = a + + module _ + (a : A) where + + open Iso + + fiberProjIso : Iso (B a) (fiber proj a) + fiberProjIso .fun b = (a , b) , refl + fiberProjIso .inv ((a' , b') , p) = subst B p b' + fiberProjIso .leftInv b i = substRefl {B = B} b i + fiberProjIso .rightInv (_ , p) i .fst .fst = p (~ i) + fiberProjIso .rightInv ((_ , b') , p) i .fst .snd = subst-filler B p b' (~ i) + fiberProjIso .rightInv (_ , p) i .snd j = p (~ i j) + + fiberProjEquiv : B a fiber proj a + fiberProjEquiv = isoToEquiv fiberProjIso + +separatedΣ : Separated A ((a : A) Separated (B a)) Separated (Σ A B) +separatedΣ {B = B} sepA sepB (a , b) (a' , b') p = ΣPathTransport→PathΣ _ _ (pA , pB) + where + pA : a a' + pA = sepA a a' q p r q (cong fst r))) + + pB : subst B pA b b' + pB = sepB _ _ _ q p r q (cong r' subst B r' b) + (Separated→isSet sepA _ _ pA (cong fst r)) snd (PathΣ→ΣPathTransport _ _ r)))) diff --git a/src/docs/Cubical.Data.Sigma.html b/src/docs/Cubical.Data.Sigma.html new file mode 100644 index 0000000..60d93c2 --- /dev/null +++ b/src/docs/Cubical.Data.Sigma.html @@ -0,0 +1,6 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Sigma where + +open import Cubical.Data.Sigma.Base public + +open import Cubical.Data.Sigma.Properties public diff --git a/src/docs/Cubical.Data.Sum.Base.html b/src/docs/Cubical.Data.Sum.Base.html new file mode 100644 index 0000000..f911b12 --- /dev/null +++ b/src/docs/Cubical.Data.Sum.Base.html @@ -0,0 +1,37 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Sum.Base where + +open import Cubical.Relation.Nullary.Base + +open import Cubical.Core.Everything + +private + variable + ℓ' : Level + A B C D : Type + +data _⊎_ (A : Type )(B : Type ℓ') : Type (ℓ-max ℓ') where + inl : A A B + inr : B A B + +rec : {C : Type } (A C) (B C) A B C +rec f _ (inl x) = f x +rec _ g (inr y) = g y + +elim : {C : A B Type } ((a : A) C (inl a)) ((b : B) C (inr b)) + (x : A B) C x +elim f _ (inl x) = f x +elim _ g (inr y) = g y + +map : (A C) (B D) A B C D +map f _ (inl x) = inl (f x) +map _ g (inr y) = inr (g y) + +_⊎?_ : {P Q : Type } Dec P Dec Q Dec (P Q) +P? ⊎? Q? with P? | Q? +... | yes p | _ = yes (inl p) +... | no _ | yes q = yes (inr q) +... | no ¬p | no ¬q = no λ + { (inl p) ¬p p + ; (inr q) ¬q q + } diff --git a/src/docs/Cubical.Data.Sum.Properties.html b/src/docs/Cubical.Data.Sum.Properties.html new file mode 100644 index 0000000..b907728 --- /dev/null +++ b/src/docs/Cubical.Data.Sum.Properties.html @@ -0,0 +1,294 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Sum.Properties where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Functions.Embedding +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Empty as +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Relation.Nullary + +open import Cubical.Data.Sum.Base as + +open Iso + + +private + variable + ℓa ℓb ℓc ℓd ℓe : Level + A : Type ℓa + B : Type ℓb + C : Type ℓc + D : Type ℓd + E : A B Type ℓe + + +-- Path space of sum type +module ⊎Path { ℓ'} {A : Type } {B : Type ℓ'} where + + Cover : A B A B Type (ℓ-max ℓ') + Cover (inl a) (inl a') = Lift {j = ℓ-max ℓ'} (a a') + Cover (inl _) (inr _) = Lift + Cover (inr _) (inl _) = Lift + Cover (inr b) (inr b') = Lift {j = ℓ-max ℓ'} (b b') + + reflCode : (c : A B) Cover c c + reflCode (inl a) = lift refl + reflCode (inr b) = lift refl + + encode : c c' c c' Cover c c' + encode c _ = J c' _ Cover c c') (reflCode c) + + encodeRefl : c encode c c refl reflCode c + encodeRefl c = JRefl c' _ Cover c c') (reflCode c) + + decode : c c' Cover c c' c c' + decode (inl a) (inl a') (lift p) = cong inl p + decode (inl a) (inr b') () + decode (inr b) (inl a') () + decode (inr b) (inr b') (lift q) = cong inr q + + decodeRefl : c decode c c (reflCode c) refl + decodeRefl (inl a) = refl + decodeRefl (inr b) = refl + + decodeEncode : c c' (p : c c') decode c c' (encode c c' p) p + decodeEncode c _ = + J c' p decode c c' (encode c c' p) p) + (cong (decode c c) (encodeRefl c) decodeRefl c) + + encodeDecode : c c' (d : Cover c c') encode c c' (decode c c' d) d + encodeDecode (inl a) (inl _) (lift d) = + J a' p encode (inl a) (inl a') (cong inl p) lift p) (encodeRefl (inl a)) d + encodeDecode (inr a) (inr _) (lift d) = + J a' p encode (inr a) (inr a') (cong inr p) lift p) (encodeRefl (inr a)) d + + Cover≃Path : c c' Cover c c' (c c') + Cover≃Path c c' = + isoToEquiv (iso (decode c c') (encode c c') (decodeEncode c c') (encodeDecode c c')) + + isOfHLevelCover : (n : HLevel) + isOfHLevel (suc (suc n)) A + isOfHLevel (suc (suc n)) B + c c' isOfHLevel (suc n) (Cover c c') + isOfHLevelCover n p q (inl a) (inl a') = isOfHLevelLift (suc n) (p a a') + isOfHLevelCover n p q (inl a) (inr b') = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p q (inr b) (inl a') = + isOfHLevelLift (suc n) (isProp→isOfHLevelSuc n isProp⊥) + isOfHLevelCover n p q (inr b) (inr b') = isOfHLevelLift (suc n) (q b b') + +isEmbedding-inl : isEmbedding (inl {A = A} {B = B}) +isEmbedding-inl w z = snd (compEquiv LiftEquiv (⊎Path.Cover≃Path (inl w) (inl z))) + +isEmbedding-inr : isEmbedding (inr {A = A} {B = B}) +isEmbedding-inr w z = snd (compEquiv LiftEquiv (⊎Path.Cover≃Path (inr w) (inr z))) + +isOfHLevel⊎ : (n : HLevel) + isOfHLevel (suc (suc n)) A + isOfHLevel (suc (suc n)) B + isOfHLevel (suc (suc n)) (A B) +isOfHLevel⊎ n lA lB c c' = + isOfHLevelRetract (suc n) + (⊎Path.encode c c') + (⊎Path.decode c c') + (⊎Path.decodeEncode c c') + (⊎Path.isOfHLevelCover n lA lB c c') + +isProp⊎ : isProp A isProp B (A B ) isProp (A B) +isProp⊎ propA _ _ (inl x) (inl y) i = inl (propA x y i) +isProp⊎ _ _ AB⊥ (inl x) (inr y) = ⊥.rec (AB⊥ x y) +isProp⊎ _ _ AB⊥ (inr x) (inl y) = ⊥.rec (AB⊥ y x) +isProp⊎ _ propB _ (inr x) (inr y) i = inr (propB x y i) + +isSet⊎ : isSet A isSet B isSet (A B) +isSet⊎ = isOfHLevel⊎ 0 + +isGroupoid⊎ : isGroupoid A isGroupoid B isGroupoid (A B) +isGroupoid⊎ = isOfHLevel⊎ 1 + +is2Groupoid⊎ : is2Groupoid A is2Groupoid B is2Groupoid (A B) +is2Groupoid⊎ = isOfHLevel⊎ 2 + +discrete⊎ : Discrete A Discrete B Discrete (A B) +discrete⊎ decA decB (inl a) (inl a') = + mapDec (cong inl) p q p (isEmbedding→Inj isEmbedding-inl _ _ q)) (decA a a') +discrete⊎ decA decB (inl a) (inr b') = no p lower (⊎Path.encode (inl a) (inr b') p)) +discrete⊎ decA decB (inr b) (inl a') = no ((λ p lower (⊎Path.encode (inr b) (inl a') p))) +discrete⊎ decA decB (inr b) (inr b') = + mapDec (cong inr) p q p (isEmbedding→Inj isEmbedding-inr _ _ q)) (decB b b') + +⊎Iso : Iso A C Iso B D Iso (A B) (C D) +fun (⊎Iso iac ibd) (inl x) = inl (iac .fun x) +fun (⊎Iso iac ibd) (inr x) = inr (ibd .fun x) +inv (⊎Iso iac ibd) (inl x) = inl (iac .inv x) +inv (⊎Iso iac ibd) (inr x) = inr (ibd .inv x) +rightInv (⊎Iso iac ibd) (inl x) = cong inl (iac .rightInv x) +rightInv (⊎Iso iac ibd) (inr x) = cong inr (ibd .rightInv x) +leftInv (⊎Iso iac ibd) (inl x) = cong inl (iac .leftInv x) +leftInv (⊎Iso iac ibd) (inr x) = cong inr (ibd .leftInv x) + +⊎-equiv : A C B D (A B) (C D) +⊎-equiv p q = isoToEquiv (⊎Iso (equivToIso p) (equivToIso q)) + +⊎-swap-Iso : Iso (A B) (B A) +fun ⊎-swap-Iso (inl x) = inr x +fun ⊎-swap-Iso (inr x) = inl x +inv ⊎-swap-Iso (inl x) = inr x +inv ⊎-swap-Iso (inr x) = inl x +rightInv ⊎-swap-Iso (inl _) = refl +rightInv ⊎-swap-Iso (inr _) = refl +leftInv ⊎-swap-Iso (inl _) = refl +leftInv ⊎-swap-Iso (inr _) = refl + +⊎-swap-≃ : A B B A +⊎-swap-≃ = isoToEquiv ⊎-swap-Iso + +⊎-assoc-Iso : Iso ((A B) C) (A (B C)) +fun ⊎-assoc-Iso (inl (inl x)) = inl x +fun ⊎-assoc-Iso (inl (inr x)) = inr (inl x) +fun ⊎-assoc-Iso (inr x) = inr (inr x) +inv ⊎-assoc-Iso (inl x) = inl (inl x) +inv ⊎-assoc-Iso (inr (inl x)) = inl (inr x) +inv ⊎-assoc-Iso (inr (inr x)) = inr x +rightInv ⊎-assoc-Iso (inl _) = refl +rightInv ⊎-assoc-Iso (inr (inl _)) = refl +rightInv ⊎-assoc-Iso (inr (inr _)) = refl +leftInv ⊎-assoc-Iso (inl (inl _)) = refl +leftInv ⊎-assoc-Iso (inl (inr _)) = refl +leftInv ⊎-assoc-Iso (inr _) = refl + +⊎-assoc-≃ : (A B) C A (B C) +⊎-assoc-≃ = isoToEquiv ⊎-assoc-Iso + +⊎-IdR-⊥-Iso : Iso (A ) A +fun ⊎-IdR-⊥-Iso (inl x) = x +inv ⊎-IdR-⊥-Iso x = inl x +rightInv ⊎-IdR-⊥-Iso _ = refl +leftInv ⊎-IdR-⊥-Iso (inl _) = refl + +⊎-IdL-⊥-Iso : Iso ( A) A +fun ⊎-IdL-⊥-Iso (inr x) = x +inv ⊎-IdL-⊥-Iso x = inr x +rightInv ⊎-IdL-⊥-Iso _ = refl +leftInv ⊎-IdL-⊥-Iso (inr _) = refl + +⊎-IdL-⊥*-Iso : ∀{} Iso (⊥* {} A) A +fun ⊎-IdL-⊥*-Iso (inr x) = x +inv ⊎-IdL-⊥*-Iso x = inr x +rightInv ⊎-IdL-⊥*-Iso _ = refl +leftInv ⊎-IdL-⊥*-Iso (inr _) = refl + +⊎-IdR-⊥*-Iso : ∀{} Iso (A ⊥* {}) A +fun ⊎-IdR-⊥*-Iso (inl x) = x +inv ⊎-IdR-⊥*-Iso x = inl x +rightInv ⊎-IdR-⊥*-Iso _ = refl +leftInv ⊎-IdR-⊥*-Iso (inl _) = refl + +⊎-IdR-⊥-≃ : A A +⊎-IdR-⊥-≃ = isoToEquiv ⊎-IdR-⊥-Iso + +⊎-IdL-⊥-≃ : A A +⊎-IdL-⊥-≃ = isoToEquiv ⊎-IdL-⊥-Iso + +⊎-IdR-⊥*-≃ : ∀{} A ⊥* {} A +⊎-IdR-⊥*-≃ = isoToEquiv ⊎-IdR-⊥*-Iso + +⊎-IdL-⊥*-≃ : ∀{} ⊥* {} A A +⊎-IdL-⊥*-≃ = isoToEquiv ⊎-IdL-⊥*-Iso + +Π⊎Iso : Iso ((x : A B) E x) (((a : A) E (inl a)) × ((b : B) E (inr b))) +fun Π⊎Iso f .fst a = f (inl a) +fun Π⊎Iso f .snd b = f (inr b) +inv Π⊎Iso (g1 , g2) (inl a) = g1 a +inv Π⊎Iso (g1 , g2) (inr b) = g2 b +rightInv Π⊎Iso (g1 , g2) i .fst a = g1 a +rightInv Π⊎Iso (g1 , g2) i .snd b = g2 b +leftInv Π⊎Iso f i (inl a) = f (inl a) +leftInv Π⊎Iso f i (inr b) = f (inr b) + +Σ⊎Iso : Iso (Σ (A B) E) ((Σ A a E (inl a))) (Σ B b E (inr b)))) +fun Σ⊎Iso (inl a , ea) = inl (a , ea) +fun Σ⊎Iso (inr b , eb) = inr (b , eb) +inv Σ⊎Iso (inl (a , ea)) = (inl a , ea) +inv Σ⊎Iso (inr (b , eb)) = (inr b , eb) +rightInv Σ⊎Iso (inl (a , ea)) = refl +rightInv Σ⊎Iso (inr (b , eb)) = refl +leftInv Σ⊎Iso (inl a , ea) = refl +leftInv Σ⊎Iso (inr b , eb) = refl + +×DistR⊎Iso : Iso (A × (B C)) ((A × B) (A × C)) +fun ×DistR⊎Iso (a , inl b) = inl (a , b) +fun ×DistR⊎Iso (a , inr c) = inr (a , c) +inv ×DistR⊎Iso (inl (a , b)) = a , inl b +inv ×DistR⊎Iso (inr (a , c)) = a , inr c +rightInv ×DistR⊎Iso (inl (a , b)) = refl +rightInv ×DistR⊎Iso (inr (a , c)) = refl +leftInv ×DistR⊎Iso (a , inl b) = refl +leftInv ×DistR⊎Iso (a , inr c) = refl + +Π⊎≃ : ((x : A B) E x) ((a : A) E (inl a)) × ((b : B) E (inr b)) +Π⊎≃ = isoToEquiv Π⊎Iso + +Σ⊎≃ : (Σ (A B) E) ((Σ A a E (inl a))) (Σ B b E (inr b)))) +Σ⊎≃ = isoToEquiv Σ⊎Iso + +⊎Monotone↪ : A C B D (A B) (C D) +⊎Monotone↪ (f , embf) (g , embg) = (map f g) , emb + where coverToMap : x y ⊎Path.Cover x y + ⊎Path.Cover (map f g x) (map f g y) + coverToMap (inl _) (inl _) cover = lift (cong f (lower cover)) + coverToMap (inr _) (inr _) cover = lift (cong g (lower cover)) + + equiv : x y isEquiv (coverToMap x y) + equiv (inl a₀) (inl a₁) + = ((invEquiv LiftEquiv) + ∙ₑ ((cong f) , (embf a₀ a₁)) + ∙ₑ LiftEquiv) .snd + equiv (inl a₀) (inr b₁) .equiv-proof () + equiv (inr b₀) (inl a₁) .equiv-proof () + equiv (inr b₀) (inr b₁) + = ((invEquiv LiftEquiv) + ∙ₑ ((cong g) , (embg b₀ b₁)) + ∙ₑ LiftEquiv) .snd + + lemma : x y + (p : x y) + cong (map f g) p + ⊎Path.decode + (map f g x) + (map f g y) + (coverToMap x y (⊎Path.encode x y p)) + lemma (inl a₀) _ + = J y p + cong (map f g) p + ⊎Path.decode (map f g (inl a₀)) + (map f g y) + (coverToMap (inl a₀) y + (⊎Path.encode (inl a₀) y p))) + (sym $ cong (cong inl) (cong (cong f) (transportRefl _))) + lemma (inr b₀) _ + = J y p + cong (map f g) p + ⊎Path.decode + (map f g (inr b₀)) + (map f g y) + (coverToMap (inr b₀) y (⊎Path.encode (inr b₀) y p))) + (sym $ cong (cong inr) (cong (cong g) (transportRefl _))) + + emb : isEmbedding (map f g) + emb x y = subst eq isEquiv eq) + (sym (funExt (lemma x y))) + ((x y ≃⟨ invEquiv (⊎Path.Cover≃Path x y) + ⊎Path.Cover x y ≃⟨ (coverToMap x y) , (equiv x y) + ⊎Path.Cover + (map f g x) + (map f g y) ≃⟨ ⊎Path.Cover≃Path + (map f g x) + (map f g y) + map f g x map f g y ) .snd) diff --git a/src/docs/Cubical.Data.Sum.html b/src/docs/Cubical.Data.Sum.html new file mode 100644 index 0000000..3864f75 --- /dev/null +++ b/src/docs/Cubical.Data.Sum.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Sum where + +open import Cubical.Data.Sum.Base public +open import Cubical.Data.Sum.Properties public diff --git a/src/docs/Cubical.Data.Unit.Base.html b/src/docs/Cubical.Data.Unit.Base.html new file mode 100644 index 0000000..19f317b --- /dev/null +++ b/src/docs/Cubical.Data.Unit.Base.html @@ -0,0 +1,25 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Unit.Base where +open import Cubical.Foundations.Prelude + + +-- Obtain Unit +open import Agda.Builtin.Unit public + renaming ( to Unit ) + +-- Universe polymorphic version +Unit* : { : Level} Type +Unit* = Lift Unit + +pattern tt* = lift tt + +-- Pointed version +Unit*∙ : {} Σ[ X Type ] X +fst Unit*∙ = Unit* +snd Unit*∙ = tt* + +-- Universe polymorphic version without definitional equality +-- Allows us to "lock" proofs. See "Locking, unlocking" in +-- https://coq.inria.fr/refman/proof-engine/ssreflect-proof-language.html +data lockUnit {} : Type where + unlock : lockUnit diff --git a/src/docs/Cubical.Data.Unit.Properties.html b/src/docs/Cubical.Data.Unit.Properties.html new file mode 100644 index 0000000..ae6756e --- /dev/null +++ b/src/docs/Cubical.Data.Unit.Properties.html @@ -0,0 +1,121 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Unit.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Data.Nat +open import Cubical.Data.Unit.Base +open import Cubical.Data.Prod.Base + +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence + +open import Cubical.Reflection.StrictEquiv + +open Iso + +private + variable + ℓ' : Level + +isContrUnit : isContr Unit +isContrUnit = tt , λ {tt refl} + +isPropUnit : isProp Unit +isPropUnit _ _ i = tt -- definitionally equal to: isContr→isProp isContrUnit + +isSetUnit : isSet Unit +isSetUnit = isProp→isSet isPropUnit + +isOfHLevelUnit : (n : HLevel) isOfHLevel n Unit +isOfHLevelUnit n = isContr→isOfHLevel n isContrUnit + +module _ (A : Type ) where + UnitToType≃ : (Unit A) A + unquoteDef UnitToType≃ = defStrictEquiv UnitToType≃ f f _) const + +UnitToTypePath : {} (A : Type ) (Unit A) A +UnitToTypePath A = ua (UnitToType≃ A) + +module _ (A : Unit Type ) where + + open Iso + + ΠUnitIso : Iso ((x : Unit) A x) (A tt) + fun ΠUnitIso f = f tt + inv ΠUnitIso a tt = a + rightInv ΠUnitIso a = refl + leftInv ΠUnitIso f = refl + + ΠUnit : ((x : Unit) A x) A tt + ΠUnit = isoToEquiv ΠUnitIso + +module _ (A : Unit* {} Type ℓ') where + + open Iso + + ΠUnit*Iso : Iso ((x : Unit*) A x) (A tt*) + fun ΠUnit*Iso f = f tt* + inv ΠUnit*Iso a tt* = a + rightInv ΠUnit*Iso a = refl + leftInv ΠUnit*Iso f = refl + + ΠUnit* : ((x : Unit*) A x) A tt* + ΠUnit* = isoToEquiv ΠUnit*Iso + +fiberUnitIso : {A : Type } Iso (fiber (a : A) tt) tt) A +fun fiberUnitIso = fst +inv fiberUnitIso a = a , refl +rightInv fiberUnitIso _ = refl +leftInv fiberUnitIso _ = refl + +isContr→Iso2 : {A : Type } {B : Type ℓ'} isContr A Iso (A B) B +fun (isContr→Iso2 iscontr) f = f (fst iscontr) +inv (isContr→Iso2 iscontr) b _ = b +rightInv (isContr→Iso2 iscontr) _ = refl +leftInv (isContr→Iso2 iscontr) f = funExt λ x cong f (snd iscontr x) + +diagonal-unit : Unit Unit × Unit +diagonal-unit = isoToPath (iso x tt , tt) x tt) {(tt , tt) i tt , tt}) λ {tt i tt}) + +fibId : (A : Type ) (fiber (x : A) tt) tt) A +fibId A = ua e + where + unquoteDecl e = declStrictEquiv e fst a a , refl) + +isContr→≃Unit : {A : Type } isContr A A Unit +isContr→≃Unit contr = isoToEquiv (iso _ tt) _ fst contr) _ refl) λ _ snd contr _) + +isContr→≡Unit : {A : Type₀} isContr A A Unit +isContr→≡Unit contr = ua (isContr→≃Unit contr) + +isContrUnit* : {} isContr (Unit* {}) +isContrUnit* = tt* , λ _ refl + +isPropUnit* : {} isProp (Unit* {}) +isPropUnit* _ _ = refl + +isSetUnit* : {} isSet (Unit* {}) +isSetUnit* _ _ _ _ = refl + +isOfHLevelUnit* : {} (n : HLevel) isOfHLevel n (Unit* {}) +isOfHLevelUnit* zero = tt* , λ _ refl +isOfHLevelUnit* (suc zero) _ _ = refl +isOfHLevelUnit* (suc (suc zero)) _ _ _ _ _ _ = tt* +isOfHLevelUnit* (suc (suc (suc n))) = isOfHLevelPlus 3 (isOfHLevelUnit* n) + +Unit≃Unit* : {} Unit Unit* {} +Unit≃Unit* = invEquiv (isContr→≃Unit isContrUnit*) + +isContr→≃Unit* : {A : Type } isContr A A Unit* {} +isContr→≃Unit* contr = compEquiv (isContr→≃Unit contr) Unit≃Unit* + +isContr→≡Unit* : {A : Type } isContr A A Unit* +isContr→≡Unit* contr = ua (isContr→≃Unit* contr) diff --git a/src/docs/Cubical.Data.Unit.html b/src/docs/Cubical.Data.Unit.html new file mode 100644 index 0000000..14cb3a6 --- /dev/null +++ b/src/docs/Cubical.Data.Unit.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Unit where + +open import Cubical.Data.Unit.Base public +open import Cubical.Data.Unit.Properties public diff --git a/src/docs/Cubical.Data.Vec.Base.html b/src/docs/Cubical.Data.Vec.Base.html new file mode 100644 index 0000000..9c6cc80 --- /dev/null +++ b/src/docs/Cubical.Data.Vec.Base.html @@ -0,0 +1,67 @@ +{- Definition of vectors. Inspired by the Agda Standard Library -} + +{-# OPTIONS --safe #-} +module Cubical.Data.Vec.Base where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat +open import Cubical.Data.FinData.Base + +private + variable + ℓ' ℓ'' : Level + A : Type + B : Type ℓ' + +infixr 5 _∷_ + +data Vec (A : Type ) : Type where + [] : Vec A zero + _∷_ : {n} (x : A) (xs : Vec A n) Vec A (suc n) + + +-- Basic functions + +length : {n} Vec A n +length {n = n} _ = n + +head : {n} Vec A (1 + n) A +head (x xs) = x + +tail : {n} Vec A (1 + n) Vec A n +tail (x xs) = xs + +map : {A : Type } {B : Type ℓ'} {n} (A B) Vec A n Vec B n +map f [] = [] +map f (x xs) = f x map f xs + +replicate : {n} {A : Type } A Vec A n +replicate {n = zero} x = [] +replicate {n = suc n} x = x replicate x + +zipWith : {A : Type } {B : Type ℓ'} {C : Type ℓ''} {n : } + (A B C) Vec A n Vec B n Vec C n +zipWith {n = zero} _∗_ [] [] = [] +zipWith {n = suc n} _∗_ (x xs) (y ys) = (x y) zipWith _∗_ xs ys + +foldr : {n : } (A B B) B Vec A n B +foldr {n = zero} _∗_ x [] = x +foldr {n = suc n} _∗_ x (y xs) = y (foldr _∗_ x xs) + +-- Concatenation + +infixr 5 _++_ + +_++_ : {m n} Vec A m Vec A n Vec A (m + n) +[] ++ ys = ys +(x xs) ++ ys = x (xs ++ ys) + +concat : {m n} Vec (Vec A m) n Vec A (n · m) +concat [] = [] +concat (xs xss) = xs ++ concat xss + +lookup : {n} {A : Type } Fin n Vec A n A +lookup zero (x xs) = x +lookup (suc i) (x xs) = lookup i xs + diff --git a/src/docs/Cubical.Data.Vec.NAry.html b/src/docs/Cubical.Data.Vec.NAry.html new file mode 100644 index 0000000..226129e --- /dev/null +++ b/src/docs/Cubical.Data.Vec.NAry.html @@ -0,0 +1,55 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Vec.NAry where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Nat +open import Cubical.Data.Vec.Base + +private + variable + ℓ' : Level + A : Type + B : Type ℓ' + +nAryLevel : Level Level Level +nAryLevel ℓ₁ ℓ₂ zero = ℓ₂ +nAryLevel ℓ₁ ℓ₂ (suc n) = ℓ-max ℓ₁ (nAryLevel ℓ₁ ℓ₂ n) + +nAryOp : (n : ) Type Type ℓ' Type (nAryLevel ℓ' n) +nAryOp zero A B = B +nAryOp (suc n) A B = A nAryOp n A B + + +_$ⁿ_ : {n} nAryOp n A B (Vec A n B) +f $ⁿ [] = f +f $ⁿ (x xs) = f x $ⁿ xs + +curryⁿ : {n} (Vec A n B) nAryOp n A B +curryⁿ {n = zero} f = f [] +curryⁿ {n = suc n} f x = curryⁿ xs f (x xs)) + +$ⁿ-curryⁿ : {n} (f : Vec A n B) _$ⁿ_ (curryⁿ f) f +$ⁿ-curryⁿ {n = zero} f = funExt λ { [] refl } +$ⁿ-curryⁿ {n = suc n} f = funExt λ { (x xs) i $ⁿ-curryⁿ {n = n} ys f (x ys)) i xs} + +curryⁿ-$ⁿ : {n} (f : nAryOp { = } {ℓ' = ℓ'} n A B) curryⁿ {A = A} {B = B} (_$ⁿ_ f) f +curryⁿ-$ⁿ {n = zero} f = refl +curryⁿ-$ⁿ {n = suc n} f = funExt λ x curryⁿ-$ⁿ {n = n} (f x) + +nAryOp≃VecFun : {n} nAryOp n A B (Vec A n B) +nAryOp≃VecFun {n = n} = isoToEquiv f + where + f : Iso (nAryOp n A B) (Vec A n B) + Iso.fun f = _$ⁿ_ + Iso.inv f = curryⁿ + Iso.rightInv f = $ⁿ-curryⁿ + Iso.leftInv f = curryⁿ-$ⁿ {n = n} + +-- In order to apply ua to nAryOp≃VecFun we probably need to change +-- the base-case of nAryLevel to "ℓ-max ℓ₁ ℓ₂". This will make it +-- necessary to add lots of Lifts in zero cases so it's not done yet, +-- but if the Path is ever needed then it might be worth to do. diff --git a/src/docs/Cubical.Data.Vec.Properties.html b/src/docs/Cubical.Data.Vec.Properties.html new file mode 100644 index 0000000..ad16f24 --- /dev/null +++ b/src/docs/Cubical.Data.Vec.Properties.html @@ -0,0 +1,141 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Vec.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Univalence + +import Cubical.Data.Empty as +open import Cubical.Data.Unit +open import Cubical.Data.Nat +open import Cubical.Data.Sigma +open import Cubical.Data.Sum +open import Cubical.Data.Vec.Base +open import Cubical.Data.FinData +open import Cubical.Relation.Nullary + +open Iso + +private + variable + : Level + A : Type + + +-- This is really cool! +-- Compare with: https://github.com/agda/agda-stdlib/blob/master/src/Data/Vec/Properties/WithK.agda#L32 +++-assoc : {m n k} (xs : Vec A m) (ys : Vec A n) (zs : Vec A k) + PathP i Vec A (+-assoc m n k (~ i))) ((xs ++ ys) ++ zs) (xs ++ ys ++ zs) +++-assoc {m = zero} [] ys zs = refl +++-assoc {m = suc m} (x xs) ys zs i = x ++-assoc xs ys zs i + + +-- Equivalence between Fin n → A and Vec A n +FinVec→Vec : {n : } FinVec A n Vec A n +FinVec→Vec {n = zero} xs = [] +FinVec→Vec {n = suc _} xs = xs zero FinVec→Vec x xs (suc x)) + +Vec→FinVec : {n : } Vec A n FinVec A n +Vec→FinVec xs f = lookup f xs + +FinVec→Vec→FinVec : {n : } (xs : FinVec A n) Vec→FinVec (FinVec→Vec xs) xs +FinVec→Vec→FinVec {n = zero} xs = funExt λ f ⊥.rec (¬Fin0 f) +FinVec→Vec→FinVec {n = suc n} xs = funExt goal + where + goal : (f : Fin (suc n)) + Vec→FinVec (xs zero FinVec→Vec x xs (suc x))) f xs f + goal zero = refl + goal (suc f) i = FinVec→Vec→FinVec x xs (suc x)) i f + +Vec→FinVec→Vec : {n : } (xs : Vec A n) FinVec→Vec (Vec→FinVec xs) xs +Vec→FinVec→Vec {n = zero} [] = refl +Vec→FinVec→Vec {n = suc n} (x xs) i = x Vec→FinVec→Vec xs i + +FinVecIsoVec : (n : ) Iso (FinVec A n) (Vec A n) +FinVecIsoVec n = iso FinVec→Vec Vec→FinVec Vec→FinVec→Vec FinVec→Vec→FinVec + +FinVec≃Vec : (n : ) FinVec A n Vec A n +FinVec≃Vec n = isoToEquiv (FinVecIsoVec n) + +FinVec≡Vec : (n : ) FinVec A n Vec A n +FinVec≡Vec n = ua (FinVec≃Vec n) + +isContrVec0 : isContr (Vec A 0) +isContrVec0 = [] , λ { [] refl } + +-- encode - decode Vec +module VecPath {A : Type } + where + + code : {n : } (v v' : Vec A n) Type + code [] [] = Unit* + code (a v) (a' v') = (a a') × (v v') + + -- encode + reflEncode : {n : } (v : Vec A n) code v v + reflEncode [] = tt* + reflEncode (a v) = refl , refl + + encode : {n : } (v v' : Vec A n) (v v') code v v' + encode v v' p = J v' _ code v v') (reflEncode v) p + + encodeRefl : {n : } (v : Vec A n) encode v v refl reflEncode v + encodeRefl v = JRefl v' _ code v v') (reflEncode v) + + -- decode + decode : {n : } (v v' : Vec A n) (r : code v v') (v v') + decode [] [] _ = refl + decode (a v) (a' v') (p , q) = cong₂ _∷_ p q + + decodeRefl : {n : } (v : Vec A n) decode v v (reflEncode v) refl + decodeRefl [] = refl + decodeRefl (a v) = refl + + -- equiv + ≡Vec≃codeVec : {n : } (v v' : Vec A n) (v v') (code v v') + ≡Vec≃codeVec v v' = isoToEquiv is + where + is : Iso (v v') (code v v') + fun is = encode v v' + inv is = decode v v' + rightInv is = sect v v' + where + sect : {n : } (v v' : Vec A n) (r : code v v') + encode v v' (decode v v' r) r + sect [] [] tt* = encodeRefl [] + sect (a v) (a' v') (p , q) = J a' p encode (a v) (a' v') (decode (a v) (a' v') (p , q)) (p , q)) + (J v' q encode (a v) (a v') (decode (a v) (a v') (refl , q)) (refl , q)) + (encodeRefl (a v)) q) p + leftInv is = retr v v' + where + retr : {n : } (v v' : Vec A n) (p : v v') + decode v v' (encode v v' p) p + retr v v' p = J v' p decode v v' (encode v v' p) p) + (cong (decode v v) (encodeRefl v) decodeRefl v) p + + + isOfHLevelVec : (h : HLevel) (n : ) + isOfHLevel (suc (suc h)) A isOfHLevel (suc (suc h)) (Vec A n) + isOfHLevelVec h zero ofLevelA [] [] = isOfHLevelRespectEquiv (suc h) (invEquiv (≡Vec≃codeVec [] [])) + (isOfHLevelUnit* (suc h)) + isOfHLevelVec h (suc n) ofLevelA (x v) (x' v') = isOfHLevelRespectEquiv (suc h) (invEquiv (≡Vec≃codeVec _ _)) + (isOfHLevelΣ (suc h) (ofLevelA x x') _ isOfHLevelVec h n ofLevelA v v')) + + + discreteA→discreteVecA : Discrete A (n : ) Discrete (Vec A n) + discreteA→discreteVecA DA zero [] [] = yes refl + discreteA→discreteVecA DA (suc n) (a v) (a' v') with (DA a a') | (discreteA→discreteVecA DA n v v') + ... | yes p | yes q = yes (invIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) (p , q)) + ... | yes p | no ¬q = no r ¬q (snd (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) + ... | no ¬p | yes q = no r ¬p (fst (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) + ... | no ¬p | no ¬q = no r ¬q (snd (funIsEq (snd (≡Vec≃codeVec (a v) (a' v'))) r))) + + ≢-∷ : {m : } (Discrete A) (a : A) (v : Vec A m) (a' : A) (v' : Vec A m) + (a v a' v' ⊥.⊥) (a a' ⊥.⊥) (v v' ⊥.⊥) + ≢-∷ {m} discreteA a v a' v' ¬r with (discreteA a a') + | (discreteA→discreteVecA discreteA m v v') + ... | yes p | yes q = ⊥.rec (¬r (cong₂ _∷_ p q)) + ... | yes p | no ¬q = inr ¬q + ... | no ¬p | y = inl ¬p diff --git a/src/docs/Cubical.Data.Vec.html b/src/docs/Cubical.Data.Vec.html new file mode 100644 index 0000000..c2df02a --- /dev/null +++ b/src/docs/Cubical.Data.Vec.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Data.Vec where + +open import Cubical.Data.Vec.Base public +open import Cubical.Data.Vec.Properties public diff --git a/src/docs/Cubical.Foundations.CartesianKanOps.html b/src/docs/Cubical.Foundations.CartesianKanOps.html new file mode 100644 index 0000000..f53387f --- /dev/null +++ b/src/docs/Cubical.Foundations.CartesianKanOps.html @@ -0,0 +1,180 @@ +-- This file derives some of the Cartesian Kan operations using transp +{-# OPTIONS --safe #-} +module Cubical.Foundations.CartesianKanOps where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Transport + +coe0→1 : {} (A : I Type ) A i0 A i1 +coe0→1 A a = transp (\ i A i) i0 a + +-- "coe filler" +coe0→i : {} (A : I Type ) (i : I) A i0 A i +coe0→i A i a = transp j A (i j)) (~ i) a + +-- Check the equations for the coe filler +coe0→i1 : {} (A : I Type ) (a : A i0) coe0→i A i1 a coe0→1 A a +coe0→i1 A a = refl + +coe0→i0 : {} (A : I Type ) (a : A i0) coe0→i A i0 a a +coe0→i0 A a = refl + +-- coe backwards +coe1→0 : {} (A : I Type ) A i1 A i0 +coe1→0 A a = transp i A (~ i)) i0 a + +-- coe backwards filler +coe1→i : {} (A : I Type ) (i : I) A i1 A i +coe1→i A i a = transp j A (i ~ j)) i a + +-- Check the equations for the coe backwards filler +coe1→i0 : {} (A : I Type ) (a : A i1) coe1→i A i0 a coe1→0 A a +coe1→i0 A a = refl + +coe1→i1 : {} (A : I Type ) (a : A i1) coe1→i A i1 a a +coe1→i1 A a = refl + +-- "squeezeNeg" +coei→0 : {} (A : I Type ) (i : I) A i A i0 +coei→0 A i a = transp j A (i ~ j)) (~ i) a + +coei0→0 : {} (A : I Type ) (a : A i0) coei→0 A i0 a a +coei0→0 A a = refl + +coei1→0 : {} (A : I Type ) (a : A i1) coei→0 A i1 a coe1→0 A a +coei1→0 A a = refl + +-- "master coe" +-- defined as the filler of coei→0, coe0→i, and coe1→i +-- unlike in cartesian cubes, we don't get coei→i = id definitionally +coei→j : {} (A : I Type ) (i j : I) A i A j +coei→j A i j a = + fill (\ i A i) + j λ { (i = i0) coe0→i A j a + ; (i = i1) coe1→i A j a + }) + (inS (coei→0 A i a)) + j + +-- "squeeze" +-- this is just defined as the composite face of the master coe +coei→1 : {} (A : I Type ) (i : I) A i A i1 +coei→1 A i a = coei→j A i i1 a + +coei0→1 : {} (A : I Type ) (a : A i0) coei→1 A i0 a coe0→1 A a +coei0→1 A a = refl + +coei1→1 : {} (A : I Type ) (a : A i1) coei→1 A i1 a a +coei1→1 A a = refl + +-- equations for "master coe" +coei→i0 : {} (A : I Type ) (i : I) (a : A i) coei→j A i i0 a coei→0 A i a +coei→i0 A i a = refl + +coei0→i : {} (A : I Type ) (i : I) (a : A i0) coei→j A i0 i a coe0→i A i a +coei0→i A i a = refl + +coei→i1 : {} (A : I Type ) (i : I) (a : A i) coei→j A i i1 a coei→1 A i a +coei→i1 A i a = refl + +coei1→i : {} (A : I Type ) (i : I) (a : A i1) coei→j A i1 i a coe1→i A i a +coei1→i A i a = refl + +-- only non-definitional equation, but definitional at the ends +coei→i : {} (A : I Type ) (i : I) (a : A i) coei→j A i i a a +coei→i A i a j = + comp k A (i (j k))) + k λ + { (i = i0) a + ; (i = i1) coe1→i A (j k) a + ; (j = i1) a }) + (transpFill {A = A i0} (~ i) t inS (A (i ~ t))) a (~ j)) + +coe0→0 : {} (A : I Type ) (a : A i0) coei→i A i0 a refl +coe0→0 A a = refl + +coe1→1 : {} (A : I Type ) (a : A i1) coei→i A i1 a refl +coe1→1 A a = refl + +-- coercion when there already exists a path +coePath : {} (A : I Type ) (p : (i : I) A i) (i j : I) coei→j A i j (p i) p j +coePath A p i j = + hcomp k λ + { (i = i0)(j = i0) rUnit refl (~ k) + ; (i = i1)(j = i1) rUnit refl (~ k) }) + (diag coei→i A j (p j)) + where + diag : coei→j A i j (p i) coei→j A j j (p j) + diag k = coei→j A _ j (p ((j (i ~ k)) (i (j k)))) + +coePathi0 : {} (A : I Type ) (p : (i : I) A i) coePath A p i0 i0 refl +coePathi0 A p = refl + +coePathi1 : {} (A : I Type ) (p : (i : I) A i) coePath A p i1 i1 refl +coePathi1 A p = refl + +-- do the same for fill + +fill1→i : {} (A : i Type ) + {φ : I} + (u : i Partial φ (A i)) + (u1 : A i1 [ φ u i1 ]) + --------------------------- + (i : I) A i +fill1→i A {φ = φ} u u1 i = + comp j A (i ~ j)) + j λ { (φ = i1) u (i ~ j) 1=1 + ; (i = i1) outS u1 }) + (outS u1) + +filli→0 : {} (A : i Type ) + {φ : I} + (u : i Partial φ (A i)) + (i : I) + (ui : A i [ φ u i ]) + --------------------------- + A i0 +filli→0 A {φ = φ} u i ui = + comp j A (i ~ j)) + j λ { (φ = i1) u (i ~ j) 1=1 + ; (i = i0) outS ui }) + (outS ui) + +filli→j : {} (A : i Type ) + {φ : I} + (u : i Partial φ (A i)) + (i : I) + (ui : A i [ φ u i ]) + --------------------------- + (j : I) A j +filli→j A {φ = φ} u i ui j = + fill (\ i A i) + j λ { (φ = i1) u j 1=1 + ; (i = i0) fill (\ i A i) (\ i u i) ui j + ; (i = i1) fill1→i A u ui j + }) + (inS (filli→0 A u i ui)) + j + +-- We can reconstruct fill from hfill, coei→j, and the path coei→i ≡ id. +-- The definition does not rely on the computational content of the coei→i path. +fill' : {} (A : I Type ) + {φ : I} + (u : i Partial φ (A i)) + (u0 : A i0 [ φ u i0 ]) + --------------------------- + (i : I) A i [ φ u i ] +fill' A {φ = φ} u u0 i = + inS (hcomp j λ {(φ = i1) coei→i A i (u i 1=1) j; (i = i0) coei→i A i (outS u0) j}) t) + where + t : A i + t = hfill {φ = φ} j v coei→j A j i (u j v)) (inS (coe0→i A i (outS u0))) i + +fill'-cap : {} (A : I Type ) + {φ : I} + (u : i Partial φ (A i)) + (u0 : A i0 [ φ u i0 ]) + --------------------------- + outS (fill' A u u0 i0) outS (u0) +fill'-cap A u u0 = refl diff --git a/src/docs/Cubical.Foundations.Equiv.Base.html b/src/docs/Cubical.Foundations.Equiv.Base.html new file mode 100644 index 0000000..9029b59 --- /dev/null +++ b/src/docs/Cubical.Foundations.Equiv.Base.html @@ -0,0 +1,65 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.Base where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude + +open import Cubical.Core.Glue public + using ( isEquiv ; equiv-proof ; _≃_ ; equivFun ; equivProof ) + +fiber : { ℓ'} {A : Type } {B : Type ℓ'} (f : A B) (y : B) Type (ℓ-max ℓ') +fiber {A = A} f y = Σ[ x A ] f x y + +-- Helper function for constructing equivalences from pairs (f,g) that cancel each other up to definitional +-- equality. For such (f,g), the result type simplifies to isContr (fiber f b). +strictContrFibers : { ℓ'} {A : Type } {B : Type ℓ'} {f : A B} (g : B A) (b : B) + Σ[ t fiber f (f (g b)) ] + ((t' : fiber f b) Path (fiber f (f (g b))) t (g (f (t' .fst)) , cong (f g) (t' .snd))) +strictContrFibers {f = f} g b .fst = (g b , refl) +strictContrFibers {f = f} g b .snd (a , p) i = (g (p (~ i)) , λ j f (g (p (~ i j)))) + +-- The identity equivalence +idIsEquiv : {} (A : Type ) isEquiv (idfun A) +idIsEquiv A .equiv-proof = strictContrFibers (idfun A) + +idEquiv : {} (A : Type ) A A +idEquiv A .fst = idfun A +idEquiv A .snd = idIsEquiv A + +-- the definition using Π-type +isEquiv' : {}{ℓ'}{A : Type }{B : Type ℓ'} (A B) Type (ℓ-max ℓ') +isEquiv' {B = B} f = (y : B) isContr (fiber f y) + +-- Transport along a line of types A, constant on some extent φ, is an equivalence. +isEquivTransp : { : I Level} (A : (i : I) Type ( i)) (φ : I) isEquiv (transp j A (φ j)) φ) +isEquivTransp A φ = u₁ where + -- NB: The transport in question is the `coei→1` or `squeeze` operation mentioned + -- at `Cubical.Foundations.CartesianKanOps` and + -- https://1lab.dev/1Lab.Path.html#coei%E2%86%921 + coei→1 : A φ A i1 + coei→1 = transp j A (φ j)) φ + + -- A line of types, interpolating between propositions: + -- (k = i0): the identity function is an equivalence + -- (k = i1): transport along A is an equivalence + γ : (k : I) Type _ + γ k = isEquiv (transp j A (φ (j k))) (φ ~ k)) + + _ : γ i0 isEquiv (idfun (A φ)) + _ = refl + + _ : γ i1 isEquiv coei→1 + _ = refl + + -- We have proof that the identity function *is* an equivalence, + u₀ : γ i0 + u₀ = idIsEquiv (A φ) + + -- and by transporting that evidence along γ, we prove that + -- transporting along A is an equivalence, too. + u₁ : γ i1 + u₁ = transp γ φ u₀ + +transpEquiv : { : I Level} (A : (i : I) Type ( i)) φ A φ A i1 +transpEquiv A φ .fst = transp j A (φ j)) φ +transpEquiv A φ .snd = isEquivTransp A φ diff --git a/src/docs/Cubical.Foundations.Equiv.Fiberwise.html b/src/docs/Cubical.Foundations.Equiv.Fiberwise.html new file mode 100644 index 0000000..7c20526 --- /dev/null +++ b/src/docs/Cubical.Foundations.Equiv.Fiberwise.html @@ -0,0 +1,99 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.Fiberwise where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma + +private + variable + ℓ' ℓ'' : Level + +module _ {A : Type } (P : A Type ℓ') (Q : A Type ℓ'') + (f : x P x Q x) + where + private + total : (Σ A P) (Σ A Q) + total = (\ p p .fst , f (p .fst) (p .snd)) + + -- Thm 4.7.6 + fibers-total : {xv} Iso (fiber total (xv)) (fiber (f (xv .fst)) (xv .snd)) + fibers-total {xv} = iso h g h-g g-h + where + h : {xv} fiber total xv fiber (f (xv .fst)) (xv .snd) + h {xv} (p , eq) = J (\ xv eq fiber (f (xv .fst)) (xv .snd)) ((p .snd) , refl) eq + g : {xv} fiber (f (xv .fst)) (xv .snd) fiber total xv + g {xv} (p , eq) = (xv .fst , p) , (\ i _ , eq i) + h-g : {xv} y h {xv} (g {xv} y) y + h-g {x , v} (p , eq) = J _ eq₁ h (g (p , eq₁)) (p , eq₁)) (JRefl xv₁ eq₁ fiber (f (xv₁ .fst)) (xv₁ .snd)) ((p , refl))) (eq) + g-h : {xv} y g {xv} (h {xv} y) y + g-h {xv} ((a , p) , eq) = J _ eq₁ g (h ((a , p) , eq₁)) ((a , p) , eq₁)) + (cong g (JRefl xv₁ eq₁ fiber (f (xv₁ .fst)) (xv₁ .snd)) (p , refl))) + eq + -- Thm 4.7.7 (fiberwise equivalences) + fiberEquiv : ([tf] : isEquiv total) + x isEquiv (f x) + fiberEquiv [tf] x .equiv-proof y = isContrRetract (fibers-total .Iso.inv) (fibers-total .Iso.fun) (fibers-total .Iso.rightInv) + ([tf] .equiv-proof (x , y)) + + totalEquiv : (fx-equiv : x isEquiv (f x)) + isEquiv total + totalEquiv fx-equiv .equiv-proof (x , v) = isContrRetract (fibers-total .Iso.fun) (fibers-total .Iso.inv) (fibers-total .Iso.leftInv) + (fx-equiv x .equiv-proof v) + + +module _ {U : Type } (_~_ : U U Type ℓ') + (idTo~ : {A B} A B A ~ B) + (c : A ∃![ X U ] (A ~ X)) + where + + isContrToUniv : {A B} isEquiv (idTo~ {A} {B}) + isContrToUniv {A} {B} + = fiberEquiv z A z) z A ~ z) (\ B idTo~ {A} {B}) + { .equiv-proof y + isContrΣ (isContrSingl _) + \ a isContr→isContrPath (c A) _ _ + }) + B + + +{- + The following is called fundamental theorem of identity types in Egbert Rijke's + introduction to homotopy type theory. +-} +recognizeId : {A : Type } {a : A} (Eq : A Type ℓ') + Eq a + isContr (Σ _ Eq) + (x : A) (a x) (Eq x) +recognizeId {A = A} {a = a} Eq eqRefl eqContr x = (fiberMap x) , (isEquivFiberMap x) + where + fiberMap : (x : A) a x Eq x + fiberMap x = J x p Eq x) eqRefl + + mapOnSigma : Σ[ x A ] a x Σ _ Eq + mapOnSigma pair = fst pair , fiberMap (fst pair) (snd pair) + + equivOnSigma : (x : A) isEquiv mapOnSigma + equivOnSigma x = isEquivFromIsContr mapOnSigma (isContrSingl a) eqContr + + isEquivFiberMap : (x : A) isEquiv (fiberMap x) + isEquivFiberMap = fiberEquiv x a x) Eq fiberMap (equivOnSigma x) + +fundamentalTheoremOfId : {A : Type } (Eq : A A Type ℓ') + ((x : A) Eq x x) + ((x : A) isContr (Σ[ y A ] Eq x y)) + (x y : A) (x y) (Eq x y) +fundamentalTheoremOfId Eq eqRefl eqContr x = recognizeId (Eq x) (eqRefl x) (eqContr x) + +fundamentalTheoremOfIdβ : + {A : Type } (Eq : A A Type ℓ') + (eqRefl : (x : A) Eq x x) + (eqContr : (x : A) isContr (Σ[ y A ] Eq x y)) + (x : A) + fst (fundamentalTheoremOfId Eq eqRefl eqContr x x) refl eqRefl x +fundamentalTheoremOfIdβ Eq eqRefl eqContr x = JRefl y p Eq x y) (eqRefl x) diff --git a/src/docs/Cubical.Foundations.Equiv.HalfAdjoint.html b/src/docs/Cubical.Foundations.Equiv.HalfAdjoint.html new file mode 100644 index 0000000..f9cd018 --- /dev/null +++ b/src/docs/Cubical.Foundations.Equiv.HalfAdjoint.html @@ -0,0 +1,138 @@ +{- + +Half adjoint equivalences ([HAEquiv]) + +- Iso to HAEquiv ([iso→HAEquiv]) +- Equiv to HAEquiv ([equiv→HAEquiv]) +- Cong is an equivalence ([congEquiv]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.HalfAdjoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws + +private + variable + ℓ' : Level + A : Type + B : Type ℓ' + +record isHAEquiv { ℓ'} {A : Type } {B : Type ℓ'} (f : A B) : Type (ℓ-max ℓ') where + field + g : B A + linv : a g (f a) a + rinv : b f (g b) b + com : a cong f (linv a) rinv (f a) + + com-op : b cong g (rinv b) linv (g b) + com-op b = subst b cong g (rinv b) linv (g b)) (rinv b) (helper (g b)) + where + helper : a cong g (rinv (f a)) linv (g (f a)) + helper a i j = hcomp k λ { (i = i0) g (rinv (f a) (j ~ k)) + ; (i = i1) linv (linv a (~ k)) j + ; (j = i0) g (com a (~ i) (~ k)) + ; (j = i1) linv a (i ~ k) }) + (linv a (i j)) + + isHAEquiv→Iso : Iso A B + Iso.fun isHAEquiv→Iso = f + Iso.inv isHAEquiv→Iso = g + Iso.rightInv isHAEquiv→Iso = rinv + Iso.leftInv isHAEquiv→Iso = linv + + isHAEquiv→isEquiv : isEquiv f + isHAEquiv→isEquiv .equiv-proof y = (g y , rinv y) , isCenter where + isCenter : xp (g y , rinv y) xp + isCenter (x , p) i = gy≡x i , ry≡p i where + gy≡x : g y x + gy≡x = sym (cong g p) ∙∙ refl ∙∙ linv x + + lem0 : Square (cong f (linv x)) p (cong f (linv x)) p + lem0 i j = invSides-filler p (sym (cong f (linv x))) (~ i) j + + ry≡p : Square (rinv y) p (cong f gy≡x) refl + ry≡p i j = hcomp k λ { (i = i0) cong rinv p k j + ; (i = i1) lem0 k j + ; (j = i0) f (doubleCompPath-filler (sym (cong g p)) refl (linv x) k i) + ; (j = i1) p k }) + (com x (~ i) j) + +open isHAEquiv using (isHAEquiv→Iso; isHAEquiv→isEquiv) public + +HAEquiv : (A : Type ) (B : Type ℓ') Type (ℓ-max ℓ') +HAEquiv A B = Σ[ f (A B) ] isHAEquiv f + +-- vogt's lemma (https://ncatlab.org/nlab/show/homotopy+equivalence#vogts_lemma) +iso→HAEquiv : Iso A B HAEquiv A B +iso→HAEquiv e = f , isHAEquivf + where + f = Iso.fun e + g = Iso.inv e + ε = Iso.rightInv e + η = Iso.leftInv e + + Hfa≡fHa : (f : A A) (H : a f a a) a H (f a) cong f (H a) + Hfa≡fHa f H = J f p a funExt⁻ (sym p) (f a) cong f (funExt⁻ (sym p) a)) + a refl) + (sym (funExt H)) + + isHAEquivf : isHAEquiv f + isHAEquiv.g isHAEquivf = g + isHAEquiv.linv isHAEquivf = η + isHAEquiv.rinv isHAEquivf b i = + hcomp j λ { (i = i0) ε (f (g b)) j + ; (i = i1) ε b j }) + (f (η (g b) i)) + isHAEquiv.com isHAEquivf a i j = + hcomp k λ { (i = i0) ε (f (η a j)) k + ; (j = i0) ε (f (g (f a))) k + ; (j = i1) ε (f a) k}) + (f (Hfa≡fHa x g (f x)) η a (~ i) j)) + +equiv→HAEquiv : A B HAEquiv A B +equiv→HAEquiv e = e .fst , λ where + .isHAEquiv.g invIsEq (snd e) + .isHAEquiv.linv retIsEq (snd e) + .isHAEquiv.rinv secIsEq (snd e) + .isHAEquiv.com a sym (commPathIsEq (snd e) a) + +congIso : {x y : A} (e : Iso A B) Iso (x y) (Iso.fun e x Iso.fun e y) +congIso {x = x} {y} e = goal + where + open isHAEquiv (iso→HAEquiv e .snd) + open Iso + + goal : Iso (x y) (Iso.fun e x Iso.fun e y) + fun goal = cong (iso→HAEquiv e .fst) + inv goal p = sym (linv x) ∙∙ cong g p ∙∙ linv y + rightInv goal p i j = + hcomp k λ { (i = i0) iso→HAEquiv e .fst + (doubleCompPath-filler (sym (linv x)) (cong g p) (linv y) k j) + ; (i = i1) rinv (p j) k + ; (j = i0) com x i k + ; (j = i1) com y i k }) + (iso→HAEquiv e .fst (g (p j))) + leftInv goal p i j = + hcomp k λ { (i = i1) p j + ; (j = i0) Iso.leftInv e x (i k) + ; (j = i1) Iso.leftInv e y (i k) }) + (Iso.leftInv e (p j) i) + +invCongFunct : {x : A} (e : Iso A B) (p : Iso.fun e x Iso.fun e x) (q : Iso.fun e x Iso.fun e x) + Iso.inv (congIso e) (p q) Iso.inv (congIso e) p Iso.inv (congIso e) q +invCongFunct {x = x} e p q = helper (Iso.inv e) _ _ _ + where + helper : {x : A} {y : B} (f : A B) (r : f x y) (p q : x x) + (sym r ∙∙ cong f (p q) ∙∙ r) (sym r ∙∙ cong f p ∙∙ r) (sym r ∙∙ cong f q ∙∙ r) + helper {x = x} f = + J y r (p q : x x) + (sym r ∙∙ cong f (p q) ∙∙ r) (sym r ∙∙ cong f p ∙∙ r) (sym r ∙∙ cong f q ∙∙ r)) + λ p q i rUnit (congFunct f p q i) (~ i)) + λ i rUnit (cong f p) i rUnit (cong f q) i + +invCongRefl : {x : A} (e : Iso A B) Iso.inv (congIso {x = x} {y = x} e) refl refl +invCongRefl {x = x} e = i j Iso.leftInv e x (i ~ j)) ∙∙ refl ∙∙ j Iso.leftInv e x (i j))) sym (rUnit refl) diff --git a/src/docs/Cubical.Foundations.Equiv.Properties.html b/src/docs/Cubical.Foundations.Equiv.Properties.html new file mode 100644 index 0000000..fb5e049 --- /dev/null +++ b/src/docs/Cubical.Foundations.Equiv.Properties.html @@ -0,0 +1,260 @@ +{- + +A couple of general facts about equivalences: + +- if f is an equivalence then (cong f) is an equivalence ([equivCong]) +- if f is an equivalence then pre- and postcomposition with f are equivalences ([preCompEquiv], [postCompEquiv]) +- if f is an equivalence then (Σ[ g ] section f g) and (Σ[ g ] retract f g) are contractible ([isContr-section], [isContr-retract]) + +- isHAEquiv is a proposition [isPropIsHAEquiv] +(these are not in 'Equiv.agda' because they need Univalence.agda (which imports Equiv.agda)) +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.HLevels + +open import Cubical.Functions.FunExtEquiv + +private + variable + ℓ' ℓ'' : Level + A B C : Type + +isEquivInvEquiv : isEquiv (e : A B) invEquiv e) +isEquivInvEquiv = isoToIsEquiv goal where + open Iso + goal : Iso (A B) (B A) + goal .fun = invEquiv + goal .inv = invEquiv + goal .rightInv g = equivEq refl + goal .leftInv f = equivEq refl + +invEquivEquiv : (A B) (B A) +invEquivEquiv = _ , isEquivInvEquiv + +isEquivCong : {x y : A} (e : A B) isEquiv (p : x y) cong (equivFun e) p) +isEquivCong e = isoToIsEquiv (congIso (equivToIso e)) + +congEquiv : {x y : A} (e : A B) (x y) (equivFun e x equivFun e y) +congEquiv e = isoToEquiv (congIso (equivToIso e)) + +equivAdjointEquiv : (e : A B) {a b} (a invEq e b) (equivFun e a b) +equivAdjointEquiv e = compEquiv (congEquiv e) (compPathrEquiv (secEq e _)) + +invEq≡→equivFun≡ : (e : A B) {a b} invEq e b a equivFun e a b +invEq≡→equivFun≡ e = equivFun (equivAdjointEquiv e) sym + +isEquivPreComp : (e : A B) isEquiv (φ : B C) φ equivFun e) +isEquivPreComp e = snd (equiv→ (invEquiv e) (idEquiv _)) + +preCompEquiv : (e : A B) (B C) (A C) +preCompEquiv e = φ φ fst e) , isEquivPreComp e + +isEquivPostComp : (e : A B) isEquiv (φ : C A) e .fst φ) +isEquivPostComp e = snd (equivΠCod _ e)) + +postCompEquiv : (e : A B) (C A) (C B) +postCompEquiv e = _ , isEquivPostComp e + +-- see also: equivΠCod for a dependent version of postCompEquiv + +hasSection : (A B) Type _ +hasSection {A = A} {B = B} f = Σ[ g (B A) ] section f g + +hasRetract : (A B) Type _ +hasRetract {A = A} {B = B} f = Σ[ g (B A) ] retract f g + +isEquiv→isContrHasSection : {f : A B} isEquiv f isContr (hasSection f) +fst (isEquiv→isContrHasSection isEq) = invIsEq isEq , secIsEq isEq +snd (isEquiv→isContrHasSection isEq) (f , ε) i = b fst (p b i)) , b snd (p b i)) + where p : b (invIsEq isEq b , secIsEq isEq b) (f b , ε b) + p b = isEq .equiv-proof b .snd (f b , ε b) + +isEquiv→hasSection : {f : A B} isEquiv f hasSection f +isEquiv→hasSection = fst isEquiv→isContrHasSection + +isContr-hasSection : (e : A B) isContr (hasSection (fst e)) +isContr-hasSection e = isEquiv→isContrHasSection (snd e) + +isEquiv→isContrHasRetract : {f : A B} isEquiv f isContr (hasRetract f) +fst (isEquiv→isContrHasRetract isEq) = invIsEq isEq , retIsEq isEq +snd (isEquiv→isContrHasRetract {f = f} isEq) (g , η) = + λ i b p b i) , a q a i) + where p : b invIsEq isEq b g b + p b = sym (η (invIsEq isEq b)) ∙' cong g (secIsEq isEq b) + -- one square from the definition of invIsEq + ieSq : a Square (cong g (secIsEq isEq (f a))) + refl + (cong (g f) (retIsEq isEq a)) + refl + ieSq a k j = g (commSqIsEq isEq a k j) + -- one square from η + ηSq : a Square (η (invIsEq isEq (f a))) + (η a) + (cong (g f) (retIsEq isEq a)) + (retIsEq isEq a) + ηSq a i j = η (retIsEq isEq a i) j + -- and one last square from the definition of p + pSq : b Square (η (invIsEq isEq b)) + refl + (cong g (secIsEq isEq b)) + (p b) + pSq b i j = compPath'-filler (sym (η (invIsEq isEq b))) (cong g (secIsEq isEq b)) j i + q : a Square (retIsEq isEq a) (η a) (p (f a)) refl + q a i j = hcomp k λ { (i = i0) ηSq a j k + ; (i = i1) η a (j k) + ; (j = i0) pSq (f a) i k + ; (j = i1) η a k + }) + (ieSq a j i) + +isEquiv→hasRetract : {f : A B} isEquiv f hasRetract f +isEquiv→hasRetract = fst isEquiv→isContrHasRetract + +isContr-hasRetract : (e : A B) isContr (hasRetract (fst e)) +isContr-hasRetract e = isEquiv→isContrHasRetract (snd e) + +isEquiv→retractIsEquiv : {f : A B} {g : B A} isEquiv f retract f g isEquiv g +isEquiv→retractIsEquiv {f = f} {g = g} isEquiv-f retract-g = subst isEquiv f⁻¹≡g (snd f⁻¹) + where f⁻¹ = invEquiv (f , isEquiv-f) + + retract-f⁻¹ : retract f (fst f⁻¹) + retract-f⁻¹ = snd (isEquiv→hasRetract isEquiv-f) + + f⁻¹≡g : fst f⁻¹ g + f⁻¹≡g = + cong fst + (isContr→isProp (isEquiv→isContrHasRetract isEquiv-f) + (fst f⁻¹ , retract-f⁻¹) + (g , retract-g)) + + +isEquiv→sectionIsEquiv : {f : A B} {g : B A} isEquiv f section f g isEquiv g +isEquiv→sectionIsEquiv {f = f} {g = g} isEquiv-f section-g = subst isEquiv f⁻¹≡g (snd f⁻¹) + where f⁻¹ = invEquiv (f , isEquiv-f) + + section-f⁻¹ : section f (fst f⁻¹) + section-f⁻¹ = snd (isEquiv→hasSection isEquiv-f) + + f⁻¹≡g : fst f⁻¹ g + f⁻¹≡g = + cong fst + (isContr→isProp (isEquiv→isContrHasSection isEquiv-f) + (fst f⁻¹ , section-f⁻¹) + (g , section-g)) + +cong≃ : (F : Type Type ℓ') (A B) F A F B +cong≃ F e = pathToEquiv (cong F (ua e)) + +cong≃-char : (F : Type Type ℓ') {A B : Type } (e : A B) ua (cong≃ F e) cong F (ua e) +cong≃-char F e = ua-pathToEquiv (cong F (ua e)) + +cong≃-idEquiv : (F : Type Type ℓ') (A : Type ) cong≃ F (idEquiv A) idEquiv (F A) +cong≃-idEquiv F A = cong≃ F (idEquiv A) ≡⟨ cong p pathToEquiv (cong F p)) uaIdEquiv + pathToEquiv refl ≡⟨ pathToEquivRefl + idEquiv (F A) + +isPropIsHAEquiv : {f : A B} isProp (isHAEquiv f) +isPropIsHAEquiv {f = f} ishaef = goal ishaef where + equivF : isEquiv f + equivF = isHAEquiv→isEquiv ishaef + + rCoh1 : (sec : hasSection f) Type _ + rCoh1 (g , ε) = Σ[ η retract f g ] x cong f (η x) ε (f x) + + rCoh2 : (sec : hasSection f) Type _ + rCoh2 (g , ε) = Σ[ η retract f g ] x Square (ε (f x)) refl (cong f (η x)) refl + + rCoh3 : (sec : hasSection f) Type _ + rCoh3 (g , ε) = x Σ[ ηx g (f x) x ] Square (ε (f x)) refl (cong f ηx) refl + + rCoh4 : (sec : hasSection f) Type _ + rCoh4 (g , ε) = x Path (fiber f (f x)) (g (f x) , ε (f x)) (x , refl) + + characterization : isHAEquiv f Σ _ rCoh4 + characterization = + isHAEquiv f + -- first convert between Σ and record + ≃⟨ isoToEquiv (iso e (e .g , e .rinv) , (e .linv , e .com)) + e record { g = e .fst .fst ; rinv = e .fst .snd + ; linv = e .snd .fst ; com = e .snd .snd }) + _ refl) λ _ refl) + Σ _ rCoh1 + -- secondly, convert the path into a dependent path for later convenience + ≃⟨ Σ-cong-equiv-snd s Σ-cong-equiv-snd + λ η equivΠCod + λ x compEquiv (flipSquareEquiv {a₀₀ = f x}) (invEquiv slideSquareEquiv)) + Σ _ rCoh2 + ≃⟨ Σ-cong-equiv-snd s invEquiv Σ-Π-≃) + Σ _ rCoh3 + ≃⟨ Σ-cong-equiv-snd s equivΠCod λ x ΣPath≃PathΣ) + Σ _ rCoh4 + + where open isHAEquiv + + goal : isProp (isHAEquiv f) + goal = subst isProp (sym (ua characterization)) + (isPropΣ (isContr→isProp (isEquiv→isContrHasSection equivF)) + λ s isPropΠ λ x isProp→isSet (isContr→isProp (equivF .equiv-proof (f x))) _ _) + +-- loop spaces connected by a path are equivalent +conjugatePathEquiv : {A : Type } {a b : A} (p : a b) (a a) (b b) +conjugatePathEquiv p = compEquiv (compPathrEquiv p) (compPathlEquiv (sym p)) + +-- composition on the right induces an equivalence of path types +compr≡Equiv : {A : Type } {a b c : A} (p q : a b) (r : b c) (p q) (p r q r) +compr≡Equiv p q r = congEquiv ((λ s s r) , compPathr-isEquiv r) + +-- composition on the left induces an equivalence of path types +compl≡Equiv : {A : Type } {a b c : A} (p : a b) (q r : b c) (q r) (p q p r) +compl≡Equiv p q r = congEquiv ((λ s p s) , (compPathl-isEquiv p)) + +isEquivFromIsContr : {A : Type } {B : Type ℓ'} + (f : A B) isContr A isContr B + isEquiv f +isEquivFromIsContr f isContrA isContrB = + subst isEquiv i x isContr→isProp isContrB (fst B≃A x) (f x) i) (snd B≃A) + where B≃A = isContr→Equiv isContrA isContrB + +isEquiv[f∘equivFunA≃B]→isEquiv[f] : {A : Type } {B : Type ℓ'} {C : Type ℓ''} + (f : B C) (A≃B : A B) + isEquiv (f equivFun A≃B) + isEquiv f +isEquiv[f∘equivFunA≃B]→isEquiv[f] f (g , gIsEquiv) f∘gIsEquiv = + precomposesToId→Equiv f _ w w' + where + w : f g equivFun (invEquiv (_ , f∘gIsEquiv)) idfun _ + w = (cong fst (invEquiv-is-linv (_ , f∘gIsEquiv))) + + w' : isEquiv (g equivFun (invEquiv (_ , f∘gIsEquiv))) + w' = (snd (compEquiv (invEquiv (_ , f∘gIsEquiv) ) (_ , gIsEquiv))) + +isEquiv[equivFunA≃B∘f]→isEquiv[f] : {A : Type } {B : Type ℓ'} {C : Type ℓ''} + (f : C A) (A≃B : A B) + isEquiv (equivFun A≃B f) + isEquiv f +isEquiv[equivFunA≃B∘f]→isEquiv[f] f (g , gIsEquiv) g∘fIsEquiv = + composesToId→Equiv _ f w w' + where + w : equivFun (invEquiv (_ , g∘fIsEquiv)) g f idfun _ + w = (cong fst (invEquiv-is-rinv (_ , g∘fIsEquiv))) + + w' : isEquiv (equivFun (invEquiv (_ , g∘fIsEquiv)) g) + w' = snd (compEquiv (_ , gIsEquiv) (invEquiv (_ , g∘fIsEquiv))) + +isPointedTarget→isEquiv→isEquiv : {A B : Type } (f : A B) + (B isEquiv f) isEquiv f +equiv-proof (isPointedTarget→isEquiv→isEquiv f hf) = + λ y equiv-proof (hf y) y diff --git a/src/docs/Cubical.Foundations.Equiv.html b/src/docs/Cubical.Foundations.Equiv.html new file mode 100644 index 0000000..b26c805 --- /dev/null +++ b/src/docs/Cubical.Foundations.Equiv.html @@ -0,0 +1,326 @@ +{- + +Theory about equivalences + +Definitions are in Core/Glue.agda but re-exported by this module + +- isEquiv is a proposition ([isPropIsEquiv]) +- Any isomorphism is an equivalence ([isoToEquiv]) + +There are more statements about equivalences in Equiv/Properties.agda: + +- if f is an equivalence then (cong f) is an equivalence +- if f is an equivalence then precomposition with f is an equivalence +- if f is an equivalence then postcomposition with f is an equivalence + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Equiv where + +open import Cubical.Foundations.Function +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism + +open import Cubical.Foundations.Equiv.Base public +open import Cubical.Data.Sigma.Base + +private + variable + ℓ' ℓ'' : Level + A B C D : Type + +infixr 30 _∙ₑ_ + +equivIsEquiv : (e : A B) isEquiv (equivFun e) +equivIsEquiv e = snd e + +equivCtr : (e : A B) (y : B) fiber (equivFun e) y +equivCtr e y = e .snd .equiv-proof y .fst + +equivCtrPath : (e : A B) (y : B) + (v : fiber (equivFun e) y) Path _ (equivCtr e y) v +equivCtrPath e y = e .snd .equiv-proof y .snd + + +-- Proof using isPropIsContr. This is slow and the direct proof below is better + +isPropIsEquiv' : (f : A B) isProp (isEquiv f) +equiv-proof (isPropIsEquiv' f u0 u1 i) y = + isPropIsContr (u0 .equiv-proof y) (u1 .equiv-proof y) i + +-- Direct proof that computes quite ok (can be optimized further if +-- necessary, see: +-- https://github.com/mortberg/cubicaltt/blob/pi4s3_dimclosures/examples/brunerie2.ctt#L562 + +isPropIsEquiv : (f : A B) isProp (isEquiv f) +equiv-proof (isPropIsEquiv f p q i) y = + let p2 = p .equiv-proof y .snd + q2 = q .equiv-proof y .snd + in p2 (q .equiv-proof y .fst) i + , λ w j hcomp k λ { (i = i0) p2 w j + ; (i = i1) q2 w (j ~ k) + ; (j = i0) p2 (q2 w (~ k)) i + ; (j = i1) w }) + (p2 w (i j)) + +equivPathP : {A : I Type } {B : I Type ℓ'} {e : A i0 B i0} {f : A i1 B i1} + (h : PathP i A i B i) (e .fst) (f .fst)) PathP i A i B i) e f +equivPathP {e = e} {f = f} h = + λ i (h i) , isProp→PathP i isPropIsEquiv (h i)) (e .snd) (f .snd) i + +equivEq : {e f : A B} (h : e .fst f .fst) e f +equivEq = equivPathP + +module _ {f : A B} (equivF : isEquiv f) where + funIsEq : A B + funIsEq = f + + invIsEq : B A + invIsEq y = equivF .equiv-proof y .fst .fst + + secIsEq : section f invIsEq + secIsEq y = equivF .equiv-proof y .fst .snd + + retIsEq : retract f invIsEq + retIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .fst + + commSqIsEq : a Square (secIsEq (f a)) refl (cong f (retIsEq a)) refl + commSqIsEq a i = equivF .equiv-proof (f a) .snd (a , refl) i .snd + + commPathIsEq : a secIsEq (f a) cong f (retIsEq a) + commPathIsEq a i j = + hcomp + k λ + { (i = i0) secIsEq (f a) j + ; (i = i1) f (retIsEq a (j ~ k)) + ; (j = i0) f (retIsEq a (i ~ k)) + ; (j = i1) f a + }) + (commSqIsEq a i j) + +module _ (w : A B) where + invEq : B A + invEq = invIsEq (snd w) + + retEq : retract (w .fst) invEq + retEq = retIsEq (snd w) + + secEq : section (w .fst) invEq + secEq = secIsEq (snd w) + +open Iso + +equivToIso : { ℓ'} {A : Type } {B : Type ℓ'} A B Iso A B +fun (equivToIso e) = e .fst +inv (equivToIso e) = invEq e +rightInv (equivToIso e) = secEq e +leftInv (equivToIso e) = retEq e + +-- TODO: there should be a direct proof of this that doesn't use equivToIso +invEquiv : A B B A +invEquiv e = isoToEquiv (invIso (equivToIso e)) + +invEquivIdEquiv : (A : Type ) invEquiv (idEquiv A) idEquiv A +invEquivIdEquiv _ = equivEq refl + +compEquiv : A B B C A C +compEquiv f g .fst = g .fst f .fst +compEquiv {A = A} {C = C} f g .snd .equiv-proof c = contr + where + contractG = g .snd .equiv-proof c .snd + + secFiller : (a : A) (p : g .fst (f .fst a) c) _ {- square in A -} + secFiller a p = compPath-filler (cong (invEq f fst) (contractG (_ , p))) (retEq f a) + + contr : isContr (fiber (g .fst f .fst) c) + contr .fst .fst = invEq f (invEq g c) + contr .fst .snd = cong (g .fst) (secEq f (invEq g c)) secEq g c + contr .snd (a , p) i .fst = secFiller a p i1 i + contr .snd (a , p) i .snd j = + hcomp + k λ + { (i = i1) fSquare k + ; (j = i0) g .fst (f .fst (secFiller a p k i)) + ; (j = i1) contractG (_ , p) i .snd k + }) + (g .fst (secEq f (contractG (_ , p) i .fst) j)) + where + fSquare : I C + fSquare k = + hcomp + l λ + { (j = i0) g .fst (f .fst (retEq f a k)) + ; (j = i1) p (k l) + ; (k = i0) g .fst (secEq f (f .fst a) j) + ; (k = i1) p (j l) + }) + (g .fst (f .snd .equiv-proof (f .fst a) .snd (a , refl) k .snd j)) + +_∙ₑ_ = compEquiv + +compEquivIdEquiv : (e : A B) compEquiv (idEquiv A) e e +compEquivIdEquiv e = equivEq refl + +compEquivEquivId : (e : A B) compEquiv e (idEquiv B) e +compEquivEquivId e = equivEq refl + +invEquiv-is-rinv : (e : A B) compEquiv e (invEquiv e) idEquiv A +invEquiv-is-rinv e = equivEq (funExt (retEq e)) + +invEquiv-is-linv : (e : A B) compEquiv (invEquiv e) e idEquiv B +invEquiv-is-linv e = equivEq (funExt (secEq e)) + +compEquiv-assoc : (f : A B) (g : B C) (h : C D) + compEquiv f (compEquiv g h) compEquiv (compEquiv f g) h +compEquiv-assoc f g h = equivEq refl + +LiftEquiv : A Lift {i = } {j = ℓ'} A +LiftEquiv .fst a .lower = a +LiftEquiv .snd .equiv-proof = strictContrFibers lower + +Lift≃Lift : (e : A B) Lift {j = ℓ'} A Lift {j = ℓ''} B +Lift≃Lift e .fst a .lower = e .fst (a .lower) +Lift≃Lift e .snd .equiv-proof b .fst .fst .lower = invEq e (b .lower) +Lift≃Lift e .snd .equiv-proof b .fst .snd i .lower = + e .snd .equiv-proof (b .lower) .fst .snd i +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .fst .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .fst +Lift≃Lift e .snd .equiv-proof b .snd (a , p) i .snd j .lower = + e .snd .equiv-proof (b .lower) .snd (a .lower , cong lower p) i .snd j + +isContr→Equiv : isContr A isContr B A B +isContr→Equiv Actr Bctr = isoToEquiv (isContr→Iso Actr Bctr) + +propBiimpl→Equiv : (Aprop : isProp A) (Bprop : isProp B) (f : A B) (g : B A) A B +propBiimpl→Equiv Aprop Bprop f g = f , hf + where + hf : isEquiv f + hf .equiv-proof y .fst = (g y , Bprop (f (g y)) y) + hf .equiv-proof y .snd h i .fst = Aprop (g y) (h .fst) i + hf .equiv-proof y .snd h i .snd = isProp→isSet' Bprop (Bprop (f (g y)) y) (h .snd) + (cong f (Aprop (g y) (h .fst))) refl i + +isEquivPropBiimpl→Equiv : isProp A isProp B + ((A B) × (B A)) (A B) +isEquivPropBiimpl→Equiv {A = A} {B = B} Aprop Bprop = isoToEquiv isom where + isom : Iso (Σ (A B) _ B A)) (A B) + isom .fun (f , g) = propBiimpl→Equiv Aprop Bprop f g + isom .inv e = equivFun e , invEq e + isom .rightInv e = equivEq refl + isom .leftInv _ = refl + +equivΠCod : {F : A Type } {G : A Type ℓ'} + ((x : A) F x G x) ((x : A) F x) ((x : A) G x) +equivΠCod k .fst f x = k x .fst (f x) +equivΠCod k .snd .equiv-proof f .fst .fst x = equivCtr (k x) (f x) .fst +equivΠCod k .snd .equiv-proof f .fst .snd i x = equivCtr (k x) (f x) .snd i +equivΠCod k .snd .equiv-proof f .snd (g , p) i .fst x = + equivCtrPath (k x) (f x) (g x , λ j p j x) i .fst +equivΠCod k .snd .equiv-proof f .snd (g , p) i .snd j x = + equivCtrPath (k x) (f x) (g x , λ k p k x) i .snd j + +equivImplicitΠCod : {F : A Type } {G : A Type ℓ'} + ({x : A} F x G x) ({x : A} F x) ({x : A} G x) +equivImplicitΠCod k .fst f {x} = k {x} .fst (f {x}) +equivImplicitΠCod k .snd .equiv-proof f .fst .fst {x} = equivCtr (k {x}) (f {x}) .fst +equivImplicitΠCod k .snd .equiv-proof f .fst .snd i {x} = equivCtr (k {x}) (f {x}) .snd i +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .fst {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ j p j {x}) i .fst +equivImplicitΠCod k .snd .equiv-proof f .snd (g , p) i .snd j {x} = + equivCtrPath (k {x}) (f {x}) (g {x} , λ k p k {x}) i .snd j + +equiv→Iso : (A B) (C D) Iso (A C) (B D) +equiv→Iso h k .Iso.fun f b = equivFun k (f (invEq h b)) +equiv→Iso h k .Iso.inv g a = invEq k (g (equivFun h a)) +equiv→Iso h k .Iso.rightInv g = funExt λ b secEq k _ cong g (secEq h b) +equiv→Iso h k .Iso.leftInv f = funExt λ a retEq k _ cong f (retEq h a) + +equiv→ : (A B) (C D) (A C) (B D) +equiv→ h k = isoToEquiv (equiv→Iso h k) + + +equivΠ' : {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A Type ℓB} {B' : A' Type ℓB'} + (eA : A A') + (eB : {a : A} {a' : A'} eA .fst a a' B a B' a') + ((a : A) B a) ((a' : A') B' a') +equivΠ' {B' = B'} eA eB = isoToEquiv isom + where + open Iso + + isom : Iso _ _ + isom .fun f a' = + eB (secEq eA a') .fst (f (invEq eA a')) + isom .inv f' a = + invEq (eB refl) (f' (eA .fst a)) + isom .rightInv f' = + funExt λ a' + J a'' p eB p .fst (invEq (eB refl) (f' (p i0))) f' a'') + (secEq (eB refl) (f' (eA .fst (invEq eA a')))) + (secEq eA a') + isom .leftInv f = + funExt λ a + subst + p invEq (eB refl) (eB p .fst (f (invEq eA (eA .fst a)))) f a) + (sym (commPathIsEq (eA .snd) a)) + (J a'' p invEq (eB refl) (eB (cong (eA .fst) p) .fst (f (invEq eA (eA .fst a)))) f a'') + (retEq (eB refl) (f (invEq eA (eA .fst a)))) + (retEq eA a)) + +equivΠ : {ℓA ℓA' ℓB ℓB'} {A : Type ℓA} {A' : Type ℓA'} + {B : A Type ℓB} {B' : A' Type ℓB'} + (eA : A A') + (eB : (a : A) B a B' (eA .fst a)) + ((a : A) B a) ((a' : A') B' a') +equivΠ {B = B} {B' = B'} eA eB = equivΠ' eA {a = a} p J a' p B a B' a') (eB a) p) + + +equivCompIso : (A B) (C D) Iso (A C) (B D) +equivCompIso h k .Iso.fun f = compEquiv (compEquiv (invEquiv h) f) k +equivCompIso h k .Iso.inv g = compEquiv (compEquiv h g) (invEquiv k) +equivCompIso h k .Iso.rightInv g = equivEq (equiv→Iso h k .Iso.rightInv (equivFun g)) +equivCompIso h k .Iso.leftInv f = equivEq (equiv→Iso h k .Iso.leftInv (equivFun f)) + +equivComp : (A B) (C D) (A C) (B D) +equivComp h k = isoToEquiv (equivCompIso h k) + +-- Some helpful notation: +_≃⟨_⟩_ : (X : Type ) (X B) (B C) (X C) +_ ≃⟨ f g = compEquiv f g + +_■ : (X : Type ) (X X) +_■ = idEquiv + +infixr 0 _≃⟨_⟩_ +infix 1 _■ + +composesToId→Equiv : (f : A B) (g : B A) f g idfun B isEquiv f isEquiv g +composesToId→Equiv f g id iseqf = + isoToIsEquiv + (iso g f + b i equiv-proof iseqf (f b) .snd (g (f b) , cong h h (f b)) id) (~ i) .fst) + ∙∙ cong x equiv-proof iseqf (f b) .fst .fst) id + ∙∙ λ i equiv-proof iseqf (f b) .snd (b , refl) i .fst) + λ a i id i a) + +precomposesToId→Equiv : (f : A B) (g : B A) f g idfun B isEquiv g isEquiv f +precomposesToId→Equiv f g id iseqg = subst isEquiv (sym f-≡-g⁻) (snd (invEquiv (_ , iseqg))) + where + g⁻ = invEq (g , iseqg) + + f-≡-g⁻ : _ + f-≡-g⁻ = cong (f ∘_ ) (cong fst (sym (invEquiv-is-linv (g , iseqg)))) cong (_∘ g⁻) id + +-- equivalence between isEquiv and isEquiv' + +isEquiv-isEquiv'-Iso : (f : A B) Iso (isEquiv f) (isEquiv' f) +isEquiv-isEquiv'-Iso f .fun p = p .equiv-proof +isEquiv-isEquiv'-Iso f .inv q .equiv-proof = q +isEquiv-isEquiv'-Iso f .rightInv q = refl +isEquiv-isEquiv'-Iso f .leftInv p i .equiv-proof = p .equiv-proof + +isEquiv≃isEquiv' : (f : A B) isEquiv f isEquiv' f +isEquiv≃isEquiv' f = isoToEquiv (isEquiv-isEquiv'-Iso f) + +-- The fact that funExt is an equivalence can be found in Cubical.Functions.FunExtEquiv + diff --git a/src/docs/Cubical.Foundations.Function.html b/src/docs/Cubical.Foundations.Function.html new file mode 100644 index 0000000..a18ac74 --- /dev/null +++ b/src/docs/Cubical.Foundations.Function.html @@ -0,0 +1,163 @@ +{- + Definitions for functions +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Function where + +open import Cubical.Foundations.Prelude + +private + variable + ℓ' ℓ'' : Level + A : Type + B : A Type + C : (a : A) B a Type + D : (a : A) (b : B a) C a b Type + E : (x : A) (y : B x) (z : C x y) (w : D x y z) Type + F : (x : A) (y : B x) (z : C x y) (w : D x y z) (u : E x y z w) Type + +-- The identity function +idfun : (A : Type ) A A +idfun _ x = x + +infixr -1 _$_ + +_$_ : { ℓ'} {A : Type } {B : A Type ℓ'} ((a : A) B a) (a : A) B a +f $ a = f a +{-# INLINE _$_ #-} + +infixr 9 _∘_ + +_∘_ : (g : {a : A} (b : B a) C a b) (f : (a : A) B a) (a : A) C a (f a) +g f = λ x g (f x) +{-# INLINE _∘_ #-} + +∘-assoc : (h : {a : A} {b : B a} (c : C a b) D a b c) + (g : {a : A} (b : B a) C a b) + (f : (a : A) B a) + (h g) f h (g f) +∘-assoc h g f i x = h (g (f x)) + +∘-idˡ : (f : (a : A) B a) f idfun A f +∘-idˡ f i x = f x + +∘-idʳ : (f : (a : A) B a) {a} idfun (B a)) f f +∘-idʳ f i x = f x + +flip : {B : Type ℓ'} {C : A B Type ℓ''} + ((x : A) (y : B) C x y) ((y : B) (x : A) C x y) +flip f y x = f x y +{-# INLINE flip #-} + +const : {B : Type } A B A +const x = λ _ x +{-# INLINE const #-} + +case_of_ : { ℓ'} {A : Type } {B : A Type ℓ'} (x : A) (∀ x B x) B x +case x of f = f x +{-# INLINE case_of_ #-} + +case_return_of_ : { ℓ'} {A : Type } (x : A) (B : A Type ℓ') (∀ x B x) B x +case x return P of f = f x +{-# INLINE case_return_of_ #-} + +uncurry : ((x : A) (y : B x) C x y) (p : Σ A B) C (fst p) (snd p) +uncurry f (x , y) = f x y + +uncurry2 : ((x : A) (y : B x) (z : C x y) D x y z) + (p : Σ A x Σ (B x) (C x))) D (p .fst) (p .snd .fst) (p .snd .snd) +uncurry2 f (x , y , z) = f x y z + +uncurry3 : ((x : A) (y : B x) (z : C x y) (w : D x y z) E x y z w) + (p : Σ A x Σ (B x) y Σ (C x y) (D x y)))) + E (p .fst) (p .snd .fst) (p .snd .snd .fst) (p .snd .snd .snd) +uncurry3 f (x , y , z , w) = f x y z w + +uncurry4 : ((x : A) (y : B x) (z : C x y) (w : D x y z) (u : E x y z w) F x y z w u) + (p : Σ A x Σ (B x) y Σ (C x y) z Σ (D x y z) (E x y z))))) + F (p .fst) (p .snd .fst) (p .snd .snd .fst) (p .snd .snd .snd .fst) (p .snd .snd .snd .snd) +uncurry4 f (x , y , z , w , u) = f x y z w u + + +curry : ((p : Σ A B) C (fst p) (snd p)) (x : A) (y : B x) C x y +curry f x y = f (x , y) + +∘diag : {B : (x y : A) Type } (∀ x y B x y) x B x x +∘diag f x = f x x + +module _ { ℓ'} {A : Type } {B : Type ℓ'} where + -- Notions of 'coherently constant' functions for low dimensions. + -- These are the properties of functions necessary to e.g. eliminate + -- from the propositional truncation. + + -- 2-Constant functions are coherently constant if B is a set. + 2-Constant : (A B) Type _ + 2-Constant f = x y f x f y + + 2-Constant-isProp : isSet B (f : A B) isProp (2-Constant f) + 2-Constant-isProp Bset f link1 link2 i x y j + = Bset (f x) (f y) (link1 x y) (link2 x y) i j + + -- 3-Constant functions are coherently constant if B is a groupoid. + record 3-Constant (f : A B) : Type (ℓ-max ℓ') where + field + link : 2-Constant f + coh₁ : x y z Square (link x y) (link x z) refl (link y z) + + coh₂ : x y z Square (link x z) (link y z) (link x y) refl + coh₂ x y z i j + = hcomp k λ + { (j = i0) link x y i + ; (i = i0) link x z (j k) + ; (j = i1) link x z (i k) + ; (i = i1) link y z j + }) + (coh₁ x y z j i) + + link≡refl : x refl link x x + link≡refl x i j + = hcomp k λ + { (i = i0) link x x (j ~ k) + ; (i = i1) link x x j + ; (j = i0) f x + ; (j = i1) link x x (~ i ~ k) + }) + (coh₁ x x x (~ i) j) + + downleft : x y Square (link x y) refl refl (link y x) + downleft x y i j + = hcomp k λ + { (i = i0) link x y j + ; (i = i1) link≡refl x (~ k) j + ; (j = i0) f x + ; (j = i1) link y x i + }) + (coh₁ x y x i j) + + link≡symlink : x y link x y sym (link y x) + link≡symlink x y i j + = hcomp k λ + { (i = i0) link x y j + ; (i = i1) link y x (~ j ~ k) + ; (j = i0) f x + ; (j = i1) link y x (i ~ k) + }) + (downleft x y i j) + +homotopyNatural : {B : Type ℓ'} {f g : A B} (H : a f a g a) {x y : A} (p : x y) + H x cong g p cong f p H y +homotopyNatural {f = f} {g = g} H {x} {y} p i j = + hcomp k λ { (i = i0) compPath-filler (H x) (cong g p) k j + ; (i = i1) compPath-filler' (cong f p) (H y) k j + ; (j = i0) cong f p (i ~ k) + ; (j = i1) cong g p (i k) }) + (H (p i) j) + +homotopySymInv : {f : A A} (H : a f a a) (a : A) + Path (f a f a) i H (H a (~ i)) i) refl +homotopySymInv {f = f} H a j i = + hcomp k λ { (i = i0) f a + ; (i = i1) H a (j ~ k) + ; (j = i0) H (H a (~ i)) i + ; (j = i1) H a (i ~ k)}) + (H (H a (~ i j)) i) diff --git a/src/docs/Cubical.Foundations.GroupoidLaws.html b/src/docs/Cubical.Foundations.GroupoidLaws.html new file mode 100644 index 0000000..5cd9b46 --- /dev/null +++ b/src/docs/Cubical.Foundations.GroupoidLaws.html @@ -0,0 +1,532 @@ +{- + +This file proves the higher groupoid structure of types +for homogeneous and heterogeneous paths + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.GroupoidLaws where + +open import Cubical.Foundations.Prelude + +private + variable + : Level + A : Type + x y z w v : A + +_⁻¹ : (x y) (y x) +x≡y ⁻¹ = sym x≡y + +infix 40 _⁻¹ + +-- homogeneous groupoid laws + +symInvo : (p : x y) p p ⁻¹ ⁻¹ +symInvo p = refl + +rUnit : (p : x y) p p refl +rUnit p j i = compPath-filler p refl j i + +-- The filler of left unit: lUnit-filler p = +-- PathP (λ i → PathP (λ j → PathP (λ k → A) x (p (~ j ∨ i))) +-- (refl i) (λ j → compPath-filler refl p i j)) (λ k i → (p (~ k ∧ i ))) (lUnit p) + +lUnit-filler : {x y : A} (p : x y) I I I A +lUnit-filler {x = x} p j k i = + hfill j λ { (i = i0) x + ; (i = i1) p (~ k j ) + ; (k = i0) p i + -- ; (k = i1) → compPath-filler refl p j i + }) (inS (p (~ k i ))) j + +lUnit : (p : x y) p refl p +lUnit p j i = lUnit-filler p i1 j i + +symRefl : refl {x = x} refl ⁻¹ +symRefl i = refl + +compPathRefl : refl {x = x} refl refl +compPathRefl = rUnit refl + +-- The filler of right cancellation: rCancel-filler p = +-- PathP (λ i → PathP (λ j → PathP (λ k → A) x (p (~ j ∧ ~ i))) +-- (λ j → compPath-filler p (p ⁻¹) i j) (refl i)) (λ j i → (p (i ∧ ~ j))) (rCancel p) + +rCancel-filler : {x y : A} (p : x y) (k j i : I) A +rCancel-filler {x = x} p k j i = + hfill k λ { (i = i0) x + ; (i = i1) p (~ k ~ j) + -- ; (j = i0) → compPath-filler p (p ⁻¹) k i + ; (j = i1) x + }) (inS (p (i ~ j))) k + +rCancel : (p : x y) p p ⁻¹ refl +rCancel {x = x} p j i = rCancel-filler p i1 j i + +rCancel-filler' : {} {A : Type } {x y : A} (p : x y) (i j k : I) A +rCancel-filler' {x = x} {y} p i j k = + hfill + i λ + { (j = i1) p (~ i k) + ; (k = i0) x + ; (k = i1) p (~ i) + }) + (inS (p k)) + (~ i) + +rCancel' : {} {A : Type } {x y : A} (p : x y) p p ⁻¹ refl +rCancel' p j k = rCancel-filler' p i0 j k + +lCancel : (p : x y) p ⁻¹ p refl +lCancel p = rCancel (p ⁻¹) + +assoc : (p : x y) (q : y z) (r : z w) + p q r (p q) r +assoc p q r k = (compPath-filler p q k) compPath-filler' q r (~ k) + + +-- heterogeneous groupoid laws + +symInvoP : {A : I Type } {x : A i0} {y : A i1} (p : PathP A x y) + PathP j PathP i symInvo i A i) j i) x y) p (symP (symP p)) +symInvoP p = refl + +rUnitP : {A : I Type } {x : A i0} {y : A i1} (p : PathP A x y) + PathP j PathP i rUnit i A i) j i) x y) p (compPathP p refl) +rUnitP p j i = compPathP-filler p refl j i + +rUnitP' : {ℓ'} {A : Type } (B : A Type ℓ') + {x y : A} {p : x y} {z : B x} {w : B y} + (q : PathP i B (p i)) z w) + PathP j PathP i B (rUnit p j i)) z w) q (compPathP' {B = B} q refl) +rUnitP' B {w = w} q j i = compPathP'-filler {B = B} q (refl {x = w}) j i + +lUnitP : {A : I Type } {x : A i0} {y : A i1} (p : PathP A x y) + PathP j PathP i lUnit i A i) j i) x y) p (compPathP refl p) +lUnitP {A = A} {x = x} p k i = + comp j lUnit-filler i A i) j k i) + j λ { (i = i0) x + ; (i = i1) p (~ k j ) + ; (k = i0) p i + }) (p (~ k i )) + +lUnitP' : {ℓ'} {A : Type } (B : A Type ℓ') + {x y : A} {p : x y} {z : B x} {w : B y} + (q : PathP i B (p i)) z w) + PathP j PathP i B (lUnit p j i)) z w) q (compPathP' {B = B} refl q) +lUnitP' B {p = p} {z = z} q k i = + comp j B (lUnit-filler p j k i)) + j λ { (i = i0) z + ; (i = i1) q (~ k j ) + ; (k = i0) q i + }) (q (~ k i )) + +rCancelP : {A : I Type } {x : A i0} {y : A i1} (p : PathP A x y) + PathP j PathP i rCancel i A i) j i) x x) (compPathP p (symP p)) refl +rCancelP {A = A} {x = x} p j i = + comp k rCancel-filler i A i) k j i) + k λ { (i = i0) x + ; (i = i1) p (~ k ~ j) + ; (j = i1) x + }) (p (i ~ j)) + +lCancelP : {A : I Type } {x : A i0} {y : A i1} (p : PathP A x y) + PathP j PathP i lCancel i A i) j i) y y) (compPathP (symP p) p) refl +lCancelP p = rCancelP (symP p) + + + +assocP : {A : I Type } {x : A i0} {y : A i1} {B_i1 : Type } {B : (A i1) B_i1} {z : B i1} + {C_i1 : Type } {C : (B i1) C_i1} {w : C i1} (p : PathP A x y) (q : PathP i B i) y z) (r : PathP i C i) z w) + PathP j PathP i assoc i A i) B C j i) x w) (compPathP p (compPathP q r)) (compPathP (compPathP p q) r) +assocP {A = A} {B = B} {C = C} p q r k i = + comp (\ j' hfill j λ { + (i = i0) A i0 + ; (i = i1) compPath-filler' i₁ B i₁) i₁ C i₁) (~ k) j }) + (inS (compPath-filler i₁ A i₁) i₁ B i₁) k i)) j') + j λ + { (i = i0) p i0 + ; (i = i1) + comp (\ j' hfill ((λ l λ + { (j = i0) B k + ; (j = i1) C l + ; (k = i1) C (j l) + })) (inS (B ( j k)) ) j') + l λ + { (j = i0) q k + ; (j = i1) r l + ; (k = i1) r (j l) + }) + (q (j k)) + }) + (compPathP-filler p q k i) + + + +-- Loic's code below + +-- some exchange law for doubleCompPath and refl + +invSides-filler : {x y z : A} (p : x y) (q : x z) Square p (sym q) q (sym p) +invSides-filler {x = x} p q i j = + hcomp k λ { (i = i0) p (k j) + ; (i = i1) q (~ j k) + ; (j = i0) q (i k) + ; (j = i1) p (~ i k)}) + x + +leftright : { : Level} {A : Type } {x y z : A} (p : x y) (q : y z) + (refl ∙∙ p ∙∙ q) (p ∙∙ q ∙∙ refl) +leftright p q i j = + hcomp t λ { (j = i0) p (i (~ t)) + ; (j = i1) q (t i) }) + (invSides-filler q (sym p) (~ i) j) + +-- equating doubleCompPath and a succession of two compPath + +split-leftright : { : Level} {A : Type } {w x y z : A} (p : w x) (q : x y) (r : y z) + (p ∙∙ q ∙∙ r) (refl ∙∙ (p ∙∙ q ∙∙ refl) ∙∙ r) +split-leftright p q r j i = + hcomp t λ { (i = i0) p (~ j ~ t) + ; (i = i1) r t }) + (doubleCompPath-filler p q refl j i) + +split-leftright' : { : Level} {A : Type } {w x y z : A} (p : w x) (q : x y) (r : y z) + (p ∙∙ q ∙∙ r) (p ∙∙ (refl ∙∙ q ∙∙ r) ∙∙ refl) +split-leftright' p q r j i = + hcomp t λ { (i = i0) p (~ t) + ; (i = i1) r (j t) }) + (doubleCompPath-filler refl q r j i) + +doubleCompPath-elim : { : Level} {A : Type } {w x y z : A} (p : w x) (q : x y) + (r : y z) (p ∙∙ q ∙∙ r) (p q) r +doubleCompPath-elim p q r = (split-leftright p q r) i (leftright p q (~ i)) r) + +doubleCompPath-elim' : { : Level} {A : Type } {w x y z : A} (p : w x) (q : x y) + (r : y z) (p ∙∙ q ∙∙ r) p (q r) +doubleCompPath-elim' p q r = (split-leftright' p q r) (sym (leftright p (q r))) + +cong-∙∙-filler : { ℓ'} {A : Type } {B : Type ℓ'} {x y z w : A} + (f : A B) (p : w x) (q : x y) (r : y z) + I I I B +cong-∙∙-filler {A = A} f p q r k j i = + hfill ((λ k λ { (j = i1) doubleCompPath-filler (cong f p) (cong f q) (cong f r) k i + ; (j = i0) f (doubleCompPath-filler p q r k i) + ; (i = i0) f (p (~ k)) + ; (i = i1) f (r k) })) + (inS (f (q i))) + k + +cong-∙∙ : {B : Type } (f : A B) (p : w x) (q : x y) (r : y z) + cong f (p ∙∙ q ∙∙ r) (cong f p) ∙∙ (cong f q) ∙∙ (cong f r) +cong-∙∙ f p q r j i = cong-∙∙-filler f p q r i1 j i + +cong-∙ : {B : Type } (f : A B) (p : x y) (q : y z) + cong f (p q) (cong f p) (cong f q) +cong-∙ f p q = cong-∙∙ f refl p q + +hcomp-unique : {} {A : Type } {φ} + (u : I Partial φ A) (u0 : A [ φ u i0 ]) + (h2 : i A [ (φ ~ i) (\ { (φ = i1) u i 1=1; (i = i0) outS u0}) ]) + (hcomp u (outS u0) outS (h2 i1)) [ φ (\ { (φ = i1) (\ i u i1 1=1)}) ] +hcomp-unique {φ = φ} u u0 h2 = inS (\ i hcomp (\ k \ { (φ = i1) u k 1=1 + ; (i = i1) outS (h2 k) }) + (outS u0)) + + +hlid-unique : {} {A : Type } {φ} + (u : I Partial φ A) (u0 : A [ φ u i0 ]) + (h1 h2 : i A [ (φ ~ i) (\ { (φ = i1) u i 1=1; (i = i0) outS u0}) ]) + (outS (h1 i1) outS (h2 i1)) [ φ (\ { (φ = i1) (\ i u i1 1=1)}) ] +hlid-unique {φ = φ} u u0 h1 h2 = inS (\ i hcomp (\ k \ { (φ = i1) u k 1=1 + ; (i = i0) outS (h1 k) + ; (i = i1) outS (h2 k) }) + (outS u0)) + +comp-unique : {} {A : I Type } {φ} + (u : (i : I) Partial φ (A i)) (u0 : A i0 [ φ u i0 ]) + (h2 : i A i [ (φ ~ i) (\ { (φ = i1) u i 1=1; (i = i0) outS u0}) ]) + (comp A u (outS u0) outS (h2 i1)) [ φ (\ { (φ = i1) (\ i u i1 1=1)}) ] +comp-unique {A = A} {φ = φ} u u0 h2 = inS (\ i comp A (\ k \ { (φ = i1) u k 1=1 + ; (i = i1) outS (h2 k) }) + (outS u0)) + +lid-unique : {} {A : I Type } {φ} + (u : (i : I) Partial φ (A i)) (u0 : A i0 [ φ u i0 ]) + (h1 h2 : i A i [ (φ ~ i) (\ { (φ = i1) u i 1=1; (i = i0) outS u0}) ]) + (outS (h1 i1) outS (h2 i1)) [ φ (\ { (φ = i1) (\ i u i1 1=1)}) ] +lid-unique {A = A} {φ = φ} u u0 h1 h2 = inS (\ i comp A (\ k \ { (φ = i1) u k 1=1 + ; (i = i0) outS (h1 k) + ; (i = i1) outS (h2 k) }) + (outS u0)) + +transp-hcomp : {} (φ : I) {A' : Type } + (A : (i : I) Type [ φ _ A') ]) (let B = \ (i : I) outS (A i)) + {ψ} (u : I Partial ψ (B i0)) (u0 : B i0 [ ψ u i0 ]) + (transp (\ i B i) φ (hcomp u (outS u0)) hcomp (\ i o transp (\ i B i) φ (u i o)) (transp (\ i B i) φ (outS u0))) + [ ψ (\ { (ψ = i1) (\ i transp (\ i B i) φ (u i1 1=1))}) ] +transp-hcomp φ A u u0 = inS (sym (outS (hcomp-unique + ((\ i o transp (\ i B i) φ (u i o))) (inS (transp (\ i B i) φ (outS u0))) + \ i inS (transp (\ i B i) φ (hfill u u0 i))))) + where + B = \ (i : I) outS (A i) + + +hcomp-cong : {} {A : Type } {φ} (u : I Partial φ A) (u0 : A [ φ u i0 ]) + (u' : I Partial φ A) (u0' : A [ φ u' i0 ]) + (ueq : i PartialP φ (\ o u i o u' i o)) (outS u0 outS u0') [ φ (\ { (φ = i1) ueq i0 1=1}) ] + (hcomp u (outS u0) hcomp u' (outS u0')) [ φ (\ { (φ = i1) ueq i1 1=1 }) ] +hcomp-cong u u0 u' u0' ueq 0eq = inS (\ j hcomp (\ i o ueq i o j) (outS 0eq j)) + + +congFunct-filler : { ℓ'} {A : Type } {B : Type ℓ'} {x y z : A} (f : A B) (p : x y) (q : y z) + I I I B +congFunct-filler {x = x} f p q i j z = + hfill k λ { (i = i0) f x + ; (i = i1) f (q k) + ; (j = i0) f (compPath-filler p q k i)}) + (inS (f (p i))) + z + +congFunct : {} {B : Type } (f : A B) (p : x y) (q : y z) cong f (p q) cong f p cong f q +congFunct f p q j i = congFunct-filler f p q i j i1 + + +-- congFunct for dependent types +congFunct-dep : { ℓ'} {A : Type } {B : A Type ℓ'} {x y z : A} (f : (a : A) B a) (p : x y) (q : y z) + PathP i PathP j B (compPath-filler p q i j)) (f x) (f (q i))) (cong f p) (cong f (p q)) +congFunct-dep {B = B} {x = x} f p q i j = f (compPath-filler p q i j) + +cong₂Funct : { ℓ'} {A : Type } {x y : A} {B : Type ℓ'} (f : A A B) + (p : x y) + {u v : A} (q : u v) + cong₂ f p q cong x f x u) p cong (f y) q +cong₂Funct {x = x} {y = y} f p {u = u} {v = v} q j i = + hcomp k λ { (i = i0) f x u + ; (i = i1) f y (q k) + ; (j = i0) f (p i) (q (i k))}) + (f (p i) u) + +symDistr-filler : {} {A : Type } {x y z : A} (p : x y) (q : y z) I I I A +symDistr-filler {A = A} {z = z} p q i j k = + hfill k λ { (i = i0) q (k j) + ; (i = i1) p (~ k j) }) + (inS (invSides-filler q (sym p) i j)) + k + +symDistr : {} {A : Type } {x y z : A} (p : x y) (q : y z) sym (p q) sym q sym p +symDistr p q i j = symDistr-filler p q j i i1 + +-- we can not write hcomp-isEquiv : {ϕ : I} → (p : I → Partial ϕ A) → isEquiv (λ (a : A [ ϕ ↦ p i0 ]) → hcomp p a) +-- due to size issues. But what we can write (compare to hfill) is: +hcomp-equivFillerSub : {ϕ : I} (p : I Partial ϕ A) (a : A [ ϕ p i0 ]) + (i : I) + A [ ϕ i ~ i { (i = i0) outS a + ; (i = i1) hcomp i p (~ i)) (hcomp p (outS a)) + ; (ϕ = i1) p i0 1=1 }) ] +hcomp-equivFillerSub {ϕ = ϕ} p a i = + inS (hcomp k λ { (i = i1) hfill j p (~ j)) (inS (hcomp p (outS a))) k + ; (i = i0) outS a + ; (ϕ = i1) p (~ k i) 1=1 }) + (hfill p a i)) + +hcomp-equivFiller : {ϕ : I} (p : I Partial ϕ A) (a : A [ ϕ p i0 ]) + (i : I) A +hcomp-equivFiller p a i = outS (hcomp-equivFillerSub p a i) + + +pentagonIdentity : (p : x y) (q : y z) (r : z w) (s : w v) + + (assoc p q (r s) assoc (p q) r s) + + cong (p ∙_) (assoc q r s) ∙∙ assoc p (q r) s ∙∙ cong (_∙ s) (assoc p q r) + +pentagonIdentity {x = x} {y} p q r s = + i + j cong (p ∙_) (assoc q r s) (i j)) + ∙∙ j lemma₀₀ i j lemma₀₁ i j) + ∙∙ j lemma₁₀ i j lemma₁₁ i j) + ) + where + + + lemma₀₀ : ( i j : I) _ _ + lemma₀₀ i j i₁ = + hcomp + k λ { (j = i0) p i₁ + ; (i₁ = i0) x + ; (i₁ = i1) hcomp + k₁ λ { (i = i0) (q (j k)) + ; (k = i0) y + ; (j = i0) y + ; (j = i1)(k = i1) r (k₁ i)}) + (q (j k)) + }) (p i₁) + + lemma₀₁ : ( i j : I) hcomp + k λ {(i = i0) q j + ; (j = i0) y + ; (j = i1) r (k i) + }) + (q j) _ + lemma₀₁ i j i₁ = (hcomp + k λ { (j = i1) hcomp + k₁ λ { (i₁ = i0) r i + ; (k = i0) r i + ; (i = i1) s (k₁ k i₁) + ; (i₁ = i1)(k = i1) s k₁ }) + (r ((i₁ k) i)) + ; (i₁ = i0) compPath-filler q r i j + ; (i₁ = i1) hcomp + k₁ λ { (k = i0) r i + ; (k = i1) s k₁ + ; (i = i1) s (k k₁)}) + (r (i k))}) + (hfill + k λ { (j = i1) r k + ; (i₁ = i1) r k + ; (i₁ = i0)(j = i0) y }) + (inS (q (i₁ j))) i)) + + lemma₁₁ : ( i j : I) (r (i j)) _ + lemma₁₁ i j i₁ = + hcomp + k λ { (i = i1) s (i₁ k) + ; (j = i1) s (i₁ k) + ; (i₁ = i0) r (i j) + ; (i₁ = i1) s k + }) (r (i j i₁)) + + + lemma₁₀-back : I I I _ + lemma₁₀-back i j i₁ = + hcomp + k λ { + (i₁ = i0) x + ; (i₁ = i1) hcomp + k₁ λ { (k = i0) q (j ~ i) + ; (k = i1) r (k₁ j) + ; (j = i0) q (k ~ i) + ; (j = i1) r (k₁ k) + ; (i = i0) r (k j k₁) + }) + (q (k j ~ i)) + ; (i = i0)(j = i0) (p q) i₁ + }) + (hcomp + k λ { (i₁ = i0) x + ; (i₁ = i1) q ((j ~ i ) k) + ; (j = i0)(i = i1) p i₁ + }) + (p i₁)) + + + lemma₁₀-front : I I I _ + lemma₁₀-front i j i₁ = + (((λ _ x) ∙∙ compPath-filler p q j ∙∙ + i₁ + hcomp + k λ { (i₁ = i0) q j + ; (i₁ = i1) r (k (j i)) + ; (j = i0)(i = i0) q i₁ + ; (j = i1) r (i₁ k) + }) + (q (j i₁)) + )) i₁) + + compPath-filler-in-filler : + (p : _ y) (q : _ _ ) + _≡_ {A = Square (p q) (p q) _ x) _ z)} + i j hcomp + i₂ + λ { (j = i0) x + ; (j = i1) q (i₂ ~ i) + ; (i = i0) (p q) j + }) + (compPath-filler p q (~ i) j)) + _ p q) + compPath-filler-in-filler p q z i j = + hcomp + k λ { + (j = i0) p i0 + ; (j = i1) q (k ~ i ~ z) + ; (i = i0) hcomp + i₂ λ { + (j = i0) p i0 + ;(j = i1) q ((k ~ z) i₂) + ;(z = i1) (k = i0) p j + }) + (p j) + ; (i = i1) compPath-filler p i₁ q (k i₁)) k j + ; (z = i0) hfill + ((λ i₂ λ { (j = i0) p i0 + ; (j = i1) q (i₂ ~ i) + ; (i = i0) (p q) j + })) + (inS ((compPath-filler p q (~ i) j))) k + ; (z = i1) compPath-filler p q k j + }) + (compPath-filler p q (~ i ~ z) j) + + + cube-comp₋₀₋ : + (c : I I I A) + {a' : Square _ _ _ _} + i i₁ c i i0 i₁) a' + (I I I A) + cube-comp₋₀₋ c p i j k = + hcomp + l λ { + (i = i0) c i0 j k + ;(i = i1) c i1 j k + ;(j = i0) p l i k + ;(j = i1) c i i1 k + ;(k = i0) c i j i0 + ;(k = i1) c i j i1 + }) + (c i j k) + + cube-comp₀₋₋ : + (c : I I I A) + {a' : Square _ _ _ _} + i i₁ c i0 i i₁) a' + (I I I A) + cube-comp₀₋₋ c p i j k = + hcomp + l λ { + (i = i0) p l j k + ;(i = i1) c i1 j k + ;(j = i0) c i i0 k + ;(j = i1) c i i1 k + ;(k = i0) c i j i0 + ;(k = i1) c i j i1 + }) + (c i j k) + + + + lemma₁₀-back' : _ + lemma₁₀-back' k j i₁ = + (cube-comp₋₀₋ (lemma₁₀-back) + (compPath-filler-in-filler p q)) k j i₁ + + + lemma₁₀ : ( i j : I) _ _ + lemma₁₀ i j i₁ = + (cube-comp₀₋₋ lemma₁₀-front (sym lemma₁₀-back')) i j i₁ + +-- misc. +∙∙lCancel-fill : {} {A : Type } {x y : A} + (p : x y) + I I I A +∙∙lCancel-fill p i j k = + hfill k λ { (i = i1) p k + ; (j = i0) p k + ; (j = i1) p k}) + (inS (p i0)) k + +∙∙lCancel : {} {A : Type } {x y : A} + (p : x y) + sym p ∙∙ refl ∙∙ p refl +∙∙lCancel p i j = ∙∙lCancel-fill p i j i1 diff --git a/src/docs/Cubical.Foundations.HLevels.html b/src/docs/Cubical.Foundations.HLevels.html new file mode 100644 index 0000000..5ac83da --- /dev/null +++ b/src/docs/Cubical.Foundations.HLevels.html @@ -0,0 +1,841 @@ +{- + +Basic theory about h-levels/n-types: + +- Basic properties of isContr, isProp and isSet (definitions are in Prelude) + +- Hedberg's theorem can be found in Cubical/Relation/Nullary/Properties + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.HLevels where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Structure +open import Cubical.Functions.FunExtEquiv +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence using (ua ; univalenceIso) + +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using (; zero; suc; _+_; +-zero; +-comm) + +HLevel : Type₀ +HLevel = + +private + variable + ℓ' ℓ'' ℓ''' ℓ'''' ℓ''''' : Level + A A' : Type + B : A Type + C : (x : A) B x Type + D : (x : A) (y : B x) C x y Type + E : (x : A) (y : B x) (z : C x y) D x y z Type + F : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) Type + G : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) (u : F x y z w v) Type + w x y z : A + n : HLevel + +isOfHLevel : HLevel Type Type +isOfHLevel 0 A = isContr A +isOfHLevel 1 A = isProp A +isOfHLevel (suc (suc n)) A = (x y : A) isOfHLevel (suc n) (x y) + +isOfHLevelFun : (n : HLevel) {A : Type } {B : Type ℓ'} (f : A B) Type (ℓ-max ℓ') +isOfHLevelFun n f = b isOfHLevel n (fiber f b) + +isOfHLevelΩ→isOfHLevel : + {} {A : Type } (n : ) + ((x : A) isOfHLevel (suc n) (x x)) isOfHLevel (2 + n) A +isOfHLevelΩ→isOfHLevel zero x y = + J y p (q : x y) p q) ( x refl) +isOfHLevelΩ→isOfHLevel (suc n) x y = + J y p (q : x y) isOfHLevel (suc n) (p q)) ( x refl) + +TypeOfHLevel : HLevel Type (ℓ-suc ) +TypeOfHLevel n = TypeWithStr (isOfHLevel n) + +hProp hSet hGroupoid h2Groupoid : Type (ℓ-suc ) +hProp = TypeOfHLevel 1 +hSet = TypeOfHLevel 2 +hGroupoid = TypeOfHLevel 3 +h2Groupoid = TypeOfHLevel 4 + +-- lower h-levels imply higher h-levels + +isOfHLevelSuc : (n : HLevel) isOfHLevel n A isOfHLevel (suc n) A +isOfHLevelSuc 0 = isContr→isProp +isOfHLevelSuc 1 = isProp→isSet +isOfHLevelSuc (suc (suc n)) h a b = isOfHLevelSuc (suc n) (h a b) + +isSet→isGroupoid : isSet A isGroupoid A +isSet→isGroupoid = isOfHLevelSuc 2 + +isGroupoid→is2Groupoid : isGroupoid A is2Groupoid A +isGroupoid→is2Groupoid = isOfHLevelSuc 3 + +isOfHLevelPlus : (m : HLevel) isOfHLevel n A isOfHLevel (m + n) A +isOfHLevelPlus zero hA = hA +isOfHLevelPlus (suc m) hA = isOfHLevelSuc _ (isOfHLevelPlus m hA) + +isContr→isOfHLevel : (n : HLevel) isContr A isOfHLevel n A +isContr→isOfHLevel zero cA = cA +isContr→isOfHLevel (suc n) cA = isOfHLevelSuc _ (isContr→isOfHLevel n cA) + +isProp→isOfHLevelSuc : (n : HLevel) isProp A isOfHLevel (suc n) A +isProp→isOfHLevelSuc zero pA = pA +isProp→isOfHLevelSuc (suc n) pA = isOfHLevelSuc _ (isProp→isOfHLevelSuc n pA) + +isOfHLevelPlus' : (m : HLevel) isOfHLevel m A isOfHLevel (m + n) A +isOfHLevelPlus' {n = n} 0 = isContr→isOfHLevel n +isOfHLevelPlus' {n = n} 1 = isProp→isOfHLevelSuc n +isOfHLevelPlus' {n = n} (suc (suc m)) hA a₀ a₁ = isOfHLevelPlus' (suc m) (hA a₀ a₁) + +-- hlevel of path types + +isProp→isContrPath : isProp A (x y : A) isContr (x y) +isProp→isContrPath h x y = h x y , isProp→isSet h x y _ + +isContr→isContrPath : isContr A (x y : A) isContr (x y) +isContr→isContrPath cA = isProp→isContrPath (isContr→isProp cA) + +isOfHLevelPath' : (n : HLevel) isOfHLevel (suc n) A (x y : A) isOfHLevel n (x y) +isOfHLevelPath' 0 = isProp→isContrPath +isOfHLevelPath' (suc n) h x y = h x y + +isOfHLevelPath'⁻ : (n : HLevel) ((x y : A) isOfHLevel n (x y)) isOfHLevel (suc n) A +isOfHLevelPath'⁻ zero h x y = h x y .fst +isOfHLevelPath'⁻ (suc n) h = h + +isOfHLevelPath : (n : HLevel) isOfHLevel n A (x y : A) isOfHLevel n (x y) +isOfHLevelPath 0 h x y = isContr→isContrPath h x y +isOfHLevelPath (suc n) h x y = isOfHLevelSuc n (isOfHLevelPath' n h x y) + +-- h-level of isOfHLevel + +isPropIsOfHLevel : (n : HLevel) isProp (isOfHLevel n A) +isPropIsOfHLevel 0 = isPropIsContr +isPropIsOfHLevel 1 = isPropIsProp +isPropIsOfHLevel (suc (suc n)) f g i a b = + isPropIsOfHLevel (suc n) (f a b) (g a b) i + +isPropIsSet : isProp (isSet A) +isPropIsSet = isPropIsOfHLevel 2 + +isPropIsGroupoid : isProp (isGroupoid A) +isPropIsGroupoid = isPropIsOfHLevel 3 + +isPropIs2Groupoid : isProp (is2Groupoid A) +isPropIs2Groupoid = isPropIsOfHLevel 4 + +TypeOfHLevel≡ : (n : HLevel) {X Y : TypeOfHLevel n} X Y X Y +TypeOfHLevel≡ n = Σ≡Prop _ isPropIsOfHLevel n) + +-- hlevels are preserved by retracts (and consequently equivalences) + +isContrRetract + : {B : Type } + (f : A B) (g : B A) + (h : retract f g) + (v : isContr B) isContr A +fst (isContrRetract f g h (b , p)) = g b +snd (isContrRetract f g h (b , p)) x = (cong g (p (f x))) (h x) + +isPropRetract + : {B : Type } + (f : A B) (g : B A) + (h : (x : A) g (f x) x) + isProp B isProp A +isPropRetract f g h p x y i = + hcomp + j λ + { (i = i0) h x j + ; (i = i1) h y j}) + (g (p (f x) (f y) i)) + +isSetRetract + : {B : Type } + (f : A B) (g : B A) + (h : (x : A) g (f x) x) + isSet B isSet A +isSetRetract f g h set x y p q i j = + hcomp k λ { (i = i0) h (p j) k + ; (i = i1) h (q j) k + ; (j = i0) h x k + ; (j = i1) h y k}) + (g (set (f x) (f y) + (cong f p) (cong f q) i j)) + +isGroupoidRetract + : {B : Type } + (f : A B) (g : B A) + (h : (x : A) g (f x) x) + isGroupoid B isGroupoid A +isGroupoidRetract f g h grp x y p q P Q i j k = + hcomp ((λ l λ { (i = i0) h (P j k) l + ; (i = i1) h (Q j k) l + ; (j = i0) h (p k) l + ; (j = i1) h (q k) l + ; (k = i0) h x l + ; (k = i1) h y l})) + (g (grp (f x) (f y) (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) i j k)) + +is2GroupoidRetract + : {B : Type } + (f : A B) (g : B A) + (h : (x : A) g (f x) x) + is2Groupoid B is2Groupoid A +is2GroupoidRetract f g h grp x y p q P Q R S i j k l = + hcomp r λ { (i = i0) h (R j k l) r + ; (i = i1) h (S j k l) r + ; (j = i0) h (P k l) r + ; (j = i1) h (Q k l) r + ; (k = i0) h (p l) r + ; (k = i1) h (q l) r + ; (l = i0) h x r + ; (l = i1) h y r}) + (g (grp (f x) (f y) (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) + (cong (cong (cong f)) R) (cong (cong (cong f)) S) i j k l)) + +isOfHLevelRetract + : (n : HLevel) {B : Type } + (f : A B) (g : B A) + (h : (x : A) g (f x) x) + isOfHLevel n B isOfHLevel n A +isOfHLevelRetract 0 = isContrRetract +isOfHLevelRetract 1 = isPropRetract +isOfHLevelRetract 2 = isSetRetract +isOfHLevelRetract 3 = isGroupoidRetract +isOfHLevelRetract 4 = is2GroupoidRetract +isOfHLevelRetract (suc (suc (suc (suc (suc n))))) f g h ofLevel x y p q P Q R S = + isOfHLevelRetract (suc n) (cong (cong (cong (cong f)))) + s i j k l + hcomp r λ { (i = i0) h (R j k l) r + ; (i = i1) h (S j k l) r + ; (j = i0) h (P k l) r + ; (j = i1) h (Q k l) r + ; (k = i0) h (p l) r + ; (k = i1) h (q l) r + ; (l = i0) h x r + ; (l = i1) h y r}) + (g (s i j k l))) + s i j k l m + hcomp n λ { (i = i1) s j k l m + ; (j = i0) h (R k l m) (i n) + ; (j = i1) h (S k l m) (i n) + ; (k = i0) h (P l m) (i n) + ; (k = i1) h (Q l m) (i n) + ; (l = i0) h (p m) (i n) + ; (l = i1) h (q m) (i n) + ; (m = i0) h x (i n) + ; (m = i1) h y (i n) }) + (h (s j k l m) i)) + (ofLevel (f x) (f y) + (cong f p) (cong f q) + (cong (cong f) P) (cong (cong f) Q) + (cong (cong (cong f)) R) (cong (cong (cong f)) S)) + +isOfHLevelRetractFromIso : {A : Type } {B : Type ℓ'} (n : HLevel) Iso A B isOfHLevel n B isOfHLevel n A +isOfHLevelRetractFromIso n e hlev = isOfHLevelRetract n (Iso.fun e) (Iso.inv e) (Iso.leftInv e) hlev + +isOfHLevelRespectEquiv : {A : Type } {B : Type ℓ'} (n : HLevel) A B isOfHLevel n A isOfHLevel n B +isOfHLevelRespectEquiv n eq = isOfHLevelRetract n (invEq eq) (eq .fst) (secEq eq) + +isContrRetractOfConstFun : {A : Type } {B : Type ℓ'} (b₀ : B) + Σ[ f (B A) ] ((x : A) (f _ b₀)) x x) + isContr A +fst (isContrRetractOfConstFun b₀ ret) = ret .fst b₀ +snd (isContrRetractOfConstFun b₀ ret) y = ret .snd y + +-- h-level of dependent path types + +isOfHLevelPathP' : {A : I Type } (n : HLevel) + isOfHLevel (suc n) (A i1) + (x : A i0) (y : A i1) isOfHLevel n (PathP A x y) +isOfHLevelPathP' {A = A} n h x y = + isOfHLevelRetractFromIso n (PathPIsoPath _ x y) (isOfHLevelPath' n h _ _) + +isOfHLevelPathP : {A : I Type } (n : HLevel) + isOfHLevel n (A i1) + (x : A i0) (y : A i1) isOfHLevel n (PathP A x y) +isOfHLevelPathP {A = A} n h x y = + isOfHLevelRetractFromIso n (PathPIsoPath _ x y) (isOfHLevelPath n h _ _) + +-- Fillers for cubes from h-level + +isSet→SquareP : + {A : I I Type } + (isSet : (i j : I) isSet (A i j)) + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP j A i0 j) a₀₀ a₀₁) + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP j A i1 j) a₁₀ a₁₁) + (a₋₀ : PathP i A i i0) a₀₀ a₁₀) (a₋₁ : PathP i A i i1) a₀₁ a₁₁) + SquareP A a₀₋ a₁₋ a₋₀ a₋₁ +isSet→SquareP isset a₀₋ a₁₋ a₋₀ a₋₁ = + PathPIsoPath _ _ _ .Iso.inv (isOfHLevelPathP' 1 (isset _ _) _ _ _ _ ) + +isGroupoid→isGroupoid' : isGroupoid A isGroupoid' A +isGroupoid→isGroupoid' {A = A} Agpd a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = + PathPIsoPath i Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ .Iso.inv + (isGroupoid→isPropSquare _ _ _ _ _ _) + where + isGroupoid→isPropSquare : + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) + (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) + isProp (Square a₀₋ a₁₋ a₋₀ a₋₁) + isGroupoid→isPropSquare a₀₋ a₁₋ a₋₀ a₋₁ = + isOfHLevelRetractFromIso 1 (PathPIsoPath i a₋₀ i a₋₁ i) a₀₋ a₁₋) (Agpd _ _ _ _) + +isGroupoid'→isGroupoid : isGroupoid' A isGroupoid A +isGroupoid'→isGroupoid Agpd' x y p q r s = Agpd' r s refl refl refl refl +-- h-level of Σ-types + +isProp∃! : isProp (∃! A B) +isProp∃! = isPropIsContr + +isContrΣ : isContr A ((x : A) isContr (B x)) isContr (Σ A B) +isContrΣ {A = A} {B = B} (a , p) q = + let h : (x : A) (y : B x) (q x) .fst y + h x y = (q x) .snd y + in (( a , q a .fst) + , ( λ x i p (x .fst) i + , h (p (x .fst) i) (transp j B (p (x .fst) (i ~ j))) i (x .snd)) i)) + +isContrΣ' : (ca : isContr A) isContr (B (fst ca)) isContr (Σ A B) +isContrΣ' ca cb = isContrΣ ca x subst _ (snd ca x) cb) + +section-Σ≡Prop + : (pB : (x : A) isProp (B x)) {u v : Σ A B} + section (Σ≡Prop pB {u} {v}) (cong fst) +section-Σ≡Prop {A = A} pB {u} {v} p j i = + (p i .fst) , isProp→PathP i isOfHLevelPath 1 (pB (fst (p i))) + (Σ≡Prop pB {u} {v} (cong fst p) i .snd) + (p i .snd) ) + refl refl i j + +isEquiv-Σ≡Prop + : (pB : (x : A) isProp (B x)) {u v : Σ A B} + isEquiv (Σ≡Prop pB {u} {v}) +isEquiv-Σ≡Prop {A = A} pB {u} {v} = isoToIsEquiv (iso (Σ≡Prop pB) (cong fst) (section-Σ≡Prop pB) _ refl)) + +isPropΣ : isProp A ((x : A) isProp (B x)) isProp (Σ A B) +isPropΣ pA pB t u = Σ≡Prop pB (pA (t .fst) (u .fst)) + +isOfHLevelΣ : n isOfHLevel n A ((x : A) isOfHLevel n (B x)) + isOfHLevel n (Σ A B) +isOfHLevelΣ 0 = isContrΣ +isOfHLevelΣ 1 = isPropΣ +isOfHLevelΣ {B = B} (suc (suc n)) h1 h2 x y = + isOfHLevelRetractFromIso (suc n) + (invIso (IsoΣPathTransportPathΣ _ _)) + (isOfHLevelΣ (suc n) (h1 (fst x) (fst y)) λ x h2 _ _ _) + +isSetΣ : isSet A ((x : A) isSet (B x)) isSet (Σ A B) +isSetΣ = isOfHLevelΣ 2 + +-- Useful consequence +isSetΣSndProp : isSet A ((x : A) isProp (B x)) isSet (Σ A B) +isSetΣSndProp h p = isSetΣ h x isProp→isSet (p x)) + +isGroupoidΣ : isGroupoid A ((x : A) isGroupoid (B x)) isGroupoid (Σ A B) +isGroupoidΣ = isOfHLevelΣ 3 + +is2GroupoidΣ : is2Groupoid A ((x : A) is2Groupoid (B x)) is2Groupoid (Σ A B) +is2GroupoidΣ = isOfHLevelΣ 4 + +-- h-level of × + +isProp× : {A : Type } {B : Type ℓ'} isProp A isProp B isProp (A × B) +isProp× pA pB = isPropΣ pA _ pB) + +isProp×2 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} + isProp A isProp B isProp C isProp (A × B × C) +isProp×2 pA pB pC = isProp× pA (isProp× pB pC) + +isProp×3 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} + isProp A isProp B isProp C isProp D isProp (A × B × C × D) +isProp×3 pA pB pC pD = isProp×2 pA pB (isProp× pC pD) + +isProp×4 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} + isProp A isProp B isProp C isProp D isProp E isProp (A × B × C × D × E) +isProp×4 pA pB pC pD pE = isProp×3 pA pB pC (isProp× pD pE) + +isProp×5 : {A : Type } {B : Type ℓ'} {C : Type ℓ''} {D : Type ℓ'''} {E : Type ℓ''''} {F : Type ℓ'''''} + isProp A isProp B isProp C isProp D isProp E isProp F + isProp (A × B × C × D × E × F) +isProp×5 pA pB pC pD pE pF = isProp×4 pA pB pC pD (isProp× pE pF) + + +isOfHLevel× : {A : Type } {B : Type ℓ'} n isOfHLevel n A isOfHLevel n B + isOfHLevel n (A × B) +isOfHLevel× n hA hB = isOfHLevelΣ n hA _ hB) + +isSet× : {A : Type } {B : Type ℓ'} isSet A isSet B isSet (A × B) +isSet× = isOfHLevel× 2 + +isGroupoid× : {A : Type } {B : Type ℓ'} isGroupoid A isGroupoid B + isGroupoid (A × B) +isGroupoid× = isOfHLevel× 3 + +is2Groupoid× : {A : Type } {B : Type ℓ'} is2Groupoid A is2Groupoid B + is2Groupoid (A × B) +is2Groupoid× = isOfHLevel× 4 + +-- h-level of Π-types + +isOfHLevelΠ : n ((x : A) isOfHLevel n (B x)) + isOfHLevel n ((x : A) B x) +isOfHLevelΠ 0 h = x fst (h x)) , λ f i y snd (h y) (f y) i +isOfHLevelΠ 1 h f g i x = (h x) (f x) (g x) i +isOfHLevelΠ 2 h f g F G i j z = h z (f z) (g z) (funExt⁻ F z) (funExt⁻ G z) i j +isOfHLevelΠ 3 h f g p q P Q i j k z = + h z (f z) (g z) + (funExt⁻ p z) (funExt⁻ q z) + (cong f funExt⁻ f z) P) (cong f funExt⁻ f z) Q) i j k +isOfHLevelΠ 4 h f g p q P Q R S i j k l z = + h z (f z) (g z) + (funExt⁻ p z) (funExt⁻ q z) + (cong f funExt⁻ f z) P) (cong f funExt⁻ f z) Q) + (cong (cong f funExt⁻ f z)) R) (cong (cong f funExt⁻ f z)) S) i j k l +isOfHLevelΠ (suc (suc (suc (suc (suc n))))) h f g p q P Q R S = + isOfHLevelRetract (suc n) + (cong (cong (cong funExt⁻))) (cong (cong (cong funExt))) _ refl) + (isOfHLevelΠ (suc (suc (suc (suc n)))) x h x (f x) (g x)) + (funExt⁻ p) (funExt⁻ q) + (cong funExt⁻ P) (cong funExt⁻ Q) + (cong (cong funExt⁻) R) (cong (cong funExt⁻) S)) + +isOfHLevelΠ2 : (n : HLevel) ((x : A) (y : B x) isOfHLevel n (C x y)) + isOfHLevel n ((x : A) (y : B x) C x y) +isOfHLevelΠ2 n f = isOfHLevelΠ n x isOfHLevelΠ n (f x)) + +isContrΠ : (h : (x : A) isContr (B x)) isContr ((x : A) B x) +isContrΠ = isOfHLevelΠ 0 + +isPropΠ : (h : (x : A) isProp (B x)) isProp ((x : A) B x) +isPropΠ = isOfHLevelΠ 1 + +isPropΠ2 : (h : (x : A) (y : B x) isProp (C x y)) + isProp ((x : A) (y : B x) C x y) +isPropΠ2 h = isPropΠ λ x isPropΠ λ y h x y + +isPropΠ3 : (h : (x : A) (y : B x) (z : C x y) isProp (D x y z)) + isProp ((x : A) (y : B x) (z : C x y) D x y z) +isPropΠ3 h = isPropΠ λ x isPropΠ λ y isPropΠ λ z h x y z + +isPropΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isProp (E x y z w)) + isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) E x y z w) +isPropΠ4 h = isPropΠ λ _ isPropΠ3 (h _) + +isPropΠ5 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) isProp (F x y z w v)) + isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) F x y z w v) +isPropΠ5 h = isPropΠ λ _ isPropΠ4 (h _) + +isPropΠ6 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) (u : F x y z w v) isProp (G x y z w v u)) + isProp ((x : A) (y : B x) (z : C x y) (w : D x y z) (v : E x y z w) (u : F x y z w v) G x y z w v u) +isPropΠ6 h = isPropΠ λ _ isPropΠ5 (h _) + +isPropImplicitΠ : (h : (x : A) isProp (B x)) isProp ({x : A} B x) +isPropImplicitΠ h f g i {x} = h x (f {x}) (g {x}) i + +isPropImplicitΠ2 : (h : (x : A) (y : B x) isProp (C x y)) isProp ({x : A} {y : B x} C x y) +isPropImplicitΠ2 h = isPropImplicitΠ x isPropImplicitΠ y h x y)) + +isPropImplicitΠ3 : (h : (x : A) (y : B x) (z : C x y) isProp (D x y z)) + isProp ({x : A} {y : B x} {z : C x y} D x y z) +isPropImplicitΠ3 h = isPropImplicitΠ x isPropImplicitΠ2 y h x y)) + +isPropImplicitΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isProp (E x y z w)) + isProp ({x : A} {y : B x} {z : C x y} {w : D x y z} E x y z w) +isPropImplicitΠ4 h = isPropImplicitΠ x isPropImplicitΠ3 y h x y)) + +isProp→ : {A : Type } {B : Type ℓ'} isProp B isProp (A B) +isProp→ pB = isPropΠ λ _ pB + +isSetΠ : ((x : A) isSet (B x)) isSet ((x : A) B x) +isSetΠ = isOfHLevelΠ 2 + +isSetImplicitΠ : (h : (x : A) isSet (B x)) isSet ({x : A} B x) +isSetImplicitΠ h f g F G i j {x} = h x (f {x}) (g {x}) i F i {x}) i G i {x}) i j + +isSetImplicitΠ2 : (h : (x : A) (y : B x) isSet (C x y)) isSet ({x : A} {y : B x} C x y) +isSetImplicitΠ2 h = isSetImplicitΠ x isSetImplicitΠ y h x y)) + +isSetImplicitΠ3 : (h : (x : A) (y : B x) (z : C x y) isSet (D x y z)) + isSet ({x : A} {y : B x} {z : C x y} D x y z) +isSetImplicitΠ3 h = isSetImplicitΠ x isSetImplicitΠ2 y λ z h x y z)) + +isSet→ : isSet A' isSet (A A') +isSet→ isSet-A' = isOfHLevelΠ 2 _ isSet-A') + +isSetΠ2 : (h : (x : A) (y : B x) isSet (C x y)) + isSet ((x : A) (y : B x) C x y) +isSetΠ2 h = isSetΠ λ x isSetΠ λ y h x y + +isSetΠ3 : (h : (x : A) (y : B x) (z : C x y) isSet (D x y z)) + isSet ((x : A) (y : B x) (z : C x y) D x y z) +isSetΠ3 h = isSetΠ λ x isSetΠ λ y isSetΠ λ z h x y z + +isGroupoidΠ : ((x : A) isGroupoid (B x)) isGroupoid ((x : A) B x) +isGroupoidΠ = isOfHLevelΠ 3 + +isGroupoidΠ2 : (h : (x : A) (y : B x) isGroupoid (C x y)) isGroupoid ((x : A) (y : B x) C x y) +isGroupoidΠ2 h = isGroupoidΠ λ _ isGroupoidΠ λ _ h _ _ + +isGroupoidΠ3 : (h : (x : A) (y : B x) (z : C x y) isGroupoid (D x y z)) + isGroupoid ((x : A) (y : B x) (z : C x y) D x y z) +isGroupoidΠ3 h = isGroupoidΠ λ _ isGroupoidΠ2 λ _ h _ _ + +isGroupoidΠ4 : (h : (x : A) (y : B x) (z : C x y) (w : D x y z) isGroupoid (E x y z w)) + isGroupoid ((x : A) (y : B x) (z : C x y) (w : D x y z) E x y z w) +isGroupoidΠ4 h = isGroupoidΠ λ _ isGroupoidΠ3 λ _ h _ _ + +is2GroupoidΠ : ((x : A) is2Groupoid (B x)) is2Groupoid ((x : A) B x) +is2GroupoidΠ = isOfHLevelΠ 4 + +isOfHLevelΠ⁻ : {A : Type } {B : Type ℓ'} n + isOfHLevel n (A B) (A isOfHLevel n B) +isOfHLevelΠ⁻ 0 h x = fst h x , λ y funExt⁻ (snd h (const y)) x +isOfHLevelΠ⁻ 1 h x y z = funExt⁻ (h (const y) (const z)) x +isOfHLevelΠ⁻ (suc (suc n)) h x y z = + isOfHLevelΠ⁻ (suc n) (isOfHLevelRetractFromIso (suc n) funExtIso (h _ _)) x + +isOfHLevel→∙ : {A : Pointed } {B : Pointed ℓ'} (n : ) + isOfHLevel n (fst B) isOfHLevel n (A →∙ B) +isOfHLevel→∙ n hlev = + isOfHLevelΣ n (isOfHLevelΠ n _ hlev)) + λ _ isOfHLevelPath n hlev _ _ + +-- h-level of A ≃ B and A ≡ B + +isOfHLevel≃ + : n {A : Type } {B : Type ℓ'} + (hA : isOfHLevel n A) (hB : isOfHLevel n B) isOfHLevel n (A B) +isOfHLevel≃ zero {A = A} {B = B} hA hB = isContr→Equiv hA hB , contr + where + contr : (y : A B) isContr→Equiv hA hB y + contr y = Σ≡Prop isPropIsEquiv (funExt a snd hB (fst y a))) + +isOfHLevel≃ (suc n) {A = A} {B = B} hA hB = + isOfHLevelΣ (suc n) (isOfHLevelΠ _ λ _ hB) + f isProp→isOfHLevelSuc n (isPropIsEquiv f)) + +isOfHLevel≡ : n {A B : Type } (hA : isOfHLevel n A) (hB : isOfHLevel n B) + isOfHLevel n (A B) +isOfHLevel≡ n hA hB = isOfHLevelRetractFromIso n univalenceIso (isOfHLevel≃ n hA hB) + +isOfHLevel⁺≃ₗ + : n {A : Type } {B : Type ℓ'} + isOfHLevel (suc n) A isOfHLevel (suc n) (A B) +isOfHLevel⁺≃ₗ zero pA e = isOfHLevel≃ 1 pA (isOfHLevelRespectEquiv 1 e pA) e +isOfHLevel⁺≃ₗ (suc n) hA e = isOfHLevel≃ m hA (isOfHLevelRespectEquiv m e hA) e + where + m = suc (suc n) + +isOfHLevel⁺≃ᵣ + : n {A : Type } {B : Type ℓ'} + isOfHLevel (suc n) B isOfHLevel (suc n) (A B) +isOfHLevel⁺≃ᵣ zero pB e + = isOfHLevel≃ 1 (isPropRetract (e .fst) (invEq e) (retEq e) pB) pB e +isOfHLevel⁺≃ᵣ (suc n) hB e + = isOfHLevel≃ m (isOfHLevelRetract m (e .fst) (invEq e) (retEq e) hB) hB e + where + m = suc (suc n) + +isOfHLevel⁺≡ₗ + : n {A B : Type } + isOfHLevel (suc n) A isOfHLevel (suc n) (A B) +isOfHLevel⁺≡ₗ zero pA P = isOfHLevel≡ 1 pA (subst isProp P pA) P +isOfHLevel⁺≡ₗ (suc n) hA P + = isOfHLevel≡ m hA (subst (isOfHLevel m) P hA) P + where + m = suc (suc n) + +isOfHLevel⁺≡ᵣ + : n {A B : Type } + isOfHLevel (suc n) B isOfHLevel (suc n) (A B) +isOfHLevel⁺≡ᵣ zero pB P = isOfHLevel≡ 1 (subst⁻ isProp P pB) pB P +isOfHLevel⁺≡ᵣ (suc n) hB P + = isOfHLevel≡ m (subst⁻ (isOfHLevel m) P hB) hB P + where + m = suc (suc n) + +-- h-level of TypeOfHLevel + +isPropHContr : isProp (TypeOfHLevel 0) +isPropHContr x y = Σ≡Prop _ isPropIsContr) (isOfHLevel≡ 0 (x .snd) (y .snd) .fst) + +isOfHLevelTypeOfHLevel : n isOfHLevel (suc n) (TypeOfHLevel n) +isOfHLevelTypeOfHLevel zero = isPropHContr +isOfHLevelTypeOfHLevel (suc n) (X , a) (Y , b) = + isOfHLevelRetract (suc n) (cong fst) (Σ≡Prop λ _ isPropIsOfHLevel (suc n)) + (section-Σ≡Prop λ _ isPropIsOfHLevel (suc n)) + (isOfHLevel≡ (suc n) a b) + +isSetHProp : isSet (hProp ) +isSetHProp = isOfHLevelTypeOfHLevel 1 + +isGroupoidHSet : isGroupoid (hSet ) +isGroupoidHSet = isOfHLevelTypeOfHLevel 2 + + +-- h-level of lifted type + +isOfHLevelLift : { ℓ'} (n : HLevel) {A : Type } isOfHLevel n A isOfHLevel n (Lift {j = ℓ'} A) +isOfHLevelLift n = isOfHLevelRetract n lower lift λ _ refl + +isOfHLevelLower : { ℓ'} (n : HLevel) {A : Type } isOfHLevel n (Lift {j = ℓ'} A) isOfHLevel n A +isOfHLevelLower n = isOfHLevelRetract n lift lower λ _ refl + +---------------------------- + +-- More consequences of isProp and isContr + +inhProp→isContr : A isProp A isContr A +inhProp→isContr x h = x , h x + +extend : isContr A (∀ φ (u : Partial φ A) Sub A φ u) +extend (x , p) φ u = inS (hcomp { j (φ = i1) p (u 1=1) j }) x) + +isContrPartial→isContr : {} {A : Type } + (extend : φ Partial φ A A) + (∀ u u (extend i1 λ { _ u})) + isContr A +isContrPartial→isContr {A = A} extend law + = ex , λ y law ex i Aux.v y i) sym (law y) + where ex = extend i0 empty + module Aux (y : A) (i : I) where + φ = ~ i i + u : Partial φ A + u = λ { (i = i0) ex ; (i = i1) y } + v = extend φ u + +-- Dependent h-level over a type + +isOfHLevelDep : HLevel {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') +isOfHLevelDep 0 {A = A} B = {a : A} Σ[ b B a ] ({a' : A} (b' : B a') (p : a a') PathP i B (p i)) b b') +isOfHLevelDep 1 {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) (p : a0 a1) PathP i B (p i)) b0 b1 +isOfHLevelDep (suc (suc n)) {A = A} B = {a0 a1 : A} (b0 : B a0) (b1 : B a1) isOfHLevelDep (suc n) {A = a0 a1} p PathP i B (p i)) b0 b1) + +isContrDep : {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') +isContrDep = isOfHLevelDep 0 + +isPropDep : {A : Type } (B : A Type ℓ') Type (ℓ-max ℓ') +isPropDep = isOfHLevelDep 1 + +isContrDep∘ + : {A' : Type } (f : A' A) isContrDep B isContrDep (B f) +isContrDep∘ f cB {a} = λ where + .fst cB .fst + .snd b' p cB .snd b' (cong f p) + +isPropDep∘ : {A' : Type } (f : A' A) isPropDep B isPropDep (B f) +isPropDep∘ f pB b0 b1 = pB b0 b1 cong f + +isOfHLevelDep→isOfHLevel : (n : HLevel) + {A : Type } {B : A Type ℓ'} isOfHLevelDep n {A = A} B (a : A) isOfHLevel n (B a) +isOfHLevelDep→isOfHLevel 0 h a = h .fst , λ b h .snd b refl +isOfHLevelDep→isOfHLevel 1 h a x y = h x y refl +isOfHLevelDep→isOfHLevel (suc (suc n)) h a x y = isOfHLevelDep→isOfHLevel (suc n) (h x y) refl + +isOfHLevel→isOfHLevelDep : (n : HLevel) + {A : Type } {B : A Type ℓ'} (h : (a : A) isOfHLevel n (B a)) isOfHLevelDep n {A = A} B +isOfHLevel→isOfHLevelDep 0 h {a} = + (h a .fst , λ b' p isProp→PathP i isContr→isProp (h (p i))) (h a .fst) b') +isOfHLevel→isOfHLevelDep 1 h = λ b0 b1 p isProp→PathP i h (p i)) b0 b1 +isOfHLevel→isOfHLevelDep (suc (suc n)) {A = A} {B} h {a0} {a1} b0 b1 = + isOfHLevel→isOfHLevelDep (suc n) p helper p) + where + helper : (p : a0 a1) + isOfHLevel (suc n) (PathP i B (p i)) b0 b1) + helper p = J a1 p b1 isOfHLevel (suc n) (PathP i B (p i)) b0 b1)) + _ h _ _ _) p b1 + +isContrDep→isPropDep : isOfHLevelDep 0 B isOfHLevelDep 1 B +isContrDep→isPropDep {B = B} Bctr {a0 = a0} b0 b1 p i + = comp k B (p (i k))) k λ where + (i = i0) Bctr .snd b0 refl k + (i = i1) Bctr .snd b1 p k) + (c0 .fst) + where + c0 = Bctr {a0} + +isPropDep→isSetDep : isOfHLevelDep 1 B isOfHLevelDep 2 B +isPropDep→isSetDep {B = B} Bprp b0 b1 b2 b3 p i j + = comp k B (p (i k) (j k))) k λ where + (j = i0) Bprp b0 b0 refl k + (i = i0) Bprp b0 (b2 j) k p i0 (j k)) k + (i = i1) Bprp b0 (b3 j) k p k (j k)) k + (j = i1) Bprp b0 b1 k p (i k) (j k)) k) + b0 + +isOfHLevelDepSuc : (n : HLevel) isOfHLevelDep n B isOfHLevelDep (suc n) B +isOfHLevelDepSuc 0 = isContrDep→isPropDep +isOfHLevelDepSuc 1 = isPropDep→isSetDep +isOfHLevelDepSuc (suc (suc n)) Blvl b0 b1 = isOfHLevelDepSuc (suc n) (Blvl b0 b1) + +isPropDep→isSetDep' + : isOfHLevelDep 1 B + {p : w x} {q : y z} {r : w y} {s : x z} + {tw : B w} {tx : B x} {ty : B y} {tz : B z} + (sq : Square p q r s) + (tp : PathP i B (p i)) tw tx) + (tq : PathP i B (q i)) ty tz) + (tr : PathP i B (r i)) tw ty) + (ts : PathP i B (s i)) tx tz) + SquareP i j B (sq i j)) tp tq tr ts +isPropDep→isSetDep' {B = B} Bprp {p} {q} {r} {s} {tw} sq tp tq tr ts i j + = comp k B (sq (i k) (j k))) k λ where + (i = i0) Bprp tw (tp j) k p (k j)) k + (i = i1) Bprp tw (tq j) k sq (i k) (j k)) k + (j = i0) Bprp tw (tr i) k r (k i)) k + (j = i1) Bprp tw (ts i) k sq (k i) (j k)) k) + tw + +isOfHLevelΣ' : n isOfHLevel n A isOfHLevelDep n B isOfHLevel n (Σ A B) +isOfHLevelΣ' 0 Actr Bctr .fst = (Actr .fst , Bctr .fst) +isOfHLevelΣ' 0 Actr Bctr .snd (x , y) i + = Actr .snd x i , Bctr .snd y (Actr .snd x) i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .fst = Alvl w x i +isOfHLevelΣ' 1 Alvl Blvl (w , y) (x , z) i .snd = Blvl y z (Alvl w x) i +isOfHLevelΣ' {A = A} {B = B} (suc (suc n)) Alvl Blvl (w , y) (x , z) + = isOfHLevelRetract (suc n) + p i p i .fst) , λ i p i .snd) + ΣPathP + x refl) + (isOfHLevelΣ' (suc n) (Alvl w x) (Blvl y z)) + +ΣSquareSet : ((x : A) isSet (B x)) {u v w x : Σ A B} + {p : u v} {q : v w} {r : x w} {s : u x} + Square (cong fst p) (cong fst r) + (cong fst s) (cong fst q) + Square p r s q +fst (ΣSquareSet pB sq i j) = sq i j +snd (ΣSquareSet {B = B} pB {p = p} {q = q} {r = r} {s = s} sq i j) = lem i j + where + lem : SquareP i j B (sq i j)) + (cong snd p) (cong snd r) (cong snd s) (cong snd q) + lem = toPathP (isOfHLevelPathP' 1 (pB _) _ _ _ _) + +module _ (isSet-A : isSet A) (isSet-A' : isSet A') where + + + isSet-SetsIso : isSet (Iso A A') + isSet-SetsIso x y p₀ p₁ = h + where + + module X = Iso x + module Y = Iso y + + f-p : i₁ (Iso.fun (p₀ i₁) , Iso.inv (p₀ i₁)) + (Iso.fun (p₁ i₁) , Iso.inv (p₁ i₁)) + fst (f-p i₁ i) a = isSet-A' (X.fun a ) (Y.fun a ) (cong _ p₀) (cong _ p₁) i i₁ + snd (f-p i₁ i) a' = isSet-A (X.inv a') (Y.inv a') (cong _ p₀) (cong _ p₁) i i₁ + + s-p : b _ + s-p b = + isSet→SquareP i j isProp→isSet (isSet-A' _ _)) + refl refl i₁ (Iso.rightInv (p₀ i₁) b)) i₁ (Iso.rightInv (p₁ i₁) b)) + + r-p : a _ + r-p a = + isSet→SquareP i j isProp→isSet (isSet-A _ _)) + refl refl i₁ (Iso.leftInv (p₀ i₁) a)) i₁ (Iso.leftInv (p₁ i₁) a)) + + + h : p₀ p₁ + Iso.fun (h i i₁) = fst (f-p i₁ i) + Iso.inv (h i i₁) = snd (f-p i₁ i) + Iso.rightInv (h i i₁) b = s-p b i₁ i + Iso.leftInv (h i i₁) a = r-p a i₁ i + + + SetsIso≡-ext : {a b : Iso A A'} + (∀ x Iso.fun a x Iso.fun b x) + (∀ x Iso.inv a x Iso.inv b x) + a b + Iso.fun (SetsIso≡-ext {a} {b} fun≡ inv≡ i) x = fun≡ x i + Iso.inv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) x = inv≡ x i + Iso.rightInv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) b₁ = + isSet→SquareP _ _ isSet-A') + (Iso.rightInv a b₁) + (Iso.rightInv b b₁) + i fun≡ (inv≡ b₁ i) i) + refl i + Iso.leftInv (SetsIso≡-ext {a} {b} fun≡ inv≡ i) a₁ = + isSet→SquareP _ _ isSet-A) + (Iso.leftInv a a₁) + (Iso.leftInv b a₁) + i inv≡ (fun≡ a₁ i) i ) + refl i + + SetsIso≡ : {a b : Iso A A'} + (Iso.fun a Iso.fun b) + (Iso.inv a Iso.inv b) + a b + SetsIso≡ p q = + SetsIso≡-ext (funExt⁻ p) (funExt⁻ q) + + isSet→Iso-Iso-≃ : Iso (Iso A A') (A A') + isSet→Iso-Iso-≃ = ww + where + open Iso + + ww : Iso _ _ + fun ww = isoToEquiv + inv ww = equivToIso + rightInv ww b = equivEq refl + leftInv ww a = SetsIso≡ refl refl + + + isSet→isEquiv-isoToPath : isEquiv isoToEquiv + isSet→isEquiv-isoToPath = isoToIsEquiv isSet→Iso-Iso-≃ + + + +isSet→Iso-Iso-≡ : (isSet-A : isSet A) (isSet-A' : isSet A') Iso (Iso A A') (A A') +isSet→Iso-Iso-≡ isSet-A isSet-A' = ww + where + open Iso + + ww : Iso _ _ + fun ww = isoToPath + inv ww = pathToIso + rightInv ww b = isInjectiveTransport (funExt λ _ transportRefl _) + leftInv ww a = SetsIso≡-ext isSet-A isSet-A' _ transportRefl (fun a _)) λ _ cong (inv a) (transportRefl _) + +hSet-Iso-Iso-≡ : (A : hSet ) (A' : hSet ) Iso (Iso (fst A) (fst A')) (A A') +hSet-Iso-Iso-≡ A A' = compIso (isSet→Iso-Iso-≡ (snd A) (snd A')) (equivToIso (_ , isEquiv-Σ≡Prop λ _ isPropIsSet)) + +module _ (B : (i j k : I) Type ) + {c₀₀₀ : B i0 i0 i0} {c₀₀₁ : B i0 i0 i1} {c₀₁₀ : B i0 i1 i0} {c₀₁₁ : B i0 i1 i1} + {c₁₀₀ : B i1 i0 i0} {c₁₀₁ : B i1 i0 i1} {c₁₁₀ : B i1 i1 i0} {c₁₁₁ : B i1 i1 i1} + {c₀₀₋ : PathP k B i0 i0 k) c₀₀₀ c₀₀₁} {c₀₁₋ : PathP k B i0 i1 k) c₀₁₀ c₀₁₁} + {c₀₋₀ : PathP i B i0 i i0) c₀₀₀ c₀₁₀} {c₀₋₁ : PathP i B i0 i i1) c₀₀₁ c₀₁₁} + {c₁₀₋ : PathP k B i1 i0 k) c₁₀₀ c₁₀₁} {c₁₁₋ : PathP k B i1 i1 k) c₁₁₀ c₁₁₁} + {c₁₋₀ : PathP i B i1 i i0) c₁₀₀ c₁₁₀} {c₁₋₁ : PathP i B i1 i i1) c₁₀₁ c₁₁₁} + {c₋₀₀ : PathP i B i i0 i0) c₀₀₀ c₁₀₀} {c₋₀₁ : PathP i B i i0 i1) c₀₀₁ c₁₀₁} + {c₋₁₀ : PathP i B i i1 i0) c₀₁₀ c₁₁₀} {c₋₁₁ : PathP i B i i1 i1) c₀₁₁ c₁₁₁} + (c₀₋₋ : SquareP j k B i0 j k) c₀₀₋ c₀₁₋ c₀₋₀ c₀₋₁) + (c₁₋₋ : SquareP j k B i1 j k) c₁₀₋ c₁₁₋ c₁₋₀ c₁₋₁) + (c₋₀₋ : SquareP i k B i i0 k) c₀₀₋ c₁₀₋ c₋₀₀ c₋₀₁) + (c₋₁₋ : SquareP i k B i i1 k) c₀₁₋ c₁₁₋ c₋₁₀ c₋₁₁) + (c₋₋₀ : SquareP i j B i j i0) c₀₋₀ c₁₋₀ c₋₀₀ c₋₁₀) + (c₋₋₁ : SquareP i j B i j i1) c₀₋₁ c₁₋₁ c₋₀₁ c₋₁₁) where + + CubeP : Type + CubeP = PathP i SquareP j k B i j k) + (c₋₀₋ i) (c₋₁₋ i) + (c₋₋₀ i) (c₋₋₁ i)) + c₀₋₋ c₁₋₋ + + isGroupoid→CubeP : isGroupoid (B i1 i1 i1) CubeP + isGroupoid→CubeP grpd = + isOfHLevelPathP' 0 (isOfHLevelPathP' 1 (isOfHLevelPathP' 2 grpd _ _) _ _) _ _ .fst diff --git a/src/docs/Cubical.Foundations.Isomorphism.html b/src/docs/Cubical.Foundations.Isomorphism.html new file mode 100644 index 0000000..d3908d1 --- /dev/null +++ b/src/docs/Cubical.Foundations.Isomorphism.html @@ -0,0 +1,225 @@ +{- + +Theory about isomorphisms + +- Definitions of [section] and [retract] +- Definition of isomorphisms ([Iso]) +- Any isomorphism is an equivalence ([isoToEquiv]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Isomorphism where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.Base + +private + variable + ℓ' : Level + A B C : Type + +-- Section and retract +module _ { ℓ'} {A : Type } {B : Type ℓ'} where + section : (f : A B) (g : B A) Type ℓ' + section f g = b f (g b) b + + -- NB: `g` is the retraction! + retract : (f : A B) (g : B A) Type + retract f g = a g (f a) a + +record Iso { ℓ'} (A : Type ) (B : Type ℓ') : Type (ℓ-max ℓ') where + no-eta-equality + constructor iso + field + fun : A B + inv : B A + rightInv : section fun inv + leftInv : retract fun inv + +isIso : (A B) Type _ +isIso {A = A} {B = B} f = Σ[ g (B A) ] Σ[ _ section f g ] retract f g + +isoFunInjective : (f : Iso A B) (x y : A) Iso.fun f x Iso.fun f y x y +isoFunInjective f x y h = sym (Iso.leftInv f x) ∙∙ cong (Iso.inv f) h ∙∙ Iso.leftInv f y + +isoInvInjective : (f : Iso A B) (x y : B) Iso.inv f x Iso.inv f y x y +isoInvInjective f x y h = sym (Iso.rightInv f x) ∙∙ cong (Iso.fun f) h ∙∙ Iso.rightInv f y + +-- Any iso is an equivalence +module _ (i : Iso A B) where + open Iso i renaming ( fun to f + ; inv to g + ; rightInv to s + ; leftInv to t) + + private + module _ (y : B) (x0 x1 : A) (p0 : f x0 y) (p1 : f x1 y) where + fill0 : I I A + fill0 i = hfill k λ { (i = i1) t x0 k + ; (i = i0) g y }) + (inS (g (p0 (~ i)))) + + fill1 : I I A + fill1 i = hfill k λ { (i = i1) t x1 k + ; (i = i0) g y }) + (inS (g (p1 (~ i)))) + + fill2 : I I A + fill2 i = hfill k λ { (i = i1) fill1 k i1 + ; (i = i0) fill0 k i1 }) + (inS (g y)) + + p : x0 x1 + p i = fill2 i i1 + + sq : I I A + sq i j = hcomp k λ { (i = i1) fill1 j (~ k) + ; (i = i0) fill0 j (~ k) + ; (j = i1) t (fill2 i i1) (~ k) + ; (j = i0) g y }) + (fill2 i j) + + sq1 : I I B + sq1 i j = hcomp k λ { (i = i1) s (p1 (~ j)) k + ; (i = i0) s (p0 (~ j)) k + ; (j = i1) s (f (p i)) k + ; (j = i0) s y k }) + (f (sq i j)) + + lemIso : (x0 , p0) (x1 , p1) + lemIso i .fst = p i + lemIso i .snd = λ j sq1 i (~ j) + + isoToIsEquiv : isEquiv f + isoToIsEquiv .equiv-proof y .fst .fst = g y + isoToIsEquiv .equiv-proof y .fst .snd = s y + isoToIsEquiv .equiv-proof y .snd z = lemIso y (g y) (fst z) (s y) (snd z) + + +isoToEquiv : Iso A B A B +isoToEquiv i .fst = i .Iso.fun +isoToEquiv i .snd = isoToIsEquiv i + +isoToPath : Iso A B A B +isoToPath {A = A} {B = B} f i = + Glue B { (i = i0) (A , isoToEquiv f) + ; (i = i1) (B , idEquiv B) }) + +open Iso + +invIso : Iso A B Iso B A +fun (invIso f) = inv f +inv (invIso f) = fun f +rightInv (invIso f) = leftInv f +leftInv (invIso f) = rightInv f + +compIso : Iso A B Iso B C Iso A C +fun (compIso i j) = fun j fun i +inv (compIso i j) = inv i inv j +rightInv (compIso i j) b = cong (fun j) (rightInv i (inv j b)) rightInv j b +leftInv (compIso i j) a = cong (inv i) (leftInv j (fun i a)) leftInv i a + +composesToId→Iso : (G : Iso A B) (g : B A) G .fun g idfun B Iso B A +fun (composesToId→Iso _ g _) = g +inv (composesToId→Iso j _ _) = fun j +rightInv (composesToId→Iso i g path) b = + sym (leftInv i (g (fun i b))) ∙∙ cong g inv i (g (fun i b))) path ∙∙ leftInv i b +leftInv (composesToId→Iso _ _ path) b i = path i b + +idIso : Iso A A +fun idIso = idfun _ +inv idIso = idfun _ +rightInv idIso _ = refl +leftInv idIso _ = refl + +compIsoIdL : (isom : Iso A B) compIso idIso isom isom +fun (compIsoIdL isom i) = fun isom +inv (compIsoIdL isom i) = inv isom +rightInv (compIsoIdL isom i) b = lUnit (isom .rightInv b) (~ i) +leftInv (compIsoIdL isom i) a = rUnit (isom .leftInv a) (~ i) + +compIsoIdR : (isom : Iso A B) compIso isom idIso isom +fun (compIsoIdR isom i) = fun isom +inv (compIsoIdR isom i) = inv isom +rightInv (compIsoIdR isom i) b = rUnit (isom .rightInv b) (~ i) +leftInv (compIsoIdR isom i) a = lUnit (isom .leftInv a) (~ i) + +LiftIso : Iso A (Lift {i = } {j = ℓ'} A) +fun LiftIso = lift +inv LiftIso = lower +rightInv LiftIso _ = refl +leftInv LiftIso _ = refl + +isContr→Iso : isContr A isContr B Iso A B +fun (isContr→Iso _ Bctr) _ = Bctr .fst +inv (isContr→Iso Actr _) _ = Actr .fst +rightInv (isContr→Iso _ Bctr) = Bctr .snd +leftInv (isContr→Iso Actr _) = Actr .snd + +isContr→Iso' : isContr A isContr B (A B) Iso A B +fun (isContr→Iso' _ Bctr f) = f +inv (isContr→Iso' Actr _ _) _ = Actr .fst +rightInv (isContr→Iso' _ Bctr f) = isContr→isProp Bctr _ +leftInv (isContr→Iso' Actr _ _) = Actr .snd + +isProp→Iso : (Aprop : isProp A) (Bprop : isProp B) (f : A B) (g : B A) Iso A B +fun (isProp→Iso _ _ f _) = f +inv (isProp→Iso _ _ _ g) = g +rightInv (isProp→Iso _ Bprop f g) b = Bprop (f (g b)) b +leftInv (isProp→Iso Aprop _ f g) a = Aprop (g (f a)) a + +domIso : {} {C : Type } Iso A B Iso (A C) (B C) +fun (domIso e) f b = f (inv e b) +inv (domIso e) f a = f (fun e a) +rightInv (domIso e) f i x = f (rightInv e x i) +leftInv (domIso e) f i x = f (leftInv e x i) + +-- Helpful notation +_Iso⟨_⟩_ : { ℓ' ℓ''} {B : Type ℓ'} {C : Type ℓ''} (X : Type ) Iso X B Iso B C Iso X C +_ Iso⟨ f g = compIso f g + +_∎Iso : {} (A : Type ) Iso A A +A ∎Iso = idIso {A = A} + +infixr 0 _Iso⟨_⟩_ +infix 1 _∎Iso + +codomainIsoDep : { ℓ' ℓ''} {A : Type } {B : A Type ℓ'} {C : A Type ℓ''} + ((a : A) Iso (B a) (C a)) + Iso ((a : A) B a) ((a : A) C a) +fun (codomainIsoDep is) f a = fun (is a) (f a) +inv (codomainIsoDep is) f a = inv (is a) (f a) +rightInv (codomainIsoDep is) f = funExt λ a rightInv (is a) (f a) +leftInv (codomainIsoDep is) f = funExt λ a leftInv (is a) (f a) + +codomainIso : { ℓ' ℓ''} {A : Type } {B : Type ℓ'} {C : Type ℓ''} + Iso B C + Iso (A B) (A C) +codomainIso z = codomainIsoDep λ _ z + +endoIso : Iso A B Iso (A A) (B B) +endoIso is = compIso (domIso is) (codomainIso is) + +binaryOpIso : Iso A B Iso (A A A) (B B B) +binaryOpIso is = compIso (domIso is) (codomainIso (endoIso is)) + +Iso≡Set : isSet A isSet B (f g : Iso A B) + ((x : A) f .fun x g .fun x) + ((x : B) f .inv x g .inv x) + f g +fun (Iso≡Set hA hB f g hfun hinv i) x = hfun x i +inv (Iso≡Set hA hB f g hfun hinv i) x = hinv x i +rightInv (Iso≡Set hA hB f g hfun hinv i) x j = + isSet→isSet' hB (rightInv f x) (rightInv g x) i hfun (hinv x i) i) refl i j +leftInv (Iso≡Set hA hB f g hfun hinv i) x j = + isSet→isSet' hA (leftInv f x) (leftInv g x) i hinv (hfun x i) i) refl i j + +transportIsoToPath : (f : Iso A B) (x : A) transport (isoToPath f) x f .fun x +transportIsoToPath f x = transportRefl _ + +transportIsoToPath⁻ : (f : Iso A B) (x : B) transport (sym (isoToPath f)) x f .inv x +transportIsoToPath⁻ f x = cong (f .inv) (transportRefl _) diff --git a/src/docs/Cubical.Foundations.Path.html b/src/docs/Cubical.Foundations.Path.html new file mode 100644 index 0000000..c0a1026 --- /dev/null +++ b/src/docs/Cubical.Foundations.Path.html @@ -0,0 +1,439 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Path where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence + +open import Cubical.Reflection.StrictEquiv + +private + variable + ℓ' : Level + A : Type + +module _ {A : I Type } {x : A i0} {y : A i1} where + toPathP⁻ : x transport⁻ i A i) y PathP A x y + toPathP⁻ p = symP (toPathP (sym p)) + + fromPathP⁻ : PathP A x y x transport⁻ i A i) y + fromPathP⁻ p = sym (fromPathP {A = λ i A (~ i)} (symP p)) + +PathP≡Path : (P : I Type ) (p : P i0) (q : P i1) + PathP P p q Path (P i1) (transport i P i) p) q +PathP≡Path P p q i = PathP j P (i j)) (transport-filler j P j) p i) q + +PathP≡Path⁻ : (P : I Type ) (p : P i0) (q : P i1) + PathP P p q Path (P i0) p (transport⁻ i P i) q) +PathP≡Path⁻ P p q i = PathP j P (~ i j)) p (transport⁻-filler j P j) q i) + +PathPIsoPath : (A : I Type ) (x : A i0) (y : A i1) Iso (PathP A x y) (transport i A i) x y) +PathPIsoPath A x y .Iso.fun = fromPathP +PathPIsoPath A x y .Iso.inv = toPathP +PathPIsoPath A x y .Iso.rightInv q k i = + hcomp + j λ + { (i = i0) slide (j ~ k) + ; (i = i1) q j + ; (k = i0) transp l A (i l)) i (fromPathPFiller j) + ; (k = i1) ∧∨Square i j + }) + (transp l A (i ~ k l)) (i ~ k) + (transp l (A (i (~ k l)))) (k i) + (transp l A (i l)) (~ i) + x))) + where + fromPathPFiller : _ + fromPathPFiller = + hfill + j λ + { (i = i0) x + ; (i = i1) q j }) + (inS (transp j A (i j)) (~ i) x)) + + slide : I _ + slide i = transp l A (i l)) i (transp l A (i l)) (~ i) x) + + ∧∨Square : I I _ + ∧∨Square i j = + hcomp + l λ + { (i = i0) slide j + ; (i = i1) q (j l) + ; (j = i0) slide i + ; (j = i1) q (i l) + }) + (slide (i j)) +PathPIsoPath A x y .Iso.leftInv q k i = + outS + (hcomp-unique + j λ + { (i = i0) x + ; (i = i1) transp l A (j l)) j (q j) + }) + (inS (transp l A (i l)) (~ i) x)) + j inS (transp l A (i (j l))) (~ i j) (q (i j))))) + k + +PathP≃Path : (A : I Type ) (x : A i0) (y : A i1) + PathP A x y (transport i A i) x y) +PathP≃Path A x y = isoToEquiv (PathPIsoPath A x y) + +PathP≡compPath : {A : Type } {x y z : A} (p : x y) (q : y z) (r : x z) + (PathP i x q i) p r) (p q r) +PathP≡compPath p q r k = PathP i p i0 q (i k)) j compPath-filler p q k j) r + +-- a quick corollary for 3-constant functions +3-ConstantCompChar : {A : Type } {B : Type ℓ'} (f : A B) (link : 2-Constant f) + (∀ x y z link x y link y z link x z) + 3-Constant f +3-Constant.link (3-ConstantCompChar f link coh₂) = link +3-Constant.coh₁ (3-ConstantCompChar f link coh₂) _ _ _ = + transport⁻ (PathP≡compPath _ _ _) (coh₂ _ _ _) + +PathP≡doubleCompPathˡ : {A : Type } {w x y z : A} (p : w y) (q : w x) (r : y z) (s : x z) + (PathP i p i s i) q r) (p ⁻¹ ∙∙ q ∙∙ s r) +PathP≡doubleCompPathˡ p q r s k = PathP i p (i k) s (i k)) + j doubleCompPath-filler (p ⁻¹) q s k j) r + +PathP≡doubleCompPathʳ : {A : Type } {w x y z : A} (p : w y) (q : w x) (r : y z) (s : x z) + (PathP i p i s i) q r) (q p ∙∙ r ∙∙ s ⁻¹) +PathP≡doubleCompPathʳ p q r s k = PathP i p (i (~ k)) s (i (~ k))) + q j doubleCompPath-filler p r (s ⁻¹) k j) + +compPathl-cancel : {} {A : Type } {x y z : A} (p : x y) (q : x z) p (sym p q) q +compPathl-cancel p q = p (sym p q) ≡⟨ assoc p (sym p) q + (p sym p) q ≡⟨ cong (_∙ q) (rCancel p) + refl q ≡⟨ sym (lUnit q) + q + +compPathr-cancel : {} {A : Type } {x y z : A} (p : z y) (q : x y) (q sym p) p q +compPathr-cancel {x = x} p q i j = + hcomp-equivFiller (doubleComp-faces _ x) (sym p) j) (inS (q j)) (~ i) + +compPathl-isEquiv : {x y z : A} (p : x y) isEquiv (q : y z) p q) +compPathl-isEquiv p = isoToIsEquiv (iso (p ∙_) (sym p ∙_) (compPathl-cancel p) (compPathl-cancel (sym p))) + +compPathlEquiv : {x y z : A} (p : x y) (y z) (x z) +compPathlEquiv p = (p ∙_) , compPathl-isEquiv p + +compPathr-isEquiv : {x y z : A} (p : y z) isEquiv (q : x y) q p) +compPathr-isEquiv p = isoToIsEquiv (iso (_∙ p) (_∙ sym p) (compPathr-cancel p) (compPathr-cancel (sym p))) + +compPathrEquiv : {x y z : A} (p : y z) (x y) (x z) +compPathrEquiv p = (_∙ p) , compPathr-isEquiv p + +-- Variations of isProp→isSet for PathP +isProp→SquareP : {B : I I Type } ((i j : I) isProp (B i j)) + {a : B i0 i0} {b : B i0 i1} {c : B i1 i0} {d : B i1 i1} + (r : PathP j B j i0) a c) (s : PathP j B j i1) b d) + (t : PathP j B i0 j) a b) (u : PathP j B i1 j) c d) + SquareP B t u r s +isProp→SquareP {B = B} isPropB {a = a} r s t u i j = + hcomp { k (i = i0) isPropB i0 j (base i0 j) (t j) k + ; k (i = i1) isPropB i1 j (base i1 j) (u j) k + ; k (j = i0) isPropB i i0 (base i i0) (r i) k + ; k (j = i1) isPropB i i1 (base i i1) (s i) k + }) (base i j) where + base : (i j : I) B i j + base i j = transport k B (i k) (j k)) a + +isProp→isPropPathP : {} {B : I Type } ((i : I) isProp (B i)) + (b0 : B i0) (b1 : B i1) + isProp (PathP i B i) b0 b1) +isProp→isPropPathP {B = B} hB b0 b1 = isProp→SquareP _ hB) refl refl + +isProp→isContrPathP : {A : I Type } (∀ i isProp (A i)) + (x : A i0) (y : A i1) + isContr (PathP A x y) +isProp→isContrPathP h x y = isProp→PathP h x y , isProp→isPropPathP h x y _ + +-- Flipping a square along its diagonal + +flipSquare : {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ a₀₁} + {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ a₁₁} + {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} + Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ +flipSquare sq i j = sq j i + +module _ {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ a₀₁} {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ a₁₁} + {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} + where + + flipSquareEquiv : Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ + unquoteDef flipSquareEquiv = defStrictEquiv flipSquareEquiv flipSquare flipSquare + + flipSquarePath : Square a₀₋ a₁₋ a₋₀ a₋₁ Square a₋₀ a₋₁ a₀₋ a₁₋ + flipSquarePath = ua flipSquareEquiv + +module _ {a₀₀ a₁₁ : A} {a₋ : a₀₀ a₁₁} + {a₁₀ : A} {a₁₋ : a₁₀ a₁₁} {a₋₀ : a₀₀ a₁₀} where + + slideSquareFaces : (i j k : I) Partial (i ~ i j ~ j) A + slideSquareFaces i j k (i = i0) = a₋ (j ~ k) + slideSquareFaces i j k (i = i1) = a₁₋ j + slideSquareFaces i j k (j = i0) = a₋₀ i + slideSquareFaces i j k (j = i1) = a₋ (i ~ k) + + slideSquare : Square a₋ a₁₋ a₋₀ refl Square refl a₁₋ a₋₀ a₋ + slideSquare sq i j = hcomp (slideSquareFaces i j) (sq i j) + + slideSquareEquiv : (Square a₋ a₁₋ a₋₀ refl) (Square refl a₁₋ a₋₀ a₋) + slideSquareEquiv = isoToEquiv (iso slideSquare slideSquareInv fillerTo fillerFrom) where + slideSquareInv : Square refl a₁₋ a₋₀ a₋ Square a₋ a₁₋ a₋₀ refl + slideSquareInv sq i j = hcomp k slideSquareFaces i j (~ k)) (sq i j) + fillerTo : p slideSquare (slideSquareInv p) p + fillerTo p k i j = hcomp-equivFiller k slideSquareFaces i j (~ k)) (inS (p i j)) (~ k) + fillerFrom : p slideSquareInv (slideSquare p) p + fillerFrom p k i j = hcomp-equivFiller (slideSquareFaces i j) (inS (p i j)) (~ k) + +-- The type of fillers of a square is equivalent to the double composition identites +Square≃doubleComp : {a₀₀ a₀₁ a₁₀ a₁₁ : A} + (a₀₋ : a₀₀ a₀₁) (a₁₋ : a₁₀ a₁₁) + (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) + Square a₀₋ a₁₋ a₋₀ a₋₁ (a₋₀ ⁻¹ ∙∙ a₀₋ ∙∙ a₋₁ a₁₋) +Square≃doubleComp a₀₋ a₁₋ a₋₀ a₋₁ = pathToEquiv (PathP≡doubleCompPathˡ a₋₀ a₀₋ a₁₋ a₋₁) + +-- Flipping a square in Ω²A is the same as inverting it +sym≡flipSquare : {x : A} (P : Square (refl {x = x}) refl refl refl) + sym P flipSquare P +sym≡flipSquare {x = x} P = sym (main refl P) + where + B : (q : x x) I Type _ + B q i = PathP j x q (i j)) k q (i k)) refl + + main : (q : x x) (p : refl q) PathP i B q i) i j p j i) (sym p) + main q = J q p PathP i B q i) i j p j i) (sym p)) refl + +-- Inverting both interval arguments of a square in Ω²A is the same as doing nothing +sym-cong-sym≡id : {x : A} (P : Square (refl {x = x}) refl refl refl) + P λ i j P (~ i) (~ j) +sym-cong-sym≡id {x = x} P = sym (main refl P) + where + B : (q : x x) I Type _ + B q i = Path (x q i) j q (i ~ j)) λ j q (i j) + + main : (q : x x) (p : refl q) PathP i B q i) i j p (~ i) (~ j)) p + main q = J q p PathP i B q i) i j p (~ i) (~ j)) p) refl + +-- Applying cong sym is the same as flipping a square in Ω²A +flipSquare≡cong-sym : {} {A : Type } {x : A} (P : Square (refl {x = x}) refl refl refl) + flipSquare P λ i j P i (~ j) +flipSquare≡cong-sym P = sym (sym≡flipSquare P) sym (sym-cong-sym≡id (cong sym P)) + +-- Applying cong sym is the same as inverting a square in Ω²A +sym≡cong-sym : {} {A : Type } {x : A} (P : Square (refl {x = x}) refl refl refl) + sym P cong sym P +sym≡cong-sym P = sym-cong-sym≡id (sym P) + +-- sym induces an equivalence on path types +symIso : {a b : A} Iso (a b) (b a) +symIso = iso sym sym _ refl) λ _ refl + +-- Vertical composition of squares (along their first dimension) +-- See Cubical.Foundations.Prelude for horizontal composition + +module _ { : Level} {A : Type } + {a₀₀ a₀₁ : A} {a₀₋ : a₀₀ a₀₁} + {a₁₀ a₁₁ : A} {a₁₋ : a₁₀ a₁₁} + {a₂₀ a₂₁ : A} {a₂₋ : a₂₀ a₂₁} + {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} + {b₋₀ : a₁₀ a₂₀} {b₋₁ : a₁₁ a₂₁} + where + + -- "Pointwise" composition + _∙v_ : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) + Square a₀₋ a₂₋ (a₋₀ b₋₀) (a₋₁ b₋₁) + (p ∙v q) i j = ((λ i p i j) i q i j)) i + + -- "Direct" composition + _∙v'_ : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) + Square a₀₋ a₂₋ (a₋₀ b₋₀) (a₋₁ b₋₁) + (p ∙v' q) i = + comp k compPath-filler a₋₀ b₋₀ k i compPath-filler a₋₁ b₋₁ k i) + where k (i = i0) a₀₋ + k (i = i1) q k) + (p i) + + -- The two ways of composing squares are equal, because they are + -- correct "lids" for the same box + ∙v≡∙v' : (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square a₁₋ a₂₋ b₋₀ b₋₁) + p ∙v q p ∙v' q + ∙v≡∙v' p q l i = outS + (comp-unique {A = λ k compPath-filler a₋₀ b₋₀ k i compPath-filler a₋₁ b₋₁ k i} + where k (i = i0) a₀₋ + k (i = i1) q k) + (inS (p i)) + k inS λ j compPath-filler i p i j) i q i j) k i)) + (~ l) + +-- Inspect + +module _ {A : Type } {B : A -> Type ℓ'} where + + record Reveal_·_is_ (f : (x : A) B x) (x : A) (y : B x) : Type (ℓ-max ℓ') where + constructor [_]ᵢ + field path : f x y + + inspect : (f : (x : A) B x) (x : A) Reveal f · x is f x + inspect f x .Reveal_·_is_.path = refl + +-- J is an equivalence +Jequiv : {x : A} (P : y x y Type ℓ') P x refl (∀ {y} (p : x y) P y p) +Jequiv P = isoToEquiv isom + where + isom : Iso _ _ + Iso.fun isom = J P + Iso.inv isom f = f refl + Iso.rightInv isom f = + implicitFunExt λ {_} + funExt λ t + J _ t J P (f refl) t f t) (JRefl P (f refl)) t + Iso.leftInv isom = JRefl P + +-- Action of PathP on equivalences (without relying on univalence) + +congPathIso : { ℓ'} {A : I Type } {B : I Type ℓ'} + (e : i A i B i) {a₀ : A i0} {a₁ : A i1} + Iso (PathP A a₀ a₁) (PathP B (e i0 .fst a₀) (e i1 .fst a₁)) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.fun p i = e i .fst (p i) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.inv q i = + hcomp + j λ + { (i = i0) retEq (e i0) a₀ j + ; (i = i1) retEq (e i1) a₁ j + }) + (invEq (e i) (q i)) +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.rightInv q k i = + hcomp + j λ + { (i = i0) commSqIsEq (e i0 .snd) a₀ j k + ; (i = i1) commSqIsEq (e i1 .snd) a₁ j k + ; (k = i0) + e i .fst + (hfill + j λ + { (i = i0) retEq (e i0) a₀ j + ; (i = i1) retEq (e i1) a₁ j + }) + (inS (invEq (e i) (q i))) + j) + ; (k = i1) q i + }) + (secEq (e i) (q i) k) + where b = commSqIsEq +congPathIso {A = A} {B} e {a₀} {a₁} .Iso.leftInv p k i = + hcomp + j λ + { (i = i0) retEq (e i0) a₀ (j k) + ; (i = i1) retEq (e i1) a₁ (j k) + ; (k = i1) p i + }) + (retEq (e i) (p i) k) + +congPathEquiv : { ℓ'} {A : I Type } {B : I Type ℓ'} + (e : i A i B i) {a₀ : A i0} {a₁ : A i1} + PathP A a₀ a₁ PathP B (e i0 .fst a₀) (e i1 .fst a₁) +congPathEquiv e = isoToEquiv (congPathIso e) + +-- Characterizations of dependent paths in path types + +doubleCompPath-filler∙ : {a b c d : A} (p : a b) (q : b c) (r : c d) + PathP i p i r (~ i)) (p q r) q +doubleCompPath-filler∙ {A = A} {b = b} p q r j i = + hcomp k λ { (i = i0) p j + ; (i = i1) side j k + ; (j = i1) q (i k)}) + (p (j i)) + where + side : I I A + side i j = + hcomp k λ { (i = i1) q j + ; (j = i0) b + ; (j = i1) r (~ i k)}) + (q j) + +PathP→compPathL : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} + PathP i p i q i) r s + sym p r q s +PathP→compPathL {p = p} {q = q} {r = r} {s = s} P j i = + hcomp k λ { (i = i0) p (j k) + ; (i = i1) q (j k) + ; (j = i0) doubleCompPath-filler∙ (sym p) r q (~ k) i + ; (j = i1) s i }) + (P j i) + +PathP→compPathR : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} + PathP i p i q i) r s + r p s sym q +PathP→compPathR {p = p} {q = q} {r = r} {s = s} P j i = + hcomp k λ { (i = i0) p (j (~ k)) + ; (i = i1) q (j (~ k)) + ; (j = i0) r i + ; (j = i1) doubleCompPath-filler∙ p s (sym q) (~ k) i}) + (P j i) + + +-- Other direction + +compPathL→PathP : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} + sym p r q s + PathP i p i q i) r s +compPathL→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp k λ { (i = i0) p (~ k j) + ; (i = i1) q (~ k j) + ; (j = i0) doubleCompPath-filler∙ (sym p) r q k i + ; (j = i1) s i}) + (P j i) + +compPathR→PathP : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} + r p s sym q + PathP i p i q i) r s +compPathR→PathP {p = p} {q = q} {r = r} {s = s} P j i = + hcomp k λ { (i = i0) p (k j) + ; (i = i1) q (k j) + ; (j = i0) r i + ; (j = i1) doubleCompPath-filler∙ p s (sym q) k i}) + (P j i) + +compPathR→PathP∙∙ : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} + r p ∙∙ s ∙∙ sym q + PathP i p i q i) r s +compPathR→PathP∙∙ {p = p} {q = q} {r = r} {s = s} P j i = + hcomp k λ { (i = i0) p (k j) + ; (i = i1) q (k j) + ; (j = i0) r i + ; (j = i1) doubleCompPath-filler p s (sym q) (~ k) i}) + (P j i) + +compPath→Square-faces : {a b c d : A} (p : a c) (q : b d) (r : a b) (s : c d) + (i j k : I) Partial (i ~ i j ~ j) A +compPath→Square-faces p q r s i j k = λ where + (i = i0) r (j k) + (i = i1) s (j ~ k) + (j = i0) compPath-filler p s (~ k) i + (j = i1) compPath-filler' r q (~ k) i + +compPath→Square : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} + p s r q Square r s p q +compPath→Square {p = p} {q = q} {r = r} {s = s} P i j = + hcomp (compPath→Square-faces p q r s i j) (P j i) + +Square→compPath : {a b c d : A} {p : a c} {q : b d} {r : a b} {s : c d} + Square r s p q p s r q +Square→compPath {p = p} {q = q} {r = r} {s = s} sq i j = + hcomp k compPath→Square-faces p q r s j i (~ k)) (sq j i) + +Square→compPathΩ² : {a : A} (sq : Square _ a) refl refl refl) + Square→compPath sq cong (_∙ refl) (flipSquare sq) +Square→compPathΩ² {a = a} sq k i j = + hcomp r λ {(i = i0) rUnit _ a) r j + ; (i = i1) rUnit _ a) r j + ; (j = i0) a + ; (j = i1) a + ; (k = i1) cong x rUnit x r) (flipSquare sq) i j}) + (sq j i) diff --git a/src/docs/Cubical.Foundations.Pointed.Base.html b/src/docs/Cubical.Foundations.Pointed.Base.html new file mode 100644 index 0000000..73d26ac --- /dev/null +++ b/src/docs/Cubical.Foundations.Pointed.Base.html @@ -0,0 +1,151 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Structure using (typ) public +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Function + +private + variable + ℓ' : Level + +Pointed : ( : Level) Type (ℓ-suc ) +Pointed = TypeWithStr x x) + +pt : {} (A∙ : Pointed ) typ A∙ +pt = str + +Pointed₀ = Pointed ℓ-zero + +{- Pointed functions -} +_→∙_ : (A : Pointed ) (B : Pointed ℓ') Type (ℓ-max ℓ') +(A , a) →∙ (B , b) = Σ[ f (A B) ] f a b + +_→∙_∙ : (A : Pointed ) (B : Pointed ℓ') Pointed (ℓ-max ℓ') +(A →∙ B ) .fst = A →∙ B +(A →∙ B ) .snd .fst x = pt B +(A →∙ B ) .snd .snd = refl + +idfun∙ : (A : Pointed ) A →∙ A +idfun∙ A .fst x = x +idfun∙ A .snd = refl + +infix 3 _≃∙_ +{-Pointed equivalences -} +_≃∙_ : (A : Pointed ) (B : Pointed ℓ') Type (ℓ-max ℓ') +A ≃∙ B = Σ[ e fst A fst B ] fst e (pt A) pt B + +{- Underlying pointed map of an equivalence -} +≃∙map : {A : Pointed } {B : Pointed ℓ'} A ≃∙ B A →∙ B +fst (≃∙map e) = fst (fst e) +snd (≃∙map e) = snd e + +invEquiv∙ : {A : Pointed } {B : Pointed ℓ'} A ≃∙ B B ≃∙ A +fst (invEquiv∙ x) = invEquiv (fst x) +snd (invEquiv∙ {A = A} x) = + sym (cong (fst (invEquiv (fst x))) (snd x)) retEq (fst x) (pt A) + +compEquiv∙ : { ℓ' ℓ''} {A : Pointed } {B : Pointed ℓ'} {C : Pointed ℓ''} + A ≃∙ B B ≃∙ C A ≃∙ C +fst (compEquiv∙ e1 e2) = compEquiv (fst e1) (fst e2) +snd (compEquiv∙ e1 e2) = cong (fst (fst e2)) (snd e1) snd e2 + +Equiv∙J : {B : Pointed } (C : (A : Pointed ) A ≃∙ B Type ℓ') + C B (idEquiv (fst B) , refl) + {A : _} (e : A ≃∙ B) C A e +Equiv∙J {} {ℓ'} {B = B} C ind {A = A} = + uncurry λ e p help e (pt A) (pt B) p C ind + where + help : {A : Type } (e : A typ B) (a : A) (b : typ B) + (p : fst e a b) + (C : (A : Pointed ) A ≃∙ (fst B , b) Type ℓ') + C (fst B , b) (idEquiv (fst B) , refl) + C (A , a) (e , p) + help = EquivJ A e (a : A) (b : typ B) + (p : fst e a b) + (C : (A : Pointed ) A ≃∙ (fst B , b) Type ℓ') + C (fst B , b) (idEquiv (fst B) , refl) + C (A , a) (e , p)) + λ a b J b p + (C : (A : Pointed ) A ≃∙ (fst B , b) Type ℓ') + C (fst B , b) + (idEquiv (fst B) , refl) + C (typ B , a) (idEquiv (typ B) , p)) + λ _ p p + +ua∙ : {A B : Pointed } (e : fst A fst B) + fst e (snd A) snd B A B +fst (ua∙ e p i) = ua e i +snd (ua∙ {A = A} e p i) = ua-gluePath e p i + +{- J for pointed function types -} +→∙J : { ℓ' ℓ''} {A : Pointed } {B : Type ℓ'} + (P : (b₀ : B) (f : A →∙ (B , b₀)) Type ℓ'') + ((f : fst A B) P (f (pt A)) (f , refl)) + {b₀ : B} (f : A →∙ (B , b₀)) P b₀ f +→∙J {A = A} P ind = + uncurry λ f J b₀ y P b₀ (f , y)) (ind f) + +{- HIT allowing for pattern matching on pointed types -} +data Pointer {} (A : Pointed ) : Type where + pt₀ : Pointer A + ⌊_⌋ : typ A Pointer A + id : pt A pt₀ + +IsoPointedPointer : {A : Pointed } Iso (typ A) (Pointer A) +Iso.fun IsoPointedPointer = ⌊_⌋ +Iso.inv (IsoPointedPointer {A = A}) pt₀ = pt A +Iso.inv IsoPointedPointer x = x +Iso.inv (IsoPointedPointer {A = A}) (id i) = pt A +Iso.rightInv IsoPointedPointer pt₀ = id +Iso.rightInv IsoPointedPointer x = refl +Iso.rightInv IsoPointedPointer (id i) j = id (i j) +Iso.leftInv IsoPointedPointer x = refl + +Pointed≡Pointer : {A : Pointed } typ A Pointer A +Pointed≡Pointer = isoToPath IsoPointedPointer + +Pointer∙ : (A : Pointed ) Pointed +Pointer∙ A .fst = Pointer A +Pointer∙ A .snd = pt₀ + +Pointed≡∙Pointer : {A : Pointed } A (Pointer A , pt₀) +Pointed≡∙Pointer {A = A} i = (Pointed≡Pointer {A = A} i) , helper i + where + helper : PathP i Pointed≡Pointer {A = A} i) (pt A) pt₀ + helper = ua-gluePath (isoToEquiv (IsoPointedPointer {A = A})) id + +pointerFun : {ℓ'} {A : Pointed } {B : Pointed ℓ'} (f : A →∙ B) + Pointer A Pointer B +pointerFun f pt₀ = pt₀ +pointerFun f x = fst f x +pointerFun f (id i) = (cong ⌊_⌋ (snd f) id) i + +pointerFun∙ : {ℓ'} {A : Pointed } {B : Pointed ℓ'} (f : A →∙ B) + Pointer∙ A →∙ Pointer∙ B +pointerFun∙ f .fst = pointerFun f +pointerFun∙ f .snd = refl + + +-- pointed identity equivalence +idEquiv∙ : (A : Pointed ) (A ≃∙ A) +idEquiv∙ A = idEquiv (typ A) , refl + +{- + Equational reasoning for pointed equivalences +-} + +infix 3 _∎≃∙ +infixr 2 _≃∙⟨_⟩_ + +_∎≃∙ : (A : Pointed ) A ≃∙ A +A ∎≃∙ = idEquiv∙ A + +_≃∙⟨_⟩_ : {B C : Pointed } (A : Pointed ) A ≃∙ B B ≃∙ C A ≃∙ C +A ≃∙⟨ f g = compEquiv∙ f g diff --git a/src/docs/Cubical.Foundations.Pointed.FunExt.html b/src/docs/Cubical.Foundations.Pointed.FunExt.html new file mode 100644 index 0000000..bb0dce2 --- /dev/null +++ b/src/docs/Cubical.Foundations.Pointed.FunExt.html @@ -0,0 +1,48 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.FunExt where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv + +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Foundations.Pointed.Homotopy + +private + variable + ℓ' : Level + +module _ {A : Pointed } {B : typ A Type ℓ'} {ptB : B (pt A)} where + + -- pointed function extensionality + funExt∙P : {f g : Π∙ A B ptB} f ∙∼P g f g + funExt∙P (h , h∙) i .fst x = h x i + funExt∙P (h , h∙) i .snd = h∙ i + + -- inverse of pointed function extensionality + funExt∙P⁻ : {f g : Π∙ A B ptB} f g f ∙∼P g + funExt∙P⁻ p .fst a i = p i .fst a + funExt∙P⁻ p .snd i = p i .snd + + -- function extensionality is an isomorphism, PathP version + funExt∙PIso : (f g : Π∙ A B ptB) Iso (f ∙∼P g) (f g) + Iso.fun (funExt∙PIso f g) = funExt∙P {f = f} {g = g} + Iso.inv (funExt∙PIso f g) = funExt∙P⁻ {f = f} {g = g} + Iso.rightInv (funExt∙PIso f g) p i j = p j + Iso.leftInv (funExt∙PIso f g) h _ = h + + -- transformed to equivalence + funExt∙P≃ : (f g : Π∙ A B ptB) (f ∙∼P g) (f g) + funExt∙P≃ f g = isoToEquiv (funExt∙PIso f g) + + -- funExt∙≃ using the other kind of pointed homotopy + funExt∙≃ : (f g : Π∙ A B ptB) (f ∙∼ g) (f g) + funExt∙≃ f g = compEquiv (∙∼≃∙∼P f g) (funExt∙P≃ f g) + + -- standard pointed function extensionality and its inverse + funExt∙ : {f g : Π∙ A B ptB} f ∙∼ g f g + funExt∙ {f = f} {g = g} = equivFun (funExt∙≃ f g) + + funExt∙⁻ : {f g : Π∙ A B ptB} f g f ∙∼ g + funExt∙⁻ {f = f} {g = g} = equivFun (invEquiv (funExt∙≃ f g)) diff --git a/src/docs/Cubical.Foundations.Pointed.Homogeneous.html b/src/docs/Cubical.Foundations.Pointed.Homogeneous.html new file mode 100644 index 0000000..97967a9 --- /dev/null +++ b/src/docs/Cubical.Foundations.Pointed.Homogeneous.html @@ -0,0 +1,216 @@ +{- + +Definition of a homogeneous pointed type, and proofs that pi, product, path, and discrete types are homogeneous + +Portions of this file adapted from Nicolai Kraus' code here: + https://bitbucket.org/nicolaikraus/agda/src/e30d70c72c6af8e62b72eefabcc57623dd921f04/trunc-inverse.lagda + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Homogeneous where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Data.Sigma +open import Cubical.Data.Empty as +open import Cubical.Relation.Nullary + +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Structures.Pointed + +{- + We might say that a type is homogeneous if its automorphism group acts transitively; + this could be phrased with a propositional truncation. + Here we demand something much stronger, namely that we are given automorphisms + that carry the base point to any given point y. + If in addition we require this automorphism to be the identity for the base point, + then we recover the notion of a left-invertible H-space, and indeed, + any homogeneous type in our sense gives rise to such, as shown in: + + Cubical.Homotopy.HSpace +-} +isHomogeneous : {} Pointed Type (ℓ-suc ) +isHomogeneous {} (A , x) = y Path (Pointed ) (A , x) (A , y) + +-- Pointed functions into a homogeneous type are equal as soon as they are equal +-- as unpointed functions +→∙Homogeneous≡ : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) f∙ .fst g∙ .fst f∙ g∙ +→∙Homogeneous≡ {A∙ = A∙@(_ , a₀)} {B∙@(B , _)} {f∙@(_ , f₀)} {g∙@(_ , g₀)} h p = + subst Q∙ PathP i A∙ →∙ Q∙ i) f∙ g∙) (sym (flipSquare fix)) badPath + where + badPath : PathP i A∙ →∙ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) f∙ g∙ + badPath i .fst = p i + badPath i .snd j = doubleCompPath-filler (sym f₀) (funExt⁻ p a₀) g₀ j i + + fix : PathP i B∙ (B , (sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) refl refl + fix i = + hcomp + j λ + { (i = i0) lCancel (h (pt B∙)) j + ; (i = i1) lCancel (h (pt B∙)) j + }) + (sym (h (pt B∙)) h ((sym f₀ ∙∙ funExt⁻ p a₀ ∙∙ g₀) i)) + +→∙HomogeneousPathP : + { ℓ'} {A∙ A∙' : Pointed } {B∙ B∙' : Pointed ℓ'} + {f∙ : A∙ →∙ B∙} {g∙ : A∙' →∙ B∙'} (p : A∙ A∙') (q : B∙ B∙') + (h : isHomogeneous B∙') + PathP i fst (p i) fst (q i)) (f∙ .fst) (g∙ .fst) + PathP i p i →∙ q i) f∙ g∙ +→∙HomogeneousPathP p q h r = toPathP (→∙Homogeneous≡ h (fromPathP r)) + +→∙Homogeneous≡Path : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) (p q : f∙ g∙) cong fst p cong fst q p q +→∙Homogeneous≡Path {A∙ = A∙@(A , a₀)} {B∙@(B , b)} {f∙@(f , f₀)} {g∙@(g , g₀)} h p q r = + transport k + PathP i + PathP j (A , a₀) →∙ newPath-refl p q r i j (~ k)) + (f , f₀) (g , g₀)) p q) + (badPath p q r) + where + newPath : (p q : f∙ g∙) (r : cong fst p cong fst q) + Square (refl {x = b}) refl refl refl + newPath p q r i j = + hcomp k λ {(i = i0) cong snd p j k + ; (i = i1) cong snd q j k + ; (j = i0) f₀ k + ; (j = i1) g₀ k}) + (r i j a₀) + + newPath-refl : (p q : f∙ g∙) (r : cong fst p cong fst q) + PathP i (PathP j B∙ (B , newPath p q r i j))) refl refl) refl refl + newPath-refl p q r i j k = + hcomp w λ { (i = i0) lCancel (h b) w k + ; (i = i1) lCancel (h b) w k + ; (j = i0) lCancel (h b) w k + ; (j = i1) lCancel (h b) w k + ; (k = i0) lCancel (h b) w k + ; (k = i1) B , newPath p q r i j}) + ((sym (h b) h (newPath p q r i j)) k) + + badPath : (p q : f∙ g∙) (r : cong fst p cong fst q) + PathP i + PathP j A∙ →∙ (B , newPath p q r i j)) + (f , f₀) (g , g₀)) + p q + fst (badPath p q r i j) = r i j + snd (badPath p q s i j) k = + hcomp r λ { (i = i0) snd (p j) (r k) + ; (i = i1) snd (q j) (r k) + ; (j = i0) f₀ (k r) + ; (j = i1) g₀ (k r) + ; (k = i0) s i j a₀}) + (s i j a₀) + +→∙HomogeneousSquare : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} {f∙ g∙ h∙ l∙ : A∙ →∙ B∙} + (h : isHomogeneous B∙) (s : f∙ h∙) (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) + Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) + Square p q s t +→∙HomogeneousSquare {f∙ = f∙} {g∙ = g∙} {h∙ = h∙} {l∙ = l∙} h = + J h∙ s (t : g∙ l∙) (p : f∙ g∙) (q : h∙ l∙) + Square (cong fst p) (cong fst q) (cong fst s) (cong fst t) + Square p q s t) + (J l∙ t (p : f∙ g∙) (q : f∙ l∙) + Square (cong fst p) (cong fst q) refl (cong fst t) + Square p q refl t) + (→∙Homogeneous≡Path {f∙ = f∙} {g∙ = g∙} h)) + +isHomogeneousPi : { ℓ'} {A : Type } {B∙ : A Pointed ℓ'} + (∀ a isHomogeneous (B∙ a)) isHomogeneous (Πᵘ∙ A B∙) +isHomogeneousPi h f i .fst = a typ (h a (f a) i) +isHomogeneousPi h f i .snd a = pt (h a (f a) i) + +isHomogeneousΠ∙ : { ℓ'} (A : Pointed ) (B : typ A Type ℓ') + (b₀ : B (pt A)) + ((a : typ A) (x : B a) isHomogeneous (B a , x)) + (f : Π∙ A B b₀) + isHomogeneous (Π∙ A B b₀ , f) +fst (isHomogeneousΠ∙ A B b₀ h f g i) = + Σ[ r ((a : typ A) fst ((h a (fst f a) (fst g a)) i)) ] + r (pt A) hcomp k λ {(i = i0) snd f k + ; (i = i1) snd g k}) + (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) +snd (isHomogeneousΠ∙ A B b₀ h f g i) = + a snd (h a (fst f a) (fst g a) i)) + , λ j hcomp k λ { (i = i0) snd f (k j) + ; (i = i1) snd g (k j) + ; (j = i0) snd (h (pt A) (fst f (pt A)) + (fst g (pt A)) i)}) + (snd (h (pt A) (fst f (pt A)) (fst g (pt A)) i)) + +isHomogeneous→∙ : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} + isHomogeneous B∙ isHomogeneous (A∙ →∙ B∙ ) +isHomogeneous→∙ {A∙ = A∙} {B∙} h f∙ = + ΣPathP + ( i Π∙ A∙ a T a i) (t₀ i)) + , PathPIsoPath _ _ _ .Iso.inv + (→∙Homogeneous≡ h + (PathPIsoPath i (a : typ A∙) T a i) _ pt B∙) _ .Iso.fun + i a pt (h (f∙ .fst a) i)))) + ) + where + T : a typ B∙ typ B∙ + T a i = typ (h (f∙ .fst a) i) + + t₀ : PathP i T (pt A∙) i) (pt B∙) (pt B∙) + t₀ = cong pt (h (f∙ .fst (pt A∙))) f∙ .snd + +isHomogeneousProd : { ℓ'} {A∙ : Pointed } {B∙ : Pointed ℓ'} + isHomogeneous A∙ isHomogeneous B∙ isHomogeneous (A∙ ×∙ B∙) +isHomogeneousProd hA hB (a , b) i .fst = typ (hA a i) × typ (hB b i) +isHomogeneousProd hA hB (a , b) i .snd .fst = pt (hA a i) +isHomogeneousProd hA hB (a , b) i .snd .snd = pt (hB b i) + +isHomogeneousPath : {} (A : Type ) {x y : A} (p : x y) isHomogeneous ((x y) , p) +isHomogeneousPath A {x} {y} p q + = pointed-sip ((x y) , p) ((x y) , q) (eqv , compPathr-cancel p q) + where eqv : (x y) (x y) + eqv = compPathlEquiv (q sym p) + +module HomogeneousDiscrete {} {A∙ : Pointed } (dA : Discrete (typ A∙)) (y : typ A∙) where + + -- switches pt A∙ with y + switch : typ A∙ typ A∙ + switch x with dA x (pt A∙) + ... | yes _ = y + ... | no _ with dA x y + ... | yes _ = pt A∙ + ... | no _ = x + + switch-ptA∙ : switch (pt A∙) y + switch-ptA∙ with dA (pt A∙) (pt A∙) + ... | yes _ = refl + ... | no ¬p = ⊥.rec (¬p refl) + + switch-idp : x switch (switch x) x + switch-idp x with dA x (pt A∙) + switch-idp x | yes p with dA y (pt A∙) + switch-idp x | yes p | yes q = q sym p + switch-idp x | yes p | no _ with dA y y + switch-idp x | yes p | no _ | yes _ = sym p + switch-idp x | yes p | no _ | no ¬p = ⊥.rec (¬p refl) + switch-idp x | no ¬p with dA x y + switch-idp x | no ¬p | yes p with dA y (pt A∙) + switch-idp x | no ¬p | yes p | yes q = ⊥.rec (¬p (p q)) + switch-idp x | no ¬p | yes p | no _ with dA (pt A∙) (pt A∙) + switch-idp x | no ¬p | yes p | no _ | yes _ = sym p + switch-idp x | no ¬p | yes p | no _ | no ¬q = ⊥.rec (¬q refl) + switch-idp x | no ¬p | no ¬q with dA x (pt A∙) + switch-idp x | no ¬p | no ¬q | yes p = ⊥.rec (¬p p) + switch-idp x | no ¬p | no ¬q | no _ with dA x y + switch-idp x | no ¬p | no ¬q | no _ | yes q = ⊥.rec (¬q q) + switch-idp x | no ¬p | no ¬q | no _ | no _ = refl + + switch-eqv : typ A∙ typ A∙ + switch-eqv = isoToEquiv (iso switch switch switch-idp switch-idp) + +isHomogeneousDiscrete : {} {A∙ : Pointed } (dA : Discrete (typ A∙)) isHomogeneous A∙ +isHomogeneousDiscrete {} {A∙} dA y + = pointed-sip (typ A∙ , pt A∙) (typ A∙ , y) (switch-eqv , switch-ptA∙) + where open HomogeneousDiscrete {} {A∙} dA y diff --git a/src/docs/Cubical.Foundations.Pointed.Homotopy.html b/src/docs/Cubical.Foundations.Pointed.Homotopy.html new file mode 100644 index 0000000..9618387 --- /dev/null +++ b/src/docs/Cubical.Foundations.Pointed.Homotopy.html @@ -0,0 +1,119 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Homotopy where + +{- + This module defines two kinds of pointed homotopies, + ∙∼ and ∙∼P, and proves their equivalence +-} + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Fiberwise +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Path +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Pointed.Properties +open import Cubical.Homotopy.Base +open import Cubical.Data.Sigma + +private + variable + ℓ' : Level + +module _ {A : Pointed } {B : typ A Type ℓ'} {ptB : B (pt A)} where + + = pt A + + -- pointed homotopy as pointed Π. This is just a Σ-type, see ∙∼Σ + _∙∼_ : (f g : Π∙ A B ptB) Type (ℓ-max ℓ') + (f₁ , f₂) ∙∼ (g₁ , g₂) = Π∙ A x f₁ x g₁ x) (f₂ g₂ ⁻¹) + + -- pointed homotopy with PathP. Also a Σ-type, see ∙∼PΣ + _∙∼P_ : (f g : Π∙ A B ptB) Type (ℓ-max ℓ') + (f₁ , f₂) ∙∼P (g₁ , g₂) = Σ[ h f₁ g₁ ] PathP i h i ptB) f₂ g₂ + + -- Proof that f ∙∼ g ≃ f ∙∼P g + -- using equivalence of the total map of φ + private + module _ (f g : Π∙ A B ptB) (H : f .fst g .fst) where + -- convenient notation + f₁ = fst f + f₂ = snd f + g₁ = fst g + g₂ = snd g + + -- P is the predicate on a homotopy H to be pointed of the ∙∼ kind + P : Type ℓ' + P = H f₂ g₂ ⁻¹ + + -- Q is the predicate on a homotopy H to be pointed of the ∙∼P kind + Q : Type ℓ' + Q = PathP i H i ptB) f₂ g₂ + + -- simplify the notation even more to see that P≡Q + -- is just a jingle of paths + p = H + r = f₂ + s = g₂ + P≡Q : P Q + P≡Q = p r s ⁻¹ + ≡⟨ isoToPath symIso + r s ⁻¹ p + ≡⟨ cong (r s ⁻¹ ≡_) (rUnit p ∙∙ cong (p ∙_) (sym (rCancel s)) ∙∙ assoc p s (s ⁻¹)) + r s ⁻¹ (p s) s ⁻¹ + ≡⟨ sym (ua (compr≡Equiv r (p s) (s ⁻¹))) + r p s + ≡⟨ ua (compl≡Equiv (p ⁻¹) r (p s)) + p ⁻¹ r p ⁻¹ (p s) + ≡⟨ cong (p ⁻¹ r ≡_ ) (assoc (p ⁻¹) p s ∙∙ (cong (_∙ s) (lCancel p)) ∙∙ sym (lUnit s)) + p ⁻¹ r s + ≡⟨ cong z p ⁻¹ z s) (rUnit r) + p ⁻¹ (r refl) s + ≡⟨ cong (_≡ s) (sym (doubleCompPath-elim' (p ⁻¹) r refl)) + p ⁻¹ ∙∙ r ∙∙ refl s + ≡⟨ sym (ua (Square≃doubleComp r s p refl)) + PathP i p i ptB) r s + + -- φ is a fiberwise transformation (H : f ∼ g) → P H → Q H + -- φ is even a fiberwise equivalence by P≡Q + φ : P Q + φ = transport P≡Q + + -- The total map corresponding to φ + totφ : (f g : Π∙ A B ptB) f ∙∼ g f ∙∼P g + totφ f g p .fst = p .fst + totφ f g p .snd = φ f g (p .fst) (p .snd) + + -- transformation of the homotopies using totφ + ∙∼→∙∼P : (f g : Π∙ A B ptB) (f ∙∼ g) (f ∙∼P g) + ∙∼→∙∼P f g = totφ f g + + -- Proof that ∙∼ and ∙∼P are equivalent using the fiberwise equivalence φ + ∙∼≃∙∼P : (f g : Π∙ A B ptB) (f ∙∼ g) (f ∙∼P g) + ∙∼≃∙∼P f g = Σ-cong-equiv-snd H pathToEquiv (P≡Q f g H)) + + -- inverse of ∙∼→∙∼P extracted from the equivalence + ∙∼P→∙∼ : {f g : Π∙ A B ptB} f ∙∼P g f ∙∼ g + ∙∼P→∙∼ {f = f} {g = g} = invEq (∙∼≃∙∼P f g) + + -- ∙∼≃∙∼P transformed to a path + ∙∼≡∙∼P : (f g : Π∙ A B ptB) (f ∙∼ g) (f ∙∼P g) + ∙∼≡∙∼P f g = ua (∙∼≃∙∼P f g) + + -- Verifies that the pointed homotopies actually correspond + -- to their Σ-type versions + _∙∼Σ_ : (f g : Π∙ A B ptB) Type (ℓ-max ℓ') + f ∙∼Σ g = Σ[ H f .fst g .fst ] (P f g H) + + _∙∼PΣ_ : (f g : Π∙ A B ptB) Type (ℓ-max ℓ') + f ∙∼PΣ g = Σ[ H f .fst g .fst ] (Q f g H) + + ∙∼≡∙∼Σ : (f g : Π∙ A B ptB) f ∙∼ g f ∙∼Σ g + ∙∼≡∙∼Σ f g = refl + + ∙∼P≡∙∼PΣ : (f g : Π∙ A B ptB) f ∙∼P g f ∙∼PΣ g + ∙∼P≡∙∼PΣ f g = refl diff --git a/src/docs/Cubical.Foundations.Pointed.Properties.html b/src/docs/Cubical.Foundations.Pointed.Properties.html new file mode 100644 index 0000000..67ae70a --- /dev/null +++ b/src/docs/Cubical.Foundations.Pointed.Properties.html @@ -0,0 +1,246 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Pointed.Base +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path + +open import Cubical.Data.Sigma + +private + variable + ℓ' ℓA ℓB ℓC ℓD : Level + +-- the default pointed Π-type: A is pointed, and B has a base point in the chosen fiber +Π∙ : (A : Pointed ) (B : typ A Type ℓ') (ptB : B (pt A)) Type (ℓ-max ℓ') +Π∙ A B ptB = Σ[ f ((a : typ A) B a) ] f (pt A) ptB + +-- the unpointed Π-type becomes a pointed type if the fibers are all pointed +Πᵘ∙ : (A : Type ) (B : A Pointed ℓ') Pointed (ℓ-max ℓ') +Πᵘ∙ A B .fst = a typ (B a) +Πᵘ∙ A B .snd a = pt (B a) + +-- if the base and all fibers are pointed, we have the pointed pointed Π-type +Πᵖ∙ : (A : Pointed ) (B : typ A Pointed ℓ') Pointed (ℓ-max ℓ') +Πᵖ∙ A B .fst = Π∙ A (typ B) (pt (B (pt A))) +Πᵖ∙ A B .snd .fst a = pt (B a) +Πᵖ∙ A B .snd .snd = refl + +-- the default pointed Σ-type is just the Σ-type, but as a pointed type +Σ∙ : (A : Pointed ) (B : typ A Type ℓ') (ptB : B (pt A)) Pointed (ℓ-max ℓ') +Σ∙ A B ptB .fst = Σ[ a typ A ] B a +Σ∙ A B ptB .snd .fst = pt A +Σ∙ A B ptB .snd .snd = ptB + +-- version if B is a family of pointed types +Σᵖ∙ : (A : Pointed ) (B : typ A Pointed ℓ') Pointed (ℓ-max ℓ') +Σᵖ∙ A B = Σ∙ A (typ B) (pt (B (pt A))) + +_×∙_ : (A∙ : Pointed ) (B∙ : Pointed ℓ') Pointed (ℓ-max ℓ') +(A∙ ×∙ B∙) .fst = (typ A∙) × (typ B∙) +(A∙ ×∙ B∙) .snd .fst = pt A∙ +(A∙ ×∙ B∙) .snd .snd = pt B∙ + +-- composition of pointed maps +_∘∙_ : {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} + (g : B →∙ C) (f : A →∙ B) (A →∙ C) +((g , g∙) ∘∙ (f , f∙)) .fst x = g (f x) +((g , g∙) ∘∙ (f , f∙)) .snd = (cong g f∙) g∙ + +-- post composition +post∘∙ : {ℓX ℓ'} (X : Pointed ℓX) {A : Pointed } {B : Pointed ℓ'} + (A →∙ B) ((X →∙ A ) →∙ (X →∙ B )) +post∘∙ X f .fst g = f ∘∙ g +post∘∙ X f .snd = + ΣPathP + ( (funExt λ _ f .snd) + , (sym (lUnit (f .snd)) λ i j f .snd (i j))) + +-- pointed identity +id∙ : (A : Pointed ℓA) (A →∙ A) +id∙ A .fst x = x +id∙ A .snd = refl + +-- constant pointed map +const∙ : (A : Pointed ℓA) (B : Pointed ℓB) (A →∙ B) +const∙ _ B .fst _ = B .snd +const∙ _ B .snd = refl + +-- left identity law for pointed maps +∘∙-idˡ : {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) f ∘∙ id∙ A f +∘∙-idˡ f = ΣPathP ( refl , (lUnit (f .snd)) ⁻¹ ) + +-- right identity law for pointed maps +∘∙-idʳ : {A : Pointed ℓA} {B : Pointed ℓB} (f : A →∙ B) id∙ B ∘∙ f f +∘∙-idʳ f = ΣPathP ( refl , (rUnit (f .snd)) ⁻¹ ) + +-- associativity for composition of pointed maps +∘∙-assoc : {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} {D : Pointed ℓD} + (h : C →∙ D) (g : B →∙ C) (f : A →∙ B) + (h ∘∙ g) ∘∙ f h ∘∙ (g ∘∙ f) +∘∙-assoc (h , h∙) (g , g∙) (f , f∙) = ΣPathP (refl , q) + where + q : (cong (h g) f∙) (cong h g∙ h∙) cong h (cong g f∙ g∙) h∙ + q = ( (cong (h g) f∙) (cong h g∙ h∙) + ≡⟨ refl + (cong h (cong g f∙)) (cong h g∙ h∙) + ≡⟨ assoc (cong h (cong g f∙)) (cong h g∙) h∙ + (cong h (cong g f∙) cong h g∙) h∙ + ≡⟨ cong p p h∙) ((cong-∙ h (cong g f∙) g∙) ⁻¹) + (cong h (cong g f∙ g∙) h∙) ) + +module _ { ℓ' : Level} {A : Pointed } {B : Pointed ℓ'} (f : A →∙ B) where + isInIm∙ : (x : typ B) Type (ℓ-max ℓ') + isInIm∙ x = Σ[ z typ A ] fst f z x + + isInKer∙ : (x : fst A) Type ℓ' + isInKer∙ x = fst f x snd B + +private module _ {ℓA ℓB ℓC : Level} (A : Pointed ℓA) (B : Pointed ℓB) (C : Pointed ℓC) (e : A ≃∙ B) where + toEq : (A →∙ C) (B →∙ C) + toEq = _∘∙ ≃∙map (invEquiv∙ e) + + fromEq : B →∙ C (A →∙ C) + fromEq = _∘∙ ≃∙map e + + toEq' : (C →∙ A) C →∙ B + toEq' = ≃∙map e ∘∙_ + + fromEq' : C →∙ B (C →∙ A) + fromEq' = ≃∙map (invEquiv∙ e) ∘∙_ + +pre∘∙equiv : {ℓA ℓB ℓC} {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} + (B ≃∙ C) Iso (A →∙ B) (A →∙ C) +Iso.fun (pre∘∙equiv {A = A} {B = B} {C = C} e) = toEq' B C A e +Iso.inv (pre∘∙equiv {A = A} {B = B} {C = C} e) = fromEq' B C A e +Iso.rightInv (pre∘∙equiv {A = A} {B = B} {C = C} e) = + J ptC p section (toEq' B (fst C , ptC) A (fst e , p)) + (fromEq' B (fst C , ptC) A (fst e , p))) + (uncurry f p ΣPathP (funExt x isHAEquiv.rinv (HAe .snd) (f x)) + , ((sym (rUnit _) + cong (cong (fst (fst e))) + λ i cong (invEq (fst e)) p + (lUnit (retEq (fst e) (pt B)) (~ i))) + cong-∙ (fst (fst e)) + (cong (invEq (fst e)) p) + (retEq (fst e) (pt B)) + refl + flipSquare (((λ _ isHAEquiv.rinv (HAe .snd) (f (pt A))) + refl) + lem _ _ _ _ (cong (isHAEquiv.rinv (HAe .snd)) p + sym (isHAEquiv.com (HAe .snd) (pt B)))))))) + (snd e) + where + HAe = equiv→HAEquiv (fst e) + lem : {} {A : Type } {x y z w : A} + (p : x y) (q : y z) (r : x w) (l : w z) + PathP i p i l i) r q + PathP i (p q) i l i) r refl + lem p q r l P i j = + hcomp k λ{ (i = i0) r j + ; (i = i1) q (j k) + ; (j = i1) l i}) + (P i j) +Iso.leftInv (pre∘∙equiv {A = A} {B = B} {C = C} e) = + J pt p retract (toEq' B (fst C , pt) A (fst e , p)) + (fromEq' B (fst C , pt) A (fst e , p))) + (uncurry f + J ptB p + fromEq' (fst B , ptB) + (fst C , fst (fst e) ptB) A (fst e , refl) + (toEq' (fst B , ptB) + (fst C , fst (fst e) ptB) A (fst e , refl) (f , p)) + (f , p)) + (ΣPathP + (funExt x retEq (fst e) (f x)) + , ((cong₂ _∙_ ((λ i cong (invEq (fst e)) (lUnit refl (~ i)))) + (sym (lUnit _) λ _ retEq (fst e) (f (pt A))) + sym (lUnit _)) + λ i j retEq (fst e) (f (pt A)) (i j)))))) + (snd e) + +post∘∙equiv : {ℓA ℓB ℓC} {A : Pointed ℓA} {B : Pointed ℓB} {C : Pointed ℓC} + (A ≃∙ B) Iso (A →∙ C) (B →∙ C) +Iso.fun (post∘∙equiv {A = A} {B = B} {C = C} e) = toEq A B C e +Iso.inv (post∘∙equiv {A = A} {B = B} {C = C} e) = fromEq A B C e +Iso.rightInv (post∘∙equiv {A = A}{B = B , ptB} {C = C} e) = + J pt p section (toEq A (B , pt) C (fst e , p)) + (fromEq A (B , pt) C (fst e , p))) + (uncurry f + J ptC p toEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) + (fromEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) + (f , p)) + (f , p)) + (ΣPathP (funExt x cong f (isHAEquiv.rinv (snd HAe) x)) + , (cong₂ _∙_ + i cong f (cong (fst (fst e)) (lUnit (retEq (fst e) (pt A)) (~ i)))) + (sym (rUnit refl)) + sym (rUnit _) + cong (cong f) (isHAEquiv.com (snd HAe) (pt A))) + λ i j f (isHAEquiv.rinv (snd HAe) (fst HAe (pt A)) (i j)))))) + (snd e) + where + HAe = equiv→HAEquiv (fst e) +Iso.leftInv (post∘∙equiv {A = A} {B = B , ptB} {C = C} e) = + J pt p retract (toEq A (B , pt) C (fst e , p)) + (fromEq A (B , pt) C (fst e , p))) + (uncurry f J ptC y + fromEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) + (toEq A (B , fst (fst e) (pt A)) (fst C , ptC) (fst e , refl) + (f , y)) + (f , y)) + (ΣPathP (funExt x cong f (retEq (fst e) x)) + , (sym (lUnit _) + sym (rUnit _) + cong (cong f) (sym (lUnit _)) + _ cong f (retEq (fst e) (pt A))) + λ i j f (retEq (fst e) (pt A) (i j))))))) + (snd e) + +flip→∙∙ : {A : Pointed } {B : Pointed ℓ'} {C : Pointed ℓA} + (A →∙ (B →∙ C )) B →∙ (A →∙ C ) +fst (fst (flip→∙∙ f) x) a = fst f a .fst x +snd (fst (flip→∙∙ f) x) i = snd f i .fst x +fst (snd (flip→∙∙ f) i) a = fst f a .snd i +snd (snd (flip→∙∙ f) i) j = snd f j .snd i + +flip→∙∙Iso : {A : Pointed } {B : Pointed ℓ'} {C : Pointed ℓA} + Iso (A →∙ (B →∙ C )) (B →∙ (A →∙ C )) +Iso.fun flip→∙∙Iso = flip→∙∙ +Iso.inv flip→∙∙Iso = flip→∙∙ +Iso.rightInv flip→∙∙Iso _ = refl +Iso.leftInv flip→∙∙Iso _ = refl + +≃∙→ret/sec∙ : {} {A B : Pointed } + (f : A ≃∙ B) ((≃∙map (invEquiv∙ f) ∘∙ ≃∙map f) idfun∙ A) + × (≃∙map f ∘∙ ≃∙map (invEquiv∙ f) idfun∙ B) +≃∙→ret/sec∙ {A = A} {B = B} = + Equiv∙J A f ((≃∙map (invEquiv∙ f) ∘∙ ≃∙map f) idfun∙ A) + × (≃∙map f ∘∙ ≃∙map (invEquiv∙ f) idfun∙ B)) + ((ΣPathP (refl , sym (lUnit _) sym (rUnit refl))) + , (ΣPathP (refl , sym (rUnit _) sym (rUnit refl)))) + +pointedSecIso : {ℓ''} {A : Pointed } {B : Pointed ℓ'} (Q : fst A Pointed ℓ'') + Iso ((a : fst A) Q a →∙ B) + (Σ[ F (Σ (fst A) (fst Q) fst B) ] + ((a : fst A) F (a , pt (Q a)) pt B)) +Iso.fun (pointedSecIso Q) F = x F (fst x) .fst (snd x)) , x F x .snd) +Iso.inv (pointedSecIso Q) F a = (fst F (a ,_)) , snd F a +Iso.rightInv (pointedSecIso Q) F = refl +Iso.leftInv (pointedSecIso Q) F = refl + +compPathrEquiv∙ : {A : Type } {a b c : A} {q : a b} (p : b c) + ((a b) , q) ≃∙ ((a c) , q p) +fst (compPathrEquiv∙ p) = compPathrEquiv p +snd (compPathrEquiv∙ p) = refl + +compPathlEquiv∙ : {A : Type } {a b c : A} {q : b c} (p : a b) + ((b c) , q) ≃∙ ((a c) , p q) +fst (compPathlEquiv∙ p) = compPathlEquiv p +snd (compPathlEquiv∙ p) = refl diff --git a/src/docs/Cubical.Foundations.Pointed.html b/src/docs/Cubical.Foundations.Pointed.html new file mode 100644 index 0000000..da34256 --- /dev/null +++ b/src/docs/Cubical.Foundations.Pointed.html @@ -0,0 +1,9 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Pointed where + +open import Cubical.Foundations.Pointed.Base public +open import Cubical.Foundations.Pointed.Properties public +open import Cubical.Foundations.Pointed.FunExt public +open import Cubical.Foundations.Pointed.Homotopy public + +open import Cubical.Foundations.Pointed.Homogeneous diff --git a/src/docs/Cubical.Foundations.Powerset.html b/src/docs/Cubical.Foundations.Powerset.html new file mode 100644 index 0000000..111a78d --- /dev/null +++ b/src/docs/Cubical.Foundations.Powerset.html @@ -0,0 +1,65 @@ +{- + +This file introduces the "powerset" of a type in the style of +Escardó's lecture notes: + +https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#propositionalextensionality + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Powerset where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence using (hPropExt) + +open import Cubical.Data.Sigma + +private + variable + : Level + X : Type + + : Type Type (ℓ-suc ) + X = X hProp _ + +isSetℙ : isSet ( X) +isSetℙ = isSetΠ λ x isSetHProp + +infix 5 _∈_ + +_∈_ : {X : Type } X X Type +x A = A x + +_⊆_ : {X : Type } X X Type +A B = x x A x B + +∈-isProp : (A : X) (x : X) isProp (x A) +∈-isProp A = snd A + +⊆-isProp : (A B : X) isProp (A B) +⊆-isProp A B = isPropΠ2 x _ ∈-isProp B x) + +⊆-refl : (A : X) A A +⊆-refl A x = idfun (x A) + +subst-∈ : (A : X) {x y : X} x y x A y A +subst-∈ A = subst (_∈ A) + +⊆-refl-consequence : (A B : X) A B (A B) × (B A) +⊆-refl-consequence A B p = subst (A ⊆_) p (⊆-refl A) + , subst (B ⊆_) (sym p) (⊆-refl B) + +⊆-extensionality : (A B : X) (A B) × (B A) A B +⊆-extensionality A B (φ , ψ) = + funExt x TypeOfHLevel≡ 1 (hPropExt (A x .snd) (B x .snd) (φ x) (ψ x))) + +⊆-extensionalityEquiv : (A B : X) (A B) × (B A) (A B) +⊆-extensionalityEquiv A B = isoToEquiv (iso (⊆-extensionality A B) + (⊆-refl-consequence A B) + _ isSetℙ A B _ _) + _ isPropΣ (⊆-isProp A B) _ ⊆-isProp B A) _ _)) diff --git a/src/docs/Cubical.Foundations.Prelude.html b/src/docs/Cubical.Foundations.Prelude.html new file mode 100644 index 0000000..305c68e --- /dev/null +++ b/src/docs/Cubical.Foundations.Prelude.html @@ -0,0 +1,616 @@ +{- + +This file proves a variety of basic results about paths: + +- refl, sym, cong and composition of paths. This is used to set up + equational reasoning. + +- Transport, subst and functional extensionality + +- J and its computation rule (up to a path) + +- Σ-types and contractibility of singletons + +- Converting PathP to and from a homogeneous path with transp + +- Direct definitions of lower h-levels + +- Export natural numbers + +- Export universe lifting + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Prelude where + +open import Cubical.Core.Primitives public + +infixr 30 _∙_ +infixr 30 _∙₂_ +infix 3 _∎ +infixr 2 step-≡ _≡⟨⟩_ +infixr 2.5 _≡⟨_⟩≡⟨_⟩_ +infixl 4 _≡$_ _≡$S_ + +-- Basic theory about paths. These proofs should typically be +-- inlined. This module also makes equational reasoning work with +-- (non-dependent) paths. + +private + variable + ℓ' ℓ'' : Level + A : Type + B : A Type + x y z w : A + +refl : x x +refl {x = x} _ = x +{-# INLINE refl #-} + +sym : x y y x +sym p i = p (~ i) +{-# INLINE sym #-} + +-- symP infers the type of its argument from the type of its output +symP : {A : I Type } {x : A i1} {y : A i0} + (p : PathP i A (~ i)) x y) PathP A y x +symP p j = p (~ j) + +-- symP infers the type of its output from the type of its argument +symP-fromGoal : {A : I Type } {x : A i0} {y : A i1} + (p : PathP A x y) PathP i A (~ i)) y x +symP-fromGoal p j = p (~ j) + +cong : (f : (a : A) B a) (p : x y) + PathP i B (p i)) (f x) (f y) +cong f p i = f (p i) +{-# INLINE cong #-} + +{- `S` stands for simply typed. Using `congS` instead of `cong` + can help Agda to solve metavariables that may otherwise remain unsolved. +-} +congS : {B : Type } (f : A B) (p : x y) f x f y +congS f p i = f (p i) +{-# INLINE congS #-} + +congP : {A : I Type } {B : (i : I) A i Type ℓ'} + (f : (i : I) (a : A i) B i a) {x : A i0} {y : A i1} + (p : PathP A x y) PathP i B i (p i)) (f i0 x) (f i1 y) +congP f p i = f i (p i) +{-# INLINE congP #-} + +cong₂ : {C : (a : A) (b : B a) Type } + (f : (a : A) (b : B a) C a b) + (p : x y) + {u : B x} {v : B y} (q : PathP i B (p i)) u v) + PathP i C (p i) (q i)) (f x u) (f y v) +cong₂ f p q i = f (p i) (q i) +{-# INLINE cong₂ #-} + +congP₂ : {A : I Type } {B : (i : I) A i Type ℓ'} + {C : (i : I) (a : A i) B i a Type ℓ''} + (f : (i : I) (a : A i) (b : B i a) C i a b) + {x : A i0} {y : A i1} {u : B i0 x} {v : B i1 y} + (p : PathP A x y) (q : PathP i B i (p i)) u v) + PathP i C i (p i) (q i)) (f i0 x u) (f i1 y v) +congP₂ f p q i = f i (p i) (q i) +{-# INLINE congP₂ #-} + +congL : {C : Type } (f : (a : A) C B a) (p : x y) + {z : C} PathP i B (p i)) (f x z) (f y z) +congL f p {z} i = f (p i) z +{-# INLINE congL #-} + +congR : {C : Type } (f : C (a : A) B a) (p : x y) + {z : C} PathP i B (p i)) (f z x) (f z y) +congR f p {z} i = f z (p i) +{-# INLINE congR #-} + +{- The most natural notion of homogenous path composition + in a cubical setting is double composition: + + x ∙ ∙ ∙ > w + ^ ^ + p⁻¹ | | r ^ + | | j | + y — — — > z ∙ — > + q i + + `p ∙∙ q ∙∙ r` gives the line at the top, + `doubleCompPath-filler p q r` gives the whole square +-} + +doubleComp-faces : {x y z w : A} (p : x y) (r : z w) + (i : I) (j : I) Partial (i ~ i) A +doubleComp-faces p r i j (i = i0) = p (~ j) +doubleComp-faces p r i j (i = i1) = r j + +_∙∙_∙∙_ : x y y z z w x w +(p ∙∙ q ∙∙ r) i = + hcomp (doubleComp-faces p r i) (q i) + +doubleCompPath-filler : (p : x y) (q : y z) (r : z w) + PathP j p (~ j) r j) q (p ∙∙ q ∙∙ r) +doubleCompPath-filler p q r j i = + hfill (doubleComp-faces p r i) (inS (q i)) j + +-- any two definitions of double composition are equal +compPath-unique : (p : x y) (q : y z) (r : z w) + (α β : Σ[ s x w ] PathP j p (~ j) r j) q s) + α β +compPath-unique p q r (α , α-filler) (β , β-filler) t + = i cb i1 i) , j i cb j i) + where cb : I I _ + cb j i = hfill j λ { (t = i0) α-filler j i + ; (t = i1) β-filler j i + ; (i = i0) p (~ j) + ; (i = i1) r j }) + (inS (q i)) j + +{- For single homogenous path composition, we take `p = refl`: + + x ∙ ∙ ∙ > z + ‖ ^ + ‖ | r ^ + ‖ | j | + x — — — > y ∙ — > + q i + + `q ∙ r` gives the line at the top, + `compPath-filler q r` gives the whole square +-} + +_∙_ : x y y z x z +p q = refl ∙∙ p ∙∙ q + +compPath-filler : (p : x y) (q : y z) PathP j x q j) p (p q) +compPath-filler p q = doubleCompPath-filler refl p q + +-- We could have also defined single composition by taking `r = refl`: + +_∙'_ : x y y z x z +p ∙' q = p ∙∙ q ∙∙ refl + +compPath'-filler : (p : x y) (q : y z) PathP j p (~ j) z) q (p ∙' q) +compPath'-filler p q = doubleCompPath-filler p q refl + +-- It's easy to show that `p ∙ q` also has such a filler: +compPath-filler' : (p : x y) (q : y z) PathP j p (~ j) z) q (p q) +compPath-filler' {z = z} p q j i = + hcomp k λ { (i = i0) p (~ j) + ; (i = i1) q k + ; (j = i0) q (i k) }) + (p (i ~ j)) +-- Note: We can omit a (j = i1) case here since when (j = i1), the whole expression is +-- definitionally equal to `p ∙ q`. (Notice that `p ∙ q` is also an hcomp.) Nevertheless, +-- we could have given `compPath-filler p q k i` as the (j = i1) case. + +-- From this, we can show that these two notions of composition are the same +compPath≡compPath' : (p : x y) (q : y z) p q p ∙' q +compPath≡compPath' p q j = + compPath-unique p q refl (p q , compPath-filler' p q) + (p ∙' q , compPath'-filler p q) j .fst + +-- Double composition agrees with iterated single composition +doubleCompPath≡compPath : {x y z w : A} + (p : x y) (q : y z) (r : z w) + p ∙∙ q ∙∙ r p q r +doubleCompPath≡compPath p q r i j = + hcomp k λ { (i = i1) compPath-filler' p (q r) k j + ; (j = i0) p (~ k) + ; (j = i1) r (i k)}) + (compPath-filler q r i j) + +-- Heterogeneous path composition and its filler: + +-- Composition in a family indexed by the interval +compPathP : {A : I Type } {x : A i0} {y : A i1} {B_i1 : Type } {B : (A i1) B_i1} {z : B i1} + PathP A x y PathP i B i) y z PathP j ((λ i A i) B) j) x z +compPathP {A = A} {x = x} {B = B} p q i = + comp j compPath-filler i A i) B j i) + j λ { (i = i0) x ; + (i = i1) q j }) + (p i) + +-- Composition in a family indexed by a type +compPathP' : {B : A Type ℓ'} {x' : B x} {y' : B y} {z' : B z} {p : x y} {q : y z} + (P : PathP i B (p i)) x' y') (Q : PathP i B (q i)) y' z') + PathP i B ((p q) i)) x' z' +compPathP' {B = B} {x' = x'} {p = p} {q = q} P Q i = + comp j B (compPath-filler p q j i)) + j λ { (i = i0) x' ; + (i = i1) Q j }) + (P i) + +compPathP-filler : {A : I Type } {x : A i0} {y : A i1} {B_i1 : Type } {B : A i1 B_i1} {z : B i1} + (p : PathP A x y) (q : PathP i B i) y z) + PathP j PathP k (compPath-filler i A i) B j k)) x (q j)) p (compPathP p q) +compPathP-filler {A = A} {x = x} {B = B} p q j i = + fill j compPath-filler i A i) B j i) + j λ { (i = i0) x ; + (i = i1) q j }) + (inS (p i)) j + +compPathP'-filler : {B : A Type ℓ'} {x' : B x} {y' : B y} {z' : B z} {p : x y} {q : y z} + (P : PathP i B (p i)) x' y') (Q : PathP i B (q i)) y' z') + PathP j PathP i B (compPath-filler p q j i)) x' (Q j)) P (compPathP' {B = B} P Q) +compPathP'-filler {B = B} {x' = x'} {p = p} {q = q} P Q j i = + fill j B (compPath-filler p q j i)) + j λ { (i = i0) x' ; + (i = i1) Q j }) + (inS (P i)) + j + +-- Syntax for chains of equational reasoning + +step-≡ : (x : A) y z x y x z +step-≡ _ p q = q p + +syntax step-≡ x y p = x ≡⟨ p y + +≡⟨⟩-syntax : (x : A) y z x y x z +≡⟨⟩-syntax = step-≡ + +infixr 2 ≡⟨⟩-syntax +syntax ≡⟨⟩-syntax x y i B) = x ≡[ i ]⟨ B y + +_≡⟨⟩_ : (x : A) x y x y +_ ≡⟨⟩ x≡y = x≡y + +≡⟨⟩⟨⟩-syntax : (x y : A) x y y z z w x w +≡⟨⟩⟨⟩-syntax x y p q r = p ∙∙ q ∙∙ r +infixr 3 ≡⟨⟩⟨⟩-syntax +syntax ≡⟨⟩⟨⟩-syntax x y B C = x ≡⟨ B ⟩≡ y ≡⟨ C ⟩≡ + +_≡⟨_⟩≡⟨_⟩_ : (x : A) x y y z z w x w +_ ≡⟨ x≡y ⟩≡⟨ y≡z z≡w = x≡y ∙∙ y≡z ∙∙ z≡w + +_∎ : (x : A) x x +_ = refl + +-- Transport and subst + +-- transport is a special case of transp +transport : {A B : Type } A B A B +transport p a = transp i p i) i0 a + +-- Transporting in a constant family is the identity function (up to a +-- path). If we would have regularity this would be definitional. +transportRefl : (x : A) transport refl x x +transportRefl {A = A} x i = transp _ A) i x + +transport-filler : {} {A B : Type } (p : A B) (x : A) + PathP i p i) x (transport p x) +transport-filler p x i = transp j p (i j)) (~ i) x + +-- We want B to be explicit in subst +subst : (B : A Type ℓ') (p : x y) B x B y +subst B p pa = transport i B (p i)) pa + +subst2 : {ℓ' ℓ''} {B : Type ℓ'} {z w : B} (C : A B Type ℓ'') + (p : x y) (q : z w) C x z C y w +subst2 B p q b = transport i B (p i) (q i)) b + +substRefl : {B : A Type } {x} (px : B x) subst B refl px px +substRefl px = transportRefl px + +subst-filler : (B : A Type ℓ') (p : x y) (b : B x) + PathP i B (p i)) b (subst B p b) +subst-filler B p = transport-filler (cong B p) + +subst2-filler : {B : Type ℓ'} {z w : B} (C : A B Type ℓ'') + (p : x y) (q : z w) (c : C x z) + PathP i C (p i) (q i)) c (subst2 C p q c) +subst2-filler C p q = transport-filler (cong₂ C p q) + +-- Function extensionality + +funExt : {B : A I Type ℓ'} + {f : (x : A) B x i0} {g : (x : A) B x i1} + ((x : A) PathP (B x) (f x) (g x)) + PathP i (x : A) B x i) f g +funExt p i x = p x i + +implicitFunExt : {B : A I Type ℓ'} + {f : {x : A} B x i0} {g : {x : A} B x i1} + ({x : A} PathP (B x) (f {x}) (g {x})) + PathP i {x : A} B x i) f g +implicitFunExt p i {x} = p {x} i + +-- the inverse to funExt (see Functions.FunExtEquiv), converting paths +-- between functions to homotopies; `funExt⁻` is called `happly` and +-- defined by path induction in the HoTT book (see function 2.9.2 in +-- section 2.9) +funExt⁻ : {B : A I Type ℓ'} + {f : (x : A) B x i0} {g : (x : A) B x i1} + PathP i (x : A) B x i) f g + ((x : A) PathP (B x) (f x) (g x)) +funExt⁻ eq x i = eq i x + +congP₂$ : {A : I Type } {B : i A i Type ℓ'} + {f : x B i0 x} {g : y B i1 y} + (p : PathP i x B i x) f g) + {x y} (p : PathP A x y) PathP i B i (p i)) (f x) (g y) +congP₂$ eq x i = eq i (x i) + +implicitFunExt⁻ : {B : A I Type ℓ'} + {f : {x : A} B x i0} {g : {x : A} B x i1} + PathP i {x : A} B x i) f g + ({x : A} PathP (B x) (f {x}) (g {x})) +implicitFunExt⁻ eq {x} i = eq i {x} + +_≡$_ = funExt⁻ + +{- `S` stands for simply typed. Using `funExtS⁻` instead of `funExt⁻` + can help Agda to solve metavariables that may otherwise remain unsolved. +-} +funExtS⁻ : {B : I Type ℓ'} + {f : (x : A) B i0} {g : (x : A) B i1} + PathP i (x : A) B i) f g + ((x : A) PathP i B i) (f x) (g x)) +funExtS⁻ eq x i = eq i x + +_≡$S_ = funExtS⁻ + +-- J for paths and its computation rule + +module _ (P : y x y Type ℓ') (d : P x refl) where + + J : (p : x y) P y p + J p = transport i P (p i) j p (i j))) d + + JRefl : J refl d + JRefl = transportRefl d + + J-∙ : (p : x y) (q : y z) + J (p q) transport i P (q i) j compPath-filler p q i j)) (J p) + J-∙ p q k = + transp + i P (q (i ~ k)) + j compPath-filler p q (i ~ k) j)) (~ k) + (J j compPath-filler p q (~ k) j)) + +-- Multi-variable versions of J + +module _ {b : B x} + (P : (y : A) (p : x y) (z : B y) (q : PathP i B (p i)) b z) Type ℓ'') + (d : P _ refl _ refl) where + + JDep : {y : A} (p : x y) {z : B y} (q : PathP i B (p i)) b z) P _ p _ q + JDep _ q = transport i P _ _ _ j q (i j))) d + + JDepRefl : JDep refl refl d + JDepRefl = transportRefl d + +module _ {x : A} + {P : (y : A) x y Type ℓ'} {d : (y : A) (p : x y) P y p} + (Q : (y : A) (p : x y) (z : P y p) d y p z Type ℓ'') + (r : Q _ refl _ refl) where + + private + ΠQ : (y : A) x y _ + ΠQ y p = z q Q y p z q + + J2 : {y : A} (p : x y) {z : P y p} (q : d y p z) Q _ p _ q + J2 p = J ΠQ _ J (Q x refl) r) p _ + + J2Refl : J2 refl refl r + J2Refl = i JRefl ΠQ _ J (Q x refl) r) i _ refl) JRefl (Q x refl) _ + +-- A prefix operator version of J that is more suitable to be nested + +module _ {P : y x y Type ℓ'} (d : P x refl) where + + J>_ : y (p : x y) P y p + J>_ _ p = transport i P (p i) j p (i j))) d + + infix 10 J>_ + +-- Converting to and from a PathP + +module _ {A : I Type } {x : A i0} {y : A i1} where + toPathP : transport i A i) x y PathP A x y + toPathP p i = hcomp j λ { (i = i0) x + ; (i = i1) p j }) + (transp j A (i j)) (~ i) x) + + fromPathP : PathP A x y transport i A i) x y + fromPathP p i = transp j A (i j)) i (p i) + +-- Whiskering a dependent path by a path +-- Double whiskering +_◁_▷_ : {} {A : I Type } {a₀ a₀' : A i0} {a₁ a₁' : A i1} + a₀ a₀' PathP A a₀' a₁ a₁ a₁' + PathP A a₀ a₁' +(p P q) i = + hcomp j λ {(i = i0) p (~ j) ; (i = i1) q j}) (P i) + +doubleWhiskFiller : + {} {A : I Type } {a₀ a₀' : A i0} {a₁ a₁' : A i1} + (p : a₀ a₀') (pq : PathP A a₀' a₁) (q : a₁ a₁') + PathP i PathP A (p (~ i)) (q i)) + pq + (p pq q) +doubleWhiskFiller p pq q k i = + hfill j λ {(i = i0) p (~ j) ; (i = i1) q j}) + (inS (pq i)) + k + +_◁_ : {} {A : I Type } {a₀ a₀' : A i0} {a₁ : A i1} + a₀ a₀' PathP A a₀' a₁ PathP A a₀ a₁ +(p q) = p q refl + +_▷_ : {} {A : I Type } {a₀ : A i0} {a₁ a₁' : A i1} + PathP A a₀ a₁ a₁ a₁' PathP A a₀ a₁' +p q = refl p q + +-- Direct definitions of lower h-levels + +isContr : Type Type +isContr A = Σ[ x A ] (∀ y x y) + +isProp : Type Type +isProp A = (x y : A) x y + +isSet : Type Type +isSet A = (x y : A) isProp (x y) + +isGroupoid : Type Type +isGroupoid A = a b isSet (Path A a b) + +is2Groupoid : Type Type +is2Groupoid A = a b isGroupoid (Path A a b) + +-- Contractibility of singletons + +singlP : (A : I Type ) (a : A i0) Type _ +singlP A a = Σ[ x A i1 ] PathP A a x + +singl : (a : A) Type _ +singl {A = A} a = singlP _ A) a + +isContrSingl : (a : A) isContr (singl a) +isContrSingl a .fst = (a , refl) +isContrSingl a .snd p i .fst = p .snd i +isContrSingl a .snd p i .snd j = p .snd (i j) + +isContrSinglP : (A : I Type ) (a : A i0) isContr (singlP A a) +isContrSinglP A a .fst = _ , transport-filler i A i) a +isContrSinglP A a .snd (x , p) i = + _ , λ j fill A j λ {(i = i0) transport-filler i A i) a j; (i = i1) p j}) (inS a) j + +-- Higher cube types + +SquareP : + (A : I I Type ) + {a₀₀ : A i0 i0} {a₀₁ : A i0 i1} (a₀₋ : PathP j A i0 j) a₀₀ a₀₁) + {a₁₀ : A i1 i0} {a₁₁ : A i1 i1} (a₁₋ : PathP j A i1 j) a₁₀ a₁₁) + (a₋₀ : PathP i A i i0) a₀₀ a₁₀) (a₋₁ : PathP i A i i1) a₀₁ a₁₁) + Type +SquareP A a₀₋ a₁₋ a₋₀ a₋₁ = PathP i PathP j A i j) (a₋₀ i) (a₋₁ i)) a₀₋ a₁₋ + +Square : + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) + (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) + Type _ +Square a₀₋ a₁₋ a₋₀ a₋₁ = PathP i a₋₀ i a₋₁ i) a₀₋ a₁₋ + +Cube : + {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ a₀₁₁} + {a₀₋₀ : a₀₀₀ a₀₁₀} {a₀₋₁ : a₀₀₁ a₀₁₁} + (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) + {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ a₁₀₁} + {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ a₁₁₁} + {a₁₋₀ : a₁₀₀ a₁₁₀} {a₁₋₁ : a₁₀₁ a₁₁₁} + (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) + {a₋₀₀ : a₀₀₀ a₁₀₀} {a₋₀₁ : a₀₀₁ a₁₀₁} + (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) + {a₋₁₀ : a₀₁₀ a₁₁₀} {a₋₁₁ : a₀₁₁ a₁₁₁} + (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) + (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) + (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) + Type _ +Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ = + PathP i Square (a₋₀₋ i) (a₋₁₋ i) (a₋₋₀ i) (a₋₋₁ i)) a₀₋₋ a₁₋₋ + +-- See HLevels.agda for CubeP + +-- Horizontal composition of squares (along their second dimension) +-- See Cubical.Foundations.Path for vertical composition + +_∙₂_ : + {a₀₀ a₀₁ a₀₂ : A} {a₀₋ : a₀₀ a₀₁} {b₀₋ : a₀₁ a₀₂} + {a₁₀ a₁₁ a₁₂ : A} {a₁₋ : a₁₀ a₁₁} {b₁₋ : a₁₁ a₁₂} + {a₋₀ : a₀₀ a₁₀} {a₋₁ : a₀₁ a₁₁} {a₋₂ : a₀₂ a₁₂} + (p : Square a₀₋ a₁₋ a₋₀ a₋₁) (q : Square b₀₋ b₁₋ a₋₁ a₋₂) + Square (a₀₋ b₀₋) (a₁₋ b₁₋) a₋₀ a₋₂ +_∙₂_ = congP₂ _ _∙_) + +-- Alternative (equivalent) definitions of hlevel n that give fillers for n-cubes instead of n-globes + +isSet' : Type Type +isSet' A = + {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ a₀₁) + {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ a₁₁) + (a₋₀ : a₀₀ a₁₀) (a₋₁ : a₀₁ a₁₁) + Square a₀₋ a₁₋ a₋₀ a₋₁ + +isSet→isSet' : isSet A isSet' A +isSet→isSet' Aset _ _ _ _ = toPathP (Aset _ _ _ _) + +isSet'→isSet : isSet' A isSet A +isSet'→isSet Aset' x y p q = Aset' p q refl refl + +isGroupoid' : Type Type +isGroupoid' A = + {a₀₀₀ a₀₀₁ : A} {a₀₀₋ : a₀₀₀ a₀₀₁} + {a₀₁₀ a₀₁₁ : A} {a₀₁₋ : a₀₁₀ a₀₁₁} + {a₀₋₀ : a₀₀₀ a₀₁₀} {a₀₋₁ : a₀₀₁ a₀₁₁} + (a₀₋₋ : Square a₀₀₋ a₀₁₋ a₀₋₀ a₀₋₁) + {a₁₀₀ a₁₀₁ : A} {a₁₀₋ : a₁₀₀ a₁₀₁} + {a₁₁₀ a₁₁₁ : A} {a₁₁₋ : a₁₁₀ a₁₁₁} + {a₁₋₀ : a₁₀₀ a₁₁₀} {a₁₋₁ : a₁₀₁ a₁₁₁} + (a₁₋₋ : Square a₁₀₋ a₁₁₋ a₁₋₀ a₁₋₁) + {a₋₀₀ : a₀₀₀ a₁₀₀} {a₋₀₁ : a₀₀₁ a₁₀₁} + (a₋₀₋ : Square a₀₀₋ a₁₀₋ a₋₀₀ a₋₀₁) + {a₋₁₀ : a₀₁₀ a₁₁₀} {a₋₁₁ : a₀₁₁ a₁₁₁} + (a₋₁₋ : Square a₀₁₋ a₁₁₋ a₋₁₀ a₋₁₁) + (a₋₋₀ : Square a₀₋₀ a₁₋₀ a₋₀₀ a₋₁₀) + (a₋₋₁ : Square a₀₋₁ a₁₋₁ a₋₀₁ a₋₁₁) + Cube a₀₋₋ a₁₋₋ a₋₀₋ a₋₁₋ a₋₋₀ a₋₋₁ + +-- Essential properties of isProp and isContr + +isProp→PathP : {B : I Type } ((i : I) isProp (B i)) + (b0 : B i0) (b1 : B i1) + PathP B b0 b1 +isProp→PathP hB b0 b1 = toPathP (hB _ _ _) + +isPropIsContr : isProp (isContr A) +isPropIsContr (c0 , h0) (c1 , h1) j .fst = h0 c1 j +isPropIsContr (c0 , h0) (c1 , h1) j .snd y i = + hcomp k λ { (i = i0) h0 (h0 c1 j) k; + (i = i1) h0 y k; + (j = i0) h0 (h0 y i) k; + (j = i1) h0 (h1 y i) k}) + c0 + +isContr→isProp : isContr A isProp A +isContr→isProp (x , p) a b = sym (p a) p b + +isProp→isSet : isProp A isSet A +isProp→isSet h a b p q j i = + hcomp k λ { (i = i0) h a a k + ; (i = i1) h a b k + ; (j = i0) h a (p i) k + ; (j = i1) h a (q i) k }) a + +isProp→isSet' : isProp A isSet' A +isProp→isSet' h {a} p q r s i j = + hcomp k λ { (i = i0) h a (p j) k + ; (i = i1) h a (q j) k + ; (j = i0) h a (r i) k + ; (j = i1) h a (s i) k}) a + +isPropIsProp : isProp (isProp A) +isPropIsProp f g i a b = isProp→isSet f a b (f a b) (g a b) i + +isPropSingl : {a : A} isProp (singl a) +isPropSingl = isContr→isProp (isContrSingl _) + +isPropSinglP : {A : I Type } {a : A i0} isProp (singlP A a) +isPropSinglP = isContr→isProp (isContrSinglP _ _) + +-- Universe lifting + +record Lift {i j} (A : Type i) : Type (ℓ-max i j) where + constructor lift + field + lower : A + +open Lift public + +liftExt : {A : Type } {a b : Lift {} {ℓ'} A} (lower a lower b) a b +liftExt x i = lift (x i) diff --git a/src/docs/Cubical.Foundations.SIP.html b/src/docs/Cubical.Foundations.SIP.html new file mode 100644 index 0000000..eec8ac3 --- /dev/null +++ b/src/docs/Cubical.Foundations.SIP.html @@ -0,0 +1,124 @@ +{- + +In this file we apply the cubical machinery to Martin Hötzel-Escardó's +structure identity principle: + +https://www.cs.bham.ac.uk/~mhe/HoTT-UF-in-Agda-Lecture-Notes/HoTT-UF-Agda.html#sns + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.SIP where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence renaming (ua-pathToEquiv to ua-pathToEquiv') +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma + +open import Cubical.Foundations.Structure public + +private + variable + ℓ₁ ℓ₂ ℓ₃ ℓ₄ ℓ₅ : Level + S : Type ℓ₁ Type ℓ₂ + +-- Note that for any equivalence (f , e) : X ≃ Y the type ι (X , s) (Y , t) (f , e) need not to be +-- a proposition. Indeed this type should correspond to the ways s and t can be identified +-- as S-structures. This we call a standard notion of structure or SNS. +-- We will use a different definition, but the two definitions are interchangeable. +SNS : (S : Type ℓ₁ Type ℓ₂) (ι : StrEquiv S ℓ₃) Type (ℓ-max (ℓ-max (ℓ-suc ℓ₁) ℓ₂) ℓ₃) +SNS {ℓ₁} S ι = {X : Type ℓ₁} (s t : S X) ι (X , s) (X , t) (idEquiv X) (s t) + +-- We introduce the notation for structure preserving equivalences a +-- bit differently, but this definition doesn't actually change from +-- Escardó's notes. +_≃[_]_ : (A : TypeWithStr ℓ₁ S) (ι : StrEquiv S ℓ₂) (B : TypeWithStr ℓ₁ S) Type (ℓ-max ℓ₁ ℓ₂) +A ≃[ ι ] B = Σ[ e typ A typ B ] (ι A B e) + + + +-- The following PathP version of SNS is a bit easier to work with +-- for the proof of the SIP +UnivalentStr : (S : Type ℓ₁ Type ℓ₂) (ι : StrEquiv S ℓ₃) Type (ℓ-max (ℓ-max (ℓ-suc ℓ₁) ℓ₂) ℓ₃) +UnivalentStr {ℓ₁} S ι = + {A B : TypeWithStr ℓ₁ S} (e : typ A typ B) + ι A B e PathP i S (ua e i)) (str A) (str B) + +-- A quick sanity-check that our definition is interchangeable with +-- Escardó's. The direction SNS→UnivalentStr corresponds more or less +-- to a dependent EquivJ formulation of Escardó's homomorphism-lemma. +UnivalentStr→SNS : (S : Type ℓ₁ Type ℓ₂) (ι : StrEquiv S ℓ₃) + UnivalentStr S ι SNS S ι +UnivalentStr→SNS S ι θ {X = X} s t = + ι (X , s) (X , t) (idEquiv X) + ≃⟨ θ (idEquiv X) + PathP i S (ua (idEquiv X) i)) s t + ≃⟨ pathToEquiv j PathP i S (uaIdEquiv {A = X} j i)) s t) + s t + + + +SNS→UnivalentStr : (ι : StrEquiv S ℓ₃) SNS S ι UnivalentStr S ι +SNS→UnivalentStr {S = S} ι θ {A = A} {B = B} e = EquivJ P C e (str A) (str B) + where + Y = typ B + + P : (X : Type _) X Y Type _ + P X e' = (s : S X) (t : S Y) ι (X , s) (Y , t) e' PathP i S (ua e' i)) s t + + C : (s t : S Y) ι (Y , s) (Y , t) (idEquiv Y) PathP i S (ua (idEquiv Y) i)) s t + C s t = + ι (Y , s) (Y , t) (idEquiv Y) + ≃⟨ θ s t + s t + ≃⟨ pathToEquiv j PathP i S (uaIdEquiv {A = Y} (~ j) i)) s t) + PathP i S (ua (idEquiv Y) i)) s t + + +TransportStr : {S : Type Type ℓ₁} (α : EquivAction S) Type (ℓ-max (ℓ-suc ) ℓ₁) +TransportStr {} {S = S} α = + {X Y : Type } (e : X Y) (s : S X) equivFun (α e) s subst S (ua e) s + +TransportStr→UnivalentStr : {S : Type Type ℓ₁} (α : EquivAction S) + TransportStr α UnivalentStr S (EquivAction→StrEquiv α) +TransportStr→UnivalentStr {S = S} α τ {X , s} {Y , t} e = + equivFun (α e) s t + ≃⟨ pathToEquiv (cong (_≡ t) (τ e s)) + subst S (ua e) s t + ≃⟨ invEquiv (PathP≃Path _ _ _) + PathP i S (ua e i)) s t + + +UnivalentStr→TransportStr : {S : Type Type ℓ₁} (α : EquivAction S) + UnivalentStr S (EquivAction→StrEquiv α) TransportStr α +UnivalentStr→TransportStr {S = S} α θ e s = + invEq (θ e) (transport-filler (cong S (ua e)) s) + +invTransportStr : {S : Type Type ℓ₂} (α : EquivAction S) (τ : TransportStr α) + {X Y : Type } (e : X Y) (t : S Y) invEq (α e) t subst⁻ S (ua e) t +invTransportStr {S = S} α τ e t = + sym (transport⁻Transport (cong S (ua e)) (invEq (α e) t)) + ∙∙ sym (cong (subst⁻ S (ua e)) (τ e (invEq (α e) t))) + ∙∙ cong (subst⁻ S (ua e)) (secEq (α e) t) + + +--- We can now define an invertible function +--- +--- sip : A ≃[ ι ] B → A ≡ B + +module _ {S : Type ℓ₁ Type ℓ₂} {ι : StrEquiv S ℓ₃} + (θ : UnivalentStr S ι) (A B : TypeWithStr ℓ₁ S) + where + + sip : A ≃[ ι ] B A B + sip (e , p) i = ua e i , θ e .fst p i + + SIP : A ≃[ ι ] B (A B) + SIP = + sip , isoToIsEquiv (compIso (Σ-cong-iso (invIso univalenceIso) (equivToIso θ)) ΣPathIsoPathΣ) + + sip⁻ : A B A ≃[ ι ] B + sip⁻ = invEq SIP diff --git a/src/docs/Cubical.Foundations.Structure.html b/src/docs/Cubical.Foundations.Structure.html new file mode 100644 index 0000000..57b7bc3 --- /dev/null +++ b/src/docs/Cubical.Foundations.Structure.html @@ -0,0 +1,47 @@ +{-# OPTIONS --safe #-} +module Cubical.Foundations.Structure where + +open import Cubical.Core.Everything +open import Cubical.Foundations.Prelude + +private + variable + ℓ' ℓ'' : Level + S : Type Type ℓ' + +-- A structure is a type-family S : Type ℓ → Type ℓ', i.e. for X : Type ℓ and s : S X, +-- the pair (X , s) : TypeWithStr ℓ S means that X is equipped with an S-structure, witnessed by s. + +TypeWithStr : ( : Level) (S : Type Type ℓ') Type (ℓ-max (ℓ-suc ) ℓ') +TypeWithStr S = Σ[ X Type ] S X + +typ : TypeWithStr S Type +typ = fst + +str : (A : TypeWithStr S) S (typ A) +str = snd + +-- Alternative notation for typ +⟨_⟩ : TypeWithStr S Type +⟨_⟩ = typ + +instance + mkTypeWithStr : {} {S : Type Type ℓ'} {X} {{S X}} TypeWithStr S + mkTypeWithStr {{i}} = _ , i + +-- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism. +-- This will be implemented by a function ι : StrEquiv S ℓ' +-- that gives us for any two types with S-structure (X , s) and (Y , t) a family: +-- ι (X , s) (Y , t) : (X ≃ Y) → Type ℓ'' +StrEquiv : (S : Type Type ℓ'') (ℓ' : Level) Type (ℓ-max (ℓ-suc (ℓ-max ℓ')) ℓ'') +StrEquiv {} S ℓ' = (A B : TypeWithStr S) typ A typ B Type ℓ' + +-- An S-structure may instead be equipped with an action on equivalences, which will +-- induce a notion of S-homomorphism + +EquivAction : (S : Type Type ℓ'') Type (ℓ-max (ℓ-suc ) ℓ'') +EquivAction {} S = {X Y : Type } X Y S X S Y + +EquivAction→StrEquiv : {S : Type Type ℓ''} + EquivAction S StrEquiv S ℓ'' +EquivAction→StrEquiv α (X , s) (Y , t) e = equivFun (α e) s t diff --git a/src/docs/Cubical.Foundations.Transport.html b/src/docs/Cubical.Foundations.Transport.html new file mode 100644 index 0000000..1112f7f --- /dev/null +++ b/src/docs/Cubical.Foundations.Transport.html @@ -0,0 +1,221 @@ +{- Basic theory about transport: + +- transport is invertible +- transport is an equivalence ([pathToEquiv]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Transport where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv as Equiv hiding (transpEquiv) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Function using (_∘_) + +-- Direct definition of transport filler, note that we have to +-- explicitly tell Agda that the type is constant (like in CHM) +transpFill : {} {A : Type } + (φ : I) + (A : (i : I) Type [ φ _ A) ]) + (u0 : outS (A i0)) + -------------------------------------- + PathP i outS (A i)) u0 (transp i outS (A i)) φ u0) +transpFill φ A u0 i = transp j outS (A (i j))) (~ i φ) u0 + +transport⁻ : {} {A B : Type } A B B A +transport⁻ p = transport i p (~ i)) + +subst⁻ : { ℓ'} {A : Type } {x y : A} (B : A Type ℓ') (p : x y) B y B x +subst⁻ B p pa = transport⁻ i B (p i)) pa + +subst⁻-filler : { ℓ'} {A : Type } {x y : A} (B : A Type ℓ') (p : x y) (b : B y) + PathP i B (p (~ i))) b (subst⁻ B p b) +subst⁻-filler B p = subst-filler B (sym p) + +transport-fillerExt : {} {A B : Type } (p : A B) + PathP i A p i) x x) (transport p) +transport-fillerExt p i x = transport-filler p x i + +transport⁻-fillerExt : {} {A B : Type } (p : A B) + PathP i p i A) x x) (transport⁻ p) +transport⁻-fillerExt p i x = transp j p (i ~ j)) (~ i) x + +transport-fillerExt⁻ : {} {A B : Type } (p : A B) + PathP i p i B) (transport p) x x) +transport-fillerExt⁻ p = symP (transport⁻-fillerExt (sym p)) + +transport⁻-fillerExt⁻ : {} {A B : Type } (p : A B) + PathP i B p i) (transport⁻ p) x x) +transport⁻-fillerExt⁻ p = symP (transport-fillerExt (sym p)) + + +transport⁻-filler : {} {A B : Type } (p : A B) (x : B) + PathP i p (~ i)) x (transport⁻ p x) +transport⁻-filler p x = transport-filler i p (~ i)) x + +transport⁻Transport : {} {A B : Type } (p : A B) (a : A) + transport⁻ p (transport p a) a +transport⁻Transport p a j = transport⁻-fillerExt p (~ j) (transport-fillerExt p (~ j) a) + +transportTransport⁻ : {} {A B : Type } (p : A B) (b : B) + transport p (transport⁻ p b) b +transportTransport⁻ p b j = transport-fillerExt⁻ p j (transport⁻-fillerExt⁻ p j b) + +subst⁻Subst : { ℓ'} {A : Type } {x y : A} (B : A Type ℓ') (p : x y) + (u : B x) subst⁻ B p (subst B p u) u +subst⁻Subst {x = x} {y = y} B p u = transport⁻Transport {A = B x} {B = B y} (cong B p) u + +substSubst⁻ : { ℓ'} {A : Type } {x y : A} (B : A Type ℓ') (p : x y) + (v : B y) subst B p (subst⁻ B p v) v +substSubst⁻ {x = x} {y = y} B p v = transportTransport⁻ {A = B x} {B = B y} (cong B p) v + +substEquiv : { ℓ'} {A : Type } {a a' : A} (P : A Type ℓ') (p : a a') P a P a' +substEquiv P p = (subst P p , isEquivTransport i P (p i))) + +liftEquiv : { ℓ'} {A B : Type } (P : Type Type ℓ') (e : A B) P A P B +liftEquiv P e = substEquiv P (ua e) + +transpEquiv : {} {A B : Type } (p : A B) i p i B +transpEquiv p = Equiv.transpEquiv i p i) +{-# WARNING_ON_USAGE transpEquiv "Deprecated: Use the more general `transpEquiv` from `Cubical.Foundations.Equiv` instead" #-} + +uaTransportη : {} {A B : Type } (P : A B) ua (pathToEquiv P) P +uaTransportη = uaη +{-# WARNING_ON_USAGE uaTransportη "Deprecated: Use `uaη` from `Cubical.Foundations.Univalence` instead of `uaTransportη`" #-} + +pathToIso : {} {A B : Type } A B Iso A B +Iso.fun (pathToIso x) = transport x +Iso.inv (pathToIso x) = transport⁻ x +Iso.rightInv (pathToIso x) = transportTransport⁻ x +Iso.leftInv (pathToIso x) = transport⁻Transport x + +substIso : { ℓ'} {A : Type } (B : A Type ℓ') {x y : A} (p : x y) Iso (B x) (B y) +substIso B p = pathToIso (cong B p) + +-- Redefining substEquiv in terms of substIso gives an explicit inverse +substEquiv' : { ℓ'} {A : Type } (B : A Type ℓ') {x y : A} (p : x y) B x B y +substEquiv' B p = isoToEquiv (substIso B p) + +isInjectiveTransport : { : Level} {A B : Type } {p q : A B} + transport p transport q p q +isInjectiveTransport {p = p} {q} α i = + hcomp + j λ + { (i = i0) retEq univalence p j + ; (i = i1) retEq univalence q j + }) + (invEq univalence ((λ a α i a) , t i)) + where + t : PathP i isEquiv a α i a)) (pathToEquiv p .snd) (pathToEquiv q .snd) + t = isProp→PathP i isPropIsEquiv a α i a)) _ _ + +transportUaInv : {} {A B : Type } (e : A B) transport (ua (invEquiv e)) transport (sym (ua e)) +transportUaInv e = cong transport (uaInvEquiv e) +-- notice that transport (ua e) would reduce, thus an alternative definition using EquivJ can give +-- refl for the case of idEquiv: +-- transportUaInv e = EquivJ (λ _ e → transport (ua (invEquiv e)) ≡ transport (sym (ua e))) refl e + +isSet-subst : { ℓ'} {A : Type } {B : A Type ℓ'} + (isSet-A : isSet A) + {a : A} + (p : a a) (x : B a) subst B p x x +isSet-subst {B = B} isSet-A p x = subst p′ subst B p′ x x) (isSet-A _ _ refl p) (substRefl {B = B} x) + +-- substituting along a composite path is equivalent to substituting twice +substComposite : { ℓ'} {A : Type } (B : A Type ℓ') + {x y z : A} (p : x y) (q : y z) (u : B x) + subst B (p q) u subst B q (subst B p u) +substComposite B p q Bx i = + transport (cong B (compPath-filler' p q (~ i))) (transport-fillerExt (cong B p) i Bx) + +-- transporting along a composite path is equivalent to transporting twice +transportComposite : {} {A B C : Type } (p : A B) (q : B C) (x : A) + transport (p q) x transport q (transport p x) +transportComposite = substComposite D D) + +-- substitution commutes with morphisms in slices +substCommSlice : { ℓ' ℓ''} {A : Type } + (B : A Type ℓ') (C : A Type ℓ'') + (F : a B a C a) + {x y : A} (p : x y) (u : B x) + subst C p (F x u) F y (subst B p u) +substCommSlice B C F p Bx a = + transport-fillerExt⁻ (cong C p) a (F _ (transport-fillerExt (cong B p) a Bx)) + +constSubstCommSlice : { ℓ' ℓ''} {A : Type } + (B : A Type ℓ') + (C : Type ℓ'') + (F : a B a C) + {x y : A} (p : x y) (u : B x) + (F x u) F y (subst B p u) +constSubstCommSlice B C F p Bx = (sym (transportRefl (F _ Bx)) substCommSlice B _ C) F p Bx) + +-- transporting over (λ i → B (p i) → C (p i)) divides the transport into +-- transports over (λ i → C (p i)) and (λ i → B (p (~ i))) +funTypeTransp : { ℓ' ℓ''} {A : Type } (B : A Type ℓ') (C : A Type ℓ'') {x y : A} (p : x y) (f : B x C x) + PathP i B (p i) C (p i)) f (subst C p f subst B (sym p)) +funTypeTransp B C {x = x} p f i b = + transp j C (p (j i))) (~ i) (f (transp j B (p (i ~ j))) (~ i) b)) + +-- transports between loop spaces preserve path composition +overPathFunct : {} {A : Type } {x y : A} (p q : x x) (P : x y) + transport i P i P i) (p q) + transport i P i P i) p transport i P i P i) q +overPathFunct p q = + J y P transport i P i P i) (p q) transport i P i P i) p transport i P i P i) q) + (transportRefl (p q) cong₂ _∙_ (sym (transportRefl p)) (sym (transportRefl q))) + +-- substition over families of paths +-- theorem 2.11.3 in The Book +substInPaths : { ℓ'} {A : Type } {B : Type ℓ'} {a a' : A} + (f g : A B) (p : a a') (q : f a g a) + subst x f x g x) p q sym (cong f p) q cong g p +substInPaths {a = a} f g p q = + J x p' (subst y f y g y) p' q) (sym (cong f p') q cong g p')) + p=refl p + where + p=refl : subst y f y g y) refl q + refl q refl + p=refl = subst y f y g y) refl q + ≡⟨ substRefl {B = y f y g y)} q q + ≡⟨ (rUnit q) lUnit (q refl) refl q refl + +flipTransport : {} {A : I Type } {x : A i0} {y : A i1} + x transport⁻ i A i) y + transport i A i) x y +flipTransport {A = A} {y = y} p = + cong (transport i A i)) p transportTransport⁻ i A i) y + +-- special cases of substInPaths from lemma 2.11.2 in The Book +module _ { : Level} {A : Type } {a x1 x2 : A} (p : x1 x2) where + substInPathsL : (q : a x1) subst x a x) p q q p + substInPathsL q = subst x a x) p q + ≡⟨ substInPaths _ a) x x) p q + sym (cong _ a) p) q cong x x) p + ≡⟨ assoc _ a) q p + (refl q) p + ≡⟨ cong (_∙ p) (sym (lUnit q)) q p + + substInPathsR : (q : x1 a) subst x x a) p q sym p q + substInPathsR q = subst x x a) p q + ≡⟨ substInPaths x x) _ a) p q + sym (cong x x) p) q cong _ a) p + ≡⟨ assoc (sym p) q refl + (sym p q) refl + ≡⟨ sym (rUnit (sym p q)) sym p q + +transport-filler-ua : {} {A B : Type } (e : A B) (a : A) + SquareP _ i ua e i) + (transport-filler (ua e) a) + (ua-gluePath e refl) + refl + (transportRefl (fst e a)) +transport-filler-ua {A = A} {B = B} (e , _) a j i = + let b = e a + tr = transportRefl b + z = tr (j ~ i) + in glue { (i = i0) a ; (i = i1) tr j }) + (hcomp k λ { (i = i0) b ; (i = i1) tr (j k) ; (j = i1) tr (~ i k) }) + (hcomp k λ { (i = i0) tr (j k) ; (i = i1) z ; (j = i1) z }) z)) diff --git a/src/docs/Cubical.Foundations.Univalence.html b/src/docs/Cubical.Foundations.Univalence.html new file mode 100644 index 0000000..cdee7c5 --- /dev/null +++ b/src/docs/Cubical.Foundations.Univalence.html @@ -0,0 +1,343 @@ +{- + +Proof of the standard formulation of the univalence theorem and +various consequences of univalence + +- Re-exports Glue types from Cubical.Core.Glue +- The ua constant and its computation rule (up to a path) +- Proof of univalence using that unglue is an equivalence ([EquivContr]) +- Equivalence induction ([EquivJ], [elimEquiv]) +- Univalence theorem ([univalence]) +- The computation rule for ua ([uaβ]) +- Isomorphism induction ([elimIso]) + +-} +{-# OPTIONS --safe #-} +module Cubical.Foundations.Univalence where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.GroupoidLaws + +open import Cubical.Data.Sigma.Base + +open import Cubical.Core.Glue public + using (Glue ; glue ; unglue) + +open import Cubical.Reflection.StrictEquiv + +private + variable + ℓ' : Level + +-- The ua constant +ua : {A B : Type } A B A B +ua {A = A} {B = B} e i = Glue B { (i = i0) (A , e) + ; (i = i1) (B , idEquiv B) }) + +uaIdEquiv : {A : Type } ua (idEquiv A) refl +uaIdEquiv {A = A} i j = Glue A {φ = i ~ j j} _ A , idEquiv A) + +-- Propositional extensionality +hPropExt : {A B : Type } isProp A isProp B (A B) (B A) A B +hPropExt Aprop Bprop f g = ua (propBiimpl→Equiv Aprop Bprop f g) + +-- the unglue and glue primitives specialized to the case of ua + +ua-unglue : {A B : Type } (e : A B) (i : I) (x : ua e i) + B {- [ _ ↦ (λ { (i = i0) → e .fst x ; (i = i1) → x }) ] -} +ua-unglue e i x = unglue (i ~ i) x + +ua-glue : {A B : Type } (e : A B) (i : I) (x : Partial (~ i) A) + (y : B [ _ { (i = i0) e .fst (x 1=1) }) ]) + ua e i {- [ _ ↦ (λ { (i = i0) → x 1=1 ; (i = i1) → outS y }) ] -} +ua-glue e i x y = glue {φ = i ~ i} { (i = i0) x 1=1 ; (i = i1) outS y }) (outS y) + +module _ {A B : Type } (e : A B) {x : A} {y : B} where + -- sometimes more useful are versions of these functions with the (i : I) factored in + + ua-ungluePath : PathP i ua e i) x y e .fst x y + ua-ungluePath p i = ua-unglue e i (p i) + + ua-gluePath : e .fst x y PathP i ua e i) x y + ua-gluePath p i = ua-glue e i { (i = i0) x }) (inS (p i)) + + -- ua-ungluePath and ua-gluePath are definitional inverses + ua-ungluePath-Equiv : (PathP i ua e i) x y) (e .fst x y) + unquoteDef ua-ungluePath-Equiv = + defStrictEquiv ua-ungluePath-Equiv ua-ungluePath ua-gluePath + +-- ua-unglue and ua-glue are also definitional inverses, in a way +-- strengthening the types of ua-unglue and ua-glue gives a nicer formulation of this, see below + +ua-unglue-glue : {A B : Type } (e : A B) (i : I) (x : Partial (~ i) A) (y : B [ _ _ ]) + ua-unglue e i (ua-glue e i x y) outS y +ua-unglue-glue _ _ _ _ = refl + +ua-glue-unglue : {A B : Type } (e : A B) (i : I) (x : ua e i) + ua-glue e i { (i = i0) x }) (inS (ua-unglue e i x)) x +ua-glue-unglue _ _ _ = refl + +-- mainly for documentation purposes, ua-unglue and ua-glue wrapped in cubical subtypes + +ua-unglueS : {A B : Type } (e : A B) (i : I) (x : A) (y : B) + ua e i [ _ { (i = i0) x ; (i = i1) y }) ] + B [ _ { (i = i0) e .fst x ; (i = i1) y }) ] +ua-unglueS e i x y s = inS (ua-unglue e i (outS s)) + +ua-glueS : {A B : Type } (e : A B) (i : I) (x : A) (y : B) + B [ _ { (i = i0) e .fst x ; (i = i1) y }) ] + ua e i [ _ { (i = i0) x ; (i = i1) y }) ] +ua-glueS e i x y s = inS (ua-glue e i { (i = i0) x }) (inS (outS s))) + +ua-unglueS-glueS : {A B : Type } (e : A B) (i : I) (x : A) (y : B) + (s : B [ _ { (i = i0) e .fst x ; (i = i1) y }) ]) + outS (ua-unglueS e i x y (ua-glueS e i x y s)) outS s +ua-unglueS-glueS _ _ _ _ _ = refl + +ua-glueS-unglueS : {A B : Type } (e : A B) (i : I) (x : A) (y : B) + (s : ua e i [ _ { (i = i0) x ; (i = i1) y }) ]) + outS (ua-glueS e i x y (ua-unglueS e i x y s)) outS s +ua-glueS-unglueS _ _ _ _ _ = refl + + +-- a version of ua-glue with a single endpoint, identical to `ua-gluePath e {x} refl i` +ua-gluePt : {A B : Type } (e : A B) (i : I) (x : A) + ua e i {- [ _ ↦ (λ { (i = i0) → x ; (i = i1) → e .fst x }) ] -} +ua-gluePt e i x = ua-glue e i { (i = i0) x }) (inS (e .fst x)) + + +-- Proof of univalence using that unglue is an equivalence: + +-- unglue is an equivalence +unglueIsEquiv : (A : Type ) (φ : I) + (f : PartialP φ o Σ[ T Type ] T A)) + isEquiv {A = Glue A f} (unglue φ) +equiv-proof (unglueIsEquiv A φ f) = λ (b : A) + let u : I Partial φ A + u i = λ{ (φ = i1) equivCtr (f 1=1 .snd) b .snd (~ i) } + ctr : fiber (unglue φ) b + ctr = ( glue { (φ = i1) equivCtr (f 1=1 .snd) b .fst }) (hcomp u b) + , λ j hfill u (inS b) (~ j)) + in ( ctr + , λ (v : fiber (unglue φ) b) i + let u' : I Partial (φ ~ i i) A + u' j = λ { (φ = i1) equivCtrPath (f 1=1 .snd) b v i .snd (~ j) + ; (i = i0) hfill u (inS b) j + ; (i = i1) v .snd (~ j) } + in ( glue { (φ = i1) equivCtrPath (f 1=1 .snd) b v i .fst }) (hcomp u' b) + , λ j hfill u' (inS b) (~ j))) + +-- Any partial family of equivalences can be extended to a total one +-- from Glue [ φ ↦ (T,f) ] A to A +unglueEquiv : (A : Type ) (φ : I) + (f : PartialP φ o Σ[ T Type ] T A)) + (Glue A f) A +unglueEquiv A φ f = ( unglue φ , unglueIsEquiv A φ f ) + + +-- The following is a formulation of univalence proposed by Martín Escardó: +-- https://groups.google.com/forum/#!msg/homotopytypetheory/HfCB_b-PNEU/Ibb48LvUMeUJ +-- See also Theorem 5.8.4 of the HoTT Book. +-- +-- The reason we have this formulation in the core library and not the +-- standard one is that this one is more direct to prove using that +-- unglue is an equivalence. The standard formulation can be found in +-- Cubical/Basics/Univalence. +-- +EquivContr : (A : Type ) ∃![ T Type ] (T A) +EquivContr { = } A = + ( (A , idEquiv A) + , idEquiv≡ ) + where + idEquiv≡ : (y : Σ (Type ) T T A)) (A , idEquiv A) y + idEquiv≡ w = \ { i .fst Glue A (f i) + ; i .snd .fst unglueEquiv _ _ (f i) .fst + ; i .snd .snd .equiv-proof unglueEquiv _ _ (f i) .snd .equiv-proof + } + where + f : i PartialP (~ i i) x Σ[ T Type ] T A) + f i = λ { (i = i0) A , idEquiv A ; (i = i1) w } + +contrSinglEquiv : {A B : Type } (e : A B) (B , idEquiv B) (A , e) +contrSinglEquiv {A = A} {B = B} e = + isContr→isProp (EquivContr B) (B , idEquiv B) (A , e) + +-- Equivalence induction +EquivJ : {A B : Type } (P : (A : Type ) (e : A B) Type ℓ') + (r : P B (idEquiv B)) (e : A B) P A e +EquivJ P r e = subst x P (x .fst) (x .snd)) (contrSinglEquiv e) r + +-- Transport along a path is an equivalence. +-- The proof is a special case of isEquivTransp where the line of types is +-- given by p, and the extend φ -- where the transport is constant -- is i0. +isEquivTransport : {A B : Type } (p : A B) isEquiv (transport p) +isEquivTransport p = isEquivTransp A φ where + A : I Type _ + A i = p i + + φ : I + φ = i0 + +pathToEquiv : {A B : Type } A B A B +pathToEquiv p .fst = transport p +pathToEquiv p .snd = isEquivTransport p + +pathToEquivRefl : {A : Type } pathToEquiv refl idEquiv A +pathToEquivRefl {A = A} = equivEq i x transp _ A) i x) + +-- The computation rule for ua. Because of "ghcomp" it is now very +-- simple compared to cubicaltt: +-- https://github.com/mortberg/cubicaltt/blob/master/examples/univalence.ctt#L202 +uaβ : {A B : Type } (e : A B) (x : A) transport (ua e) x equivFun e x +uaβ e x = transportRefl (equivFun e x) + +uaη : {A B : Type } (P : A B) ua (pathToEquiv P) P +uaη {A = A} {B = B} P i j = Glue B {φ = φ} sides where + -- Adapted from a proof by @dolio, cf. commit e42a6fa1 + φ = i j ~ j + + sides : Partial φ (Σ[ T Type _ ] T B) + sides (i = i1) = P j , transpEquiv k P k) j + sides (j = i0) = A , pathToEquiv P + sides (j = i1) = B , idEquiv B + +pathToEquiv-ua : {A B : Type } (e : A B) pathToEquiv (ua e) e +pathToEquiv-ua e = equivEq (funExt (uaβ e)) + +ua-pathToEquiv : {A B : Type } (p : A B) ua (pathToEquiv p) p +ua-pathToEquiv = uaη + +-- Univalence +univalenceIso : {A B : Type } Iso (A B) (A B) +univalenceIso .Iso.fun = pathToEquiv +univalenceIso .Iso.inv = ua +univalenceIso .Iso.rightInv = pathToEquiv-ua +univalenceIso .Iso.leftInv = ua-pathToEquiv + +isEquivPathToEquiv : {A B : Type } isEquiv (pathToEquiv {A = A} {B = B}) +isEquivPathToEquiv = isoToIsEquiv univalenceIso + +univalence : {A B : Type } (A B) (A B) +univalence .fst = pathToEquiv +univalence .snd = isEquivPathToEquiv + +-- Assuming that we have an inverse to ua we can easily prove univalence +module Univalence (au : {} {A B : Type } A B A B) + (aurefl : {} {A : Type } au refl idEquiv A) where + + ua-au : {A B : Type } (p : A B) ua (au p) p + ua-au {B = B} = J _ p ua (au p) p) + (cong ua aurefl uaIdEquiv) + + au-ua : {A B : Type } (e : A B) au (ua e) e + au-ua {B = B} = EquivJ _ f au (ua f) f) + (subst r au r idEquiv _) (sym uaIdEquiv) aurefl) + + isoThm : {} {A B : Type } Iso (A B) (A B) + isoThm .Iso.fun = au + isoThm .Iso.inv = ua + isoThm .Iso.rightInv = au-ua + isoThm .Iso.leftInv = ua-au + + thm : {} {A B : Type } isEquiv au + thm {A = A} {B = B} = isoToIsEquiv {B = A B} isoThm + +-- The original map from UniMath/Foundations +eqweqmap : {A B : Type } A B A B +eqweqmap {A = A} e = J X _ A X) (idEquiv A) e + +eqweqmapid : {A : Type } eqweqmap refl idEquiv A +eqweqmapid {A = A} = JRefl X _ A X) (idEquiv A) + +univalenceStatement : {A B : Type } isEquiv (eqweqmap {} {A} {B}) +univalenceStatement = Univalence.thm eqweqmap eqweqmapid + +univalenceUAH : {A B : Type } (A B) (A B) +univalenceUAH = ( _ , univalenceStatement ) + +univalencePath : {A B : Type } (A B) Lift (A B) +univalencePath = ua (compEquiv univalence LiftEquiv) + +-- Lemmas for constructing and destructing dependent paths in a function type where the domain is ua. +ua→ : { ℓ'} {A₀ A₁ : Type } {e : A₀ A₁} {B : (i : I) Type ℓ'} + {f₀ : A₀ B i0} {f₁ : A₁ B i1} + ((a : A₀) PathP B (f₀ a) (f₁ (e .fst a))) + PathP i ua e i B i) f₀ f₁ +ua→ {e = e} {f₀ = f₀} {f₁} h i a = + hcomp + j λ + { (i = i0) f₀ a + ; (i = i1) f₁ (lem a j) + }) + (h (transp j ua e (~ j i)) (~ i) a) i) + where + lem : a₁ e .fst (transport (sym (ua e)) a₁) a₁ + lem a₁ = secEq e _ transportRefl _ + +ua→⁻ : { ℓ'} {A₀ A₁ : Type } {e : A₀ A₁} {B : (i : I) Type ℓ'} + {f₀ : A₀ B i0} {f₁ : A₁ B i1} + PathP i ua e i B i) f₀ f₁ + ((a : A₀) PathP B (f₀ a) (f₁ (e .fst a))) +ua→⁻ {e = e} {f₀ = f₀} {f₁} p a i = + hcomp + k λ + { (i = i0) f₀ a + ; (i = i1) f₁ (uaβ e a k) + }) + (p i (transp j ua e (j i)) (~ i) a)) + +ua→2 : { ℓ' ℓ''} {A₀ A₁ : Type } {e₁ : A₀ A₁} + {B₀ B₁ : Type ℓ'} {e₂ : B₀ B₁} + {C : (i : I) Type ℓ''} + {f₀ : A₀ B₀ C i0} {f₁ : A₁ B₁ C i1} + (∀ a b PathP C (f₀ a b) (f₁ (e₁ .fst a) (e₂ .fst b))) + PathP i ua e₁ i ua e₂ i C i) f₀ f₁ +ua→2 h = ua→ (ua→ h) + +-- Useful lemma for unfolding a transported function over ua +-- If we would have regularity this would be refl +transportUAop₁ : {A B : Type } (e : A B) (f : A A) (x : B) + transport i ua e i ua e i) f x equivFun e (f (invEq e x)) +transportUAop₁ e f x i = transportRefl (equivFun e (f (invEq e (transportRefl x i)))) i + +-- Binary version +transportUAop₂ : {} {A B : Type } (e : A B) (f : A A A) (x y : B) + transport i ua e i ua e i ua e i) f x y + equivFun e (f (invEq e x) (invEq e y)) +transportUAop₂ e f x y i = + transportRefl (equivFun e (f (invEq e (transportRefl x i)) + (invEq e (transportRefl y i)))) i + +-- Alternative version of EquivJ that only requires a predicate on functions +elimEquivFun : {A B : Type } (P : (A : Type ) (A B) Type ℓ') + (r : P B (idfun B)) (e : A B) P A (e .fst) +elimEquivFun P r e = subst x P (x .fst) (x .snd .fst)) (contrSinglEquiv e) r + +-- Isomorphism induction +elimIso : {B : Type } (Q : {A : Type } (A B) (B A) Type ℓ') + (h : Q (idfun B) (idfun B)) {A : Type } + (f : A B) (g : B A) section f g retract f g Q f g +elimIso {} {ℓ'} {B} Q h {A} f g sfg rfg = rem1 f g sfg rfg + where + P : (A : Type ) (f : A B) Type (ℓ-max ℓ' ) + P A f = (g : B A) section f g retract f g Q f g + + rem : P B (idfun B) + rem g sfg rfg = subst (Q (idfun B)) i b (sfg b) (~ i)) h + + rem1 : {A : Type } (f : A B) P A f + rem1 f g sfg rfg = elimEquivFun P rem (f , isoToIsEquiv (iso f g sfg rfg)) g sfg rfg + + +uaInvEquiv : {A B : Type } (e : A B) ua (invEquiv e) sym (ua e) +uaInvEquiv {B = B} = EquivJ _ e ua (invEquiv e) sym (ua e)) + (cong ua (invEquivIdEquiv B)) + +uaCompEquiv : {A B C : Type } (e : A B) (f : B C) ua (compEquiv e f) ua e ua f +uaCompEquiv {B = B} {C} = EquivJ _ e (f : B C) ua (compEquiv e f) ua e ua f) + f cong ua (compEquivIdEquiv f) + sym (cong x x ua f) uaIdEquiv + sym (lUnit (ua f)))) diff --git a/src/docs/Cubical.Functions.Embedding.html b/src/docs/Cubical.Functions.Embedding.html new file mode 100644 index 0000000..997805a --- /dev/null +++ b/src/docs/Cubical.Functions.Embedding.html @@ -0,0 +1,476 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Embedding where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Transport +open import Cubical.Foundations.Univalence using (ua; univalence; pathToEquiv) +open import Cubical.Functions.Fibration + +open import Cubical.Data.Sigma +open import Cubical.Functions.Fibration +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Nullary using (Discrete; yes; no) +open import Cubical.Structures.Axioms + +open import Cubical.Reflection.StrictEquiv + +open import Cubical.Data.Nat using (; zero; suc) +open import Cubical.Data.Sigma + +private + variable + ℓ' ℓ'' : Level + A B C : Type + f h : A B + w x : A + y z : B + +-- Embeddings are generalizations of injections. The usual +-- definition of injection as: +-- +-- f x ≡ f y → x ≡ y +-- +-- is not well-behaved with higher h-levels, while embeddings +-- are. +isEmbedding : (A B) Type _ +isEmbedding f = w x isEquiv {A = w x} (cong f) + +isPropIsEmbedding : isProp (isEmbedding f) +isPropIsEmbedding {f = f} = isPropΠ2 λ _ _ isPropIsEquiv (cong f) + +-- Embedding is injection in the aforementioned sense: +isEmbedding→Inj + : {f : A B} + isEmbedding f + w x f w f x w x +isEmbedding→Inj {f = f} embb w x p + = equiv-proof (embb w x) p .fst .fst + +-- The converse implication holds if B is an h-set, see injEmbedding below. + + +-- If `f` is an embedding, we'd expect the fibers of `f` to be +-- propositions, like an injective function. +hasPropFibers : (A B) Type _ +hasPropFibers f = y isProp (fiber f y) + +-- This can be relaxed to having all prop fibers over the image, see [hasPropFibersOfImage→isEmbedding] +hasPropFibersOfImage : (A B) Type _ +hasPropFibersOfImage f = x isProp (fiber f (f x)) + +-- some notation +_↪_ : Type ℓ' Type ℓ'' Type (ℓ-max ℓ' ℓ'') +A B = Σ[ f (A B) ] isEmbedding f + +hasPropFibersIsProp : isProp (hasPropFibers f) +hasPropFibersIsProp = isPropΠ _ isPropIsProp) + +private + lemma₀ : (p : y z) fiber f y fiber f z + lemma₀ {f = f} p = λ i fiber f (p i) + + lemma₁ : isEmbedding f x isContr (fiber f (f x)) + lemma₁ {f = f} iE x = value , path + where + value : fiber f (f x) + value = (x , refl) + + path : ∀(fi : fiber f (f x)) value fi + path (w , p) i + = case equiv-proof (iE w x) p of λ + { ((q , sq) , _) + hfill j λ { (i = i0) (x , refl) + ; (i = i1) (w , sq j) + }) + (inS (q (~ i) , λ j f (q (~ i j)))) + i1 + } + +isEmbedding→hasPropFibers : isEmbedding f hasPropFibers f +isEmbedding→hasPropFibers iE y (x , p) + = subst f isProp f) (lemma₀ p) (isContr→isProp (lemma₁ iE x)) (x , p) + +private + fibCong→PathP + : {f : A B} + (p : f w f x) + (fi : fiber (cong f) p) + PathP i fiber f (p i)) (w , refl) (x , refl) + fibCong→PathP p (q , r) i = q i , λ j r j i + + PathP→fibCong + : {f : A B} + (p : f w f x) + (pp : PathP i fiber f (p i)) (w , refl) (x , refl)) + fiber (cong f) p + PathP→fibCong p pp = i fst (pp i)) , j i snd (pp i) j) + +PathP≡fibCong + : {f : A B} + (p : f w f x) + PathP i fiber f (p i)) (w , refl) (x , refl) fiber (cong f) p +PathP≡fibCong p + = isoToPath (iso (PathP→fibCong p) (fibCong→PathP p) _ refl) _ refl)) + +hasPropFibers→isEmbedding : hasPropFibers f isEmbedding f +hasPropFibers→isEmbedding {f = f} iP w x .equiv-proof p + = subst isContr (PathP≡fibCong p) (isProp→isContrPathP i iP (p i)) fw fx) + where + fw : fiber f (f w) + fw = (w , refl) + + fx : fiber f (f x) + fx = (x , refl) + +hasPropFibersOfImage→hasPropFibers : hasPropFibersOfImage f hasPropFibers f +hasPropFibersOfImage→hasPropFibers {f = f} fibImg y a b = + subst y isProp (fiber f y)) (snd a) (fibImg (fst a)) a b + +hasPropFibersOfImage→isEmbedding : hasPropFibersOfImage f isEmbedding f +hasPropFibersOfImage→isEmbedding = hasPropFibers→isEmbedding hasPropFibersOfImage→hasPropFibers + +isEmbedding≡hasPropFibers : isEmbedding f hasPropFibers f +isEmbedding≡hasPropFibers + = isoToPath + (iso isEmbedding→hasPropFibers + hasPropFibers→isEmbedding + _ hasPropFibersIsProp _ _) + _ isPropIsEmbedding _ _)) + +-- We use the characterization as hasPropFibers to show that naive injectivity +-- implies isEmbedding as long as B is an h-set. +module _ + {f : A B} + (isSetB : isSet B) + where + + module _ + (inj : ∀{w x} f w f x w x) + where + + injective→hasPropFibers : hasPropFibers f + injective→hasPropFibers y (x , fx≡y) (x' , fx'≡y) = + Σ≡Prop + _ isSetB _ _) + (inj (fx≡y sym (fx'≡y))) + + injEmbedding : isEmbedding f + injEmbedding = hasPropFibers→isEmbedding injective→hasPropFibers + + retractableIntoSet→isEmbedding : hasRetract f isEmbedding f + retractableIntoSet→isEmbedding (g , ret) = injEmbedding inj + where + inj : f w f x w x + inj {w = w} {x = x} p = sym (ret w) ∙∙ cong g p ∙∙ ret x + +isEquiv→hasPropFibers : isEquiv f hasPropFibers f +isEquiv→hasPropFibers e b = isContr→isProp (equiv-proof e b) + +isEquiv→isEmbedding : isEquiv f isEmbedding f +isEquiv→isEmbedding e = λ _ _ congEquiv (_ , e) .snd + +Equiv→Embedding : A B A B +Equiv→Embedding (f , isEquivF) = (f , isEquiv→isEmbedding isEquivF) + +id↪ : {} (A : Type ) A A +id↪ A = Equiv→Embedding (idEquiv A) + +iso→isEmbedding : {} {A B : Type } + (isom : Iso A B) + ------------------------------- + isEmbedding (Iso.fun isom) +iso→isEmbedding {A = A} {B} isom = (isEquiv→isEmbedding (equivIsEquiv (isoToEquiv isom))) + +isEmbedding→Injection : + {} {A B C : Type } + (a : A B) + (e : isEmbedding a) + ---------------------- + {f g : C A} + x (a (f x) a (g x)) (f x g x) +isEmbedding→Injection a e {f = f} {g} x = sym (ua (cong a , e (f x) (g x))) + +Embedding-into-Discrete→Discrete : A B Discrete B Discrete A +Embedding-into-Discrete→Discrete (f , isEmbeddingF) _≟_ x y with f x f y +... | yes p = yes (invIsEq (isEmbeddingF x y) p) +... | no ¬p = no (¬p cong f) + +Embedding-into-isProp→isProp : A B isProp B isProp A +Embedding-into-isProp→isProp (f , isEmbeddingF) isProp-B x y + = invIsEq (isEmbeddingF x y) (isProp-B (f x) (f y)) + +Embedding-into-isSet→isSet : A B isSet B isSet A +Embedding-into-isSet→isSet (f , isEmbeddingF) isSet-B x y p q = + p ≡⟨ sym (retIsEq isEquiv-cong-f p) + cong-f⁻¹ (cong f p) ≡⟨ cong cong-f⁻¹ cong-f-p≡cong-f-q + cong-f⁻¹ (cong f q) ≡⟨ retIsEq isEquiv-cong-f q + q + where + cong-f-p≡cong-f-q = isSet-B (f x) (f y) (cong f p) (cong f q) + isEquiv-cong-f = isEmbeddingF x y + cong-f⁻¹ = invIsEq isEquiv-cong-f + +Embedding-into-hLevel→hLevel + : n A B isOfHLevel (suc n) B isOfHLevel (suc n) A +Embedding-into-hLevel→hLevel zero = Embedding-into-isProp→isProp +Embedding-into-hLevel→hLevel (suc n) (f , isEmbeddingF) Blvl x y + = isOfHLevelRespectEquiv (suc n) (invEquiv equiv) subLvl + where + equiv : (x y) (f x f y) + equiv .fst = cong f + equiv .snd = isEmbeddingF x y + subLvl : isOfHLevel (suc n) (f x f y) + subLvl = Blvl (f x) (f y) + +-- We now show that the powerset is the subtype classifier +-- i.e. ℙ X ≃ Σ[A ∈ Type ℓ] (A ↪ X) +Embedding→Subset : {X : Type } Σ[ A Type ] (A X) X +Embedding→Subset (_ , f , isEmbeddingF) x = fiber f x , isEmbedding→hasPropFibers isEmbeddingF x + +Subset→Embedding : {X : Type } X Σ[ A Type ] (A X) +Subset→Embedding {X = X} A = D , fst , Ψ + where + D = Σ[ x X ] x A + + Ψ : isEmbedding fst + Ψ w x = isEmbeddingFstΣProp (∈-isProp A) + +Subset→Embedding→Subset : {X : Type } section (Embedding→Subset {} {X}) (Subset→Embedding {} {X}) +Subset→Embedding→Subset _ = funExt λ x Σ≡Prop _ isPropIsProp) (ua (FiberIso.fiberEquiv _ x)) + +Embedding→Subset→Embedding : {X : Type } retract (Embedding→Subset {} {X}) (Subset→Embedding {} {X}) +Embedding→Subset→Embedding { = } {X = X} (A , f , ψ) = + cong (equivFun Σ-assoc-≃) (Σ≡Prop _ isPropIsEmbedding) (retEq (fibrationEquiv X ) (A , f))) + +Subset≃Embedding : {X : Type } X (Σ[ A Type ] (A X)) +Subset≃Embedding = isoToEquiv (iso Subset→Embedding Embedding→Subset + Embedding→Subset→Embedding Subset→Embedding→Subset) + +Subset≡Embedding : {X : Type } X (Σ[ A Type ] (A X)) +Subset≡Embedding = ua Subset≃Embedding + +isEmbedding-∘ : isEmbedding f isEmbedding h isEmbedding (f h) +isEmbedding-∘ {f = f} {h = h} Embf Embh w x + = compEquiv (cong h , Embh w x) (cong f , Embf (h w) (h x)) .snd + +compEmbedding : (B C) (A B) (A C) +(compEmbedding (g , _ ) (f , _ )).fst = g f +(compEmbedding (_ , g↪) (_ , f↪)).snd = isEmbedding-∘ g↪ f↪ + +isEmbedding→embedsFibersIntoSingl + : isEmbedding f + z fiber f z singl z +isEmbedding→embedsFibersIntoSingl {f = f} isE z = e , isEmbE where + e : fiber f z singl z + e x = f (fst x) , sym (snd x) + + isEmbE : isEmbedding e + isEmbE u v = goal where + -- "adjust" ΣeqCf by trivial equivalences that hold judgementally, which should save compositions + Dom′ : u v Type _ + Dom′ u v = Σ[ p fst u fst v ] PathP i f (p i) z) (snd u) (snd v) + Cod′ : u v Type _ + Cod′ u v = Σ[ p f (fst u) f (fst v) ] PathP i p i z) (snd u) (snd v) + ΣeqCf : Dom′ u v Cod′ u v + ΣeqCf = Σ-cong-equiv-fst (_ , isE _ _) + + dom→ : u v Dom′ u v + dom→ p = cong fst p , cong snd p + dom← : Dom′ u v u v + dom← p i = p .fst i , p .snd i + + cod→ : e u e v Cod′ u v + cod→ p = cong fst p , cong (sym snd) p + cod← : Cod′ u v e u e v + cod← p i = p .fst i , sym (p .snd i) + + goal : isEquiv (cong e) + goal .equiv-proof x .fst .fst = + dom← (equivCtr ΣeqCf (cod→ x) .fst) + goal .equiv-proof x .fst .snd j = + cod← (equivCtr ΣeqCf (cod→ x) .snd j) + goal .equiv-proof x .snd (g , p) i .fst = + dom← (equivCtrPath ΣeqCf (cod→ x) (dom→ g , cong cod→ p) i .fst) + goal .equiv-proof x .snd (g , p) i .snd j = + cod← (equivCtrPath ΣeqCf (cod→ x) (dom→ g , cong cod→ p) i .snd j) + +isEmbedding→hasPropFibers′ : isEmbedding f hasPropFibers f +isEmbedding→hasPropFibers′ {f = f} iE z = + Embedding-into-isProp→isProp (isEmbedding→embedsFibersIntoSingl iE z) isPropSingl + +universeEmbedding : + { ℓ' : Level} + (F : Type Type ℓ') + (∀ X F X X) + isEmbedding F +universeEmbedding F liftingEquiv = hasPropFibersOfImage→isEmbedding propFibersF where + lemma : A B (F A F B) (B A) + lemma A B = (F A F B) ≃⟨ univalence + (F A F B) ≃⟨ equivComp (liftingEquiv A) (liftingEquiv B) + (A B) ≃⟨ invEquivEquiv + (B A) ≃⟨ invEquiv univalence + (B A) + fiberSingl : X fiber F (F X) singl X + fiberSingl X = Σ-cong-equiv-snd _ lemma _ _) + propFibersF : hasPropFibersOfImage F + propFibersF X = Embedding-into-isProp→isProp (Equiv→Embedding (fiberSingl X)) isPropSingl + +liftEmbedding : ( ℓ' : Level) + isEmbedding (Lift {i = } {j = ℓ'}) +liftEmbedding ℓ' = universeEmbedding (Lift {j = ℓ'}) _ invEquiv LiftEquiv) + +module FibrationIdentityPrinciple {B : Type } {ℓ'} where + -- note that fibrationEquiv (for good reason) uses ℓ' = ℓ-max ℓ ℓ', so we have to work + -- some universe magic to achieve good universe polymorphism + + -- First, prove it for the case that's dealt with in fibrationEquiv + Fibration′ = Fibration B (ℓ-max ℓ') + + module Lifted (f g : Fibration′) where + f≃g′ : Type (ℓ-max ℓ') + f≃g′ = b fiber (f .snd) b fiber (g .snd) b + + Fibration′IP : f≃g′ (f g) + Fibration′IP = + f≃g′ + ≃⟨ equivΠCod _ invEquiv univalence) + (∀ b fiber (f .snd) b fiber (g .snd) b) + ≃⟨ funExtEquiv + fiber (f .snd) fiber (g .snd) + ≃⟨ invEquiv (congEquiv (fibrationEquiv B ℓ')) + f g + + + -- Then embed into the above case by lifting the type + L : Type _ Type _ -- local synonym fixing the levels of Lift + L = Lift {i = ℓ'} {j = } + + liftFibration : Fibration B ℓ' Fibration′ + liftFibration (A , f) = L A , f lower + + hasPropFibersLiftFibration : hasPropFibers liftFibration + hasPropFibersLiftFibration (A , f) = + Embedding-into-isProp→isProp (Equiv→Embedding fiberChar) + (isPropΣ (isEmbedding→hasPropFibers (liftEmbedding _ _) A) + λ _ isEquiv→hasPropFibers (snd (invEquiv (preCompEquiv LiftEquiv))) _) + where + fiberChar : fiber liftFibration (A , f) + (Σ[ (E , eq) fiber L A ] fiber (_∘ lower) (transport⁻ i eq i B) f)) + fiberChar = + fiber liftFibration (A , f) + ≃⟨ Σ-cong-equiv-snd _ invEquiv ΣPath≃PathΣ) + (Σ[ (E , g) Fibration B ℓ' ] Σ[ eq (L E A) ] PathP i eq i B) (g lower) f) + ≃⟨ boringSwap + (Σ[ (E , eq) fiber L A ] Σ[ g (E B) ] PathP i eq i B) (g lower) f) + ≃⟨ Σ-cong-equiv-snd _ Σ-cong-equiv-snd λ _ pathToEquiv (PathP≡Path⁻ _ _ _)) + (Σ[ (E , eq) fiber L A ] fiber (_∘ lower) (transport⁻ i eq i B) f)) + where + unquoteDecl boringSwap = + declStrictEquiv boringSwap + ((E , g) , (eq , p)) ((E , eq) , (g , p))) + ((E , g) , (eq , p)) ((E , eq) , (g , p))) + + isEmbeddingLiftFibration : isEmbedding liftFibration + isEmbeddingLiftFibration = hasPropFibers→isEmbedding hasPropFibersLiftFibration + + -- and finish off + module _ (f g : Fibration B ℓ') where + open Lifted (liftFibration f) (liftFibration g) + f≃g : Type (ℓ-max ℓ') + f≃g = b fiber (f .snd) b fiber (g .snd) b + + FibrationIP : f≃g (f g) + FibrationIP = + f≃g ≃⟨ equivΠCod b equivComp (Σ-cong-equiv-fst LiftEquiv) + (Σ-cong-equiv-fst LiftEquiv)) + f≃g′ ≃⟨ Fibration′IP + (liftFibration f liftFibration g) ≃⟨ invEquiv (_ , isEmbeddingLiftFibration _ _) + (f g) + +_≃Fib_ : {B : Type } (f g : Fibration B ℓ') Type (ℓ-max ℓ') +_≃Fib_ = FibrationIdentityPrinciple.f≃g + +FibrationIP : {B : Type } (f g : Fibration B ℓ') f ≃Fib g (f g) +FibrationIP = FibrationIdentityPrinciple.FibrationIP + +Embedding : (B : Type ℓ') ( : Level) Type (ℓ-max ℓ' (ℓ-suc )) +Embedding B = Σ[ A Type ] A B + +module EmbeddingIdentityPrinciple {B : Type } {ℓ'} (f g : Embedding B ℓ') where + open Σ f renaming (fst to F) + open Σ g renaming (fst to G) + open Σ (f .snd) renaming (fst to ffun; snd to isEmbF) + open Σ (g .snd) renaming (fst to gfun; snd to isEmbG) + f≃g : Type _ + f≃g = (∀ b fiber ffun b fiber gfun b) × + (∀ b fiber gfun b fiber ffun b) + toFibr : Embedding B ℓ' Fibration B ℓ' + toFibr (A , (f , _)) = (A , f) + + isEmbeddingToFibr : isEmbedding toFibr + isEmbeddingToFibr w x = fullEquiv .snd where + -- carefully managed such that (cong toFibr) is the equivalence + fullEquiv : (w x) (toFibr w toFibr x) + fullEquiv = compEquiv (congEquiv (invEquiv Σ-assoc-≃)) (invEquiv (Σ≡PropEquiv _ isPropIsEmbedding))) + + EmbeddingIP : f≃g (f g) + EmbeddingIP = + f≃g + ≃⟨ strictIsoToEquiv (invIso toProdIso) + (∀ b (fiber ffun b fiber gfun b) × (fiber gfun b fiber ffun b)) + ≃⟨ equivΠCod _ isEquivPropBiimpl→Equiv (isEmbedding→hasPropFibers isEmbF _) + (isEmbedding→hasPropFibers isEmbG _)) + (∀ b (fiber (f .snd .fst) b) (fiber (g .snd .fst) b)) + ≃⟨ FibrationIP (toFibr f) (toFibr g) + (toFibr f toFibr g) + ≃⟨ invEquiv (_ , isEmbeddingToFibr _ _) + f g + + +_≃Emb_ : {B : Type } (f g : Embedding B ℓ') Type _ +_≃Emb_ = EmbeddingIdentityPrinciple.f≃g + +EmbeddingIP : {B : Type } (f g : Embedding B ℓ') f ≃Emb g (f g) +EmbeddingIP = EmbeddingIdentityPrinciple.EmbeddingIP + +-- Cantor's theorem for sets +Set-Embedding-into-Powerset : {A : Type } isSet A A A +Set-Embedding-into-Powerset {A = A} setA + = fun , (injEmbedding isSetℙ y sym (H₃ (H₂ y)))) + where fun : A A + fun a b = (a b) , (setA a b) + + H₂ : {a b : A} fun a fun b a (fun b) + H₂ {a} fa≡fb = transport (cong (fst (_$ a)) fa≡fb) refl + + H₃ : {a b : A} b (fun a) a b + H₃ b∈fa = b∈fa + +×Monotone↪ : {ℓa ℓb ℓc ℓd} + {A : Type ℓa} {B : Type ℓb} {C : Type ℓc} {D : Type ℓd} + A C B D (A × B) (C × D) +×Monotone↪ {A = A} {B = B} {C = C} {D = D} (f , embf) (g , embg) + = (map-× f g) , emb + where apmap : x y x y map-× f g x map-× f g y + apmap x y x≡y = ΣPathP (cong (f fst) x≡y , cong (g snd) x≡y) + + equiv : x y isEquiv (apmap x y) + equiv x y = ((invEquiv ΣPathP≃PathPΣ) + ∙ₑ (≃-× ((cong f) , (embf (fst x) (fst y))) + ((cong g) , (embg (snd x) (snd y)))) + ∙ₑ ΣPathP≃PathPΣ) .snd + + emb : isEmbedding (map-× f g) + emb x y = equiv x y + +EmbeddingΣProp : {A : Type } {B : A Type ℓ'} (∀ a isProp (B a)) Σ A B A +EmbeddingΣProp f = fst , _ _ isEmbeddingFstΣProp f) diff --git a/src/docs/Cubical.Functions.Fibration.html b/src/docs/Cubical.Functions.Fibration.html new file mode 100644 index 0000000..6acbc69 --- /dev/null +++ b/src/docs/Cubical.Functions.Fibration.html @@ -0,0 +1,111 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.Fibration where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels using (isOfHLevel→isOfHLevelDep) +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Path +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Transport +open import Cubical.Data.Sigma + +private + variable + ℓb : Level + B : Type ℓb + +module FiberIso {} (p⁻¹ : B Type ) (x : B) where + + p : Σ B p⁻¹ B + p = fst + + fwd : fiber p x p⁻¹ x + fwd ((x' , y) , q) = subst z p⁻¹ z) q y + + bwd : p⁻¹ x fiber p x + bwd y = (x , y) , refl + + fwd-bwd : x fwd (bwd x) x + fwd-bwd y = transportRefl y + + bwd-fwd : x bwd (fwd x) x + bwd-fwd ((x' , y) , q) i = h (r i) + where h : Σ[ s singl x ] p⁻¹ (s .fst) fiber p x + h ((x , p) , y) = (x , y) , sym p + r : Path (Σ[ s singl x ] p⁻¹ (s .fst)) + ((x , refl ) , subst p⁻¹ q y) + ((x' , sym q) , y ) + r = ΣPathP (isContrSingl x .snd (x' , sym q) + , toPathP (transport⁻Transport i p⁻¹ (q i)) y)) + + -- HoTT Lemma 4.8.1 + fiberEquiv : fiber p x p⁻¹ x + fiberEquiv = isoToEquiv (iso fwd bwd fwd-bwd bwd-fwd) + +open FiberIso using (fiberEquiv) public + +module _ {} {E : Type } (p : E B) where + + -- HoTT Lemma 4.8.2 + totalEquiv : E Σ B (fiber p) + totalEquiv = isoToEquiv isom + where isom : Iso E (Σ B (fiber p)) + Iso.fun isom x = p x , x , refl + Iso.inv isom (b , x , q) = x + Iso.leftInv isom x i = x + Iso.rightInv isom (b , x , q) i = q i , x , λ j q (i j) + +module _ (B : Type ℓb) ( : Level) where + private + ℓ' = ℓ-max ℓb + + -- HoTT Theorem 4.8.3 + fibrationEquiv : (Σ[ E Type ℓ' ] (E B)) (B Type ℓ') + fibrationEquiv = isoToEquiv isom + where isom : Iso (Σ[ E Type ℓ' ] (E B)) (B Type ℓ') + Iso.fun isom (E , p) = fiber p + Iso.inv isom p⁻¹ = Σ B p⁻¹ , fst + Iso.rightInv isom p⁻¹ i x = ua (fiberEquiv p⁻¹ x) i + Iso.leftInv isom (E , p) i = ua e (~ i) , fst ua-unglue e (~ i) + where e = totalEquiv p + + +module ForSets {E : Type } {isSetB : isSet B} (f : E B) where + module _ {x x'} {px : x x'} {a' : fiber f x} {b' : fiber f x'} where + -- fibers are equal when their representatives are equal + fibersEqIfRepsEq : fst a' fst b' + PathP i fiber f (px i)) a' b' + fibersEqIfRepsEq p = ΣPathP (p , + (isOfHLevel→isOfHLevelDep 1 + (v , w) isSetB (f v) w) + (snd a') (snd b') + i (p i , px i)))) +-- The path type in a fiber of f is equivalent to a fiber of (cong f) +open import Cubical.Foundations.Function + +fiberPath : { ℓ'} {A : Type } {B : Type ℓ'} {f : A B} {b : B} (h h' : fiber f b) + (Σ[ p (fst h fst h') ] (PathP i f (p i) b) (snd h) (snd h'))) + fiber (cong f) (h .snd ∙∙ refl ∙∙ sym (h' .snd)) +fiberPath h h' = cong (Σ (h .fst h' .fst)) (funExt λ p flipSquarePath PathP≡doubleCompPathʳ _ _ _ _) + +fiber≡ : { ℓ'} {A : Type } {B : Type ℓ'} {f : A B} {b : B} (h h' : fiber f b) + (h h') fiber (cong f) (h .snd ∙∙ refl ∙∙ sym (h' .snd)) +fiber≡ {f = f} {b = b} h h' = + ΣPath≡PathΣ ⁻¹ + fiberPath h h' + +fiberCong : { ℓ'} {A : Type } {B : Type ℓ'} (f : A B) {a₀ a₁ : A} (q : f a₀ f a₁) + fiber (cong f) q Path (fiber f (f a₁)) (a₀ , q) (a₁ , refl) +fiberCong f q = + cong (fiber (cong f)) (cong sym (lUnit (sym q))) + sym (fiber≡ (_ , q) (_ , refl)) + +FibrationStr : (B : Type ℓb) Type Type (ℓ-max ℓb) +FibrationStr B A = A B + +Fibration : (B : Type ℓb) ( : Level) Type (ℓ-max ℓb (ℓ-suc )) +Fibration {ℓb = ℓb} B = Σ[ A Type ] FibrationStr B A diff --git a/src/docs/Cubical.Functions.Fixpoint.html b/src/docs/Cubical.Functions.Fixpoint.html new file mode 100644 index 0000000..2e332e8 --- /dev/null +++ b/src/docs/Cubical.Functions.Fixpoint.html @@ -0,0 +1,43 @@ +{- + Definition of function fixpoint and Kraus' lemma +-} +{-# OPTIONS --safe #-} +module Cubical.Functions.Fixpoint where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.GroupoidLaws + +private + variable + : Level + A : Type + +Fixpoint : (A A) Type _ +Fixpoint {A = A} f = Σ A x f x x) + +fixpoint : {f : A A} Fixpoint f A +fixpoint = fst + +fixpointPath : {f : A A} (p : Fixpoint f) f (fixpoint p) fixpoint p +fixpointPath = snd + +-- Kraus' lemma +-- a version not using cubical features can be found at +-- https://www.cs.bham.ac.uk/~mhe/GeneralizedHedberg/html/GeneralizedHedberg.html#21576 +2-Constant→isPropFixpoint : (f : A A) 2-Constant f isProp (Fixpoint f) +2-Constant→isPropFixpoint f fconst (x , p) (y , q) i = s i , t i where + noose : x y f x f y + noose x y = sym (fconst x x) fconst x y + -- the main idea is that for any path p, cong f p does not depend on p + -- but only on its endpoints and the structure of 2-Constant f + KrausInsight : {x y} (p : x y) noose x y cong f p + KrausInsight {x} = J y p noose x y cong f p) (lCancel (fconst x x)) + -- Need to solve for a path s : x ≡ y, such that: + -- transport (λ i → cong f s i ≡ s i) p ≡ q + s : x y + s = sym p ∙∙ noose x y ∙∙ q + t' : PathP i noose x y i s i) p q + t' i j = doubleCompPath-filler (sym p) (noose x y) q j i + t : PathP i cong f s i s i) p q + t = subst kraus PathP i kraus i s i) p q) (KrausInsight s) t' diff --git a/src/docs/Cubical.Functions.FunExtEquiv.html b/src/docs/Cubical.Functions.FunExtEquiv.html new file mode 100644 index 0000000..a7e5903 --- /dev/null +++ b/src/docs/Cubical.Functions.FunExtEquiv.html @@ -0,0 +1,193 @@ +{-# OPTIONS --safe #-} +module Cubical.Functions.FunExtEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.CartesianKanOps +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Vec.Base +open import Cubical.Data.Vec.NAry +open import Cubical.Data.Nat + +open import Cubical.Reflection.StrictEquiv + +private + variable + ℓ₁ ℓ₂ ℓ₃ : Level + +-- Function extensionality is an equivalence +module _ {A : Type } {B : A I Type ℓ₁} + {f : (x : A) B x i0} {g : (x : A) B x i1} where + + funExtEquiv : (∀ x PathP (B x) (f x) (g x)) PathP i x B x i) f g + unquoteDef funExtEquiv = defStrictEquiv funExtEquiv funExt funExt⁻ + + funExtPath : (∀ x PathP (B x) (f x) (g x)) PathP i x B x i) f g + funExtPath = ua funExtEquiv + + funExtIso : Iso (∀ x PathP (B x) (f x) (g x)) (PathP i x B x i) f g) + funExtIso = iso funExt funExt⁻ x refl {x = x}) x refl {x = x}) + +-- Function extensionality for binary functions +funExt₂ : {A : Type } {B : A Type ℓ₁} {C : (x : A) B x I Type ℓ₂} + {f : (x : A) (y : B x) C x y i0} + {g : (x : A) (y : B x) C x y i1} + ((x : A) (y : B x) PathP (C x y) (f x y) (g x y)) + PathP i x y C x y i) f g +funExt₂ p i x y = p x y i + +-- Function extensionality for binary functions is an equivalence +module _ {A : Type } {B : A Type ℓ₁} {C : (x : A) B x I Type ℓ₂} + {f : (x : A) (y : B x) C x y i0} + {g : (x : A) (y : B x) C x y i1} where + private + appl₂ : PathP i x y C x y i) f g x y PathP (C x y) (f x y) (g x y) + appl₂ eq x y i = eq i x y + + funExt₂Equiv : (∀ x y PathP (C x y) (f x y) (g x y)) (PathP i x y C x y i) f g) + unquoteDef funExt₂Equiv = defStrictEquiv funExt₂Equiv funExt₂ appl₂ + + funExt₂Path : (∀ x y PathP (C x y) (f x y) (g x y)) (PathP i x y C x y i) f g) + funExt₂Path = ua funExt₂Equiv + + +-- Function extensionality for ternary functions +funExt₃ : {A : Type } {B : A Type ℓ₁} {C : (x : A) B x Type ℓ₂} + {D : (x : A) (y : B x) C x y I Type ℓ₃} + {f : (x : A) (y : B x) (z : C x y) D x y z i0} + {g : (x : A) (y : B x) (z : C x y) D x y z i1} + ((x : A) (y : B x) (z : C x y) PathP (D x y z) (f x y z) (g x y z)) + PathP i x y z D x y z i) f g +funExt₃ p i x y z = p x y z i + +-- Function extensionality for ternary functions is an equivalence +module _ {A : Type } {B : A Type ℓ₁} {C : (x : A) B x Type ℓ₂} + {D : (x : A) (y : B x) C x y I Type ℓ₃} + {f : (x : A) (y : B x) (z : C x y) D x y z i0} + {g : (x : A) (y : B x) (z : C x y) D x y z i1} where + private + appl₃ : PathP i x y z D x y z i) f g x y z PathP (D x y z) (f x y z) (g x y z) + appl₃ eq x y z i = eq i x y z + + funExt₃Equiv : (∀ x y z PathP (D x y z) (f x y z) (g x y z)) (PathP i x y z D x y z i) f g) + unquoteDef funExt₃Equiv = defStrictEquiv funExt₃Equiv funExt₃ appl₃ + + funExt₃Path : (∀ x y z PathP (D x y z) (f x y z) (g x y z)) (PathP i x y z D x y z i) f g) + funExt₃Path = ua funExt₃Equiv + + +-- n-ary non-dependent funext +nAryFunExt : {X : Type } {Y : I Type ℓ₁} (n : ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + ((xs : Vec X n) PathP Y (fX $ⁿ xs) (fY $ⁿ map x x) xs)) + PathP i nAryOp n X (Y i)) fX fY +nAryFunExt zero fX fY p = p [] +nAryFunExt (suc n) fX fY p i x = nAryFunExt n (fX x) (fY x) xs p (x xs)) i + +-- n-ary funext⁻ +nAryFunExt⁻ : (n : ) {X : Type } {Y : I Type ℓ₁} (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + PathP i nAryOp n X (Y i)) fX fY + ((xs : Vec X n) PathP Y (fX $ⁿ xs) (fY $ⁿ map x x) xs)) +nAryFunExt⁻ zero fX fY p [] = p +nAryFunExt⁻ (suc n) fX fY p (x xs) = nAryFunExt⁻ n (fX x) (fY x) i p i x) xs + +nAryFunExtEquiv : (n : ) {X : Type } {Y : I Type ℓ₁} (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + ((xs : Vec X n) PathP Y (fX $ⁿ xs) (fY $ⁿ map x x) xs)) PathP i nAryOp n X (Y i)) fX fY +nAryFunExtEquiv n {X} {Y} fX fY = isoToEquiv (iso (nAryFunExt n fX fY) (nAryFunExt⁻ n fX fY) + (linv n fX fY) (rinv n fX fY)) + where + linv : (n : ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + (p : PathP i nAryOp n X (Y i)) fX fY) + nAryFunExt n fX fY (nAryFunExt⁻ n fX fY p) p + linv zero fX fY p = refl + linv (suc n) fX fY p i j x = linv n (fX x) (fY x) k p k x) i j + + rinv : (n : ) (fX : nAryOp n X (Y i0)) (fY : nAryOp n X (Y i1)) + (p : (xs : Vec X n) PathP Y (fX $ⁿ xs) (fY $ⁿ map x x) xs)) + nAryFunExt⁻ n fX fY (nAryFunExt n fX fY p) p + rinv zero fX fY p i [] = p [] + rinv (suc n) fX fY p i (x xs) = rinv n (fX x) (fY x) ys i p (x ys) i) i xs + +-- Funext when the domain also depends on the interval + +funExtDep : {A : I Type } {B : (i : I) A i Type ℓ₁} + {f : (x : A i0) B i0 x} {g : (x : A i1) B i1 x} + ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) PathP i B i (p i)) (f x₀) (g x₁)) + PathP i (x : A i) B i x) f g +funExtDep {A = A} {B} {f} {g} h i x = + comp + k B i (coei→i A i x k)) + k λ + { (i = i0) f (coei→i A i0 x k) + ; (i = i1) g (coei→i A i1 x k) + }) + (h j coei→j A i j x) i) + +funExtDep⁻ : {A : I Type } {B : (i : I) A i Type ℓ₁} + {f : (x : A i0) B i0 x} {g : (x : A i1) B i1 x} + PathP i (x : A i) B i x) f g + ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) PathP i B i (p i)) (f x₀) (g x₁)) +funExtDep⁻ q p i = q i (p i) + +funExtDepEquiv : {A : I Type } {B : (i : I) A i Type ℓ₁} + {f : (x : A i0) B i0 x} {g : (x : A i1) B i1 x} + ({x₀ : A i0} {x₁ : A i1} (p : PathP A x₀ x₁) PathP i B i (p i)) (f x₀) (g x₁)) + PathP i (x : A i) B i x) f g +funExtDepEquiv {A = A} {B} {f} {g} = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun = funExtDep + isom .inv = funExtDep⁻ + isom .rightInv q m i x = + comp + k B i (coei→i A i x (k m))) + k λ + { (i = i0) f (coei→i A i0 x (k m)) + ; (i = i1) g (coei→i A i1 x (k m)) + ; (m = i1) q i x + }) + (q i (coei→i A i x m)) + isom .leftInv h m p i = + comp + k B i (lemi→i m k)) + k λ + { (i = i0) f (lemi→i m k) + ; (i = i1) g (lemi→i m k) + ; (m = i1) h p i + }) + (h j lemi→j j m) i) + where + lemi→j : j coei→j A i j (p i) p j + lemi→j j = + coei→j k coei→j A i k (p i) p k) i j (coei→i A i (p i)) + + lemi→i : PathP m lemi→j i m p i) (coei→i A i (p i)) refl + lemi→i = + sym (coei→i k coei→j A i k (p i) p k) i (coei→i A i (p i))) + λ m k lemi→j i (m k) + +heteroHomotopy≃Homotopy : {A : I Type } {B : (i : I) Type ℓ₁} + {f : A i0 B i0} {g : A i1 B i1} + ({x₀ : A i0} {x₁ : A i1} PathP A x₀ x₁ PathP B (f x₀) (g x₁)) + ((x₀ : A i0) PathP B (f x₀) (g (transport i A i) x₀))) +heteroHomotopy≃Homotopy {A = A} {B} {f} {g} = isoToEquiv isom + where + open Iso + isom : Iso _ _ + isom .fun h x₀ = h (isContrSinglP A x₀ .fst .snd) + isom .inv k {x₀} {x₁} p = + subst fib PathP B (f x₀) (g (fib .fst))) (isContrSinglP A x₀ .snd (x₁ , p)) (k x₀) + isom .rightInv k = funExt λ x₀ + cong α subst fib PathP B (f x₀) (g (fib .fst))) α (k x₀)) + (isProp→isSet isPropSinglP (isContrSinglP A x₀ .fst) _ + (isContrSinglP A x₀ .snd (isContrSinglP A x₀ .fst)) + refl) + transportRefl (k x₀) + isom .leftInv h j {x₀} {x₁} p = + transp + i PathP B (f x₀) (g (isContrSinglP A x₀ .snd (x₁ , p) (i j) .fst))) + j + (h (isContrSinglP A x₀ .snd (x₁ , p) j .snd)) diff --git a/src/docs/Cubical.Functions.Involution.html b/src/docs/Cubical.Functions.Involution.html new file mode 100644 index 0000000..42cde46 --- /dev/null +++ b/src/docs/Cubical.Functions.Involution.html @@ -0,0 +1,42 @@ +{-# OPTIONS --safe #-} + +module Cubical.Functions.Involution where + +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Univalence + +isInvolution : ∀{} {A : Type } (A A) Type _ +isInvolution f = x f (f x) x + +module _ {} {A : Type } {f : A A} (invol : isInvolution f) where + open Iso + + involIso : Iso A A + involIso .fun = f + involIso .inv = f + involIso .rightInv = invol + involIso .leftInv = invol + + involIsEquiv : isEquiv f + involIsEquiv = isoToIsEquiv involIso + + involEquiv : A A + involEquiv = f , involIsEquiv + + involPath : A A + involPath = ua involEquiv + + involEquivComp : compEquiv involEquiv involEquiv idEquiv A + involEquivComp + = equivEq i x invol x i) + + involPathComp : involPath involPath refl + involPathComp + = sym (uaCompEquiv involEquiv involEquiv) ∙∙ cong ua involEquivComp ∙∙ uaIdEquiv + + involPath² : Square involPath refl refl involPath + involPath² + = subst s Square involPath s refl involPath) + involPathComp (compPath-filler involPath involPath) diff --git a/src/docs/Cubical.Functions.Logic.html b/src/docs/Cubical.Functions.Logic.html new file mode 100644 index 0000000..4869f2d --- /dev/null +++ b/src/docs/Cubical.Functions.Logic.html @@ -0,0 +1,290 @@ +-- Various functions for manipulating hProps. +-- +-- This file used to be part of Foundations, but it turned out to be +-- not very useful so moved here. Feel free to upstream content. +-- +-- Note that it is often a bad idea to use hProp instead of having the +-- isProp proof separate. The reason is that Agda can rarely infer +-- isProp proofs making it easier to just give them explicitly instead +-- of having them bundled up with the type. +-- +{-# OPTIONS --safe #-} +module Cubical.Functions.Logic where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Structure +open import Cubical.Foundations.Function +open import Cubical.Foundations.Univalence using (hPropExt) + +import Cubical.Data.Empty as +open import Cubical.Data.Sum as using (_⊎_) +open import Cubical.Data.Unit +open import Cubical.Data.Sigma + +open import Cubical.HITs.PropositionalTruncation as PropTrunc + +open import Cubical.Relation.Nullary hiding (¬_) + + +-------------------------------------------------------------------------------- +-- The type hProp of mere propositions +-- the definition hProp is given in Foundations.HLevels +-- hProp ℓ = Σ (Type ℓ) isProp + +private + variable + ℓ' ℓ'' : Level + P Q R : hProp + A B C : Type + +infix 10 ¬_ +infixr 8 _⊔_ +infixr 8 _⊔′_ +infixr 8 _⊓_ +infixr 8 _⊓′_ +infixr 6 _⇒_ +infixr 4 _⇔_ +infix 30 _≡ₚ_ +infix 30 _≢ₚ_ + +infix 2 ∃[]-syntax +infix 2 ∃[∶]-syntax + +infix 2 ∀[∶]-syntax +infix 2 ∀[]-syntax + +infix 2 ⇒∶_⇐∶_ +infix 2 ⇐∶_⇒∶_ + +∥_∥ₚ : Type hProp + A ∥ₚ = A ∥₁ , isPropPropTrunc + +_≡ₚ_ : (x y : A) hProp _ +x ≡ₚ y = x y ∥ₚ + +hProp≡ : P Q P Q +hProp≡ = TypeOfHLevel≡ 1 + +isProp⟨⟩ : (A : hProp ) isProp A +isProp⟨⟩ = snd + +-------------------------------------------------------------------------------- +-- Logical implication of mere propositions + +_⇒_ : (A : hProp ) (B : hProp ℓ') hProp _ +A B = ( A B ) , isPropΠ λ _ isProp⟨⟩ B + +⇔toPath : P Q Q P P Q +⇔toPath {P = P} {Q = Q} P⇒Q Q⇒P = hProp≡ (hPropExt (isProp⟨⟩ P) (isProp⟨⟩ Q) P⇒Q Q⇒P) + +pathTo⇒ : P Q P Q +pathTo⇒ p x = subst fst p x + +pathTo⇐ : P Q Q P +pathTo⇐ p x = subst fst (sym p) x + +substₚ : {x y : A} (B : A hProp ) x ≡ₚ y B x B y +substₚ {x = x} {y = y} B = PropTrunc.elim _ isPropΠ λ _ isProp⟨⟩ (B y)) (subst (fst B)) + +-------------------------------------------------------------------------------- +-- Mixfix notations for ⇔-toPath +-- see ⊔-identityˡ and ⊔-identityʳ for the difference + +⇒∶_⇐∶_ : P Q Q P P Q +⇒∶_⇐∶_ = ⇔toPath + +⇐∶_⇒∶_ : Q P P Q P Q +⇐∶ g ⇒∶ f = ⇔toPath f g +-------------------------------------------------------------------------------- +-- False and True + + : hProp _ + = ⊥.⊥ , λ () + + : {} hProp + = Unit* , _ _ _ tt*) + +-------------------------------------------------------------------------------- +-- Pseudo-complement of mere propositions + +¬_ : hProp hProp _ +¬ A = ( A ⊥.⊥) , isPropΠ λ _ ⊥.isProp⊥ + +_≢ₚ_ : (x y : A) hProp _ +x ≢ₚ y = ¬ x ≡ₚ y + +-------------------------------------------------------------------------------- +-- Disjunction of mere propositions + +_⊔′_ : Type Type ℓ' Type _ +A ⊔′ B = A B ∥₁ + +_⊔_ : hProp hProp ℓ' hProp _ +P Q = P Q ∥ₚ + +inl : A A ⊔′ B +inl x = ⊎.inl x ∣₁ + +inr : B A ⊔′ B +inr x = ⊎.inr x ∣₁ + +⊔-elim : (P : hProp ) (Q : hProp ℓ') (R : P Q hProp ℓ'') + (∀ x R (inl x) ) (∀ y R (inr y) ) (∀ z R z ) +⊔-elim _ _ R P⇒R Q⇒R = PropTrunc.elim (snd R) (⊎.elim P⇒R Q⇒R) + +-------------------------------------------------------------------------------- +-- Conjunction of mere propositions +_⊓′_ : Type Type ℓ' Type _ +A ⊓′ B = A × B + +_⊓_ : hProp hProp ℓ' hProp _ +A B = A ⊓′ B , isOfHLevelΣ 1 (isProp⟨⟩ A) (\ _ isProp⟨⟩ B) + +⊓-intro : (P : hProp ) (Q : P hProp ℓ') (R : P hProp ℓ'') + (∀ a Q a ) (∀ a R a ) (∀ (a : P ) Q a R a ) +⊓-intro _ _ _ = \ f g a f a , g a + +-------------------------------------------------------------------------------- +-- Logical bi-implication of mere propositions + +_⇔_ : hProp hProp ℓ' hProp _ +A B = (A B) (B A) + +⇔-id : (P : hProp ) P P +⇔-id P = (idfun P ) , (idfun P ) + +-------------------------------------------------------------------------------- +-- Universal Quantifier + + +∀[∶]-syntax : (A hProp ) hProp _ +∀[∶]-syntax {A = A} P = (∀ x P x ) , isPropΠ (isProp⟨⟩ P) + +∀[]-syntax : (A hProp ) hProp _ +∀[]-syntax {A = A} P = (∀ x P x ) , isPropΠ (isProp⟨⟩ P) + +syntax ∀[∶]-syntax {A = A} a P) = ∀[ a A ] P +syntax ∀[]-syntax a P) = ∀[ a ] P + +-------------------------------------------------------------------------------- +-- Existential Quantifier + +∃[]-syntax : (A hProp ) hProp _ +∃[]-syntax {A = A} P = Σ A (⟨_⟩ P) ∥ₚ + +∃[∶]-syntax : (A hProp ) hProp _ +∃[∶]-syntax {A = A} P = Σ A (⟨_⟩ P) ∥ₚ + +syntax ∃[∶]-syntax {A = A} x P) = ∃[ x A ] P +syntax ∃[]-syntax x P) = ∃[ x ] P + +-------------------------------------------------------------------------------- +-- Decidable mere proposition + +Decₚ : (P : hProp ) hProp +Decₚ P = Dec P , isPropDec (isProp⟨⟩ P) + +-------------------------------------------------------------------------------- +-- Negation commutes with truncation + +∥¬A∥≡¬∥A∥ : (A : Type ) (A ⊥.⊥) ∥ₚ (¬ A ∥ₚ) +∥¬A∥≡¬∥A∥ _ = + ⇒∶ ¬A A PropTrunc.elim _ ⊥.isProp⊥) + (PropTrunc.elim _ isPropΠ λ _ ⊥.isProp⊥) ¬p p ¬p p) ¬A) A) + ⇐∶ λ ¬p p ¬p p ∣₁) ∣₁ + +-------------------------------------------------------------------------------- +-- (hProp, ⊔, ⊥) is a bounded ⊔-semilattice + +⊔-assoc : (P : hProp ) (Q : hProp ℓ') (R : hProp ℓ'') + P (Q R) (P Q) R +⊔-assoc P Q R = + ⇒∶ ⊔-elim P (Q R) _ (P Q) R) + (inl inl) + (⊔-elim Q R _ (P Q) R) (inl inr) inr) + + ⇐∶ assoc2 + where + assoc2 : (A ⊔′ B) ⊔′ C A ⊔′ (B ⊔′ C) + assoc2 ⊎.inr a ∣₁ = ⊎.inr ⊎.inr a ∣₁ ∣₁ + assoc2 ⊎.inl ⊎.inr b ∣₁ ∣₁ = ⊎.inr ⊎.inl b ∣₁ ∣₁ + assoc2 ⊎.inl ⊎.inl c ∣₁ ∣₁ = ⊎.inl c ∣₁ + assoc2 ⊎.inl (squash₁ x y i) ∣₁ = isPropPropTrunc (assoc2 ⊎.inl x ∣₁) (assoc2 ⊎.inl y ∣₁) i + assoc2 (squash₁ x y i) = isPropPropTrunc (assoc2 x) (assoc2 y) i + +⊔-idem : (P : hProp ) P P P +⊔-idem P = + ⇒∶ (⊔-elim P P (\ _ P) (\ x x) (\ x x)) + ⇐∶ inl + +⊔-comm : (P : hProp ) (Q : hProp ℓ') P Q Q P +⊔-comm P Q = + ⇒∶ (⊔-elim P Q (\ _ (Q P)) inr inl) + ⇐∶ (⊔-elim Q P (\ _ (P Q)) inr inl) + +⊔-identityˡ : (P : hProp ) P P +⊔-identityˡ P = + ⇒∶ (⊔-elim P _ P) ()) x x)) + ⇐∶ inr + +⊔-identityʳ : (P : hProp ) P P +⊔-identityʳ P = ⇔toPath (⊔-elim P _ P) x x) λ ()) inl + +-------------------------------------------------------------------------------- +-- (hProp, ⊓, ⊤) is a bounded ⊓-lattice + +⊓-assoc : (P : hProp ) (Q : hProp ℓ') (R : hProp ℓ'') + P Q R (P Q) R +⊓-assoc _ _ _ = + ⇒∶ {(x , (y , z)) (x , y) , z}) + ⇐∶ {((x , y) , z) x , (y , z) }) + +⊓-comm : (P : hProp ) (Q : hProp ℓ') P Q Q P +⊓-comm _ _ = ⇔toPath (\ p p .snd , p .fst) (\ p p .snd , p .fst) + +⊓-idem : (P : hProp ) P P P +⊓-idem _ = ⇔toPath fst x x , x) + +⊓-identityˡ : (P : hProp ) {} P P +⊓-identityˡ _ = ⇔toPath snd λ x tt* , x + +⊓-identityʳ : (P : hProp ) P {} P +⊓-identityʳ _ = ⇔toPath fst λ x x , tt* + +-------------------------------------------------------------------------------- +-- Distributive laws + +⇒-⊓-distrib : (P : hProp ) (Q : hProp ℓ')(R : hProp ℓ'') + P (Q R) (P Q) (P R) +⇒-⊓-distrib _ _ _ = + ⇒∶ f (fst f) , (snd f)) + ⇐∶ { (f , g) x f x , g x}) + +⊓-⊔-distribˡ : (P : hProp ) (Q : hProp ℓ')(R : hProp ℓ'') + P (Q R) (P Q) (P R) +⊓-⊔-distribˡ P Q R = + ⇒∶ { (x , a) ⊔-elim Q R _ (P Q) (P R)) + y ⊎.inl (x , y) ∣₁ ) + z ⊎.inr (x , z) ∣₁ ) a }) + + ⇐∶ ⊔-elim (P Q) (P R) _ P Q R) + y fst y , inl (snd y)) + z fst z , inr (snd z)) + +⊔-⊓-distribˡ : (P : hProp ) (Q : hProp ℓ')(R : hProp ℓ'') + P (Q R) (P Q) (P R) +⊔-⊓-distribˡ P Q R = + ⇒∶ ⊔-elim P (Q R) _ (P Q) (P R) ) + (\ x inl x , inl x) (\ p inr (p .fst) , inr (p .snd)) + + ⇐∶ { (x , y) ⊔-elim P R _ P Q R) inl + z ⊔-elim P Q _ P Q R) inl y inr (y , z)) x) y }) + +⊓-∀-distrib : (P : A hProp ) (Q : A hProp ℓ') + (∀[ a A ] P a) (∀[ a A ] Q a) (∀[ a A ] (P a Q a)) +⊓-∀-distrib P Q = + ⇒∶ {(p , q) a p a , q a}) + ⇐∶ λ pq (fst pq ) , (snd pq) diff --git a/src/docs/Cubical.HITs.PropositionalTruncation.Base.html b/src/docs/Cubical.HITs.PropositionalTruncation.Base.html new file mode 100644 index 0000000..ccaecd8 --- /dev/null +++ b/src/docs/Cubical.HITs.PropositionalTruncation.Base.html @@ -0,0 +1,17 @@ +{- + +This file contains: + +- Definition of propositional truncation + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.PropositionalTruncation.Base where + +open import Cubical.Core.Primitives + +-- Propositional truncation as a higher inductive type: + +data ∥_∥₁ {} (A : Type ) : Type where + ∣_∣₁ : A A ∥₁ + squash₁ : (x y : A ∥₁) x y diff --git a/src/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html b/src/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html new file mode 100644 index 0000000..0c7abb7 --- /dev/null +++ b/src/docs/Cubical.HITs.PropositionalTruncation.MagicTrick.html @@ -0,0 +1,88 @@ +{- + +Based on Nicolai Kraus' blog post: + The Truncation Map |_| : ℕ -> ‖ℕ‖ is nearly Invertible + https://homotopytypetheory.org/2013/10/28/the-truncation-map-_-ℕ-‖ℕ‖-is-nearly-invertible/ + +Defines [recover], which definitionally satisfies `recover ∣ x ∣ ≡ x` ([recover∣∣]) for homogeneous types + +Also see the follow-up post by Jason Gross: + Composition is not what you think it is! Why “nearly invertible” isn’t. + https://homotopytypetheory.org/2014/02/24/composition-is-not-what-you-think-it-is-why-nearly-invertible-isnt/ + +-} +{-# OPTIONS --safe #-} + +module Cubical.HITs.PropositionalTruncation.MagicTrick where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Path +open import Cubical.Foundations.Pointed +open import Cubical.Foundations.Pointed.Homogeneous + +open import Cubical.HITs.PropositionalTruncation.Base +open import Cubical.HITs.PropositionalTruncation.Properties + +module Recover {} (A∙ : Pointed ) (h : isHomogeneous A∙) where + private + A = typ A∙ + a = pt A∙ + + toEquivPtd : A ∥₁ Σ[ B∙ Pointed ] (A , a) B∙ + toEquivPtd = rec isPropSingl x (A , x) , h x) + private + B∙ : A ∥₁ Pointed + B∙ tx = fst (toEquivPtd tx) + + -- the key observation is that B∙ ∣ x ∣₁ is definitionally equal to (A , x) + private + obvs : x B∙ x ∣₁ (A , x) + obvs x = refl -- try it: `C-c C-n B∙ ∣ x ∣₁` gives `(A , x)` + + -- thus any truncated element (of a homogeneous type) can be recovered by agda's normalizer! + + recover : (tx : A ∥₁) typ (B∙ tx) + recover tx = pt (B∙ tx) + + recover∣∣ : (x : A) recover x ∣₁ x + recover∣∣ x = refl -- try it: `C-c C-n recover ∣ x ∣₁` gives `x` + + private + -- notice that the following typechecks because typ (B∙ ∣ x ∣₁) is definitionally equal to to A, but + -- `recover : ∥ A ∥₁ → A` does not because typ (B∙ tx) is not definitionally equal to A (though it is + -- judegmentally equal to A by cong typ (snd (toEquivPtd tx)) : A ≡ typ (B∙ tx)) + obvs2 : A A + obvs2 = recover ∣_∣₁ + + -- one might wonder if (cong recover (squash₁ ∣ x ∣₁ ∣ y ∣₁)) therefore has type x ≡ y, but thankfully + -- typ (B∙ (squash₁ ∣ x ∣₁ ∣ y ∣₁ i)) is *not* A (it's a messy hcomp involving h x and h y) + recover-squash₁ : x y -- x ≡ y -- this raises an error + PathP i typ (B∙ (squash₁ x ∣₁ y ∣₁ i))) x y + recover-squash₁ x y = cong recover (squash₁ x ∣₁ y ∣₁) + + +-- Demo, adapted from: +-- https://bitbucket.org/nicolaikraus/agda/src/e30d70c72c6af8e62b72eefabcc57623dd921f04/trunc-inverse.lagda + +private + open import Cubical.Data.Nat + open Recover ( , zero) (isHomogeneousDiscrete discreteℕ) + + -- only `∣hidden∣` is exported, `hidden` is no longer in scope + module _ where + private + hidden : + hidden = 17 + + ∣hidden∣ : ∥₁ + ∣hidden∣ = hidden ∣₁ + + -- we can still recover the value, even though agda can no longer see `hidden`! + test : recover ∣hidden∣ 17 + test = refl -- try it: `C-c C-n recover ∣hidden∣` gives `17` + -- `C-c C-n hidden` gives an error + + -- Finally, note that the definition of recover is independent of the proof that A is homogeneous. Thus we + -- still can definitionally recover information hidden by ∣_∣₁ as long as we permit holes. Try replacing + -- `isHomogeneousDiscrete discreteℕ` above with a hole (`?`) and notice that everything still works diff --git a/src/docs/Cubical.HITs.PropositionalTruncation.Properties.html b/src/docs/Cubical.HITs.PropositionalTruncation.Properties.html new file mode 100644 index 0000000..2cb1b40 --- /dev/null +++ b/src/docs/Cubical.HITs.PropositionalTruncation.Properties.html @@ -0,0 +1,564 @@ +{- + +This file contains: + +- Eliminator for propositional truncation + +-} +{-# OPTIONS --safe #-} +module Cubical.HITs.PropositionalTruncation.Properties where + +open import Cubical.Core.Everything + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence + +open import Cubical.Data.Sigma +open import Cubical.Data.Sum hiding (rec ; elim ; map) +open import Cubical.Data.Nat using ( ; zero ; suc) +open import Cubical.Data.FinData using (Fin ; zero ; suc) + +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + ℓ' : Level + A B C : Type + A′ : Type ℓ' + +∥∥-isPropDep : (P : A Type ) isOfHLevelDep 1 x P x ∥₁) +∥∥-isPropDep P = isOfHLevel→isOfHLevelDep 1 _ squash₁) + +rec : {P : Type } isProp P (A P) A ∥₁ P +rec Pprop f x ∣₁ = f x +rec Pprop f (squash₁ x y i) = Pprop (rec Pprop f x) (rec Pprop f y) i + +rec2 : {P : Type } isProp P (A B P) A ∥₁ B ∥₁ P +rec2 Pprop f x ∣₁ y ∣₁ = f x y +rec2 Pprop f x ∣₁ (squash₁ y z i) = Pprop (rec2 Pprop f x ∣₁ y) (rec2 Pprop f x ∣₁ z) i +rec2 Pprop f (squash₁ x y i) z = Pprop (rec2 Pprop f x z) (rec2 Pprop f y z) i + +rec3 : {P : Type } isProp P (A B C P) A ∥₁ B ∥₁ C ∥₁ P +rec3 Pprop f x ∣₁ y ∣₁ z ∣₁ = f x y z +rec3 Pprop f x ∣₁ y ∣₁ (squash₁ z w i) = Pprop (rec3 Pprop f x ∣₁ y ∣₁ z) (rec3 Pprop f x ∣₁ y ∣₁ w) i +rec3 Pprop f x ∣₁ (squash₁ y z i) w = Pprop (rec3 Pprop f x ∣₁ y w) (rec3 Pprop f x ∣₁ z w) i +rec3 Pprop f (squash₁ x y i) z w = Pprop (rec3 Pprop f x z w) (rec3 Pprop f y z w) i + +-- Old version +-- rec2 : ∀ {P : Type ℓ} → isProp P → (A → A → P) → ∥ A ∥ → ∥ A ∥ → P +-- rec2 Pprop f = rec (isProp→ Pprop) (λ a → rec Pprop (f a)) + +-- n-ary recursor, stated using a dependent FinVec +recFin : {m : } {P : Fin m Type } + {B : Type ℓ'} (isPropB : isProp B) + ((∀ i P i) B) + --------------------- + ((∀ i P i ∥₁) B) +recFin {m = zero} _ untruncHyp _ = untruncHyp ()) +recFin {m = suc m} {P = P} {B = B} isPropB untruncHyp truncFam = + curriedishTrunc (truncFam zero) (truncFam suc) + where + curriedish : P zero (∀ i P (suc i) ∥₁) B + curriedish p₀ = recFin isPropB + famSuc untruncHyp { zero p₀ ; (suc i) famSuc i })) + + curriedishTrunc : P zero ∥₁ (∀ i P (suc i) ∥₁) B + curriedishTrunc = rec (isProp→ isPropB) curriedish + +recFin2 : {m1 m2 : } {P : Fin m1 Fin m2 Type } + {B : Type ℓ'} (isPropB : isProp B) + ((∀ i j P i j) B) + -------------------------- + (∀ i j P i j ∥₁) + B +recFin2 {m1 = zero} _ untruncHyp _ = untruncHyp λ () +recFin2 {m1 = suc m1} {P = P} {B = B} isPropB untruncHyp truncFam = + curriedishTrunc (truncFam zero) (truncFam suc) + where + curriedish : (∀ j P zero j) (∀ i j P (suc i) j ∥₁) B + curriedish p₀ truncFamSuc = recFin2 isPropB + famSuc untruncHyp λ { zero p₀ ; (suc i) famSuc i }) + truncFamSuc + + curriedishTrunc : (∀ j P zero j ∥₁) (∀ i j P (suc i) j ∥₁) B + curriedishTrunc = recFin (isProp→ isPropB) curriedish + + +elim : {P : A ∥₁ Type } ((a : A ∥₁) isProp (P a)) + ((x : A) P x ∣₁) (a : A ∥₁) P a +elim Pprop f x ∣₁ = f x +elim Pprop f (squash₁ x y i) = + isOfHLevel→isOfHLevelDep 1 Pprop + (elim Pprop f x) (elim Pprop f y) (squash₁ x y) i + +elim2 : {P : A ∥₁ B ∥₁ Type } + (Pprop : (x : A ∥₁) (y : B ∥₁) isProp (P x y)) + (f : (a : A) (b : B) P a ∣₁ b ∣₁) + (x : A ∥₁) (y : B ∥₁) P x y +elim2 Pprop f = + elim _ isPropΠ _ Pprop _ _)) + a elim _ Pprop _ _) (f a)) + +elim3 : {P : A ∥₁ B ∥₁ C ∥₁ Type } + (Pprop : ((x : A ∥₁) (y : B ∥₁) (z : C ∥₁) isProp (P x y z))) + (g : (a : A) (b : B) (c : C) P ( a ∣₁) b ∣₁ c ∣₁) + (x : A ∥₁) (y : B ∥₁) (z : C ∥₁) P x y z +elim3 Pprop g = elim2 _ _ isPropΠ _ Pprop _ _ _)) + a b elim _ Pprop _ _ _) (g a b)) + +-- n-ary eliminator, stated using a dependent FinVec +elimFin : {m : } {P : Fin m Type } + {B : (∀ i P i ∥₁) Type ℓ'} + (isPropB : x isProp (B x)) + ((x : i P i) B i x i ∣₁)) + ---------------------------------------- + ((x : i P i ∥₁) B x) +elimFin {m = zero} {B = B} _ untruncHyp _ = subst B (funExt ())) (untruncHyp ())) +elimFin {m = suc m} {P = P} {B = B} isPropB untruncHyp x = + subst B (funExt { zero refl ; (suc i) refl})) + (curriedishTrunc (x zero) (x suc)) + where + curriedish : (x₀ : P zero) (xₛ : i P (suc i) ∥₁) + B { zero x₀ ∣₁ ; (suc i) xₛ i}) + curriedish x₀ xₛ = subst B (funExt { zero refl ; (suc i) refl})) + (elimFin xₛ isPropB { zero x₀ ∣₁ ; (suc i) xₛ i})) + y subst B (funExt { zero refl ; (suc i) refl})) + (untruncHyp { zero x₀ ; (suc i) y i }))) xₛ) + + curriedishTrunc : (x₀ : P zero ∥₁) (xₛ : i P (suc i) ∥₁) + B { zero x₀ ; (suc i) xₛ i}) + curriedishTrunc = elim _ isPropΠ λ _ isPropB _) + λ x₀ xₛ subst B (funExt { zero refl ; (suc i) refl})) + (curriedish x₀ xₛ) + +isPropPropTrunc : isProp A ∥₁ +isPropPropTrunc x y = squash₁ x y + +propTrunc≃ : A B A ∥₁ B ∥₁ +propTrunc≃ e = + propBiimpl→Equiv + isPropPropTrunc isPropPropTrunc + (rec isPropPropTrunc a e .fst a ∣₁)) + (rec isPropPropTrunc b invEq e b ∣₁)) + +propTruncIdempotent≃ : isProp A A ∥₁ A +propTruncIdempotent≃ {A = A} hA = isoToEquiv f + where + f : Iso A ∥₁ A + Iso.fun f = rec hA (idfun A) + Iso.inv f x = x ∣₁ + Iso.rightInv f _ = refl + Iso.leftInv f = elim _ isProp→isSet isPropPropTrunc _ _) _ refl) + +propTruncIdempotent : isProp A A ∥₁ A +propTruncIdempotent hA = ua (propTruncIdempotent≃ hA) + +-- We could also define the eliminator using the recursor +elim' : {P : A ∥₁ Type } ((a : A ∥₁) isProp (P a)) + ((x : A) P x ∣₁) (a : A ∥₁) P a +elim' {P = P} Pprop f a = + rec (Pprop a) x transp i P (squash₁ x ∣₁ a i)) i0 (f x)) a + +map : (A B) ( A ∥₁ B ∥₁) +map f = rec squash₁ (∣_∣₁ f) + +map2 : (A B C) ( A ∥₁ B ∥₁ C ∥₁) +map2 f = rec (isPropΠ λ _ squash₁) (map f) + +-- The propositional truncation can be eliminated into non-propositional +-- types as long as the function used in the eliminator is 'coherently +-- constant.' The details of this can be found in the following paper: +-- +-- https://arxiv.org/pdf/1411.2682.pdf +module SetElim (Bset : isSet B) where + Bset' : isSet' B + Bset' = isSet→isSet' Bset + + rec→Set : (f : A B) (kf : 2-Constant f) A ∥₁ B + helper : (f : A B) (kf : 2-Constant f) (t u : A ∥₁) + rec→Set f kf t rec→Set f kf u + + rec→Set f kf x ∣₁ = f x + rec→Set f kf (squash₁ t u i) = helper f kf t u i + + helper f kf x ∣₁ y ∣₁ = kf x y + helper f kf (squash₁ t u i) v + = Bset' (helper f kf t v) (helper f kf u v) (helper f kf t u) refl i + helper f kf t (squash₁ u v i) + = Bset' (helper f kf t u) (helper f kf t v) refl (helper f kf u v) i + + kcomp : (f : A ∥₁ B) 2-Constant (f ∣_∣₁) + kcomp f x y = cong f (squash₁ x ∣₁ y ∣₁) + + Fset : isSet (A B) + Fset = isSetΠ (const Bset) + + Kset : (f : A B) isSet (2-Constant f) + Kset f = isSetΠ _ isSetΠ _ isProp→isSet (Bset _ _))) + + setRecLemma + : (f : A ∥₁ B) + rec→Set (f ∣_∣₁) (kcomp f) f + setRecLemma f i t + = elim {P = λ t rec→Set (f ∣_∣₁) (kcomp f) t f t} + t Bset _ _) x refl) t i + + mkKmap : ( A ∥₁ B) Σ (A B) 2-Constant + mkKmap f = f ∣_∣₁ , kcomp f + + fib : (g : Σ (A B) 2-Constant) fiber mkKmap g + fib (g , kg) = rec→Set g kg , refl + + eqv : (g : Σ (A B) 2-Constant) fi fib g fi + eqv g (f , p) = + Σ≡Prop f isOfHLevelΣ 2 Fset Kset _ _) + (cong (uncurry rec→Set) (sym p) setRecLemma f) + + trunc→Set≃ : ( A ∥₁ B) (Σ (A B) 2-Constant) + trunc→Set≃ .fst = mkKmap + trunc→Set≃ .snd .equiv-proof g = fib g , eqv g + + -- The strategy of this equivalence proof follows the paper more closely. + -- It is used further down for the groupoid version, because the above + -- strategy does not generalize so easily. + e : B Σ (A B) 2-Constant + e b = const b , λ _ _ refl + + eval : A (γ : Σ (A B) 2-Constant) B + eval a₀ (g , _) = g a₀ + + e-eval : (a₀ : A) γ e (eval a₀ γ) γ + e-eval a₀ (g , kg) i .fst a₁ = kg a₀ a₁ i + e-eval a₀ (g , kg) i .snd a₁ a₂ = Bset' refl (kg a₁ a₂) (kg a₀ a₁) (kg a₀ a₂) i + + e-isEquiv : A isEquiv (e {A = A}) + e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ refl) + + preEquiv₁ : A ∥₁ B Σ (A B) 2-Constant + preEquiv₁ t = e , rec (isPropIsEquiv e) e-isEquiv t + + preEquiv₂ : ( A ∥₁ Σ (A B) 2-Constant) Σ (A B) 2-Constant + preEquiv₂ = isoToEquiv (iso to const _ refl) retr) + where + to : ( A ∥₁ Σ (A B) 2-Constant) Σ (A B) 2-Constant + to f .fst x = f x ∣₁ .fst x + to f .snd x y i = f (squash₁ x ∣₁ y ∣₁ i) .snd x y i + + retr : retract to const + retr f i t .fst x = f (squash₁ x ∣₁ t i) .fst x + retr f i t .snd x y + = Bset' + j f (squash₁ x ∣₁ y ∣₁ j) .snd x y j) + (f t .snd x y) + j f (squash₁ x ∣₁ t j) .fst x) + j f (squash₁ y ∣₁ t j) .fst y) + i + + trunc→Set≃₂ : ( A ∥₁ B) Σ (A B) 2-Constant + trunc→Set≃₂ = compEquiv (equivΠCod preEquiv₁) preEquiv₂ + +open SetElim public using (rec→Set; trunc→Set≃) + +elim→Set + : {P : A ∥₁ Type } + (∀ t isSet (P t)) + (f : (x : A) P x ∣₁) + (kf : x y PathP i P (squash₁ x ∣₁ y ∣₁ i)) (f x) (f y)) + (t : A ∥₁) P t +elim→Set {A = A} {P = P} Pset f kf t + = rec→Set (Pset t) g gk t + where + g : A P t + g x = transp i P (squash₁ x ∣₁ t i)) i0 (f x) + + gk : 2-Constant g + gk x y i = transp j P (squash₁ (squash₁ x ∣₁ y ∣₁ i) t j)) i0 (kf x y i) + +elim2→Set : + {P : A ∥₁ B ∥₁ Type } + (∀ t u isSet (P t u)) + (f : (x : A) (y : B) P x ∣₁ y ∣₁) + (kf₁ : x y v PathP i P (squash₁ x ∣₁ y ∣₁ i) v ∣₁) (f x v) (f y v)) + (kf₂ : x v w PathP i P x ∣₁ (squash₁ v ∣₁ w ∣₁ i)) (f x v) (f x w)) + (sf : x y v w SquareP i j P (squash₁ x ∣₁ y ∣₁ i) (squash₁ v ∣₁ w ∣₁ j)) + (kf₂ x v w) (kf₂ y v w) (kf₁ x y v) (kf₁ x y w)) + (t : A ∥₁) (u : B ∥₁) P t u +elim2→Set {A = A} {B = B} {P = P} Pset f kf₁ kf₂ sf = + elim→Set _ isSetΠ _ Pset _ _)) mapHelper squareHelper + where + mapHelper : (x : A) (u : B ∥₁) P x ∣₁ u + mapHelper x = elim→Set _ Pset _ _) (f x) (kf₂ x) + + squareHelper : (x y : A) + PathP i (u : B ∥₁) P (squash₁ x ∣₁ y ∣₁ i) u) (mapHelper x) (mapHelper y) + squareHelper x y i = elim→Set _ Pset _ _) v kf₁ x y v i) λ v w sf x y v w i + +RecHProp : (P : A hProp ) (kP : x y P x P y) A ∥₁ hProp +RecHProp P kP = rec→Set isSetHProp P kP + +squash₁ᵗ + : ∀(x y z : A) + Square (squash₁ x ∣₁ y ∣₁) (squash₁ x ∣₁ z ∣₁) refl (squash₁ y ∣₁ z ∣₁) +squash₁ᵗ x y z i = squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) + +module _ (B : A ∥₁ Type ) + (B-gpd : (a : _) isGroupoid (B a)) + (f : (a : A) B a ∣₁) + (f-coh : (x y : A) PathP i B (squash₁ x ∣₁ y ∣₁ i)) (f x) (f y)) + (f-coh-coh : (x y z : A) SquareP + i j B (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j)) + (f-coh x y) (f-coh x z) refl (f-coh y z)) + where + elim→Gpd : (t : A ∥₁) B t + private + pathHelper : (t u : A ∥₁) PathP i B (squash₁ t u i)) (elim→Gpd t) (elim→Gpd u) + triHelper₁ + : (t u v : A ∥₁) + SquareP i j B (squash₁ t (squash₁ u v i) j)) + (pathHelper t u) (pathHelper t v) + refl (pathHelper u v) + triHelper₂ + : (t u v : A ∥₁) + SquareP i j B (squash₁ (squash₁ t u i) v j)) + (pathHelper t v) (pathHelper u v) + (pathHelper t u) refl + triHelper₂Cube : (x y z : A ∥₁) + Cube j k squash₁ x z (k j)) + j k squash₁ y z j) + i k squash₁ x y i) + i k squash₁ x z (i k)) + i j squash₁ x (squash₁ y z j) i) + i j squash₁ (squash₁ x y i) z j) + + elim→Gpd x ∣₁ = f x + elim→Gpd (squash₁ t u i) = pathHelper t u i + triHelper₂Cube x y z = + isProp→PathP _ isOfHLevelPathP 1 (isOfHLevelPath 1 squash₁ _ _) _ _) _ _ + + pathHelper x ∣₁ y ∣₁ = f-coh x y + pathHelper (squash₁ t u j) v = triHelper₂ t u v j + pathHelper x ∣₁ (squash₁ u v j) = triHelper₁ x ∣₁ u v j + + triHelper₁ x ∣₁ y ∣₁ z ∣₁ = f-coh-coh x y z + triHelper₁ (squash₁ s t i) u v + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ s t i) (squash₁ u v i₁) j)) + (triHelper₁ s u v) (triHelper₁ t u v) + (triHelper₂ s t u) + (triHelper₂ s t v) + i j pathHelper s t i) + i j pathHelper u v j) + (B-gpd v) i + + triHelper₁ x ∣₁ (squash₁ t u i) v + = isGroupoid→CubeP i i₁ j B (squash₁ x ∣₁ (squash₁ (squash₁ t u i) v i₁) j)) + (triHelper₁ x ∣₁ t v) (triHelper₁ x ∣₁ u v) + (triHelper₁ x ∣₁ t u) + i j pathHelper x ∣₁ v j) + refl (triHelper₂ t u v) + (B-gpd v) i + triHelper₁ x ∣₁ y ∣₁ (squash₁ u v i) + = isGroupoid→CubeP i i₁ j B (squash₁ x ∣₁ (squash₁ y ∣₁ (squash₁ u v i) i₁) j)) + (triHelper₁ x ∣₁ y ∣₁ u) (triHelper₁ x ∣₁ y ∣₁ v) + i j f-coh x y j) (triHelper₁ x ∣₁ u v) + refl (triHelper₁ y ∣₁ u v) + (B-gpd v) i + triHelper₂ x ∣₁ y ∣₁ z ∣₁ i j = + comp k B (triHelper₂Cube x ∣₁ y ∣₁ z ∣₁ i j k)) + k λ {(i = i0) f-coh x z (k j) + ; (i = i1) f-coh y z j + ; (j = i0) f-coh x y i + ; (j = i1) f-coh x z (i k)}) + (f-coh-coh x y z j i) + triHelper₂ (squash₁ s t i) u v + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ (squash₁ s t i) u i₁) v j)) + (triHelper₂ s u v) (triHelper₂ t u v) + (triHelper₂ s t v) i j pathHelper u v j) + (triHelper₂ s t u) refl + (B-gpd v) i + triHelper₂ x ∣₁ (squash₁ t u i) v + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ x ∣₁ (squash₁ t u i) i₁) v j)) + (triHelper₂ x ∣₁ t v) (triHelper₂ x ∣₁ u v) + i j pathHelper x ∣₁ v j) (triHelper₂ t u v) + (triHelper₁ x ∣₁ t u) refl + (B-gpd v) i + triHelper₂ x ∣₁ y ∣₁ (squash₁ u v i) + = isGroupoid→CubeP i i₁ j B (squash₁ (squash₁ x ∣₁ y ∣₁ i₁) (squash₁ u v i) j)) + (triHelper₂ x ∣₁ y ∣₁ u) (triHelper₂ x ∣₁ y ∣₁ v) + (triHelper₁ x ∣₁ u v) (triHelper₁ y ∣₁ u v) + refl i j pathHelper u v i) + (B-gpd v) i + + +module GpdElim (Bgpd : isGroupoid B) where + Bgpd' : isGroupoid' B + Bgpd' = isGroupoid→isGroupoid' Bgpd + + module _ (f : A B) (3kf : 3-Constant f) where + open 3-Constant 3kf + + rec→Gpd : A ∥₁ B + rec→Gpd = elim→Gpd _ B) _ Bgpd) f link coh₁ + + preEquiv₁ : ( A ∥₁ Σ (A B) 3-Constant) Σ (A B) 3-Constant + preEquiv₁ = isoToEquiv (iso fn const _ refl) retr) + where + open 3-Constant + + fn : ( A ∥₁ Σ (A B) 3-Constant) Σ (A B) 3-Constant + fn f .fst x = f x ∣₁ .fst x + fn f .snd .link x y i = f (squash₁ x ∣₁ y ∣₁ i) .snd .link x y i + fn f .snd .coh₁ x y z i j + = f (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j) .snd .coh₁ x y z i j + + retr : retract fn const + retr f i t .fst x = f (squash₁ x ∣₁ t i) .fst x + retr f i t .snd .link x y j + = f (squash₁ (squash₁ x ∣₁ y ∣₁ j) t i) .snd .link x y j + retr f i t .snd .coh₁ x y z + = Bgpd' + k j f (cb k j i0) .snd .coh₁ x y z k j ) + k j f (cb k j i1) .snd .coh₁ x y z k j) + k j f (cb i0 j k) .snd .link x y j) + k j f (cb i1 j k) .snd .link x z j) + _ refl) + k j f (cb j i1 k) .snd .link y z j) + i + where + cb : I I I _ ∥₁ + cb i j k = squash₁ (squash₁ x ∣₁ (squash₁ y ∣₁ z ∣₁ i) j) t k + + e : B Σ (A B) 3-Constant + e b .fst _ = b + e b .snd = record + { link = λ _ _ _ b + ; coh₁ = λ _ _ _ _ _ b + } + + eval : A Σ (A B) 3-Constant B + eval a₀ (g , _) = g a₀ + + module _ where + open 3-Constant + e-eval : ∀(a₀ : A) γ e (eval a₀ γ) γ + e-eval a₀ (g , 3kg) i .fst x = 3kg .link a₀ x i + e-eval a₀ (g , 3kg) i .snd .link x y = λ j 3kg .coh₁ a₀ x y j i + e-eval a₀ (g , 3kg) i .snd .coh₁ x y z + = Bgpd' + _ _ g a₀) + (3kg .coh₁ x y z) + k j 3kg .coh₁ a₀ x y j k) + k j 3kg .coh₁ a₀ x z j k) + _ refl) + k j 3kg .coh₁ a₀ y z j k) + i + + e-isEquiv : A isEquiv (e {A = A}) + e-isEquiv a₀ = isoToIsEquiv (iso e (eval a₀) (e-eval a₀) λ _ refl) + + preEquiv₂ : A ∥₁ B Σ (A B) 3-Constant + preEquiv₂ t = e , rec (isPropIsEquiv e) e-isEquiv t + + trunc→Gpd≃ : ( A ∥₁ B) Σ (A B) 3-Constant + trunc→Gpd≃ = compEquiv (equivΠCod preEquiv₂) preEquiv₁ + +open GpdElim using (rec→Gpd; trunc→Gpd≃) public + +RecHSet : (P : A TypeOfHLevel 2) 3-Constant P A ∥₁ TypeOfHLevel 2 +RecHSet P 3kP = rec→Gpd (isOfHLevelTypeOfHLevel 2) P 3kP + +∥∥-IdempotentL-⊎-≃ : A ∥₁ A′ ∥₁ A A′ ∥₁ +∥∥-IdempotentL-⊎-≃ = isoToEquiv ∥∥-IdempotentL-⊎-Iso + where ∥∥-IdempotentL-⊎-Iso : Iso ( A ∥₁ A′ ∥₁) ( A A′ ∥₁) + Iso.fun ∥∥-IdempotentL-⊎-Iso x = rec squash₁ lem x + where lem : A ∥₁ A′ A A′ ∥₁ + lem (inl x) = map a inl a) x + lem (inr x) = inr x ∣₁ + Iso.inv ∥∥-IdempotentL-⊎-Iso x = map lem x + where lem : A A′ A ∥₁ A′ + lem (inl x) = inl x ∣₁ + lem (inr x) = inr x + Iso.rightInv ∥∥-IdempotentL-⊎-Iso x = squash₁ (Iso.fun ∥∥-IdempotentL-⊎-Iso (Iso.inv ∥∥-IdempotentL-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-⊎-Iso x = squash₁ (Iso.inv ∥∥-IdempotentL-⊎-Iso (Iso.fun ∥∥-IdempotentL-⊎-Iso x)) x + +∥∥-IdempotentL-⊎ : A ∥₁ A′ ∥₁ A A′ ∥₁ +∥∥-IdempotentL-⊎ = ua ∥∥-IdempotentL-⊎-≃ + +∥∥-IdempotentR-⊎-≃ : A A′ ∥₁ ∥₁ A A′ ∥₁ +∥∥-IdempotentR-⊎-≃ = isoToEquiv ∥∥-IdempotentR-⊎-Iso + where ∥∥-IdempotentR-⊎-Iso : Iso ( A A′ ∥₁ ∥₁) ( A A′ ∥₁) + Iso.fun ∥∥-IdempotentR-⊎-Iso x = rec squash₁ lem x + where lem : A A′ ∥₁ A A′ ∥₁ + lem (inl x) = inl x ∣₁ + lem (inr x) = map a inr a) x + Iso.inv ∥∥-IdempotentR-⊎-Iso x = map lem x + where lem : A A′ A A′ ∥₁ + lem (inl x) = inl x + lem (inr x) = inr x ∣₁ + Iso.rightInv ∥∥-IdempotentR-⊎-Iso x = squash₁ (Iso.fun ∥∥-IdempotentR-⊎-Iso (Iso.inv ∥∥-IdempotentR-⊎-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-⊎-Iso x = squash₁ (Iso.inv ∥∥-IdempotentR-⊎-Iso (Iso.fun ∥∥-IdempotentR-⊎-Iso x)) x + +∥∥-IdempotentR-⊎ : A A′ ∥₁ ∥₁ A A′ ∥₁ +∥∥-IdempotentR-⊎ = ua ∥∥-IdempotentR-⊎-≃ + +∥∥-Idempotent-⊎ : {A : Type } {A′ : Type ℓ'} A ∥₁ A′ ∥₁ ∥₁ A A′ ∥₁ +∥∥-Idempotent-⊎ {A = A} {A′} = A ∥₁ A′ ∥₁ ∥₁ ≡⟨ ∥∥-IdempotentR-⊎ + A ∥₁ A′ ∥₁ ≡⟨ ∥∥-IdempotentL-⊎ + A A′ ∥₁ + +∥∥-IdempotentL-×-≃ : A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-IdempotentL-×-≃ = isoToEquiv ∥∥-IdempotentL-×-Iso + where ∥∥-IdempotentL-×-Iso : Iso ( A ∥₁ × A′ ∥₁) ( A × A′ ∥₁) + Iso.fun ∥∥-IdempotentL-×-Iso x = rec squash₁ lem x + where lem : A ∥₁ × A′ A × A′ ∥₁ + lem (a , a′) = map2 a a′ a , a′) a a′ ∣₁ + Iso.inv ∥∥-IdempotentL-×-Iso x = map lem x + where lem : A × A′ A ∥₁ × A′ + lem (a , a′) = a ∣₁ , a′ + Iso.rightInv ∥∥-IdempotentL-×-Iso x = squash₁ (Iso.fun ∥∥-IdempotentL-×-Iso (Iso.inv ∥∥-IdempotentL-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentL-×-Iso x = squash₁ (Iso.inv ∥∥-IdempotentL-×-Iso (Iso.fun ∥∥-IdempotentL-×-Iso x)) x + +∥∥-IdempotentL-× : A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-IdempotentL-× = ua ∥∥-IdempotentL-×-≃ + +∥∥-IdempotentR-×-≃ : A × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-IdempotentR-×-≃ = isoToEquiv ∥∥-IdempotentR-×-Iso + where ∥∥-IdempotentR-×-Iso : Iso ( A × A′ ∥₁ ∥₁) ( A × A′ ∥₁) + Iso.fun ∥∥-IdempotentR-×-Iso x = rec squash₁ lem x + where lem : A × A′ ∥₁ A × A′ ∥₁ + lem (a , a′) = map2 a a′ a , a′) a ∣₁ a′ + Iso.inv ∥∥-IdempotentR-×-Iso x = map lem x + where lem : A × A′ A × A′ ∥₁ + lem (a , a′) = a , a′ ∣₁ + Iso.rightInv ∥∥-IdempotentR-×-Iso x = squash₁ (Iso.fun ∥∥-IdempotentR-×-Iso (Iso.inv ∥∥-IdempotentR-×-Iso x)) x + Iso.leftInv ∥∥-IdempotentR-×-Iso x = squash₁ (Iso.inv ∥∥-IdempotentR-×-Iso (Iso.fun ∥∥-IdempotentR-×-Iso x)) x + +∥∥-IdempotentR-× : A × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-IdempotentR-× = ua ∥∥-IdempotentR-×-≃ + +∥∥-Idempotent-× : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-Idempotent-× {A = A} {A′} = A ∥₁ × A′ ∥₁ ∥₁ ≡⟨ ∥∥-IdempotentR-× + A ∥₁ × A′ ∥₁ ≡⟨ ∥∥-IdempotentL-× + A × A′ ∥₁ + +∥∥-Idempotent-×-≃ : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ ∥₁ A × A′ ∥₁ +∥∥-Idempotent-×-≃ {A = A} {A′} = compEquiv ∥∥-IdempotentR-×-≃ ∥∥-IdempotentL-×-≃ + +∥∥-×-≃ : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-×-≃ {A = A} {A′} = compEquiv (invEquiv (propTruncIdempotent≃ (isProp× isPropPropTrunc isPropPropTrunc))) ∥∥-Idempotent-×-≃ + +∥∥-× : {A : Type } {A′ : Type ℓ'} A ∥₁ × A′ ∥₁ A × A′ ∥₁ +∥∥-× = ua ∥∥-×-≃ + +-- using this we get a convenient recursor/eliminator for binary functions into sets +rec2→Set : {A B C : Type } (Cset : isSet C) + (f : A B C) + (∀ (a a' : A) (b b' : B) f a b f a' b') + A ∥₁ B ∥₁ C +rec2→Set {A = A} {B = B} {C = C} Cset f fconst = curry (g ∥∥-×-≃ .fst) + where + g : A × B ∥₁ C + g = rec→Set Cset (uncurry f) λ x y fconst (fst x) (fst y) (snd x) (snd y) diff --git a/src/docs/Cubical.HITs.PropositionalTruncation.html b/src/docs/Cubical.HITs.PropositionalTruncation.html new file mode 100644 index 0000000..c07e2c1 --- /dev/null +++ b/src/docs/Cubical.HITs.PropositionalTruncation.html @@ -0,0 +1,7 @@ +{-# OPTIONS --safe #-} +module Cubical.HITs.PropositionalTruncation where + +open import Cubical.HITs.PropositionalTruncation.Base public +open import Cubical.HITs.PropositionalTruncation.Properties public + +open import Cubical.HITs.PropositionalTruncation.MagicTrick diff --git a/src/docs/Cubical.Homotopy.Base.html b/src/docs/Cubical.Homotopy.Base.html new file mode 100644 index 0000000..d5420d7 --- /dev/null +++ b/src/docs/Cubical.Homotopy.Base.html @@ -0,0 +1,19 @@ +{-# OPTIONS --safe #-} + +module Cubical.Homotopy.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv.Properties + +private + variable + ℓ' : Level + +_∼_ : {X : Type } {Y : X Type ℓ'} (f g : (x : X) Y x) Type (ℓ-max ℓ') +_∼_ {X = X} f g = (x : X) f x g x + +funExt∼ : {X : Type } {Y : X Type ℓ'} {f g : (x : X) Y x} (H : f g) f g +funExt∼ = funExt + +∼-refl : {X : Type } {Y : X Type ℓ'} {f : (x : X) Y x} f f +∼-refl {f = f} = λ x refl {x = f x} diff --git a/src/docs/Cubical.Induction.WellFounded.html b/src/docs/Cubical.Induction.WellFounded.html new file mode 100644 index 0000000..2d4e3c9 --- /dev/null +++ b/src/docs/Cubical.Induction.WellFounded.html @@ -0,0 +1,47 @@ +{-# OPTIONS --safe #-} + +module Cubical.Induction.WellFounded where + +open import Cubical.Foundations.Prelude + +module _ { ℓ'} {A : Type } (_<_ : A A Type ℓ') where + WFRec : ∀{ℓ''} (A Type ℓ'') A Type _ + WFRec P x = y y < x P y + + data Acc (x : A) : Type (ℓ-max ℓ') where + acc : WFRec Acc x Acc x + + WellFounded : Type _ + WellFounded = x Acc x + + +module _ { ℓ'} {A : Type } {_<_ : A A Type ℓ'} where + isPropAcc : x isProp (Acc _<_ x) + isPropAcc x (acc p) (acc q) + = λ i acc y y<x isPropAcc y (p y y<x) (q y y<x) i) + + isPropWellFounded : isProp (WellFounded _<_) + isPropWellFounded p q i a = isPropAcc a (p a) (q a) i + + access : ∀{x} Acc _<_ x WFRec _<_ (Acc _<_) x + access (acc r) = r + + private + wfi : ∀{ℓ''} {P : A Type ℓ''} + x (wf : Acc _<_ x) + (∀ x (∀ y y < x P y) P x) + P x + wfi x (acc p) e = e x λ y y<x wfi y (p y y<x) e + + module WFI (wf : WellFounded _<_) where + module _ {ℓ''} {P : A Type ℓ''} (e : x (∀ y y < x P y) P x) where + private + wfi-compute : x ax wfi x ax e e x y _ wfi y (wf y) e) + wfi-compute x (acc p) + = λ i e x y y<x wfi y (isPropAcc y (p y y<x) (wf y) i) e) + + induction : x P x + induction x = wfi x (wf x) e + + induction-compute : x induction x (e x λ y _ induction y) + induction-compute x = wfi-compute x (wf x) diff --git a/src/docs/Cubical.Reflection.Base.html b/src/docs/Cubical.Reflection.Base.html new file mode 100644 index 0000000..7a47868 --- /dev/null +++ b/src/docs/Cubical.Reflection.Base.html @@ -0,0 +1,50 @@ +{- + +Some basic utilities for reflection + +-} +{-# OPTIONS --no-exact-split --safe #-} +module Cubical.Reflection.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Data.List.Base +open import Cubical.Data.Nat.Base + +import Agda.Builtin.Reflection as R +open import Agda.Builtin.String + +_>>=_ = R.bindTC +_<|>_ = R.catchTC + +_>>_ : { ℓ'} {A : Type } {B : Type ℓ'} R.TC A R.TC B R.TC B +f >> g = f >>= λ _ g + +infixl 4 _>>=_ _>>_ _<|>_ + +liftTC : { ℓ'} {A : Type } {B : Type ℓ'} (A B) R.TC A R.TC B +liftTC f ta = ta >>= λ a R.returnTC (f a) + +v : R.Term +v n = R.var n [] + +pattern varg t = R.arg (R.arg-info R.visible (R.modality R.relevant R.quantity-ω)) t +pattern harg {q = q} t = R.arg (R.arg-info R.hidden (R.modality R.relevant q)) t +pattern _v∷_ a l = varg a l +pattern _h∷_ a l = harg a l + +infixr 5 _v∷_ _h∷_ + +vlam : String R.Term R.Term +vlam str t = R.lam R.visible (R.abs str t) + +hlam : String R.Term R.Term +hlam str t = R.lam R.hidden (R.abs str t) + +newMeta = R.checkType R.unknown + +makeAuxiliaryDef : String R.Type R.Term R.TC R.Term +makeAuxiliaryDef s ty term = + R.freshName s >>= λ name + R.declareDef (varg name) ty >> + R.defineFun name [ R.clause [] [] term ] >> + R.returnTC (R.def name []) diff --git a/src/docs/Cubical.Reflection.StrictEquiv.html b/src/docs/Cubical.Reflection.StrictEquiv.html new file mode 100644 index 0000000..b1baaf4 --- /dev/null +++ b/src/docs/Cubical.Reflection.StrictEquiv.html @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} +module Cubical.Reflection.StrictEquiv where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv.Base +open import Cubical.Foundations.Isomorphism + +open import Cubical.Data.List.Base +open import Cubical.Data.Unit.Base + +import Agda.Builtin.Reflection as R +open import Cubical.Reflection.Base + +strictEquivClauses : R.Term R.Term List R.Clause +strictEquivClauses f g = + R.clause [] + (R.proj (quote fst) v∷ []) + f + R.clause [] + (R.proj (quote snd) v∷ R.proj (quote equiv-proof) v∷ []) + (R.def (quote strictContrFibers) (g v∷ [])) + [] + +strictEquivTerm : R.Term R.Term R.Term +strictEquivTerm f g = R.pat-lam (strictEquivClauses f g) [] + +strictEquivMacro : { ℓ'} {A : Type } {B : Type ℓ'} + (A B) (B A) R.Term R.TC Unit +strictEquivMacro {A = A} {B} f g hole = + R.quoteTC (A B) >>= λ equivTy + R.checkType hole equivTy >> + R.quoteTC f >>= λ `f` + R.quoteTC g >>= λ `g` + R.unify (strictEquivTerm `f` `g`) hole + +strictIsoToEquivMacro : { ℓ'} {A : Type } {B : Type ℓ'} + Iso A B R.Term R.TC Unit +strictIsoToEquivMacro isom = + strictEquivMacro (Iso.fun isom) (Iso.inv isom) + +-- For use with unquoteDef + +defStrictEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + R.Name (A B) (B A) R.TC Unit +defStrictEquiv idName f g = + R.quoteTC f >>= λ `f` + R.quoteTC g >>= λ `g` + R.defineFun idName (strictEquivClauses `f` `g`) + +defStrictIsoToEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + R.Name Iso A B R.TC Unit +defStrictIsoToEquiv idName isom = + defStrictEquiv idName (Iso.fun isom) (Iso.inv isom) + +-- For use with unquoteDef + +declStrictEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + R.Name (A B) (B A) R.TC Unit +declStrictEquiv {A = A} {B = B} idName f g = + R.quoteTC (A B) >>= λ ty + R.declareDef (varg idName) ty >> + defStrictEquiv idName f g + +declStrictIsoToEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + R.Name Iso A B R.TC Unit +declStrictIsoToEquiv idName isom = + declStrictEquiv idName (Iso.fun isom) (Iso.inv isom) + +macro + -- (f : A → B) → (g : B → A) → (A ≃ B) + -- Assumes that `f` and `g` are inverse up to definitional equality + strictEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + (A B) (B A) R.Term R.TC Unit + strictEquiv = strictEquivMacro + + -- (isom : Iso A B) → (A ≃ B) + -- Assumes that the functions defining `isom` are inverse up to definitional equality + strictIsoToEquiv : { ℓ'} {A : Type } {B : Type ℓ'} + Iso A B R.Term R.TC Unit + strictIsoToEquiv = strictIsoToEquivMacro diff --git a/src/docs/Cubical.Relation.Nullary.Base.html b/src/docs/Cubical.Relation.Nullary.Base.html new file mode 100644 index 0000000..ba2f57a --- /dev/null +++ b/src/docs/Cubical.Relation.Nullary.Base.html @@ -0,0 +1,70 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary.Base where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Functions.Fixpoint + +open import Cubical.Data.Empty as +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + : Level + A : Type + +-- Negation +infix 3 ¬_ + +¬_ : Type Type +¬ A = A + +-- Decidable types (inspired by standard library) +data Dec (P : Type ) : Type where + yes : ( p : P) Dec P + no : (¬p : ¬ P) Dec P + +decRec : { ℓ'} {P : Type } {A : Type ℓ'} (P A) (¬ P A) (Dec P) A +decRec ifyes ifno (yes p) = ifyes p +decRec ifyes ifno (no ¬p) = ifno ¬p + +NonEmpty : Type Type +NonEmpty A = ¬ ¬ A + +Stable : Type Type +Stable A = NonEmpty A A + +-- reexport propositional truncation for uniformity +open Cubical.HITs.PropositionalTruncation.Base + using (∥_∥₁) public + +SplitSupport : Type Type +SplitSupport A = A ∥₁ A + +Collapsible : Type Type +Collapsible A = Σ[ f (A A) ] 2-Constant f + +Populated ⟪_⟫ : Type Type +Populated A = (f : Collapsible A) Fixpoint (f .fst) +⟪_⟫ = Populated + +PStable : Type Type +PStable A = A A + +onAllPaths : (Type Type ) Type Type +onAllPaths S A = (x y : A) S (x y) + +Separated : Type Type +Separated = onAllPaths Stable + +HSeparated : Type Type +HSeparated = onAllPaths SplitSupport + +Collapsible≡ : Type Type +Collapsible≡ = onAllPaths Collapsible + +PStable≡ : Type Type +PStable≡ = onAllPaths PStable + +Discrete : Type Type +Discrete = onAllPaths Dec diff --git a/src/docs/Cubical.Relation.Nullary.Properties.html b/src/docs/Cubical.Relation.Nullary.Properties.html new file mode 100644 index 0000000..5f0a148 --- /dev/null +++ b/src/docs/Cubical.Relation.Nullary.Properties.html @@ -0,0 +1,203 @@ +{-# OPTIONS --safe #-} +{- + +Properties of nullary relations, i.e. structures on types. + +Includes several results from [Notions of Anonymous Existence in Martin-Löf Type Theory](https://lmcs.episciences.org/3217) +by Altenkirch, Coquand, Escardo, Kraus. + +-} +module Cubical.Relation.Nullary.Properties where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Function +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Functions.Fixpoint + +open import Cubical.Data.Empty as +open import Cubical.Data.Sigma.Base using (_×_) + +open import Cubical.Relation.Nullary.Base +open import Cubical.HITs.PropositionalTruncation.Base + +private + variable + : Level + A B : Type + P : A -> Type + +-- Functions with a section preserve discreteness. +sectionDiscrete + : (f : A B) (g : B A) section f g Discrete A Discrete B +sectionDiscrete f g sect dA x y with dA (g x) (g y) +... | yes p = yes (sym (sect x) ∙∙ cong f p ∙∙ sect y) +... | no ¬p = no p ¬p (cong g p)) + +isoPresDiscrete : Iso A B Discrete A Discrete B +isoPresDiscrete e = sectionDiscrete fun inv rightInv + where open Iso e + +EquivPresDiscrete : { ℓ'}{A : Type } {B : Type ℓ'} A B + Discrete A Discrete B +EquivPresDiscrete e = isoPresDiscrete (equivToIso e) + +isProp¬ : (A : Type ) isProp (¬ A) +isProp¬ A p q i x = isProp⊥ (p x) (q x) i + +Stable¬ : Stable (¬ A) +Stable¬ ¬¬¬a a = ¬¬¬a ¬¬a + where + ¬¬a = λ ¬a ¬a a + +StableΠ : (∀ x Stable (P x)) -> Stable (∀ x P x) +StableΠ Ps e x = Ps x λ k e λ f k (f x) + +Stable→ : Stable B Stable (A B) +Stable→ Bs = StableΠ _ Bs) + +Stable× : Stable A -> Stable B -> Stable (A × B) +Stable× As Bs e = λ where + .fst As λ k e (k fst) + .snd Bs λ k e (k snd) + +fromYes : A Dec A A +fromYes _ (yes a) = a +fromYes a (no _) = a + +discreteDec : (Adis : Discrete A) Discrete (Dec A) +discreteDec Adis (yes p) (yes p') = decideYes (Adis p p') -- TODO: monad would simply stuff + where + decideYes : Dec (p p') Dec (yes p yes p') + decideYes (yes eq) = yes (cong yes eq) + decideYes (no ¬eq) = no λ eq ¬eq (cong (fromYes p) eq) +discreteDec Adis (yes p) (no ¬p) = ⊥.rec (¬p p) +discreteDec Adis (no ¬p) (yes p) = ⊥.rec (¬p p) +discreteDec {A = A} Adis (no ¬p) (no ¬p') = yes (cong no (isProp¬ A ¬p ¬p')) + +isPropDec : (Aprop : isProp A) isProp (Dec A) +isPropDec Aprop (yes a) (yes a') = cong yes (Aprop a a') +isPropDec Aprop (yes a) (no ¬a) = ⊥.rec (¬a a) +isPropDec Aprop (no ¬a) (yes a) = ⊥.rec (¬a a) +isPropDec {A = A} Aprop (no ¬a) (no ¬a') = cong no (isProp¬ A ¬a ¬a') + +mapDec : {B : Type } (A B) (¬ A ¬ B) Dec A Dec B +mapDec f _ (yes p) = yes (f p) +mapDec _ f (no ¬p) = no (f ¬p) + +EquivPresDec : { ℓ'}{A : Type } {B : Type ℓ'} A B + Dec A Dec B +EquivPresDec p = mapDec (p .fst) f f invEq p) + +¬→¬∥∥ : ¬ A ¬ A ∥₁ +¬→¬∥∥ ¬p a ∣₁ = ¬p a +¬→¬∥∥ ¬p (squash₁ x y i) = isProp⊥ (¬→¬∥∥ ¬p x) (¬→¬∥∥ ¬p y) i + +Dec∥∥ : Dec A Dec A ∥₁ +Dec∥∥ = mapDec ∣_∣₁ ¬→¬∥∥ + +-- we have the following implications +-- X ── ∣_∣ ─→ ∥ X ∥ +-- ∥ X ∥ ── populatedBy ─→ ⟪ X ⟫ +-- ⟪ X ⟫ ── notEmptyPopulated ─→ NonEmpty X + +-- reexport propositional truncation for uniformity +open Cubical.HITs.PropositionalTruncation.Base + +populatedBy : A ∥₁ A +populatedBy {A = A} a (f , fIsConst) = h a where + h : A ∥₁ Fixpoint f + h a ∣₁ = f a , fIsConst (f a) a + h (squash₁ a b i) = 2-Constant→isPropFixpoint f fIsConst (h a) (h b) i + +notEmptyPopulated : A NonEmpty A +notEmptyPopulated {A = A} pop u = u (fixpoint (pop (h , hIsConst))) where + h : A A + h a = ⊥.elim (u a) + hIsConst : x y h x h y + hIsConst x y i = ⊥.elim (isProp⊥ (u x) (u y) i) + +-- these implications induce the following for different kinds of stability, gradually weakening the assumption +Dec→Stable : Dec A Stable A +Dec→Stable (yes x) = λ _ x +Dec→Stable (no x) = λ f ⊥.elim (f x) + +Stable→PStable : Stable A PStable A +Stable→PStable st = st notEmptyPopulated + +PStable→SplitSupport : PStable A SplitSupport A +PStable→SplitSupport pst = pst populatedBy + +-- Although SplitSupport and Collapsible are not properties, their path versions, HSeparated and Collapsible≡, are. +-- Nevertheless they are logically equivalent +SplitSupport→Collapsible : SplitSupport A Collapsible A +SplitSupport→Collapsible {A = A} hst = h , hIsConst where + h : A A + h p = hst p ∣₁ + hIsConst : 2-Constant h + hIsConst p q i = hst (squash₁ p ∣₁ q ∣₁ i) + +Collapsible→SplitSupport : Collapsible A SplitSupport A +Collapsible→SplitSupport f x = fixpoint (populatedBy x f) + +HSeparated→Collapsible≡ : HSeparated A Collapsible≡ A +HSeparated→Collapsible≡ hst x y = SplitSupport→Collapsible (hst x y) + +Collapsible≡→HSeparated : Collapsible≡ A HSeparated A +Collapsible≡→HSeparated col x y = Collapsible→SplitSupport (col x y) + +-- stability of path space under truncation or path collapsability are necessary and sufficient properties +-- for a a type to be a set. +Collapsible≡→isSet : Collapsible≡ A isSet A +Collapsible≡→isSet {A = A} col a b p q = p≡q where + f : (x : A) a x a x + f x = col a x .fst + fIsConst : (x : A) (p q : a x) f x p f x q + fIsConst x = col a x .snd + rem : (p : a b) PathP i a p i) (f a refl) (f b p) + rem p j = f (p j) i p (i j)) + p≡q : p q + p≡q j i = hcomp k λ { (i = i0) f a refl k + ; (i = i1) fIsConst b p q j k + ; (j = i0) rem p i k + ; (j = i1) rem q i k }) a + +HSeparated→isSet : HSeparated A isSet A +HSeparated→isSet = Collapsible≡→isSet HSeparated→Collapsible≡ + +isSet→HSeparated : isSet A HSeparated A +isSet→HSeparated setA x y = extract where + extract : x y ∥₁ x y + extract p ∣₁ = p + extract (squash₁ p q i) = setA x y (extract p) (extract q) i + +-- by the above more sufficient conditions to inhibit isSet A are given +PStable≡→HSeparated : PStable≡ A HSeparated A +PStable≡→HSeparated pst x y = PStable→SplitSupport (pst x y) + +PStable≡→isSet : PStable≡ A isSet A +PStable≡→isSet = HSeparated→isSet PStable≡→HSeparated + +Separated→PStable≡ : Separated A PStable≡ A +Separated→PStable≡ st x y = Stable→PStable (st x y) + +Separated→isSet : Separated A isSet A +Separated→isSet = PStable≡→isSet Separated→PStable≡ + +SeparatedΠ : (∀ x Separated (P x)) -> Separated ((x : A) -> P x) +SeparatedΠ Ps f g e i x = Ps x (f x) (g x) k e (k cong f f x))) i + +Separated→ : Separated B -> Separated (A B) +Separated→ Bs = SeparatedΠ _ Bs) + +Separated× : Separated A -> Separated B -> Separated (A × B) +Separated× As Bs p q e i = λ where + .fst As (fst p) (fst q) k e λ r k (cong fst r)) i + .snd Bs (snd p) (snd q) k e λ r k (cong snd r)) i + +-- Proof of Hedberg's theorem: a type with decidable equality is an h-set +Discrete→Separated : Discrete A Separated A +Discrete→Separated d x y = Dec→Stable (d x y) + +Discrete→isSet : Discrete A isSet A +Discrete→isSet = Separated→isSet Discrete→Separated diff --git a/src/docs/Cubical.Relation.Nullary.html b/src/docs/Cubical.Relation.Nullary.html new file mode 100644 index 0000000..c1ff6cf --- /dev/null +++ b/src/docs/Cubical.Relation.Nullary.html @@ -0,0 +1,5 @@ +{-# OPTIONS --safe #-} +module Cubical.Relation.Nullary where + +open import Cubical.Relation.Nullary.Base public +open import Cubical.Relation.Nullary.Properties public diff --git a/src/docs/Cubical.Structures.Axioms.html b/src/docs/Cubical.Structures.Axioms.html new file mode 100644 index 0000000..9cb1f2b --- /dev/null +++ b/src/docs/Cubical.Structures.Axioms.html @@ -0,0 +1,69 @@ +{- + +Add axioms (i.e., propositions) to a structure S without changing the definition of structured equivalence. + +X ↦ Σ[ s ∈ S X ] (P X s) where (P X s) is a proposition for all X and s. + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Axioms where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Function +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Path +open import Cubical.Foundations.SIP +open import Cubical.Data.Sigma + +private + variable + ℓ₁ ℓ₁' ℓ₂ : Level + +AxiomsStructure : (S : Type Type ℓ₁) + (axioms : (X : Type ) S X Type ℓ₂) + Type Type (ℓ-max ℓ₁ ℓ₂) +AxiomsStructure S axioms X = Σ[ s S X ] (axioms X s) + +AxiomsEquivStr : {S : Type Type ℓ₁} (ι : StrEquiv S ℓ₁') + (axioms : (X : Type ) S X Type ℓ₂) + StrEquiv (AxiomsStructure S axioms) ℓ₁' +AxiomsEquivStr ι axioms (X , (s , a)) (Y , (t , b)) e = ι (X , s) (Y , t) e + +axiomsUnivalentStr : {S : Type Type ℓ₁} + (ι : (A B : TypeWithStr S) A .fst B .fst Type ℓ₁') + {axioms : (X : Type ) S X Type ℓ₂} + (axioms-are-Props : (X : Type ) (s : S X) isProp (axioms X s)) + (θ : UnivalentStr S ι) + UnivalentStr (AxiomsStructure S axioms) (AxiomsEquivStr ι axioms) +axiomsUnivalentStr {S = S} ι {axioms = axioms} axioms-are-Props θ {X , s , a} {Y , t , b} e = + ι (X , s) (Y , t) e + ≃⟨ θ e + PathP i S (ua e i)) s t + ≃⟨ invEquiv (Σ-contractSnd λ _ isOfHLevelPathP' 0 (axioms-are-Props _ _) _ _) + Σ[ p PathP i S (ua e i)) s t ] PathP i axioms (ua e i) (p i)) a b + ≃⟨ ΣPath≃PathΣ + PathP i AxiomsStructure S axioms (ua e i)) (s , a) (t , b) + + +inducedStructure : {S : Type Type ℓ₁} + {ι : (A B : TypeWithStr S) A .fst B .fst Type ℓ₁'} + (θ : UnivalentStr S ι) + {axioms : (X : Type ) S X Type ℓ₂} + (A : TypeWithStr (AxiomsStructure S axioms)) (B : TypeWithStr S) + (typ A , str A .fst) ≃[ ι ] B + TypeWithStr (AxiomsStructure S axioms) +inducedStructure θ {axioms} A B eqv = + B .fst , B .snd , subst (uncurry axioms) (sip θ _ _ eqv) (A .snd .snd) + +transferAxioms : {S : Type Type ℓ₁} + {ι : (A B : TypeWithStr S) A .fst B .fst Type ℓ₁'} + (θ : UnivalentStr S ι) + {axioms : (X : Type ) S X Type ℓ₂} + (A : TypeWithStr (AxiomsStructure S axioms)) (B : TypeWithStr S) + (typ A , str A .fst) ≃[ ι ] B + axioms (fst B) (snd B) +transferAxioms θ {axioms} A B eqv = + subst (uncurry axioms) (sip θ _ _ eqv) (A .snd .snd) diff --git a/src/docs/Cubical.Structures.Pointed.html b/src/docs/Cubical.Structures.Pointed.html new file mode 100644 index 0000000..f42c80d --- /dev/null +++ b/src/docs/Cubical.Structures.Pointed.html @@ -0,0 +1,59 @@ +{- + +Pointed structure: X ↦ X + +-} +{-# OPTIONS --safe #-} +module Cubical.Structures.Pointed where + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.SIP + +open import Cubical.Foundations.Pointed.Base + +private + variable + : Level + +-- Structured isomorphisms + +PointedStructure : Type Type +PointedStructure X = X + +PointedEquivStr : StrEquiv PointedStructure +PointedEquivStr A B f = equivFun f (pt A) pt B + +pointedUnivalentStr : UnivalentStr {} PointedStructure PointedEquivStr +pointedUnivalentStr f = invEquiv (ua-ungluePath-Equiv f) + +pointedSIP : (A B : Pointed ) A ≃[ PointedEquivStr ] B (A B) +pointedSIP = SIP pointedUnivalentStr + +pointed-sip : (A B : Pointed ) A ≃[ PointedEquivStr ] B (A B) +pointed-sip A B = equivFun (pointedSIP A B) -- ≡ λ (e , p) i → ua e i , ua-gluePath e p i + +pointed-sip-idEquiv∙ : (A : Pointed ) pointed-sip A A (idEquiv∙ A) refl +fst (pointed-sip-idEquiv∙ A i j) = uaIdEquiv i j +snd (pointed-sip-idEquiv∙ A i j) = glue {φ = i ~ j j} _ pt A) (pt A) + +{- + The following terms have huge normal forms, so they are abstract to avoid + type checking speed problems, for example in + + Cubical.Homotopy.HSpace +-} +abstract + pointed-sip⁻ : (A B : Pointed ) (A B) A ≃[ PointedEquivStr ] B + pointed-sip⁻ A B = invEq (pointedSIP A B) + + pointed-sip⁻-refl : (A : Pointed ) pointed-sip⁻ A A refl idEquiv∙ A + pointed-sip⁻-refl A = sym (invEq (equivAdjointEquiv (pointedSIP A A)) (pointed-sip-idEquiv∙ A)) + +pointedEquivAction : EquivAction {} PointedStructure +pointedEquivAction e = e + +pointedTransportStr : TransportStr {} pointedEquivAction +pointedTransportStr e s = sym (transportRefl _) diff --git a/src/docs/Cubical.Tactics.NatSolver.EvalHom.html b/src/docs/Cubical.Tactics.NatSolver.EvalHom.html new file mode 100644 index 0000000..c50ba31 --- /dev/null +++ b/src/docs/Cubical.Tactics.NatSolver.EvalHom.html @@ -0,0 +1,196 @@ +{-# OPTIONS --safe #-} +module Cubical.Tactics.NatSolver.EvalHom where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat +open import Cubical.Data.FinData +open import Cubical.Data.Vec + +open import Cubical.Tactics.NatSolver.HornerForms + +private + variable + : Level + +module HomomorphismProperties where + open IteratedHornerOperations + + evalHom+0 : {n : } (P : IteratedHornerForms n) (xs : Vec n) + eval (0ₕ +ₕ P) xs eval P xs + evalHom+0 (const x) [] = refl + evalHom+0 _ (x xs) = refl + + eval0H : {n : } (xs : Vec n) + eval 0ₕ xs 0 + eval0H [] = refl + eval0H (x xs) = refl + + eval1ₕ : {n : } (xs : Vec n) + eval 1ₕ xs 1 + eval1ₕ [] = refl + eval1ₕ (x xs) = + eval 1ₕ (x xs) ≡⟨ refl + eval (0H ·X+ 1ₕ) (x xs) ≡⟨ refl + eval 0H (x xs) · x + eval 1ₕ xs ≡⟨ cong u u · x + eval 1ₕ xs) + (eval0H (x xs)) + 0 · x + eval 1ₕ xs ≡⟨ cong u 0 · x + u) + (eval1ₕ xs) + 1 + + + +ShufflePairs : (a b c d : ) + (a + b) + (c + d) (a + c) + (b + d) + +ShufflePairs a b c d = + (a + b) + (c + d) ≡⟨ +-assoc (a + b) c d + ((a + b) + c) + d ≡⟨ cong u u + d) (sym (+-assoc a b c)) + (a + (b + c)) + d ≡⟨ cong u (a + u) + d) (+-comm b c) + (a + (c + b)) + d ≡⟨ cong u u + d) (+-assoc a c b) + ((a + c) + b) + d ≡⟨ sym (+-assoc (a + c) b d) + (a + c) + (b + d) + + +Homeval : + {n : } (P Q : IteratedHornerForms n) (xs : Vec n) + eval (P +ₕ Q) xs (eval P xs) + (eval Q xs) + +Homeval (const x) (const y) [] = refl + +Homeval 0H Q xs = + eval (0H +ₕ Q) xs ≡⟨ refl + 0 + eval Q xs ≡⟨ cong u u + eval Q xs) (sym (eval0H xs)) + eval 0H xs + eval Q xs + +Homeval (P ·X+ Q) 0H xs = + eval ((P ·X+ Q) +ₕ 0H) xs ≡⟨ refl + eval (P ·X+ Q) xs ≡⟨ sym (+-zero _) + eval (P ·X+ Q) xs + 0 ≡⟨ cong u eval (P ·X+ Q) xs + u) + (sym (eval0H xs)) + eval (P ·X+ Q) xs + eval 0H xs + +Homeval (P ·X+ Q) (S ·X+ T) (x xs) = + eval ((P ·X+ Q) +ₕ (S ·X+ T)) (x xs) + ≡⟨ refl + eval ((P +ₕ S) ·X+ (Q +ₕ T)) (x xs) + ≡⟨ refl + (eval (P +ₕ S) (x xs)) · x + eval (Q +ₕ T) xs + ≡⟨ cong u (eval (P +ₕ S) (x xs)) · x + u) (+Homeval Q T xs) + (eval (P +ₕ S) (x xs)) · x + (eval Q xs + eval T xs) + ≡⟨ cong u u · x + (eval Q xs + eval T xs)) (+Homeval P S (x xs)) + (eval P (x xs) + eval S (x xs)) · x + + (eval Q xs + eval T xs) + ≡⟨ cong u u + (eval Q xs + eval T xs)) + (sym (·-distribʳ (eval P (x xs)) (eval S (x xs)) x)) + (eval P (x xs)) · x + (eval S (x xs)) · x + + (eval Q xs + eval T xs) + ≡⟨ +ShufflePairs ((eval P (x xs)) · x) ((eval S (x xs)) · x) (eval Q xs) (eval T xs) + ((eval P (x xs)) · x + eval Q xs) + + ((eval S (x xs)) · x + eval T xs) + + + ⋆Homeval : {n : } + (r : IteratedHornerForms n) + (P : IteratedHornerForms (ℕ.suc n)) (x : ) (xs : Vec n) + eval (r P) (x xs) eval r xs · eval P (x xs) + + + ⋆0LeftAnnihilates : + {n : } (P : IteratedHornerForms (ℕ.suc n)) (xs : Vec (ℕ.suc n)) + eval (0ₕ P) xs 0 + + ·Homeval : {n : } (P Q : IteratedHornerForms n) (xs : Vec n) + eval (P ·ₕ Q) xs (eval P xs) · (eval Q xs) + + ⋆0LeftAnnihilates 0H xs = eval0H xs + ⋆0LeftAnnihilates (P ·X+ Q) (x xs) = + eval (0ₕ (P ·X+ Q)) (x xs) ≡⟨ refl + eval ((0ₕ P) ·X+ (0ₕ ·ₕ Q)) (x xs) ≡⟨ refl + (eval (0ₕ P) (x xs)) · x + eval (0ₕ ·ₕ Q) xs + ≡⟨ cong u (u · x) + eval (0ₕ ·ₕ Q) _) (⋆0LeftAnnihilates P (x xs)) + 0 · x + eval (0ₕ ·ₕ Q) xs + ≡⟨ ·Homeval 0ₕ Q _ + eval 0ₕ xs · eval Q xs + ≡⟨ cong u u · eval Q xs) (eval0H xs) + 0 · eval Q xs + + ⋆Homeval r 0H x xs = + eval (r 0H) (x xs) ≡⟨ refl + 0 ≡⟨ 0≡m·0 (eval r xs) + eval r xs · 0 ≡⟨ refl + eval r xs · eval 0H (x xs) + ⋆Homeval r (P ·X+ Q) x xs = + eval (r (P ·X+ Q)) (x xs) ≡⟨ refl + eval ((r P) ·X+ (r ·ₕ Q)) (x xs) ≡⟨ refl + (eval (r P) (x xs)) · x + eval (r ·ₕ Q) xs + ≡⟨ cong u u · x + eval (r ·ₕ Q) xs) (⋆Homeval r P x xs) + (eval r xs · eval P (x xs)) · x + eval (r ·ₕ Q) xs + ≡⟨ cong u (eval r xs · eval P (x xs)) · x + u) (·Homeval r Q xs) + (eval r xs · eval P (x xs)) · x + eval r xs · eval Q xs + ≡⟨ cong u u + eval r xs · eval Q xs) (sym (·-assoc (eval r xs) (eval P (x xs)) x)) + eval r xs · (eval P (x xs) · x) + eval r xs · eval Q xs + ≡⟨ ·-distribˡ (eval r xs) ((eval P (x xs) · x)) (eval Q xs) + eval r xs · ((eval P (x xs) · x) + eval Q xs) + ≡⟨ refl + eval r xs · eval (P ·X+ Q) (x xs) + + combineCases : + {n : } (Q : IteratedHornerForms n) (P S : IteratedHornerForms (ℕ.suc n)) + (xs : Vec (ℕ.suc n)) + eval ((P ·X+ Q) ·ₕ S) xs eval (((P ·ₕ S) ·X+ 0ₕ) +ₕ (Q S)) xs + combineCases Q P S (x xs) with (P ·ₕ S) + ... | 0H = + eval (Q S) (x xs) ≡⟨ refl + 0 + eval (Q S) (x xs) ≡⟨ cong u u + eval (Q S) (x xs)) lemma + eval (0H ·X+ 0ₕ) (x xs) + + eval (Q S) (x xs) ≡⟨ sym (+Homeval + (0H ·X+ 0ₕ) (Q S) (x xs)) + eval ((0H ·X+ 0ₕ) +ₕ (Q S)) (x xs) + where lemma : 0 eval (0H ·X+ 0ₕ) (x xs) + lemma = 0 + ≡⟨ refl + 0 + 0 + ≡⟨ cong u u + 0) refl + 0 · x + 0 + ≡⟨ cong u 0 · x + u) (sym (eval0H xs)) + 0 · x + eval 0ₕ xs + ≡⟨ cong u u · x + eval 0ₕ xs) (sym (eval0H (x xs))) + eval 0H (x xs) · x + eval 0ₕ xs + ≡⟨ refl + eval (0H ·X+ 0ₕ) (x xs) + ... | (_ ·X+ _) = refl + + ·Homeval (const x) (const y) [] = refl + ·Homeval 0H Q xs = + eval (0H ·ₕ Q) xs ≡⟨ eval0H xs + 0 ≡⟨ refl + 0 · eval Q xs ≡⟨ cong u u · eval Q xs) (sym (eval0H xs)) + eval 0H xs · eval Q xs + ·Homeval (P ·X+ Q) S (x xs) = + eval ((P ·X+ Q) ·ₕ S) (x xs) + ≡⟨ combineCases Q P S (x xs) + eval (((P ·ₕ S) ·X+ 0ₕ) +ₕ (Q S)) (x xs) + ≡⟨ +Homeval ((P ·ₕ S) ·X+ 0ₕ) (Q S) (x xs) + eval ((P ·ₕ S) ·X+ 0ₕ) (x xs) + eval (Q S) (x xs) + ≡⟨ refl + (eval (P ·ₕ S) (x xs) · x + eval 0ₕ xs) + + eval (Q S) (x xs) + ≡⟨ cong u u + eval (Q S) (x xs)) + ((eval (P ·ₕ S) (x xs) · x + eval 0ₕ xs) + ≡⟨ cong u eval (P ·ₕ S) (x xs) · x + u) (eval0H xs) + (eval (P ·ₕ S) (x xs) · x + 0) + ≡⟨ +-zero _ + (eval (P ·ₕ S) (x xs) · x) + ≡⟨ cong u u · x) (·Homeval P S (x xs)) + ((eval P (x xs) · eval S (x xs)) · x) + ≡⟨ sym (·-assoc (eval P (x xs)) (eval S (x xs)) x) + (eval P (x xs) · (eval S (x xs) · x)) + ≡⟨ cong u eval P (x xs) · u) (·-comm _ x) + (eval P (x xs) · (x · eval S (x xs))) + ≡⟨ ·-assoc (eval P (x xs)) x (eval S (x xs)) + (eval P (x xs) · x) · eval S (x xs) + ) + (eval P (x xs) · x) · eval S (x xs) + + eval (Q S) (x xs) + ≡⟨ cong u (eval P (x xs) · x) · eval S (x xs) + u) + (⋆Homeval Q S x xs) + (eval P (x xs) · x) · eval S (x xs) + + eval Q xs · eval S (x xs) + ≡⟨ ·-distribʳ (eval P (x xs) · x) (eval Q xs) (eval S (x xs)) + ((eval P (x xs) · x) + eval Q xs) · eval S (x xs) + ≡⟨ refl + eval (P ·X+ Q) (x xs) · eval S (x xs) diff --git a/src/docs/Cubical.Tactics.NatSolver.HornerForms.html b/src/docs/Cubical.Tactics.NatSolver.HornerForms.html new file mode 100644 index 0000000..7c2e6b8 --- /dev/null +++ b/src/docs/Cubical.Tactics.NatSolver.HornerForms.html @@ -0,0 +1,100 @@ +{-# OPTIONS --safe #-} +module Cubical.Tactics.NatSolver.HornerForms where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.Nat hiding (isZero) +open import Cubical.Data.FinData +open import Cubical.Data.Vec +open import Cubical.Data.Bool using (Bool; true; false; if_then_else_) + +private + variable + : Level + +{- + This defines the type of multivariate Polynomials over ℕ. + The construction is based on the algebraic fact + + ℕ[X₀][X₁]⋯[Xₙ] ≅ ℕ[X₀,⋯,Xₙ] + + BUT: Contrary to algebraic convetions, we will give 'Xₙ' the lowest index + in the definition of 'Variable' below. So if 'Variable n k' is identified + with 'Xₖ', what we construct should rather be denoted with + + ℕ[Xₙ][Xₙ₋₁]⋯[X₀] + + or, to be precise about the evaluation order: + + (⋯((ℕ[Xₙ])[Xₙ₋₁])⋯)[X₀] + +-} + +data IteratedHornerForms : Type ℓ-zero where + const : IteratedHornerForms ℕ.zero + 0H : {n : } IteratedHornerForms (ℕ.suc n) + _·X+_ : {n : } IteratedHornerForms (ℕ.suc n) IteratedHornerForms n + IteratedHornerForms (ℕ.suc n) + +eval : {n : } (P : IteratedHornerForms n) + Vec n +eval (const r) [] = r +eval 0H (_ _) = 0 +eval (P ·X+ Q) (x xs) = + (eval P (x xs)) · x + eval Q xs + +module IteratedHornerOperations where + + private + 1H' : (n : ) IteratedHornerForms n + 1H' ℕ.zero = const 1 + 1H' (ℕ.suc n) = 0H ·X+ 1H' n + + 0H' : (n : ) IteratedHornerForms n + 0H' ℕ.zero = const 0 + 0H' (ℕ.suc n) = 0H + + 1ₕ : {n : } IteratedHornerForms n + 1ₕ {n = n} = 1H' n + + 0ₕ : {n : } IteratedHornerForms n + 0ₕ {n = n} = 0H' n + + X : (n : ) (k : Fin n) IteratedHornerForms n + X (ℕ.suc m) zero = 1ₕ ·X+ 0ₕ + X (ℕ.suc m) (suc k) = 0ₕ ·X+ X m k + + _+ₕ_ : {n : } IteratedHornerForms n IteratedHornerForms n + IteratedHornerForms n + (const r) +ₕ (const s) = const (r + s) + 0H +ₕ Q = Q + (P ·X+ r) +ₕ 0H = P ·X+ r + (P ·X+ r) +ₕ (Q ·X+ s) = (P +ₕ Q) ·X+ (r +ₕ s) + + isZero : {n : } IteratedHornerForms (ℕ.suc n) + Bool + isZero 0H = true + isZero (P ·X+ P₁) = false + + _⋆_ : {n : } IteratedHornerForms n IteratedHornerForms (ℕ.suc n) + IteratedHornerForms (ℕ.suc n) + _·ₕ_ : {n : } IteratedHornerForms n IteratedHornerForms n + IteratedHornerForms n + r 0H = 0H + r (P ·X+ Q) = (r P) ·X+ (r ·ₕ Q) + + const x ·ₕ const y = const (x · y) + 0H ·ₕ Q = 0H + (P ·X+ Q) ·ₕ S = + let + z = (P ·ₕ S) + in if (isZero z) + then (Q S) + else (z ·X+ 0ₕ) +ₕ (Q S) + +Variable : (n : ) (k : Fin n) IteratedHornerForms n +Variable n k = IteratedHornerOperations.X n k + +Constant : (n : ) (r : ) IteratedHornerForms n +Constant ℕ.zero r = const r +Constant (ℕ.suc n) r = IteratedHornerOperations.0ₕ ·X+ Constant n r diff --git a/src/docs/Cubical.Tactics.NatSolver.NatExpression.html b/src/docs/Cubical.Tactics.NatSolver.NatExpression.html new file mode 100644 index 0000000..4c418fb --- /dev/null +++ b/src/docs/Cubical.Tactics.NatSolver.NatExpression.html @@ -0,0 +1,28 @@ +{-# OPTIONS --safe #-} +module Cubical.Tactics.NatSolver.NatExpression where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order using (zero-≤) +open import Cubical.Data.Vec.Base + +infixl 6 _+'_ +infixl 7 _·'_ + +-- Expression in a ring on A with n variables +data Expr (n : ) : Type ℓ-zero where + K : Expr n + : Fin n Expr n + _+'_ : Expr n Expr n Expr n + _·'_ : Expr n Expr n Expr n + +module Eval where + open import Cubical.Data.Vec + + ⟦_⟧ : {n} Expr n Vec n + K r v = r + k v = lookup k v + x +' y v = x v + y v + x ·' y v = x v · y v diff --git a/src/docs/Cubical.Tactics.NatSolver.Reflection.html b/src/docs/Cubical.Tactics.NatSolver.Reflection.html new file mode 100644 index 0000000..f4d9867 --- /dev/null +++ b/src/docs/Cubical.Tactics.NatSolver.Reflection.html @@ -0,0 +1,145 @@ +{-# OPTIONS --safe #-} +{- + This is inspired by/copied from: + https://github.com/agda/agda-stdlib/blob/master/src/Tactic/MonoidSolver.agda + and the 1lab. + + Boilerplate code for calling the ring solver is constructed automatically + with agda's reflection features. +-} +module Cubical.Tactics.NatSolver.Reflection where + +open import Cubical.Foundations.Prelude hiding (Type) +open import Cubical.Functions.Logic + +open import Agda.Builtin.Reflection hiding (Type) +open import Agda.Builtin.String + +open import Cubical.Reflection.Base + +open import Cubical.Data.Maybe +open import Cubical.Data.Sigma +open import Cubical.Data.List +open import Cubical.Data.Nat +open import Cubical.Data.Bool +open import Cubical.Data.Bool.SwitchStatement +open import Cubical.Data.Vec using (Vec) renaming ([] to emptyVec; _∷_ to _∷vec_) public + +open import Cubical.Tactics.NatSolver.NatExpression +open import Cubical.Tactics.NatSolver.Solver + +open import Cubical.Tactics.Reflection +open import Cubical.Tactics.Reflection.Variables +open import Cubical.Tactics.Reflection.Utilities + +open EqualityToNormalform renaming (solve to natSolve) +private + variable + : Level + + private + solverCallAsTerm : Arg Term Term Term Term + solverCallAsTerm varList lhs rhs = + def + (quote natSolve) + (varg lhs varg rhs + varList + varg (def (quote refl) []) []) + + solverCallWithVars : Vars Term Term Term + solverCallWithVars n vars lhs rhs = + solverCallAsTerm (variableList vars) lhs rhs + where + variableList : Vars Arg Term + variableList [] = varg (con (quote emptyVec) []) + variableList (t ts) + = varg (con (quote _∷vec_) (t v∷ (variableList ts) [])) + +module pr {n : } where + 0' : Expr n + 0' = K 0 + + 1' : Expr n + 1' = K 1 + +module NatSolverReflection where + open pr + + buildExpression : Term TC (Template × Vars) + + op2 : Name Term Term TC (Template × Vars) + op2 op x y = do + r1 buildExpression x + r2 buildExpression y + returnTC ((λ ass con op (fst r1 ass v∷ fst r2 ass v∷ [])) , + appendWithoutRepetition (snd r1) (snd r2)) + + `_·_` : List (Arg Term) TC (Template × Vars) + `_·_` (_ h∷ xs) = `_·_` xs + `_·_` (x v∷ y v∷ []) = op2 (quote _·'_) x y + `_·_` ts = errorOut ts + + `_+_` : List (Arg Term) TC (Template × Vars) + `_+_` (_ h∷ xs) = `_+_` xs + `_+_` (x v∷ y v∷ []) = op2 (quote _+'_) x y + `_+_` ts = errorOut ts + + `1+_` : List (Arg Term) TC (Template × Vars) + `1+_` (x v∷ []) = do + r1 buildExpression x + returnTC ((λ ass con (quote _+'_) ((def (quote 1') []) v∷ fst r1 ass v∷ [])) , + snd r1) + `1+_` ts = errorOut ts + + K' : List (Arg Term) TC (Template × Vars) + K' xs = returnTC ((λ _ con (quote K) xs) , []) + + polynomialVariable : Maybe Term + polynomialVariable (just n) = con (quote ) (finiteNumberAsTerm (just n) v∷ []) + polynomialVariable nothing = unknown + + buildExpression v@(var _ _) = + returnTC ((λ ass polynomialVariable (ass v)) , + v []) + buildExpression t@(lit n) = K' (t v∷ []) + buildExpression t@(def n xs) = + switch (n ==_) cases + case (quote _·_) `_·_` xs break + case (quote _+_) `_+_` xs break + default⇒ (K' xs) + buildExpression t@(con n xs) = + switch (n ==_) cases + case (quote suc) `1+_` xs break + default⇒ (K' xs) + buildExpression t = errorOut' t + + toNatExpression : Term × Term TC (Term × Term × Vars) + toNatExpression (lhs , rhs) = do + r1 buildExpression lhs + r2 buildExpression rhs + vars returnTC (appendWithoutRepetition (snd r1) (snd r2)) + returnTC ( + let ass : VarAss + ass n = indexOf n vars + in (fst r1 ass , fst r2 ass , vars )) + +private + + solve!-macro : Term TC Unit + solve!-macro hole = + do + goal inferType hole >>= normalise + + just (lhs , rhs) get-boundary goal + where + nothing + typeError(strErr "The NatSolver failed to parse the goal " + termErr goal []) + + (lhs' , rhs' , vars) NatSolverReflection.toNatExpression (lhs , rhs) + let solution = solverCallWithVars (length vars) vars lhs' rhs' + unify hole solution + +macro + solveℕ! : Term TC _ + solveℕ! = solve!-macro diff --git a/src/docs/Cubical.Tactics.NatSolver.Solver.html b/src/docs/Cubical.Tactics.NatSolver.Solver.html new file mode 100644 index 0000000..fb8dce4 --- /dev/null +++ b/src/docs/Cubical.Tactics.NatSolver.Solver.html @@ -0,0 +1,123 @@ +{-# OPTIONS --safe #-} +module Cubical.Tactics.NatSolver.Solver where + +open import Cubical.Foundations.Prelude + +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Nat.Order using (zero-≤) +open import Cubical.Data.Vec.Base +open import Cubical.Tactics.NatSolver.NatExpression +open import Cubical.Tactics.NatSolver.HornerForms +open import Cubical.Tactics.NatSolver.EvalHom + +private + variable + : Level + +module EqualityToNormalform where + open Eval + open IteratedHornerOperations + open HomomorphismProperties + + normalize : {n : } Expr n IteratedHornerForms n + normalize {n = n} (K r) = Constant n r + normalize {n = n} ( k) = Variable n k + normalize (x +' y) = + (normalize x) +ₕ (normalize y) + normalize (x ·' y) = + (normalize x) ·ₕ (normalize y) + + isEqualToNormalform : + {n : } + (e : Expr n) (xs : Vec n) + eval (normalize e) xs e xs + isEqualToNormalform (K r) [] = refl + isEqualToNormalform {n = ℕ.suc n} (K r) (x xs) = + eval (Constant (ℕ.suc n) r) (x xs) ≡⟨ refl + eval (0ₕ ·X+ Constant n r) (x xs) ≡⟨ refl + eval 0ₕ (x xs) · x + eval (Constant n r) xs + ≡⟨ cong u u · x + eval (Constant n r) xs) (eval0H (x xs)) + 0 · x + eval (Constant n r) xs + ≡⟨ refl + eval (Constant n r) xs + ≡⟨ isEqualToNormalform (K r) xs + r + + isEqualToNormalform ( zero) (x xs) = + eval (1ₕ ·X+ 0ₕ) (x xs) ≡⟨ refl + eval 1ₕ (x xs) · x + eval 0ₕ xs ≡⟨ cong u u · x + eval 0ₕ xs) + (eval1ₕ (x xs)) + 1 · x + eval 0ₕ xs ≡⟨ cong u 1 · x + u ) (eval0H xs) + 1 · x + 0 ≡⟨ +-zero _ + 1 · x ≡⟨ ·-identityˡ _ + x + isEqualToNormalform {n = ℕ.suc n} ( (suc k)) (x xs) = + eval (0ₕ ·X+ Variable n k) (x xs) ≡⟨ refl + eval 0ₕ (x xs) · x + eval (Variable n k) xs + ≡⟨ cong u u · x + eval (Variable n k) xs) (eval0H (x xs)) + 0 · x + eval (Variable n k) xs + ≡⟨ refl + eval (Variable n k) xs + ≡⟨ isEqualToNormalform ( k) xs + (suc k) (x xs) + + isEqualToNormalform (e +' e₁) [] = + eval (normalize e +ₕ normalize e₁) [] + ≡⟨ +Homeval (normalize e) _ [] + eval (normalize e) [] + + eval (normalize e₁) [] + ≡⟨ cong u u + eval (normalize e₁) []) + (isEqualToNormalform e []) + e [] + + eval (normalize e₁) [] + ≡⟨ cong u e [] + u) (isEqualToNormalform e₁ []) + e [] + e₁ [] + isEqualToNormalform (e +' e₁) (x xs) = + eval (normalize e + +ₕ normalize e₁) (x xs) + ≡⟨ +Homeval (normalize e) _ (x xs) + eval (normalize e) (x xs) + + eval (normalize e₁) (x xs) + ≡⟨ cong u u + eval (normalize e₁) (x xs)) + (isEqualToNormalform e (x xs)) + e (x xs) + + eval (normalize e₁) (x xs) + ≡⟨ cong u e (x xs) + u) + (isEqualToNormalform e₁ (x xs)) + e (x xs) + e₁ (x xs) + + isEqualToNormalform (e ·' e₁) [] = + eval (normalize e ·ₕ normalize e₁) [] + ≡⟨ ·Homeval (normalize e) _ [] + eval (normalize e) [] + · eval (normalize e₁) [] + ≡⟨ cong u u · eval (normalize e₁) []) + (isEqualToNormalform e []) + e [] + · eval (normalize e₁) [] + ≡⟨ cong u e [] · u) (isEqualToNormalform e₁ []) + e [] · e₁ [] + + isEqualToNormalform (e ·' e₁) (x xs) = + eval (normalize e ·ₕ normalize e₁) (x xs) + ≡⟨ ·Homeval (normalize e) _ (x xs) + eval (normalize e) (x xs) + · eval (normalize e₁) (x xs) + ≡⟨ cong u u · eval (normalize e₁) (x xs)) + (isEqualToNormalform e (x xs)) + e (x xs) + · eval (normalize e₁) (x xs) + ≡⟨ cong u e (x xs) · u) + (isEqualToNormalform e₁ (x xs)) + e (x xs) · e₁ (x xs) + + solve : + {n : } (e₁ e₂ : Expr n) (xs : Vec n) + (p : eval (normalize e₁) xs eval (normalize e₂) xs) + e₁ xs e₂ xs + solve e₁ e₂ xs p = + e₁ xs ≡⟨ sym (isEqualToNormalform e₁ xs) + eval (normalize e₁) xs ≡⟨ p + eval (normalize e₂) xs ≡⟨ isEqualToNormalform e₂ xs + e₂ xs diff --git a/src/docs/Cubical.Tactics.NatSolver.html b/src/docs/Cubical.Tactics.NatSolver.html new file mode 100644 index 0000000..05b3812 --- /dev/null +++ b/src/docs/Cubical.Tactics.NatSolver.html @@ -0,0 +1,12 @@ +{-# OPTIONS --safe #-} +{- + This is inspired by/copied from: + https://github.com/agda/agda-stdlib/blob/master/src/Tactic/MonoidSolver.agda + and the 1lab. + + Boilerplate code for calling the ring solver is constructed automatically + with agda's reflection features. +-} +module Cubical.Tactics.NatSolver where + +open import Cubical.Tactics.NatSolver.Reflection public diff --git a/src/docs/Cubical.Tactics.Reflection.Utilities.html b/src/docs/Cubical.Tactics.Reflection.Utilities.html new file mode 100644 index 0000000..4691b3c --- /dev/null +++ b/src/docs/Cubical.Tactics.Reflection.Utilities.html @@ -0,0 +1,35 @@ +{-# OPTIONS --safe #-} +module Cubical.Tactics.Reflection.Utilities where + +open import Cubical.Foundations.Prelude hiding (Type) + +open import Agda.Builtin.Reflection hiding (Type) +open import Agda.Builtin.String +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_) + +open import Cubical.Reflection.Base +open import Cubical.Data.List +open import Cubical.Data.Maybe +open import Cubical.Data.FinData using () renaming (zero to fzero; suc to fsuc) +open import Cubical.Data.Sigma +open import Cubical.Data.Nat using () + +open import Cubical.Tactics.Reflection +open import Cubical.Tactics.Reflection.Variables + + +finiteNumberAsTerm : Maybe Term +finiteNumberAsTerm (just ℕ.zero) = con (quote fzero) [] +finiteNumberAsTerm (just (ℕ.suc n)) = con (quote fsuc) (finiteNumberAsTerm (just n) v∷ []) +finiteNumberAsTerm _ = unknown + +-- error message helper +errorOut : List (Arg Term) TC (Template × Vars) +errorOut something = typeError (strErr "Don't know what to do with " map {(arg _ t) termErr t}) something) + +errorOut' : Term TC (Template × Vars) +errorOut' something = typeError (strErr "Don't know what to do with " termErr something []) + + +_==_ = primQNameEquality +{-# INLINE _==_ #-} diff --git a/src/docs/Cubical.Tactics.Reflection.Variables.html b/src/docs/Cubical.Tactics.Reflection.Variables.html new file mode 100644 index 0000000..72a2054 --- /dev/null +++ b/src/docs/Cubical.Tactics.Reflection.Variables.html @@ -0,0 +1,81 @@ +{-# OPTIONS --safe #-} +{- + This code contains some helper functions for solvers. + Variables in the sense of this files are things that are treated like variables by a solver. + A ring solver might want to treat "f x" in an equation "f x + 0 ≡ f x" like a variable "y". + During the inspection of the lhs and rhs of an equation, terms like "f x" are found and saved + and later, indices are assigned to them. These indices will be the indices of the variables + in the normal forms the solver uses. +-} +module Cubical.Tactics.Reflection.Variables where + +open import Cubical.Foundations.Prelude hiding (Type) + +open import Agda.Builtin.Reflection hiding (Type) +open import Agda.Builtin.String +open import Agda.Builtin.Float +open import Agda.Builtin.Word +open import Agda.Builtin.Char +open import Agda.Builtin.Nat using () renaming (_==_ to _=ℕ_) + +open import Cubical.Reflection.Base +open import Cubical.Data.Bool +open import Cubical.Data.List +open import Cubical.Data.Maybe +open import Cubical.Data.Nat using () + +open import Cubical.Tactics.Reflection + +private + variable + : Level + + +Vars = List Term +VarAss = Term Maybe +Template = VarAss Term + +private + _=N_ = primQNameEquality + _=M_ = primMetaEquality + + _=L_ : Literal Literal Bool + nat n =L nat m = n =ℕ m + word64 n =L word64 m = primWord64ToNat n =ℕ primWord64ToNat m + float x =L float y = primFloatEquality x y + char c =L char c' = primCharEquality c c' + string s =L string s' = primStringEquality s s' + name x =L name y = x =N y + meta x =L meta y = x =M y + _ =L _ = false + + compareVargs : List (Arg Term) List (Arg Term) Bool + + _=T_ : Term Term Bool -- this should be a TC, since it should error out sometimes + var index args =T var index' args' = (index =ℕ index') and compareVargs args args' + con c args =T con c' args' = (c =N c') and compareVargs args args' + def f args =T def f' args' = (f =N f') and compareVargs args args' + lit l =T lit l' = l =L l' + meta x args =T meta x' args' = (x =M x') and compareVargs args args' + _ =T _ = false -- this should be fixed + +compareVargs [] [] = true +compareVargs (x v∷ l) (x' v∷ l') = (x =T x') and compareVargs l l' +compareVargs (_ h∷ l) (_ h∷ l') = compareVargs l l' -- ignore hargs, not sure this is good +compareVargs _ _ = false + +addWithoutRepetition : Term Vars Vars +addWithoutRepetition t l@(t' l') = if (t =T t') then l else t' addWithoutRepetition t l' +addWithoutRepetition t [] = t [] + +appendWithoutRepetition : Vars Vars Vars +appendWithoutRepetition (x l) l' = appendWithoutRepetition l (addWithoutRepetition x l') +appendWithoutRepetition [] l' = l' + +-- this can be used to get a map from variables to numbers 0,...,n +indexOf : Term Vars Maybe +indexOf t (t' l) = + if (t =T t') + then just 0 + else map-Maybe k ℕ.suc k) (indexOf t l) +indexOf t [] = nothing diff --git a/src/docs/Cubical.Tactics.Reflection.html b/src/docs/Cubical.Tactics.Reflection.html new file mode 100644 index 0000000..0b61580 --- /dev/null +++ b/src/docs/Cubical.Tactics.Reflection.html @@ -0,0 +1,114 @@ +-- SPDX-License-Identifier: BSD-3-Clause +{-# OPTIONS --safe #-} +module Cubical.Tactics.Reflection where + +{- Utilities common to different reflection solvers. + + Most of these are copied/adapted from the 1Lab +-} + +open import Cubical.Foundations.Prelude + +open import Agda.Builtin.Reflection hiding (Type) + +open import Cubical.Data.Bool +open import Cubical.Data.List +open import Cubical.Data.Maybe +open import Cubical.Data.Sigma +open import Cubical.Data.Unit + +open import Cubical.Reflection.Base + +private + variable + ℓ' : Level + +_<$>_ : { ℓ'} {A : Type }{B : Type ℓ'} (A B) TC A TC B +f <$> t = t >>= λ x returnTC (f x) + +_<*>_ : { ℓ'} {A : Type }{B : Type ℓ'} TC (A B) TC A TC B +s <*> t = s >>= λ f t >>= λ x returnTC (f x) + +wait-for-args : List (Arg Term) TC (List (Arg Term)) +wait-for-type : Term TC Term + +wait-for-type (var x args) = var x <$> wait-for-args args +wait-for-type (con c args) = con c <$> wait-for-args args +wait-for-type (def f args) = def f <$> wait-for-args args +wait-for-type (lam v (abs x t)) = returnTC (lam v (abs x t)) +wait-for-type (pat-lam cs args) = returnTC (pat-lam cs args) +wait-for-type (pi (arg i a) (abs x b)) = do + a wait-for-type a + b wait-for-type b + returnTC (pi (arg i a) (abs x b)) +wait-for-type (agda-sort s) = returnTC (agda-sort s) +wait-for-type (lit l) = returnTC (lit l) +wait-for-type (meta x x₁) = blockOnMeta x +wait-for-type unknown = returnTC unknown + +wait-for-args [] = returnTC [] +wait-for-args (arg i a xs) = + (_∷_ <$> (arg i <$> wait-for-type a)) <*> wait-for-args xs + +unapply-path : Term TC (Maybe (Term × Term × Term)) +unapply-path red@(def (quote PathP) (l h∷ T v∷ x v∷ y v∷ [])) = do + domain newMeta (def (quote Type) (l v∷ [])) + ty returnTC (def (quote Path) (domain v∷ x v∷ y v∷ [])) + debugPrint "tactic" 50 + (strErr "(no reduction) unapply-path: got a " termErr red + strErr " but I really want it to be " termErr ty []) + unify red ty + returnTC (just (domain , x , y)) +unapply-path tm = reduce tm >>= λ where + tm@(meta _ _) do + dom newMeta (def (quote Type) []) + l newMeta dom + r newMeta dom + unify tm (def (quote Type) (dom v∷ l v∷ r v∷ [])) + wait-for-type l + wait-for-type r + returnTC (just (dom , l , r)) + red@(def (quote PathP) (l h∷ T v∷ x v∷ y v∷ [])) do + domain newMeta (def (quote Type) (l v∷ [])) + ty returnTC (def (quote Path) (domain v∷ x v∷ y v∷ [])) + debugPrint "tactic" 50 + (strErr "unapply-path: got a " termErr red + strErr " but I really want it to be " termErr ty []) + unify red ty + returnTC (just (domain , x , y)) + _ returnTC nothing + +{- + get-boundary maps a term 'x ≡ y' to the pair '(x,y)' +-} +get-boundary : Term TC (Maybe (Term × Term)) +get-boundary tm = unapply-path tm >>= λ where + (just (_ , x , y)) returnTC (just (x , y)) + nothing returnTC nothing + +equation-solver : List Name (Term -> Term -> TC Term) Bool Term TC Unit +equation-solver don't-Reduce mk-call debug hole = + withNormalisation false ( + withReduceDefs (false , don't-Reduce) ( + do + -- | First we normalize the goal + goal inferType hole >>= reduce + -- | Then we parse the goal into an AST + just (lhs , rhs) get-boundary goal + where + nothing + typeError(strErr "The functor solver failed to parse the goal" + termErr goal []) + -- | Then we invoke the solver + -- | And we unify the result of the solver with the original hole. + elhs normalise lhs + erhs normalise rhs + call mk-call elhs erhs + let unify-with-goal = (unify hole call) + noConstraints ( + if debug + then unify-with-goal + else ( + unify-with-goal <|> + typeError ((strErr "Could not equate the following expressions:\n " + termErr elhs strErr "\nAnd\n " termErr erhs [])))))) diff --git a/src/docs/Realizability.ApplicativeStructure.html b/src/docs/Realizability.ApplicativeStructure.html new file mode 100644 index 0000000..02b83b1 --- /dev/null +++ b/src/docs/Realizability.ApplicativeStructure.html @@ -0,0 +1,231 @@ +

Applicative Structure

+

In this module we define the notion of an applicative +structure.

+

A type A has applicative +structure if it has an “application operation” (here represented by +_⨾_) and is a set.

+
+
open import Cubical.Core.Everything
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Relation.Nullary
+open import Cubical.Data.Nat
+open import Cubical.Data.Nat.Order
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Empty renaming (elim to ⊥elim)
+open import Cubical.Tactics.NatSolver
+
+module Realizability.ApplicativeStructure where
+
+private module _ {} {A : Type } where
+  -- Taken from Data.Vec.Base from agda-stdlib
+  foldlOp :  {ℓ'} (B :   Type ℓ')  Type _
+  foldlOp B =  {n}  B n  A  B (suc n)
+
+  opaque
+    foldl :  {ℓ'} {n : } (B :   Type ℓ')  foldlOp B  B zero  Vec A n  B n
+    foldl {ℓ'} {.zero} B op acc emptyVec = acc
+    foldl {ℓ'} {.(suc _)} B op acc (x ∷vec vec) = foldl  n  B (suc n)) op (op acc x) vec
+
+  opaque
+    reverse :  {n}  Vec A n  Vec A n
+    reverse vec = foldl  n  Vec A n)  acc curr  curr  acc) [] vec
+
+  opaque
+    chain :  {n}  Vec A (suc n)  (A  A  A)  A
+    chain {n} (x ∷vec vec) op = foldl  _  A)  acc curr  op acc curr) x vec
+
+
+
record ApplicativeStructure {} (A : Type ) : Type  where
+  infixl 20 _⨾_
+  field
+    isSetA : isSet A
+    _⨾_ : A  A  A
+
+

Since being a set is a property - we will generally not have to think +about it too much. Also, since A is a set - we can drop the +relevance of paths and simply talk about “equality”.

+We can define the notion of a term over an applicative structure. +
module _ {} {A : Type } (as : ApplicativeStructure A) where
+  open ApplicativeStructure as
+  infix 23 `_
+  infixl 22 _̇_
+  data Term (n : ) : Type  where
+    # : Fin n  Term n
+    `_ : A  Term n
+    _̇_ : Term n  Term n  Term n
+
+

These terms can be evaluated into A if we can give the values of the +free variables.

+
  ⟦_⟧ :  {n}  Term n  Vec A n  A
+  ⟦_⟧ (` a) _ = a
+  ⟦_⟧ {n} (# k) subs = lookup k subs
+  ⟦_⟧ (a ̇ b) subs = ( a  subs)  ( b  subs)
+
+  applicationChain :  {n m}  Vec (Term m) (suc n)  Term m
+  applicationChain {n} {m} vec = chain vec  x y  x ̇ y)
+
+  apply :  {n}  A  Vec A n  A
+  apply {n} a vec = chain (a  vec)  x y  x  y)
+
+
+
  private
+    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+      applyWorks :  K a b  apply K (a  b  [])  K  a  b
+      applyWorks K a b = refl
+
+
+

On an applicative structure we can define Feferman structure (or SK +structure). We call an applicative structure endowed with Feferman +structure a combinatory algebra.

+
  record Feferman : Type  where
+    field
+      s : A
+      k : A
+      kab≡a :  a b  k  a  b  a
+      sabc≡ac_bc :  a b c  s  a  b  c  (a  c)  (b  c)
+
+

Feferman structure allows us to construct combinatorial +completeness structure.

+

Imagine we have a term t : Term n (for some +n : ℕ). We can ask if A has a “copy” of +t so that application would correspond to subsitution. That +is, we may ask if we can find an a : A such that +a ⨾ < a¹ ⨾ a² ⨾ .... ⨾ aⁿ > (here the +< ... > notation represents a chain of applications) +would be equal to +t [a¹ / # 0 , a² / # 1 , .... , aⁿ / # (pred n) ]. If the +applicative structure additionally can be endowed with Feferman +structure - then the answer is yes.

+

To get to such a term, we first need to define a function that takes +Term (suc n) to Term n by “abstracting” the +free variable represented by the index # 0.

+

We will call this λ*abst and this will turn out to +behave very similar to λ abstraction - and we will also show that it +validates a kind of β reduction rule.

+
  module ComputationRules (feferman : Feferman) where
+    open Feferman feferman
+
+    opaque
+      λ*abst :  {n}  (e : Term (suc n))  Term n
+      λ*abst {n} (# zero) = ` s ̇ ` k ̇ ` k
+      λ*abst {n} (# (suc x)) = ` k ̇ # x
+      λ*abst {n} (` x) = ` k ̇ ` x
+      λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁
+
+

Remark : It is important to note that in general, +realizability is developed using partial combinatory +algebras and partial applicative structures. +In this case, λ*abst is not particularly well-behaved. The +β reduction-esque rule we derive also does not behave +completely like β reduction. See Jonh Longley’s PhD thesis +“Realizability Toposes and Language Semantics” Theorem 1.1.9.

+

Remark : We declare the definition as +opaque - this is important. If we let Agda unfold this +definition all the way we ocassionally end up with unreadable goals +containing a mess of s and k.

+

We define meta-syntactic sugar for some of the more common cases +:

+
    λ* : Term one  A
+    λ* t =  λ*abst t  []
+
+    λ*2 : Term two  A
+    λ*2 t =  λ*abst (λ*abst t)  []
+
+    λ*3 : Term three  A
+    λ*3 t =  λ*abst (λ*abst (λ*abst t))  []
+
+    λ*4 : Term four  A
+    λ*4 t =  λ*abst (λ*abst (λ*abst (λ*abst t)))  []
+
+

We now show that we have a β-reduction-esque operation. We proceed by +induction on the structure of the term and the number of free +variables.

+

For the particular combinatory algebra Λ/β (terms of the untyped λ +calculus quotiented by β equality) - this β reduction actually coincides +with the “actual” β reduction. TODO : Prove this.

+
    opaque
+      unfolding λ*abst
+      βreduction :  {n}  (body : Term (suc n))  (prim : A)  (subs : Vec A n)   λ*abst body ̇ ` prim  subs   body  (prim  subs)
+      βreduction {n} (# zero) prim subs =
+        s  k  k  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+        k  prim  (k  prim)
+          ≡⟨ kab≡a _ _ 
+        prim
+          
+      βreduction {n} (# (suc x)) prim subs = kab≡a _ _
+      βreduction {n} (` x) prim subs = kab≡a _ _
+      βreduction {n} (rator ̇ rand) prim subs =
+        s   λ*abst rator  subs   λ*abst rand  subs  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+         λ*abst rator  subs  prim  ( λ*abst rand  subs  prim)
+          ≡⟨ cong₂  x y  x  y) (βreduction rator prim subs) (βreduction rand prim subs) 
+         rator  (prim  subs)   rand  (prim  subs)
+          
+
+
+
    λ*chainTerm :  n  Term n  Term zero
+    λ*chainTerm zero t = t
+    λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t)
+
+    λ*chain :  {n}  Term n  A
+    λ*chain {n} t =  λ*chainTerm n t  []
+
+
+

We provide useful reasoning combinators that are useful and +frequent.

+
    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+
+      λ*ComputationRule :  (t : Term 1) (a : A)  λ* t  a   t  (a  [])
+      λ*ComputationRule t a =
+         λ*abst t  []  a
+          ≡⟨ βreduction t a [] 
+         t  (a  [])
+          
+
+      λ*2ComputationRule :  (t : Term 2) (a b : A)  λ*2 t  a  b   t  (b  a  [])
+      λ*2ComputationRule t a b =
+         λ*abst (λ*abst t)  []  a  b
+          ≡⟨ refl 
+         λ*abst (λ*abst t)  []   ` a  []   ` b  []
+          ≡⟨ refl 
+         λ*abst (λ*abst t) ̇ ` a  []   ` b  []
+          ≡⟨ cong  x  x  b) (βreduction (λ*abst t) a []) 
+         λ*abst t  (a  [])  b
+          ≡⟨ βreduction t b (a  []) 
+         t  (b  a  [])
+          
+          
+      λ*3ComputationRule :  (t : Term 3) (a b c : A)  λ*3 t  a  b  c   t  (c  b  a  [])
+      λ*3ComputationRule t a b c =
+         λ*abst (λ*abst (λ*abst t))  []   ` a  []   ` b  []   ` c  []
+          ≡⟨ cong  x  x  b  c) (βreduction (λ*abst (λ*abst t)) a []) 
+         λ*abst (λ*abst t)  (a  [])   ` b  (a  [])   ` c  (a  [])
+          ≡⟨ cong  x  x  c) (βreduction (λ*abst t) b (a  [])) 
+         λ*abst t  (b  a  [])   ` c  (b  a  [])
+          ≡⟨ βreduction t c (b  a  []) 
+         t  (c  b  a  [])
+          
+
+      λ*4ComputationRule :  (t : Term 4) (a b c d : A)  λ*4 t  a  b  c  d   t  (d  c  b  a  [])
+      λ*4ComputationRule t a b c d =
+         λ*abst (λ*abst (λ*abst (λ*abst t)))  []   ` a  []   ` b  []   ` c  []   ` d  []
+          ≡⟨ cong  x  x  b  c  d) (βreduction (λ*abst (λ*abst (λ*abst t))) a []) 
+         λ*abst (λ*abst (λ*abst t))  (a  [])   ` b  (a  [])   ` c  (a  [])   ` d  (a  [])
+          ≡⟨ cong  x  x  c  d) (βreduction (λ*abst (λ*abst t)) b (a  [])) 
+         λ*abst (λ*abst t)  (b  a  [])   ` c  (b  a  [])   ` d  (b  a  [])
+          ≡⟨ cong  x  x  d) (βreduction (λ*abst t) c (b  a  [])) 
+         λ*abst t  (c  b  a  [])   ` d  (c  b  a  [])
+          ≡⟨ βreduction t d (c  b  a  []) 
+         t  (d  c  b  a  [])
+          
+
diff --git a/src/docs/Realizability.ApplicativeStructure.md b/src/docs/Realizability.ApplicativeStructure.md new file mode 100644 index 0000000..62775da --- /dev/null +++ b/src/docs/Realizability.ApplicativeStructure.md @@ -0,0 +1,215 @@ +--- +title: Applicative Structure +author: Rahul Chhabra +--- +# Applicative Structure + +In this module we define the notion of an _applicative structure_. + +A type $A$ has applicative structure if it has an "application operation" (here represented by `_⨾_`) and is a set. + +
+
open import Cubical.Core.Everything
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Relation.Nullary
+open import Cubical.Data.Nat
+open import Cubical.Data.Nat.Order
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Empty renaming (elim to ⊥elim)
+open import Cubical.Tactics.NatSolver
+
+module Realizability.ApplicativeStructure where
+
+private module _ {} {A : Type } where
+  -- Taken from Data.Vec.Base from agda-stdlib
+  foldlOp :  {ℓ'} (B :   Type ℓ')  Type _
+  foldlOp B =  {n}  B n  A  B (suc n)
+
+  opaque
+    foldl :  {ℓ'} {n : } (B :   Type ℓ')  foldlOp B  B zero  Vec A n  B n
+    foldl {ℓ'} {.zero} B op acc emptyVec = acc
+    foldl {ℓ'} {.(suc _)} B op acc (x ∷vec vec) = foldl  n  B (suc n)) op (op acc x) vec
+
+  opaque
+    reverse :  {n}  Vec A n  Vec A n
+    reverse vec = foldl  n  Vec A n)  acc curr  curr  acc) [] vec
+
+  opaque
+    chain :  {n}  Vec A (suc n)  (A  A  A)  A
+    chain {n} (x ∷vec vec) op = foldl  _  A)  acc curr  op acc curr) x vec
+
+ +
record ApplicativeStructure {} (A : Type ) : Type  where
+  infixl 20 _⨾_
+  field
+    isSetA : isSet A
+    _⨾_ : A  A  A
+
+Since being a set is a property - we will generally not have to think about it too much. Also, since `A` is a set - we can drop the relevance of paths and simply talk about "equality". + +We can define the notion of a term over an applicative structure. +
module _ {} {A : Type } (as : ApplicativeStructure A) where
+  open ApplicativeStructure as
+  infix 23 `_
+  infixl 22 _̇_
+  data Term (n : ) : Type  where
+    # : Fin n  Term n
+    `_ : A  Term n
+    _̇_ : Term n  Term n  Term n
+
+These terms can be evaluated into $A$ if we can give the values of the free variables. + +
  ⟦_⟧ :  {n}  Term n  Vec A n  A
+  ⟦_⟧ (` a) _ = a
+  ⟦_⟧ {n} (# k) subs = lookup k subs
+  ⟦_⟧ (a ̇ b) subs = ( a  subs)  ( b  subs)
+
+  applicationChain :  {n m}  Vec (Term m) (suc n)  Term m
+  applicationChain {n} {m} vec = chain vec  x y  x ̇ y)
+
+  apply :  {n}  A  Vec A n  A
+  apply {n} a vec = chain (a  vec)  x y  x  y)
+
+
+
  private
+    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+      applyWorks :  K a b  apply K (a  b  [])  K  a  b
+      applyWorks K a b = refl
+
+ +On an applicative structure we can define Feferman structure (or SK structure). We call an applicative structure endowed with Feferman structure a **combinatory algebra**. + +
  record Feferman : Type  where
+    field
+      s : A
+      k : A
+      kab≡a :  a b  k  a  b  a
+      sabc≡ac_bc :  a b c  s  a  b  c  (a  c)  (b  c)
+
+Feferman structure allows us to construct **combinatorial completeness** structure. + +Imagine we have a term `t : Term n` (for some `n : ℕ`). We can ask if `A` has a "copy" of `t` so that application would correspond to subsitution. That is, we may ask if we can find an `a : A` such that +`a ⨾ < a¹ ⨾ a² ⨾ .... ⨾ aⁿ >` (here the `< ... >` notation represents a chain of applications) would be equal to `t [a¹ / # 0 , a² / # 1 , .... , aⁿ / # (pred n) ]`. If the applicative structure additionally can be endowed with Feferman structure - then the answer is yes. + +To get to such a term, we first need to define a function that takes `Term (suc n)` to `Term n` by "abstracting" the free variable represented by the index `# 0`. + +We will call this `λ*abst` and this will turn out to behave very similar to λ abstraction - and we will also show that it validates a kind of β reduction rule. + +
  module ComputationRules (feferman : Feferman) where
+    open Feferman feferman
+
+    opaque
+      λ*abst :  {n}  (e : Term (suc n))  Term n
+      λ*abst {n} (# zero) = ` s ̇ ` k ̇ ` k
+      λ*abst {n} (# (suc x)) = ` k ̇ # x
+      λ*abst {n} (` x) = ` k ̇ ` x
+      λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁
+
+**Remark** : It is important to note that in general, realizability is developed using **partial combinatory algebras** and **partial applicative structures**. In this case, `λ*abst` is not particularly well-behaved. The β reduction-esque rule we derive also does not behave *completely* like β reduction. See Jonh Longley's PhD thesis "Realizability Toposes and Language Semantics" Theorem 1.1.9. + +**Remark** : We declare the definition as `opaque` - this is important. If we let Agda unfold this definition all the way we ocassionally end up with unreadable goals containing a mess of `s` and `k`. + +We define meta-syntactic sugar for some of the more common cases : + +
    λ* : Term one  A
+    λ* t =  λ*abst t  []
+
+    λ*2 : Term two  A
+    λ*2 t =  λ*abst (λ*abst t)  []
+
+    λ*3 : Term three  A
+    λ*3 t =  λ*abst (λ*abst (λ*abst t))  []
+
+    λ*4 : Term four  A
+    λ*4 t =  λ*abst (λ*abst (λ*abst (λ*abst t)))  []
+
+We now show that we have a β-reduction-esque operation. We proceed by induction on the structure of the term and the number of free variables. + +For the particular combinatory algebra Λ/β (terms of the untyped λ calculus quotiented by β equality) - this β reduction actually coincides with the "actual" β reduction. +TODO : Prove this. + +
    opaque
+      unfolding λ*abst
+      βreduction :  {n}  (body : Term (suc n))  (prim : A)  (subs : Vec A n)   λ*abst body ̇ ` prim  subs   body  (prim  subs)
+      βreduction {n} (# zero) prim subs =
+        s  k  k  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+        k  prim  (k  prim)
+          ≡⟨ kab≡a _ _ 
+        prim
+          
+      βreduction {n} (# (suc x)) prim subs = kab≡a _ _
+      βreduction {n} (` x) prim subs = kab≡a _ _
+      βreduction {n} (rator ̇ rand) prim subs =
+        s   λ*abst rator  subs   λ*abst rand  subs  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+         λ*abst rator  subs  prim  ( λ*abst rand  subs  prim)
+          ≡⟨ cong₂  x y  x  y) (βreduction rator prim subs) (βreduction rand prim subs) 
+         rator  (prim  subs)   rand  (prim  subs)
+          
+
+
+
    λ*chainTerm :  n  Term n  Term zero
+    λ*chainTerm zero t = t
+    λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t)
+
+    λ*chain :  {n}  Term n  A
+    λ*chain {n} t =  λ*chainTerm n t  []
+
+ +We provide useful reasoning combinators that are useful and frequent. + +
    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+
+      λ*ComputationRule :  (t : Term 1) (a : A)  λ* t  a   t  (a  [])
+      λ*ComputationRule t a =
+         λ*abst t  []  a
+          ≡⟨ βreduction t a [] 
+         t  (a  [])
+          
+
+      λ*2ComputationRule :  (t : Term 2) (a b : A)  λ*2 t  a  b   t  (b  a  [])
+      λ*2ComputationRule t a b =
+         λ*abst (λ*abst t)  []  a  b
+          ≡⟨ refl 
+         λ*abst (λ*abst t)  []   ` a  []   ` b  []
+          ≡⟨ refl 
+         λ*abst (λ*abst t) ̇ ` a  []   ` b  []
+          ≡⟨ cong  x  x  b) (βreduction (λ*abst t) a []) 
+         λ*abst t  (a  [])  b
+          ≡⟨ βreduction t b (a  []) 
+         t  (b  a  [])
+          
+          
+      λ*3ComputationRule :  (t : Term 3) (a b c : A)  λ*3 t  a  b  c   t  (c  b  a  [])
+      λ*3ComputationRule t a b c =
+         λ*abst (λ*abst (λ*abst t))  []   ` a  []   ` b  []   ` c  []
+          ≡⟨ cong  x  x  b  c) (βreduction (λ*abst (λ*abst t)) a []) 
+         λ*abst (λ*abst t)  (a  [])   ` b  (a  [])   ` c  (a  [])
+          ≡⟨ cong  x  x  c) (βreduction (λ*abst t) b (a  [])) 
+         λ*abst t  (b  a  [])   ` c  (b  a  [])
+          ≡⟨ βreduction t c (b  a  []) 
+         t  (c  b  a  [])
+          
+
+      λ*4ComputationRule :  (t : Term 4) (a b c d : A)  λ*4 t  a  b  c  d   t  (d  c  b  a  [])
+      λ*4ComputationRule t a b c d =
+         λ*abst (λ*abst (λ*abst (λ*abst t)))  []   ` a  []   ` b  []   ` c  []   ` d  []
+          ≡⟨ cong  x  x  b  c  d) (βreduction (λ*abst (λ*abst (λ*abst t))) a []) 
+         λ*abst (λ*abst (λ*abst t))  (a  [])   ` b  (a  [])   ` c  (a  [])   ` d  (a  [])
+          ≡⟨ cong  x  x  c  d) (βreduction (λ*abst (λ*abst t)) b (a  [])) 
+         λ*abst (λ*abst t)  (b  a  [])   ` c  (b  a  [])   ` d  (b  a  [])
+          ≡⟨ cong  x  x  d) (βreduction (λ*abst t) c (b  a  [])) 
+         λ*abst t  (c  b  a  [])   ` d  (c  b  a  [])
+          ≡⟨ βreduction t d (c  b  a  []) 
+         t  (d  c  b  a  [])
+          
+
\ No newline at end of file From 26be6f07ba92128146fc9fe9c0448fbccea5e220 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 6 May 2024 17:41:14 +0530 Subject: [PATCH 43/61] Define PERs and refactor assembly modules --- src/Realizability/Assembly/Base.agda | 59 ++++-- src/Realizability/Assembly/BinCoproducts.agda | 21 +-- src/Realizability/Assembly/BinProducts.agda | 131 +++---------- src/Realizability/Assembly/Coequalizers.agda | 33 ++-- src/Realizability/Assembly/Equalizers.agda | 2 +- src/Realizability/Assembly/Everything.agda | 3 +- src/Realizability/Assembly/Exponentials.agda | 17 +- src/Realizability/Assembly/Morphism.agda | 54 +++--- src/Realizability/PERs/PER.agda | 178 +++++++++++++----- 9 files changed, 272 insertions(+), 226 deletions(-) diff --git a/src/Realizability/Assembly/Base.agda b/src/Realizability/Assembly/Base.agda index 21be2d4..8ddc335 100644 --- a/src/Realizability/Assembly/Base.agda +++ b/src/Realizability/Assembly/Base.agda @@ -2,6 +2,10 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Functions.FunExtEquiv open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation open import Cubical.Reflection.RecordEquiv @@ -11,32 +15,55 @@ module Realizability.Assembly.Base {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra record Assembly (X : Type ℓ) : Type (ℓ-suc ℓ) where infix 25 _⊩_ field - isSetX : isSet X _⊩_ : A → X → Type ℓ + isSetX : isSet X ⊩isPropValued : ∀ a x → isProp (a ⊩ x) ⊩surjective : ∀ x → ∃[ a ∈ A ] a ⊩ x + open Assembly public + _⊩[_]_ : ∀ {X : Type ℓ} → A → Assembly X → X → Type ℓ + a ⊩[ A ] x = A ._⊩_ a x AssemblyΣ : Type ℓ → Type _ AssemblyΣ X = - Σ[ isSetX ∈ isSet X ] Σ[ _⊩_ ∈ (A → X → hProp ℓ) ] - (∀ x → ∃[ a ∈ A ] ⟨ a ⊩ x ⟩) - - AssemblyΣX→isSetX : ∀ X → AssemblyΣ X → isSet X - AssemblyΣX→isSetX X (isSetX , _ , _) = isSetX - - AssemblyΣX→⊩ : ∀ X → AssemblyΣ X → (A → X → hProp ℓ) - AssemblyΣX→⊩ X (_ , ⊩ , _) = ⊩ - - AssemblyΣX→⊩surjective : ∀ X → (asm : AssemblyΣ X) → (∀ x → ∃[ a ∈ A ] ⟨ AssemblyΣX→⊩ X asm a x ⟩) - AssemblyΣX→⊩surjective X (_ , _ , ⊩surjective) = ⊩surjective + (∀ x → ∃[ a ∈ A ] ⟨ a ⊩ x ⟩) × + (isSet X) isSetAssemblyΣ : ∀ X → isSet (AssemblyΣ X) - isSetAssemblyΣ X = isSetΣ (isProp→isSet isPropIsSet) λ isSetX → isSetΣ (isSetΠ (λ a → isSetΠ λ x → isSetHProp)) λ _⊩_ → isSetΠ λ x → isProp→isSet isPropPropTrunc + isSetAssemblyΣ X = isSetΣ (isSetΠ2 λ _ _ → isSetHProp) (λ rel → isSet× (isSetΠ λ x → isProp→isSet isPropPropTrunc) (isProp→isSet isPropIsSet)) - unquoteDecl AssemblyIsoΣ = declareRecordIsoΣ AssemblyIsoΣ (quote Assembly) + AssemblyΣ≡Equiv : ∀ X → (a b : AssemblyΣ X) → (a ≡ b) ≃ (∀ r x → ⟨ a .fst r x ⟩ ≃ ⟨ b .fst r x ⟩) + AssemblyΣ≡Equiv X a b = + a ≡ b + ≃⟨ invEquiv (Σ≡PropEquiv (λ rel → isProp× (isPropΠ λ x → isPropPropTrunc) isPropIsSet) {u = a} {v = b}) ⟩ + a .fst ≡ b .fst + ≃⟨ invEquiv (funExt₂Equiv {f = a .fst} {g = b .fst}) ⟩ + (∀ (r : A) (x : X) → a .fst r x ≡ b .fst r x) + ≃⟨ + equivΠCod + (λ r → + equivΠCod + λ x → + compEquiv + (invEquiv (Σ≡PropEquiv (λ _ → isPropIsProp) {u = a .fst r x} {v = b .fst r x})) + (univalence {A = a .fst r x .fst} {B = b .fst r x .fst})) + ⟩ + (∀ (r : A) (x : X) → ⟨ a .fst r x ⟩ ≃ ⟨ b .fst r x ⟩) + ■ - open Assembly public + -- definitional isomorphism + AssemblyΣIsoAssembly : ∀ X → Iso (AssemblyΣ X) (Assembly X) + _⊩_ (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) a x = ⟨ rel a x ⟩ + Assembly.isSetX (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) = isSetX + ⊩isPropValued (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) a x = str (rel a x) + ⊩surjective (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) x = surj x + Iso.inv (AssemblyΣIsoAssembly X) asm = (λ a x → (a ⊩[ asm ] x) , (asm .⊩isPropValued a x)) , (λ x → asm .⊩surjective x) , asm .isSetX + Iso.rightInv (AssemblyΣIsoAssembly X) asm = refl + Iso.leftInv (AssemblyΣIsoAssembly X) (rel , surj , isSetX) = refl - + AssemblyΣ≃Assembly : ∀ X → AssemblyΣ X ≃ Assembly X + AssemblyΣ≃Assembly X = isoToEquiv (AssemblyΣIsoAssembly X) + + isSetAssembly : ∀ X → isSet (Assembly X) + isSetAssembly X = isOfHLevelRespectEquiv 2 (AssemblyΣ≃Assembly X) (isSetAssemblyΣ X) diff --git a/src/Realizability/Assembly/BinCoproducts.agda b/src/Realizability/Assembly/BinCoproducts.agda index 502cfa9..4648b7d 100644 --- a/src/Realizability/Assembly/BinCoproducts.agda +++ b/src/Realizability/Assembly/BinCoproducts.agda @@ -3,7 +3,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Data.Sum hiding (map) open import Cubical.Data.Sigma -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Nat open import Cubical.Data.Vec hiding (map) open import Cubical.HITs.PropositionalTruncation hiding (map) @@ -11,7 +11,7 @@ open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Categories.Category open import Cubical.Categories.Limits.BinCoproduct open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-chain to `λ*; λ*-naturality to `λ*ComputationRule) hiding (λ*) +open import Realizability.ApplicativeStructure module Realizability.Assembly.BinCoproducts {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where @@ -20,9 +20,6 @@ open import Realizability.Assembly.Base ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) open import Realizability.Assembly.Morphism ca -λ* = `λ* as fefermanStructure -λ*ComputationRule = `λ*ComputationRule as fefermanStructure - infixl 23 _⊕_ _⊕_ : {A B : Type ℓ} → Assembly A → Assembly B → Assembly (A ⊎ B) (as ⊕ bs) .isSetX = isSet⊎ (as .isSetX) (bs .isSetX) @@ -61,12 +58,13 @@ _⊕_ : {A B : Type ℓ} → Assembly A → Assembly B → Assembly (A ⊎ B) [ f , g ] .map (inr y) = g .map y [_,_] {asmZ = asmZ} f g .tracker = do + -- these are not considered structurally smaller since these are in the propositional truncation (f~ , f~tracks) ← f .tracker (g~ , g~tracks) ← g .tracker -- if (pr₁ r) then f (pr₂ r) else g (pr₂ r) let tracker : Term as (suc zero) - tracker = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` f~ ̇ (` pr₂ ̇ (# fzero))) ̇ (` g~ ̇ (` pr₂ ̇ (# fzero))) + tracker = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` f~ ̇ (` pr₂ ̇ (# zero))) ̇ (` g~ ̇ (` pr₂ ̇ (# zero))) return (λ* tracker , λ { (inl x) r r⊩inl → @@ -79,7 +77,7 @@ _⊕_ : {A B : Type ℓ} → Assembly A → Assembly B → Assembly (A ⊎ B) (λ r → asmZ ._⊩_ r ([ f , g ] .map (inl x))) (sym (λ* tracker ⨾ r - ≡⟨ λ*ComputationRule tracker (r ∷ []) ⟩ + ≡⟨ λ*ComputationRule tracker r ⟩ Id ⨾ (pr₁ ⨾ r) ⨾ (f~ ⨾ (pr₂ ⨾ r)) ⨾ (g~ ⨾ (pr₂ ⨾ r)) ≡⟨ cong (λ r → Id ⨾ (pr₁ ⨾ r) ⨾ (f~ ⨾ (pr₂ ⨾ r)) ⨾ (g~ ⨾ (pr₂ ⨾ r))) r≡pair⨾true⨾rₓ ⟩ Id ⨾ (pr₁ ⨾ (pair ⨾ true ⨾ rₓ)) ⨾ (f~ ⨾ (pr₂ ⨾ (pair ⨾ true ⨾ rₓ))) ⨾ (g~ ⨾ (pr₂ ⨾ (pair ⨾ true ⨾ rₓ))) @@ -101,7 +99,7 @@ _⊕_ : {A B : Type ℓ} → Assembly A → Assembly B → Assembly (A ⊎ B) (λ r → asmZ ._⊩_ r ([ f , g ] .map (inr y))) (sym ((λ* tracker ⨾ r - ≡⟨ λ*ComputationRule tracker (r ∷ []) ⟩ + ≡⟨ λ*ComputationRule tracker r ⟩ Id ⨾ (pr₁ ⨾ r) ⨾ (f~ ⨾ (pr₂ ⨾ r)) ⨾ (g~ ⨾ (pr₂ ⨾ r)) ≡⟨ cong (λ r → Id ⨾ (pr₁ ⨾ r) ⨾ (f~ ⨾ (pr₂ ⨾ r)) ⨾ (g~ ⨾ (pr₂ ⨾ r))) r≡pair⨾false⨾yᵣ ⟩ Id ⨾ (pr₁ ⨾ (pair ⨾ false ⨾ yᵣ)) ⨾ (f~ ⨾ (pr₂ ⨾ (pair ⨾ false ⨾ yᵣ))) ⨾ (g~ ⨾ (pr₂ ⨾ (pair ⨾ false ⨾ yᵣ))) @@ -126,6 +124,7 @@ BinCoproductsASM (X , asmX) (Y , asmY) .univProp {Z , asmZ} f g = (λ ! → isProp× (isSetAssemblyMorphism _ _ _ _) (isSetAssemblyMorphism _ _ _ _)) λ ! (κ₁⊚!≡f , κ₂⊚!≡g) → AssemblyMorphism≡ _ _ (funExt λ { (inl x) i → κ₁⊚!≡f (~ i) .map x ; (inr y) i → κ₂⊚!≡g (~ i) .map y }) +-- I have no idea why I did these since this can be derived from the universal property of the coproduct anyway? module _ {X Y : Type ℓ} (asmX : Assembly X) @@ -142,7 +141,7 @@ module _ do let tracker : Term as 1 - tracker = ` Id ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # fzero)) ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ # fzero)) + tracker = ` Id ̇ (` pr₁ ̇ # zero) ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # zero)) ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ # zero)) return ((λ* tracker) , (λ { (inl x) r r⊩inl → @@ -154,7 +153,7 @@ module _ λ*trackerEq : λ* tracker ⨾ r ≡ pair ⨾ false ⨾ rₓ λ*trackerEq = λ* tracker ⨾ r - ≡⟨ λ*ComputationRule tracker (r ∷ []) ⟩ + ≡⟨ λ*ComputationRule tracker r ⟩ Id ⨾ (pr₁ ⨾ r) ⨾ (pair ⨾ false ⨾ (pr₂ ⨾ r)) ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ r)) ≡⟨ cong (λ r → Id ⨾ (pr₁ ⨾ r) ⨾ (pair ⨾ false ⨾ (pr₂ ⨾ r)) ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ r))) r≡pair⨾true⨾rₓ ⟩ Id ⨾ (pr₁ ⨾ (pair ⨾ true ⨾ rₓ)) ⨾ (pair ⨾ false ⨾ (pr₂ ⨾ (pair ⨾ true ⨾ rₓ))) ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ (pair ⨾ true ⨾ rₓ))) @@ -175,7 +174,7 @@ module _ λ*trackerEq : λ* tracker ⨾ r ≡ pair ⨾ true ⨾ yᵣ λ*trackerEq = λ* tracker ⨾ r - ≡⟨ λ*ComputationRule tracker (r ∷ []) ⟩ + ≡⟨ λ*ComputationRule tracker r ⟩ Id ⨾ (pr₁ ⨾ r) ⨾ (pair ⨾ false ⨾ (pr₂ ⨾ r)) ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ r)) ≡⟨ cong (λ r → Id ⨾ (pr₁ ⨾ r) ⨾ (pair ⨾ false ⨾ (pr₂ ⨾ r)) ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ r))) r≡pair⨾false⨾yᵣ ⟩ Id ⨾ (pr₁ ⨾ (pair ⨾ false ⨾ yᵣ)) ⨾ (pair ⨾ false ⨾ (pr₂ ⨾ (pair ⨾ false ⨾ yᵣ))) ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ (pair ⨾ false ⨾ yᵣ))) diff --git a/src/Realizability/Assembly/BinProducts.agda b/src/Realizability/Assembly/BinProducts.agda index 36c7976..b076635 100644 --- a/src/Realizability/Assembly/BinProducts.agda +++ b/src/Realizability/Assembly/BinProducts.agda @@ -1,11 +1,13 @@ -{-# OPTIONS --cubical --allow-unsolved-metas #-} +{-# OPTIONS --cubical #-} open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Data.Sigma +open import Cubical.Data.FinData open import Cubical.HITs.PropositionalTruncation hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Categories.Limits.BinProduct open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure module Realizability.Assembly.BinProducts {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where @@ -40,52 +42,19 @@ _⊗_ : {A B : Type ℓ} → Assembly A → Assembly B → Assembly (A × B) (g : AssemblyMorphism zs ws) → AssemblyMorphism (xs ⊗ zs) (ys ⊗ ws) ⟪ f , g ⟫ .map (x , z) = f .map x , g .map z -⟪_,_⟫ {ys = ys} {ws = ws} f g .tracker = (do - (f~ , f~tracks) ← f .tracker - (g~ , g~tracks) ← g .tracker - return (s ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ (s ⨾ (k ⨾ pr₁) ⨾ Id))) ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id)) - , λ xz r r⊩xz → - ( subst (λ y → ys ._⊩_ y (f .map (xz .fst))) - (sym (subst _ - (sym (t⨾r≡pair_fg f~ g~ r)) - (pr₁pxy≡x (f~ ⨾ (pr₁ ⨾ r)) (g~ ⨾ (pr₂ ⨾ r))))) - (f~tracks (xz .fst) (pr₁ ⨾ r) (r⊩xz .fst))) - , subst (λ y → ws ._⊩_ y (g .map (xz .snd))) - (sym (subst _ - (sym (t⨾r≡pair_fg f~ g~ r)) - (pr₂pxy≡y (f~ ⨾ (pr₁ ⨾ r)) (g~ ⨾ (pr₂ ⨾ r))))) - (g~tracks (xz .snd) (pr₂ ⨾ r) (r⊩xz .snd)))) - where - module _ (f~ g~ r : A) where - subf≡fprr : ∀ f pr → (s ⨾ (k ⨾ f) ⨾ (s ⨾ (k ⨾ pr) ⨾ Id) ⨾ r) ≡ (f ⨾ (pr ⨾ r)) - subf≡fprr f pr = - s ⨾ (k ⨾ f) ⨾ (s ⨾ (k ⨾ pr) ⨾ Id) ⨾ r - ≡⟨ sabc≡ac_bc _ _ _ ⟩ - (k ⨾ f ⨾ r) ⨾ (s ⨾ (k ⨾ pr) ⨾ Id ⨾ r) - ≡⟨ cong (λ x → x ⨾ _) (kab≡a f r) ⟩ - f ⨾ (s ⨾ (k ⨾ pr) ⨾ Id ⨾ r) - ≡⟨ cong (λ x → f ⨾ x) (sabc≡ac_bc _ _ _) ⟩ - f ⨾ (k ⨾ pr ⨾ r ⨾ (Id ⨾ r)) - ≡⟨ cong (λ x → f ⨾ (x ⨾ (Id ⨾ r))) (kab≡a _ _ ) ⟩ - f ⨾ (pr ⨾ (Id ⨾ r)) - ≡⟨ cong (λ x → f ⨾ (pr ⨾ x)) (Ida≡a r) ⟩ - f ⨾ (pr ⨾ r) - ∎ - t⨾r≡pair_fg : - s ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ (s ⨾ (k ⨾ pr₁) ⨾ Id))) ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id)) ⨾ r - ≡ pair ⨾ (f~ ⨾ (pr₁ ⨾ r)) ⨾ (g~ ⨾ (pr₂ ⨾ r)) - t⨾r≡pair_fg = - s ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ (s ⨾ (k ⨾ pr₁) ⨾ Id))) ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id)) ⨾ r - ≡⟨ sabc≡ac_bc _ _ _ ⟩ - s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ (s ⨾ (k ⨾ pr₁) ⨾ Id)) ⨾ r ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id) ⨾ r) - ≡⟨ cong (λ x → x ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id) ⨾ r)) (sabc≡ac_bc _ _ _) ⟩ - k ⨾ pair ⨾ r ⨾ (s ⨾ (k ⨾ f~) ⨾ (s ⨾ (k ⨾ pr₁) ⨾ Id) ⨾ r) ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id) ⨾ r) - ≡⟨ cong (λ x → x ⨾ (s ⨾ (k ⨾ f~) ⨾ (s ⨾ (k ⨾ pr₁) ⨾ Id) ⨾ r) ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id) ⨾ r)) - (kab≡a pair r) ⟩ - pair ⨾ (s ⨾ (k ⨾ f~) ⨾ (s ⨾ (k ⨾ pr₁) ⨾ Id) ⨾ r) ⨾ (s ⨾ (k ⨾ g~) ⨾ (s ⨾ (k ⨾ pr₂) ⨾ Id) ⨾ r) - ≡⟨ cong₂ (λ x y → pair ⨾ x ⨾ y) (subf≡fprr f~ pr₁) (subf≡fprr g~ pr₂) ⟩ - pair ⨾ (f~ ⨾ (pr₁ ⨾ r)) ⨾ (g~ ⨾ (pr₂ ⨾ r)) - ∎ +⟪_,_⟫ {ys = ys} {ws = ws} f g .tracker = + do + (f~ , f~⊩isTrackedF) ← f .tracker + (g~ , g~⊩isTrackedG) ← g .tracker + let + realizer : Term as 1 + realizer = ` pair ̇ (` f~ ̇ (` pr₁ ̇ # zero)) ̇ (` g~ ̇ (` pr₂ ̇ # zero)) + return + (λ* realizer , + (λ { (x , z) r (pr₁r⊩x , pr₂r⊩z) → + subst (λ r' → r' ⊩[ ys ] (f .map x)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₁pxy≡x _ _)) (f~⊩isTrackedF x (pr₁ ⨾ r) pr₁r⊩x) , + subst (λ r' → r' ⊩[ ws ] (g .map z)) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₂pxy≡y _ _)) (g~⊩isTrackedG z (pr₂ ⨾ r) pr₂r⊩z) })) + π₁ : {A B : Type ℓ} {as : Assembly A} {bs : Assembly B} → AssemblyMorphism (as ⊗ bs) as π₁ .map (a , b) = a π₁ .tracker = ∣ pr₁ , (λ (a , b) p (goal , _) → goal) ∣₁ @@ -100,63 +69,19 @@ _⊗_ : {A B : Type ℓ} → Assembly A → Assembly B → Assembly (A × B) → AssemblyMorphism zs ys → AssemblyMorphism zs (xs ⊗ ys) ⟨ f , g ⟩ .map z = f .map z , g .map z -⟨_,_⟩ {X} {Y} {Z} {xs} {ys} {zs} f g .tracker = map2 untruncated (f .tracker) (g .tracker) where - module _ - ((f~ , f~tracks) : Σ[ f~ ∈ A ] tracks {xs = zs} {ys = xs} f~ (f .map)) - ((g~ , g~tracks) : Σ[ g~ ∈ A ] tracks {xs = zs} {ys = ys} g~ (g .map)) where - - _⊩X_ = xs ._⊩_ - _⊩Y_ = ys ._⊩_ - _⊩Z_ = zs ._⊩_ - - t = s ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ Id)) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id) - untruncated : Σ[ t ∈ A ] (∀ z zᵣ zᵣ⊩z → ((pr₁ ⨾ (t ⨾ zᵣ)) ⊩X (f .map z)) × ((pr₂ ⨾ (t ⨾ zᵣ)) ⊩Y (g .map z))) - untruncated = t , λ z zᵣ zᵣ⊩z → goal₁ z zᵣ zᵣ⊩z , goal₂ z zᵣ zᵣ⊩z where - module _ (z : Z) (zᵣ : A) (zᵣ⊩z : zᵣ ⊩Z z) where - - pr₁⨾tracker⨾zᵣ≡f~⨾zᵣ : pr₁ ⨾ (t ⨾ zᵣ) ≡ f~ ⨾ zᵣ - pr₁⨾tracker⨾zᵣ≡f~⨾zᵣ = - pr₁ ⨾ (s ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ Id)) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id) ⨾ zᵣ) - ≡⟨ cong (λ x → pr₁ ⨾ x) (sabc≡ac_bc _ _ _) ⟩ - pr₁ ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ Id) ⨾ zᵣ ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ)) - ≡⟨ cong (λ x → pr₁ ⨾ (x ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ))) (sabc≡ac_bc _ _ _) ⟩ - pr₁ ⨾ (k ⨾ pair ⨾ zᵣ ⨾ (s ⨾ (k ⨾ f~) ⨾ Id ⨾ zᵣ) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ)) - ≡⟨ cong (λ x → pr₁ ⨾ (x ⨾ (s ⨾ (k ⨾ f~) ⨾ Id ⨾ zᵣ) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ))) (kab≡a _ _) ⟩ - pr₁ ⨾ (pair ⨾ (s ⨾ (k ⨾ f~) ⨾ Id ⨾ zᵣ) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ)) - ≡⟨ pr₁pxy≡x _ _ ⟩ - s ⨾ (k ⨾ f~) ⨾ Id ⨾ zᵣ - ≡⟨ sabc≡ac_bc _ _ _ ⟩ - k ⨾ f~ ⨾ zᵣ ⨾ (Id ⨾ zᵣ) - ≡⟨ cong (λ x → x ⨾ (Id ⨾ zᵣ)) (kab≡a _ _) ⟩ - f~ ⨾ (Id ⨾ zᵣ) - ≡⟨ cong (λ x → f~ ⨾ x) (Ida≡a _) ⟩ - f~ ⨾ zᵣ - ∎ - - pr₂⨾tracker⨾zᵣ≡g~⨾zᵣ : pr₂ ⨾ (t ⨾ zᵣ) ≡ g~ ⨾ zᵣ - pr₂⨾tracker⨾zᵣ≡g~⨾zᵣ = - pr₂ ⨾ (s ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ Id)) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id) ⨾ zᵣ) - ≡⟨ cong (λ x → pr₂ ⨾ x) (sabc≡ac_bc _ _ _) ⟩ - pr₂ ⨾ (s ⨾ (k ⨾ pair) ⨾ (s ⨾ (k ⨾ f~) ⨾ Id) ⨾ zᵣ ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ)) - ≡⟨ cong (λ x → pr₂ ⨾ (x ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ))) (sabc≡ac_bc _ _ _) ⟩ - pr₂ ⨾ (k ⨾ pair ⨾ zᵣ ⨾ (s ⨾ (k ⨾ f~) ⨾ Id ⨾ zᵣ) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ)) - ≡⟨ cong (λ x → pr₂ ⨾ (x ⨾ (s ⨾ (k ⨾ f~) ⨾ Id ⨾ zᵣ) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ))) (kab≡a _ _) ⟩ - pr₂ ⨾ (pair ⨾ (s ⨾ (k ⨾ f~) ⨾ Id ⨾ zᵣ) ⨾ (s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ)) - ≡⟨ pr₂pxy≡y _ _ ⟩ - s ⨾ (k ⨾ g~) ⨾ Id ⨾ zᵣ - ≡⟨ sabc≡ac_bc _ _ _ ⟩ - k ⨾ g~ ⨾ zᵣ ⨾ (Id ⨾ zᵣ) - ≡⟨ cong (λ x → x ⨾ (Id ⨾ zᵣ)) (kab≡a _ _) ⟩ - g~ ⨾ (Id ⨾ zᵣ) - ≡⟨ cong (λ x → g~ ⨾ x) (Ida≡a _) ⟩ - g~ ⨾ zᵣ - ∎ - - goal₁ : (pr₁ ⨾ (t ⨾ zᵣ)) ⊩X (f .map z) - goal₁ = subst (λ y → y ⊩X (f .map z)) (sym pr₁⨾tracker⨾zᵣ≡f~⨾zᵣ) (f~tracks z zᵣ zᵣ⊩z) +⟨_,_⟩ {X} {Y} {Z} {xs} {ys} {zs} f g .tracker = + do + (f~ , f~⊩isTrackedF) ← f .tracker + (g~ , g~⊩isTrackedG) ← g .tracker + let + realizer : Term as 1 + realizer = ` pair ̇ (` f~ ̇ # zero) ̇ (` g~ ̇ # zero) + return + (λ* realizer , + (λ z r r⊩z → + subst (λ r' → r' ⊩[ xs ] (f .map z)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₁pxy≡x _ _)) (f~⊩isTrackedF z r r⊩z) , + subst (λ r' → r' ⊩[ ys ] (g .map z)) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₂pxy≡y _ _)) (g~⊩isTrackedG z r r⊩z))) - goal₂ : (pr₂ ⨾ (t ⨾ zᵣ)) ⊩Y (g .map z) - goal₂ = subst (λ y → y ⊩Y (g .map z)) (sym pr₂⨾tracker⨾zᵣ≡g~⨾zᵣ) (g~tracks z zᵣ zᵣ⊩z) module _ {X Y : Type ℓ} (xs : Assembly X) (ys : Assembly Y) where theπ₁ = π₁ {A = X} {B = Y} {as = xs} {bs = ys} theπ₂ = π₂ {A = X} {B = Y} {as = xs} {bs = ys} diff --git a/src/Realizability/Assembly/Coequalizers.agda b/src/Realizability/Assembly/Coequalizers.agda index 1d7e0b8..c49ad62 100644 --- a/src/Realizability/Assembly/Coequalizers.agda +++ b/src/Realizability/Assembly/Coequalizers.agda @@ -5,11 +5,12 @@ open import Cubical.HITs.SetCoequalizer renaming (rec to setCoequalizerRec; elim open import Cubical.HITs.PropositionalTruncation hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Data.Sigma -open import Cubical.Categories.Limits.Coequalizers open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure module Realizability.Assembly.Coequalizers {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open CombinatoryAlgebra ca open import Realizability.Assembly.Base ca open import Realizability.Assembly.Morphism ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -62,24 +63,28 @@ module _ → (f ⊚ ι' ≡ g ⊚ ι') → ∃![ ! ∈ AssemblyMorphism coequalizer zs ] (ιcoequalizer ⊚ ! ≡ ι') coequalizerFactors (Z , zs) ι' f⊚ι'≡g⊚ι' = + uniqueExists + (let + map = (λ x → setCoequalizerRec (zs .isSetX) (ι' .map) (λ x i → f⊚ι'≡g⊚ι' i .map x) x) + in + makeAssemblyMorphism + map + (do + (ι'~ , ι'~⊩isTrackedι') ← ι' .tracker + return + (ι'~ , + (λ x r r⊩x → setCoequalizerElimProp {C = λ x → ∀ (r : A) → r ⊩[ coequalizer ] x → (ι'~ ⨾ r) ⊩[ zs ] (map x)} {!!} (λ y r r⊩y → {!!}) x r r⊩x)))) + {!!} + {!!} + {!!} + {- uniqueExists (λ where .map → setCoequalizerRec (zs .isSetX) (ι' .map) λ x → λ i → f⊚ι'≡g⊚ι' i .map x - .tracker → {!!}) + .tracker → return ({!!} , (λ x r r⊩x → {!setCoequalizerElimProp {C = λ x → !}))) (AssemblyMorphism≡ _ _ (funExt λ x → refl)) (λ ! → isSetAssemblyMorphism ys zs (ιcoequalizer ⊚ !) ι') λ ! ιcoequalizer⊚!≡ι' → AssemblyMorphism≡ _ _ (funExt λ x → setCoequalizerElimProp {C = λ x → setCoequalizerRec (zs .isSetX) (ι' .map) (λ x₁ i → f⊚ι'≡g⊚ι' i .map x₁) x ≡ ! .map x} - (λ x → zs .isSetX _ _) (λ y → λ i → ιcoequalizer⊚!≡ι' (~ i) .map y) x) - open Coequalizer - open IsCoequalizer - - ιIsCoequalizer : IsCoequalizer {C = ASM} f g ιcoequalizer - ιIsCoequalizer .glues = AssemblyMorphism≡ _ _ (funExt λ x → SetCoequalizer.coeq x) - ιIsCoequalizer .univProp q qGlues = coequalizerFactors _ q qGlues - - ASMCoequalizer : Coequalizer {C = ASM} f g - ASMCoequalizer .coapex = (SetCoequalizer (f .map) (g .map)) , coequalizer - Coequalizer.coeq ASMCoequalizer = ιcoequalizer - ASMCoequalizer .isCoequalizer = ιIsCoequalizer + (λ x → zs .isSetX _ _) (λ y → λ i → ιcoequalizer⊚!≡ι' (~ i) .map y) x) -} diff --git a/src/Realizability/Assembly/Equalizers.agda b/src/Realizability/Assembly/Equalizers.agda index 3bee734..67e634e 100644 --- a/src/Realizability/Assembly/Equalizers.agda +++ b/src/Realizability/Assembly/Equalizers.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --cubical --allow-unsolved-metas #-} +{-# OPTIONS --cubical #-} open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Data.Sigma diff --git a/src/Realizability/Assembly/Everything.agda b/src/Realizability/Assembly/Everything.agda index d91566f..549c4a2 100644 --- a/src/Realizability/Assembly/Everything.agda +++ b/src/Realizability/Assembly/Everything.agda @@ -7,4 +7,5 @@ open import Realizability.Assembly.Coequalizers open import Realizability.Assembly.Equalizers open import Realizability.Assembly.Exponentials open import Realizability.Assembly.Morphism -open import Realizability.Assembly.Regular.Everything +-- TODO : Fix regular structure modules +-- open import Realizability.Assembly.Regular.Everything diff --git a/src/Realizability/Assembly/Exponentials.agda b/src/Realizability/Assembly/Exponentials.agda index 901ab18..114ed1a 100644 --- a/src/Realizability/Assembly/Exponentials.agda +++ b/src/Realizability/Assembly/Exponentials.agda @@ -1,9 +1,11 @@ {-# OPTIONS --cubical --allow-unsolved-metas #-} open import Cubical.Foundations.Prelude open import Cubical.Data.Sigma +open import Cubical.Data.FinData hiding (eq) open import Cubical.HITs.PropositionalTruncation hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure module Realizability.Assembly.Exponentials {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where @@ -81,7 +83,20 @@ module _ {X Y Z : Type ℓ} .tracker → do (f~ , f~tracker) ← f .tracker -- λ* x. λ* y. f~ ⨾ (pair ⨾ x ⨾ y) - return ({!!} , (λ z zᵣ zᵣ⊩z x xᵣ xᵣ⊩x → {!!}))) + let + realizer : Term as 2 + realizer = ` f~ ̇ (` pair ̇ # one ̇ # zero) + return + (λ*2 realizer , + (λ z a a⊩z x b b⊩x → + subst + (λ r' → r' ⊩[ ys ] (f .map (z , x))) + (sym (λ*2ComputationRule realizer a b)) + (f~tracker + (z , x) + (pair ⨾ a ⨾ b) + ((subst (λ r' → r' ⊩[ zs ] z) (sym (pr₁pxy≡x _ _)) a⊩z) , + (subst (λ r' → r' ⊩[ xs ] x) (sym (pr₂pxy≡y _ _)) b⊩x)))))) (AssemblyMorphism≡ _ _ (funExt (λ (z , x) → refl))) (λ g → isSetAssemblyMorphism _ _ (⟪ g , identityMorphism xs ⟫ ⊚ theEval) f) λ g g×id⊚eval≡f → AssemblyMorphism≡ _ _ diff --git a/src/Realizability/Assembly/Morphism.agda b/src/Realizability/Assembly/Morphism.agda index bac0774..9cf92b7 100644 --- a/src/Realizability/Assembly/Morphism.agda +++ b/src/Realizability/Assembly/Morphism.agda @@ -3,11 +3,15 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv open import Cubical.Data.Sigma +open import Cubical.Data.FinData open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Reflection.RecordEquiv open import Cubical.Categories.Category open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure module Realizability.Assembly.Morphism {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where @@ -31,6 +35,8 @@ module _ {X Y : Type ℓ} {xs : Assembly X} {ys : Assembly Y} (t : A) (f : X → ys .⊩isPropValued (t ⨾ aₓ) (f x) record AssemblyMorphism {X Y : Type ℓ} (as : Assembly X) (bs : Assembly Y) : Type ℓ where + no-eta-equality + constructor makeAssemblyMorphism open Assembly as renaming (_⊩_ to _⊩X_) open Assembly bs renaming (_⊩_ to _⊩Y_) field @@ -55,6 +61,9 @@ module _ {X Y : Type ℓ} (xs : Assembly X) (ys : Assembly Y) where isSetAssemblyMorphism : isSet (AssemblyMorphism xs ys) isSetAssemblyMorphism = subst (λ t → isSet t) (sym AssemblyMorphism≡Σ) isSetAssemblyMorphismΣ +AssemblyMorphism≡Equiv : ∀ {X Y : Type ℓ} {xs : Assembly X} {ys : Assembly Y} (f g : AssemblyMorphismΣ xs ys) → (f .fst ≡ g .fst) ≃ (f ≡ g) +AssemblyMorphism≡Equiv {X} {Y} {xs} {ys} f g = Σ≡PropEquiv λ _ → isPropPropTrunc + AssemblyMorphismΣ≡ : {X Y : Type ℓ} {xs : Assembly X} {ys : Assembly Y} @@ -73,7 +82,11 @@ module _ {X Y : Type ℓ} thePath = AssemblyMorphismΣ≡ {X = X} {Y = Y} {xs = xs} {ys = ys} open Iso AssemblyMorphism≡ : (f .map ≡ g .map) → f ≡ g - AssemblyMorphism≡ fmap≡gmap i = theIso .inv (thePath (theIso .fun f) (theIso .fun g) (fmap≡gmap) i) + map (AssemblyMorphism≡ fmap≡gmap i) x = fmap≡gmap i x + tracker (AssemblyMorphism≡ fmap≡gmap i) = + isProp→PathP + (λ i → isPropPropTrunc {A = Σ[ t ∈ A ] (∀ x aₓ → aₓ ⊩[ xs ] x → (t ⨾ aₓ) ⊩[ ys ] (fmap≡gmap i x))}) + (f .tracker) (g .tracker) i identityMorphism : {X : Type ℓ} (as : Assembly X) → AssemblyMorphism as as identityMorphism as .map x = x @@ -84,34 +97,17 @@ compositeMorphism : {X Y Z : Type ℓ} {xs : Assembly X} {ys : Assembly Y} {zs : → (g : AssemblyMorphism ys zs) → AssemblyMorphism xs zs compositeMorphism f g .map x = g .map (f .map x) -compositeMorphism {X} {Y} {Z} {xs} {ys} {zs} f g .tracker = map2 untruncated (f .tracker) (g .tracker) where - open Assembly xs renaming (_⊩_ to _⊩X_) - open Assembly ys renaming (_⊩_ to _⊩Y_) - open Assembly zs renaming (_⊩_ to _⊩Z_) - module _ (fTracker : Σ[ f~ ∈ A ] tracks {xs = xs} {ys = ys} f~ (f .map)) - (gTracker : Σ[ g~ ∈ A ] tracks {xs = ys} {ys = zs} g~ (g .map)) where - - f~ = fTracker .fst - f~tracks = fTracker .snd - - g~ = gTracker .fst - g~tracks = gTracker .snd - - easierVariant : ∀ x aₓ aₓ⊩x → (g~ ⨾ (f~ ⨾ aₓ)) ⊩Z g .map (f .map x) - easierVariant x aₓ aₓ⊩x = g~tracks (f .map x) (f~ ⨾ aₓ) (f~tracks x aₓ aₓ⊩x) - - goal : ∀ (x : X) (aₓ : A) (aₓ⊩x : aₓ ⊩X x) - → (B g~ f~ ⨾ aₓ) ⊩Z (compositeMorphism f g .map x) - goal x aₓ aₓ⊩x = subst (λ y → y ⊩Z g .map (f .map x)) - (sym (Ba≡gfa g~ f~ aₓ)) - (easierVariant x aₓ aₓ⊩x) - - untruncated : Σ[ t ∈ A ] - ((x : X) (aₓ : A) - → aₓ ⊩X x - → (t ⨾ aₓ) ⊩Z (compositeMorphism f g) .map x) - untruncated = B g~ f~ , goal - +compositeMorphism {X} {Y} {Z} {xs} {ys} {zs} f g .tracker = + do + (f~ , isTrackedF) ← f .tracker + (g~ , isTrackedG) ← g .tracker + let + realizer : Term as 1 + realizer = ` g~ ̇ (` f~ ̇ # zero) + return + (λ* realizer , + (λ x aₓ aₓ⊩x → subst (λ r' → r' ⊩[ zs ] (g .map (f .map x))) (sym (λ*ComputationRule realizer aₓ)) (isTrackedG (f .map x) (f~ ⨾ aₓ) (isTrackedF x aₓ aₓ⊩x)))) + infixl 23 _⊚_ _⊚_ : {X Y Z : Type ℓ} → {xs : Assembly X} {ys : Assembly Y} {zs : Assembly Z} → AssemblyMorphism xs ys diff --git a/src/Realizability/PERs/PER.agda b/src/Realizability/PERs/PER.agda index a686f24..38bcb96 100644 --- a/src/Realizability/PERs/PER.agda +++ b/src/Realizability/PERs/PER.agda @@ -6,15 +6,28 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure using (⟨_⟩; str) open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Powerset open import Cubical.Functions.FunExtEquiv open import Cubical.Relation.Binary open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category module Realizability.PERs.PER {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open CombinatoryAlgebra ca +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) + +module BR = BinaryRelation + isPartialEquivalenceRelation : PropRel A A ℓ → Type _ -isPartialEquivalenceRelation (rel , isPropValuedRel) = BinaryRelation.isSym rel × BinaryRelation.isTrans rel +isPartialEquivalenceRelation (rel , isPropValuedRel) = BR.isSym rel × BR.isTrans rel isPropIsPartialEquivalenceRelation : ∀ r → isProp (isPartialEquivalenceRelation r) isPropIsPartialEquivalenceRelation (rel , isPropValuedRel) = @@ -26,57 +39,122 @@ record PER : Type (ℓ-suc ℓ) where no-eta-equality constructor makePER field - relation : A → A → Type ℓ - isPropValuedRelation : ∀ x y → isProp (relation x y) - isPER : isPartialEquivalenceRelation (relation , isPropValuedRelation) + relation : PropRel A A ℓ + isPER : isPartialEquivalenceRelation relation + ∣_∣ = relation .fst isSymmetric = isPER .fst isTransitive = isPER .snd + isPropValued = relation .snd open PER -PERΣ : Type (ℓ-suc ℓ) -PERΣ = Σ[ relation ∈ (A → A → hProp ℓ) ] isPartialEquivalenceRelation ((λ a b → ⟨ relation a b ⟩) , λ a b → str (relation a b)) - -IsoPERΣPER : Iso PERΣ PER -PER.relation (Iso.fun IsoPERΣPER (relation , isPER)) x y = ⟨ relation x y ⟩ -PER.isPropValuedRelation (Iso.fun IsoPERΣPER (relation , isPER)) x y = str (relation x y) -PER.isPER (Iso.fun IsoPERΣPER (relation , isPER)) = isPER -Iso.inv IsoPERΣPER per = (λ x y → per .relation x y , per .isPropValuedRelation x y) , (isSymmetric per) , (isTransitive per) --- refl does not work because of no-eta-equality maybe? -relation (Iso.rightInv IsoPERΣPER per i) = λ a b → per .relation a b -isPropValuedRelation (Iso.rightInv IsoPERΣPER per i) = λ a b → per .isPropValuedRelation a b -isPER (Iso.rightInv IsoPERΣPER per i) = (isSymmetric per) , (isTransitive per) -Iso.leftInv IsoPERΣPER perΣ = - Σ≡Prop - (λ x → isPropIsPartialEquivalenceRelation ((λ a b → ⟨ x a b ⟩) , (λ a b → str (x a b)))) - (funExt₂ λ a b → Σ≡Prop (λ x → isPropIsProp) refl) - -PERΣ≃PER : PERΣ ≃ PER -PERΣ≃PER = isoToEquiv IsoPERΣPER - -isSetPERΣ : isSet PERΣ -isSetPERΣ = isSetΣ (isSet→ (isSet→ isSetHProp)) (λ rel → isProp→isSet (isPropIsPartialEquivalenceRelation ((λ a b → ⟨ rel a b ⟩) , (λ a b → str (rel a b))))) - -isSetPER : isSet PER -isSetPER = isOfHLevelRespectEquiv 2 PERΣ≃PER isSetPERΣ - -module ResizedPER (resizing : hPropResizing ℓ) where - private - smallHProp = resizing .fst - hProp≃smallHProp = resizing .snd - smallHProp≃hProp = invEquiv hProp≃smallHProp - - ResizedPER : Type ℓ - ResizedPER = Σ[ relation ∈ (A → A → smallHProp) ] isPartialEquivalenceRelation ((λ a b → ⟨ equivFun (smallHProp≃hProp) (relation a b) ⟩) , λ a b → str (equivFun (smallHProp≃hProp) (relation a b))) - - PERΣ≃ResizedPER : PERΣ ≃ ResizedPER - PERΣ≃ResizedPER = - Σ-cong-equiv-prop - (equiv→ (idEquiv A) (equiv→ (idEquiv A) hProp≃smallHProp)) - (λ relation → isPropIsPartialEquivalenceRelation ((λ a b → ⟨ relation a b ⟩) , (λ a b → str (relation a b)))) - (λ resizedRelation → isPropIsPartialEquivalenceRelation ((λ a b → ⟨ equivFun (smallHProp≃hProp) (resizedRelation a b) ⟩) , λ a b → str (equivFun smallHProp≃hProp (resizedRelation a b)))) - (λ relation isPERRelation → (λ a b aRb → {!!}) , λ a b c aRb bRc → {!!}) - λ relation isPERRelation → {!!} - - PER≃ResizedPER : PER ≃ ResizedPER - PER≃ResizedPER = compEquiv (invEquiv PERΣ≃PER) PERΣ≃ResizedPER +module _ (R : PER) where + Quotient = A / ∣ R ∣ + + -- mimics the proof at Cubical.HITs.SetQuotients.Properties + effectiveOnDomain : ∀ a b → ∣ R ∣ a a → [ a ] ≡ [ b ] → ∣ R ∣ a b + effectiveOnDomain a b aRa [a]≡[b] = transport aRa≡aRb aRa where + helper : Quotient → hProp _ + helper x = + SQ.rec + isSetHProp + (λ c → (∣ R ∣ a c) , (isPropValued R a c)) + (λ c d cRd → + Σ≡Prop + (λ _ → isPropIsProp) + (hPropExt + (isPropValued R a c) + (isPropValued R a d) + (λ aRc → isTransitive R a c d aRc cRd) + (λ aRd → isTransitive R a d c aRd (isSymmetric R c d cRd)))) + x + + aRa≡aRb : ∣ R ∣ a a ≡ ∣ R ∣ a b + aRa≡aRb i = helper ([a]≡[b] i) .fst + +record PERMorphism (R S : PER) : Type ℓ where + no-eta-equality + constructor makePERMorphism + field + map : Quotient R → Quotient S + isTracked : ∃[ e ∈ A ] (∀ (a : A) → ∣ R ∣ a a → (map [ a ] ≡ [ e ⨾ a ]) × ∣ S ∣ (e ⨾ a) (e ⨾ a)) + +open PERMorphism + +unquoteDecl PERMorphismIsoΣ = declareRecordIsoΣ PERMorphismIsoΣ (quote PERMorphism) + +PERMorphismΣ : PER → PER → Type ℓ +PERMorphismΣ R S = Σ[ map ∈ (Quotient R → Quotient S) ] ∃[ e ∈ A ] (∀ (a : A) → ∣ R ∣ a a → (map [ a ] ≡ [ e ⨾ a ]) × ∣ S ∣ (e ⨾ a) (e ⨾ a)) + +isSetPERMorphismΣ : ∀ {R S} → isSet (PERMorphismΣ R S) +isSetPERMorphismΣ {R} {S} = isSetΣ (isSet→ squash/) (λ map → isProp→isSet isPropPropTrunc) + +isSetPERMorphism : ∀ {R S} → isSet (PERMorphism R S) +isSetPERMorphism {R} {S} = subst (λ type → isSet type) (sym (isoToPath PERMorphismIsoΣ)) (isSetPERMorphismΣ {R = R} {S = S}) + +PERMorphism≡Iso : ∀ {R S} → (f g : PERMorphism R S) → Iso (f ≡ g) (f .map ≡ g .map) +Iso.fun (PERMorphism≡Iso {R} {S} f g) f≡g i = f≡g i .map +map (Iso.inv (PERMorphism≡Iso {R} {S} f g) fMap≡gMap i) = fMap≡gMap i +isTracked (Iso.inv (PERMorphism≡Iso {R} {S} f g) fMap≡gMap i) = + isProp→PathP + (λ i → + isPropPropTrunc + {A = Σ[ e ∈ A ] (∀ (a : A) → ∣ R ∣ a a → ((fMap≡gMap i) [ a ] ≡ [ e ⨾ a ]) × ∣ S ∣ (e ⨾ a) (e ⨾ a))}) + (f .isTracked) (g .isTracked) i +Iso.rightInv (PERMorphism≡Iso {R} {S} f g) fMap≡gMap = refl +Iso.leftInv (PERMorphism≡Iso {R} {S} f g) f≡g = isSetPERMorphism f g (Iso.inv (PERMorphism≡Iso f g) (λ i → f≡g i .map)) f≡g + +PERMorphism≡ : ∀ {R S} → (f g : PERMorphism R S) → f .map ≡ g .map → f ≡ g +PERMorphism≡ {R} {S} f g fMap≡gMap = Iso.inv (PERMorphism≡Iso f g) fMap≡gMap + +idPERMorphism : ∀ {R : PER} → PERMorphism R R +map (idPERMorphism {R}) x = x +isTracked (idPERMorphism {R}) = + return (Id , (λ a aRa → (subst (λ r → [ a ] ≡ [ r ]) (sym (Ida≡a a)) refl) , (subst (λ r → ∣ R ∣ r r) (sym (Ida≡a a)) aRa))) + +composePERMorphism : ∀ {R S T : PER} → PERMorphism R S → PERMorphism S T → PERMorphism R T +map (composePERMorphism {R} {S} {T} f g) x = g .map (f .map x) +isTracked (composePERMorphism {R} {S} {T} f g) = + do + (f~ , isTrackedF) ← f .isTracked + (g~ , isTrackedG) ← g .isTracked + let + realizer : Term as 1 + realizer = ` g~ ̇ (` f~ ̇ # zero) + return + (λ* realizer , + (λ a aRa → + (g .map (f .map [ a ]) + ≡⟨ cong (g .map) (isTrackedF a aRa .fst) ⟩ + g .map [ f~ ⨾ a ] + ≡⟨ isTrackedG (f~ ⨾ a) (isTrackedF a aRa .snd) .fst ⟩ + [ g~ ⨾ (f~ ⨾ a) ] + ≡⟨ cong [_] (sym (λ*ComputationRule realizer a)) ⟩ + [ λ* realizer ⨾ a ] + ∎) , + (subst (λ r' → ∣ T ∣ r' r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~ ⨾ a) (isTrackedF a aRa .snd) .snd)))) + +-- all definitional +idLPERMorphism : ∀ {R S} → (f : PERMorphism R S) → composePERMorphism idPERMorphism f ≡ f +idLPERMorphism {R} {S} f = PERMorphism≡ _ _ refl + +idRPERMorphism : ∀ {R S} → (f : PERMorphism R S) → composePERMorphism f idPERMorphism ≡ f +idRPERMorphism {R} {S} f = PERMorphism≡ _ _ refl + +assocPERMorphism : + ∀ {R S T U} + → (f : PERMorphism R S) + → (g : PERMorphism S T) + → (h : PERMorphism T U) + → composePERMorphism (composePERMorphism f g) h ≡ composePERMorphism f (composePERMorphism g h) +assocPERMorphism {R} {S} {T} {U} f g h = PERMorphism≡ _ _ refl + +PERCat : Category (ℓ-suc ℓ) ℓ +Category.ob PERCat = PER +Category.Hom[_,_] PERCat R S = PERMorphism R S +Category.id PERCat {R} = idPERMorphism +Category._⋆_ PERCat {R} {S} {T} f g = composePERMorphism f g +Category.⋆IdL PERCat f = idLPERMorphism f +Category.⋆IdR PERCat f = idRPERMorphism f +Category.⋆Assoc PERCat f g h = assocPERMorphism f g h +Category.isSetHom PERCat = isSetPERMorphism From bd065f5ed75999b3b072dc3d73ee2b019ae25abf Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 15 May 2024 17:14:39 +0530 Subject: [PATCH 44/61] Start working towards LCCC on Mod --- src/Realizability/Assembly/Base.agda | 1 + src/Realizability/Assembly/Pullbacks.agda | 159 ++++++++++++++++++ .../Assembly/SetsReflectiveSubcategory.agda | 62 +++++++ src/Realizability/Assembly/Terminal.agda | 49 ++++++ src/Realizability/Modest/Base.agda | 78 +++++++++ src/Realizability/Modest/FamilyOverAsm.agda | 63 +++++++ .../Modest/PartialSurjection.agda | 150 +++++++++++++++++ src/Realizability/PERs/#TerminalObject.agda# | 43 +++++ src/Realizability/PERs/.#TerminalObject.agda | 1 + src/Realizability/PERs/PER.agda | 17 ++ src/Realizability/PERs/TerminalObject.agda | 43 +++++ src/Realizability/PropResizing.agda | 54 ++++++ 12 files changed, 720 insertions(+) create mode 100644 src/Realizability/Assembly/Pullbacks.agda create mode 100644 src/Realizability/Assembly/SetsReflectiveSubcategory.agda create mode 100644 src/Realizability/Assembly/Terminal.agda create mode 100644 src/Realizability/Modest/Base.agda create mode 100644 src/Realizability/Modest/FamilyOverAsm.agda create mode 100644 src/Realizability/Modest/PartialSurjection.agda create mode 100644 src/Realizability/PERs/#TerminalObject.agda# create mode 120000 src/Realizability/PERs/.#TerminalObject.agda create mode 100644 src/Realizability/PERs/TerminalObject.agda diff --git a/src/Realizability/Assembly/Base.agda b/src/Realizability/Assembly/Base.agda index 8ddc335..ae68e6c 100644 --- a/src/Realizability/Assembly/Base.agda +++ b/src/Realizability/Assembly/Base.agda @@ -13,6 +13,7 @@ open import Realizability.CombinatoryAlgebra module Realizability.Assembly.Base {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where record Assembly (X : Type ℓ) : Type (ℓ-suc ℓ) where + constructor makeAssembly infix 25 _⊩_ field _⊩_ : A → X → Type ℓ diff --git a/src/Realizability/Assembly/Pullbacks.agda b/src/Realizability/Assembly/Pullbacks.agda new file mode 100644 index 0000000..d2bb815 --- /dev/null +++ b/src/Realizability/Assembly/Pullbacks.agda @@ -0,0 +1,159 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor +open import Cubical.Categories.Constructions.Slice +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure + +module Realizability.Assembly.Pullbacks {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) hiding (Z) +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca + +module _ (cspn : Cospan ASM) where + open Cospan cspn + + X = l .fst + xs = l .snd + + Y = m .fst + ys = m .snd + + Z = r .fst + zs = r .snd + + f = s₁ + g = s₂ + + opaque + pullbackType : Type ℓ + pullbackType = (Σ[ x ∈ X ] Σ[ z ∈ Z ] (f .map x ≡ g .map z)) + + opaque + unfolding pullbackType + pullbackAsm : Assembly pullbackType + Assembly._⊩_ pullbackAsm = λ { r (x , z , fx≡gz) → (pr₁ ⨾ r) ⊩[ xs ] x × ((pr₂ ⨾ r) ⊩[ zs ] z) } + Assembly.isSetX pullbackAsm = isSetΣ (xs .isSetX) (λ _ → isSetΣ (zs .isSetX) (λ _ → isProp→isSet (ys .isSetX _ _))) + Assembly.⊩isPropValued pullbackAsm = λ { r (x , z , fx≡gz) → isProp× (xs .⊩isPropValued _ _) (zs .⊩isPropValued _ _) } + Assembly.⊩surjective pullbackAsm = + (λ { (x , z , fx≡gz) → + do + (a , a⊩x) ← xs .⊩surjective x + (b , b⊩z) ← zs .⊩surjective z + return + (pair ⨾ a ⨾ b , + subst (λ r' → r' ⊩[ xs ] x) (sym (pr₁pxy≡x _ _)) a⊩x , + subst (λ r' → r' ⊩[ zs ] z) (sym (pr₂pxy≡y _ _)) b⊩z) }) + + opaque + unfolding pullbackType + unfolding pullbackAsm + pullbackPr₁ : AssemblyMorphism pullbackAsm xs + AssemblyMorphism.map pullbackPr₁ (x , z , fx≡gz) = x + AssemblyMorphism.tracker pullbackPr₁ = + return (pr₁ , (λ { (x , z , fx≡gz) r (pr₁r⊩x , pr₂r⊩z) → pr₁r⊩x })) + + opaque + unfolding pullbackType + unfolding pullbackAsm + pullbackPr₂ : AssemblyMorphism pullbackAsm zs + AssemblyMorphism.map pullbackPr₂ (x , z , fx≡gz) = z + AssemblyMorphism.tracker pullbackPr₂ = + return (pr₂ , (λ { (x , z , fx≡gz) r (pr₁r⊩x , pr₂r⊩z) → pr₂r⊩z })) + + opaque + unfolding pullbackAsm + unfolding pullbackPr₁ + unfolding pullbackPr₂ + pullback : Pullback ASM cspn + Pullback.pbOb pullback = pullbackType , pullbackAsm + Pullback.pbPr₁ pullback = pullbackPr₁ + Pullback.pbPr₂ pullback = pullbackPr₂ + Pullback.pbCommutes pullback = AssemblyMorphism≡ _ _ (funExt λ { (x , z , fx≡gz) → fx≡gz }) + Pullback.univProp pullback {D , ds} p q pf≡qg = + uniqueExists + uniqueMorphism + ((AssemblyMorphism≡ _ _ (funExt (λ d → refl))) , (AssemblyMorphism≡ _ _ (funExt (λ d → refl)))) + (λ !' → isProp× (isSetAssemblyMorphism _ _ p (!' ⊚ pullbackPr₁)) (isSetAssemblyMorphism _ _ q (!' ⊚ pullbackPr₂))) + λ { !' (p≡!'*pr₁ , q≡!'*pr₂) → + AssemblyMorphism≡ + _ _ + (funExt + λ d → + ΣPathP + ((λ i → p≡!'*pr₁ i .map d) , + ΣPathPProp + {u = q .map d , λ i → pf≡qg i .map d} + {v = !' .map d .snd} + (λ z → ys .isSetX _ _) + λ i → q≡!'*pr₂ i .map d)) } + where + uniqueMap : D → pullbackType + uniqueMap d = p .map d , q .map d , λ i → pf≡qg i .map d + + uniqueMorphism : AssemblyMorphism ds pullbackAsm + uniqueMorphism = + (makeAssemblyMorphism + uniqueMap + (do + (p~ , isTrackedP) ← p .tracker + (q~ , isTrackedQ) ← q .tracker + let + realizer : Term as 1 + realizer = ` pair ̇ (` p~ ̇ # zero) ̇ (` q~ ̇ # zero) + return + (λ* realizer , + λ d r r⊩d → + subst (λ r' → r' ⊩[ xs ] (p .map d)) (sym (cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₁pxy≡x _ _)) (isTrackedP d r r⊩d) , + subst (λ r' → r' ⊩[ zs ] (q .map d)) (sym (cong (λ x → pr₂ ⨾ x) (λ*ComputationRule realizer r) ∙ pr₂pxy≡y _ _)) (isTrackedQ d r r⊩d)))) + +PullbacksASM : Pullbacks ASM +PullbacksASM cspn = pullback cspn + +-- Using pullbacks we get a functor that sends any f : X → Y to f* : Asm / Y → Asm / X +module _ {X Y : Type ℓ} (asmX : Assembly X) (asmY : Assembly Y) (f : AssemblyMorphism asmX asmY) where + open Pullback + opaque + unfolding pullbackAsm + unfolding pullbackPr₁ + unfolding pullback + pullbackFunctor : Functor (SliceCat ASM (Y , asmY)) (SliceCat ASM (X , asmX)) + Functor.F-ob pullbackFunctor sliceY = sliceob (pullback (cospan (X , asmX) (Y , asmY) (sliceY .S-ob) f (sliceY .S-arr)) .pbPr₁) + Functor.F-hom pullbackFunctor {ySliceA} {ySliceB} sliceHomAB = + let + sliceACospan : Cospan ASM + sliceACospan = cospan (X , asmX) (Y , asmY) (ySliceA .S-ob) f (ySliceA .S-arr) + p = pullbackPr₂ sliceACospan + m = ySliceA .S-arr + n = ySliceB .S-arr + f*m = pullbackPr₁ sliceACospan + h = sliceHomAB .S-hom + m≡h⊚n = sliceHomAB .S-comm + f*m⊚f≡p⊚m = pbCommutes (pullback sliceACospan) + arrow = + pullbackArrow + ASM + (pullback (cospan (X , asmX) (Y , asmY) (ySliceB .S-ob) f (ySliceB .S-arr))) (pullbackPr₁ sliceACospan) (p ⊚ h) + (AssemblyMorphism≡ _ _ + (funExt + λ { (x , a , fx≡ma) → + f .map x + ≡⟨ fx≡ma ⟩ + m .map a + ≡[ i ]⟨ m≡h⊚n (~ i) .map a ⟩ + n .map (h .map a) + ∎ })) + in + slicehom + arrow + {!!} + Functor.F-id pullbackFunctor = {!!} + Functor.F-seq pullbackFunctor = {!!} diff --git a/src/Realizability/Assembly/SetsReflectiveSubcategory.agda b/src/Realizability/Assembly/SetsReflectiveSubcategory.agda new file mode 100644 index 0000000..8a2f295 --- /dev/null +++ b/src/Realizability/Assembly/SetsReflectiveSubcategory.agda @@ -0,0 +1,62 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Adjoint +open import Cubical.Categories.NaturalTransformation +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure + +module Realizability.Assembly.SetsReflectiveSubcategory {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) + +forgetfulFunctor : Functor ASM (SET ℓ) +Functor.F-ob forgetfulFunctor (X , asmX) = X , asmX .isSetX +Functor.F-hom forgetfulFunctor {X , asmX} {Y , asmY} f = f .map +Functor.F-id forgetfulFunctor = refl +Functor.F-seq forgetfulFunctor {X , asmX} {Y , asmY} {Z , asmZ} f g = refl + +∇ : Functor (SET ℓ) ASM +Functor.F-ob ∇ (X , isSetX) = X , makeAssembly (λ a x → Unit*) isSetX (λ _ _ → isPropUnit*) λ x → ∣ k , tt* ∣₁ +Functor.F-hom ∇ {X , isSetX} {Y , isSetY} f = makeAssemblyMorphism f (return (k , (λ _ _ _ → tt*))) +Functor.F-id ∇ {X , isSetX} = AssemblyMorphism≡ _ _ refl +Functor.F-seq ∇ {X , isSetX} {Y , isSetY} {Z , isSetZ} f g = AssemblyMorphism≡ _ _ refl + +module _ where + open UnitCounit + + adjointUnitCounit : forgetfulFunctor ⊣ ∇ + NatTrans.N-ob (_⊣_.η adjointUnitCounit) (X , asmX) = makeAssemblyMorphism (λ x → x) (return (k , (λ _ _ _ → tt*))) + NatTrans.N-hom (_⊣_.η adjointUnitCounit) {X , asmX} {Y , asmY} f = AssemblyMorphism≡ _ _ refl + NatTrans.N-ob (_⊣_.ε adjointUnitCounit) (X , isSetX) x = x + NatTrans.N-hom (_⊣_.ε adjointUnitCounit) {X , isSetX} {Y , isSetY} f = refl + TriangleIdentities.Δ₁ (_⊣_.triangleIdentities adjointUnitCounit) (X , asmX) = refl + TriangleIdentities.Δ₂ (_⊣_.triangleIdentities adjointUnitCounit) (X , isSetX) = AssemblyMorphism≡ _ _ refl + +module _ where + open NaturalBijection + + adjointNaturalBijection : forgetfulFunctor ⊣ ∇ + Iso.fun (_⊣_.adjIso adjointNaturalBijection) f = makeAssemblyMorphism f (return (k , (λ x r r⊩x → tt*))) + Iso.inv (_⊣_.adjIso adjointNaturalBijection) f = f .map + Iso.rightInv (_⊣_.adjIso adjointNaturalBijection) b = AssemblyMorphism≡ _ _ refl + Iso.leftInv (_⊣_.adjIso adjointNaturalBijection) a = refl + _⊣_.adjNatInD adjointNaturalBijection {X , isSetX} {Y , isSetY} f g = AssemblyMorphism≡ _ _ refl + _⊣_.adjNatInC adjointNaturalBijection {X , asmX} {Y , asmY} f g = refl + diff --git a/src/Realizability/Assembly/Terminal.agda b/src/Realizability/Assembly/Terminal.agda new file mode 100644 index 0000000..64fb1c6 --- /dev/null +++ b/src/Realizability/Assembly/Terminal.agda @@ -0,0 +1,49 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Categories.Limits.Terminal +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure + +module Realizability.Assembly.Terminal {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open CombinatoryAlgebra ca +open Assembly +open AssemblyMorphism + +terminalAsm : Assembly Unit* +(Assembly._⊩_ terminalAsm) a tt* = Unit* +Assembly.isSetX terminalAsm = isSetUnit* +(Assembly.⊩isPropValued terminalAsm) a tt* = isPropUnit* +Assembly.⊩surjective terminalAsm tt* = ∣ k , tt* ∣₁ + +isTerminalTerminalAsm : isTerminal ASM (Unit* , terminalAsm) +isTerminalTerminalAsm (X , asmX) = + inhProp→isContr + (makeAssemblyMorphism + (λ x → tt*) + (return + (k , (λ x r r⊩x → tt*)))) + (λ f g → + AssemblyMorphism≡ _ _ (funExt λ x → refl)) + +TerminalASM : Terminal ASM +fst TerminalASM = Unit* , terminalAsm +snd TerminalASM = isTerminalTerminalAsm + +-- global element +module _ {X : Type ℓ} (asmX : Assembly X) (x : X) (r : A) (r⊩x : r ⊩[ asmX ] x) where + + globalElement : AssemblyMorphism terminalAsm asmX + AssemblyMorphism.map globalElement tt* = x + AssemblyMorphism.tracker globalElement = + return + ((k ⨾ r) , + (λ { tt* a tt* → subst (λ r' → r' ⊩[ asmX ] x) (sym (kab≡a _ _)) r⊩x })) diff --git a/src/Realizability/Modest/Base.agda b/src/Realizability/Modest/Base.agda new file mode 100644 index 0000000..e2e9b1c --- /dev/null +++ b/src/Realizability/Modest/Base.agda @@ -0,0 +1,78 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.Base {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) + +module _ {X : Type ℓ} (asmX : Assembly X) where + + isModest : Type _ + isModest = ∀ (x y : X) (a : A) → a ⊩[ asmX ] x → a ⊩[ asmX ] y → x ≡ y + + isPropIsModest : isProp isModest + isPropIsModest = isPropΠ3 λ x y a → isProp→ (isProp→ (asmX .isSetX x y)) + + isUniqueRealised : isModest → ∀ (a : A) → isProp (Σ[ x ∈ X ] (a ⊩[ asmX ] x)) + isUniqueRealised isMod a (x , a⊩x) (y , a⊩y) = Σ≡Prop (λ x' → asmX .⊩isPropValued a x') (isMod x y a a⊩x a⊩y) + +ModestSet : Type ℓ → Type (ℓ-suc ℓ) +ModestSet X = Σ[ xs ∈ Assembly X ] isModest xs + +MOD : Category (ℓ-suc ℓ) ℓ +MOD = ΣPropCat ASM λ { (X , asmX) → (Lift (isModest asmX)) , (isOfHLevelRespectEquiv 1 (LiftEquiv {A = isModest asmX}) (isPropIsModest asmX)) } + +-- Every modest set is isomorphic to a canonically modest set +module Canonical (X : Type ℓ) (asmX : Assembly X) (isModestAsmX : isModest asmX) (resizing : hPropResizing ℓ) where + open ResizedPowerset resizing + -- Replace every term of X by it's set of realisers + realisersOf : X → ℙ A + realisersOf x a = (a ⊩[ asmX ] x) , (asmX .⊩isPropValued a x) + + resizedRealisersOf : X → 𝓟 A + resizedRealisersOf x = ℙ→𝓟 A (realisersOf x) + + realiserSet : Type ℓ + realiserSet = Σ[ P ∈ 𝓟 A ] ∃[ x ∈ X ] P ≡ resizedRealisersOf x + + canonicalModestSet : Assembly realiserSet + Assembly._⊩_ canonicalModestSet r (P , ∃x) = r ϵ P + Assembly.isSetX canonicalModestSet = isSetΣ (isSet𝓟 A) (λ P → isProp→isSet isPropPropTrunc) + Assembly.⊩isPropValued canonicalModestSet r (P , ∃x) = isPropϵ r P + Assembly.⊩surjective canonicalModestSet (P , ∃x) = + do + (x , P≡⊩x) ← ∃x + (a , a⊩x) ← asmX .⊩surjective x + return + (a , + (subst + (λ P → a ϵ P) + (sym P≡⊩x) + (subst (λ P → a ∈ P) (sym (compIsIdFunc (realisersOf x))) a⊩x))) + {- + isModestCanonicalModestSet : isModest canonicalModestSet + isModestCanonicalModestSet x y a a⊩x a⊩y = + Σ≡Prop + (λ _ → isPropPropTrunc) + (𝓟≡ (x .fst) (y .fst) (⊆-extensionality (𝓟→ℙ A (x .fst)) (𝓟→ℙ A (y .fst)) ((λ b b⊩x → {!!}) , {!!}))) -} + + diff --git a/src/Realizability/Modest/FamilyOverAsm.agda b/src/Realizability/Modest/FamilyOverAsm.agda new file mode 100644 index 0000000..07acadf --- /dev/null +++ b/src/Realizability/Modest/FamilyOverAsm.agda @@ -0,0 +1,63 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Limits.Pullback +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.FamilyOverAsm {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Pullbacks ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Modest.Base ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Pullback + +module _ {X : Type ℓ} (asmX : Assembly X) (Y : Type ℓ) (asmY : Assembly Y) where + record FamilyOver : Type (ℓ-suc ℓ) where + no-eta-equality + constructor makeFamilyOver + field + fibration : AssemblyMorphism asmY asmX + {- + A[x] ------> A + |_| | + | | + | | fibration + | | + ↓ ↓ + 1 --------> X + x + + For any x the fiber over x is modest + -} + fiberOver : + ∀ (x : X) (a : A) → (a⊩x : a ⊩[ asmX ] x) → + Pullback + ASM + (cospan + (Unit* , terminalAsm) + (X , asmX) + (Y , asmY) + (globalElement asmX x a a⊩x) + fibration) + isModestFibreOver : ∀ (x : X) (a : A) (a⊩x : a ⊩[ asmX ] x) → isModest (fiberOver x a a⊩x .pbOb .snd) + diff --git a/src/Realizability/Modest/PartialSurjection.agda b/src/Realizability/Modest/PartialSurjection.agda new file mode 100644 index 0000000..7fb7aac --- /dev/null +++ b/src/Realizability/Modest/PartialSurjection.agda @@ -0,0 +1,150 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Univalence +open import Cubical.Functions.Surjection +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.PartialSurjection {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) (resizing : hPropResizing ℓ) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Modest.Base ca resizing + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open ResizedPowerset resizing + +record PartialSurjection (X : Type ℓ) : Type (ℓ-suc ℓ) where + no-eta-equality + constructor makePartialSurjection + field + support : A → Type ℓ + enumeration : Σ[ a ∈ A ] (support a) → X + isPropSupport : ∀ a → isProp (support a) + isSurjectionEnumeration : isSurjection enumeration + isSetX : isSet X -- potentially redundant? +open PartialSurjection + +-- let us derive a structure of identity principle for partial surjections +module _ (X : Type ℓ) where + + PartialSurjection≡Iso : + ∀ (p q : PartialSurjection X) + → Iso + (Σ[ suppPath ∈ p .support ≡ q .support ] + PathP (λ i → Σ[ a ∈ A ] (suppPath i a) → X) (p .enumeration) (q .enumeration)) + (p ≡ q) + support (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z = suppPath i z + enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) (a , enum) = enumPath i (a , enum) + isPropSupport (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z = + isProp→PathP {B = λ j → isProp (suppPath j z)} (λ j → isPropIsProp) (p .isPropSupport z) (q .isPropSupport z) i + isSurjectionEnumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) b = + isProp→PathP + {B = λ j → ∥ fiber (enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) j)) b ∥₁} + (λ j → isPropPropTrunc) + (p .isSurjectionEnumeration b) (q .isSurjectionEnumeration b) i + isSetX (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) = isPropIsSet (p .isSetX) (q .isSetX) i + Iso.inv (PartialSurjection≡Iso p q) p≡q = (λ i → p≡q i .support) , (λ i → p≡q i .enumeration) + Iso.rightInv (PartialSurjection≡Iso p q) p≡q = {!!} + Iso.leftInv (PartialSurjection≡Iso p q) (suppPath , enumPath) = ΣPathP (refl , refl) + + PartialSurjection≡ : ∀ (p q : PartialSurjection X) → Σ[ suppPath ∈ p .support ≡ q .support ] PathP (λ i → Σ[ a ∈ A ] (suppPath i a) → X) (p .enumeration) (q .enumeration) → p ≡ q + PartialSurjection≡ p q (suppPath , enumPath) = Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) + +-- the type of partial surjections is equivalent to the type of modest assemblies on X +module _ (X : Type ℓ) where + + {-# TERMINATING #-} + ModestSet→PartialSurjection : ModestSet X → PartialSurjection X + support (ModestSet→PartialSurjection (xs , isModestXs)) r = ∃[ x ∈ X ] (r ⊩[ xs ] x) + enumeration (ModestSet→PartialSurjection (xs , isModestXs)) (r , ∃x) = + let + answer : Σ[ x ∈ X ] (r ⊩[ xs ] x) + answer = PT.rec (isUniqueRealised xs isModestXs r) (λ t → t) ∃x + in fst answer + isPropSupport (ModestSet→PartialSurjection (xs , isModestXs)) r = isPropPropTrunc + isSurjectionEnumeration (ModestSet→PartialSurjection (xs , isModestXs)) x = + do + (a , a⊩x) ← xs .⊩surjective x + return ((a , ∣ x , a⊩x ∣₁) , refl) + isSetX (ModestSet→PartialSurjection (xs , isModestXs)) = xs .isSetX + + PartialSurjection→ModestSet : PartialSurjection X → ModestSet X + Assembly._⊩_ (fst (PartialSurjection→ModestSet surj)) r x = + Σ[ s ∈ surj .support r ] surj .enumeration (r , s) ≡ x + Assembly.isSetX (fst (PartialSurjection→ModestSet surj)) = surj .isSetX + Assembly.⊩isPropValued (fst (PartialSurjection→ModestSet surj)) a x (s , ≡x) (t , ≡x') = + Σ≡Prop (λ u → surj .isSetX (surj .enumeration (a , u)) x) (surj .isPropSupport a s t) + Assembly.⊩surjective (fst (PartialSurjection→ModestSet surj)) x = + do + ((a , s) , ≡x) ← surj .isSurjectionEnumeration x + return (a , (s , ≡x)) + snd (PartialSurjection→ModestSet surj) x y r (s , ≡x) (t , ≡x') = + x + ≡⟨ sym ≡x ⟩ + surj .enumeration (r , s) + ≡⟨ cong (λ s → surj .enumeration (r , s)) (surj .isPropSupport r s t) ⟩ + surj .enumeration (r , t) + ≡⟨ ≡x' ⟩ + y + ∎ + + opaque + rightInv : ∀ surj → ModestSet→PartialSurjection (PartialSurjection→ModestSet surj) ≡ surj + rightInv surj = + PartialSurjection≡ + X (ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)) surj + (funExt supportEq , + funExtDep + {A = λ i → Σ-syntax A (funExt supportEq i)} + {B = λ _ _ → X} + {f = ModestSet→PartialSurjection (PartialSurjection→ModestSet surj) .enumeration} + {g = surj .enumeration} + λ { {r , ∃x} {s , supp} p → + PT.elim + {P = λ ∃x → fst + (PT.rec + (isUniqueRealised (fst (PartialSurjection→ModestSet surj)) + (snd (PartialSurjection→ModestSet surj)) r) + (λ t → t) ∃x) + ≡ surj .enumeration (s , supp)} + (λ ∃x → surj .isSetX _ _) + (λ { (x , r⊩x) → + let + ∃x' = transport (sym (supportEq s)) supp + in + equivFun + (propTruncIdempotent≃ {!!}) + (do + (x' , suppS , ≡x') ← ∃x' + return {!!}) }) + ∃x }) where + supportEq : ∀ r → (∃[ x ∈ X ] (Σ[ supp ∈ surj .support r ] (surj .enumeration (r , supp) ≡ x))) ≡ support surj r + supportEq = + (λ r → + hPropExt + isPropPropTrunc + (surj .isPropSupport r) + (λ ∃x → PT.rec (surj .isPropSupport r) (λ { (x , supp , ≡x) → supp }) ∃x) + (λ supp → return (surj .enumeration (r , supp) , supp , refl))) + + IsoModestSetPartialSurjection : Iso (ModestSet X) (PartialSurjection X) + Iso.fun IsoModestSetPartialSurjection = ModestSet→PartialSurjection + Iso.inv IsoModestSetPartialSurjection = PartialSurjection→ModestSet + Iso.rightInv IsoModestSetPartialSurjection = rightInv + Iso.leftInv IsoModestSetPartialSurjection mod = {!!} diff --git a/src/Realizability/PERs/#TerminalObject.agda# b/src/Realizability/PERs/#TerminalObject.agda# new file mode 100644 index 0000000..b573343 --- /dev/null +++ b/src/Realizability/PERs/#TerminalObject.agda# @@ -0,0 +1,43 @@ +open import Realizability.ApplicativeStructure +open import Realizability.CombinatoryAlgebra +open import Realizability.PropResizing +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Powerset +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Binary +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.Terminal + +module Realizability.PERs.TerminalObject + {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.PERs.PER ca +open CombinatoryAlgebra ca +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open PERMorphism + +terminalPER : PER +PER.relation terminalPER = (λ x y → Unit*) , λ x y → isPropUnit* +fst (PER.isPER terminalPER) = λ a b x → tt* +snd (PER.isPER terminalPER) = λ a b c x x₁ → tt* + +isTerminalTerminalPER : isTerminal PERCat terminalPER +isTerminalTerminalPER X = + inhProp→isContr + (makePERMorphism (λ x → [ k ]) (return (Id , (λ a aXa → (eq/ k (Id ⨾ a) tt*) , tt*)))) + λ x y → PERMorphism≡ x y (funExt λ q → {!effectiveOnDomain!}) + +terminal : Terminal PERCat +terminal = terminalPER , {!!} diff --git a/src/Realizability/PERs/.#TerminalObject.agda b/src/Realizability/PERs/.#TerminalObject.agda new file mode 120000 index 0000000..53126fa --- /dev/null +++ b/src/Realizability/PERs/.#TerminalObject.agda @@ -0,0 +1 @@ +rahul@Rahuls-MacBook-Air.local.1114:1710392191 \ No newline at end of file diff --git a/src/Realizability/PERs/PER.agda b/src/Realizability/PERs/PER.agda index 38bcb96..bd5adb7 100644 --- a/src/Realizability/PERs/PER.agda +++ b/src/Realizability/PERs/PER.agda @@ -17,10 +17,14 @@ open import Cubical.HITs.PropositionalTruncation as PT hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.HITs.SetQuotients as SQ open import Cubical.Categories.Category +open import Cubical.Categories.Functor hiding (Id) module Realizability.PERs.PER {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca + open CombinatoryAlgebra ca open Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -158,3 +162,16 @@ Category.⋆IdL PERCat f = idLPERMorphism f Category.⋆IdR PERCat f = idRPERMorphism f Category.⋆Assoc PERCat f g h = assocPERMorphism f g h Category.isSetHom PERCat = isSetPERMorphism + +open Assembly + +inclusion : Functor PERCat ASM +fst (Functor.F-ob inclusion per) = Σ[ a ∈ A ] ∣ per ∣ a a +(snd (Functor.F-ob inclusion per)) ._⊩_ r (a , aRa) = ∣ per ∣ r a +isSetX (snd (Functor.F-ob inclusion per)) = isSetΣ isSetA (λ x → isProp→isSet (isPropValued per x x)) +⊩isPropValued (snd (Functor.F-ob inclusion per)) r (a , aRa) = isPropValued per r a +⊩surjective (snd (Functor.F-ob inclusion per)) (a , aRa) = return (a , aRa) +AssemblyMorphism.map (Functor.F-hom inclusion morphism) (a , aRa) = {!!} +AssemblyMorphism.tracker (Functor.F-hom inclusion morphism) = {!!} +Functor.F-id inclusion = {!!} +Functor.F-seq inclusion = {!!} diff --git a/src/Realizability/PERs/TerminalObject.agda b/src/Realizability/PERs/TerminalObject.agda new file mode 100644 index 0000000..c3a6576 --- /dev/null +++ b/src/Realizability/PERs/TerminalObject.agda @@ -0,0 +1,43 @@ +open import Realizability.ApplicativeStructure +open import Realizability.CombinatoryAlgebra +open import Realizability.PropResizing +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Powerset +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Binary +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.Terminal + +module Realizability.PERs.TerminalObject + {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.PERs.PER ca +open CombinatoryAlgebra ca +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open PERMorphism + +terminalPER : PER +PER.relation terminalPER = (λ x y → Unit*) , λ x y → isPropUnit* +fst (PER.isPER terminalPER) = λ a b x → tt* +snd (PER.isPER terminalPER) = λ a b c x x₁ → tt* + +isTerminalTerminalPER : isTerminal PERCat terminalPER +isTerminalTerminalPER X = + inhProp→isContr + (makePERMorphism (λ x → [ k ]) (return (Id , (λ a aXa → (eq/ k (Id ⨾ a) tt*) , tt*)))) + λ x y → PERMorphism≡ x y (funExt λ q → {!!}) + +terminal : Terminal PERCat +terminal = terminalPER , {!!} diff --git a/src/Realizability/PropResizing.agda b/src/Realizability/PropResizing.agda index 5ba4076..1063210 100644 --- a/src/Realizability/PropResizing.agda +++ b/src/Realizability/PropResizing.agda @@ -1,7 +1,9 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Equiv.Properties open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure +open import Cubical.Foundations.Powerset open import Cubical.Data.Sigma module Realizability.PropResizing where @@ -21,3 +23,55 @@ copyEquiv = snd -- This we call hPropResizing hPropResizing : ∀ ℓ → Type _ hPropResizing ℓ = copyOf (hProp ℓ) ℓ + +-- We obtain a copy of the powerset using hPropResizing +module ResizedPowerset {ℓ} (resizing : hPropResizing ℓ) where + + smallHProp = resizing .fst + hProp≃smallHProp = resizing .snd + smallHProp≃hProp = invEquiv hProp≃smallHProp + + 𝓟 : Type ℓ → Type ℓ + 𝓟 X = X → smallHProp + + ℙ≃𝓟 : ∀ X → ℙ X ≃ 𝓟 X + ℙ≃𝓟 X = + ℙ X + ≃⟨ idEquiv (ℙ X) ⟩ + (X → hProp ℓ) + ≃⟨ equiv→ (idEquiv X) hProp≃smallHProp ⟩ + (X → smallHProp) + ≃⟨ idEquiv (𝓟 X) ⟩ + 𝓟 X + ■ + + 𝓟≃ℙ : ∀ X → 𝓟 X ≃ ℙ X + 𝓟≃ℙ X = invEquiv (ℙ≃𝓟 X) + + ℙ→𝓟 : ∀ X → ℙ X → 𝓟 X + ℙ→𝓟 X = equivFun (ℙ≃𝓟 X) + + 𝓟→ℙ : ∀ X → 𝓟 X → ℙ X + 𝓟→ℙ X = equivFun (invEquiv (ℙ≃𝓟 X)) + + compIsIdEquiv : ∀ X → compEquiv (ℙ≃𝓟 X) (invEquiv (ℙ≃𝓟 X)) ≡ idEquiv (ℙ X) + compIsIdEquiv X = invEquiv-is-rinv (ℙ≃𝓟 X) + + compIsIdFunc : ∀ {X} (p : ℙ X) → 𝓟→ℙ X (ℙ→𝓟 X p) ≡ p + compIsIdFunc {X} p i = equivFun (compIsIdEquiv X i) p + + _ϵ_ : ∀ {X} → X → 𝓟 X → Type ℓ + _ϵ_ {X} x P = x ∈ 𝓟→ℙ X P + + isPropϵ : ∀ {X} (x : X) P → isProp (x ϵ P) + isPropϵ {X} x P = ∈-isProp (𝓟→ℙ X P) x + + isSet𝓟 : ∀ X → isSet (𝓟 X) + isSet𝓟 X = isOfHLevelRespectEquiv 2 (ℙ≃𝓟 X) isSetℙ + + 𝓟≡Equiv : ∀ {X} (P Q : 𝓟 X) → (P ≡ Q) ≃ (𝓟→ℙ X P ≡ 𝓟→ℙ X Q) + 𝓟≡Equiv {X} P Q = congEquiv {x = P} {y = Q} (𝓟≃ℙ X) + + 𝓟≡ : ∀ {X} (P Q : 𝓟 X) → 𝓟→ℙ X P ≡ 𝓟→ℙ X Q → P ≡ Q + 𝓟≡ {X} P Q equ = equivFun (invEquiv (𝓟≡Equiv P Q)) equ + From 56dbf845817c995fc100463f39f00641e3599ad3 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 15 May 2024 17:23:19 +0530 Subject: [PATCH 45/61] Remove files not caught in gitignore --- src/Realizability/PERs/#TerminalObject.agda# | 43 -------------------- src/Realizability/PERs/.#TerminalObject.agda | 1 - 2 files changed, 44 deletions(-) delete mode 100644 src/Realizability/PERs/#TerminalObject.agda# delete mode 120000 src/Realizability/PERs/.#TerminalObject.agda diff --git a/src/Realizability/PERs/#TerminalObject.agda# b/src/Realizability/PERs/#TerminalObject.agda# deleted file mode 100644 index b573343..0000000 --- a/src/Realizability/PERs/#TerminalObject.agda# +++ /dev/null @@ -1,43 +0,0 @@ -open import Realizability.ApplicativeStructure -open import Realizability.CombinatoryAlgebra -open import Realizability.PropResizing -open import Cubical.Foundations.Prelude -open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Structure using (⟨_⟩; str) -open import Cubical.Foundations.Isomorphism -open import Cubical.Foundations.Equiv -open import Cubical.Foundations.Univalence -open import Cubical.Foundations.Powerset -open import Cubical.Functions.FunExtEquiv -open import Cubical.Relation.Binary -open import Cubical.Data.Sigma -open import Cubical.Data.FinData -open import Cubical.Data.Unit -open import Cubical.Reflection.RecordEquiv -open import Cubical.HITs.PropositionalTruncation as PT hiding (map) -open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.HITs.SetQuotients as SQ -open import Cubical.Categories.Category -open import Cubical.Categories.Limits.Terminal - -module Realizability.PERs.TerminalObject - {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where - -open import Realizability.PERs.PER ca -open CombinatoryAlgebra ca -open Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open PERMorphism - -terminalPER : PER -PER.relation terminalPER = (λ x y → Unit*) , λ x y → isPropUnit* -fst (PER.isPER terminalPER) = λ a b x → tt* -snd (PER.isPER terminalPER) = λ a b c x x₁ → tt* - -isTerminalTerminalPER : isTerminal PERCat terminalPER -isTerminalTerminalPER X = - inhProp→isContr - (makePERMorphism (λ x → [ k ]) (return (Id , (λ a aXa → (eq/ k (Id ⨾ a) tt*) , tt*)))) - λ x y → PERMorphism≡ x y (funExt λ q → {!effectiveOnDomain!}) - -terminal : Terminal PERCat -terminal = terminalPER , {!!} diff --git a/src/Realizability/PERs/.#TerminalObject.agda b/src/Realizability/PERs/.#TerminalObject.agda deleted file mode 120000 index 53126fa..0000000 --- a/src/Realizability/PERs/.#TerminalObject.agda +++ /dev/null @@ -1 +0,0 @@ -rahul@Rahuls-MacBook-Air.local.1114:1710392191 \ No newline at end of file From a5d81497ee84ce309e36b8f0f394f8d7e20f87cc Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 20 May 2024 19:01:32 +0530 Subject: [PATCH 46/61] Partial Surjections stuff --- src/Realizability/Assembly/SIP.agda | 26 ++-- src/Realizability/Modest/Base.agda | 4 +- .../Modest/PartialSurjection.agda | 116 ++++++++++++++++-- 3 files changed, 124 insertions(+), 22 deletions(-) diff --git a/src/Realizability/Assembly/SIP.agda b/src/Realizability/Assembly/SIP.agda index c9782fa..72f37b6 100644 --- a/src/Realizability/Assembly/SIP.agda +++ b/src/Realizability/Assembly/SIP.agda @@ -1,10 +1,9 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Structure -open import Cubical.Foundations.SIP open import Cubical.Foundations.Equiv open import Cubical.Foundations.HLevels open import Cubical.Foundations.Univalence -open import Cubical.Foundations.CartesianKanOps +open import Cubical.Foundations.Isomorphism open import Cubical.Functions.FunExtEquiv open import Cubical.Data.Sigma open import Cubical.HITs.PropositionalTruncation @@ -16,10 +15,19 @@ open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a t open import Realizability.Assembly.Base ca open import Realizability.Assembly.Morphism ca -Assembly≡ : ∀ {X Y : Type ℓ} - → (asmX : Assembly X) - → (asmY : Assembly Y) - → (P : X ≡ Y) - → (⊩overP : PathP (λ i → ∀ (a : A) (p : P i) → Type ℓ) (asmX ._⊩_) (asmY ._⊩_)) - → PathP (λ i → Assembly (P i)) asmX asmY -Assembly≡ {X} {Y} asmX asmY P ⊩overP = {!AssemblyIsoΣ!} +module _ {X : Type ℓ} where + + Assembly≡Iso : ∀ (asmA asmB : Assembly X) → Iso (asmA ≡ asmB) (∀ r x → r ⊩[ asmA ] x ≡ r ⊩[ asmB ] x) + Iso.fun (Assembly≡Iso asmA asmB) path r x i = r ⊩[ path i ] x + Assembly._⊩_ (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) r x = pointwisePath r x i + Assembly.isSetX (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) = isPropIsSet {A = X} (asmA .isSetX) (asmB .isSetX) i + Assembly.⊩isPropValued (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) r x = isProp→PathP {B = λ j → isProp (pointwisePath r x j)} (λ j → isPropIsProp) (asmA .⊩isPropValued r x) (asmB .⊩isPropValued r x) i + Assembly.⊩surjective (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) x = isProp→PathP {B = λ j → ∃[ a ∈ A ] (pointwisePath a x j)} (λ j → isPropPropTrunc) (asmA .⊩surjective x) (asmB .⊩surjective x) i + Iso.rightInv (Assembly≡Iso asmA asmB) pointwise = funExt₂ λ r x → refl + Iso.leftInv (Assembly≡Iso asmA asmB) path = isSetAssembly X asmA asmB _ _ + + Assembly≡Equiv : ∀ (asmA asmB : Assembly X) → (asmA ≡ asmB) ≃ (∀ r x → r ⊩[ asmA ] x ≡ r ⊩[ asmB ] x) + Assembly≡Equiv asmA asmB = isoToEquiv (Assembly≡Iso asmA asmB) + + Assembly≡ : ∀ (asmA asmB : Assembly X) → (∀ r x → r ⊩[ asmA ] x ≡ r ⊩[ asmB ] x) → (asmA ≡ asmB) + Assembly≡ asmA asmB pointwise = Iso.inv (Assembly≡Iso asmA asmB) pointwise diff --git a/src/Realizability/Modest/Base.agda b/src/Realizability/Modest/Base.agda index e2e9b1c..8295067 100644 --- a/src/Realizability/Modest/Base.agda +++ b/src/Realizability/Modest/Base.agda @@ -32,8 +32,8 @@ module _ {X : Type ℓ} (asmX : Assembly X) where isPropIsModest : isProp isModest isPropIsModest = isPropΠ3 λ x y a → isProp→ (isProp→ (asmX .isSetX x y)) - isUniqueRealised : isModest → ∀ (a : A) → isProp (Σ[ x ∈ X ] (a ⊩[ asmX ] x)) - isUniqueRealised isMod a (x , a⊩x) (y , a⊩y) = Σ≡Prop (λ x' → asmX .⊩isPropValued a x') (isMod x y a a⊩x a⊩y) + isUniqueRealized : isModest → ∀ (a : A) → isProp (Σ[ x ∈ X ] (a ⊩[ asmX ] x)) + isUniqueRealized isMod a (x , a⊩x) (y , a⊩y) = Σ≡Prop (λ x' → asmX .⊩isPropValued a x') (isMod x y a a⊩x a⊩y) ModestSet : Type ℓ → Type (ℓ-suc ℓ) ModestSet X = Σ[ xs ∈ Assembly X ] isModest xs diff --git a/src/Realizability/Modest/PartialSurjection.agda b/src/Realizability/Modest/PartialSurjection.agda index 7fb7aac..8ca4c13 100644 --- a/src/Realizability/Modest/PartialSurjection.agda +++ b/src/Realizability/Modest/PartialSurjection.agda @@ -22,7 +22,8 @@ module Realizability.Modest.PartialSurjection {ℓ} {A : Type ℓ} (ca : Combina open import Realizability.Assembly.Base ca open import Realizability.Assembly.Morphism ca -open import Realizability.Modest.Base ca resizing +open import Realizability.Assembly.SIP ca +open import Realizability.Modest.Base ca open Assembly open CombinatoryAlgebra ca @@ -40,8 +41,36 @@ record PartialSurjection (X : Type ℓ) : Type (ℓ-suc ℓ) where isSetX : isSet X -- potentially redundant? open PartialSurjection +module _ (X : Type ℓ) (isCorrectHLevel : isSet X) where + -- first we need a Σ type equivalent to partial surjections + -- we could use RecordEquiv but this does not give hProps and hSets and + -- that causes problems when trying to compute the hlevel + + PartialSurjectionΣ : Type (ℓ-suc ℓ) + PartialSurjectionΣ = Σ[ support ∈ (A → hProp ℓ) ] Σ[ enumeration ∈ ((Σ[ a ∈ A ] ⟨ support a ⟩) → X) ] isSurjection enumeration × isSet X + + isSetPartialSurjectionΣ : isSet PartialSurjectionΣ + isSetPartialSurjectionΣ = isSetΣ (isSet→ isSetHProp) (λ support → isSetΣ (isSet→ isCorrectHLevel) (λ enum → isSet× (isProp→isSet isPropIsSurjection) (isProp→isSet isPropIsSet))) + + PartialSurjectionIsoΣ : Iso (PartialSurjection X) PartialSurjectionΣ + Iso.fun PartialSurjectionIsoΣ surj = + (λ a → (surj .support a) , (surj .isPropSupport a)) , (λ { (a , suppA) → surj .enumeration (a , suppA) }) , surj .isSurjectionEnumeration , PartialSurjection.isSetX surj + Iso.inv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = makePartialSurjection (λ a → ⟨ support a ⟩) enumeration (λ a → str (support a)) isSurjectionEnumeration isSetX + Iso.rightInv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = refl + support (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .support a + enumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) (a , suppA) = surj .enumeration (a , suppA) + isPropSupport (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .isPropSupport a + isSurjectionEnumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSurjectionEnumeration + isSetX (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSetX + + PartialSurjection≡Σ : PartialSurjection X ≡ PartialSurjectionΣ + PartialSurjection≡Σ = isoToPath PartialSurjectionIsoΣ + + isSetPartialSurjection : isSet (PartialSurjection X) + isSetPartialSurjection = subst isSet (sym PartialSurjection≡Σ) isSetPartialSurjectionΣ + -- let us derive a structure of identity principle for partial surjections -module _ (X : Type ℓ) where +module SIP (X : Type ℓ) (isCorrectHLevel : isSet X) where PartialSurjection≡Iso : ∀ (p q : PartialSurjection X) @@ -60,14 +89,16 @@ module _ (X : Type ℓ) where (p .isSurjectionEnumeration b) (q .isSurjectionEnumeration b) i isSetX (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) = isPropIsSet (p .isSetX) (q .isSetX) i Iso.inv (PartialSurjection≡Iso p q) p≡q = (λ i → p≡q i .support) , (λ i → p≡q i .enumeration) - Iso.rightInv (PartialSurjection≡Iso p q) p≡q = {!!} + Iso.rightInv (PartialSurjection≡Iso p q) p≡q = isSetPartialSurjection X isCorrectHLevel _ _ _ _ Iso.leftInv (PartialSurjection≡Iso p q) (suppPath , enumPath) = ΣPathP (refl , refl) PartialSurjection≡ : ∀ (p q : PartialSurjection X) → Σ[ suppPath ∈ p .support ≡ q .support ] PathP (λ i → Σ[ a ∈ A ] (suppPath i a) → X) (p .enumeration) (q .enumeration) → p ≡ q PartialSurjection≡ p q (suppPath , enumPath) = Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) -- the type of partial surjections is equivalent to the type of modest assemblies on X -module _ (X : Type ℓ) where +module ModestSetIso (X : Type ℓ) (isCorrectHLevel : isSet X) where + + open SIP X isCorrectHLevel {-# TERMINATING #-} ModestSet→PartialSurjection : ModestSet X → PartialSurjection X @@ -75,7 +106,7 @@ module _ (X : Type ℓ) where enumeration (ModestSet→PartialSurjection (xs , isModestXs)) (r , ∃x) = let answer : Σ[ x ∈ X ] (r ⊩[ xs ] x) - answer = PT.rec (isUniqueRealised xs isModestXs r) (λ t → t) ∃x + answer = PT.rec (isUniqueRealized xs isModestXs r) (λ t → t) ∃x in fst answer isPropSupport (ModestSet→PartialSurjection (xs , isModestXs)) r = isPropPropTrunc isSurjectionEnumeration (ModestSet→PartialSurjection (xs , isModestXs)) x = @@ -108,7 +139,7 @@ module _ (X : Type ℓ) where rightInv : ∀ surj → ModestSet→PartialSurjection (PartialSurjection→ModestSet surj) ≡ surj rightInv surj = PartialSurjection≡ - X (ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)) surj + (ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)) surj (funExt supportEq , funExtDep {A = λ i → Σ-syntax A (funExt supportEq i)} @@ -119,20 +150,28 @@ module _ (X : Type ℓ) where PT.elim {P = λ ∃x → fst (PT.rec - (isUniqueRealised (fst (PartialSurjection→ModestSet surj)) + (isUniqueRealized (fst (PartialSurjection→ModestSet surj)) (snd (PartialSurjection→ModestSet surj)) r) (λ t → t) ∃x) ≡ surj .enumeration (s , supp)} (λ ∃x → surj .isSetX _ _) - (λ { (x , r⊩x) → + (λ { (x , suppR , ≡x) → let ∃x' = transport (sym (supportEq s)) supp + r≡s : r ≡ s + r≡s = PathPΣ p .fst in equivFun - (propTruncIdempotent≃ {!!}) + (propTruncIdempotent≃ (surj .isSetX x (surj .enumeration (s , supp)))) (do (x' , suppS , ≡x') ← ∃x' - return {!!}) }) + return + (x + ≡⟨ sym ≡x ⟩ + surj .enumeration (r , suppR) + ≡⟨ cong (surj .enumeration) (ΣPathP (r≡s , (isProp→PathP (λ i → surj .isPropSupport (PathPΣ p .fst i)) suppR supp))) ⟩ + surj .enumeration (s , supp) + ∎)) }) ∃x }) where supportEq : ∀ r → (∃[ x ∈ X ] (Σ[ supp ∈ surj .support r ] (surj .enumeration (r , supp) ≡ x))) ≡ support surj r supportEq = @@ -143,8 +182,63 @@ module _ (X : Type ℓ) where (λ ∃x → PT.rec (surj .isPropSupport r) (λ { (x , supp , ≡x) → supp }) ∃x) (λ supp → return (surj .enumeration (r , supp) , supp , refl))) + leftInv : ∀ mod → PartialSurjection→ModestSet (ModestSet→PartialSurjection mod) ≡ mod + leftInv (asmX , isModestAsmX) = + Σ≡Prop + isPropIsModest + (Assembly≡ _ _ + λ r x → + hPropExt + (isPropΣ isPropPropTrunc (λ ∃x → asmX .isSetX _ _)) + (asmX .⊩isPropValued r x) + (λ { (∃x , ≡x) → + let + (x' , r⊩x') = PT.rec (isUniqueRealized asmX isModestAsmX r) (λ t → t) ∃x + in subst (λ x' → r ⊩[ asmX ] x') ≡x r⊩x'}) + λ r⊩x → ∣ x , r⊩x ∣₁ , refl) + IsoModestSetPartialSurjection : Iso (ModestSet X) (PartialSurjection X) Iso.fun IsoModestSetPartialSurjection = ModestSet→PartialSurjection Iso.inv IsoModestSetPartialSurjection = PartialSurjection→ModestSet Iso.rightInv IsoModestSetPartialSurjection = rightInv - Iso.leftInv IsoModestSetPartialSurjection mod = {!!} + Iso.leftInv IsoModestSetPartialSurjection = leftInv + +record PartialSurjectionMorphism {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) : Type ℓ where + no-eta-equality + constructor makePartialSurjectionMorphism + field + map : X → Y + {- + The following "diagram" commutes + + Xˢ -----------> X + | | + | | + | | + | | + | | + ↓ ↓ + Yˢ -----------> Y + -} + isTracked : ∃[ t ∈ A ] (∀ (a : A) (sᵃ : psX .support a) → Σ[ sᵇ ∈ (psY .support (t ⨾ a)) ] map (psX .enumeration (a , sᵃ)) ≡ psY .enumeration ((t ⨾ a) , sᵇ)) + +-- SIP +module MorphismSIP {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) where + open PartialSurjectionMorphism + PartialSurjectionMorphism≡Iso : ∀ (f g : PartialSurjectionMorphism psX psY) → Iso (f ≡ g) (f .map ≡ g .map) + Iso.fun (PartialSurjectionMorphism≡Iso f g) f≡g i = f≡g i .map + map (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) = fMap≡gMap i + isTracked (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) = + isProp→PathP + -- Agda can't infer the type B + {B = λ j → ∃-syntax A + (λ t → + (a : A) (sᵃ : psX .support a) → + Σ-syntax (psY .support (t ⨾ a)) + (λ sᵇ → + fMap≡gMap j (psX .enumeration (a , sᵃ)) ≡ + psY .enumeration (t ⨾ a , sᵇ)))} + (λ j → isPropPropTrunc) + (f .isTracked) (g .isTracked) i + Iso.rightInv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap = refl + Iso.leftInv (PartialSurjectionMorphism≡Iso f g) f≡g = {!!} From b40f374708f2d5928e7f47a6c0aa0fde7ac5ca24 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Mon, 20 May 2024 23:13:35 +0530 Subject: [PATCH 47/61] Show modest set morphisms are equivalent to partial surjection morphisms --- .../Modest/PartialSurjection.agda | 70 ++++++++++++++++++- 1 file changed, 67 insertions(+), 3 deletions(-) diff --git a/src/Realizability/Modest/PartialSurjection.agda b/src/Realizability/Modest/PartialSurjection.agda index 8ca4c13..87bf0e6 100644 --- a/src/Realizability/Modest/PartialSurjection.agda +++ b/src/Realizability/Modest/PartialSurjection.agda @@ -54,8 +54,12 @@ module _ (X : Type ℓ) (isCorrectHLevel : isSet X) where PartialSurjectionIsoΣ : Iso (PartialSurjection X) PartialSurjectionΣ Iso.fun PartialSurjectionIsoΣ surj = - (λ a → (surj .support a) , (surj .isPropSupport a)) , (λ { (a , suppA) → surj .enumeration (a , suppA) }) , surj .isSurjectionEnumeration , PartialSurjection.isSetX surj - Iso.inv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = makePartialSurjection (λ a → ⟨ support a ⟩) enumeration (λ a → str (support a)) isSurjectionEnumeration isSetX + (λ a → (surj .support a) , (surj .isPropSupport a)) , + (λ { (a , suppA) → surj .enumeration (a , suppA) }) , + surj .isSurjectionEnumeration , + PartialSurjection.isSetX surj + Iso.inv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = + makePartialSurjection (λ a → ⟨ support a ⟩) enumeration (λ a → str (support a)) isSurjectionEnumeration isSetX Iso.rightInv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = refl support (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .support a enumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) (a , suppA) = surj .enumeration (a , suppA) @@ -221,6 +225,22 @@ record PartialSurjectionMorphism {X Y : Type ℓ} (psX : PartialSurjection X) (p Yˢ -----------> Y -} isTracked : ∃[ t ∈ A ] (∀ (a : A) (sᵃ : psX .support a) → Σ[ sᵇ ∈ (psY .support (t ⨾ a)) ] map (psX .enumeration (a , sᵃ)) ≡ psY .enumeration ((t ⨾ a) , sᵇ)) +open PartialSurjectionMorphism + +unquoteDecl PartialSurjectionMorphismIsoΣ = declareRecordIsoΣ PartialSurjectionMorphismIsoΣ (quote PartialSurjectionMorphism) + +PartialSurjectionMorphismΣ : {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) → Type ℓ +PartialSurjectionMorphismΣ {X} {Y} psX psY = + Σ[ f ∈ (X → Y) ] ∃[ t ∈ A ] ((∀ (a : A) (sᵃ : psX .support a) → Σ[ sᵇ ∈ (psY .support (t ⨾ a)) ] f (psX .enumeration (a , sᵃ)) ≡ psY .enumeration ((t ⨾ a) , sᵇ))) + +isSetPartialSurjectionMorphismΣ : {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) → isSet (PartialSurjectionMorphismΣ psX psY) +isSetPartialSurjectionMorphismΣ {X} {Y} psX psY = isSetΣ (isSet→ (psY .isSetX)) (λ f → isProp→isSet isPropPropTrunc) + +PartialSurjectionMorphismΣ≡ : {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) → PartialSurjectionMorphism psX psY ≡ PartialSurjectionMorphismΣ psX psY +PartialSurjectionMorphismΣ≡ {X} {Y} psX psY = isoToPath PartialSurjectionMorphismIsoΣ + +isSetPartialSurjectionMorphism : {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) → isSet (PartialSurjectionMorphism psX psY) +isSetPartialSurjectionMorphism {X} {Y} psX psY = subst isSet (sym (PartialSurjectionMorphismΣ≡ psX psY)) (isSetPartialSurjectionMorphismΣ psX psY) -- SIP module MorphismSIP {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) where @@ -241,4 +261,48 @@ module MorphismSIP {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSu (λ j → isPropPropTrunc) (f .isTracked) (g .isTracked) i Iso.rightInv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap = refl - Iso.leftInv (PartialSurjectionMorphism≡Iso f g) f≡g = {!!} + Iso.leftInv (PartialSurjectionMorphism≡Iso f g) f≡g = isSetPartialSurjectionMorphism psX psY f g _ _ + + PartialSurjectionMorphism≡ : ∀ {f g : PartialSurjectionMorphism psX psY} → (f .map ≡ g .map) → f ≡ g + PartialSurjectionMorphism≡ {f} {g} fMap≡gMap = Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap + +-- morphisms between partial surjections are equivalent to assembly morphisms between corresponding modest assemblies +module + _ + {X Y : Type ℓ} + (psX : PartialSurjection X) + (psY : PartialSurjection Y) where + open ModestSetIso + open MorphismSIP psX psY + + asmX = PartialSurjection→ModestSet X (psX .isSetX) psX .fst + isModestAsmX = PartialSurjection→ModestSet X (psX .isSetX) psX .snd + + asmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .fst + isModestAsmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .snd + + partialSurjectionHomModestSetHomIso : Iso (AssemblyMorphism asmX asmY) (PartialSurjectionMorphism psX psY) + map (Iso.fun partialSurjectionHomModestSetHomIso asmHom) = asmHom .map + isTracked (Iso.fun partialSurjectionHomModestSetHomIso asmHom) = + do + (map~ , isTrackedMap) ← asmHom .tracker + return + (map~ , + λ a aSuppX → + let + worker : (map~ ⨾ a) ⊩[ asmY ] (asmHom .map (psX .enumeration (a , aSuppX))) + worker = isTrackedMap (psX .enumeration (a , aSuppX)) a (aSuppX , refl) + in + (worker .fst) , + (sym (worker .snd))) + AssemblyMorphism.map (Iso.inv partialSurjectionHomModestSetHomIso surjHom) = surjHom .map + AssemblyMorphism.tracker (Iso.inv partialSurjectionHomModestSetHomIso surjHom) = + do + (t , isTrackedMap) ← surjHom .isTracked + return + (t , + (λ { x a (aSuppX , ≡x) → + (isTrackedMap a aSuppX .fst) , + (sym (cong (surjHom .map) (sym ≡x) ∙ isTrackedMap a aSuppX .snd)) })) + Iso.rightInv partialSurjectionHomModestSetHomIso surjHom = PartialSurjectionMorphism≡ refl + Iso.leftInv partialSurjectionHomModestSetHomIso asmHom = AssemblyMorphism≡ _ _ refl From f06f4a10c63599cc1211b04858e8711060d02341 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sat, 25 May 2024 14:06:04 +0530 Subject: [PATCH 48/61] Right adjoint to pullback functor ASM --- .../ApplicativeStructure.lagda.md | 31 +++--- src/Realizability/Assembly/Pullbacks.agda | 4 +- src/Realizability/Modest/Base.agda | 9 +- .../Modest/PartialSurjection.agda | 103 ++++++++++++++++-- 4 files changed, 121 insertions(+), 26 deletions(-) diff --git a/src/Realizability/ApplicativeStructure.lagda.md b/src/Realizability/ApplicativeStructure.lagda.md index 2287fef..3eeb7dd 100644 --- a/src/Realizability/ApplicativeStructure.lagda.md +++ b/src/Realizability/ApplicativeStructure.lagda.md @@ -60,7 +60,7 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where infix 23 `_ infixl 22 _̇_ data Term (n : ℕ) : Type ℓ where - # : Fin n → Term n + # : (Vec A n → A) → Term n `_ : A → Term n _̇_ : Term n → Term n → Term n ``` @@ -69,8 +69,8 @@ These terms can be evaluated into $A$ if we can give the values of the free vari ```agda ⟦_⟧ : ∀ {n} → Term n → Vec A n → A - ⟦_⟧ (` a) _ = a - ⟦_⟧ {n} (# k) subs = lookup k subs + ⟦_⟧ (` a) subs = a + ⟦_⟧ {n} (# k) subs = k subs ⟦_⟧ (a ̇ b) subs = (⟦ a ⟧ subs) ⨾ (⟦ b ⟧ subs) applicationChain : ∀ {n m} → Vec (Term m) (suc n) → Term m @@ -115,13 +115,12 @@ We will call this `λ*abst` and this will turn out to behave very similar to λ ```agda module ComputationRules (feferman : Feferman) where open Feferman feferman - + opaque λ*abst : ∀ {n} → (e : Term (suc n)) → Term n - λ*abst {n} (# zero) = ` s ̇ ` k ̇ ` k - λ*abst {n} (# (suc x)) = ` k ̇ # x - λ*abst {n} (` x) = ` k ̇ ` x - λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁ + λ*abst {n} (# x) = {!!} + λ*abst {n} (` x) = {!!} + λ*abst {n} (e ̇ e₁) = {!!} ``` **Remark** : It is important to note that in general, realizability is developed using **partial combinatory algebras** and **partial applicative structures**. In this case, `λ*abst` is not particularly well-behaved. The β reduction-esque rule we derive also does not behave *completely* like β reduction. See Jonh Longley's PhD thesis "Realizability Toposes and Language Semantics" Theorem 1.1.9. @@ -131,7 +130,7 @@ We will call this `λ*abst` and this will turn out to behave very similar to λ We define meta-syntactic sugar for some of the more common cases : ```agda - λ* : Term one → A + {-λ* : Term one → A λ* t = ⟦ λ*abst t ⟧ [] λ*2 : Term two → A @@ -141,7 +140,7 @@ We define meta-syntactic sugar for some of the more common cases : λ*3 t = ⟦ λ*abst (λ*abst (λ*abst t)) ⟧ [] λ*4 : Term four → A - λ*4 t = ⟦ λ*abst (λ*abst (λ*abst (λ*abst t))) ⟧ [] + λ*4 t = ⟦ λ*abst (λ*abst (λ*abst (λ*abst t))) ⟧ [] -} ``` We now show that we have a β-reduction-esque operation. We proceed by induction on the structure of the term and the number of free variables. @@ -150,7 +149,7 @@ For the particular combinatory algebra Λ/β (terms of the untyped λ calculus q TODO : Prove this. ```agda - opaque + {- opaque unfolding λ*abst βreduction : ∀ {n} → (body : Term (suc n)) → (prim : A) → (subs : Vec A n) → ⟦ λ*abst body ̇ ` prim ⟧ subs ≡ ⟦ body ⟧ (prim ∷ subs) βreduction {n} (# zero) prim subs = @@ -168,24 +167,24 @@ TODO : Prove this. ⟦ λ*abst rator ⟧ subs ⨾ prim ⨾ (⟦ λ*abst rand ⟧ subs ⨾ prim) ≡⟨ cong₂ (λ x y → x ⨾ y) (βreduction rator prim subs) (βreduction rand prim subs) ⟩ ⟦ rator ⟧ (prim ∷ subs) ⨾ ⟦ rand ⟧ (prim ∷ subs) - ∎ + ∎ -} ```
```agda - λ*chainTerm : ∀ n → Term n → Term zero + {-λ*chainTerm : ∀ n → Term n → Term zero λ*chainTerm zero t = t λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t) λ*chain : ∀ {n} → Term n → A - λ*chain {n} t = ⟦ λ*chainTerm n t ⟧ [] + λ*chain {n} t = ⟦ λ*chainTerm n t ⟧ [] -} ```
We provide useful reasoning combinators that are useful and frequent. ```agda - opaque + {- opaque unfolding reverse unfolding foldl unfolding chain @@ -232,5 +231,5 @@ We provide useful reasoning combinators that are useful and frequent. ⟦ λ*abst t ⟧ (c ∷ b ∷ a ∷ []) ⨾ ⟦ ` d ⟧ (c ∷ b ∷ a ∷ []) ≡⟨ βreduction t d (c ∷ b ∷ a ∷ []) ⟩ ⟦ t ⟧ (d ∷ c ∷ b ∷ a ∷ []) - ∎ + ∎ -} ``` diff --git a/src/Realizability/Assembly/Pullbacks.agda b/src/Realizability/Assembly/Pullbacks.agda index d2bb815..d138559 100644 --- a/src/Realizability/Assembly/Pullbacks.agda +++ b/src/Realizability/Assembly/Pullbacks.agda @@ -128,7 +128,7 @@ module _ {X Y : Type ℓ} (asmX : Assembly X) (asmY : Assembly Y) (f : AssemblyM pullbackFunctor : Functor (SliceCat ASM (Y , asmY)) (SliceCat ASM (X , asmX)) Functor.F-ob pullbackFunctor sliceY = sliceob (pullback (cospan (X , asmX) (Y , asmY) (sliceY .S-ob) f (sliceY .S-arr)) .pbPr₁) Functor.F-hom pullbackFunctor {ySliceA} {ySliceB} sliceHomAB = - let + {-let sliceACospan : Cospan ASM sliceACospan = cospan (X , asmX) (Y , asmY) (ySliceA .S-ob) f (ySliceA .S-arr) p = pullbackPr₂ sliceACospan @@ -153,7 +153,7 @@ module _ {X Y : Type ℓ} (asmX : Assembly X) (asmY : Assembly Y) (f : AssemblyM ∎ })) in slicehom - arrow + arrow-} {!!} Functor.F-id pullbackFunctor = {!!} Functor.F-seq pullbackFunctor = {!!} diff --git a/src/Realizability/Modest/Base.agda b/src/Realizability/Modest/Base.agda index 8295067..4dab0f7 100644 --- a/src/Realizability/Modest/Base.agda +++ b/src/Realizability/Modest/Base.agda @@ -39,7 +39,14 @@ ModestSet : Type ℓ → Type (ℓ-suc ℓ) ModestSet X = Σ[ xs ∈ Assembly X ] isModest xs MOD : Category (ℓ-suc ℓ) ℓ -MOD = ΣPropCat ASM λ { (X , asmX) → (Lift (isModest asmX)) , (isOfHLevelRespectEquiv 1 (LiftEquiv {A = isModest asmX}) (isPropIsModest asmX)) } +Category.ob MOD = Σ[ X ∈ Type ℓ ] Σ[ asmX ∈ Assembly X ] isModest asmX +Category.Hom[_,_] MOD (X , asmX , isModestAsmX) (Y , asmY , isModestAsmY) = AssemblyMorphism asmX asmY +Category.id MOD {X , asmX , isModestAsmX} = identityMorphism asmX +Category._⋆_ MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} {Z , asmZ , isModestAsmZ} f g = compositeMorphism f g +Category.⋆IdL MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = ⊚idL asmX asmY f +Category.⋆IdR MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = ⊚idR asmY asmX f +Category.⋆Assoc MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} {Z , asmZ , isModestAsmZ} {W , asmW , isModestAsmW} f g h = ⊚assoc asmX asmY asmZ asmW f g h +Category.isSetHom MOD {X , asmX , idModestAsmX} {Y , asmY , isModestAsmY} = isSetAssemblyMorphism asmX asmY -- Every modest set is isomorphic to a canonically modest set module Canonical (X : Type ℓ) (asmX : Assembly X) (isModestAsmX : isModest asmX) (resizing : hPropResizing ℓ) where diff --git a/src/Realizability/Modest/PartialSurjection.agda b/src/Realizability/Modest/PartialSurjection.agda index 87bf0e6..a475d78 100644 --- a/src/Realizability/Modest/PartialSurjection.agda +++ b/src/Realizability/Modest/PartialSurjection.agda @@ -14,6 +14,7 @@ open import Cubical.HITs.PropositionalTruncation as PT hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Reflection.RecordEquiv open import Cubical.Categories.Category +open import Cubical.Categories.Functor.Base hiding (Id) open import Realizability.CombinatoryAlgebra open import Realizability.ApplicativeStructure open import Realizability.PropResizing @@ -207,6 +208,9 @@ module ModestSetIso (X : Type ℓ) (isCorrectHLevel : isSet X) where Iso.rightInv IsoModestSetPartialSurjection = rightInv Iso.leftInv IsoModestSetPartialSurjection = leftInv + ModestSet≡PartialSurjection : ModestSet X ≡ PartialSurjection X + ModestSet≡PartialSurjection = isoToPath IsoModestSetPartialSurjection + record PartialSurjectionMorphism {X Y : Type ℓ} (psX : PartialSurjection X) (psY : PartialSurjection Y) : Type ℓ where no-eta-equality constructor makePartialSurjectionMorphism @@ -281,9 +285,9 @@ module asmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .fst isModestAsmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .snd - partialSurjectionHomModestSetHomIso : Iso (AssemblyMorphism asmX asmY) (PartialSurjectionMorphism psX psY) - map (Iso.fun partialSurjectionHomModestSetHomIso asmHom) = asmHom .map - isTracked (Iso.fun partialSurjectionHomModestSetHomIso asmHom) = + PartialSurjectionHomModestSetHomIso : Iso (AssemblyMorphism asmX asmY) (PartialSurjectionMorphism psX psY) + map (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) = asmHom .map + isTracked (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) = do (map~ , isTrackedMap) ← asmHom .tracker return @@ -295,8 +299,8 @@ module in (worker .fst) , (sym (worker .snd))) - AssemblyMorphism.map (Iso.inv partialSurjectionHomModestSetHomIso surjHom) = surjHom .map - AssemblyMorphism.tracker (Iso.inv partialSurjectionHomModestSetHomIso surjHom) = + AssemblyMorphism.map (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) = surjHom .map + AssemblyMorphism.tracker (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) = do (t , isTrackedMap) ← surjHom .isTracked return @@ -304,5 +308,90 @@ module (λ { x a (aSuppX , ≡x) → (isTrackedMap a aSuppX .fst) , (sym (cong (surjHom .map) (sym ≡x) ∙ isTrackedMap a aSuppX .snd)) })) - Iso.rightInv partialSurjectionHomModestSetHomIso surjHom = PartialSurjectionMorphism≡ refl - Iso.leftInv partialSurjectionHomModestSetHomIso asmHom = AssemblyMorphism≡ _ _ refl + Iso.rightInv PartialSurjectionHomModestSetHomIso surjHom = PartialSurjectionMorphism≡ refl + Iso.leftInv PartialSurjectionHomModestSetHomIso asmHom = AssemblyMorphism≡ _ _ refl + + PartialSurjectionHom≡ModestSetHom : AssemblyMorphism asmX asmY ≡ PartialSurjectionMorphism psX psY + PartialSurjectionHom≡ModestSetHom = isoToPath PartialSurjectionHomModestSetHomIso + +-- the category of partial surjections + +idPartSurjMorphism : ∀ {X : Type ℓ} → (psX : PartialSurjection X) → PartialSurjectionMorphism psX psX +map (idPartSurjMorphism {X} psX) x = x +isTracked (idPartSurjMorphism {X} psX) = + return (Id , (λ a aSuppX → (subst (λ r → psX .support r) (sym (Ida≡a a)) aSuppX) , (cong (psX .enumeration) (Σ≡Prop (λ b → psX .isPropSupport b) (sym (Ida≡a a)))))) + +composePartSurjMorphism : + ∀ {X Y Z : Type ℓ} {psX : PartialSurjection X} {psY : PartialSurjection Y} {psZ : PartialSurjection Z} + → (f : PartialSurjectionMorphism psX psY) + → (g : PartialSurjectionMorphism psY psZ) + → PartialSurjectionMorphism psX psZ +map (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) x = g .map (f .map x) +isTracked (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) = + do + (f~ , isTrackedF) ← f .isTracked + (g~ , isTrackedG) ← g .isTracked + let + realizer : Term as 1 + realizer = ` g~ ̇ (` f~ ̇ # zero) + return + (λ* realizer , + (λ a aSuppX → + subst (λ r' → psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~ ⨾ a) (isTrackedF a aSuppX .fst) .fst) , + (g .map (f .map (psX .enumeration (a , aSuppX))) + ≡⟨ cong (g .map) (isTrackedF a aSuppX .snd) ⟩ + g .map (psY .enumeration (f~ ⨾ a , fst (isTrackedF a aSuppX))) + ≡⟨ isTrackedG (f~ ⨾ a) (fst (isTrackedF a aSuppX)) .snd ⟩ + psZ .enumeration (g~ ⨾ (f~ ⨾ a) , fst (isTrackedG (f~ ⨾ a) (fst (isTrackedF a aSuppX)))) + ≡⟨ cong (psZ .enumeration) (Σ≡Prop (λ z → psZ .isPropSupport z) (sym (λ*ComputationRule realizer a))) ⟩ + psZ .enumeration + (λ* realizer ⨾ a , + subst (λ r' → psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~ ⨾ a) (isTrackedF a aSuppX .fst) .fst)) + ∎))) + +idLPartSurjMorphism : + ∀ {X Y : Type ℓ} + → {psX : PartialSurjection X} + → {psY : PartialSurjection Y} + → (f : PartialSurjectionMorphism psX psY) + → composePartSurjMorphism (idPartSurjMorphism psX) f ≡ f +idLPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl + +idRPartSurjMorphism : + ∀ {X Y : Type ℓ} + → {psX : PartialSurjection X} + → {psY : PartialSurjection Y} + → (f : PartialSurjectionMorphism psX psY) + → composePartSurjMorphism f (idPartSurjMorphism psY) ≡ f +idRPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl + +assocComposePartSurjMorphism : + ∀ {X Y Z W : Type ℓ} + → {psX : PartialSurjection X} + → {psY : PartialSurjection Y} + → {psZ : PartialSurjection Z} + → {psW : PartialSurjection W} + → (f : PartialSurjectionMorphism psX psY) + → (g : PartialSurjectionMorphism psY psZ) + → (h : PartialSurjectionMorphism psZ psW) + → composePartSurjMorphism (composePartSurjMorphism f g) h ≡ composePartSurjMorphism f (composePartSurjMorphism g h) +assocComposePartSurjMorphism {X} {Y} {Z} {W} {psX} {psY} {psZ} {psW} f g h = MorphismSIP.PartialSurjectionMorphism≡ psX psW refl + +PARTSURJ : Category (ℓ-suc ℓ) ℓ +Category.ob PARTSURJ = Σ[ X ∈ Type ℓ ] PartialSurjection X +Category.Hom[_,_] PARTSURJ (X , surjX) (Y , surjY) = PartialSurjectionMorphism surjX surjY +Category.id PARTSURJ {X , surjX} = idPartSurjMorphism surjX +Category._⋆_ PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} f g = composePartSurjMorphism f g +Category.⋆IdL PARTSURJ {X , surjX} {Y , surjY} f = idLPartSurjMorphism f +Category.⋆IdR PARTSURJ {X , surjX} {Y , surjY} f = idRPartSurjMorphism f +Category.⋆Assoc PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} {W , surjW} f g h = assocComposePartSurjMorphism f g h +Category.isSetHom PARTSURJ {X , surjX} {Y , surjY} = isSetPartialSurjectionMorphism surjX surjY + +open Category +open ModestSetIso + +L : Functor MOD PARTSURJ +Functor.F-ob L (X , modX) = X , (ModestSet→PartialSurjection X (modX .fst .isSetX) modX) +Functor.F-hom L {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = {!!} +Functor.F-id L = {!!} +Functor.F-seq L = {!!} From d644d36931f405a3abea72cda40c698ac410529d Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 30 May 2024 23:05:45 +0530 Subject: [PATCH 49/61] Archive --- src/Categories/PullbackFunctor.agda | 97 +++++++++++ .../ApplicativeStructure.lagda.md | 31 ++-- .../Assembly/LocallyCartesianClosed.agda | 155 ++++++++++++++++++ src/Realizability/Assembly/Pullbacks.agda | 40 ----- src/Realizability/Assembly/Reindexing.agda | 43 +++++ 5 files changed, 311 insertions(+), 55 deletions(-) create mode 100644 src/Categories/PullbackFunctor.agda create mode 100644 src/Realizability/Assembly/LocallyCartesianClosed.agda create mode 100644 src/Realizability/Assembly/Reindexing.agda diff --git a/src/Categories/PullbackFunctor.agda b/src/Categories/PullbackFunctor.agda new file mode 100644 index 0000000..caea860 --- /dev/null +++ b/src/Categories/PullbackFunctor.agda @@ -0,0 +1,97 @@ +{-# OPTIONS --allow-unsolved-metas #-} +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Constructions.Slice + +module Categories.PullbackFunctor where + +private + variable + ℓ ℓ' : Level +module _ (C : Category ℓ ℓ') (pullbacks : Pullbacks C) where + open Category C + open Pullback + open SliceOb + open SliceHom + module _ {X Y : ob} (f : Hom[ X , Y ]) where + module TransformMaps {A B : ob} (m : Hom[ A , Y ]) (n : Hom[ B , Y ]) (k : Hom[ A , B ]) (tri : k ⋆ n ≡ m) where + cospanA : Cospan C + cospanA = cospan X Y A f m + + cospanB : Cospan C + cospanB = cospan X Y B f n + + P : ob + P = pullbacks cospanA .pbOb + + Q : ob + Q = pullbacks cospanB .pbOb + + f*m : Hom[ P , X ] + f*m = pullbacks cospanA .pbPr₁ + + g : Hom[ P , A ] + g = pullbacks cospanA .pbPr₂ + + f*n : Hom[ Q , X ] + f*n = pullbacks cospanB .pbPr₁ + + h : Hom[ Q , B ] + h = pullbacks cospanB .pbPr₂ + + f*m⋆f≡g⋆m : f*m ⋆ f ≡ g ⋆ m + f*m⋆f≡g⋆m = pullbacks cospanA .pbCommutes + + g⋆k : Hom[ P , B ] + g⋆k = g ⋆ k + + g⋆k⋆n≡f*m⋆f : (g ⋆ k) ⋆ n ≡ f*m ⋆ f + g⋆k⋆n≡f*m⋆f = + (g ⋆ k) ⋆ n + ≡⟨ ⋆Assoc g k n ⟩ + g ⋆ (k ⋆ n) + ≡⟨ cong (λ x → g ⋆ x) tri ⟩ + g ⋆ m + ≡⟨ sym f*m⋆f≡g⋆m ⟩ + f*m ⋆ f + ∎ + + arrow : Hom[ P , Q ] + arrow = pullbackArrow C (pullbacks cospanB) f*m g⋆k (sym g⋆k⋆n≡f*m⋆f) + + arrow⋆f*n≡f*m : arrow ⋆ f*n ≡ f*m + arrow⋆f*n≡f*m = sym (pullbackArrowPr₁ C (pullbacks cospanB) f*m g⋆k (sym g⋆k⋆n≡f*m⋆f)) + + reindexFunctor : Functor (SliceCat C Y) (SliceCat C X) + Functor.F-ob reindexFunctor (sliceob {A} m) = sliceob (pullbacks (cospan X Y A f m) .pbPr₁) + Functor.F-hom reindexFunctor {sliceob {A} m} {sliceob {B} n} (slicehom k tri) = slicehom arrow arrow⋆f*n≡f*m where open TransformMaps m n k tri + Functor.F-id reindexFunctor {sliceob {A} m} = SliceHom-≡-intro' C X (pullbackArrowUnique C (pullbacks cospanB) f*m g⋆k (sym g⋆k⋆n≡f*m⋆f) id (sym (⋆IdL f*m)) (⋆IdR g ∙ sym (⋆IdL g))) where open TransformMaps m m id (⋆IdL m) + Functor.F-seq reindexFunctor {sliceob {A} m} {sliceob {B} n} {sliceob {C'} o} (slicehom j jComm) (slicehom k kComm) = SliceHom-≡-intro' C X {!!} where + module ABTransform = TransformMaps m n j jComm + module BCTransform = TransformMaps n o k kComm + module ACTransform = TransformMaps m o (j ⋆ k) (⋆Assoc j k o ∙ cong (λ x → j ⋆ x) kComm ∙ jComm) + + P : ob + P = ABTransform.P + + f*m : Hom[ P , X ] + f*m = ABTransform.f*m + + Q : ob + Q = ABTransform.Q + + R : ob + R = BCTransform.P + + f*n : Hom[ Q , X ] + f*n = ABTransform.f*n + + g : Hom[ P , A ] + g = ABTransform.g + + f*m⋆f≡g⋆m : f*m ⋆ f ≡ g ⋆ m + f*m⋆f≡g⋆m = ABTransform.f*m⋆f≡g⋆m + + diff --git a/src/Realizability/ApplicativeStructure.lagda.md b/src/Realizability/ApplicativeStructure.lagda.md index 3eeb7dd..2287fef 100644 --- a/src/Realizability/ApplicativeStructure.lagda.md +++ b/src/Realizability/ApplicativeStructure.lagda.md @@ -60,7 +60,7 @@ module _ {ℓ} {A : Type ℓ} (as : ApplicativeStructure A) where infix 23 `_ infixl 22 _̇_ data Term (n : ℕ) : Type ℓ where - # : (Vec A n → A) → Term n + # : Fin n → Term n `_ : A → Term n _̇_ : Term n → Term n → Term n ``` @@ -69,8 +69,8 @@ These terms can be evaluated into $A$ if we can give the values of the free vari ```agda ⟦_⟧ : ∀ {n} → Term n → Vec A n → A - ⟦_⟧ (` a) subs = a - ⟦_⟧ {n} (# k) subs = k subs + ⟦_⟧ (` a) _ = a + ⟦_⟧ {n} (# k) subs = lookup k subs ⟦_⟧ (a ̇ b) subs = (⟦ a ⟧ subs) ⨾ (⟦ b ⟧ subs) applicationChain : ∀ {n m} → Vec (Term m) (suc n) → Term m @@ -115,12 +115,13 @@ We will call this `λ*abst` and this will turn out to behave very similar to λ ```agda module ComputationRules (feferman : Feferman) where open Feferman feferman - + opaque λ*abst : ∀ {n} → (e : Term (suc n)) → Term n - λ*abst {n} (# x) = {!!} - λ*abst {n} (` x) = {!!} - λ*abst {n} (e ̇ e₁) = {!!} + λ*abst {n} (# zero) = ` s ̇ ` k ̇ ` k + λ*abst {n} (# (suc x)) = ` k ̇ # x + λ*abst {n} (` x) = ` k ̇ ` x + λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁ ``` **Remark** : It is important to note that in general, realizability is developed using **partial combinatory algebras** and **partial applicative structures**. In this case, `λ*abst` is not particularly well-behaved. The β reduction-esque rule we derive also does not behave *completely* like β reduction. See Jonh Longley's PhD thesis "Realizability Toposes and Language Semantics" Theorem 1.1.9. @@ -130,7 +131,7 @@ We will call this `λ*abst` and this will turn out to behave very similar to λ We define meta-syntactic sugar for some of the more common cases : ```agda - {-λ* : Term one → A + λ* : Term one → A λ* t = ⟦ λ*abst t ⟧ [] λ*2 : Term two → A @@ -140,7 +141,7 @@ We define meta-syntactic sugar for some of the more common cases : λ*3 t = ⟦ λ*abst (λ*abst (λ*abst t)) ⟧ [] λ*4 : Term four → A - λ*4 t = ⟦ λ*abst (λ*abst (λ*abst (λ*abst t))) ⟧ [] -} + λ*4 t = ⟦ λ*abst (λ*abst (λ*abst (λ*abst t))) ⟧ [] ``` We now show that we have a β-reduction-esque operation. We proceed by induction on the structure of the term and the number of free variables. @@ -149,7 +150,7 @@ For the particular combinatory algebra Λ/β (terms of the untyped λ calculus q TODO : Prove this. ```agda - {- opaque + opaque unfolding λ*abst βreduction : ∀ {n} → (body : Term (suc n)) → (prim : A) → (subs : Vec A n) → ⟦ λ*abst body ̇ ` prim ⟧ subs ≡ ⟦ body ⟧ (prim ∷ subs) βreduction {n} (# zero) prim subs = @@ -167,24 +168,24 @@ TODO : Prove this. ⟦ λ*abst rator ⟧ subs ⨾ prim ⨾ (⟦ λ*abst rand ⟧ subs ⨾ prim) ≡⟨ cong₂ (λ x y → x ⨾ y) (βreduction rator prim subs) (βreduction rand prim subs) ⟩ ⟦ rator ⟧ (prim ∷ subs) ⨾ ⟦ rand ⟧ (prim ∷ subs) - ∎ -} + ∎ ```
```agda - {-λ*chainTerm : ∀ n → Term n → Term zero + λ*chainTerm : ∀ n → Term n → Term zero λ*chainTerm zero t = t λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t) λ*chain : ∀ {n} → Term n → A - λ*chain {n} t = ⟦ λ*chainTerm n t ⟧ [] -} + λ*chain {n} t = ⟦ λ*chainTerm n t ⟧ [] ```
We provide useful reasoning combinators that are useful and frequent. ```agda - {- opaque + opaque unfolding reverse unfolding foldl unfolding chain @@ -231,5 +232,5 @@ We provide useful reasoning combinators that are useful and frequent. ⟦ λ*abst t ⟧ (c ∷ b ∷ a ∷ []) ⨾ ⟦ ` d ⟧ (c ∷ b ∷ a ∷ []) ≡⟨ βreduction t d (c ∷ b ∷ a ∷ []) ⟩ ⟦ t ⟧ (d ∷ c ∷ b ∷ a ∷ []) - ∎ -} + ∎ ``` diff --git a/src/Realizability/Assembly/LocallyCartesianClosed.agda b/src/Realizability/Assembly/LocallyCartesianClosed.agda new file mode 100644 index 0000000..8e94d88 --- /dev/null +++ b/src/Realizability/Assembly/LocallyCartesianClosed.agda @@ -0,0 +1,155 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Isomorphism +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor +open import Cubical.Categories.Constructions.Slice +open import Cubical.Categories.Adjoint +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Categories.PullbackFunctor + +module Realizability.Assembly.LocallyCartesianClosed {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) hiding (Z) +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Pullbacks ca +open import Realizability.Assembly.Reindexing ca +open NaturalBijection +open SliceOb + +module _ {X Y : Type ℓ} {asmX : Assembly X} {asmY : Assembly Y} (f : AssemblyMorphism asmX asmY) where + module SliceDomain {V : Type ℓ} {asmV : Assembly V} (h : AssemblyMorphism asmV asmX) where + + D : Type ℓ + D = Σ[ y ∈ Y ] Σ[ s ∈ (fiber (f .map) y → V)] (∀ (yFiberF : fiber (f .map) y) → h .map (s yFiberF) ≡ yFiberF .fst) + + isSetD : isSet D + isSetD = + isSetΣ + (asmY .isSetX) + (λ y → + isSetΣ + (isSet→ (asmV .isSetX)) + (λ s → + isSetΠ λ yFiberF → isProp→isSet (asmX .isSetX _ _))) + + _⊩D_ : A → D → Type ℓ + r ⊩D (y , s , coh) = ((pr₁ ⨾ r) ⊩[ asmY ] y) × (∀ (yFiberF : fiber (f .map) y) (a : A) → a ⊩[ asmX ] (yFiberF .fst) → (pr₂ ⨾ r ⨾ a) ⊩[ asmV ] (s yFiberF)) + + isProp⊩D : ∀ r d → isProp (r ⊩D d) + isProp⊩D r d = + isProp× + (asmY .⊩isPropValued _ _) + (isPropΠ + λ yFiberF → + isPropΠ + λ a → + isProp→ (asmV .⊩isPropValued _ _)) + + asmD : Assembly (Σ[ d ∈ D ] ∃[ r ∈ A ] (r ⊩D d)) + Assembly._⊩_ asmD r (d@(y , s , coh) , ∃r) = r ⊩D d + Assembly.isSetX asmD = isSetΣ isSetD (λ d → isProp→isSet isPropPropTrunc) + Assembly.⊩isPropValued asmD r (d@(y , s , coh) , ∃a) = isProp⊩D r d + Assembly.⊩surjective asmD (d , ∃a) = ∃a + + projection : AssemblyMorphism asmD asmY + AssemblyMorphism.map projection (d@(y , s , coh) , ∃r) = y + AssemblyMorphism.tracker projection = + return + (pr₁ , + (λ { (d@(y , s , coh) , ∃a) r r⊩d → r⊩d .fst })) + + private module SliceDomainHom {V W : Type ℓ} {asmV : Assembly V} {asmW : Assembly W} (g : AssemblyMorphism asmV asmX) (h : AssemblyMorphism asmW asmX) (j : AssemblyMorphism asmV asmW) (comm : j ⊚ h ≡ g) where + + rawMap : SliceDomain.D g → SliceDomain.D h + rawMap (y , s , coh) = + y , + (λ fib → j .map (s fib)) , + λ { fib@(x , fx≡y) → + h .map (j .map (s fib)) + ≡[ i ]⟨ comm i .map (s fib) ⟩ + g .map (s fib) + ≡⟨ coh fib ⟩ + x + ∎ } + + transformRealizers : ∀ (d : SliceDomain.D g) → Σ[ b ∈ A ] (SliceDomain._⊩D_ g b d) → Σ[ j~ ∈ A ] (tracks {xs = asmV} {ys = asmW} j~ (j .map)) → Σ[ r ∈ A ] (SliceDomain._⊩D_ h r (rawMap d)) + transformRealizers d@(y , s , coh) (e , pr₁e⊩y , pr₂e⊩) (j~ , isTrackedJ) = + let + realizer : A + realizer = pair ⨾ (pr₁ ⨾ e) ⨾ λ* (` j~ ̇ (` pr₂ ̇ ` e ̇ # zero)) + + pr₁realizer≡pr₁e : pr₁ ⨾ realizer ≡ pr₁ ⨾ e + pr₁realizer≡pr₁e = + pr₁ ⨾ realizer + ≡⟨ refl ⟩ -- unfolding + pr₁ ⨾ (pair ⨾ (pr₁ ⨾ e) ⨾ λ* (` j~ ̇ (` pr₂ ̇ ` e ̇ # zero))) + ≡⟨ pr₁pxy≡x _ _ ⟩ + pr₁ ⨾ e + ∎ + + pr₂realizerEq : (a : A) → pr₂ ⨾ realizer ⨾ a ≡ j~ ⨾ (pr₂ ⨾ e ⨾ a) + pr₂realizerEq a = + pr₂ ⨾ realizer ⨾ a + ≡⟨ refl ⟩ + pr₂ ⨾ (pair ⨾ (pr₁ ⨾ e) ⨾ λ* (` j~ ̇ (` pr₂ ̇ ` e ̇ # zero))) ⨾ a + ≡⟨ cong (λ x → x ⨾ a) (pr₂pxy≡y _ _) ⟩ + λ* (` j~ ̇ (` pr₂ ̇ ` e ̇ # zero)) ⨾ a + ≡⟨ λ*ComputationRule (` j~ ̇ (` pr₂ ̇ ` e ̇ # zero)) a ⟩ + j~ ⨾ (pr₂ ⨾ e ⨾ a) + ∎ + in + (realizer , + (subst (λ r' → r' ⊩[ asmY ] y) (sym pr₁realizer≡pr₁e) pr₁e⊩y , + (λ { fib@(x , fx≡y) a a⊩x → + subst (λ r' → r' ⊩[ asmW ] (j .map (s fib))) (sym (pr₂realizerEq a)) (isTrackedJ (s fib) (pr₂ ⨾ e ⨾ a) (pr₂e⊩ fib a a⊩x)) }))) + + morphism : AssemblyMorphism (SliceDomain.asmD g) (SliceDomain.asmD h) + AssemblyMorphism.map morphism (d@(y , s , coh) , ∃r) = rawMap d , PT.rec2 isPropPropTrunc (λ r⊩d j~ → ∣ transformRealizers d r⊩d j~ ∣₁) ∃r (j .tracker) + AssemblyMorphism.tracker morphism = + do + (j~ , isTrackedJ) ← j .tracker + let + closure : Term as 2 + closure = (` j~ ̇ (` pr₂ ̇ # one ̇ # zero)) + + realizer : Term as 1 + realizer = ` pair ̇ (` pr₁ ̇ # zero) ̇ (λ*abst closure) + return + (λ* realizer , + (λ { (d@(y , s , coh) , ∃r) a (pr₁a⊩y , pr₂a⊩) → + let + pr₁Eq : pr₁ ⨾ (λ* realizer ⨾ a) ≡ pr₁ ⨾ a + pr₁Eq = + pr₁ ⨾ (λ* realizer ⨾ a) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer a) ⟩ + pr₁ ⨾ (pair ⨾ (pr₁ ⨾ a) ⨾ _) + ≡⟨ pr₁pxy≡x _ _ ⟩ + pr₁ ⨾ a + ∎ + in + subst (λ r' → r' ⊩[ asmY ] y) (sym pr₁Eq) pr₁a⊩y , + (λ { fib@(x , fx≡y) b b⊩x → {!isTrackedJ !} }) })) + + Π[_] : Functor (SliceCat ASM (X , asmX)) (SliceCat ASM (Y , asmY)) + Functor.F-ob Π[_] (sliceob {V , asmV} h) = sliceob (SliceDomain.projection h) + Functor.F-hom Π[_] {sliceob {V , asmV} g} {sliceob {W , asmW} h} (slicehom j comm) = {!!} + Functor.F-id Π[_] = {!!} + Functor.F-seq Π[_] = {!!} + + reindex⊣Π[_] : reindexFunctor ASM PullbacksASM f ⊣ Π[_] + Iso.fun (_⊣_.adjIso reindex⊣Π[_]) = λ fo → slicehom {!!} {!!} + Iso.inv (_⊣_.adjIso reindex⊣Π[_]) = {!!} + Iso.rightInv (_⊣_.adjIso reindex⊣Π[_]) = {!!} + Iso.leftInv (_⊣_.adjIso reindex⊣Π[_]) = {!!} + _⊣_.adjNatInD reindex⊣Π[_] = {!!} + _⊣_.adjNatInC reindex⊣Π[_] = {!!} diff --git a/src/Realizability/Assembly/Pullbacks.agda b/src/Realizability/Assembly/Pullbacks.agda index d138559..9a5f39f 100644 --- a/src/Realizability/Assembly/Pullbacks.agda +++ b/src/Realizability/Assembly/Pullbacks.agda @@ -117,43 +117,3 @@ module _ (cspn : Cospan ASM) where PullbacksASM : Pullbacks ASM PullbacksASM cspn = pullback cspn - --- Using pullbacks we get a functor that sends any f : X → Y to f* : Asm / Y → Asm / X -module _ {X Y : Type ℓ} (asmX : Assembly X) (asmY : Assembly Y) (f : AssemblyMorphism asmX asmY) where - open Pullback - opaque - unfolding pullbackAsm - unfolding pullbackPr₁ - unfolding pullback - pullbackFunctor : Functor (SliceCat ASM (Y , asmY)) (SliceCat ASM (X , asmX)) - Functor.F-ob pullbackFunctor sliceY = sliceob (pullback (cospan (X , asmX) (Y , asmY) (sliceY .S-ob) f (sliceY .S-arr)) .pbPr₁) - Functor.F-hom pullbackFunctor {ySliceA} {ySliceB} sliceHomAB = - {-let - sliceACospan : Cospan ASM - sliceACospan = cospan (X , asmX) (Y , asmY) (ySliceA .S-ob) f (ySliceA .S-arr) - p = pullbackPr₂ sliceACospan - m = ySliceA .S-arr - n = ySliceB .S-arr - f*m = pullbackPr₁ sliceACospan - h = sliceHomAB .S-hom - m≡h⊚n = sliceHomAB .S-comm - f*m⊚f≡p⊚m = pbCommutes (pullback sliceACospan) - arrow = - pullbackArrow - ASM - (pullback (cospan (X , asmX) (Y , asmY) (ySliceB .S-ob) f (ySliceB .S-arr))) (pullbackPr₁ sliceACospan) (p ⊚ h) - (AssemblyMorphism≡ _ _ - (funExt - λ { (x , a , fx≡ma) → - f .map x - ≡⟨ fx≡ma ⟩ - m .map a - ≡[ i ]⟨ m≡h⊚n (~ i) .map a ⟩ - n .map (h .map a) - ∎ })) - in - slicehom - arrow-} - {!!} - Functor.F-id pullbackFunctor = {!!} - Functor.F-seq pullbackFunctor = {!!} diff --git a/src/Realizability/Assembly/Reindexing.agda b/src/Realizability/Assembly/Reindexing.agda new file mode 100644 index 0000000..13ddc62 --- /dev/null +++ b/src/Realizability/Assembly/Reindexing.agda @@ -0,0 +1,43 @@ +{-# OPTIONS --allow-unsolved-metas #-} +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Category.Base +open import Cubical.Categories.Functor +open import Cubical.Categories.Constructions.Slice +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Categories.PullbackFunctor + +module Realizability.Assembly.Reindexing {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) hiding (Z) +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Pullbacks ca +-- Using pullbacks we get a functor that sends any f : X → Y to f* : Asm / Y → Asm / X +module _ {X Y : Type ℓ} {asmX : Assembly X} {asmY : Assembly Y} (f : AssemblyMorphism asmX asmY) where + + open TransformMaps ASM PullbacksASM f + open SliceOb + open SliceHom + open Pullback + + opaque + reindexPb : {A : Type ℓ} (asmA : Assembly A) (m : AssemblyMorphism asmA asmY) → ASM .Category.ob + reindexPb {A} asmA m = pullback (cospan (X , asmX) (Y , asmY) (A , asmA) f m) .pbOb + + opaque + reindexOb : SliceOb ASM (Y , asmY) → SliceOb ASM (X , asmX) + reindexOb (sliceob {A , asmA} m) = sliceob (pullback (cospan (X , asmX) (Y , asmY) (A , asmA) f m) .pbPr₁) + + opaque + unfolding reindexOb + reindexHom : (m n : SliceOb ASM (Y , asmY)) → SliceHom ASM (Y , asmY) m n → SliceHom ASM (X , asmX) (reindexOb m) (reindexOb n) + reindexHom (sliceob {A , asmA} m) (sliceob {B , asmB} n) (slicehom k tri) = slicehom (arrow m n k tri) (arrow⋆f*n≡f*m m n k tri) + From 827bf93f016736ea941de06cf40f744135336732 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 5 Jun 2024 23:07:32 +0530 Subject: [PATCH 50/61] Archive --- .../Assembly/LocallyCartesianClosed.agda | 221 ++++++++++++++++-- src/Realizability/Modest/UniformFamily.agda | 69 ++++++ 2 files changed, 277 insertions(+), 13 deletions(-) create mode 100644 src/Realizability/Modest/UniformFamily.agda diff --git a/src/Realizability/Assembly/LocallyCartesianClosed.agda b/src/Realizability/Assembly/LocallyCartesianClosed.agda index 8e94d88..669ce5f 100644 --- a/src/Realizability/Assembly/LocallyCartesianClosed.agda +++ b/src/Realizability/Assembly/LocallyCartesianClosed.agda @@ -4,6 +4,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Isomorphism open import Cubical.Data.Sigma open import Cubical.Data.FinData +open import Cubical.Data.Vec hiding (map) open import Cubical.HITs.PropositionalTruncation as PT hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Categories.Limits.Pullback @@ -136,20 +137,214 @@ module _ {X Y : Type ℓ} {asmX : Assembly X} {asmY : Assembly Y} (f : AssemblyM ≡⟨ pr₁pxy≡x _ _ ⟩ pr₁ ⨾ a ∎ + + pr₂Eq : ∀ b → pr₂ ⨾ (λ* realizer ⨾ a) ⨾ b ≡ j~ ⨾ (pr₂ ⨾ a ⨾ b) + pr₂Eq b = + pr₂ ⨾ (λ* realizer ⨾ a) ⨾ b + ≡⟨ cong (λ x → pr₂ ⨾ x ⨾ b) (λ*ComputationRule realizer a) ⟩ + pr₂ ⨾ (pair ⨾ (pr₁ ⨾ a) ⨾ _) ⨾ b + ≡⟨ cong (λ x → x ⨾ b) (pr₂pxy≡y _ _) ⟩ + βreduction closure b (a ∷ []) in subst (λ r' → r' ⊩[ asmY ] y) (sym pr₁Eq) pr₁a⊩y , - (λ { fib@(x , fx≡y) b b⊩x → {!isTrackedJ !} }) })) - + (λ { fib@(x , fx≡y) b b⊩x → subst (λ r' → r' ⊩[ asmW ] j .map (s fib)) (sym (pr₂Eq b)) (isTrackedJ (s fib) (pr₂ ⨾ a ⨾ b) (pr₂a⊩ fib b b⊩x)) }) })) + -- + {-# TERMINATING #-} Π[_] : Functor (SliceCat ASM (X , asmX)) (SliceCat ASM (Y , asmY)) Functor.F-ob Π[_] (sliceob {V , asmV} h) = sliceob (SliceDomain.projection h) - Functor.F-hom Π[_] {sliceob {V , asmV} g} {sliceob {W , asmW} h} (slicehom j comm) = {!!} - Functor.F-id Π[_] = {!!} - Functor.F-seq Π[_] = {!!} - - reindex⊣Π[_] : reindexFunctor ASM PullbacksASM f ⊣ Π[_] - Iso.fun (_⊣_.adjIso reindex⊣Π[_]) = λ fo → slicehom {!!} {!!} - Iso.inv (_⊣_.adjIso reindex⊣Π[_]) = {!!} - Iso.rightInv (_⊣_.adjIso reindex⊣Π[_]) = {!!} - Iso.leftInv (_⊣_.adjIso reindex⊣Π[_]) = {!!} - _⊣_.adjNatInD reindex⊣Π[_] = {!!} - _⊣_.adjNatInC reindex⊣Π[_] = {!!} + Functor.F-hom Π[_] {sliceob {V , asmV} g} {sliceob {W , asmW} h} (slicehom j comm) = slicehom (SliceDomainHom.morphism g h j comm) (AssemblyMorphism≡ _ _ (funExt λ { (d@(y , s , coh) , ∃r) → refl } )) + Functor.F-id Π[_] {sliceob {V , asmV} h} = SliceHom-≡-intro' ASM (Y , asmY) (AssemblyMorphism≡ _ _ (funExt λ { (d@(y , s , coh) , ∃r) → Σ≡Prop (λ d → isPropPropTrunc) (ΣPathP (refl , (ΣPathP (refl , (funExt (λ { (x , fx≡y) → asmX .isSetX _ _ _ _ })))))) })) + -- this call is marked as problematic + -- it is not even recursive :( + Functor.F-seq Π[_] {sliceob {U , asmU} g} {sliceob {V , asmV} h} {sliceob {W , asmW} i} (slicehom j jComm) (slicehom k kComm) = SliceHom-≡-intro' ASM ((Y , asmY)) (AssemblyMorphism≡ _ _ (funExt (λ { (d@(y , s , coh) , ∃r) → Σ≡Prop (λ d → isPropPropTrunc) (ΣPathP (refl , (ΣPathP (refl , (funExt (λ { (x , fx≡y) → asmX .isSetX _ _ _ _ })))))) }))) + + module ForwardIso {V P : Type ℓ} {asmV : Assembly V} {asmP : Assembly P} (g : AssemblyMorphism asmV asmY) (m : AssemblyMorphism asmP asmX) (h : AssemblyMorphism ((reindexFunctor ASM PullbacksASM f ⟅ sliceob g ⟆) .S-ob .snd) asmP) (hComm : h ⊚ m ≡ (reindexFunctor ASM PullbacksASM f ⟅ sliceob g ⟆) .S-arr) where + opaque + unfolding pullback + answerMap : V → (SliceDomain.D m) + answerMap v = + g .map v , + (λ { fib@(x , fx≡gv) → h .map (x , v , fx≡gv) }) , + (λ { (x , fx≡gv) → + m .map (h .map (x , v , fx≡gv)) ≡[ i ]⟨ hComm i .map (x , v , fx≡gv) ⟩ x ∎ }) + + answerTracker : ∃[ t ∈ A ] (∀ (v : V) (b : A) → b ⊩[ asmV ] v → SliceDomain._⊩D_ m (t ⨾ b) (answerMap v)) + answerTracker = + do + (g~ , isTrackedG) ← g .tracker + (h~ , isTrackedH) ← h .tracker + let + closure : Term as 2 + closure = ` h~ ̇ (` pair ̇ # zero ̇ # one) + + realizer : Term as 1 + realizer = ` pair ̇ (` g~ ̇ # zero) ̇ λ*abst closure + return + (λ* realizer , + (λ v a a⊩v → + let + pr₁eq : pr₁ ⨾ (λ* realizer ⨾ a) ≡ g~ ⨾ a + pr₁eq = + pr₁ ⨾ (λ* realizer ⨾ a) + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule realizer a) ⟩ + pr₁ ⨾ (pair ⨾ (g~ ⨾ a) ⨾ _) + ≡⟨ pr₁pxy≡x _ _ ⟩ + g~ ⨾ a + ∎ + + pr₂eq : ∀ b → pr₂ ⨾ (λ* realizer ⨾ a) ⨾ b ≡ h~ ⨾ (pair ⨾ b ⨾ a) + pr₂eq b = + pr₂ ⨾ (λ* realizer ⨾ a) ⨾ b + ≡⟨ cong (λ x → pr₂ ⨾ x ⨾ b) (λ*ComputationRule realizer a) ⟩ + pr₂ ⨾ (pair ⨾ (g~ ⨾ a) ⨾ _) ⨾ b + ≡⟨ cong (λ x → x ⨾ b) (pr₂pxy≡y _ _) ⟩ + βreduction closure b (a ∷ []) + in + subst (λ r' → r' ⊩[ asmY ] (g .map v)) (sym pr₁eq) (isTrackedG v a a⊩v) , + λ { (x , fx≡gv) b b⊩x → + subst + (λ r' → r' ⊩[ asmP ] h .map (x , v , fx≡gv)) + (sym (pr₂eq b)) + (isTrackedH + (x , v , fx≡gv) + (pair ⨾ b ⨾ a) + (subst (λ r' → r' ⊩[ asmX ] x) (sym (pr₁pxy≡x _ _)) b⊩x , + subst (λ r' → r' ⊩[ asmV ] v) (sym (pr₂pxy≡y _ _)) a⊩v)) })) + + answer : AssemblyMorphism asmV (SliceDomain.asmD m) + AssemblyMorphism.map answer v = + answerMap v , + do + (tr , isTrackedAnswer) ← answerTracker + (b , b⊩v) ← asmV .⊩surjective v + return (tr ⨾ b , isTrackedAnswer v b b⊩v) + AssemblyMorphism.tracker answer = answerTracker + + answerEq : answer ⊚ ((Π[_] ⟅ sliceob m ⟆) .S-arr) ≡ g + answerEq = AssemblyMorphism≡ _ _ (funExt λ v → refl) + + + module BackwardIso {V P : Type ℓ} {asmV : Assembly V} {asmP : Assembly P} (g : AssemblyMorphism asmV asmY) (m : AssemblyMorphism asmP asmX) (h : AssemblyMorphism asmV ((Π[_] ⟅ sliceob m ⟆) .S-ob .snd)) (hComm : h ⊚ ((Π[_] ⟅ sliceob m ⟆) .S-arr) ≡ g) where + Q : ASM .Category.ob + Q = (reindexFunctor ASM PullbacksASM f ⟅ sliceob g ⟆) .S-ob + + typeQ : Type ℓ + typeQ = Q .fst + + asmQ : Assembly typeQ + asmQ = Q .snd + + isFiberOfY : (x : X) → (v : V) → f .map x ≡ g .map v → h .map v .fst .fst ≡ f .map x + isFiberOfY x v fx≡gv = + h .map v .fst .fst + ≡[ i ]⟨ hComm i .map v ⟩ + g .map v + ≡⟨ sym fx≡gv ⟩ + f .map x + ∎ + + opaque + unfolding pullback + answerMap : Q .fst → P + answerMap (x , v , fx≡gv) = + h .map v .fst .snd .fst + (x , + sym (isFiberOfY x v fx≡gv)) + + opaque + unfolding pullback + unfolding answerMap + answer : AssemblyMorphism asmQ asmP + AssemblyMorphism.map answer = answerMap + AssemblyMorphism.tracker answer = + do + (h~ , isTrackedH) ← h .tracker + let + realizer : Term as 1 + realizer = ` pr₂ ̇ (` h~ ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero) + return + (λ* realizer , + (λ { q@(x , v , fx≡gv) a (pr₁a⊩x , pr₂a⊩v) → + subst + (λ r' → r' ⊩[ asmP ] (answerMap (x , v , fx≡gv))) + (sym (λ*ComputationRule realizer a)) + (isTrackedH v (pr₂ ⨾ a) pr₂a⊩v .snd (x , sym (isFiberOfY x v fx≡gv)) (pr₁ ⨾ a) pr₁a⊩x) })) + + opaque + unfolding pullback + reindex⊣Π[_] : reindexFunctor ASM PullbacksASM f ⊣ Π[_] + Iso.fun (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = slicehom (ForwardIso.answer g m h hComm) (ForwardIso.answerEq g m h hComm) + Iso.inv (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = + slicehom answer (AssemblyMorphism≡ _ _ (funExt λ { (x , v , fx≡gv) → answerEq x v fx≡gv })) where + Π[f]m : AssemblyMorphism (S-ob (Π[_] ⟅ sliceob m ⟆) .snd) asmY + Π[f]m = (Π[_] ⟅ sliceob m ⟆) .S-arr + + Q : ASM .Category.ob + Q = (reindexFunctor ASM PullbacksASM f ⟅ sliceob g ⟆) .S-ob + + typeQ : Type ℓ + typeQ = Q .fst + + asmQ : Assembly typeQ + asmQ = Q .snd + + isFiberOfY : (x : X) → (v : V) → f .map x ≡ g .map v → h .map v .fst .fst ≡ f .map x + isFiberOfY x v fx≡gv = + h .map v .fst .fst + ≡[ i ]⟨ hComm i .map v ⟩ + g .map v + ≡⟨ sym fx≡gv ⟩ + f .map x + ∎ + + answerMap : typeQ → P + answerMap (x , v , fx≡gv) = + h .map v .fst .snd .fst + (x , + sym (isFiberOfY x v fx≡gv)) + + answerEq : (x : X) (v : V) (fx≡gv : f .map x ≡ g .map v) → m .map (answerMap (x , v , fx≡gv)) ≡ x + answerEq x v fx≡gv = + m .map (answerMap (x , v , fx≡gv)) + ≡⟨ refl ⟩ + m .map (h .map v .fst .snd .fst (x , sym (isFiberOfY x v fx≡gv))) + ≡⟨ h .map v .fst .snd .snd (x , sym (isFiberOfY x v fx≡gv)) ⟩ + x + ∎ + + answer : AssemblyMorphism asmQ asmP + AssemblyMorphism.map answer = answerMap + AssemblyMorphism.tracker answer = + do + (h~ , isTrackedH) ← h .tracker + let + realizer : Term as 1 + realizer = ` pr₂ ̇ (` h~ ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero) + return + (λ* realizer , + (λ { q@(x , v , fx≡gv) a (pr₁a⊩x , pr₂a⊩v) → + subst + (λ r' → r' ⊩[ asmP ] (answerMap (x , v , fx≡gv))) + (sym (λ*ComputationRule realizer a)) + (isTrackedH v (pr₂ ⨾ a) pr₂a⊩v .snd (x , sym (isFiberOfY x v fx≡gv)) (pr₁ ⨾ a) pr₁a⊩x) })) + Iso.rightInv (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = + SliceHom-≡-intro' + ASM + (Y , asmY) + (AssemblyMorphism≡ _ _ + (funExt + λ v → + Σ≡Prop + (λ d → isPropPropTrunc) + (ΣPathP + ({!h!} , {!!})))) + Iso.leftInv (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = + SliceHom-≡-intro' + ASM + (X , asmX) + (AssemblyMorphism≡ _ _ + (funExt + λ { (x , v , fx≡gv) → + {!!} })) + _⊣_.adjNatInD reindex⊣Π[_] = {!!} + _⊣_.adjNatInC reindex⊣Π[_] = {!!} diff --git a/src/Realizability/Modest/UniformFamily.agda b/src/Realizability/Modest/UniformFamily.agda new file mode 100644 index 0000000..8dfb38a --- /dev/null +++ b/src/Realizability/Modest/UniformFamily.agda @@ -0,0 +1,69 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor +open import Cubical.Categories.Constructions.Slice +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.UniformFamily {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.LocallyCartesianClosed ca +open import Realizability.Modest.Base ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) + +UNIMOD : Categoryᴰ ASM (ℓ-suc ℓ) ℓ +Categoryᴰ.ob[ UNIMOD ] (X , asmX) = Σ[ Y ∈ Type ℓ ] Σ[ asmY ∈ Assembly Y ] isModest asmY × AssemblyMorphism asmY asmX +Categoryᴰ.Hom[_][_,_] UNIMOD {X , asmX} {Y , asmY} reindex (V , asmV , isModestAsmV , m) (W , asmW , isModestAsmW , n) = Σ[ reindexᵈ ∈ (AssemblyMorphism asmV asmW) ] m ⊚ reindex ≡ reindexᵈ ⊚ n +Categoryᴰ.idᴰ UNIMOD {X , asmX} {V , asmV , isModestAsmV , m} = (identityMorphism asmV) , (Category.⋆IdR ASM m ∙ sym (Category.⋆IdL ASM m)) +Categoryᴰ._⋆ᴰ_ UNIMOD + {X , asmX} {Y , asmY} {Z , asmZ} {f} {g} + {U , asmU , isModestU , m} {V , asmV , isModestV , n} {W , asmW , isModestW , o} + (fᵈ , fᵈcomm) (gᵈ , gᵈcomm) = + (fᵈ ⊚ gᵈ) , + (m ⊚ (f ⊚ g) + ≡⟨ sym (Category.⋆Assoc ASM m f g) ⟩ + (m ⊚ f) ⊚ g + ≡⟨ cong (λ x → x ⊚ g) fᵈcomm ⟩ + fᵈ ⊚ n ⊚ g + ≡⟨ Category.⋆Assoc ASM fᵈ n g ⟩ + fᵈ ⊚ (n ⊚ g) + ≡⟨ cong (fᵈ ⊚_) gᵈcomm ⟩ + fᵈ ⊚ (gᵈ ⊚ o) + ≡⟨ sym (Category.⋆Assoc ASM fᵈ gᵈ o) ⟩ + fᵈ ⊚ gᵈ ⊚ o + ∎) +Categoryᴰ.⋆IdLᴰ UNIMOD {X , asmX} {Y , asmY} {f} {V , asmV , isModestAsmV , m} {W , asmW , isModestAsmW , n} (fᵈ , fᵈcomm) = + ΣPathPProp (λ fᵈ → isSetAssemblyMorphism asmV asmY _ _) (Category.⋆IdL ASM fᵈ) +Categoryᴰ.⋆IdRᴰ UNIMOD {X , asmX} {Y , asmY} {f} {V , asmV , isModestAsmV , m} {W , asmW , isModestAsmW , n} (fᵈ , fᵈcomm) = + ΣPathPProp (λ fᵈ → isSetAssemblyMorphism asmV asmY _ _) (Category.⋆IdR ASM fᵈ) +Categoryᴰ.⋆Assocᴰ UNIMOD + {X , asmX} {Y , asmY} {Z , asmZ} {W , asmW} {f} {g} {h} + {Xᴰ , asmXᴰ , isModestAsmXᴰ , pX} {Yᴰ , asmYᴰ , isModestAsmYᴰ , pY} {Zᴰ , asmZᴰ , isModestAsmZᴰ , pZ} {Wᴰ , asmWᴰ , isModestAsmWᴰ , pW} + (fᵈ , fᵈcomm) (gᵈ , gᵈcomm) (hᵈ , hᵈcomm) = + ΣPathPProp (λ comp → isSetAssemblyMorphism asmXᴰ asmW _ _) (Category.⋆Assoc ASM fᵈ gᵈ hᵈ ) +Categoryᴰ.isSetHomᴰ UNIMOD = isSetΣ (isSetAssemblyMorphism _ _) (λ f → isProp→isSet (isSetAssemblyMorphism _ _ _ _)) + +UniformFamily : {X : Type ℓ} → (asmX : Assembly X) → Type (ℓ-suc ℓ) +UniformFamily {X} asmX = UNIMOD .Categoryᴰ.ob[_] (X , asmX) From 41d4f1856b036394e31e3a7805aea014dca3d1b0 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Tue, 25 Jun 2024 19:34:12 +0530 Subject: [PATCH 51/61] Finite Limits for PERs --- .../Assembly/LocallyCartesianClosed.agda | 70 ++--- src/Realizability/Modest/UniformFamily.agda | 2 + src/Realizability/PERs/BinProducts.agda | 196 ++++++++++++++ src/Realizability/PERs/PER.agda | 240 +++++++++--------- src/Realizability/PERs/TerminalObject.agda | 12 +- src/Utils/SQElim.agda | 52 ++++ 6 files changed, 395 insertions(+), 177 deletions(-) create mode 100644 src/Realizability/PERs/BinProducts.agda create mode 100644 src/Utils/SQElim.agda diff --git a/src/Realizability/Assembly/LocallyCartesianClosed.agda b/src/Realizability/Assembly/LocallyCartesianClosed.agda index 669ce5f..81b0170 100644 --- a/src/Realizability/Assembly/LocallyCartesianClosed.agda +++ b/src/Realizability/Assembly/LocallyCartesianClosed.agda @@ -270,63 +270,27 @@ module _ {X Y : Type ℓ} {asmX : Assembly X} {asmY : Assembly Y} (f : AssemblyM (sym (λ*ComputationRule realizer a)) (isTrackedH v (pr₂ ⨾ a) pr₂a⊩v .snd (x , sym (isFiberOfY x v fx≡gv)) (pr₁ ⨾ a) pr₁a⊩x) })) + opaque + unfolding pullback + unfolding answer + unfolding answerMap + answerEq : answer ⊚ m ≡ (reindexFunctor ASM PullbacksASM f ⟅ sliceob g ⟆) .S-arr + answerEq = + AssemblyMorphism≡ _ _ + (funExt + λ { q@(x , v , fx≡gv) → + m .map (answerMap q) + ≡⟨ refl ⟩ + m .map (h .map v .fst .snd .fst (x , sym (isFiberOfY x v fx≡gv))) + ≡⟨ h .map v .fst .snd .snd (x , sym (isFiberOfY x v fx≡gv)) ⟩ + x + ∎ }) + opaque unfolding pullback reindex⊣Π[_] : reindexFunctor ASM PullbacksASM f ⊣ Π[_] Iso.fun (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = slicehom (ForwardIso.answer g m h hComm) (ForwardIso.answerEq g m h hComm) - Iso.inv (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = - slicehom answer (AssemblyMorphism≡ _ _ (funExt λ { (x , v , fx≡gv) → answerEq x v fx≡gv })) where - Π[f]m : AssemblyMorphism (S-ob (Π[_] ⟅ sliceob m ⟆) .snd) asmY - Π[f]m = (Π[_] ⟅ sliceob m ⟆) .S-arr - - Q : ASM .Category.ob - Q = (reindexFunctor ASM PullbacksASM f ⟅ sliceob g ⟆) .S-ob - - typeQ : Type ℓ - typeQ = Q .fst - - asmQ : Assembly typeQ - asmQ = Q .snd - - isFiberOfY : (x : X) → (v : V) → f .map x ≡ g .map v → h .map v .fst .fst ≡ f .map x - isFiberOfY x v fx≡gv = - h .map v .fst .fst - ≡[ i ]⟨ hComm i .map v ⟩ - g .map v - ≡⟨ sym fx≡gv ⟩ - f .map x - ∎ - - answerMap : typeQ → P - answerMap (x , v , fx≡gv) = - h .map v .fst .snd .fst - (x , - sym (isFiberOfY x v fx≡gv)) - - answerEq : (x : X) (v : V) (fx≡gv : f .map x ≡ g .map v) → m .map (answerMap (x , v , fx≡gv)) ≡ x - answerEq x v fx≡gv = - m .map (answerMap (x , v , fx≡gv)) - ≡⟨ refl ⟩ - m .map (h .map v .fst .snd .fst (x , sym (isFiberOfY x v fx≡gv))) - ≡⟨ h .map v .fst .snd .snd (x , sym (isFiberOfY x v fx≡gv)) ⟩ - x - ∎ - - answer : AssemblyMorphism asmQ asmP - AssemblyMorphism.map answer = answerMap - AssemblyMorphism.tracker answer = - do - (h~ , isTrackedH) ← h .tracker - let - realizer : Term as 1 - realizer = ` pr₂ ̇ (` h~ ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero) - return - (λ* realizer , - (λ { q@(x , v , fx≡gv) a (pr₁a⊩x , pr₂a⊩v) → - subst - (λ r' → r' ⊩[ asmP ] (answerMap (x , v , fx≡gv))) - (sym (λ*ComputationRule realizer a)) - (isTrackedH v (pr₂ ⨾ a) pr₂a⊩v .snd (x , sym (isFiberOfY x v fx≡gv)) (pr₁ ⨾ a) pr₁a⊩x) })) + Iso.inv (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = slicehom (BackwardIso.answer g m h hComm) (BackwardIso.answerEq g m h hComm) Iso.rightInv (_⊣_.adjIso reindex⊣Π[_] {sliceob {V , asmV} g} {sliceob {P , asmP} m}) (slicehom h hComm) = SliceHom-≡-intro' ASM diff --git a/src/Realizability/Modest/UniformFamily.agda b/src/Realizability/Modest/UniformFamily.agda index 8dfb38a..805446a 100644 --- a/src/Realizability/Modest/UniformFamily.agda +++ b/src/Realizability/Modest/UniformFamily.agda @@ -65,5 +65,7 @@ Categoryᴰ.⋆Assocᴰ UNIMOD ΣPathPProp (λ comp → isSetAssemblyMorphism asmXᴰ asmW _ _) (Category.⋆Assoc ASM fᵈ gᵈ hᵈ ) Categoryᴰ.isSetHomᴰ UNIMOD = isSetΣ (isSetAssemblyMorphism _ _) (λ f → isProp→isSet (isSetAssemblyMorphism _ _ _ _)) +open Categoryᴰ UNIMOD + UniformFamily : {X : Type ℓ} → (asmX : Assembly X) → Type (ℓ-suc ℓ) UniformFamily {X} asmX = UNIMOD .Categoryᴰ.ob[_] (X , asmX) diff --git a/src/Realizability/PERs/BinProducts.agda b/src/Realizability/PERs/BinProducts.agda new file mode 100644 index 0000000..0f9ae26 --- /dev/null +++ b/src/Realizability/PERs/BinProducts.agda @@ -0,0 +1,196 @@ +open import Realizability.ApplicativeStructure +open import Realizability.CombinatoryAlgebra +open import Realizability.PropResizing +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Powerset +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Binary +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.BinProduct +open import Utils.SQElim as SQElim + +module Realizability.PERs.BinProducts + {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.PERs.PER ca +open CombinatoryAlgebra ca +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open PER +open Category PERCat + +module _ (R S : PER) where + binProdObPER : PER + PER.relation binProdObPER = + (λ a b → (pr₁ ⨾ a) ~[ R ] (pr₁ ⨾ b) × (pr₂ ⨾ a) ~[ S ] (pr₂ ⨾ b)) , λ a b → isProp× (isProp~ (pr₁ ⨾ a) R (pr₁ ⨾ b)) (isProp~ (pr₂ ⨾ a) S (pr₂ ⨾ b)) + fst (isPER binProdObPER) a b (pr₁a~[R]pr₁b , pr₂a~[S]pr₂b) = + (isSymmetric R (pr₁ ⨾ a) (pr₁ ⨾ b) pr₁a~[R]pr₁b) , (isSymmetric S (pr₂ ⨾ a) (pr₂ ⨾ b) pr₂a~[S]pr₂b) + snd (isPER binProdObPER) a b c (pr₁a~[R]pr₁b , pr₂a~[S]pr₂b) (pr₁b~[R]pr₁c , pr₂b~[S]pr₂c) = + (isTransitive R (pr₁ ⨾ a) (pr₁ ⨾ b) (pr₁ ⨾ c) pr₁a~[R]pr₁b pr₁b~[R]pr₁c) , (isTransitive S (pr₂ ⨾ a) (pr₂ ⨾ b) (pr₂ ⨾ c) pr₂a~[S]pr₂b pr₂b~[S]pr₂c) + + isTrackerPr₁ : isTracker binProdObPER R pr₁ + isTrackerPr₁ = λ r r' (pr₁r~[R]pr₁r' , pr₂r~[S]pr₂r') → pr₁r~[R]pr₁r' + + binProdPr₁Tracker : perTracker binProdObPER R + binProdPr₁Tracker = pr₁ , isTrackerPr₁ + + binProdPr₁PER : perMorphism binProdObPER R + binProdPr₁PER = [ binProdPr₁Tracker ] + + isTrackerPr₂ : isTracker binProdObPER S pr₂ + isTrackerPr₂ = λ { r r' (pr₁r~[R]pr₁r' , pr₂r~[S]pr₂r') → pr₂r~[S]pr₂r' } + + binProdPr₂Tracker : perTracker binProdObPER S + binProdPr₂Tracker = pr₂ , isTrackerPr₂ + + binProdPr₂PER : perMorphism binProdObPER S + binProdPr₂PER = [ binProdPr₂Tracker ] + + binProdUnivPropPER : + (T : PER) + (f : perMorphism T R) + (g : perMorphism T S) → + ∃![ ! ∈ perMorphism T binProdObPER ] ((composePerMorphism T binProdObPER R ! binProdPr₁PER ≡ f) × (composePerMorphism T binProdObPER S ! binProdPr₂PER ≡ g)) + binProdUnivPropPER T f g = + inhProp→isContr + map + (isPropMapType f g) where + + mapEqProj1 : ∀ ! f' → Type _ + mapEqProj1 ! f' = composePerMorphism T binProdObPER R ! binProdPr₁PER ≡ f' + + mapEqProj2 : ∀ ! g' → Type _ + mapEqProj2 ! g' = composePerMorphism T binProdObPER S ! binProdPr₂PER ≡ g' + + mapEqs : ∀ ! f' g' → Type _ + mapEqs ! f' g' = (composePerMorphism T binProdObPER R ! binProdPr₁PER ≡ f') × (composePerMorphism T binProdObPER S ! binProdPr₂PER ≡ g') + + isPropMapEqs : ∀ ! f' g' → isProp (mapEqs ! f' g') + isPropMapEqs ! f' g' = isProp× (squash/ _ _) (squash/ _ _) + + mapType : ∀ f' g' → Type _ + mapType f' g' = Σ[ ! ∈ perMorphism T binProdObPER ] (mapEqs ! f' g') + + mapRealizer : ∀ a b → Term as 1 + mapRealizer a b = ` pair ̇ (` a ̇ # zero) ̇ (` b ̇ # zero) + + isSetMapType : ∀ f' g' → isSet (mapType f' g') + isSetMapType f' g' = isSetΣ squash/ (λ ! → isSet× (isProp→isSet (squash/ _ _)) (isProp→isSet (squash/ _ _))) + + isPropMapType : ∀ f' g' → isProp (mapType f' g') + isPropMapType f' g' (! , !⋆π₁≡f , !⋆π₂≡g) (!' , !'⋆π₁≡f , !'⋆π₂≡g) = + Σ≡Prop + (λ ! → isPropMapEqs ! f' g') + (SQ.elimProp4 + {P = motive} + isPropMotive + goal + f' g' ! !' + !⋆π₁≡f + !⋆π₂≡g + !'⋆π₁≡f + !'⋆π₂≡g) where + + motive : ∀ f' g' ! !' → Type _ + motive f' g' ! !' = ∀ (!proj1 : mapEqProj1 ! f') (!proj2 : mapEqProj2 ! g') (!'proj1 : mapEqProj1 !' f') (!'proj2 : mapEqProj2 !' g') → ! ≡ !' + + isPropMotive : ∀ f' g' ! !' → isProp (motive f' g' ! !') + isPropMotive f' g' ! !' = + isPropΠ4 λ _ _ _ _ → squash/ _ _ + + goal : ∀ f' g' ! !' → motive [ f' ] [ g' ] [ ! ] [ !' ] + goal (f , f⊩f) (g , g⊩g) (a , a⊩!) (b , b⊩!') !proj1 !proj2 !'proj1 !'proj2 = + eq/ _ _ + λ r r~r → + let + pr₁Equiv : (pr₁ ⨾ (a ⨾ r)) ~[ R ] (pr₁ ⨾ (b ⨾ r)) + pr₁Equiv = + isTransitive + R (pr₁ ⨾ (a ⨾ r)) (f ⨾ r) (pr₁ ⨾ (b ⨾ r)) + (subst (_~[ R ] (f ⨾ r)) (λ*ComputationRule (` pr₁ ̇ (` a ̇ # zero)) r) (!proj1Equiv r r~r)) + (isSymmetric R (pr₁ ⨾ (b ⨾ r)) (f ⨾ r) (subst (_~[ R ] (f ⨾ r)) (λ*ComputationRule (` pr₁ ̇ (` b ̇ # zero)) r) (!'proj1Equiv r r~r))) + + pr₂Equiv : (pr₂ ⨾ (a ⨾ r)) ~[ S ] (pr₂ ⨾ (b ⨾ r)) + pr₂Equiv = + isTransitive + S (pr₂ ⨾ (a ⨾ r)) (g ⨾ r) (pr₂ ⨾ (b ⨾ r)) + (subst (_~[ S ] (g ⨾ r)) (λ*ComputationRule (` pr₂ ̇ (` a ̇ # zero)) r) (!proj2Equiv r r~r)) + (isSymmetric S (pr₂ ⨾ (b ⨾ r)) (g ⨾ r) (subst (_~[ S ] (g ⨾ r)) (λ*ComputationRule (` pr₂ ̇ (` b ̇ # zero)) r) (!'proj2Equiv r r~r))) + in + pr₁Equiv , + pr₂Equiv where + !proj1Equiv : isEquivTracker T R (composePerTracker T binProdObPER R (a , a⊩!) binProdPr₁Tracker) (f , f⊩f) + !proj1Equiv = effectiveIsEquivTracker T R _ _ !proj1 + + !proj2Equiv : isEquivTracker T S (composePerTracker T binProdObPER S (a , a⊩!) binProdPr₂Tracker) (g , g⊩g) + !proj2Equiv = effectiveIsEquivTracker T S _ _ !proj2 + + !'proj1Equiv : isEquivTracker T R (composePerTracker T binProdObPER R (b , b⊩!') binProdPr₁Tracker) (f , f⊩f) + !'proj1Equiv = effectiveIsEquivTracker T R _ _ !'proj1 + + !'proj2Equiv : isEquivTracker T S (composePerTracker T binProdObPER S (b , b⊩!') binProdPr₂Tracker) (g , g⊩g) + !'proj2Equiv = effectiveIsEquivTracker T S _ _ !'proj2 + + coreMap : ∀ a b → mapType [ a ] [ b ] + coreMap (a , a⊩f) (b , b⊩g) = + [ (λ* (mapRealizer a b)) , + (λ r r' r~r' → + subst2 + (λ abr abr' → abr ~[ binProdObPER ] abr') + (sym (λ*ComputationRule (mapRealizer a b) r)) + (sym (λ*ComputationRule (mapRealizer a b) r')) + (subst2 + (λ ar ar' → ar ~[ R ] ar') + (sym (pr₁pxy≡x _ _)) + (sym (pr₁pxy≡x _ _)) + (a⊩f r r' r~r') , + subst2 + (λ br br' → br ~[ S ] br') + (sym (pr₂pxy≡y _ _)) + (sym (pr₂pxy≡y _ _)) + (b⊩g r r' r~r'))) ] , + eq/ _ _ + (λ r r~r → + subst + (_~[ R ] (a ⨾ r)) + (sym + (λ*ComputationRule (` pr₁ ̇ (` λ* (mapRealizer a b) ̇ # zero)) r ∙ + cong (pr₁ ⨾_) (λ*ComputationRule (mapRealizer a b) r))) + (subst (_~[ R ] (a ⨾ r)) (sym (pr₁pxy≡x _ _)) (a⊩f r r r~r))) , + eq/ _ _ + λ r r~r → + subst + (_~[ S ] (b ⨾ r)) + (sym + (λ*ComputationRule (` pr₂ ̇ (` λ* (mapRealizer a b) ̇ # zero)) r ∙ + cong (pr₂ ⨾_) (λ*ComputationRule (mapRealizer a b) r) ∙ + pr₂pxy≡y _ _)) + (b⊩g r r r~r) + + map : mapType f g + map = + SQ.elimProp2 + {P = mapType} + isPropMapType + coreMap + f g + +BinProductPER : (R S : PER) → BinProduct PERCat R S +BinProduct.binProdOb (BinProductPER R S) = binProdObPER R S +BinProduct.binProdPr₁ (BinProductPER R S) = binProdPr₁PER R S +BinProduct.binProdPr₂ (BinProductPER R S) = binProdPr₂PER R S +BinProduct.univProp (BinProductPER R S) {T} f g = binProdUnivPropPER R S T f g + +BinProductsPER : BinProducts PERCat +BinProductsPER R S = BinProductPER R S diff --git a/src/Realizability/PERs/PER.agda b/src/Realizability/PERs/PER.agda index bd5adb7..90ca204 100644 --- a/src/Realizability/PERs/PER.agda +++ b/src/Realizability/PERs/PER.agda @@ -12,6 +12,7 @@ open import Cubical.Functions.FunExtEquiv open import Cubical.Relation.Binary open import Cubical.Data.Sigma open import Cubical.Data.FinData +open import Cubical.Data.Vec open import Cubical.Reflection.RecordEquiv open import Cubical.HITs.PropositionalTruncation as PT hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad @@ -52,126 +53,125 @@ record PER : Type (ℓ-suc ℓ) where open PER -module _ (R : PER) where - Quotient = A / ∣ R ∣ - - -- mimics the proof at Cubical.HITs.SetQuotients.Properties - effectiveOnDomain : ∀ a b → ∣ R ∣ a a → [ a ] ≡ [ b ] → ∣ R ∣ a b - effectiveOnDomain a b aRa [a]≡[b] = transport aRa≡aRb aRa where - helper : Quotient → hProp _ - helper x = - SQ.rec - isSetHProp - (λ c → (∣ R ∣ a c) , (isPropValued R a c)) - (λ c d cRd → - Σ≡Prop - (λ _ → isPropIsProp) - (hPropExt - (isPropValued R a c) - (isPropValued R a d) - (λ aRc → isTransitive R a c d aRc cRd) - (λ aRd → isTransitive R a d c aRd (isSymmetric R c d cRd)))) - x - - aRa≡aRb : ∣ R ∣ a a ≡ ∣ R ∣ a b - aRa≡aRb i = helper ([a]≡[b] i) .fst - -record PERMorphism (R S : PER) : Type ℓ where - no-eta-equality - constructor makePERMorphism - field - map : Quotient R → Quotient S - isTracked : ∃[ e ∈ A ] (∀ (a : A) → ∣ R ∣ a a → (map [ a ] ≡ [ e ⨾ a ]) × ∣ S ∣ (e ⨾ a) (e ⨾ a)) - -open PERMorphism - -unquoteDecl PERMorphismIsoΣ = declareRecordIsoΣ PERMorphismIsoΣ (quote PERMorphism) - -PERMorphismΣ : PER → PER → Type ℓ -PERMorphismΣ R S = Σ[ map ∈ (Quotient R → Quotient S) ] ∃[ e ∈ A ] (∀ (a : A) → ∣ R ∣ a a → (map [ a ] ≡ [ e ⨾ a ]) × ∣ S ∣ (e ⨾ a) (e ⨾ a)) - -isSetPERMorphismΣ : ∀ {R S} → isSet (PERMorphismΣ R S) -isSetPERMorphismΣ {R} {S} = isSetΣ (isSet→ squash/) (λ map → isProp→isSet isPropPropTrunc) - -isSetPERMorphism : ∀ {R S} → isSet (PERMorphism R S) -isSetPERMorphism {R} {S} = subst (λ type → isSet type) (sym (isoToPath PERMorphismIsoΣ)) (isSetPERMorphismΣ {R = R} {S = S}) - -PERMorphism≡Iso : ∀ {R S} → (f g : PERMorphism R S) → Iso (f ≡ g) (f .map ≡ g .map) -Iso.fun (PERMorphism≡Iso {R} {S} f g) f≡g i = f≡g i .map -map (Iso.inv (PERMorphism≡Iso {R} {S} f g) fMap≡gMap i) = fMap≡gMap i -isTracked (Iso.inv (PERMorphism≡Iso {R} {S} f g) fMap≡gMap i) = - isProp→PathP - (λ i → - isPropPropTrunc - {A = Σ[ e ∈ A ] (∀ (a : A) → ∣ R ∣ a a → ((fMap≡gMap i) [ a ] ≡ [ e ⨾ a ]) × ∣ S ∣ (e ⨾ a) (e ⨾ a))}) - (f .isTracked) (g .isTracked) i -Iso.rightInv (PERMorphism≡Iso {R} {S} f g) fMap≡gMap = refl -Iso.leftInv (PERMorphism≡Iso {R} {S} f g) f≡g = isSetPERMorphism f g (Iso.inv (PERMorphism≡Iso f g) (λ i → f≡g i .map)) f≡g - -PERMorphism≡ : ∀ {R S} → (f g : PERMorphism R S) → f .map ≡ g .map → f ≡ g -PERMorphism≡ {R} {S} f g fMap≡gMap = Iso.inv (PERMorphism≡Iso f g) fMap≡gMap - -idPERMorphism : ∀ {R : PER} → PERMorphism R R -map (idPERMorphism {R}) x = x -isTracked (idPERMorphism {R}) = - return (Id , (λ a aRa → (subst (λ r → [ a ] ≡ [ r ]) (sym (Ida≡a a)) refl) , (subst (λ r → ∣ R ∣ r r) (sym (Ida≡a a)) aRa))) - -composePERMorphism : ∀ {R S T : PER} → PERMorphism R S → PERMorphism S T → PERMorphism R T -map (composePERMorphism {R} {S} {T} f g) x = g .map (f .map x) -isTracked (composePERMorphism {R} {S} {T} f g) = - do - (f~ , isTrackedF) ← f .isTracked - (g~ , isTrackedG) ← g .isTracked - let - realizer : Term as 1 - realizer = ` g~ ̇ (` f~ ̇ # zero) - return - (λ* realizer , - (λ a aRa → - (g .map (f .map [ a ]) - ≡⟨ cong (g .map) (isTrackedF a aRa .fst) ⟩ - g .map [ f~ ⨾ a ] - ≡⟨ isTrackedG (f~ ⨾ a) (isTrackedF a aRa .snd) .fst ⟩ - [ g~ ⨾ (f~ ⨾ a) ] - ≡⟨ cong [_] (sym (λ*ComputationRule realizer a)) ⟩ - [ λ* realizer ⨾ a ] - ∎) , - (subst (λ r' → ∣ T ∣ r' r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~ ⨾ a) (isTrackedF a aRa .snd) .snd)))) - --- all definitional -idLPERMorphism : ∀ {R S} → (f : PERMorphism R S) → composePERMorphism idPERMorphism f ≡ f -idLPERMorphism {R} {S} f = PERMorphism≡ _ _ refl - -idRPERMorphism : ∀ {R S} → (f : PERMorphism R S) → composePERMorphism f idPERMorphism ≡ f -idRPERMorphism {R} {S} f = PERMorphism≡ _ _ refl - -assocPERMorphism : - ∀ {R S T U} - → (f : PERMorphism R S) - → (g : PERMorphism S T) - → (h : PERMorphism T U) - → composePERMorphism (composePERMorphism f g) h ≡ composePERMorphism f (composePERMorphism g h) -assocPERMorphism {R} {S} {T} {U} f g h = PERMorphism≡ _ _ refl +_~[_]_ : A → PER → A → Type ℓ +a ~[ R ] b = R .relation .fst a b + +isProp~ : ∀ a R b → isProp (a ~[ R ] b) +isProp~ a R b = R .relation .snd a b + +isTracker : (R S : PER) → A → Type ℓ +isTracker R S a = ∀ r r' → r ~[ R ] r' → (a ⨾ r) ~[ S ] (a ⨾ r') + +perTracker : (R S : PER) → Type ℓ +perTracker R S = Σ[ a ∈ A ] isTracker R S a + +isEquivTracker : (R S : PER) → perTracker R S → perTracker R S → Type ℓ +isEquivTracker R S (a , _) (b , _) = (∀ r → r ~[ R ] r → (a ⨾ r) ~[ S ] (b ⨾ r)) + +isEquivRelIsEquivTracker : (R S : PER) → BR.isEquivRel (isEquivTracker R S) +BinaryRelation.isEquivRel.reflexive (isEquivRelIsEquivTracker R S) (a , isTrackerA) r r~r = isTrackerA r r r~r +BinaryRelation.isEquivRel.symmetric (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) a~b r r~r = isSymmetric S (a ⨾ r) (b ⨾ r) (a~b r r~r) +BinaryRelation.isEquivRel.transitive (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) (c , isTrackerC) a~b b~c r r~r = isTransitive S (a ⨾ r) (b ⨾ r) (c ⨾ r) (a~b r r~r) (b~c r r~r) + +isPropIsEquivTracker : ∀ R S a b → isProp (isEquivTracker R S a b) +isPropIsEquivTracker R S (a , isTrackerA) (b , isTrackerB) = isPropΠ2 λ r r~r → isPropValued S (a ⨾ r) (b ⨾ r) + +isEffectiveIsEquivTracker : ∀ R S → BR.isEffective (isEquivTracker R S) +isEffectiveIsEquivTracker R S = isEquivRel→isEffective (isPropIsEquivTracker R S) (isEquivRelIsEquivTracker R S) + +perMorphism : (R S : PER) → Type ℓ +perMorphism R S = perTracker R S / (isEquivTracker R S) + +effectiveIsEquivTracker : ∀ R S a b → [ a ] ≡ [ b ] → isEquivTracker R S a b +effectiveIsEquivTracker R S a b eq' = effective (isPropIsEquivTracker R S) (isEquivRelIsEquivTracker R S) a b eq' + +isSetPerMorphism : ∀ R S → isSet (perMorphism R S) +isSetPerMorphism R S = squash/ + +idPerMorphism : (R : PER) → perMorphism R R +idPerMorphism R = [ Id , (λ r r' r~r' → subst2 (λ r r' → r ~[ R ] r') (sym (Ida≡a r)) (sym (Ida≡a r')) r~r') ] + +composePerTracker : (R S T : PER) → perTracker R S → perTracker S T → perTracker R T +composePerTracker R S T (a , a⊩f) (b , b⊩g) = + let + realizer : Term as 1 + realizer = ` b ̇ (` a ̇ # zero) + in + λ* realizer , + λ r r' r~r' → + subst2 + _~[ T ]_ + (sym (λ*ComputationRule realizer r)) + (sym (λ*ComputationRule realizer r')) + (b⊩g (a ⨾ r) (a ⨾ r') (a⊩f r r' r~r')) + +composePerMorphism : (R S T : PER) → perMorphism R S → perMorphism S T → perMorphism R T +composePerMorphism R S T f g = + SQ.rec2 + squash/ + (λ { (a , a⊩f) (b , b⊩g) → + [ composePerTracker R S T (a , a⊩f) (b , b⊩g) ] }) + (λ { (a , a⊩f) (b , b⊩f) (c , c⊩g) a~b → + eq/ _ _ + λ r r~r → + subst2 + (λ car cbr → car ~[ T ] cbr) + (sym (λ*ComputationRule (` c ̇ (` a ̇ # zero)) r)) + (sym (λ*ComputationRule (` c ̇ (` b ̇ # zero)) r)) + (c⊩g (a ⨾ r) (b ⨾ r) (a~b r r~r)) }) + (λ { (a , a⊩f) (b , b⊩g) (c , c⊩g) b~c → + eq/ _ _ + λ r r~r → + subst2 + (λ bar car → bar ~[ T ] car) + (sym (λ*ComputationRule (` b ̇ (` a ̇ # zero)) r)) + (sym (λ*ComputationRule (` c ̇ (` a ̇ # zero)) r)) + (b~c (a ⨾ r) (a⊩f r r r~r)) }) + f g + +idLPerMorphism : ∀ R S f → composePerMorphism R R S (idPerMorphism R) f ≡ f +idLPerMorphism R S f = + SQ.elimProp + (λ f → squash/ (composePerMorphism R R S (idPerMorphism R) f) f) + (λ { (a , a⊩f) → + eq/ _ _ + λ r r~r → + subst + (λ ar → ar ~[ S ] (a ⨾ r)) + (sym (λ*ComputationRule (` a ̇ (` Id ̇ # zero)) r ∙ cong (λ x → a ⨾ x) (Ida≡a r))) + (a⊩f r r r~r) }) + f + +idRPerMorphism : ∀ R S f → composePerMorphism R S S f (idPerMorphism S) ≡ f +idRPerMorphism R S f = + SQ.elimProp + (λ f → squash/ (composePerMorphism R S S f (idPerMorphism S)) f) + (λ { (a , a⊩f) → + eq/ _ _ + λ r r~r → + subst (λ ar → ar ~[ S ] (a ⨾ r)) (sym (λ*ComputationRule (` Id ̇ (` a ̇ # zero)) r ∙ Ida≡a (a ⨾ r))) (a⊩f r r r~r) }) + f + +assocPerMorphism : ∀ R S T U f g h → composePerMorphism R T U (composePerMorphism R S T f g) h ≡ composePerMorphism R S U f (composePerMorphism S T U g h) +assocPerMorphism R S T U f g h = + SQ.elimProp3 + (λ f g h → squash/ (composePerMorphism R T U (composePerMorphism R S T f g) h) (composePerMorphism R S U f (composePerMorphism S T U g h))) + (λ { (a , a⊩f) (b , b⊩g) (c , c⊩h) → + eq/ _ _ + λ r r~r → + subst2 + (λ cba cba' → cba ~[ U ] cba') + (sym (λ*ComputationRule (` c ̇ (` ⟦ as ⟧ (λ*abst (` b ̇ (` a ̇ # zero))) [] ̇ # zero)) r ∙ cong (λ bar → c ⨾ bar) (λ*ComputationRule (` b ̇ (` a ̇ # zero)) r))) + (sym (λ*ComputationRule (` ⟦ as ⟧ (λ*abst (` c ̇ (` b ̇ # zero))) [] ̇ (` a ̇ # zero)) r ∙ λ*ComputationRule (` c ̇ (` b ̇ # zero)) (a ⨾ r))) + (c⊩h (b ⨾ (a ⨾ r)) (b ⨾ (a ⨾ r)) (b⊩g (a ⨾ r) (a ⨾ r) (a⊩f r r r~r)) ) }) + f g h PERCat : Category (ℓ-suc ℓ) ℓ Category.ob PERCat = PER -Category.Hom[_,_] PERCat R S = PERMorphism R S -Category.id PERCat {R} = idPERMorphism -Category._⋆_ PERCat {R} {S} {T} f g = composePERMorphism f g -Category.⋆IdL PERCat f = idLPERMorphism f -Category.⋆IdR PERCat f = idRPERMorphism f -Category.⋆Assoc PERCat f g h = assocPERMorphism f g h -Category.isSetHom PERCat = isSetPERMorphism - -open Assembly - -inclusion : Functor PERCat ASM -fst (Functor.F-ob inclusion per) = Σ[ a ∈ A ] ∣ per ∣ a a -(snd (Functor.F-ob inclusion per)) ._⊩_ r (a , aRa) = ∣ per ∣ r a -isSetX (snd (Functor.F-ob inclusion per)) = isSetΣ isSetA (λ x → isProp→isSet (isPropValued per x x)) -⊩isPropValued (snd (Functor.F-ob inclusion per)) r (a , aRa) = isPropValued per r a -⊩surjective (snd (Functor.F-ob inclusion per)) (a , aRa) = return (a , aRa) -AssemblyMorphism.map (Functor.F-hom inclusion morphism) (a , aRa) = {!!} -AssemblyMorphism.tracker (Functor.F-hom inclusion morphism) = {!!} -Functor.F-id inclusion = {!!} -Functor.F-seq inclusion = {!!} +Category.Hom[_,_] PERCat R S = perMorphism R S +Category.id PERCat {R} = idPerMorphism R +Category._⋆_ PERCat {R} {S} {T} f g = composePerMorphism R S T f g +Category.⋆IdL PERCat {R} {S} f = idLPerMorphism R S f +Category.⋆IdR PERCat {R} {S} f = idRPerMorphism R S f +Category.⋆Assoc PERCat {R} {S} {T} {U} f g h = assocPerMorphism R S T U f g h +Category.isSetHom PERCat {R} {S} = isSetPerMorphism R S diff --git a/src/Realizability/PERs/TerminalObject.agda b/src/Realizability/PERs/TerminalObject.agda index c3a6576..5f2bd62 100644 --- a/src/Realizability/PERs/TerminalObject.agda +++ b/src/Realizability/PERs/TerminalObject.agda @@ -26,7 +26,6 @@ module Realizability.PERs.TerminalObject open import Realizability.PERs.PER ca open CombinatoryAlgebra ca open Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open PERMorphism terminalPER : PER PER.relation terminalPER = (λ x y → Unit*) , λ x y → isPropUnit* @@ -36,8 +35,13 @@ snd (PER.isPER terminalPER) = λ a b c x x₁ → tt* isTerminalTerminalPER : isTerminal PERCat terminalPER isTerminalTerminalPER X = inhProp→isContr - (makePERMorphism (λ x → [ k ]) (return (Id , (λ a aXa → (eq/ k (Id ⨾ a) tt*) , tt*)))) - λ x y → PERMorphism≡ x y (funExt λ q → {!!}) + [ k , (λ r r' r~r' → tt*) ] + λ ! !' → + SQ.elimProp2 + (λ ! !' → squash/ ! !') + (λ { (a , a⊩!) (b , b⊩!) → + eq/ _ _ λ r r~r → tt* }) + ! !' terminal : Terminal PERCat -terminal = terminalPER , {!!} +terminal = terminalPER , isTerminalTerminalPER diff --git a/src/Utils/SQElim.agda b/src/Utils/SQElim.agda new file mode 100644 index 0000000..d138409 --- /dev/null +++ b/src/Utils/SQElim.agda @@ -0,0 +1,52 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.HITs.SetQuotients as SQ + +module Utils.SQElim where + +private + variable + ℓ ℓ' ℓ'' : Level + A B C Q : Type ℓ + R S T W : A → A → Type ℓ + +elim2 : {P : A / R → B / S → Type ℓ} + → (∀ x y → isSet (P x y)) + → (f : ∀ a b → P [ a ] [ b ]) + → (∀ a b c → (s : S b c) → PathP (λ i → P [ a ] (eq/ b c s i)) (f a b) (f a c)) + → (∀ a b c → (r : R a b) → PathP (λ i → P (eq/ a b r i) [ c ]) (f a c) (f b c)) + → ∀ x y → P x y +elim2 {P = P} isSetP f coh1 coh2 x y = + SQ.elim + {P = λ x → ∀ y → P x y} + (λ x → isSetΠ λ y → isSetP x y) + (λ a y → + SQ.elim + {P = λ y → P [ a ] y} + (λ y → isSetP [ a ] y) + (λ b → f a b) + (λ b c r → coh1 a b c r) + y) + (λ a b r → + funExt + λ z → + SQ.elimProp + {P = λ z → PathP (λ z₁ → P (eq/ a b r z₁) z) (elim (λ y₁ → isSetP [ a ] y₁) (λ b₁ → f a b₁) (λ b₁ c r₁ → coh1 a b₁ c r₁) z) (elim (λ y₁ → isSetP [ b ] y₁) (λ b₁ → f b b₁) (λ b₁ c r₁ → coh1 b b₁ c r₁) z)} + (λ z p q i j → + isSet→SquareP + {A = λ i' j' → P (eq/ a b r j') z} + (λ i' j' → isSetP (eq/ a b r j') z) + {a₀₀ = elim (isSetP [ a ]) (f a) (coh1 a) z} + {a₀₁ = elim (isSetP [ b ]) (f b) (coh1 b) z} + p + {a₁₀ = elim (isSetP [ a ]) (f a) (coh1 a) z} + {a₁₁ = elim (isSetP [ b ]) (f b) (coh1 b) z} + q + refl + refl + i j) + (λ c → coh2 a b c r) + z) + x y + From 73d04b5d95afbc6742a07a4b8359e1d893a05be8 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Wed, 26 Jun 2024 22:02:10 +0530 Subject: [PATCH 52/61] Families of modest sets as displayed category --- src/Realizability/Modest/UniformFamily.agda | 220 +++++++++++++++--- src/Realizability/PERs/PER.agda | 7 +- src/Realizability/PERs/ResizedPER.agda | 91 ++++++++ src/Realizability/PERs/SystemF.agda | 6 + .../PERs/UniformFamiliesOverAsm.agda | 78 +++++++ 5 files changed, 362 insertions(+), 40 deletions(-) create mode 100644 src/Realizability/PERs/ResizedPER.agda create mode 100644 src/Realizability/PERs/UniformFamiliesOverAsm.agda diff --git a/src/Realizability/Modest/UniformFamily.agda b/src/Realizability/Modest/UniformFamily.agda index 805446a..d2aa9aa 100644 --- a/src/Realizability/Modest/UniformFamily.agda +++ b/src/Realizability/Modest/UniformFamily.agda @@ -4,6 +4,7 @@ open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function open import Cubical.Foundations.Equiv open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path open import Cubical.Foundations.Structure using (⟨_⟩; str) open import Cubical.Data.Sigma open import Cubical.Data.FinData @@ -13,8 +14,9 @@ open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Reflection.RecordEquiv open import Cubical.Categories.Category open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning open import Cubical.Categories.Limits.Pullback -open import Cubical.Categories.Functor +open import Cubical.Categories.Functor hiding (Id) open import Cubical.Categories.Constructions.Slice open import Realizability.CombinatoryAlgebra open import Realizability.ApplicativeStructure @@ -22,50 +24,194 @@ open import Realizability.PropResizing module Realizability.Modest.UniformFamily {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where - open import Realizability.Assembly.Base ca open import Realizability.Assembly.Morphism ca open import Realizability.Assembly.Terminal ca -open import Realizability.Assembly.LocallyCartesianClosed ca open import Realizability.Modest.Base ca open Assembly open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -UNIMOD : Categoryᴰ ASM (ℓ-suc ℓ) ℓ -Categoryᴰ.ob[ UNIMOD ] (X , asmX) = Σ[ Y ∈ Type ℓ ] Σ[ asmY ∈ Assembly Y ] isModest asmY × AssemblyMorphism asmY asmX -Categoryᴰ.Hom[_][_,_] UNIMOD {X , asmX} {Y , asmY} reindex (V , asmV , isModestAsmV , m) (W , asmW , isModestAsmW , n) = Σ[ reindexᵈ ∈ (AssemblyMorphism asmV asmW) ] m ⊚ reindex ≡ reindexᵈ ⊚ n -Categoryᴰ.idᴰ UNIMOD {X , asmX} {V , asmV , isModestAsmV , m} = (identityMorphism asmV) , (Category.⋆IdR ASM m ∙ sym (Category.⋆IdL ASM m)) -Categoryᴰ._⋆ᴰ_ UNIMOD - {X , asmX} {Y , asmY} {Z , asmZ} {f} {g} - {U , asmU , isModestU , m} {V , asmV , isModestV , n} {W , asmW , isModestW , o} - (fᵈ , fᵈcomm) (gᵈ , gᵈcomm) = - (fᵈ ⊚ gᵈ) , - (m ⊚ (f ⊚ g) - ≡⟨ sym (Category.⋆Assoc ASM m f g) ⟩ - (m ⊚ f) ⊚ g - ≡⟨ cong (λ x → x ⊚ g) fᵈcomm ⟩ - fᵈ ⊚ n ⊚ g - ≡⟨ Category.⋆Assoc ASM fᵈ n g ⟩ - fᵈ ⊚ (n ⊚ g) - ≡⟨ cong (fᵈ ⊚_) gᵈcomm ⟩ - fᵈ ⊚ (gᵈ ⊚ o) - ≡⟨ sym (Category.⋆Assoc ASM fᵈ gᵈ o) ⟩ - fᵈ ⊚ gᵈ ⊚ o - ∎) -Categoryᴰ.⋆IdLᴰ UNIMOD {X , asmX} {Y , asmY} {f} {V , asmV , isModestAsmV , m} {W , asmW , isModestAsmW , n} (fᵈ , fᵈcomm) = - ΣPathPProp (λ fᵈ → isSetAssemblyMorphism asmV asmY _ _) (Category.⋆IdL ASM fᵈ) -Categoryᴰ.⋆IdRᴰ UNIMOD {X , asmX} {Y , asmY} {f} {V , asmV , isModestAsmV , m} {W , asmW , isModestAsmW , n} (fᵈ , fᵈcomm) = - ΣPathPProp (λ fᵈ → isSetAssemblyMorphism asmV asmY _ _) (Category.⋆IdR ASM fᵈ) -Categoryᴰ.⋆Assocᴰ UNIMOD - {X , asmX} {Y , asmY} {Z , asmZ} {W , asmW} {f} {g} {h} - {Xᴰ , asmXᴰ , isModestAsmXᴰ , pX} {Yᴰ , asmYᴰ , isModestAsmYᴰ , pY} {Zᴰ , asmZᴰ , isModestAsmZᴰ , pZ} {Wᴰ , asmWᴰ , isModestAsmWᴰ , pW} - (fᵈ , fᵈcomm) (gᵈ , gᵈcomm) (hᵈ , hᵈcomm) = - ΣPathPProp (λ comp → isSetAssemblyMorphism asmXᴰ asmW _ _) (Category.⋆Assoc ASM fᵈ gᵈ hᵈ ) -Categoryᴰ.isSetHomᴰ UNIMOD = isSetΣ (isSetAssemblyMorphism _ _) (λ f → isProp→isSet (isSetAssemblyMorphism _ _ _ _)) +record UniformFamily {I : Type ℓ} (asmI : Assembly I) : Type (ℓ-suc ℓ) where + no-eta-equality + field + carriers : I → Type ℓ + assemblies : ∀ i → Assembly (carriers i) + isModestFamily : ∀ i → isModest (assemblies i) +open UniformFamily +record DisplayedUFamMap {I J : Type ℓ} (asmI : Assembly I) (asmJ : Assembly J) (u : AssemblyMorphism asmI asmJ) (X : UniformFamily asmI) (Y : UniformFamily asmJ) : Type ℓ where + no-eta-equality + field + fibrewiseMap : ∀ i → X .carriers i → Y .carriers (u .map i) + isTracked : ∃[ e ∈ A ] (∀ (i : I) (a : A) (a⊩i : a ⊩[ asmI ] i) (x : X .carriers i) (b : A) (b⊩x : b ⊩[ X .assemblies i ] x) → (e ⨾ a ⨾ b) ⊩[ Y .assemblies (u .map i) ] (fibrewiseMap i x)) + +open DisplayedUFamMap + +DisplayedUFamMapPathP : + ∀ {I J} (asmI : Assembly I) (asmJ : Assembly J) → + ∀ u v X Y + → (uᴰ : DisplayedUFamMap asmI asmJ u X Y) + → (vᴰ : DisplayedUFamMap asmI asmJ v X Y) + → (p : u ≡ v) + → (∀ (i : I) (x : X .carriers i) → PathP (λ j → Y .carriers (p j .map i)) (uᴰ .fibrewiseMap i x) (vᴰ .fibrewiseMap i x)) + ----------------------------------------------------------------------------------------------------------------------- + → PathP (λ i → DisplayedUFamMap asmI asmJ (p i) X Y) uᴰ vᴰ +fibrewiseMap (DisplayedUFamMapPathP {I} {J} asmI asmJ u v X Y uᴰ vᴰ p pᴰ dimI) i x = pᴰ i x dimI +isTracked (DisplayedUFamMapPathP {I} {J} asmI asmJ u v X Y uᴰ vᴰ p pᴰ dimI) = + isProp→PathP + {B = λ dimJ → ∃[ e ∈ A ] ((i : I) (a : A) → a ⊩[ asmI ] i → (x : X .carriers i) (b : A) → b ⊩[ X .assemblies i ] x → (e ⨾ a ⨾ b) ⊩[ Y .assemblies (p dimJ .map i) ] pᴰ i x dimJ)} + (λ dimJ → isPropPropTrunc) + (uᴰ .isTracked) + (vᴰ .isTracked) + dimI + +isSetDisplayedUFamMap : ∀ {I J} (asmI : Assembly I) (asmJ : Assembly J) → ∀ u X Y → isSet (DisplayedUFamMap asmI asmJ u X Y) +fibrewiseMap (isSetDisplayedUFamMap {I} {J} asmI asmJ u X Y f g p q dimI dimJ) i x = + isSet→isSet' + (Y .assemblies (u .map i) .isSetX) + {a₀₀ = fibrewiseMap f i x} + {a₀₁ = fibrewiseMap f i x} + refl + {a₁₀ = fibrewiseMap g i x} + {a₁₁ = fibrewiseMap g i x} + refl + (λ dimK → fibrewiseMap (p dimK) i x) + (λ dimK → fibrewiseMap (q dimK) i x) + dimJ dimI +isTracked (isSetDisplayedUFamMap {I} {J} asmI asmJ u X Y f g p q dimI dimJ) = + isProp→SquareP + {B = λ dimI dimJ → + ∃[ e ∈ A ] + ((i : I) (a : A) → + a ⊩[ asmI ] i → + (x : X .carriers i) (b : A) → + b ⊩[ X .assemblies i ] x → + (e ⨾ a ⨾ b) ⊩[ Y .assemblies (u .map i) ] + isSet→isSet' + (Y .assemblies + (u .map i) + .isSetX) + (λ _ → fibrewiseMap f i x) (λ _ → fibrewiseMap g i x) + (λ dimK → fibrewiseMap (p dimK) i x) + (λ dimK → fibrewiseMap (q dimK) i x) dimJ dimI)} + (λ dimI dimJ → isPropPropTrunc) + {a = isTracked f} + {b = isTracked g} + {c = isTracked f} + {d = isTracked g} + refl + refl + (λ dimK → isTracked (p dimK)) + (λ dimK → isTracked (q dimK)) + dimI dimJ -open Categoryᴰ UNIMOD +DisplayedUFamMapPathPIso : + ∀ {I J} (asmI : Assembly I) (asmJ : Assembly J) → + ∀ u v X Y + → (uᴰ : DisplayedUFamMap asmI asmJ u X Y) + → (vᴰ : DisplayedUFamMap asmI asmJ v X Y) + → (p : u ≡ v) + → Iso + (∀ (i : I) (x : X .carriers i) → PathP (λ dimI → Y .carriers (p dimI .map i)) (uᴰ .fibrewiseMap i x) (vᴰ .fibrewiseMap i x)) + (PathP (λ dimI → DisplayedUFamMap asmI asmJ (p dimI) X Y) uᴰ vᴰ) +Iso.fun (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) pᴰ = DisplayedUFamMapPathP asmI asmJ u v X Y uᴰ vᴰ p pᴰ +Iso.inv (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) uᴰ≡vᴰ i x dimI = (uᴰ≡vᴰ dimI) .fibrewiseMap i x +Iso.rightInv (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) uᴰ≡vᴰ dimI dimJ = + isSet→SquareP + {A = λ dimK dimL → DisplayedUFamMap asmI asmJ (p dimL) X Y} + (λ dimI dimJ → isSetDisplayedUFamMap asmI asmJ (p dimJ) X Y) + {a₀₀ = uᴰ} + {a₀₁ = vᴰ} + (λ dimK → DisplayedUFamMapPathP asmI asmJ u v X Y uᴰ vᴰ p (λ i x dimL → uᴰ≡vᴰ dimL .fibrewiseMap i x) dimK) + {a₁₀ = uᴰ} + {a₁₁ = vᴰ} + uᴰ≡vᴰ + refl + refl dimI dimJ +Iso.leftInv (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) pᴰ = refl -UniformFamily : {X : Type ℓ} → (asmX : Assembly X) → Type (ℓ-suc ℓ) -UniformFamily {X} asmX = UNIMOD .Categoryᴰ.ob[_] (X , asmX) +idDisplayedUFamMap : ∀ {I} (asmI : Assembly I) (p : UniformFamily asmI) → DisplayedUFamMap asmI asmI (identityMorphism asmI) p p +DisplayedUFamMap.fibrewiseMap (idDisplayedUFamMap {I} asmI p) i pi = pi +DisplayedUFamMap.isTracked (idDisplayedUFamMap {I} asmI p) = + return + (λ*2 realizer , + λ i a a⊩i x b b⊩x → + subst + (λ r → r ⊩[ p .assemblies i ] x) + (sym (λ*2ComputationRule realizer a b)) + b⊩x) where + realizer : Term as 2 + realizer = # zero + +module _ + {I J K : Type ℓ} + (asmI : Assembly I) + (asmJ : Assembly J) + (asmK : Assembly K) + (f : AssemblyMorphism asmI asmJ) + (g : AssemblyMorphism asmJ asmK) + (X : UniformFamily asmI) + (Y : UniformFamily asmJ) + (Z : UniformFamily asmK) + (fᴰ : DisplayedUFamMap asmI asmJ f X Y) + (gᴰ : DisplayedUFamMap asmJ asmK g Y Z) where + + composeDisplayedUFamMap : DisplayedUFamMap asmI asmK (f ⊚ g) X Z + DisplayedUFamMap.fibrewiseMap composeDisplayedUFamMap i Xi = gᴰ .fibrewiseMap (f .map i) (fᴰ .fibrewiseMap i Xi) + DisplayedUFamMap.isTracked composeDisplayedUFamMap = + do + (gᴰ~ , isTrackedGᴰ) ← gᴰ .isTracked + (fᴰ~ , isTrackedFᴰ) ← fᴰ .isTracked + (f~ , isTrackedF) ← f .tracker + let + realizer : Term as 2 + realizer = ` gᴰ~ ̇ (` f~ ̇ # one) ̇ (` fᴰ~ ̇ # one ̇ # zero) + return + (λ*2 realizer , + (λ i a a⊩i x b b⊩x → + subst + (_⊩[ Z .assemblies (g .map (f .map i)) ] _) + (sym (λ*2ComputationRule realizer a b)) + (isTrackedGᴰ (f .map i) (f~ ⨾ a) (isTrackedF i a a⊩i) (fᴰ .fibrewiseMap i x) (fᴰ~ ⨾ a ⨾ b) (isTrackedFᴰ i a a⊩i x b b⊩x)))) + +UNIMOD : Categoryᴰ ASM (ℓ-suc ℓ) ℓ +Categoryᴰ.ob[ UNIMOD ] (I , asmI) = UniformFamily asmI +Categoryᴰ.Hom[_][_,_] UNIMOD {I , asmI} {J , asmJ} u X Y = DisplayedUFamMap asmI asmJ u X Y +Categoryᴰ.idᴰ UNIMOD {I , asmI} {X} = idDisplayedUFamMap asmI X +Categoryᴰ._⋆ᴰ_ UNIMOD {I , asmI} {J , asmJ} {K , asmK} {f} {g} {X} {Y} {Z} fᴰ gᴰ = composeDisplayedUFamMap asmI asmJ asmK f g X Y Z fᴰ gᴰ +Categoryᴰ.⋆IdLᴰ UNIMOD {I , asmI} {J , asmJ} {f} {X} {Y} fᴰ = + DisplayedUFamMapPathP + asmI asmJ + (identityMorphism asmI ⊚ f) f + X Y + (composeDisplayedUFamMap asmI asmI asmJ (Category.id ASM) f X X Y (idDisplayedUFamMap asmI X) fᴰ) + fᴰ + (Category.⋆IdL ASM f) + (λ i x → refl) +Categoryᴰ.⋆IdRᴰ UNIMOD {I , asmI} {J , asmJ} {f} {X} {Y} fᴰ = + DisplayedUFamMapPathP + asmI asmJ + (f ⊚ identityMorphism asmJ) f + X Y + (composeDisplayedUFamMap asmI asmJ asmJ f (Category.id ASM) X Y Y fᴰ (idDisplayedUFamMap asmJ Y)) + fᴰ + (Category.⋆IdR ASM f) + λ i x → refl +Categoryᴰ.⋆Assocᴰ UNIMOD {I , asmI} {J , asmJ} {K , asmK} {L , asmL} {f} {g} {h} {X} {Y} {Z} {W} fᴰ gᴰ hᴰ = + DisplayedUFamMapPathP + asmI asmL + ((f ⊚ g) ⊚ h) (f ⊚ (g ⊚ h)) + X W + (composeDisplayedUFamMap + asmI asmK asmL + (f ⊚ g) h X Z W + (composeDisplayedUFamMap asmI asmJ asmK f g X Y Z fᴰ gᴰ) + hᴰ) + (composeDisplayedUFamMap + asmI asmJ asmL + f (g ⊚ h) X Y W + fᴰ (composeDisplayedUFamMap asmJ asmK asmL g h Y Z W gᴰ hᴰ)) + (Category.⋆Assoc ASM f g h) + λ i x → refl +Categoryᴰ.isSetHomᴰ UNIMOD {I , asmI} {J , asmJ} {f} {X} {Y} = isSetDisplayedUFamMap asmI asmJ f X Y diff --git a/src/Realizability/PERs/PER.agda b/src/Realizability/PERs/PER.agda index 90ca204..236b9b4 100644 --- a/src/Realizability/PERs/PER.agda +++ b/src/Realizability/PERs/PER.agda @@ -1,6 +1,5 @@ open import Realizability.ApplicativeStructure open import Realizability.CombinatoryAlgebra -open import Realizability.PropResizing open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure using (⟨_⟩; str) @@ -70,8 +69,10 @@ isEquivTracker R S (a , _) (b , _) = (∀ r → r ~[ R ] r → (a ⨾ r) ~[ S ] isEquivRelIsEquivTracker : (R S : PER) → BR.isEquivRel (isEquivTracker R S) BinaryRelation.isEquivRel.reflexive (isEquivRelIsEquivTracker R S) (a , isTrackerA) r r~r = isTrackerA r r r~r -BinaryRelation.isEquivRel.symmetric (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) a~b r r~r = isSymmetric S (a ⨾ r) (b ⨾ r) (a~b r r~r) -BinaryRelation.isEquivRel.transitive (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) (c , isTrackerC) a~b b~c r r~r = isTransitive S (a ⨾ r) (b ⨾ r) (c ⨾ r) (a~b r r~r) (b~c r r~r) +BinaryRelation.isEquivRel.symmetric (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) a~b r r~r = + isSymmetric S (a ⨾ r) (b ⨾ r) (a~b r r~r) +BinaryRelation.isEquivRel.transitive (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) (c , isTrackerC) a~b b~c r r~r = + isTransitive S (a ⨾ r) (b ⨾ r) (c ⨾ r) (a~b r r~r) (b~c r r~r) isPropIsEquivTracker : ∀ R S a b → isProp (isEquivTracker R S a b) isPropIsEquivTracker R S (a , isTrackerA) (b , isTrackerB) = isPropΠ2 λ r r~r → isPropValued S (a ⨾ r) (b ⨾ r) diff --git a/src/Realizability/PERs/ResizedPER.agda b/src/Realizability/PERs/ResizedPER.agda new file mode 100644 index 0000000..4575301 --- /dev/null +++ b/src/Realizability/PERs/ResizedPER.agda @@ -0,0 +1,91 @@ +open import Realizability.ApplicativeStructure +open import Realizability.CombinatoryAlgebra +open import Realizability.PropResizing +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Binary +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Vec +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Functor hiding (Id) + +module Realizability.PERs.ResizedPER + {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) (resizing : hPropResizing ℓ) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.PERs.PER ca + +open CombinatoryAlgebra ca +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) + +smallHProp = resizing .fst +hProp≃smallHProp = resizing .snd +smallHProp≃hProp = invEquiv hProp≃smallHProp + +hPropIsoSmallHProp : Iso (hProp ℓ) smallHProp +hPropIsoSmallHProp = equivToIso hProp≃smallHProp + +shrink : hProp ℓ → smallHProp +shrink = Iso.fun hPropIsoSmallHProp + +enlarge : smallHProp → hProp ℓ +enlarge = Iso.inv hPropIsoSmallHProp + +enlarge⋆shrink≡id : section shrink enlarge +enlarge⋆shrink≡id = Iso.rightInv hPropIsoSmallHProp + +shrink⋆enlarge≡id : retract shrink enlarge +shrink⋆enlarge≡id = Iso.leftInv hPropIsoSmallHProp + +extractType : smallHProp → Type ℓ +extractType p = ⟨ enlarge p ⟩ + +extractFromShrunk : ∀ p isPropP → extractType (shrink (p , isPropP)) ≡ p +extractFromShrunk p isPropP = + extractType (shrink (p , isPropP)) + ≡⟨ refl ⟩ + ⟨ enlarge (shrink (p , isPropP)) ⟩ + ≡⟨ cong ⟨_⟩ (shrink⋆enlarge≡id (p , isPropP)) ⟩ + p + ∎ + +record ResizedPER : Type ℓ where + no-eta-equality + constructor makeResizedPER + field + relation : A → A → smallHProp + isSymmetric : ∀ a b → extractType (relation a b) → extractType (relation b a) + isTransitive : ∀ a b c → extractType (relation a b) → extractType (relation b c) → extractType (relation a c) + +open ResizedPER + +ResizedPERHAEquivPER : HAEquiv ResizedPER PER +PER.relation (fst ResizedPERHAEquivPER resized) = + (λ a b → extractType (resized .relation a b)) , + λ a b → str (enlarge (resized .relation a b)) +fst (PER.isPER (fst ResizedPERHAEquivPER resized)) a b a~b = resized .isSymmetric a b a~b +snd (PER.isPER (fst ResizedPERHAEquivPER resized)) a b c a~b b~c = resized .isTransitive a b c a~b b~c +relation (isHAEquiv.g (snd ResizedPERHAEquivPER) per) a b = + shrink ((per .PER.relation .fst a b) , (per .PER.relation .snd a b)) +isSymmetric (isHAEquiv.g (snd ResizedPERHAEquivPER) per) a b a~b = + subst _ (sym (extractFromShrunk (per .PER.relation .fst b a) (per .PER.relation .snd b a))) {!per .PER.isPER .fst a b a~b!} +isTransitive (isHAEquiv.g (snd ResizedPERHAEquivPER) per) a b c a~b b~c = {!!} +isHAEquiv.linv (snd ResizedPERHAEquivPER) resized = {!!} +isHAEquiv.rinv (snd ResizedPERHAEquivPER) per = {!!} +isHAEquiv.com (snd ResizedPERHAEquivPER) resized = {!!} + +ResizedPER≃PER : ResizedPER ≃ PER +ResizedPER≃PER = ResizedPERHAEquivPER .fst , isHAEquiv→isEquiv (ResizedPERHAEquivPER .snd) diff --git a/src/Realizability/PERs/SystemF.agda b/src/Realizability/PERs/SystemF.agda index 70a186b..5bca5ba 100644 --- a/src/Realizability/PERs/SystemF.agda +++ b/src/Realizability/PERs/SystemF.agda @@ -33,6 +33,12 @@ module Syntax where Ren* : TypeCtxt → TypeCtxt → Type Ren* Γ Δ = ∀ {K} → K ∈* Γ → K ∈* Δ + lift* : ∀ {Γ Δ k} → Ren* Γ Δ → Ren* (Γ , k) (Δ , k) + lift* {Γ} {Δ} {k} ρ {.k} here = here + lift* {Γ} {Δ} {k} ρ {K} (there inm) = there (ρ inm) + + ren* : ∀ {Γ Δ} → + data _∈_ : ∀ {Γ} → Tp Γ tp → TermCtxt Γ → Type where here : ∀ {Γ} {A : Tp Γ tp} {Θ : TermCtxt Γ} → A ∈ (Θ , A) thereType : ∀ {Γ} {A B : Tp Γ tp} {Θ : TermCtxt Γ} → A ∈ Θ → A ∈ (Θ , B) diff --git a/src/Realizability/PERs/UniformFamiliesOverAsm.agda b/src/Realizability/PERs/UniformFamiliesOverAsm.agda new file mode 100644 index 0000000..5b9bf3f --- /dev/null +++ b/src/Realizability/PERs/UniformFamiliesOverAsm.agda @@ -0,0 +1,78 @@ +open import Realizability.ApplicativeStructure +open import Realizability.CombinatoryAlgebra +open import Realizability.PropResizing +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Powerset +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Binary +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base + +module Realizability.PERs.UniformFamiliesOverAsm + {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open CombinatoryAlgebra ca +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open import Realizability.PERs.PER ca +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca + +module _ + {I J : Type ℓ} {asmI : Assembly I} {asmJ : Assembly J} (u : AssemblyMorphism asmI asmJ) (R : I → PER) (S : J → PER) where + + isDisplayedTracker : A → Type _ + isDisplayedTracker a = ∀ (i : I) (b : A) → b ⊩[ asmI ] i → isTracker (R i) (S (u .map i)) (a ⨾ b) + + isPropIsDisplayedTracker : ∀ a → isProp (isDisplayedTracker a) + isPropIsDisplayedTracker a = isPropΠ3 λ i b b⊩i → isPropΠ3 λ r r' r~r' → isProp~ _ (S (u .map i)) _ + + displayedTracker : Type _ + displayedTracker = Σ[ a ∈ A ] isDisplayedTracker a + + isSetDisplayedTracker : isSet displayedTracker + isSetDisplayedTracker = isSetΣ isSetA (λ a → isProp→isSet (isPropIsDisplayedTracker a)) + + isEquivDisplayedTracker : displayedTracker → displayedTracker → Type _ + isEquivDisplayedTracker (f , f⊩f) (g , g⊩g) = ∀ (i : I) (a : A) (a⊩i : a ⊩[ asmI ] i) → isEquivTracker (R i) (S (u .map i)) (f ⨾ a , f⊩f i a a⊩i) (g ⨾ a , g⊩g i a a⊩i) + + displayedPerMorphism : Type _ + displayedPerMorphism = displayedTracker / isEquivDisplayedTracker + +idDisplayedTracker : {I : Type ℓ} → (asmI : Assembly I) → (R : I → PER) → displayedTracker (identityMorphism asmI) R R +idDisplayedTracker {I} asmI R = λ*2 (# zero) , (λ i a a⊩i r r' r~r' → subst2 _~[ R i ]_ (sym (λ*2ComputationRule (# zero) a r)) (sym (λ*2ComputationRule (# zero) a r')) r~r' ) + +open Category ASM +module _ + {I J K : Type ℓ} + {asmI : Assembly I} + {asmJ : Assembly J} + {asmK : Assembly K} + (f : AssemblyMorphism asmI asmJ) + (g : AssemblyMorphism asmJ asmK) + (R : I → PER) + (S : J → PER) + (T : K → PER) + (fᴰ : displayedPerMorphism f R S) + (gᴰ : displayedPerMorphism g S T) where + +UFAMPER : Categoryᴰ ASM (ℓ-suc ℓ) ℓ +Categoryᴰ.ob[ UFAMPER ] (X , asmX) = X → PER +Categoryᴰ.Hom[_][_,_] UFAMPER {I , asmI} {J , asmJ} u R S = displayedPerMorphism u R S +Categoryᴰ.idᴰ UFAMPER {I , asmI} {R} = [ idDisplayedTracker asmI R ] +Categoryᴰ._⋆ᴰ_ UFAMPER {I , asmI} {J , asmJ} {K , asmK} {f} {g} {R} {S} {T} fᴰ gᴰ = {!!} +Categoryᴰ.⋆IdLᴰ UFAMPER = {!!} +Categoryᴰ.⋆IdRᴰ UFAMPER = {!!} +Categoryᴰ.⋆Assocᴰ UFAMPER = {!!} +Categoryᴰ.isSetHomᴰ UFAMPER = {!!} From 1c1ecc390e7ce0f4b1135e7b2e3cd6d8597c8b45 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 27 Jun 2024 20:26:42 +0530 Subject: [PATCH 53/61] Define Small PERs and Equivalence to ordinary PERs --- src/Realizability/PERs/PER.agda | 64 +++++++++++++-- src/Realizability/PERs/ResizedPER.agda | 109 +++++++++++++++++++++---- 2 files changed, 147 insertions(+), 26 deletions(-) diff --git a/src/Realizability/PERs/PER.agda b/src/Realizability/PERs/PER.agda index 236b9b4..923f455 100644 --- a/src/Realizability/PERs/PER.agda +++ b/src/Realizability/PERs/PER.agda @@ -30,11 +30,11 @@ open Combinators ca renaming (i to Id; ia≡a to Ida≡a) module BR = BinaryRelation -isPartialEquivalenceRelation : PropRel A A ℓ → Type _ -isPartialEquivalenceRelation (rel , isPropValuedRel) = BR.isSym rel × BR.isTrans rel +isPartialEquivalenceRelation : (A → A → Type ℓ) → Type _ +isPartialEquivalenceRelation rel = BR.isSym rel × BR.isTrans rel -isPropIsPartialEquivalenceRelation : ∀ r → isProp (isPartialEquivalenceRelation r) -isPropIsPartialEquivalenceRelation (rel , isPropValuedRel) = +isPropIsPartialEquivalenceRelation : ∀ r → (∀ a b → isProp (r a b)) → isProp (isPartialEquivalenceRelation r) +isPropIsPartialEquivalenceRelation rel isPropValuedRel = isProp× (isPropΠ (λ x → isPropΠ λ y → isProp→ (isPropValuedRel y x))) (isPropΠ λ x → isPropΠ λ y → isPropΠ λ z → isProp→ (isProp→ (isPropValuedRel x z))) @@ -43,20 +43,66 @@ record PER : Type (ℓ-suc ℓ) where no-eta-equality constructor makePER field - relation : PropRel A A ℓ + relation : A → A → Type ℓ + isPropValued : ∀ a b → isProp (relation a b) isPER : isPartialEquivalenceRelation relation - ∣_∣ = relation .fst isSymmetric = isPER .fst isTransitive = isPER .snd - isPropValued = relation .snd open PER +PERΣ : Type (ℓ-suc ℓ) +PERΣ = Σ[ relation ∈ (A → A → hProp ℓ) ] isPartialEquivalenceRelation λ a b → ⟨ relation a b ⟩ + +isSetPERΣ : isSet PERΣ +isSetPERΣ = + isSetΣ + (isSet→ (isSet→ isSetHProp)) + (λ relation → + isProp→isSet + (isPropIsPartialEquivalenceRelation + (λ a b → ⟨ relation a b ⟩) + (λ a b → str (relation a b)))) + +PER≡ : ∀ (R S : PER) → (R .relation ≡ S .relation) → R ≡ S +relation (PER≡ R S rel≡ i) = rel≡ i +isPropValued (PER≡ R S rel≡ i) a b = + isProp→PathP + {B = λ j → isProp (rel≡ j a b)} + (λ j → isPropIsProp) + (R .isPropValued a b) + (S .isPropValued a b) i +isPER (PER≡ R S rel≡ i) = + isProp→PathP + {B = λ j → isPartialEquivalenceRelation (rel≡ j)} + (λ j → isPropIsPartialEquivalenceRelation (rel≡ j) λ a b → isPropRelJ a b j) + (R .isPER) + (S .isPER) i where + isPropRelJ : ∀ a b j → isProp (rel≡ j a b) + isPropRelJ a b j = isProp→PathP {B = λ k → isProp (rel≡ k a b)} (λ k → isPropIsProp) (R .isPropValued a b) (S .isPropValued a b) j + +PERIsoΣ : Iso PER PERΣ +Iso.fun PERIsoΣ per = (λ a b → per .relation a b , per .isPropValued a b) , per .isPER +relation (Iso.inv PERIsoΣ perΣ) a b = ⟨ perΣ .fst a b ⟩ +isPropValued (Iso.inv PERIsoΣ perΣ) a b = str (perΣ .fst a b) +isPER (Iso.inv PERIsoΣ perΣ) = perΣ .snd +Iso.rightInv PERIsoΣ perΣ = refl +Iso.leftInv PERIsoΣ per = PER≡ _ _ refl + +isSetPER : isSet PER +isSetPER = isOfHLevelRetractFromIso 2 PERIsoΣ isSetPERΣ + +PER≡Iso : ∀ (R S : PER) → Iso (R ≡ S) (R .relation ≡ S .relation) +Iso.fun (PER≡Iso R S) R≡S i = R≡S i .relation +Iso.inv (PER≡Iso R S) rel≡ = PER≡ R S rel≡ +Iso.rightInv (PER≡Iso R S) rel≡ = refl +Iso.leftInv (PER≡Iso R S) R≡S = isSetPER R S _ _ + _~[_]_ : A → PER → A → Type ℓ -a ~[ R ] b = R .relation .fst a b +a ~[ R ] b = R .relation a b isProp~ : ∀ a R b → isProp (a ~[ R ] b) -isProp~ a R b = R .relation .snd a b +isProp~ a R b = R .isPropValued a b isTracker : (R S : PER) → A → Type ℓ isTracker R S a = ∀ r r' → r ~[ R ] r' → (a ⨾ r) ~[ S ] (a ⨾ r') diff --git a/src/Realizability/PERs/ResizedPER.agda b/src/Realizability/PERs/ResizedPER.agda index 4575301..c8b5b16 100644 --- a/src/Realizability/PERs/ResizedPER.agda +++ b/src/Realizability/PERs/ResizedPER.agda @@ -8,7 +8,7 @@ open import Cubical.Foundations.Equiv open import Cubical.Foundations.Univalence open import Cubical.Foundations.Powerset open import Cubical.Foundations.HLevels -open import Cubical.Foundations.Equiv.HalfAdjoint +open import Cubical.Foundations.Path open import Cubical.Functions.FunExtEquiv open import Cubical.Relation.Binary open import Cubical.Data.Sigma @@ -35,6 +35,9 @@ smallHProp = resizing .fst hProp≃smallHProp = resizing .snd smallHProp≃hProp = invEquiv hProp≃smallHProp +isSetSmallHProp : isSet smallHProp +isSetSmallHProp = isOfHLevelRespectEquiv 2 hProp≃smallHProp isSetHProp + hPropIsoSmallHProp : Iso (hProp ℓ) smallHProp hPropIsoSmallHProp = equivToIso hProp≃smallHProp @@ -53,6 +56,9 @@ shrink⋆enlarge≡id = Iso.leftInv hPropIsoSmallHProp extractType : smallHProp → Type ℓ extractType p = ⟨ enlarge p ⟩ +isPropExtractType : ∀ p → isProp (extractType p) +isPropExtractType p = str (enlarge p) + extractFromShrunk : ∀ p isPropP → extractType (shrink (p , isPropP)) ≡ p extractFromShrunk p isPropP = extractType (shrink (p , isPropP)) @@ -62,6 +68,15 @@ extractFromShrunk p isPropP = p ∎ +shrinkFromExtracted : ∀ p → shrink (extractType p , isPropExtractType p) ≡ p +shrinkFromExtracted p = + shrink (extractType p , isPropExtractType p) + ≡⟨ refl ⟩ + shrink (enlarge p) + ≡⟨ enlarge⋆shrink≡id p ⟩ + p + ∎ + record ResizedPER : Type ℓ where no-eta-equality constructor makeResizedPER @@ -69,23 +84,83 @@ record ResizedPER : Type ℓ where relation : A → A → smallHProp isSymmetric : ∀ a b → extractType (relation a b) → extractType (relation b a) isTransitive : ∀ a b c → extractType (relation a b) → extractType (relation b c) → extractType (relation a c) - + open ResizedPER -ResizedPERHAEquivPER : HAEquiv ResizedPER PER -PER.relation (fst ResizedPERHAEquivPER resized) = - (λ a b → extractType (resized .relation a b)) , - λ a b → str (enlarge (resized .relation a b)) -fst (PER.isPER (fst ResizedPERHAEquivPER resized)) a b a~b = resized .isSymmetric a b a~b -snd (PER.isPER (fst ResizedPERHAEquivPER resized)) a b c a~b b~c = resized .isTransitive a b c a~b b~c -relation (isHAEquiv.g (snd ResizedPERHAEquivPER) per) a b = - shrink ((per .PER.relation .fst a b) , (per .PER.relation .snd a b)) -isSymmetric (isHAEquiv.g (snd ResizedPERHAEquivPER) per) a b a~b = - subst _ (sym (extractFromShrunk (per .PER.relation .fst b a) (per .PER.relation .snd b a))) {!per .PER.isPER .fst a b a~b!} -isTransitive (isHAEquiv.g (snd ResizedPERHAEquivPER) per) a b c a~b b~c = {!!} -isHAEquiv.linv (snd ResizedPERHAEquivPER) resized = {!!} -isHAEquiv.rinv (snd ResizedPERHAEquivPER) per = {!!} -isHAEquiv.com (snd ResizedPERHAEquivPER) resized = {!!} +unquoteDecl ResizedPERIsoΣ = declareRecordIsoΣ ResizedPERIsoΣ (quote ResizedPER) + +ResizedPERΣ : Type ℓ +ResizedPERΣ = + Σ[ relation ∈ (A → A → smallHProp) ] + (∀ a b → extractType (relation a b) → extractType (relation b a)) × + (∀ a b c → extractType (relation a b) → extractType (relation b c) → extractType (relation a c)) + +isSetResizedPERΣ : isSet ResizedPERΣ +isSetResizedPERΣ = + isSetΣ + (isSet→ (isSet→ isSetSmallHProp)) + (λ relation → isProp→isSet (isProp× (isPropΠ3 λ _ _ _ → isPropExtractType _) (isPropΠ5 λ _ _ _ _ _ → isPropExtractType _))) + +isSetResizedPER : isSet ResizedPER +isSetResizedPER = isOfHLevelRetractFromIso 2 ResizedPERIsoΣ isSetResizedPERΣ + +ResizedPER≡Iso : ∀ (R S : ResizedPER) → Iso (R ≡ S) (∀ a b → R .relation a b ≡ S .relation a b) +Iso.fun (ResizedPER≡Iso R S) R≡S a b i = (R≡S i) .relation a b +relation (Iso.inv (ResizedPER≡Iso R S) pointwise i) a b = pointwise a b i +isSymmetric (Iso.inv (ResizedPER≡Iso R S) pointwise i) = + isProp→PathP + {B = λ j → (a b : A) → extractType (pointwise a b j) → extractType (pointwise b a j)} + (λ j → isPropΠ3 λ _ _ _ → isPropExtractType _) + (isSymmetric R) + (isSymmetric S) i +isTransitive (Iso.inv (ResizedPER≡Iso R S) pointwise i) = + isProp→PathP + {B = λ j → (a b c : A) → extractType (pointwise a b j) → extractType (pointwise b c j) → extractType (pointwise a c j)} + (λ j → isPropΠ5 λ _ _ _ _ _ → isPropExtractType _) + (R .isTransitive) + (S .isTransitive) + i +Iso.rightInv (ResizedPER≡Iso R S) pointwise = refl +Iso.leftInv (ResizedPER≡Iso R S) R≡S = isSetResizedPER R S _ _ + +ResizedPER≡ : ∀ (R S : ResizedPER) → (∀ a b → R .relation a b ≡ S .relation a b) → R ≡ S +ResizedPER≡ R S pointwise = Iso.inv (ResizedPER≡Iso R S) pointwise + +ResizedPERIsoPER : Iso ResizedPER PER +PER.relation (Iso.fun ResizedPERIsoPER resized) a b = extractType (resized .relation a b) +PER.isPropValued (Iso.fun ResizedPERIsoPER resized) a b = isPropExtractType _ +fst (PER.isPER (Iso.fun ResizedPERIsoPER resized)) a b a~b = resized .isSymmetric a b a~b +snd (PER.isPER (Iso.fun ResizedPERIsoPER resized)) a b c a~b b~c = resized .isTransitive a b c a~b b~c +relation (Iso.inv ResizedPERIsoPER per) a b = shrink (per .PER.relation a b , per .PER.isPropValued a b) +isSymmetric (Iso.inv ResizedPERIsoPER per) a b a~[resized]b = b~[resized]a where + a~b : per .PER.relation a b + a~b = transport (extractFromShrunk _ _) a~[resized]b + + b~a : per .PER.relation b a + b~a = per .PER.isPER .fst a b a~b + + b~[resized]a : extractType (shrink (per .PER.relation b a , per .PER.isPropValued b a)) + b~[resized]a = transport (sym (extractFromShrunk _ _)) b~a +isTransitive (Iso.inv ResizedPERIsoPER per) a b c a~[resized]b b~[resized]c = a~[resized]c where + a~b : per .PER.relation a b + a~b = transport (extractFromShrunk _ _) a~[resized]b + + b~c : per .PER.relation b c + b~c = transport (extractFromShrunk _ _) b~[resized]c + + a~c : per .PER.relation a c + a~c = per .PER.isPER .snd a b c a~b b~c + + a~[resized]c : extractType (shrink (per .PER.relation a c , per .PER.isPropValued a c)) + a~[resized]c = transport (sym (extractFromShrunk _ _)) a~c +Iso.rightInv ResizedPERIsoPER per = + PER≡ _ _ + (funExt₂ + λ a b → + extractFromShrunk (per .PER.relation a b) (per .PER.isPropValued a b)) +Iso.leftInv ResizedPERIsoPER resizedPer = + ResizedPER≡ _ _ + λ a b → shrinkFromExtracted (resizedPer .relation a b) ResizedPER≃PER : ResizedPER ≃ PER -ResizedPER≃PER = ResizedPERHAEquivPER .fst , isHAEquiv→isEquiv (ResizedPERHAEquivPER .snd) +ResizedPER≃PER = isoToEquiv ResizedPERIsoPER From 2d33f50d6f9690618ee6b52b274af1f6fb43323a Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 4 Jul 2024 23:03:33 +0530 Subject: [PATCH 54/61] Subquotient modest set and canonical PER of a modest set --- src/Categories/CartesianMorphism.agda | 102 +++++++++++ src/Categories/FamiliesFibration.agda | 108 ++++++++++++ src/Categories/GenericObject.agda | 28 +++ src/Realizability/Modest/CanonicalPER.agda | 66 +++++++ .../Modest/GenericUniformFamily.agda | 164 ++++++++++++++++++ src/Realizability/PERs/ResizedPER.agda | 29 ++++ src/Realizability/PERs/SubQuotient.agda | 82 +++++++++ .../PERs/SubQuotientOfCanonicalPER.agda | 72 ++++++++ 8 files changed, 651 insertions(+) create mode 100644 src/Categories/CartesianMorphism.agda create mode 100644 src/Categories/FamiliesFibration.agda create mode 100644 src/Categories/GenericObject.agda create mode 100644 src/Realizability/Modest/CanonicalPER.agda create mode 100644 src/Realizability/Modest/GenericUniformFamily.agda create mode 100644 src/Realizability/PERs/SubQuotient.agda create mode 100644 src/Realizability/PERs/SubQuotientOfCanonicalPER.agda diff --git a/src/Categories/CartesianMorphism.agda b/src/Categories/CartesianMorphism.agda new file mode 100644 index 0000000..b7d9188 --- /dev/null +++ b/src/Categories/CartesianMorphism.agda @@ -0,0 +1,102 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation + +module Categories.CartesianMorphism where + +module Contravariant + {ℓB ℓ'B ℓE ℓ'E} + {B : Category ℓB ℓ'B} + (E : Categoryᴰ B ℓE ℓ'E) where + + open Category B + open Categoryᴰ E + + opaque + isCartesian : + {a b : ob} (f : B [ a , b ]) + {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} + (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) → + Type (ℓ-max (ℓ-max ℓB ℓ'B) (ℓ-max ℓE ℓ'E)) + isCartesian {a} {b} f {aᴰ} {bᴰ} fᴰ = + ∀ {c : ob} {cᴰ : ob[ c ]} (g : B [ c , a ]) → isEquiv λ (gᴰ : Hom[ g ][ cᴰ , aᴰ ]) → gᴰ ⋆ᴰ fᴰ + + opaque + unfolding isCartesian + isPropIsCartesian : + {a b : ob} (f : B [ a , b ]) + {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} + (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) → + isProp (isCartesian f fᴰ) + isPropIsCartesian f fᴰ = isPropImplicitΠ2 λ c cᴰ → isPropΠ λ g → isPropIsEquiv (_⋆ᴰ fᴰ) + + opaque + isCartesian' : + {a b : ob} (f : B [ a , b ]) + {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} + (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) → + Type (ℓ-max (ℓ-max ℓB ℓ'B) (ℓ-max ℓE ℓ'E)) + isCartesian' {a} {b} f {aᴰ} {bᴰ} fᴰ = + ∀ {c : ob} {cᴰ : ob[ c ]} (g : B [ c , a ]) → + (∀ (hᴰ : Hom[ g ⋆ f ][ cᴰ , bᴰ ]) → ∃![ gᴰ ∈ Hom[ g ][ cᴰ , aᴰ ] ] gᴰ ⋆ᴰ fᴰ ≡ hᴰ) + + opaque + unfolding isCartesian' + isPropIsCartesian' : + {a b : ob} {f : B [ a , b ]} + {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} + (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) → + isProp (isCartesian' f fᴰ) + isPropIsCartesian' {a} {b} {f} {aᴰ} {bᴰ} fᴰ = + isPropImplicitΠ2 λ c cᴰ → isPropΠ2 λ g hᴰ → isPropIsContr + + opaque + unfolding isCartesian + unfolding isCartesian' + isCartesian→isCartesian' : + {a b : ob} (f : B [ a , b ]) + {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} + (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) → + isCartesian f fᴰ → + isCartesian' f fᴰ + isCartesian→isCartesian' {a} {b} f {aᴰ} {bᴰ} fᴰ cartfᴰ g hᴰ = + ((invIsEq (cartfᴰ g) hᴰ) , (secIsEq (cartfᴰ g) hᴰ)) , + (λ { (gᴰ , gᴰ⋆fᴰ≡hᴰ) → cartfᴰ g .equiv-proof hᴰ .snd (gᴰ , gᴰ⋆fᴰ≡hᴰ) }) + + opaque + unfolding isCartesian + unfolding isCartesian' + isCartesian'→isCartesian : + {a b : ob} (f : B [ a , b ]) + {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} + (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) → + isCartesian' f fᴰ → + isCartesian f fᴰ + equiv-proof (isCartesian'→isCartesian {a} {b} f {aᴰ} {bᴰ} fᴰ cart'fᴰ g) hᴰ = (cart'fᴰ g hᴰ .fst) , (cart'fᴰ g hᴰ .snd) + + isCartesian≃isCartesian' : + {a b : ob} (f : B [ a , b ]) + {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} + (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) → + isCartesian f fᴰ ≃ isCartesian' f fᴰ + isCartesian≃isCartesian' {a} {b} f {aᴰ} {bᴰ} fᴰ = + propBiimpl→Equiv (isPropIsCartesian f fᴰ) (isPropIsCartesian' fᴰ) (isCartesian→isCartesian' f fᴰ) (isCartesian'→isCartesian f fᴰ) + + cartesianLift : {a b : ob} (f : B [ a , b ]) (bᴰ : ob[ b ]) → Type _ + cartesianLift {a} {b} f bᴰ = Σ[ aᴰ ∈ ob[ a ] ] Σ[ fᴰ ∈ Hom[ f ][ aᴰ , bᴰ ] ] isCartesian f fᴰ + + isCartesianFibration : Type _ + isCartesianFibration = + ∀ {a b : ob} {bᴰ : ob[ b ]} + → (f : Hom[ a , b ]) + → ∥ cartesianLift f bᴰ ∥₁ + + isPropIsCartesianFibration : isProp isCartesianFibration + isPropIsCartesianFibration = isPropImplicitΠ3 λ a b bᴰ → isPropΠ λ f → isPropPropTrunc + + cleavage : Type _ + cleavage = {a b : ob} (f : Hom[ a , b ]) (bᴰ : ob[ b ]) → cartesianLift f bᴰ diff --git a/src/Categories/FamiliesFibration.agda b/src/Categories/FamiliesFibration.agda new file mode 100644 index 0000000..8b2bfd1 --- /dev/null +++ b/src/Categories/FamiliesFibration.agda @@ -0,0 +1,108 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Instances.Sets +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation +open import Categories.CartesianMorphism +open import Categories.GenericObject +module Categories.FamiliesFibration where + +module FamiliesFibration + {ℓ ℓ'} + (C : Category ℓ ℓ') + (ℓ'' : Level) where + open Category C + Families : Categoryᴰ (SET ℓ'') (ℓ-max ℓ ℓ'') (ℓ-max ℓ' ℓ'') + Categoryᴰ.ob[ Families ] (I , isSetI) = I → ob + Categoryᴰ.Hom[_][_,_] Families {I , isSetI} {J , isSetJ} u X Y = (i : I) → Hom[ X i , Y (u i) ] + Categoryᴰ.idᴰ Families {I , isSetI} {X} i = id + Categoryᴰ._⋆ᴰ_ Families {I , isSetI} {J , isSetJ} {K , isSetK} {f} {g} {X} {Y} {Z} fᴰ gᴰ i = + fᴰ i ⋆ gᴰ (f i) + Categoryᴰ.⋆IdLᴰ Families {I , isSetI} {J , isSetJ} {f} {X} {Y} fᴰ = + funExt λ i → ⋆IdL (fᴰ i) + Categoryᴰ.⋆IdRᴰ Families {I , isSetI} {J , isSetJ} {f} {X} {Y} fᴰ = + funExt λ i → ⋆IdR (fᴰ i) + Categoryᴰ.⋆Assocᴰ Families {I , isSetI} {J , isSetJ} {K , isSetK} {L , isSetL} {f} {g} {h} {X} {Y} {Z} {W} fᴰ gᴰ hᴰ = + funExt λ i → ⋆Assoc (fᴰ i) (gᴰ (f i)) (hᴰ (g (f i))) + Categoryᴰ.isSetHomᴰ Families {I , isSetI} {J , isSetJ} {f} {X} {Y} = isSetΠ λ i → isSetHom + + open Contravariant Families + open Categoryᴰ Families + cleavageFamilies : cleavage + cleavageFamilies a@{I , isSetI} b@{J , isSetJ} f Y = + X , + fᴰ , + cart where + + X : I → ob + X i = Y (f i) + + fᴰ : (i : I) → Hom[ X i , Y (f i) ] + fᴰ i = id + + opaque + unfolding isCartesian' + cart' : isCartesian' {a = a} {b = b} f {bᴰ = Y} fᴰ + cart' k@{K , isSetK} {Z} g hᴰ = + (hᴰ , (funExt (λ k → ⋆IdR (hᴰ k)))) , + λ { (! , !comm) → + Σ≡Prop + (λ ! → isSetHomᴰ {x = k} {y = b} {xᴰ = Z} {yᴰ = Y} (λ k → ! k ⋆ fᴰ (g k)) hᴰ) + (funExt + (λ k → + sym + (! k + ≡⟨ sym (⋆IdR (! k)) ⟩ + (! k ⋆ fᴰ (g k)) + ≡[ i ]⟨ !comm i k ⟩ + hᴰ k + ∎))) } + + cart : isCartesian {a = a} {b = b} f {bᴰ = Y} fᴰ + cart = isCartesian'→isCartesian {a = a} {b = b} f {bᴰ = Y} fᴰ cart' + +module GenericFamily + {ℓ ℓ'} + (C : Category ℓ ℓ') + (isSetCOb : isSet (C .Category.ob)) where + + open FamiliesFibration C ℓ + open Category C + open Categoryᴰ Families + open Contravariant Families + genericFamily : GenericObject Families + GenericObject.base genericFamily = ob , isSetCOb + GenericObject.displayed genericFamily = λ x → x + GenericObject.isGeneric genericFamily i@{I , isSetI} X = + f , + fᴰ , + cart where + f : I → ob + f = X + + fᴰ : Hom[_][_,_] {x = i} {y = ob , isSetCOb} f X (λ x → x) + fᴰ i = id + + opaque + unfolding isCartesian' + cart' : isCartesian' {a = i} {b = ob , isSetCOb} f {bᴰ = λ x → x} fᴰ + cart' {J , isSetJ} {Z} g hᴰ = + (hᴰ , funExt λ j → ⋆IdR (hᴰ j)) , + λ { (! , !comm) → + Σ≡Prop + (λ ! → isSetHomᴰ {x = J , isSetJ} {y = i} {f = g} {xᴰ = Z} {yᴰ = f} (λ j → ! j ⋆ fᴰ (g j)) hᴰ) + (funExt + λ j → + sym + (! j + ≡⟨ sym (⋆IdR (! j)) ⟩ + ! j ⋆ fᴰ (g j) + ≡[ i ]⟨ !comm i j ⟩ + hᴰ j + ∎)) } + + cart : isCartesian {a = i} {b = ob , isSetCOb} f {bᴰ = λ x → x} fᴰ + cart = isCartesian'→isCartesian f fᴰ cart' diff --git a/src/Categories/GenericObject.agda b/src/Categories/GenericObject.agda new file mode 100644 index 0000000..0412bb1 --- /dev/null +++ b/src/Categories/GenericObject.agda @@ -0,0 +1,28 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.HLevels +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation +open import Categories.CartesianMorphism + +module Categories.GenericObject where + +module _ + {ℓB ℓ'B ℓE ℓ'E} + {B : Category ℓB ℓ'B} + (E : Categoryᴰ B ℓE ℓ'E) where + + open Category B + open Categoryᴰ E + open Contravariant E + + record GenericObject : Type (ℓ-max (ℓ-max ℓB ℓ'B) (ℓ-max ℓE ℓ'E)) where + constructor makeGenericObject + field + base : ob + displayed : ob[ base ] + isGeneric : + ∀ {t : ob} (tᴰ : ob[ t ]) + → Σ[ f ∈ Hom[ t , base ] ] Σ[ fᴰ ∈ Hom[ f ][ tᴰ , displayed ] ] isCartesian f fᴰ diff --git a/src/Realizability/Modest/CanonicalPER.agda b/src/Realizability/Modest/CanonicalPER.agda new file mode 100644 index 0000000..30ade69 --- /dev/null +++ b/src/Realizability/Modest/CanonicalPER.agda @@ -0,0 +1,66 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.CanonicalPER {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +module _ + {X : Type ℓ} + (asmX : Assembly X) + (isModestAsmX : isModest asmX) where + + canonicalPER : PER + PER.relation canonicalPER a b = ∃[ x ∈ X ] a ⊩[ asmX ] x × b ⊩[ asmX ] x + PER.isPropValued canonicalPER a b = isPropPropTrunc + fst (PER.isPER canonicalPER) a b ∃x = PT.map (λ { (x , a⊩x , b⊩x) → x , b⊩x , a⊩x }) ∃x + snd (PER.isPER canonicalPER) a b c ∃x ∃x' = + do + (x , a⊩x , b⊩x) ← ∃x + (x' , b⊩x' , c⊩x') ← ∃x' + let + x≡x' : x ≡ x' + x≡x' = isModestAsmX x x' b b⊩x b⊩x' + + c⊩x : c ⊩[ asmX ] x + c⊩x = subst (c ⊩[ asmX ]_) (sym x≡x') c⊩x' + return (x , a⊩x , c⊩x) + + diff --git a/src/Realizability/Modest/GenericUniformFamily.agda b/src/Realizability/Modest/GenericUniformFamily.agda new file mode 100644 index 0000000..50388b7 --- /dev/null +++ b/src/Realizability/Modest/GenericUniformFamily.agda @@ -0,0 +1,164 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.GenericUniformFamily {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) (resizing : hPropResizing ℓ) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.ResizedPER ca resizing +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +uniformFamilyCleavage : cleavage +uniformFamilyCleavage {X , asmX} {Y , asmY} f N = + N' , fᴰ , cartfᴰ where + N' : UniformFamily asmX + UniformFamily.carriers N' x = N .carriers (f .map x) + UniformFamily.assemblies N' x = N .assemblies (f .map x) + UniformFamily.isModestFamily N' x = N .isModestFamily (f .map x) + + fᴰ : DisplayedUFamMap asmX asmY f N' N + DisplayedUFamMap.fibrewiseMap fᴰ x Nfx = Nfx + DisplayedUFamMap.isTracked fᴰ = + do + let + realizer : Term as 2 + realizer = # zero + return + (λ*2 realizer , + (λ x a a⊩x Nfx b b⊩Nfx → + subst + (_⊩[ N .assemblies (f .map x) ] Nfx) + (sym (λ*2ComputationRule realizer a b)) + b⊩Nfx)) + + opaque + unfolding isCartesian' + cart'fᴰ : isCartesian' f fᴰ + cart'fᴰ {Z , asmZ} {M} g hᴰ = + (! , !⋆fᴰ≡hᴰ) , + λ { (!' , !'comm) → + Σ≡Prop + (λ ! → UNIMOD .Categoryᴰ.isSetHomᴰ _ _) + (DisplayedUFamMapPathP + _ _ _ _ _ _ _ _ _ + λ z Mz → + sym + (!' .fibrewiseMap z Mz + ≡[ i ]⟨ !'comm i .fibrewiseMap z Mz ⟩ + hᴰ .fibrewiseMap z Mz + ∎)) } where + ! : DisplayedUFamMap asmZ asmX g M N' + DisplayedUFamMap.fibrewiseMap ! z Mz = hᴰ .fibrewiseMap z Mz + DisplayedUFamMap.isTracked ! = hᴰ .isTracked + + !⋆fᴰ≡hᴰ : composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ ≡ hᴰ + !⋆fᴰ≡hᴰ = + DisplayedUFamMapPathP + asmZ asmY _ _ + M N + (composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ) hᴰ refl + λ z Mz → refl + + cartfᴰ : isCartesian f fᴰ + cartfᴰ = isCartesian'→isCartesian f fᴰ cart'fᴰ + +∇PER = ∇ ⟅ ResizedPER , isSetResizedPER ⟆ +asm∇PER = ∇PER .snd + +genericUniformFamily : GenericObject UNIMOD +GenericObject.base genericUniformFamily = ∇PER +UniformFamily.carriers (GenericObject.displayed genericUniformFamily) per = subQuotient (enlargePER per) +UniformFamily.assemblies (GenericObject.displayed genericUniformFamily) per = subQuotientAssembly (enlargePER per) +UniformFamily.isModestFamily (GenericObject.displayed genericUniformFamily) per = isModestSubQuotientAssembly (enlargePER per) +GenericObject.isGeneric genericUniformFamily {X , asmX} M = + f , fᴰ , {!!} where + + f : AssemblyMorphism asmX asm∇PER + AssemblyMorphism.map f x = shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x)) + AssemblyMorphism.tracker f = return (k , (λ _ _ _ → tt*)) + + subQuotientTargetType : (x : X) → M .carriers x → Type ℓ + subQuotientTargetType x Mx = subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + + fᴰ : DisplayedUFamMap asmX asm∇PER f M (genericUniformFamily .GenericObject.displayed) + DisplayedUFamMap.fibrewiseMap fᴰ x Mx = + subst (λ x → subQuotient x) (sym equation) target module CartesianMapDefinition where + targetType : Type ℓ + targetType = subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + + opaque + equation : enlargePER (shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x))) ≡ canonicalPER (M .assemblies x) (M .isModestFamily x) + equation = enlargePER⋆shrinkPER≡id (canonicalPER (M .assemblies x) (M .isModestFamily x)) + + mainMap : Σ[ a ∈ A ] (a ⊩[ M .assemblies x ] Mx) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + mainMap (a , a⊩Mx) = [ a ] + + mainMap2Constant : 2-Constant mainMap + mainMap2Constant (a , a⊩Mx) (b , b⊩Mx) = eq/ a b (return (Mx , a⊩Mx , b⊩Mx)) + + opaque + target : targetType + target = + PT.rec→Set + squash/ + mainMap + mainMap2Constant + (M .assemblies x .⊩surjective Mx) + + opaque + isTrackedCartesianMap : + ∀ (a : A) + → a ⊩[ asmX ] x + → (b : A) + → b ⊩[ M .assemblies x ] Mx + → ⟨ + subQuotientRealizability + (enlargePER (shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x)))) + (λ*2 (# one) ⨾ a ⨾ b) + (subst subQuotient (sym equation) target) + ⟩ + isTrackedCartesianMap a a⊩x b b⊩Mx = + {!!} + DisplayedUFamMap.isTracked fᴰ = + do + return + (λ*2 (# one) , + λ x a a⊩x Mx b b⊩Mx → {!CartesianMapDefinition.isTrackedCartesianMap x Mx a a⊩x b b⊩Mx!}) + diff --git a/src/Realizability/PERs/ResizedPER.agda b/src/Realizability/PERs/ResizedPER.agda index c8b5b16..dc43ef1 100644 --- a/src/Realizability/PERs/ResizedPER.agda +++ b/src/Realizability/PERs/ResizedPER.agda @@ -27,6 +27,7 @@ module Realizability.PERs.ResizedPER open import Realizability.Assembly.Base ca open import Realizability.Assembly.Morphism ca open import Realizability.PERs.PER ca +open import Realizability.Modest.Base ca open CombinatoryAlgebra ca open Combinators ca renaming (i to Id; ia≡a to Ida≡a) @@ -162,5 +163,33 @@ Iso.leftInv ResizedPERIsoPER resizedPer = ResizedPER≡ _ _ λ a b → shrinkFromExtracted (resizedPer .relation a b) +opaque + shrinkPER : PER → ResizedPER + shrinkPER = ResizedPERIsoPER .Iso.inv + +opaque + enlargePER : ResizedPER → PER + enlargePER = ResizedPERIsoPER .Iso.fun + +opaque + unfolding shrinkPER + unfolding enlargePER + shrinkPER⋆enlargePER≡id : ∀ resized → shrinkPER (enlargePER resized) ≡ resized + shrinkPER⋆enlargePER≡id resized = ResizedPERIsoPER .Iso.leftInv resized + +opaque + unfolding shrinkPER + unfolding enlargePER + enlargePER⋆shrinkPER≡id : ∀ per → enlargePER (shrinkPER per) ≡ per + enlargePER⋆shrinkPER≡id per = ResizedPERIsoPER .Iso.rightInv per + ResizedPER≃PER : ResizedPER ≃ PER ResizedPER≃PER = isoToEquiv ResizedPERIsoPER + +opaque + transportFromSmall : ∀ {ℓ'} {P : ResizedPER → Type ℓ'} → (∀ per → P (shrinkPER per)) → ∀ resized → P resized + transportFromSmall {ℓ'} {P} small resized = subst P (shrinkPER⋆enlargePER≡id resized) (small (enlargePER resized)) + +opaque + transportFromLarge : ∀ {ℓ'} {P : PER → Type ℓ'} → (∀ resized → P (enlargePER resized)) → ∀ per → P per + transportFromLarge {ℓ'} {P} large per = subst P (enlargePER⋆shrinkPER≡id per) (large (shrinkPER per)) diff --git a/src/Realizability/PERs/SubQuotient.agda b/src/Realizability/PERs/SubQuotient.agda new file mode 100644 index 0000000..e28420a --- /dev/null +++ b/src/Realizability/PERs/SubQuotient.agda @@ -0,0 +1,82 @@ +open import Realizability.ApplicativeStructure +open import Realizability.CombinatoryAlgebra +open import Realizability.PropResizing +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Path +open import Cubical.Functions.FunExtEquiv +open import Cubical.Relation.Binary +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Vec +open import Cubical.Reflection.RecordEquiv +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Categories.Category +open import Cubical.Categories.Functor hiding (Id) + +module Realizability.PERs.SubQuotient + {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.PERs.PER ca +open import Realizability.Modest.Base ca + +open CombinatoryAlgebra ca +open Combinators ca renaming (i to Id; ia≡a to Ida≡a) + +module _ + (per : PER) where + + subQuotient : Type ℓ + subQuotient = A / per .PER.relation + + subQuotientRealizability : A → subQuotient → hProp ℓ + subQuotientRealizability r [a] = + SQ.rec + isSetHProp + (λ a → ([ a ] ≡ [ r ]) , squash/ [ a ] [ r ]) + (λ a b a~b → + Σ≡Prop + (λ _ → isPropIsProp) + (hPropExt (squash/ [ a ] [ r ]) (squash/ [ b ] [ r ]) (λ [a]≡[r] → sym (eq/ a b a~b) ∙ [a]≡[r]) λ [b]≡[r] → sym (eq/ b a (per .PER.isPER .fst a b a~b)) ∙ [b]≡[r])) + [a] + + subQuotientAssembly : Assembly subQuotient + Assembly._⊩_ subQuotientAssembly r [a] = ⟨ subQuotientRealizability r [a] ⟩ + Assembly.isSetX subQuotientAssembly = squash/ + Assembly.⊩isPropValued subQuotientAssembly r [a] = str (subQuotientRealizability r [a]) + Assembly.⊩surjective subQuotientAssembly [a] = + do + (a , [a]≡[a]) ← []surjective [a] + return + (a , (subst (λ [a] → ⟨ subQuotientRealizability a [a] ⟩) [a]≡[a] refl)) + + isModestSubQuotientAssembly : isModest subQuotientAssembly + isModestSubQuotientAssembly x y a a⊩x a⊩y = + SQ.elimProp2 + {P = motive} + isPropMotive + coreMap + x y a a⊩x a⊩y where + motive : subQuotient → subQuotient → Type ℓ + motive x y = (a : A) → a ⊩[ subQuotientAssembly ] x → a ⊩[ subQuotientAssembly ] y → x ≡ y + + isPropMotive : ∀ x y → isProp (motive x y) + isPropMotive x y = isPropΠ3 λ _ _ _ → squash/ x y + + coreMap : (r s : A) → motive [ r ] [ s ] + coreMap r s a a⊩[r] a⊩[s] = + [ r ] + ≡⟨ a⊩[r] ⟩ + [ a ] + ≡⟨ sym a⊩[s] ⟩ + [ s ] + ∎ diff --git a/src/Realizability/PERs/SubQuotientOfCanonicalPER.agda b/src/Realizability/PERs/SubQuotientOfCanonicalPER.agda new file mode 100644 index 0000000..2037d7d --- /dev/null +++ b/src/Realizability/PERs/SubQuotientOfCanonicalPER.agda @@ -0,0 +1,72 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Foundations.Univalence +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.PERs.SubQuotientOfCanonicalPER {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.SubQuotient ca + +module _ (per : PER) where + + theCanonicalPER : PER + theCanonicalPER = canonicalPER (subQuotientAssembly per) (isModestSubQuotientAssembly per) + + opaque + canonicalPEROfSubQuotientIsId : theCanonicalPER ≡ per + canonicalPEROfSubQuotientIsId = + PER≡ _ _ + (funExt₂ pointwise) where + opaque + effectiveness : (x : subQuotient per) (a b : A) → a ⊩[ subQuotientAssembly per ] x → b ⊩[ subQuotientAssembly per ] x → per .PER.relation a b + effectiveness x a b a⊩x b⊩x = {!!} + + opaque + dir1 : ∀ a b → theCanonicalPER .PER.relation a b → per .PER.relation a b + dir1 a b ∃x = + equivFun + (propTruncIdempotent≃ (per .PER.isPropValued a b)) + (do + (x , a⊩x , b⊩x) ← ∃x + return (effectiveness x a b a⊩x b⊩x)) + + opaque + pointwise : ∀ a b → theCanonicalPER .PER.relation a b ≡ per .PER.relation a b + pointwise a b = + hPropExt + (theCanonicalPER .PER.isPropValued a b) + (per .PER.isPropValued a b) + (dir1 a b) + {!!} + From 0c8e6265c73a54c249e453b197e9bafd01dc0bc3 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Fri, 5 Jul 2024 23:39:25 +0530 Subject: [PATCH 55/61] Subquotient and Canonical PER Operations --- .../Modest/GenericUniformFamily.agda | 63 +++++-------- .../SubQuotientCanonicalPERToOriginal.agda | 79 +++++++++++++++++ .../Modest/UnresizedGeneric.agda | 88 +++++++++++++++++++ src/Realizability/PERs/SubQuotient.agda | 51 ++++++----- 4 files changed, 216 insertions(+), 65 deletions(-) create mode 100644 src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda create mode 100644 src/Realizability/Modest/UnresizedGeneric.agda diff --git a/src/Realizability/Modest/GenericUniformFamily.agda b/src/Realizability/Modest/GenericUniformFamily.agda index 50388b7..27aa6e6 100644 --- a/src/Realizability/Modest/GenericUniformFamily.agda +++ b/src/Realizability/Modest/GenericUniformFamily.agda @@ -34,6 +34,7 @@ open import Realizability.Assembly.SetsReflectiveSubcategory ca open import Realizability.Modest.Base ca open import Realizability.Modest.UniformFamily ca open import Realizability.Modest.CanonicalPER ca +open import Realizability.Modest.UnresizedGeneric ca open import Realizability.PERs.PER ca open import Realizability.PERs.ResizedPER ca resizing open import Realizability.PERs.SubQuotient ca @@ -99,6 +100,7 @@ uniformFamilyCleavage {X , asmX} {Y , asmY} f N = cartfᴰ : isCartesian f fᴰ cartfᴰ = isCartesian'→isCartesian f fᴰ cart'fᴰ + ∇PER = ∇ ⟅ ResizedPER , isSetResizedPER ⟆ asm∇PER = ∇PER .snd @@ -108,57 +110,36 @@ UniformFamily.carriers (GenericObject.displayed genericUniformFamily) per = subQ UniformFamily.assemblies (GenericObject.displayed genericUniformFamily) per = subQuotientAssembly (enlargePER per) UniformFamily.isModestFamily (GenericObject.displayed genericUniformFamily) per = isModestSubQuotientAssembly (enlargePER per) GenericObject.isGeneric genericUniformFamily {X , asmX} M = - f , fᴰ , {!!} where + f , fᴰ , isCartesian'→isCartesian f fᴰ cart' where f : AssemblyMorphism asmX asm∇PER AssemblyMorphism.map f x = shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x)) AssemblyMorphism.tracker f = return (k , (λ _ _ _ → tt*)) - subQuotientTargetType : (x : X) → M .carriers x → Type ℓ - subQuotientTargetType x Mx = subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + subQuotientCod≡ : ∀ per → subQuotient (enlargePER (shrinkPER per)) ≡ subQuotient per + subQuotientCod≡ per = cong subQuotient (enlargePER⋆shrinkPER≡id per) fᴰ : DisplayedUFamMap asmX asm∇PER f M (genericUniformFamily .GenericObject.displayed) DisplayedUFamMap.fibrewiseMap fᴰ x Mx = - subst (λ x → subQuotient x) (sym equation) target module CartesianMapDefinition where - targetType : Type ℓ - targetType = subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) - - opaque - equation : enlargePER (shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x))) ≡ canonicalPER (M .assemblies x) (M .isModestFamily x) - equation = enlargePER⋆shrinkPER≡id (canonicalPER (M .assemblies x) (M .isModestFamily x)) - - mainMap : Σ[ a ∈ A ] (a ⊩[ M .assemblies x ] Mx) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) - mainMap (a , a⊩Mx) = [ a ] - - mainMap2Constant : 2-Constant mainMap - mainMap2Constant (a , a⊩Mx) (b , b⊩Mx) = eq/ a b (return (Mx , a⊩Mx , b⊩Mx)) - - opaque - target : targetType - target = - PT.rec→Set - squash/ - mainMap - mainMap2Constant - (M .assemblies x .⊩surjective Mx) - - opaque - isTrackedCartesianMap : - ∀ (a : A) - → a ⊩[ asmX ] x - → (b : A) - → b ⊩[ M .assemblies x ] Mx - → ⟨ - subQuotientRealizability - (enlargePER (shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x)))) - (λ*2 (# one) ⨾ a ⨾ b) - (subst subQuotient (sym equation) target) - ⟩ - isTrackedCartesianMap a a⊩x b b⊩Mx = - {!!} + subst + subQuotient + (sym + (enlargePER⋆shrinkPER≡id + (canonicalPER (M .assemblies x) (M .isModestFamily x)))) + (Unresized.mainMap resizing asmX M x Mx) DisplayedUFamMap.isTracked fᴰ = do return (λ*2 (# one) , - λ x a a⊩x Mx b b⊩Mx → {!CartesianMapDefinition.isTrackedCartesianMap x Mx a a⊩x b b⊩Mx!}) + (λ x a a⊩x Mx b b⊩Mx → + equivFun + (propTruncIdempotent≃ {!!}) + (do + (r , r⊩mainMap) ← Unresized.isTrackedMainMap resizing asmX M + return {!!}))) + + opaque + unfolding isCartesian' + cart' : isCartesian' f fᴰ + cart' {Y , asmY} {N} g hᴰ = {!!} diff --git a/src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda b/src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda new file mode 100644 index 0000000..319a0ad --- /dev/null +++ b/src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda @@ -0,0 +1,79 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.SubQuotientCanonicalPERToOriginal {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +module + _ {X : Type ℓ} + (asmX : Assembly X) + (isModestAsmX : isModest asmX) where + + theCanonicalPER : PER + theCanonicalPER = canonicalPER asmX isModestAsmX + + theSubQuotient : Assembly (subQuotient theCanonicalPER) + theSubQuotient = subQuotientAssembly theCanonicalPER + + invert : AssemblyMorphism theSubQuotient asmX + AssemblyMorphism.map invert sq = SQ.rec (asmX .isSetX) mainMap mainMapCoh sq where + + mainMap : Σ[ a ∈ A ] (theCanonicalPER .PER.relation a a) → X + mainMap (a , a~a) = PT.rec→Set (asmX .isSetX) action 2ConstantAction a~a where + action : Σ[ x ∈ X ] ((a ⊩[ asmX ] x) × (a ⊩[ asmX ] x)) → X + action (x , _ , _) = x + + 2ConstantAction : 2-Constant action + 2ConstantAction (x , a⊩x , _) (x' , a⊩x' , _) = isModestAsmX x x' a a⊩x a⊩x' + + mainMapCoh : ∀ a b a~b → mainMap a ≡ mainMap b + mainMapCoh (a , a~a) (b , b~b) a~b = + PT.elim3 + {P = λ a~a b~b a~b → mainMap (a , a~a) ≡ mainMap (b , b~b)} + (λ _ _ _ → asmX .isSetX _ _) + (λ { (x , a⊩x , _) (x' , b⊩x' , _) (x'' , a⊩x'' , b⊩x'') → + {!isModestAsmX x x' !} }) + a~a + b~b + a~b + AssemblyMorphism.tracker invert = {!!} diff --git a/src/Realizability/Modest/UnresizedGeneric.agda b/src/Realizability/Modest/UnresizedGeneric.agda new file mode 100644 index 0000000..3ab06c1 --- /dev/null +++ b/src/Realizability/Modest/UnresizedGeneric.agda @@ -0,0 +1,88 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.UnresizedGeneric {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) (resizing : hPropResizing ℓ) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.ResizedPER ca resizing +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +module Unresized + {X : Type ℓ} + (asmX : Assembly X) + (M : UniformFamily asmX) where + + theCanonicalPER : ∀ x → PER + theCanonicalPER x = canonicalPER (M . assemblies x) (M .isModestFamily x) + + elimRealizerForMx : ∀ {x : X} {Mx : M .carriers x} → Σ[ a ∈ A ] (a ⊩[ M .assemblies x ] Mx) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + elimRealizerForMx {x} {Mx} (a , a⊩Mx) = [ a , (return (Mx , a⊩Mx , a⊩Mx)) ] + + elimRealizerForMx2Constant : ∀ {x Mx} → 2-Constant (elimRealizerForMx {x} {Mx}) + elimRealizerForMx2Constant {x} {Mx} (a , a⊩Mx) (b , b⊩Mx) = + eq/ + (a , (return (Mx , a⊩Mx , a⊩Mx))) + (b , return (Mx , b⊩Mx , b⊩Mx)) + (return (Mx , a⊩Mx , b⊩Mx)) + + mainMap : (x : X) (Mx : M .carriers x) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + mainMap x Mx = + PT.rec→Set + squash/ + (elimRealizerForMx {x = x} {Mx = Mx}) + (elimRealizerForMx2Constant {x = x} {Mx = Mx}) + (M .assemblies x .⊩surjective Mx) + + opaque + isTrackedMainMap : ∃[ r ∈ A ] (∀ (x : X) (a : A) → a ⊩[ asmX ] x → (Mx : M .carriers x) → (b : A) → b ⊩[ M .assemblies x ] Mx → (r ⨾ a ⨾ b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] (mainMap x Mx)) + isTrackedMainMap = + return + ((λ*2 (# zero)) , + (λ x a a⊩x Mx b b⊩Mx → + PT.elim + {P = λ MxRealizer → (λ*2 (# zero) ⨾ a ⨾ b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] (PT.rec→Set squash/ (elimRealizerForMx {x = x} {Mx = Mx}) (elimRealizerForMx2Constant {x = x} {Mx = Mx}) MxRealizer)} + (λ ⊩Mx → subQuotientAssembly (theCanonicalPER x) .⊩isPropValued (λ*2 (# zero) ⨾ a ⨾ b) (rec→Set squash/ elimRealizerForMx elimRealizerForMx2Constant ⊩Mx)) + (λ { (c , c⊩Mx) → + subst + (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx (c , c⊩Mx))) + (sym (λ*2ComputationRule (# zero) a b)) + (return (Mx , b⊩Mx , c⊩Mx))}) + (M .assemblies x .⊩surjective Mx))) diff --git a/src/Realizability/PERs/SubQuotient.agda b/src/Realizability/PERs/SubQuotient.agda index e28420a..0c12f24 100644 --- a/src/Realizability/PERs/SubQuotient.agda +++ b/src/Realizability/PERs/SubQuotient.agda @@ -35,48 +35,51 @@ open Combinators ca renaming (i to Id; ia≡a to Ida≡a) module _ (per : PER) where + domain : Type ℓ + domain = Σ[ a ∈ A ] (per .PER.relation a a) + subQuotient : Type ℓ - subQuotient = A / per .PER.relation + subQuotient = domain / λ { (a , _) (b , _) → per .PER.relation a b } subQuotientRealizability : A → subQuotient → hProp ℓ subQuotientRealizability r [a] = SQ.rec isSetHProp - (λ a → ([ a ] ≡ [ r ]) , squash/ [ a ] [ r ]) - (λ a b a~b → + (λ { (a , a~a) → r ~[ per ] a , isProp~ r per a }) + (λ { (a , a~a) (b , b~b) a~b → Σ≡Prop - (λ _ → isPropIsProp) - (hPropExt (squash/ [ a ] [ r ]) (squash/ [ b ] [ r ]) (λ [a]≡[r] → sym (eq/ a b a~b) ∙ [a]≡[r]) λ [b]≡[r] → sym (eq/ b a (per .PER.isPER .fst a b a~b)) ∙ [b]≡[r])) + (λ x → isPropIsProp) + (hPropExt + (isProp~ r per a) + (isProp~ r per b) + (λ r~a → PER.isTransitive per r a b r~a a~b) + (λ r~b → PER.isTransitive per r b a r~b (PER.isSymmetric per a b a~b))) }) [a] - + + subQuotientAssembly : Assembly subQuotient Assembly._⊩_ subQuotientAssembly r [a] = ⟨ subQuotientRealizability r [a] ⟩ Assembly.isSetX subQuotientAssembly = squash/ Assembly.⊩isPropValued subQuotientAssembly r [a] = str (subQuotientRealizability r [a]) Assembly.⊩surjective subQuotientAssembly [a] = - do - (a , [a]≡[a]) ← []surjective [a] - return - (a , (subst (λ [a] → ⟨ subQuotientRealizability a [a] ⟩) [a]≡[a] refl)) + SQ.elimProp + {P = λ [a] → ∃[ r ∈ A ] ⟨ subQuotientRealizability r [a] ⟩} + (λ [a] → isPropPropTrunc) + (λ { (a , a~a) → return (a , a~a) }) + [a] + isModestSubQuotientAssembly : isModest subQuotientAssembly isModestSubQuotientAssembly x y a a⊩x a⊩y = SQ.elimProp2 - {P = motive} + {P = λ x y → motive x y} isPropMotive - coreMap - x y a a⊩x a⊩y where - motive : subQuotient → subQuotient → Type ℓ - motive x y = (a : A) → a ⊩[ subQuotientAssembly ] x → a ⊩[ subQuotientAssembly ] y → x ≡ y + (λ { (x , x~x) (y , y~y) a a~x a~y → + eq/ (x , x~x) (y , y~y) (PER.isTransitive per x a y (PER.isSymmetric per a x a~x) a~y) }) + x y + a a⊩x a⊩y where + motive : ∀ (x y : subQuotient) → Type ℓ + motive x y = ∀ (a : A) (a⊩x : a ⊩[ subQuotientAssembly ] x) (a⊩y : a ⊩[ subQuotientAssembly ] y) → x ≡ y isPropMotive : ∀ x y → isProp (motive x y) isPropMotive x y = isPropΠ3 λ _ _ _ → squash/ x y - - coreMap : (r s : A) → motive [ r ] [ s ] - coreMap r s a a⊩[r] a⊩[s] = - [ r ] - ≡⟨ a⊩[r] ⟩ - [ a ] - ≡⟨ sym a⊩[s] ⟩ - [ s ] - ∎ From 75c6f59213942ffff4c5193a98a9964f40d2996d Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sun, 7 Jul 2024 05:56:56 +0530 Subject: [PATCH 56/61] Fix performance issues --- .../Modest/GenericUniformFamily.agda | 99 +++-------------- .../Modest/SubquotientUniformFamily.agda | 100 ++++++++++++++++++ .../Modest/UniformFamilyCleavage.agda | 100 ++++++++++++++++++ .../Modest/UnresizedGeneric.agda | 66 ++++++------ 4 files changed, 250 insertions(+), 115 deletions(-) create mode 100644 src/Realizability/Modest/SubquotientUniformFamily.agda create mode 100644 src/Realizability/Modest/UniformFamilyCleavage.agda diff --git a/src/Realizability/Modest/GenericUniformFamily.agda b/src/Realizability/Modest/GenericUniformFamily.agda index 27aa6e6..2a37bef 100644 --- a/src/Realizability/Modest/GenericUniformFamily.agda +++ b/src/Realizability/Modest/GenericUniformFamily.agda @@ -35,6 +35,7 @@ open import Realizability.Modest.Base ca open import Realizability.Modest.UniformFamily ca open import Realizability.Modest.CanonicalPER ca open import Realizability.Modest.UnresizedGeneric ca +open import Realizability.Modest.SubquotientUniformFamily ca resizing open import Realizability.PERs.PER ca open import Realizability.PERs.ResizedPER ca resizing open import Realizability.PERs.SubQuotient ca @@ -46,100 +47,30 @@ open Contravariant UNIMOD open UniformFamily open DisplayedUFamMap -uniformFamilyCleavage : cleavage -uniformFamilyCleavage {X , asmX} {Y , asmY} f N = - N' , fᴰ , cartfᴰ where - N' : UniformFamily asmX - UniformFamily.carriers N' x = N .carriers (f .map x) - UniformFamily.assemblies N' x = N .assemblies (f .map x) - UniformFamily.isModestFamily N' x = N .isModestFamily (f .map x) - - fᴰ : DisplayedUFamMap asmX asmY f N' N - DisplayedUFamMap.fibrewiseMap fᴰ x Nfx = Nfx - DisplayedUFamMap.isTracked fᴰ = - do - let - realizer : Term as 2 - realizer = # zero - return - (λ*2 realizer , - (λ x a a⊩x Nfx b b⊩Nfx → - subst - (_⊩[ N .assemblies (f .map x) ] Nfx) - (sym (λ*2ComputationRule realizer a b)) - b⊩Nfx)) - - opaque - unfolding isCartesian' - cart'fᴰ : isCartesian' f fᴰ - cart'fᴰ {Z , asmZ} {M} g hᴰ = - (! , !⋆fᴰ≡hᴰ) , - λ { (!' , !'comm) → - Σ≡Prop - (λ ! → UNIMOD .Categoryᴰ.isSetHomᴰ _ _) - (DisplayedUFamMapPathP - _ _ _ _ _ _ _ _ _ - λ z Mz → - sym - (!' .fibrewiseMap z Mz - ≡[ i ]⟨ !'comm i .fibrewiseMap z Mz ⟩ - hᴰ .fibrewiseMap z Mz - ∎)) } where - ! : DisplayedUFamMap asmZ asmX g M N' - DisplayedUFamMap.fibrewiseMap ! z Mz = hᴰ .fibrewiseMap z Mz - DisplayedUFamMap.isTracked ! = hᴰ .isTracked - - !⋆fᴰ≡hᴰ : composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ ≡ hᴰ - !⋆fᴰ≡hᴰ = - DisplayedUFamMapPathP - asmZ asmY _ _ - M N - (composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ) hᴰ refl - λ z Mz → refl - - cartfᴰ : isCartesian f fᴰ - cartfᴰ = isCartesian'→isCartesian f fᴰ cart'fᴰ - - -∇PER = ∇ ⟅ ResizedPER , isSetResizedPER ⟆ -asm∇PER = ∇PER .snd +private + ∇PER = ∇ ⟅ ResizedPER , isSetResizedPER ⟆ + asm∇PER = ∇PER .snd genericUniformFamily : GenericObject UNIMOD GenericObject.base genericUniformFamily = ∇PER -UniformFamily.carriers (GenericObject.displayed genericUniformFamily) per = subQuotient (enlargePER per) -UniformFamily.assemblies (GenericObject.displayed genericUniformFamily) per = subQuotientAssembly (enlargePER per) -UniformFamily.isModestFamily (GenericObject.displayed genericUniformFamily) per = isModestSubQuotientAssembly (enlargePER per) +GenericObject.displayed genericUniformFamily = subQuotientUniformFamily GenericObject.isGeneric genericUniformFamily {X , asmX} M = f , fᴰ , isCartesian'→isCartesian f fᴰ cart' where f : AssemblyMorphism asmX asm∇PER - AssemblyMorphism.map f x = shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x)) - AssemblyMorphism.tracker f = return (k , (λ _ _ _ → tt*)) + f = classifyingMap M - subQuotientCod≡ : ∀ per → subQuotient (enlargePER (shrinkPER per)) ≡ subQuotient per - subQuotientCod≡ per = cong subQuotient (enlargePER⋆shrinkPER≡id per) - - fᴰ : DisplayedUFamMap asmX asm∇PER f M (genericUniformFamily .GenericObject.displayed) - DisplayedUFamMap.fibrewiseMap fᴰ x Mx = - subst - subQuotient - (sym - (enlargePER⋆shrinkPER≡id - (canonicalPER (M .assemblies x) (M .isModestFamily x)))) - (Unresized.mainMap resizing asmX M x Mx) - DisplayedUFamMap.isTracked fᴰ = - do - return - (λ*2 (# one) , - (λ x a a⊩x Mx b b⊩Mx → - equivFun - (propTruncIdempotent≃ {!!}) - (do - (r , r⊩mainMap) ← Unresized.isTrackedMainMap resizing asmX M - return {!!}))) + opaque + unfolding Unresized.mainMap + fᴰ : DisplayedUFamMap asmX asm∇PER f M (genericUniformFamily .GenericObject.displayed) + fᴰ = classifyingMapᴰ M opaque unfolding isCartesian' cart' : isCartesian' f fᴰ - cart' {Y , asmY} {N} g hᴰ = {!!} + cart' {Y , asmY} {N} g hᴰ = + (! , {!!}) , {!!} where + ! : DisplayedUFamMap asmY asmX g N M + DisplayedUFamMap.fibrewiseMap ! y Ny = {!hᴰ .fibrewiseMap y Ny!} + DisplayedUFamMap.isTracked ! = {!!} diff --git a/src/Realizability/Modest/SubquotientUniformFamily.agda b/src/Realizability/Modest/SubquotientUniformFamily.agda new file mode 100644 index 0000000..775e753 --- /dev/null +++ b/src/Realizability/Modest/SubquotientUniformFamily.agda @@ -0,0 +1,100 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.SubquotientUniformFamily {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) (resizing : hPropResizing ℓ) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.Modest.UnresizedGeneric ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.ResizedPER ca resizing +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +private + ∇PER = ∇ ⟅ ResizedPER , isSetResizedPER ⟆ + asm∇PER = ∇PER .snd + +-- G over ∇PER +subQuotientUniformFamily : UniformFamily asm∇PER +UniformFamily.carriers subQuotientUniformFamily per = subQuotient (enlargePER per) +UniformFamily.assemblies subQuotientUniformFamily per = subQuotientAssembly (enlargePER per) +UniformFamily.isModestFamily subQuotientUniformFamily per = isModestSubQuotientAssembly (enlargePER per) + +-- For any uniform family M over X +-- there is a map to the canonical uniform family G over ∇PER +module _ + {X : Type ℓ} + {asmX : Assembly X} + (M : UniformFamily asmX) where + + classifyingMap : AssemblyMorphism asmX asm∇PER + AssemblyMorphism.map classifyingMap x = shrinkPER (canonicalPER (M .assemblies x) (M .isModestFamily x)) + AssemblyMorphism.tracker classifyingMap = return (k , (λ _ _ _ → tt*)) + + opaque + unfolding Unresized.mainMap + classifyingMapᴰ : DisplayedUFamMap asmX asm∇PER classifyingMap M subQuotientUniformFamily + DisplayedUFamMap.fibrewiseMap classifyingMapᴰ x Mx = + subst + subQuotient + (sym + (enlargePER⋆shrinkPER≡id + (canonicalPER (M .assemblies x) (M .isModestFamily x)))) + (Unresized.mainMap resizing asmX M x Mx .fst) + DisplayedUFamMap.isTracked classifyingMapᴰ = + return + ((λ*2 (# zero)) , + (λ x a a⊩x Mx b b⊩Mx → + -- Is there a way to do this that avoids transp ? + -- This really eats type-checking time + transp + (λ i → + ⟨ + subQuotientRealizability + (enlargePER⋆shrinkPER≡id + (canonicalPER (M .assemblies x) (M .isModestFamily x)) (~ i)) + (λ*2 (# zero) ⨾ a ⨾ b) + (subst-filler + subQuotient + (sym + (enlargePER⋆shrinkPER≡id + (canonicalPER (M .assemblies x) (M .isModestFamily x)))) + (Unresized.mainMap resizing asmX M x Mx .fst) i) + ⟩) i0 (Unresized.mainMap resizing asmX M x Mx .snd a a⊩x b b⊩Mx))) + diff --git a/src/Realizability/Modest/UniformFamilyCleavage.agda b/src/Realizability/Modest/UniformFamilyCleavage.agda new file mode 100644 index 0000000..2989477 --- /dev/null +++ b/src/Realizability/Modest/UniformFamilyCleavage.agda @@ -0,0 +1,100 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.UniformFamilyCleavage {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.Modest.UnresizedGeneric ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +uniformFamilyCleavage : cleavage +uniformFamilyCleavage {X , asmX} {Y , asmY} f N = + N' , fᴰ , cartfᴰ where + N' : UniformFamily asmX + UniformFamily.carriers N' x = N .carriers (f .map x) + UniformFamily.assemblies N' x = N .assemblies (f .map x) + UniformFamily.isModestFamily N' x = N .isModestFamily (f .map x) + + fᴰ : DisplayedUFamMap asmX asmY f N' N + DisplayedUFamMap.fibrewiseMap fᴰ x Nfx = Nfx + DisplayedUFamMap.isTracked fᴰ = + do + let + realizer : Term as 2 + realizer = # zero + return + (λ*2 realizer , + (λ x a a⊩x Nfx b b⊩Nfx → + subst + (_⊩[ N .assemblies (f .map x) ] Nfx) + (sym (λ*2ComputationRule realizer a b)) + b⊩Nfx)) + + opaque + unfolding isCartesian' + cart'fᴰ : isCartesian' f fᴰ + cart'fᴰ {Z , asmZ} {M} g hᴰ = + (! , !⋆fᴰ≡hᴰ) , + λ { (!' , !'comm) → + Σ≡Prop + (λ ! → UNIMOD .Categoryᴰ.isSetHomᴰ _ _) + (DisplayedUFamMapPathP + _ _ _ _ _ _ _ _ _ + λ z Mz → + sym + (!' .fibrewiseMap z Mz + ≡[ i ]⟨ !'comm i .fibrewiseMap z Mz ⟩ + hᴰ .fibrewiseMap z Mz + ∎)) } where + ! : DisplayedUFamMap asmZ asmX g M N' + DisplayedUFamMap.fibrewiseMap ! z Mz = hᴰ .fibrewiseMap z Mz + DisplayedUFamMap.isTracked ! = hᴰ .isTracked + + !⋆fᴰ≡hᴰ : composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ ≡ hᴰ + !⋆fᴰ≡hᴰ = + DisplayedUFamMapPathP + asmZ asmY _ _ + M N + (composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ) hᴰ refl + λ z Mz → refl + + cartfᴰ : isCartesian f fᴰ + cartfᴰ = isCartesian'→isCartesian f fᴰ cart'fᴰ diff --git a/src/Realizability/Modest/UnresizedGeneric.agda b/src/Realizability/Modest/UnresizedGeneric.agda index 3ab06c1..4c97614 100644 --- a/src/Realizability/Modest/UnresizedGeneric.agda +++ b/src/Realizability/Modest/UnresizedGeneric.agda @@ -53,36 +53,40 @@ module Unresized theCanonicalPER : ∀ x → PER theCanonicalPER x = canonicalPER (M . assemblies x) (M .isModestFamily x) - elimRealizerForMx : ∀ {x : X} {Mx : M .carriers x} → Σ[ a ∈ A ] (a ⊩[ M .assemblies x ] Mx) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) - elimRealizerForMx {x} {Mx} (a , a⊩Mx) = [ a , (return (Mx , a⊩Mx , a⊩Mx)) ] - - elimRealizerForMx2Constant : ∀ {x Mx} → 2-Constant (elimRealizerForMx {x} {Mx}) - elimRealizerForMx2Constant {x} {Mx} (a , a⊩Mx) (b , b⊩Mx) = - eq/ - (a , (return (Mx , a⊩Mx , a⊩Mx))) - (b , return (Mx , b⊩Mx , b⊩Mx)) - (return (Mx , a⊩Mx , b⊩Mx)) + elimRealizerForMx : ∀ (x : X) (Mx : M .carriers x) → Σ[ a ∈ A ] (a ⊩[ M .assemblies x ] Mx) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + elimRealizerForMx x Mx (a , a⊩Mx) = [ a , (return (Mx , a⊩Mx , a⊩Mx)) ] - mainMap : (x : X) (Mx : M .carriers x) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) - mainMap x Mx = - PT.rec→Set - squash/ - (elimRealizerForMx {x = x} {Mx = Mx}) - (elimRealizerForMx2Constant {x = x} {Mx = Mx}) - (M .assemblies x .⊩surjective Mx) - opaque - isTrackedMainMap : ∃[ r ∈ A ] (∀ (x : X) (a : A) → a ⊩[ asmX ] x → (Mx : M .carriers x) → (b : A) → b ⊩[ M .assemblies x ] Mx → (r ⨾ a ⨾ b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] (mainMap x Mx)) - isTrackedMainMap = - return - ((λ*2 (# zero)) , - (λ x a a⊩x Mx b b⊩Mx → - PT.elim - {P = λ MxRealizer → (λ*2 (# zero) ⨾ a ⨾ b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] (PT.rec→Set squash/ (elimRealizerForMx {x = x} {Mx = Mx}) (elimRealizerForMx2Constant {x = x} {Mx = Mx}) MxRealizer)} - (λ ⊩Mx → subQuotientAssembly (theCanonicalPER x) .⊩isPropValued (λ*2 (# zero) ⨾ a ⨾ b) (rec→Set squash/ elimRealizerForMx elimRealizerForMx2Constant ⊩Mx)) - (λ { (c , c⊩Mx) → - subst - (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx (c , c⊩Mx))) - (sym (λ*2ComputationRule (# zero) a b)) - (return (Mx , b⊩Mx , c⊩Mx))}) - (M .assemblies x .⊩surjective Mx))) + elimRealizerForMx2Constant : ∀ x Mx → 2-Constant (elimRealizerForMx x Mx) + elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx) = + eq/ + (a , (return (Mx , a⊩Mx , a⊩Mx))) + (b , return (Mx , b⊩Mx , b⊩Mx)) + (return (Mx , a⊩Mx , b⊩Mx)) + + mainMapType : Type _ + mainMapType = + ∀ (x : X) (Mx : M .carriers x) → + Σ[ out ∈ (subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x))) ] + (∀ (a : A) → a ⊩[ asmX ] x → (b : A) → b ⊩[ M .assemblies x ] Mx → (λ*2 (# zero) ⨾ a ⨾ b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] out) + + opaque + mainMap : mainMapType + mainMap x Mx = + PT.rec→Set + (isSetΣ + squash/ + (λ out → + isSetΠ3 + λ a a⊩x b → + isSet→ + (isProp→isSet + (str + (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero) ⨾ a ⨾ b) out))))) + ((λ { (c , c⊩Mx) → + (elimRealizerForMx x Mx (c , c⊩Mx)) , + (λ a a⊩x b b⊩Mx → + subst (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx x Mx (c , c⊩Mx))) (sym (λ*2ComputationRule (# zero) a b)) (return (Mx , b⊩Mx , c⊩Mx))) })) + (λ { (a , a⊩Mx) (b , b⊩Mx) → + Σ≡Prop (λ out → isPropΠ4 λ a a⊩x b b⊩Mx → str (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero) ⨾ a ⨾ b) out)) (elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx)) }) + (M .assemblies x .⊩surjective Mx) From 1b95e78dbb71f76c475bab62b648e6af9f043a79 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sat, 17 Aug 2024 00:01:17 +0530 Subject: [PATCH 57/61] Isomorphism of PERs and Modest Sets --- .../#SubQuotientCanonicalPERToOriginal.agda# | 53 +++++++ .../.#SubQuotientCanonicalPERToOriginal.agda | 1 + src/Realizability/Modest/CanonicalPER.agda | 22 +-- .../Modest/GenericUniformFamily.agda | 17 +- .../Modest/SubQuotientCanonicalPERIso.agda | 147 ++++++++++++++++++ .../SubQuotientCanonicalPERToOriginal.agda | 49 +++--- .../Modest/UnresizedGeneric.agda | 10 +- 7 files changed, 261 insertions(+), 38 deletions(-) create mode 100644 src/Realizability/Modest/#SubQuotientCanonicalPERToOriginal.agda# create mode 120000 src/Realizability/Modest/.#SubQuotientCanonicalPERToOriginal.agda create mode 100644 src/Realizability/Modest/SubQuotientCanonicalPERIso.agda diff --git a/src/Realizability/Modest/#SubQuotientCanonicalPERToOriginal.agda# b/src/Realizability/Modest/#SubQuotientCanonicalPERToOriginal.agda# new file mode 100644 index 0000000..c6ceb57 --- /dev/null +++ b/src/Realizability/Modest/#SubQuotientCanonicalPERToOriginal.agda# @@ -0,0 +1,53 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.SubQuotientCanonicalPERToOriginal {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +module + _ {X : Type ℓ} + (asmX : Assembly X) + (isModestAsmX : isModest asmX) where + + + diff --git a/src/Realizability/Modest/.#SubQuotientCanonicalPERToOriginal.agda b/src/Realizability/Modest/.#SubQuotientCanonicalPERToOriginal.agda new file mode 120000 index 0000000..4751d10 --- /dev/null +++ b/src/Realizability/Modest/.#SubQuotientCanonicalPERToOriginal.agda @@ -0,0 +1 @@ +rahul@Rahuls-MacBook-Air.local.13251:1720167879 \ No newline at end of file diff --git a/src/Realizability/Modest/CanonicalPER.agda b/src/Realizability/Modest/CanonicalPER.agda index 30ade69..f0d0682 100644 --- a/src/Realizability/Modest/CanonicalPER.agda +++ b/src/Realizability/Modest/CanonicalPER.agda @@ -48,19 +48,13 @@ module _ (isModestAsmX : isModest asmX) where canonicalPER : PER - PER.relation canonicalPER a b = ∃[ x ∈ X ] a ⊩[ asmX ] x × b ⊩[ asmX ] x - PER.isPropValued canonicalPER a b = isPropPropTrunc - fst (PER.isPER canonicalPER) a b ∃x = PT.map (λ { (x , a⊩x , b⊩x) → x , b⊩x , a⊩x }) ∃x - snd (PER.isPER canonicalPER) a b c ∃x ∃x' = - do - (x , a⊩x , b⊩x) ← ∃x - (x' , b⊩x' , c⊩x') ← ∃x' - let - x≡x' : x ≡ x' - x≡x' = isModestAsmX x x' b b⊩x b⊩x' - - c⊩x : c ⊩[ asmX ] x - c⊩x = subst (c ⊩[ asmX ]_) (sym x≡x') c⊩x' - return (x , a⊩x , c⊩x) + PER.relation canonicalPER a b = Σ[ x ∈ X ] a ⊩[ asmX ] x × b ⊩[ asmX ] x + PER.isPropValued canonicalPER a b (x , a⊩x , b⊩x) (x' , a⊩x' , b⊩x') = + Σ≡Prop + (λ x → isProp× (asmX .⊩isPropValued a x) (asmX .⊩isPropValued b x)) + (isModestAsmX x x' a a⊩x a⊩x') + fst (PER.isPER canonicalPER) a b (x , a⊩x , b⊩x) = x , b⊩x , a⊩x + snd (PER.isPER canonicalPER) a b c (x , a⊩x , b⊩x) (x' , b⊩x' , c⊩x') = + x' , subst (a ⊩[ asmX ]_) (isModestAsmX x x' b b⊩x b⊩x') a⊩x , c⊩x' diff --git a/src/Realizability/Modest/GenericUniformFamily.agda b/src/Realizability/Modest/GenericUniformFamily.agda index 2a37bef..0a0f881 100644 --- a/src/Realizability/Modest/GenericUniformFamily.agda +++ b/src/Realizability/Modest/GenericUniformFamily.agda @@ -70,7 +70,22 @@ GenericObject.isGeneric genericUniformFamily {X , asmX} M = cart' : isCartesian' f fᴰ cart' {Y , asmY} {N} g hᴰ = (! , {!!}) , {!!} where + ! : DisplayedUFamMap asmY asmX g N M - DisplayedUFamMap.fibrewiseMap ! y Ny = {!hᴰ .fibrewiseMap y Ny!} + DisplayedUFamMap.fibrewiseMap ! y Ny = lhsIsoRhs .fst .map {!hᴰ .fibrewiseMap y Ny!} module !Definition where + gy : X + gy = g .map y + + canonicalPERMgy : PER + canonicalPERMgy = canonicalPER (M .assemblies gy) (M .isModestFamily gy) + + lhsModestSet : MOD .Category.ob + lhsModestSet = subQuotient canonicalPERMgy , subQuotientAssembly canonicalPERMgy , isModestSubQuotientAssembly canonicalPERMgy + + rhsModestSet : MOD .Category.ob + rhsModestSet = M .carriers gy , M .assemblies gy , M .isModestFamily gy + + lhsIsoRhs : CatIso MOD lhsModestSet rhsModestSet + lhsIsoRhs = {!!} DisplayedUFamMap.isTracked ! = {!!} diff --git a/src/Realizability/Modest/SubQuotientCanonicalPERIso.agda b/src/Realizability/Modest/SubQuotientCanonicalPERIso.agda new file mode 100644 index 0000000..82cb630 --- /dev/null +++ b/src/Realizability/Modest/SubQuotientCanonicalPERIso.agda @@ -0,0 +1,147 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.SubQuotientCanonicalPERIso {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +module + _ {X : Type ℓ} + (asmX : Assembly X) + (isModestAsmX : isModest asmX) where + + theCanonicalPER : PER + theCanonicalPER = canonicalPER asmX isModestAsmX + + theSubQuotient : Assembly (subQuotient theCanonicalPER) + theSubQuotient = subQuotientAssembly theCanonicalPER + + invert : AssemblyMorphism theSubQuotient asmX + AssemblyMorphism.map invert sq = SQ.rec (asmX .isSetX) reprAction reprActionCoh sq module Invert where + + reprAction : Σ[ a ∈ A ] (a ~[ theCanonicalPER ] a) → X + reprAction (a , x , a⊩x , _) = x + + reprActionCoh : ∀ a b a~b → reprAction a ≡ reprAction b + reprActionCoh (a , x , a⊩x , _) (b , x' , b⊩x' , _) (x'' , a⊩x'' , b⊩x'') = + x + ≡⟨ isModestAsmX x x'' a a⊩x a⊩x'' ⟩ + x'' + ≡⟨ isModestAsmX x'' x' b b⊩x'' b⊩x' ⟩ + x' + ∎ + AssemblyMorphism.tracker invert = return (Id , (λ sq a a⊩sq → goal sq a a⊩sq)) where + realizability : (sq : subQuotient theCanonicalPER) → (a : A) → a ⊩[ theSubQuotient ] sq → a ⊩[ asmX ] (invert .map sq) + realizability sq a a⊩sq = + SQ.elimProp + {P = motive} + isPropMotive + elemMotive + sq a a⊩sq where + + motive : (sq : subQuotient theCanonicalPER) → Type _ + motive sq = ∀ (a : A) → a ⊩[ theSubQuotient ] sq → a ⊩[ asmX ] (invert .map sq) + + isPropMotive : ∀ sq → isProp (motive sq) + isPropMotive sq = isPropΠ2 λ a a⊩sq → asmX .⊩isPropValued _ _ + + elemMotive : (x : domain theCanonicalPER) → motive [ x ] + elemMotive (r , x , r⊩x , _) a (x' , a⊩x' , r⊩x') = subst (a ⊩[ asmX ]_) (isModestAsmX x' x r r⊩x' r⊩x) a⊩x' + + goal : (sq : subQuotient theCanonicalPER) → (a : A) → a ⊩[ theSubQuotient ] sq → (Id ⨾ a) ⊩[ asmX ] (invert .map sq) + goal sq a a⊩sq = subst (_⊩[ asmX ] _) (sym (Ida≡a a)) (realizability sq a a⊩sq) + + forward : AssemblyMorphism asmX theSubQuotient + AssemblyMorphism.map forward x = subquot module Forward where + mainMap : Σ[ a ∈ A ] (a ⊩[ asmX ] x) → subQuotient theCanonicalPER + mainMap (a , a⊩x) = [ a , x , a⊩x , a⊩x ] + + mainMap2Constant : 2-Constant mainMap + mainMap2Constant (a , a⊩x) (b , b⊩x) = eq/ _ _ (x , a⊩x , b⊩x) + + subquot : subQuotient theCanonicalPER + subquot = PT.rec→Set squash/ mainMap mainMap2Constant (asmX .⊩surjective x) + AssemblyMorphism.tracker forward = + return + (Id , + (λ x a a⊩x → + PT.elim + {P = λ surj → (Id ⨾ a) ⊩[ theSubQuotient ] (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)} + (λ surj → theSubQuotient .⊩isPropValued (Id ⨾ a) (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)) + (λ { (b , b⊩x) → x , subst (_⊩[ asmX ] x) (sym (Ida≡a a)) a⊩x , b⊩x }) + (asmX .⊩surjective x))) + + subQuotientCanonicalIso : CatIso MOD (X , asmX , isModestAsmX) (subQuotient theCanonicalPER , theSubQuotient , isModestSubQuotientAssembly theCanonicalPER) + fst subQuotientCanonicalIso = forward + isIso.inv (snd subQuotientCanonicalIso) = invert + isIso.sec (snd subQuotientCanonicalIso) = goal where + opaque + pointwise : ∀ sq → (invert ⊚ forward) .map sq ≡ sq + pointwise sq = + SQ.elimProp + (λ sq → squash/ (forward .map (invert .map sq)) sq) + (λ { d@(a , x , a⊩x , a⊩'x) → + PT.elim + {P = λ surj → PT.rec→Set squash/ (Forward.mainMap (Invert.reprAction [ d ] d)) (Forward.mainMap2Constant (Invert.reprAction [ d ] d)) surj ≡ [ d ]} + (λ surj → squash/ _ _) + (λ { (b , b⊩x) → eq/ _ _ (x , b⊩x , a⊩x) }) + (asmX .⊩surjective x) }) + sq + + goal : invert ⊚ forward ≡ identityMorphism theSubQuotient + goal = AssemblyMorphism≡ _ _ (funExt pointwise) + isIso.ret (snd subQuotientCanonicalIso) = goal where + opaque + pointwise : ∀ x → (forward ⊚ invert) .map x ≡ x + pointwise x = + PT.elim + {P = + λ surj → + invert .map + (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj) + ≡ x} + (λ surj → asmX .isSetX _ _) + (λ { (a , a⊩x) → refl }) + (asmX .⊩surjective x) + + goal : forward ⊚ invert ≡ identityMorphism asmX + goal = AssemblyMorphism≡ _ _ (funExt pointwise) diff --git a/src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda b/src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda index 319a0ad..e492d8c 100644 --- a/src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda +++ b/src/Realizability/Modest/SubQuotientCanonicalPERToOriginal.agda @@ -56,24 +56,37 @@ module theSubQuotient = subQuotientAssembly theCanonicalPER invert : AssemblyMorphism theSubQuotient asmX - AssemblyMorphism.map invert sq = SQ.rec (asmX .isSetX) mainMap mainMapCoh sq where + AssemblyMorphism.map invert sq = SQ.rec (asmX .isSetX) reprAction reprActionCoh sq module Invert where - mainMap : Σ[ a ∈ A ] (theCanonicalPER .PER.relation a a) → X - mainMap (a , a~a) = PT.rec→Set (asmX .isSetX) action 2ConstantAction a~a where - action : Σ[ x ∈ X ] ((a ⊩[ asmX ] x) × (a ⊩[ asmX ] x)) → X - action (x , _ , _) = x + reprAction : Σ[ a ∈ A ] (a ~[ theCanonicalPER ] a) → X + reprAction (a , x , a⊩x , _) = x - 2ConstantAction : 2-Constant action - 2ConstantAction (x , a⊩x , _) (x' , a⊩x' , _) = isModestAsmX x x' a a⊩x a⊩x' + reprActionCoh : ∀ a b a~b → reprAction a ≡ reprAction b + reprActionCoh (a , x , a⊩x , _) (b , x' , b⊩x' , _) (x'' , a⊩x'' , b⊩x'') = + x + ≡⟨ isModestAsmX x x'' a a⊩x a⊩x'' ⟩ + x'' + ≡⟨ isModestAsmX x'' x' b b⊩x'' b⊩x' ⟩ + x' + ∎ + AssemblyMorphism.tracker invert = return (Id , (λ sq a a⊩sq → goal sq a a⊩sq)) where + realizability : (sq : subQuotient theCanonicalPER) → (a : A) → a ⊩[ theSubQuotient ] sq → a ⊩[ asmX ] (invert .map sq) + realizability sq a a⊩sq = + SQ.elimProp + {P = motive} + isPropMotive + elemMotive + sq a a⊩sq where + + motive : (sq : subQuotient theCanonicalPER) → Type _ + motive sq = ∀ (a : A) → a ⊩[ theSubQuotient ] sq → a ⊩[ asmX ] (invert .map sq) + + isPropMotive : ∀ sq → isProp (motive sq) + isPropMotive sq = isPropΠ2 λ a a⊩sq → asmX .⊩isPropValued _ _ + + elemMotive : (x : domain theCanonicalPER) → motive [ x ] + elemMotive (r , x , r⊩x , _) a (x' , a⊩x' , r⊩x') = subst (a ⊩[ asmX ]_) (isModestAsmX x' x r r⊩x' r⊩x) a⊩x' + + goal : (sq : subQuotient theCanonicalPER) → (a : A) → a ⊩[ theSubQuotient ] sq → (Id ⨾ a) ⊩[ asmX ] (invert .map sq) + goal sq a a⊩sq = subst (_⊩[ asmX ] _) (sym (Ida≡a a)) (realizability sq a a⊩sq) - mainMapCoh : ∀ a b a~b → mainMap a ≡ mainMap b - mainMapCoh (a , a~a) (b , b~b) a~b = - PT.elim3 - {P = λ a~a b~b a~b → mainMap (a , a~a) ≡ mainMap (b , b~b)} - (λ _ _ _ → asmX .isSetX _ _) - (λ { (x , a⊩x , _) (x' , b⊩x' , _) (x'' , a⊩x'' , b⊩x'') → - {!isModestAsmX x x' !} }) - a~a - b~b - a~b - AssemblyMorphism.tracker invert = {!!} diff --git a/src/Realizability/Modest/UnresizedGeneric.agda b/src/Realizability/Modest/UnresizedGeneric.agda index 4c97614..4330640 100644 --- a/src/Realizability/Modest/UnresizedGeneric.agda +++ b/src/Realizability/Modest/UnresizedGeneric.agda @@ -54,15 +54,15 @@ module Unresized theCanonicalPER x = canonicalPER (M . assemblies x) (M .isModestFamily x) elimRealizerForMx : ∀ (x : X) (Mx : M .carriers x) → Σ[ a ∈ A ] (a ⊩[ M .assemblies x ] Mx) → subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) - elimRealizerForMx x Mx (a , a⊩Mx) = [ a , (return (Mx , a⊩Mx , a⊩Mx)) ] + elimRealizerForMx x Mx (a , a⊩Mx) = [ a , Mx , a⊩Mx , a⊩Mx ] opaque elimRealizerForMx2Constant : ∀ x Mx → 2-Constant (elimRealizerForMx x Mx) elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx) = eq/ - (a , (return (Mx , a⊩Mx , a⊩Mx))) - (b , return (Mx , b⊩Mx , b⊩Mx)) - (return (Mx , a⊩Mx , b⊩Mx)) + (a , Mx , a⊩Mx , a⊩Mx) + (b , Mx , b⊩Mx , b⊩Mx) + (Mx , a⊩Mx , b⊩Mx) mainMapType : Type _ mainMapType = @@ -86,7 +86,7 @@ module Unresized ((λ { (c , c⊩Mx) → (elimRealizerForMx x Mx (c , c⊩Mx)) , (λ a a⊩x b b⊩Mx → - subst (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx x Mx (c , c⊩Mx))) (sym (λ*2ComputationRule (# zero) a b)) (return (Mx , b⊩Mx , c⊩Mx))) })) + subst (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx x Mx (c , c⊩Mx))) (sym (λ*2ComputationRule (# zero) a b)) (Mx , b⊩Mx , c⊩Mx)) })) (λ { (a , a⊩Mx) (b , b⊩Mx) → Σ≡Prop (λ out → isPropΠ4 λ a a⊩x b b⊩Mx → str (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero) ⨾ a ⨾ b) out)) (elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx)) }) (M .assemblies x .⊩surjective Mx) From a84a79ba81af3a524d645cffd4511ba8af20038e Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sun, 25 Aug 2024 19:31:51 +0530 Subject: [PATCH 58/61] Subquotient functor --- src/Realizability/PERs/SubQuotient.agda | 97 ++++++++++++++++++++++++- 1 file changed, 95 insertions(+), 2 deletions(-) diff --git a/src/Realizability/PERs/SubQuotient.agda b/src/Realizability/PERs/SubQuotient.agda index 0c12f24..ae9720f 100644 --- a/src/Realizability/PERs/SubQuotient.agda +++ b/src/Realizability/PERs/SubQuotient.agda @@ -67,8 +67,7 @@ module _ (λ [a] → isPropPropTrunc) (λ { (a , a~a) → return (a , a~a) }) [a] - - + isModestSubQuotientAssembly : isModest subQuotientAssembly isModestSubQuotientAssembly x y a a⊩x a⊩y = SQ.elimProp2 @@ -83,3 +82,97 @@ module _ isPropMotive : ∀ x y → isProp (motive x y) isPropMotive x y = isPropΠ3 λ _ _ _ → squash/ x y + +module _ (R S : PER) (f : perMorphism R S) where + + subQuotientAssemblyMorphism : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S) + subQuotientAssemblyMorphism = + SQ.rec + (isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)) + mainMap + mainMapCoherence + f where + + mainMap : perTracker R S → AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S) + AssemblyMorphism.map (mainMap (f , fIsTracker)) sqR = + SQ.rec + squash/ + mainMapRepr + mainMapReprCoherence + sqR module MainMapDefn where + mainMapRepr : domain R → subQuotient S + mainMapRepr (r , r~r) = [ f ⨾ r , fIsTracker r r r~r ] + + mainMapReprCoherence : (a b : domain R) → R .PER.relation (a .fst) (b .fst) → mainMapRepr a ≡ mainMapRepr b + mainMapReprCoherence (a , a~a) (b , b~b) a~b = eq/ _ _ (fIsTracker a b a~b) + + AssemblyMorphism.tracker (mainMap (f , fIsTracker)) = + do + return + (f , + (λ sqR s s⊩sqR → + SQ.elimProp + {P = + λ sqR + → ∀ (s : A) + → s ⊩[ subQuotientAssembly R ] sqR + → ⟨ subQuotientRealizability S (f ⨾ s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR) ⟩} + (λ sqR → isPropΠ2 λ s s⊩sqR → str (subQuotientRealizability S (f ⨾ s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR))) + (λ { (a , a~a) s s~a → fIsTracker s a s~a }) + sqR s s⊩sqR)) + + mainMapCoherence : (a b : perTracker R S) → isEquivTracker R S a b → mainMap a ≡ mainMap b + mainMapCoherence (a , a~a) (b , b~b) a~b = + AssemblyMorphism≡ _ _ + (funExt + λ sq → + SQ.elimProp + {P = + λ sq → + SQ.rec + squash/ + (MainMapDefn.mainMapRepr a a~a sq) + (MainMapDefn.mainMapReprCoherence a a~a sq) sq + ≡ + SQ.rec + squash/ + (MainMapDefn.mainMapRepr b b~b sq) + (MainMapDefn.mainMapReprCoherence b b~b sq) sq} + (λ sq → squash/ _ _) + (λ { (r , r~r) → eq/ _ _ (a~b r r~r) }) + sq) + + +subQuotientModestSet : PER → MOD .Category.ob +subQuotientModestSet R = subQuotient R , subQuotientAssembly R , isModestSubQuotientAssembly R + +subQuotientFunctor : Functor PERCat MOD +Functor.F-ob subQuotientFunctor R = subQuotientModestSet R +Functor.F-hom subQuotientFunctor {R} {S} f = subQuotientAssemblyMorphism R S f +Functor.F-id subQuotientFunctor {R} = + AssemblyMorphism≡ _ _ + (funExt + λ sqR → + SQ.elimProp + {P = λ sqR → subQuotientAssemblyMorphism R R (PERCat .Category.id {R}) .AssemblyMorphism.map sqR ≡ identityMorphism (subQuotientAssembly R) .AssemblyMorphism.map sqR} + (λ sqR → squash/ _ _) + (λ { (a , a~a) → + eq/ _ _ + (subst (_~[ R ] a) (sym (Ida≡a a)) a~a) }) + sqR) +Functor.F-seq subQuotientFunctor {R} {S} {T} f g = + AssemblyMorphism≡ _ _ + (funExt + λ sq → + SQ.elimProp3 + {P = λ sqR f g → + subQuotientAssemblyMorphism R T (seq' PERCat {R} {S} {T} f g) .AssemblyMorphism.map sqR ≡ + seq' MOD + {x = subQuotientModestSet R} + {y = subQuotientModestSet S} + {z = subQuotientModestSet T} + (subQuotientAssemblyMorphism R S f) (subQuotientAssemblyMorphism S T g) .AssemblyMorphism.map sqR} + (λ sq f g → squash/ _ _) + (λ { (a , a~a) (b , bIsTracker) (c , cIsTracker) → + eq/ _ _ (subst (_~[ T ] (c ⨾ (b ⨾ a))) (sym (λ*ComputationRule (` c ̇ (` b ̇ # zero)) a)) (cIsTracker (b ⨾ a) (b ⨾ a) (bIsTracker a a a~a))) }) + sq f g) From 02fc36027ff69f774907773092a931a9254d1246 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Thu, 19 Sep 2024 22:44:03 +0530 Subject: [PATCH 59/61] Render HTML --- docs/Categories.CartesianMorphism.html | 104 +++++ docs/Categories.GenericObject.html | 30 ++ docs/Cubical.Categories.Displayed.Base.html | 131 ++++++ ...ubical.Categories.Displayed.Reasoning.html | 110 +++++ ...ubical.HITs.SetCoequalizer.Properties.html | 14 +- docs/Realizability.ApplicativeStructure.html | 407 ++++++++++-------- docs/Realizability.Assembly.Base.html | 90 ++-- .../Realizability.Assembly.BinCoproducts.html | 361 ++++++++-------- docs/Realizability.Assembly.BinProducts.html | 305 +++++-------- docs/Realizability.Assembly.Coequalizers.html | 141 +++--- docs/Realizability.Assembly.Equalizers.html | 68 +-- docs/Realizability.Assembly.Everything.html | 3 +- docs/Realizability.Assembly.Exponentials.html | 245 ++++++----- docs/Realizability.Assembly.Morphism.html | 270 ++++++------ docs/Realizability.Assembly.SIP.html | 35 ++ ...ty.Assembly.SetsReflectiveSubcategory.html | 64 +++ docs/Realizability.Assembly.Terminal.html | 51 +++ docs/Realizability.CombinatoryAlgebra.html | 186 ++++---- docs/Realizability.Modest.Base.html | 87 ++++ docs/Realizability.Modest.CanonicalPER.html | 62 +++ docs/Realizability.Modest.Everything.html | 11 + ...ealizability.Modest.PartialSurjection.html | 390 +++++++++++++++++ ...ity.Modest.SubQuotientCanonicalPERIso.html | 149 +++++++ docs/Realizability.Modest.UniformFamily.html | 219 ++++++++++ ...zability.Modest.UniformFamilyCleavage.html | 102 +++++ ...Realizability.Modest.UnresizedGeneric.html | 94 ++++ docs/Realizability.PERs.Everything.html | 7 + docs/Realizability.PERs.PER.html | 226 ++++++++++ docs/Realizability.PERs.ResizedPER.html | 197 +++++++++ docs/Realizability.PERs.SubQuotient.html | 271 ++++++++++++ docs/Realizability.PropResizing.html | 86 +++- docs/Realizability.Topos.BinProducts.html | 242 +++++------ docs/Realizability.Topos.Equalizer.html | 196 ++++----- docs/Realizability.Topos.Everything.html | 21 +- ...ealizability.Topos.FunctionalRelation.html | 182 ++++---- .../Realizability.Topos.MonicReprFuncRel.html | 96 ++--- docs/Realizability.Topos.Object.html | 8 +- .../Realizability.Topos.ResizedPredicate.html | 10 +- docs/Realizability.Topos.StrictRelation.html | 186 ++++---- ...alizability.Topos.SubobjectClassifier.html | 396 ++++++++--------- docs/Realizability.Topos.TerminalObject.html | 28 +- ...Tripos.Prealgebra.Meets.Commutativity.html | 28 +- ...lity.Tripos.Prealgebra.Meets.Identity.html | 18 +- ...ripos.Prealgebra.Predicate.Properties.html | 126 +++--- docs/index.html | 5 +- src/Realizability/Modest/Everything.agda | 9 + .../Modest/PartialSurjection.agda | 9 - src/Realizability/PERs/Everything.agda | 5 + src/Realizability/PERs/SubQuotient.agda | 93 +++- .../Topos/SubobjectClassifier.agda | 30 -- src/index.agda | 3 + 51 files changed, 4354 insertions(+), 1853 deletions(-) create mode 100644 docs/Categories.CartesianMorphism.html create mode 100644 docs/Categories.GenericObject.html create mode 100644 docs/Cubical.Categories.Displayed.Base.html create mode 100644 docs/Cubical.Categories.Displayed.Reasoning.html create mode 100644 docs/Realizability.Assembly.SIP.html create mode 100644 docs/Realizability.Assembly.SetsReflectiveSubcategory.html create mode 100644 docs/Realizability.Assembly.Terminal.html create mode 100644 docs/Realizability.Modest.Base.html create mode 100644 docs/Realizability.Modest.CanonicalPER.html create mode 100644 docs/Realizability.Modest.Everything.html create mode 100644 docs/Realizability.Modest.PartialSurjection.html create mode 100644 docs/Realizability.Modest.SubQuotientCanonicalPERIso.html create mode 100644 docs/Realizability.Modest.UniformFamily.html create mode 100644 docs/Realizability.Modest.UniformFamilyCleavage.html create mode 100644 docs/Realizability.Modest.UnresizedGeneric.html create mode 100644 docs/Realizability.PERs.Everything.html create mode 100644 docs/Realizability.PERs.PER.html create mode 100644 docs/Realizability.PERs.ResizedPER.html create mode 100644 docs/Realizability.PERs.SubQuotient.html create mode 100644 src/Realizability/Modest/Everything.agda create mode 100644 src/Realizability/PERs/Everything.agda diff --git a/docs/Categories.CartesianMorphism.html b/docs/Categories.CartesianMorphism.html new file mode 100644 index 0000000..900f768 --- /dev/null +++ b/docs/Categories.CartesianMorphism.html @@ -0,0 +1,104 @@ + +Categories.CartesianMorphism
open import Cubical.Foundations.Prelude
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+
+module Categories.CartesianMorphism where
+
+module Contravariant
+  {ℓB ℓ'B ℓE ℓ'E}
+  {B : Category ℓB ℓ'B}
+  (E : Categoryᴰ B ℓE ℓ'E) where
+
+  open Category B
+  open Categoryᴰ E
+
+  opaque
+    isCartesian :
+      {a b : ob} (f : B [ a , b ])
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) 
+      Type (ℓ-max (ℓ-max ℓB ℓ'B) (ℓ-max ℓE ℓ'E))
+    isCartesian {a} {b} f {aᴰ} {bᴰ} fᴰ =
+       {c : ob} {cᴰ : ob[ c ]} (g : B [ c , a ])  isEquiv λ (gᴰ : Hom[ g ][ cᴰ , aᴰ ])  gᴰ ⋆ᴰ fᴰ
+
+  opaque
+    unfolding isCartesian
+    isPropIsCartesian :
+      {a b : ob} (f : B [ a , b ])
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) 
+      isProp (isCartesian f fᴰ)
+    isPropIsCartesian f fᴰ = isPropImplicitΠ2 λ c cᴰ  isPropΠ λ g  isPropIsEquiv (_⋆ᴰ fᴰ)
+
+  opaque
+    isCartesian' :
+      {a b : ob} (f : B [ a , b ])
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) 
+      Type (ℓ-max (ℓ-max ℓB ℓ'B) (ℓ-max ℓE ℓ'E))
+    isCartesian' {a} {b} f {aᴰ} {bᴰ} fᴰ =
+       {c : ob} {cᴰ : ob[ c ]} (g : B [ c , a ]) 
+        (∀ (hᴰ : Hom[ g  f ][ cᴰ , bᴰ ])  ∃![ gᴰ  Hom[ g ][ cᴰ , aᴰ ] ] gᴰ ⋆ᴰ fᴰ  hᴰ)
+
+  opaque
+    unfolding isCartesian'
+    isPropIsCartesian' :
+      {a b : ob} {f : B [ a , b ]}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) 
+      isProp (isCartesian' f fᴰ)
+    isPropIsCartesian' {a} {b} {f} {aᴰ} {bᴰ} fᴰ =
+      isPropImplicitΠ2 λ c cᴰ  isPropΠ2 λ g hᴰ  isPropIsContr
+
+  opaque
+    unfolding isCartesian
+    unfolding isCartesian'
+    isCartesian→isCartesian' :
+      {a b : ob} (f : B [ a , b ])
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) 
+      isCartesian f fᴰ 
+      isCartesian' f fᴰ
+    isCartesian→isCartesian' {a} {b} f {aᴰ} {bᴰ} fᴰ cartfᴰ g hᴰ =
+      ((invIsEq (cartfᴰ g) hᴰ) , (secIsEq (cartfᴰ g) hᴰ)) ,
+       { (gᴰ , gᴰ⋆fᴰ≡hᴰ)  cartfᴰ g .equiv-proof hᴰ .snd (gᴰ , gᴰ⋆fᴰ≡hᴰ) })
+
+  opaque
+    unfolding isCartesian
+    unfolding isCartesian'
+    isCartesian'→isCartesian :
+      {a b : ob} (f : B [ a , b ])
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) 
+      isCartesian' f fᴰ 
+      isCartesian f fᴰ
+    equiv-proof (isCartesian'→isCartesian {a} {b} f {aᴰ} {bᴰ} fᴰ cart'fᴰ g) hᴰ = (cart'fᴰ g hᴰ .fst) , (cart'fᴰ g hᴰ .snd)
+
+  isCartesian≃isCartesian' :
+    {a b : ob} (f : B [ a , b ])
+    {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+    (fᴰ : Hom[ f ][ aᴰ , bᴰ ]) 
+    isCartesian f fᴰ  isCartesian' f fᴰ
+  isCartesian≃isCartesian' {a} {b} f {aᴰ} {bᴰ} fᴰ =
+    propBiimpl→Equiv (isPropIsCartesian f fᴰ) (isPropIsCartesian' fᴰ) (isCartesian→isCartesian' f fᴰ) (isCartesian'→isCartesian f fᴰ)
+
+  cartesianLift : {a b : ob} (f : B [ a , b ]) (bᴰ : ob[ b ])  Type _
+  cartesianLift {a} {b} f bᴰ = Σ[ aᴰ  ob[ a ] ] Σ[ fᴰ  Hom[ f ][ aᴰ , bᴰ ] ] isCartesian f fᴰ
+
+  isCartesianFibration : Type _
+  isCartesianFibration =
+     {a b : ob} {bᴰ : ob[ b ]}
+     (f : Hom[ a , b ])
+      cartesianLift f bᴰ ∥₁
+
+  isPropIsCartesianFibration : isProp isCartesianFibration
+  isPropIsCartesianFibration = isPropImplicitΠ3 λ a b bᴰ  isPropΠ λ f  isPropPropTrunc
+
+  cleavage : Type _
+  cleavage = {a b : ob} (f : Hom[ a , b ]) (bᴰ : ob[ b ])  cartesianLift f bᴰ
+
\ No newline at end of file diff --git a/docs/Categories.GenericObject.html b/docs/Categories.GenericObject.html new file mode 100644 index 0000000..7ca3cbb --- /dev/null +++ b/docs/Categories.GenericObject.html @@ -0,0 +1,30 @@ + +Categories.GenericObject
open import Cubical.Foundations.Prelude
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Categories.CartesianMorphism
+
+module Categories.GenericObject where
+
+module _
+  {ℓB ℓ'B ℓE ℓ'E}
+  {B : Category ℓB ℓ'B}
+  (E : Categoryᴰ B ℓE ℓ'E) where
+
+  open Category B
+  open Categoryᴰ E
+  open Contravariant E
+
+  record GenericObject : Type (ℓ-max (ℓ-max ℓB ℓ'B) (ℓ-max ℓE ℓ'E)) where
+    constructor makeGenericObject
+    field
+      base : ob
+      displayed : ob[ base ]
+      isGeneric :
+         {t : ob} (tᴰ : ob[ t ])
+         Σ[ f  Hom[ t , base ] ] Σ[ fᴰ  Hom[ f ][ tᴰ , displayed ] ] isCartesian f fᴰ
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Displayed.Base.html b/docs/Cubical.Categories.Displayed.Base.html new file mode 100644 index 0000000..ef08d5e --- /dev/null +++ b/docs/Cubical.Categories.Displayed.Base.html @@ -0,0 +1,131 @@ + +Cubical.Categories.Displayed.Base
{-
+  Definition of a category displayed over another category.
+  Some definitions were guided by those at https://1lab.dev
+-}
+{-# OPTIONS --safe #-}
+module Cubical.Categories.Displayed.Base where
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Sigma
+open import Cubical.Categories.Category.Base
+
+private
+  variable
+    ℓC ℓC' ℓCᴰ ℓCᴰ' ℓDᴰ ℓDᴰ' : Level
+
+-- Displayed categories with hom-sets
+record Categoryᴰ (C : Category ℓC ℓC') ℓCᴰ ℓCᴰ' : Type (ℓ-suc (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓCᴰ ℓCᴰ'))) where
+  no-eta-equality
+  open Category C
+  field
+    ob[_] : ob  Type ℓCᴰ
+    Hom[_][_,_] : {x y : ob}  Hom[ x , y ]  ob[ x ]  ob[ y ]  Type ℓCᴰ'
+    idᴰ :  {x} {p : ob[ x ]}  Hom[ id ][ p , p ]
+    _⋆ᴰ_ :  {x y z} {f : Hom[ x , y ]} {g : Hom[ y , z ]} {xᴰ yᴰ zᴰ}
+       Hom[ f ][ xᴰ , yᴰ ]  Hom[ g ][ yᴰ , zᴰ ]  Hom[ f  g ][ xᴰ , zᴰ ]
+
+  infixr 9 _⋆ᴰ_
+  infixr 9 _∘ᴰ_
+
+  _≡[_]_ :  {x y xᴰ yᴰ} {f g : Hom[ x , y ]}  Hom[ f ][ xᴰ , yᴰ ]  f  g  Hom[ g ][ xᴰ , yᴰ ]  Type ℓCᴰ'
+  _≡[_]_ {x} {y} {xᴰ} {yᴰ} fᴰ p gᴰ = PathP  i  Hom[ p i ][ xᴰ , yᴰ ]) fᴰ gᴰ
+
+  infix 2 _≡[_]_
+
+  field
+    ⋆IdLᴰ :  {x y} {f : Hom[ x , y ]} {xᴰ yᴰ} (fᴰ : Hom[ f ][ xᴰ , yᴰ ])  idᴰ ⋆ᴰ fᴰ ≡[ ⋆IdL f ] fᴰ
+    ⋆IdRᴰ :  {x y} {f : Hom[ x , y ]} {xᴰ yᴰ} (fᴰ : Hom[ f ][ xᴰ , yᴰ ])  fᴰ ⋆ᴰ idᴰ ≡[ ⋆IdR f ] fᴰ
+    ⋆Assocᴰ :  {x y z w} {f : Hom[ x , y ]} {g : Hom[ y , z ]}  {h : Hom[ z , w ]} {xᴰ yᴰ zᴰ wᴰ}
+      (fᴰ : Hom[ f ][ xᴰ , yᴰ ]) (gᴰ : Hom[ g ][ yᴰ , zᴰ ]) (hᴰ : Hom[ h ][ zᴰ , wᴰ ])
+       (fᴰ ⋆ᴰ gᴰ) ⋆ᴰ hᴰ ≡[ ⋆Assoc f g h ] fᴰ ⋆ᴰ (gᴰ ⋆ᴰ hᴰ)
+    isSetHomᴰ :  {x y} {f : Hom[ x , y ]} {xᴰ yᴰ}  isSet Hom[ f ][ xᴰ , yᴰ ]
+
+  -- composition: alternative to diagramatic order
+  _∘ᴰ_ :  {x y z} {f : Hom[ x , y ]} {g : Hom[ y , z ]} {xᴰ yᴰ zᴰ}
+       Hom[ g ][ yᴰ , zᴰ ]  Hom[ f ][ xᴰ , yᴰ ]  Hom[ f  g ][ xᴰ , zᴰ ]
+  g ∘ᴰ f = f ⋆ᴰ g
+
+-- Helpful syntax/notation
+_[_][_,_] = Categoryᴰ.Hom[_][_,_]
+
+-- Total category of a displayed category
+module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where
+
+  open Category
+  open Categoryᴰ Cᴰ
+  private
+    module C = Category C
+
+  ∫C : Category (ℓ-max ℓC ℓCᴰ) (ℓ-max ℓC' ℓCᴰ')
+  ∫C .ob = Σ _ ob[_]
+  ∫C .Hom[_,_] (_ , xᴰ) (_ , yᴰ) = Σ _ Hom[_][ xᴰ , yᴰ ]
+  ∫C .id = _ , idᴰ
+  ∫C ._⋆_ (_ , fᴰ) (_ , gᴰ) = _ , fᴰ ⋆ᴰ gᴰ
+  ∫C .⋆IdL _ = ΣPathP (_ , ⋆IdLᴰ _)
+  ∫C .⋆IdR _ = ΣPathP (_ , ⋆IdRᴰ _)
+  ∫C .⋆Assoc _ _ _ = ΣPathP (_ , ⋆Assocᴰ _ _ _)
+  ∫C .isSetHom = isSetΣ C.isSetHom  _  isSetHomᴰ)
+
+-- Displayed total category, i.e. Σ for displayed categories
+module _ {C : Category ℓC ℓC'}
+  (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ')
+  (Dᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ')
+  where
+
+  open Categoryᴰ
+  private
+    module Cᴰ = Categoryᴰ Cᴰ
+    module Dᴰ = Categoryᴰ Dᴰ
+
+  ∫Cᴰ : Categoryᴰ C (ℓ-max ℓCᴰ ℓDᴰ) (ℓ-max ℓCᴰ' ℓDᴰ')
+  ∫Cᴰ .ob[_] x = Σ[ xᴰ  Cᴰ.ob[ x ] ] Dᴰ.ob[ x , xᴰ ]
+  ∫Cᴰ .Hom[_][_,_] f (_ , zᴰ) (_ , wᴰ) = Σ[ fᴰ  Cᴰ.Hom[ f ][ _ , _ ] ] Dᴰ.Hom[ f , fᴰ ][ zᴰ , wᴰ ]
+  ∫Cᴰ .idᴰ = Cᴰ.idᴰ , Dᴰ.idᴰ
+  ∫Cᴰ ._⋆ᴰ_ (_ , hᴰ) (_ , kᴰ) = _ , hᴰ Dᴰ.⋆ᴰ kᴰ
+  ∫Cᴰ .⋆IdLᴰ _ = ΣPathP (_ , Dᴰ.⋆IdLᴰ _)
+  ∫Cᴰ .⋆IdRᴰ _ = ΣPathP (_ , Dᴰ.⋆IdRᴰ _)
+  ∫Cᴰ .⋆Assocᴰ _ _ _ = ΣPathP (_ , Dᴰ.⋆Assocᴰ _ _ _)
+  ∫Cᴰ .isSetHomᴰ = isSetΣ Cᴰ.isSetHomᴰ  _  Dᴰ.isSetHomᴰ)
+
+module _ {C : Category ℓC ℓC'}
+  (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ')
+  (Dᴰ : Categoryᴰ C ℓDᴰ ℓDᴰ')
+  where
+
+  open Categoryᴰ
+  private
+    module Dᴰ = Categoryᴰ Dᴰ
+
+  weakenᴰ : Categoryᴰ (∫C Cᴰ) ℓDᴰ ℓDᴰ'
+  weakenᴰ .ob[_] (x , _) = Dᴰ.ob[ x ]
+  weakenᴰ .Hom[_][_,_] (f , _) = Dᴰ.Hom[ f ][_,_]
+  weakenᴰ .idᴰ = Dᴰ.idᴰ
+  weakenᴰ ._⋆ᴰ_ = Dᴰ._⋆ᴰ_
+  weakenᴰ .⋆IdLᴰ = Dᴰ.⋆IdLᴰ
+  weakenᴰ .⋆IdRᴰ = Dᴰ.⋆IdRᴰ
+  weakenᴰ .⋆Assocᴰ = Dᴰ.⋆Assocᴰ
+  weakenᴰ .isSetHomᴰ = Dᴰ.isSetHomᴰ
+
+module _ {C : Category ℓC ℓC'} (Cᴰ : Categoryᴰ C ℓCᴰ ℓCᴰ') where
+  open Category C
+  open Categoryᴰ Cᴰ
+
+  record isIsoᴰ {a b : ob} {f : C [ a , b ]} (f-isIso : isIso C f)
+    {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} (fᴰ : Hom[ f ][ aᴰ , bᴰ ])
+    : Type ℓCᴰ'
+    where
+    constructor isisoᴰ
+    open isIso f-isIso
+    field
+      invᴰ : Hom[ inv ][ bᴰ , aᴰ ]
+      secᴰ : invᴰ ⋆ᴰ fᴰ ≡[ sec ] idᴰ
+      retᴰ : fᴰ ⋆ᴰ invᴰ ≡[ ret ] idᴰ
+
+  CatIsoᴰ : {a b : ob}  CatIso C a b  ob[ a ]  ob[ b ]  Type ℓCᴰ'
+  CatIsoᴰ (f , f-isIso) aᴰ bᴰ = Σ[ fᴰ  Hom[ f ][ aᴰ , bᴰ ] ] isIsoᴰ f-isIso fᴰ
+
+  idᴰCatIsoᴰ : {x : ob} {xᴰ : ob[ x ]}  CatIsoᴰ idCatIso xᴰ xᴰ
+  idᴰCatIsoᴰ = idᴰ , isisoᴰ idᴰ (⋆IdLᴰ idᴰ) (⋆IdLᴰ idᴰ)
+
\ No newline at end of file diff --git a/docs/Cubical.Categories.Displayed.Reasoning.html b/docs/Cubical.Categories.Displayed.Reasoning.html new file mode 100644 index 0000000..6e7e7da --- /dev/null +++ b/docs/Cubical.Categories.Displayed.Reasoning.html @@ -0,0 +1,110 @@ + +Cubical.Categories.Displayed.Reasoning
{-# OPTIONS --safe #-}
+
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.GroupoidLaws
+open import Cubical.Foundations.Transport
+
+open import Cubical.Categories.Category.Base
+open import Cubical.Categories.Displayed.Base
+
+module Cubical.Categories.Displayed.Reasoning
+  {ℓC ℓ'C ℓCᴰ ℓ'Cᴰ : Level}
+  {C : Category ℓC ℓ'C}
+  (Cᴰ : Categoryᴰ C ℓCᴰ ℓ'Cᴰ)
+  where
+
+  open Categoryᴰ Cᴰ
+  private module C = Category C
+  open Category hiding (_∘_)
+
+  reind : {a b : C.ob} {f g : C [ a , b ]} (p : f  g)
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+     Hom[ f ][ aᴰ , bᴰ ]  Hom[ g ][ aᴰ , bᴰ ]
+  reind p = subst Hom[_][ _ , _ ] p
+
+  reind-filler : {a b : C.ob} {f g : C [ a , b ]} (p : f  g)
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (f : Hom[ f ][ aᴰ , bᴰ ])
+     f ≡[ p ] reind p f
+  reind-filler p = subst-filler Hom[_][ _ , _ ] p
+
+  reind-filler-sym : {a b : C.ob} {f g : C [ a , b ]} (p : f  g)
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      (f : Hom[ g ][ aᴰ , bᴰ ])
+     reind (sym p) f ≡[ p ] f
+  reind-filler-sym p = symP  reind-filler (sym p)
+
+  ≡[]-rectify : {a b : C.ob} {f g : C [ a , b ]} {p p' : f  g}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]}
+      {gᴰ : Hom[ g ][ aᴰ , bᴰ ]}
+     fᴰ ≡[ p ] gᴰ  fᴰ ≡[ p' ] gᴰ
+  ≡[]-rectify {fᴰ = fᴰ} {gᴰ = gᴰ} = subst (fᴰ ≡[_] gᴰ) (C.isSetHom _ _ _ _)
+
+  ≡[]∙ : {a b : C.ob} {f g h : C [ a , b ]}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]}
+      {gᴰ : Hom[ g ][ aᴰ , bᴰ ]}
+      {hᴰ : Hom[ h ][ aᴰ , bᴰ ]}
+      (p : f  g) (q : g  h)
+     fᴰ ≡[ p ] gᴰ
+     gᴰ ≡[ q ] hᴰ  fᴰ ≡[ p  q ] hᴰ
+  ≡[]∙ {fᴰ = fᴰ} {hᴰ = hᴰ} p q eq1 eq2 =
+    subst
+       p  PathP  i  p i) fᴰ hᴰ)
+      (sym $ congFunct Hom[_][ _ , _ ] p q)
+      (compPathP eq1 eq2)
+
+  infixr 30 ≡[]∙
+  syntax ≡[]∙ p q eq1 eq2 = eq1 [ p ]∙[ q ] eq2
+
+  ≡[]⋆ : {a b c : C.ob} {f g : C [ a , b ]} {h i : C [ b , c ]}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} {cᴰ : ob[ c ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]}
+      {gᴰ : Hom[ g ][ aᴰ , bᴰ ]}
+      {hᴰ : Hom[ h ][ bᴰ , cᴰ ]}
+      {iᴰ : Hom[ i ][ bᴰ , cᴰ ]}
+      (p : f  g) (q : h  i)
+     fᴰ ≡[ p ] gᴰ  hᴰ ≡[ q ] iᴰ  fᴰ ⋆ᴰ hᴰ ≡[ cong₂ C._⋆_ p q ] gᴰ ⋆ᴰ iᴰ
+  ≡[]⋆ _ _ = congP₂  _  _⋆ᴰ_)
+
+  infixr 30 ≡[]⋆
+  syntax ≡[]⋆ p q eq1 eq2 = eq1 [ p ]⋆[ q ] eq2
+
+  reind-rectify : {a b : C.ob} {f g : C [ a , b ]} {p p' : f  g}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]}
+     reind p fᴰ  reind p' fᴰ
+  reind-rectify {fᴰ = fᴰ} = cong  p  reind p fᴰ) (C.isSetHom _ _ _ _)
+
+  reind-contractʳ : {a b c : C.ob} {f g : C [ a , b ]} {h : C [ b , c ]}
+      {p : f  g}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} {cᴰ : ob[ c ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]} {hᴰ : Hom[ h ][ bᴰ , cᴰ ]}
+     reind (cong (C._⋆ h) p) (fᴰ ⋆ᴰ hᴰ)  reind p fᴰ ⋆ᴰ hᴰ
+  reind-contractʳ {hᴰ = hᴰ} = fromPathP $
+    congP  _  _⋆ᴰ hᴰ) (transport-filler _ _)
+
+  reind-comp : {a b : C.ob} {f g h : C [ a , b ]} {p : f  g} {q : g  h}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]}
+     reind (p  q) fᴰ  reind q (reind p fᴰ)
+  reind-comp = substComposite Hom[_][ _ , _ ] _ _ _
+
+  reind-contractˡ : {a b c : C.ob} {f : C [ a , b ]} {g h : C [ b , c ]}
+      {p : g  h}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]} {cᴰ : ob[ c ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]} {gᴰ : Hom[ g ][ bᴰ , cᴰ ]}
+     reind (cong (f C.⋆_) p) (fᴰ ⋆ᴰ gᴰ)  fᴰ ⋆ᴰ reind p gᴰ
+  reind-contractˡ {fᴰ = fᴰ} = fromPathP $
+    congP  _  fᴰ ⋆ᴰ_) (transport-filler _ _)
+
+  ≡→≡[] : {a b : C.ob} {f g : C [ a , b ]} {p : f  g}
+      {aᴰ : ob[ a ]} {bᴰ : ob[ b ]}
+      {fᴰ : Hom[ f ][ aᴰ , bᴰ ]}
+      {gᴰ : Hom[ g ][ aᴰ , bᴰ ]}
+     reind p fᴰ  gᴰ  fᴰ ≡[ p ] gᴰ
+  ≡→≡[] = toPathP
+
\ No newline at end of file diff --git a/docs/Cubical.HITs.SetCoequalizer.Properties.html b/docs/Cubical.HITs.SetCoequalizer.Properties.html index 5c15678..1422113 100644 --- a/docs/Cubical.HITs.SetCoequalizer.Properties.html +++ b/docs/Cubical.HITs.SetCoequalizer.Properties.html @@ -24,21 +24,21 @@ -- Some helpful lemmas, similar to those in Cubical/HITs/SetQuotients/Properties.agda elimProp : {f g : A B} {C : SetCoequalizer f g Type } - (Cprop : (x : SetCoequalizer f g) isProp (C x)) + (Cprop : (x : SetCoequalizer f g) isProp (C x)) (Cinc : (b : B) C (inc b)) (x : SetCoequalizer f g) C x elimProp Cprop Cinc (inc x) = Cinc x elimProp {f = f} {g = g} Cprop Cinc (coeq a i) = - isProp→PathP i Cprop (coeq a i)) (Cinc (f a)) (Cinc (g a)) i + isProp→PathP i Cprop (coeq a i)) (Cinc (f a)) (Cinc (g a)) i elimProp Cprop Cinc (squash x y p q i j) = - isOfHLevel→isOfHLevelDep - 2 x isProp→isSet (Cprop x)) (g x) (g y) (cong g p) (cong g q) (squash x y p q) i j + isOfHLevel→isOfHLevelDep + 2 x isProp→isSet (Cprop x)) (g x) (g y) (cong g p) (cong g q) (squash x y p q) i j where g = elimProp Cprop Cinc elimProp2 : {A' : Type } {B' : Type ℓ'} {f g : A B} {f' g' : A' B'} {C : SetCoequalizer f g SetCoequalizer f' g' Type (ℓ-max ℓ')} - (Cprop : (x : SetCoequalizer f g) (y : SetCoequalizer f' g') isProp (C x y)) + (Cprop : (x : SetCoequalizer f g) (y : SetCoequalizer f' g') isProp (C x y)) (Cinc : (b : B) (b' : B') C (inc b) (inc b')) (x : SetCoequalizer f g) (y : SetCoequalizer f' g') C x y elimProp2 Cprop Cinc = elimProp x isPropΠ y Cprop x y)) @@ -51,7 +51,7 @@ (Cprop : (x : SetCoequalizer f g) (y : SetCoequalizer f' g') (z : SetCoequalizer f'' g'') - isProp (C x y z)) + isProp (C x y z)) (Cinc : (b : B) (b' : B') (b'' : B'') C (inc b) (inc b') (inc b'')) (x : SetCoequalizer f g) (y : SetCoequalizer f' g') (z : SetCoequalizer f'' g'') C x y z @@ -77,7 +77,7 @@ SetCoequalizer f g SetCoequalizer f' g' C rec2 Cset h hcoeqsl hcoeqsr = rec - (isSetΠ _ Cset)) + (isSetΠ _ Cset)) b rec Cset b' h b b') a' hcoeqsr a' b)) a funExt (elimProp _ Cset _ _) b' hcoeqsl a b'))) diff --git a/docs/Realizability.ApplicativeStructure.html b/docs/Realizability.ApplicativeStructure.html index 8499598..7fcf1d3 100644 --- a/docs/Realizability.ApplicativeStructure.html +++ b/docs/Realizability.ApplicativeStructure.html @@ -1,173 +1,238 @@ -Realizability.ApplicativeStructure
open import Cubical.Core.Everything
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Relation.Nullary
-open import Cubical.Data.Nat
-open import Cubical.Data.Nat.Order
-open import Cubical.Data.FinData
-open import Cubical.Data.Vec
-open import Cubical.Data.Empty renaming (elim to ⊥elim)
-open import Cubical.Tactics.NatSolver
-
-module Realizability.ApplicativeStructure where
-
-private module _ {} {A : Type } where
-  -- Taken from Data.Vec.Base from agda-stdlib
-  foldlOp :  {ℓ'} (B :   Type ℓ')  Type _
-  foldlOp B =  {n}  B n  A  B (suc n)
-
-  opaque
-    foldl :  {ℓ'} {n : } (B :   Type ℓ')  foldlOp B  B zero  Vec A n  B n
-    foldl {ℓ'} {.zero} B op acc emptyVec = acc
-    foldl {ℓ'} {.(suc _)} B op acc (x ∷vec vec) = foldl  n  B (suc n)) op (op acc x) vec
-
-  opaque
-    reverse :  {n}  Vec A n  Vec A n
-    reverse vec = foldl  n  Vec A n)  acc curr  curr  acc) [] vec
-
-  opaque
-    chain :  {n}  Vec A (suc n)  (A  A  A)  A
-    chain {n} (x ∷vec vec) op = foldl  _  A)  acc curr  op acc curr) x vec
-
-record ApplicativeStructure {} (A : Type ) : Type  where
-  infixl 20 _⨾_
-  field
-    isSetA : isSet A
-    _⨾_ : A  A  A
-
-module _ {} {A : Type } (as : ApplicativeStructure A) where
-  open ApplicativeStructure as
-  infix 23 `_
-  infixl 22 _̇_
-  data Term (n : ) : Type  where
-    # : Fin n  Term n
-    `_ : A  Term n
-    _̇_ : Term n  Term n  Term n
-
-  ⟦_⟧ :  {n}  Term n  Vec A n  A
-  ⟦_⟧ (` a) _ = a
-  ⟦_⟧ {n} (# k) subs = lookup k subs
-  ⟦_⟧ (a ̇ b) subs = ( a  subs)  ( b  subs)
-
-  applicationChain :  {n m}  Vec (Term m) (suc n)  Term m
-  applicationChain {n} {m} vec = chain vec  x y  x ̇ y)
-
-  apply :  {n}  A  Vec A n  A
-  apply {n} a vec = chain (a  vec)  x y  x  y)
-  
-  private
-    opaque
-      unfolding reverse
-      unfolding foldl
-      unfolding chain
-      applyWorks :  K a b  apply K (a  b  [])  K  a  b
-      applyWorks K a b = refl
-
-  record Feferman : Type  where
-    field
-      s : A
-      k : A
-      kab≡a :  a b  k  a  b  a
-      sabc≡ac_bc :  a b c  s  a  b  c  (a  c)  (b  c)
-      
-  module ComputationRules (feferman : Feferman) where
-    open Feferman feferman
-
-    opaque
-      λ*abst :  {n}  (e : Term (suc n))  Term n
-      λ*abst {n} (# zero) = ` s ̇ ` k ̇ ` k
-      λ*abst {n} (# (suc x)) = ` k ̇ # x
-      λ*abst {n} (` x) = ` k ̇ ` x
-      λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁
-
-    -- Some shortcuts
-
-    λ* : Term one  A
-    λ* t =  λ*abst t  []
-
-    λ*2 : Term two  A
-    λ*2 t =  λ*abst (λ*abst t)  []
-
-    λ*3 : Term three  A
-    λ*3 t =  λ*abst (λ*abst (λ*abst t))  []
-
-    λ*4 : Term four  A
-    λ*4 t =  λ*abst (λ*abst (λ*abst (λ*abst t)))  []
-
-    opaque
-      unfolding λ*abst
-      βreduction :  {n}  (body : Term (suc n))  (prim : A)  (subs : Vec A n)   λ*abst body ̇ ` prim  subs   body  (prim  subs)
-      βreduction {n} (# zero) prim subs =
-        s  k  k  prim
-          ≡⟨ sabc≡ac_bc _ _ _ 
-        k  prim  (k  prim)
-          ≡⟨ kab≡a _ _ 
-        prim
-          
-      βreduction {n} (# (suc x)) prim subs = kab≡a _ _
-      βreduction {n} (` x) prim subs = kab≡a _ _
-      βreduction {n} (rator ̇ rand) prim subs =
-        s   λ*abst rator  subs   λ*abst rand  subs  prim
-          ≡⟨ sabc≡ac_bc _ _ _ 
-         λ*abst rator  subs  prim  ( λ*abst rand  subs  prim)
-          ≡⟨ cong₂  x y  x  y) (βreduction rator prim subs) (βreduction rand prim subs) 
-         rator  (prim  subs)   rand  (prim  subs)
-          
-
-    λ*chainTerm :  n  Term n  Term zero
-    λ*chainTerm zero t = t
-    λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t)
-
-    λ*chain :  {n}  Term n  A
-    λ*chain {n} t =  λ*chainTerm n t  []
-
-    opaque
-      unfolding reverse
-      unfolding foldl
-      unfolding chain
-
-      λ*ComputationRule :  (t : Term 1) (a : A)  λ* t  a   t  (a  [])
-      λ*ComputationRule t a =
-         λ*abst t  []  a
-          ≡⟨ βreduction t a [] 
-         t  (a  [])
-          
-
-      λ*2ComputationRule :  (t : Term 2) (a b : A)  λ*2 t  a  b   t  (b  a  [])
-      λ*2ComputationRule t a b =
-         λ*abst (λ*abst t)  []  a  b
-          ≡⟨ refl 
-         λ*abst (λ*abst t)  []   ` a  []   ` b  []
-          ≡⟨ refl 
-         λ*abst (λ*abst t) ̇ ` a  []   ` b  []
-          ≡⟨ cong  x  x  b) (βreduction (λ*abst t) a []) 
-         λ*abst t  (a  [])  b
-          ≡⟨ βreduction t b (a  []) 
-         t  (b  a  [])
-          
+Realizability.ApplicativeStructure
---
+title: Applicative Structure
+author: Rahul Chhabra
+---
+# Applicative Structure
+
+In this module we define the notion of an _applicative structure_.
+
+A type $A$ has applicative structure if it has an "application operation" (here represented by `_⨾_`) and is a set.
+
+<details>
+```agda
+open import Cubical.Core.Everything
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Relation.Nullary
+open import Cubical.Data.Nat
+open import Cubical.Data.Nat.Order
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Empty renaming (elim to ⊥elim)
+open import Cubical.Tactics.NatSolver
+
+module Realizability.ApplicativeStructure where
+
+private module _ {} {A : Type } where
+  -- Taken from Data.Vec.Base from agda-stdlib
+  foldlOp :  {ℓ'} (B :   Type ℓ')  Type _
+  foldlOp B =  {n}  B n  A  B (suc n)
+
+  opaque
+    foldl :  {ℓ'} {n : } (B :   Type ℓ')  foldlOp B  B zero  Vec A n  B n
+    foldl {ℓ'} {.zero} B op acc emptyVec = acc
+    foldl {ℓ'} {.(suc _)} B op acc (x ∷vec vec) = foldl  n  B (suc n)) op (op acc x) vec
+
+  opaque
+    reverse :  {n}  Vec A n  Vec A n
+    reverse vec = foldl  n  Vec A n)  acc curr  curr  acc) [] vec
+
+  opaque
+    chain :  {n}  Vec A (suc n)  (A  A  A)  A
+    chain {n} (x ∷vec vec) op = foldl  _  A)  acc curr  op acc curr) x vec
+```
+</details>
+
+```agda
+record ApplicativeStructure {} (A : Type ) : Type  where
+  infixl 20 _⨾_
+  field
+    isSetA : isSet A
+    _⨾_ : A  A  A
+```
+
+Since being a set is a property - we will generally not have to think about it too much. Also, since `A` is a set - we can drop the relevance of paths and simply talk about "equality".
+
+We can define the notion of a term over an applicative structure.
+```agda
+module _ {} {A : Type } (as : ApplicativeStructure A) where
+  open ApplicativeStructure as
+  infix 23 `_
+  infixl 22 _̇_
+  data Term (n : ) : Type  where
+    # : Fin n  Term n
+    `_ : A  Term n
+    _̇_ : Term n  Term n  Term n
+```
+
+These terms can be evaluated into $A$ if we can give the values of the free variables.
+
+```agda
+  ⟦_⟧ :  {n}  Term n  Vec A n  A
+  ⟦_⟧ (` a) _ = a
+  ⟦_⟧ {n} (# k) subs = lookup k subs
+  ⟦_⟧ (a ̇ b) subs = ( a  subs)  ( b  subs)
+
+  applicationChain :  {n m}  Vec (Term m) (suc n)  Term m
+  applicationChain {n} {m} vec = chain vec  x y  x ̇ y)
+
+  apply :  {n}  A  Vec A n  A
+  apply {n} a vec = chain (a  vec)  x y  x  y)
+```
+
+<details>
+```agda
+  private
+    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+      applyWorks :  K a b  apply K (a  b  [])  K  a  b
+      applyWorks K a b = refl
+```
+</details>
+
+On an applicative structure we can define Feferman structure (or SK structure). We call an applicative structure endowed with Feferman structure a **combinatory algebra**.
+
+```agda
+  record Feferman : Type  where
+    field
+      s : A
+      k : A
+      kab≡a :  a b  k  a  b  a
+      sabc≡ac_bc :  a b c  s  a  b  c  (a  c)  (b  c)
+```
+
+Feferman structure allows us to construct **combinatorial completeness** structure.
+
+Imagine we have a term `t : Term n` (for some `n : ℕ`). We can ask if `A` has a "copy" of `t` so that application would correspond to subsitution. That is, we may ask if we can find an `a : A` such that
+`a ⨾ < a¹ ⨾ a² ⨾ .... ⨾ aⁿ >` (here the `< ... >` notation represents a chain of applications) would be equal to `t [a¹ / # 0 , a² / # 1 , .... , aⁿ / # (pred n) ]`. If the applicative structure additionally can be endowed with Feferman structure - then the answer is yes. 
+
+To get to such a term, we first need to define a function that takes `Term (suc n)` to `Term n` by "abstracting" the free variable represented by the index `# 0`.
+
+We will call this `λ*abst` and this will turn out to behave very similar to λ abstraction - and we will also show that it validates a kind of β reduction rule.
+
+```agda
+  module ComputationRules (feferman : Feferman) where
+    open Feferman feferman
+
+    opaque
+      λ*abst :  {n}  (e : Term (suc n))  Term n
+      λ*abst {n} (# zero) = ` s ̇ ` k ̇ ` k
+      λ*abst {n} (# (suc x)) = ` k ̇ # x
+      λ*abst {n} (` x) = ` k ̇ ` x
+      λ*abst {n} (e ̇ e₁) = ` s ̇ λ*abst e ̇ λ*abst e₁
+```
+
+**Remark** : It is important to note that in general, realizability is developed using **partial combinatory algebras** and **partial applicative structures**. In this case, `λ*abst` is not particularly well-behaved. The β reduction-esque rule we derive also does not behave *completely* like β reduction. See Jonh Longley's PhD thesis "Realizability Toposes and Language Semantics" Theorem 1.1.9.
+
+**Remark** : We declare the definition as `opaque` - this is important. If we let Agda unfold this definition all the way we ocassionally end up with unreadable goals containing a mess of `s` and `k`. 
+
+We define meta-syntactic sugar for some of the more common cases :
+
+```agda
+    λ* : Term one  A
+    λ* t =  λ*abst t  []
+
+    λ*2 : Term two  A
+    λ*2 t =  λ*abst (λ*abst t)  []
+
+    λ*3 : Term three  A
+    λ*3 t =  λ*abst (λ*abst (λ*abst t))  []
+
+    λ*4 : Term four  A
+    λ*4 t =  λ*abst (λ*abst (λ*abst (λ*abst t)))  []
+```
+
+We now show that we have a β-reduction-esque operation. We proceed by induction on the structure of the term and the number of free variables.
+
+For the particular combinatory algebra Λ/β (terms of the untyped λ calculus quotiented by β equality) - this β reduction actually coincides with the "actual" β reduction.
+TODO : Prove this.
+
+```agda
+    opaque
+      unfolding λ*abst
+      βreduction :  {n}  (body : Term (suc n))  (prim : A)  (subs : Vec A n)   λ*abst body ̇ ` prim  subs   body  (prim  subs)
+      βreduction {n} (# zero) prim subs =
+        s  k  k  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+        k  prim  (k  prim)
+          ≡⟨ kab≡a _ _ 
+        prim
+          
+      βreduction {n} (# (suc x)) prim subs = kab≡a _ _
+      βreduction {n} (` x) prim subs = kab≡a _ _
+      βreduction {n} (rator ̇ rand) prim subs =
+        s   λ*abst rator  subs   λ*abst rand  subs  prim
+          ≡⟨ sabc≡ac_bc _ _ _ 
+         λ*abst rator  subs  prim  ( λ*abst rand  subs  prim)
+          ≡⟨ cong₂  x y  x  y) (βreduction rator prim subs) (βreduction rand prim subs) 
+         rator  (prim  subs)   rand  (prim  subs)
+          
+```
+
+<details>
+```agda
+    λ*chainTerm :  n  Term n  Term zero
+    λ*chainTerm zero t = t
+    λ*chainTerm (suc n) t = λ*chainTerm n (λ*abst t)
+
+    λ*chain :  {n}  Term n  A
+    λ*chain {n} t =  λ*chainTerm n t  []
+```
+</details>
+
+We provide useful reasoning combinators that are useful and frequent.
+
+```agda
+    opaque
+      unfolding reverse
+      unfolding foldl
+      unfolding chain
+
+      λ*ComputationRule :  (t : Term 1) (a : A)  λ* t  a   t  (a  [])
+      λ*ComputationRule t a =
+         λ*abst t  []  a
+          ≡⟨ βreduction t a [] 
+         t  (a  [])
+          
+
+      λ*2ComputationRule :  (t : Term 2) (a b : A)  λ*2 t  a  b   t  (b  a  [])
+      λ*2ComputationRule t a b =
+         λ*abst (λ*abst t)  []  a  b
+          ≡⟨ refl 
+         λ*abst (λ*abst t)  []   ` a  []   ` b  []
+          ≡⟨ refl 
+         λ*abst (λ*abst t) ̇ ` a  []   ` b  []
+          ≡⟨ cong  x  x  b) (βreduction (λ*abst t) a []) 
+         λ*abst t  (a  [])  b
+          ≡⟨ βreduction t b (a  []) 
+         t  (b  a  [])
+          
           
-      λ*3ComputationRule :  (t : Term 3) (a b c : A)  λ*3 t  a  b  c   t  (c  b  a  [])
-      λ*3ComputationRule t a b c =
-         λ*abst (λ*abst (λ*abst t))  []   ` a  []   ` b  []   ` c  []
-          ≡⟨ cong  x  x  b  c) (βreduction (λ*abst (λ*abst t)) a []) 
-         λ*abst (λ*abst t)  (a  [])   ` b  (a  [])   ` c  (a  [])
-          ≡⟨ cong  x  x  c) (βreduction (λ*abst t) b (a  [])) 
-         λ*abst t  (b  a  [])   ` c  (b  a  [])
-          ≡⟨ βreduction t c (b  a  []) 
-         t  (c  b  a  [])
-          
-
-      λ*4ComputationRule :  (t : Term 4) (a b c d : A)  λ*4 t  a  b  c  d   t  (d  c  b  a  [])
-      λ*4ComputationRule t a b c d =
-         λ*abst (λ*abst (λ*abst (λ*abst t)))  []   ` a  []   ` b  []   ` c  []   ` d  []
-          ≡⟨ cong  x  x  b  c  d) (βreduction (λ*abst (λ*abst (λ*abst t))) a []) 
-         λ*abst (λ*abst (λ*abst t))  (a  [])   ` b  (a  [])   ` c  (a  [])   ` d  (a  [])
-          ≡⟨ cong  x  x  c  d) (βreduction (λ*abst (λ*abst t)) b (a  [])) 
-         λ*abst (λ*abst t)  (b  a  [])   ` c  (b  a  [])   ` d  (b  a  [])
-          ≡⟨ cong  x  x  d) (βreduction (λ*abst t) c (b  a  [])) 
-         λ*abst t  (c  b  a  [])   ` d  (c  b  a  [])
-          ≡⟨ βreduction t d (c  b  a  []) 
-         t  (d  c  b  a  [])
-          
-
\ No newline at end of file + λ*3ComputationRule : (t : Term 3) (a b c : A) λ*3 t a b c t (c b a []) + λ*3ComputationRule t a b c = + λ*abst (λ*abst (λ*abst t)) [] ` a [] ` b [] ` c [] + ≡⟨ cong x x b c) (βreduction (λ*abst (λ*abst t)) a []) + λ*abst (λ*abst t) (a []) ` b (a []) ` c (a []) + ≡⟨ cong x x c) (βreduction (λ*abst t) b (a [])) + λ*abst t (b a []) ` c (b a []) + ≡⟨ βreduction t c (b a []) + t (c b a []) + + + λ*4ComputationRule : (t : Term 4) (a b c d : A) λ*4 t a b c d t (d c b a []) + λ*4ComputationRule t a b c d = + λ*abst (λ*abst (λ*abst (λ*abst t))) [] ` a [] ` b [] ` c [] ` d [] + ≡⟨ cong x x b c d) (βreduction (λ*abst (λ*abst (λ*abst t))) a []) + λ*abst (λ*abst (λ*abst t)) (a []) ` b (a []) ` c (a []) ` d (a []) + ≡⟨ cong x x c d) (βreduction (λ*abst (λ*abst t)) b (a [])) + λ*abst (λ*abst t) (b a []) ` c (b a []) ` d (b a []) + ≡⟨ cong x x d) (βreduction (λ*abst t) c (b a [])) + λ*abst t (c b a []) ` d (c b a []) + ≡⟨ βreduction t d (c b a []) + t (d c b a []) + +``` +
\ No newline at end of file diff --git a/docs/Realizability.Assembly.Base.html b/docs/Realizability.Assembly.Base.html index 79b7f6d..92ea81c 100644 --- a/docs/Realizability.Assembly.Base.html +++ b/docs/Realizability.Assembly.Base.html @@ -3,42 +3,70 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Structure -open import Cubical.Data.Sigma -open import Cubical.HITs.PropositionalTruncation -open import Cubical.Reflection.RecordEquiv -open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Univalence +open import Cubical.Foundations.Isomorphism +open import Cubical.Functions.FunExtEquiv +open import Cubical.Data.Sigma +open import Cubical.HITs.PropositionalTruncation +open import Cubical.Reflection.RecordEquiv +open import Realizability.CombinatoryAlgebra -module Realizability.Assembly.Base {} {A : Type } (ca : CombinatoryAlgebra A) where - record Assembly (X : Type ) : Type (ℓ-suc ) where - infix 25 _⊩_ - field - isSetX : isSet X - _⊩_ : A X Type - ⊩isPropValued : a x isProp (a x) - ⊩surjective : x ∃[ a A ] a x +module Realizability.Assembly.Base {} {A : Type } (ca : CombinatoryAlgebra A) where + record Assembly (X : Type ) : Type (ℓ-suc ) where + constructor makeAssembly + infix 25 _⊩_ + field + _⊩_ : A X Type + isSetX : isSet X + ⊩isPropValued : a x isProp (a x) + ⊩surjective : x ∃[ a A ] a x + open Assembly public + _⊩[_]_ : {X : Type } A Assembly X X Type + a ⊩[ A ] x = A ._⊩_ a x - AssemblyΣ : Type Type _ - AssemblyΣ X = - Σ[ isSetX isSet X ] - Σ[ _⊩_ (A X hProp ) ] - (∀ x ∃[ a A ] a x ) + AssemblyΣ : Type Type _ + AssemblyΣ X = + Σ[ _⊩_ (A X hProp ) ] + (∀ x ∃[ a A ] a x ) × + (isSet X) - AssemblyΣX→isSetX : X AssemblyΣ X isSet X - AssemblyΣX→isSetX X (isSetX , _ , _) = isSetX - - AssemblyΣX→⊩ : X AssemblyΣ X (A X hProp ) - AssemblyΣX→⊩ X (_ , , _) = - - AssemblyΣX→⊩surjective : X (asm : AssemblyΣ X) (∀ x ∃[ a A ] AssemblyΣX→⊩ X asm a x ) - AssemblyΣX→⊩surjective X (_ , _ , ⊩surjective) = ⊩surjective - - isSetAssemblyΣ : X isSet (AssemblyΣ X) - isSetAssemblyΣ X = isSetΣ (isProp→isSet isPropIsSet) λ isSetX isSetΣ (isSetΠ a isSetΠ λ x isSetHProp)) λ _⊩_ isSetΠ λ x isProp→isSet isPropPropTrunc + isSetAssemblyΣ : X isSet (AssemblyΣ X) + isSetAssemblyΣ X = isSetΣ (isSetΠ2 λ _ _ isSetHProp) rel isSet× (isSetΠ λ x isProp→isSet isPropPropTrunc) (isProp→isSet isPropIsSet)) - unquoteDecl AssemblyIsoΣ = declareRecordIsoΣ AssemblyIsoΣ (quote Assembly) + AssemblyΣ≡Equiv : X (a b : AssemblyΣ X) (a b) (∀ r x a .fst r x b .fst r x ) + AssemblyΣ≡Equiv X a b = + a b + ≃⟨ invEquiv (Σ≡PropEquiv rel isProp× (isPropΠ λ x isPropPropTrunc) isPropIsSet) {u = a} {v = b}) + a .fst b .fst + ≃⟨ invEquiv (funExt₂Equiv {f = a .fst} {g = b .fst}) + (∀ (r : A) (x : X) a .fst r x b .fst r x) + ≃⟨ + equivΠCod + r + equivΠCod + λ x + compEquiv + (invEquiv (Σ≡PropEquiv _ isPropIsProp) {u = a .fst r x} {v = b .fst r x})) + (univalence {A = a .fst r x .fst} {B = b .fst r x .fst})) + + (∀ (r : A) (x : X) a .fst r x b .fst r x ) + - open Assembly public + -- definitional isomorphism + AssemblyΣIsoAssembly : X Iso (AssemblyΣ X) (Assembly X) + _⊩_ (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) a x = rel a x + Assembly.isSetX (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) = isSetX + ⊩isPropValued (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) a x = str (rel a x) + ⊩surjective (Iso.fun (AssemblyΣIsoAssembly X) (rel , surj , isSetX)) x = surj x + Iso.inv (AssemblyΣIsoAssembly X) asm = a x (a ⊩[ asm ] x) , (asm .⊩isPropValued a x)) , x asm .⊩surjective x) , asm .isSetX + Iso.rightInv (AssemblyΣIsoAssembly X) asm = refl + Iso.leftInv (AssemblyΣIsoAssembly X) (rel , surj , isSetX) = refl - + AssemblyΣ≃Assembly : X AssemblyΣ X Assembly X + AssemblyΣ≃Assembly X = isoToEquiv (AssemblyΣIsoAssembly X) + + isSetAssembly : X isSet (Assembly X) + isSetAssembly X = isOfHLevelRespectEquiv 2 (AssemblyΣ≃Assembly X) (isSetAssemblyΣ X)
\ No newline at end of file diff --git a/docs/Realizability.Assembly.BinCoproducts.html b/docs/Realizability.Assembly.BinCoproducts.html index 2dc0932..b6dc069 100644 --- a/docs/Realizability.Assembly.BinCoproducts.html +++ b/docs/Realizability.Assembly.BinCoproducts.html @@ -4,197 +4,196 @@ open import Cubical.Foundations.HLevels open import Cubical.Data.Sum hiding (map) open import Cubical.Data.Sigma -open import Cubical.Data.Fin -open import Cubical.Data.Nat -open import Cubical.Data.Vec hiding (map) -open import Cubical.HITs.PropositionalTruncation hiding (map) -open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.Categories.Category -open import Cubical.Categories.Limits.BinCoproduct -open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-chain to `λ*; λ*-naturality to `λ*ComputationRule) hiding (λ*) +open import Cubical.Data.FinData +open import Cubical.Data.Nat +open import Cubical.Data.Vec hiding (map) +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Categories.Category +open import Cubical.Categories.Limits.BinCoproduct +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure -module Realizability.Assembly.BinCoproducts {} {A : Type } (ca : CombinatoryAlgebra A) where +module Realizability.Assembly.BinCoproducts {} {A : Type } (ca : CombinatoryAlgebra A) where -open CombinatoryAlgebra ca -open import Realizability.Assembly.Base ca -open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -open import Realizability.Assembly.Morphism ca +open CombinatoryAlgebra ca +open import Realizability.Assembly.Base ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open import Realizability.Assembly.Morphism ca -λ* = `λ* as fefermanStructure -λ*ComputationRule = `λ*ComputationRule as fefermanStructure - -infixl 23 _⊕_ -_⊕_ : {A B : Type } Assembly A Assembly B Assembly (A B) -(as bs) .isSetX = isSet⊎ (as .isSetX) (bs .isSetX) -(as bs) ._⊩_ r (inl a) = ∃[ aᵣ A ] (as ._⊩_ aᵣ a) × (r pair true aᵣ) -(as bs) ._⊩_ r (inr b) = ∃[ bᵣ A ] (bs ._⊩_ bᵣ b) × (r pair false bᵣ) -(as bs) .⊩isPropValued r (inl a) = squash₁ -(as bs) .⊩isPropValued r (inr b) = squash₁ -(as bs) .⊩surjective (inl a) = - do - (a~ , a~realizes) as .⊩surjective a - return ( pair true a~ - , a~ - , a~realizes - , refl ∣₁ - ) -(as bs) .⊩surjective (inr b) = - do - (b~ , b~realizes) bs .⊩surjective b - return ( pair false b~ - , b~ - , b~realizes - , refl ∣₁ - ) +infixl 23 _⊕_ +_⊕_ : {A B : Type } Assembly A Assembly B Assembly (A B) +(as bs) .isSetX = isSet⊎ (as .isSetX) (bs .isSetX) +(as bs) ._⊩_ r (inl a) = ∃[ aᵣ A ] (as ._⊩_ aᵣ a) × (r pair true aᵣ) +(as bs) ._⊩_ r (inr b) = ∃[ bᵣ A ] (bs ._⊩_ bᵣ b) × (r pair false bᵣ) +(as bs) .⊩isPropValued r (inl a) = squash₁ +(as bs) .⊩isPropValued r (inr b) = squash₁ +(as bs) .⊩surjective (inl a) = + do + (a~ , a~realizes) as .⊩surjective a + return ( pair true a~ + , a~ + , a~realizes + , refl ∣₁ + ) +(as bs) .⊩surjective (inr b) = + do + (b~ , b~realizes) bs .⊩surjective b + return ( pair false b~ + , b~ + , b~realizes + , refl ∣₁ + ) -κ₁ : {A B : Type } {as : Assembly A} {bs : Assembly B} AssemblyMorphism as (as bs) -κ₁ .map = inl -κ₁ .tracker = pair true , x aₓ aₓ⊩x aₓ , aₓ⊩x , refl ∣₁) ∣₁ +κ₁ : {A B : Type } {as : Assembly A} {bs : Assembly B} AssemblyMorphism as (as bs) +κ₁ .map = inl +κ₁ .tracker = pair true , x aₓ aₓ⊩x aₓ , aₓ⊩x , refl ∣₁) ∣₁ -κ₂ : {A B : Type } {as : Assembly A} {bs : Assembly B} AssemblyMorphism bs (as bs) -κ₂ .map b = inr b -κ₂ .tracker = pair false , x bₓ bₓ⊩x bₓ , bₓ⊩x , refl ∣₁) ∣₁ +κ₂ : {A B : Type } {as : Assembly A} {bs : Assembly B} AssemblyMorphism bs (as bs) +κ₂ .map b = inr b +κ₂ .tracker = pair false , x bₓ bₓ⊩x bₓ , bₓ⊩x , refl ∣₁) ∣₁ -{-# TERMINATING #-} -[_,_] : {X Y Z : Type } {asmX : Assembly X} {asmY : Assembly Y} {asmZ : Assembly Z} (f : AssemblyMorphism asmX asmZ) (g : AssemblyMorphism asmY asmZ) AssemblyMorphism (asmX asmY) asmZ -[ f , g ] .map (inl x) = f .map x -[ f , g ] .map (inr y) = g .map y -[_,_] {asmZ = asmZ} f g .tracker = - do - (f~ , f~tracks) f .tracker - (g~ , g~tracks) g .tracker - -- if (pr₁ r) then f (pr₂ r) else g (pr₂ r) - let - tracker : Term as (suc zero) - tracker = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` f~ ̇ (` pr₂ ̇ (# fzero))) ̇ (` g~ ̇ (` pr₂ ̇ (# fzero))) - return - (λ* tracker , - λ { (inl x) r r⊩inl - transport - (propTruncIdempotent (asmZ .⊩isPropValued _ _)) - (do - (rₓ , rₓ⊩x , r≡pair⨾true⨾rₓ) r⊩inl - return - (subst - r asmZ ._⊩_ r ([ f , g ] .map (inl x))) - (sym - (λ* tracker r - ≡⟨ λ*ComputationRule tracker (r []) - Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r)) - ≡⟨ cong r Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r))) r≡pair⨾true⨾rₓ - Id (pr₁ (pair true rₓ)) (f~ (pr₂ (pair true rₓ))) (g~ (pr₂ (pair true rₓ))) - ≡⟨ cong x Id x (f~ (pr₂ (pair true rₓ))) (g~ (pr₂ (pair true rₓ)))) (pr₁pxy≡x _ _) - Id true (f~ (pr₂ (pair true rₓ))) (g~ (pr₂ (pair true rₓ))) - ≡⟨ ifTrueThen _ _ - f~ (pr₂ (pair true rₓ)) - ≡⟨ cong x f~ x) (pr₂pxy≡y _ _) - f~ rₓ - )) - (f~tracks x rₓ rₓ⊩x))) - ; (inr y) r r⊩inr - transport - (propTruncIdempotent (asmZ .⊩isPropValued _ _)) - (do - (yᵣ , yᵣ⊩y , r≡pair⨾false⨾yᵣ) r⊩inr - return - (subst - r asmZ ._⊩_ r ([ f , g ] .map (inr y))) - (sym - ((λ* tracker r - ≡⟨ λ*ComputationRule tracker (r []) - Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r)) - ≡⟨ cong r Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r))) r≡pair⨾false⨾yᵣ - Id (pr₁ (pair false yᵣ)) (f~ (pr₂ (pair false yᵣ))) (g~ (pr₂ (pair false yᵣ))) - ≡⟨ cong x Id x (f~ (pr₂ (pair false yᵣ))) (g~ (pr₂ (pair false yᵣ)))) (pr₁pxy≡x _ _) - Id false (f~ (pr₂ (pair false yᵣ))) (g~ (pr₂ (pair false yᵣ))) - ≡⟨ ifFalseElse _ _ - g~ (pr₂ (pair false yᵣ)) - ≡⟨ cong x g~ x) (pr₂pxy≡y _ _) - g~ yᵣ - ))) - (g~tracks y yᵣ yᵣ⊩y))) }) +{-# TERMINATING #-} +[_,_] : {X Y Z : Type } {asmX : Assembly X} {asmY : Assembly Y} {asmZ : Assembly Z} (f : AssemblyMorphism asmX asmZ) (g : AssemblyMorphism asmY asmZ) AssemblyMorphism (asmX asmY) asmZ +[ f , g ] .map (inl x) = f .map x +[ f , g ] .map (inr y) = g .map y +[_,_] {asmZ = asmZ} f g .tracker = + do + -- these are not considered structurally smaller since these are in the propositional truncation + (f~ , f~tracks) f .tracker + (g~ , g~tracks) g .tracker + -- if (pr₁ r) then f (pr₂ r) else g (pr₂ r) + let + tracker : Term as (suc zero) + tracker = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` f~ ̇ (` pr₂ ̇ (# zero))) ̇ (` g~ ̇ (` pr₂ ̇ (# zero))) + return + (λ* tracker , + λ { (inl x) r r⊩inl + transport + (propTruncIdempotent (asmZ .⊩isPropValued _ _)) + (do + (rₓ , rₓ⊩x , r≡pair⨾true⨾rₓ) r⊩inl + return + (subst + r asmZ ._⊩_ r ([ f , g ] .map (inl x))) + (sym + (λ* tracker r + ≡⟨ λ*ComputationRule tracker r + Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r)) + ≡⟨ cong r Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r))) r≡pair⨾true⨾rₓ + Id (pr₁ (pair true rₓ)) (f~ (pr₂ (pair true rₓ))) (g~ (pr₂ (pair true rₓ))) + ≡⟨ cong x Id x (f~ (pr₂ (pair true rₓ))) (g~ (pr₂ (pair true rₓ)))) (pr₁pxy≡x _ _) + Id true (f~ (pr₂ (pair true rₓ))) (g~ (pr₂ (pair true rₓ))) + ≡⟨ ifTrueThen _ _ + f~ (pr₂ (pair true rₓ)) + ≡⟨ cong x f~ x) (pr₂pxy≡y _ _) + f~ rₓ + )) + (f~tracks x rₓ rₓ⊩x))) + ; (inr y) r r⊩inr + transport + (propTruncIdempotent (asmZ .⊩isPropValued _ _)) + (do + (yᵣ , yᵣ⊩y , r≡pair⨾false⨾yᵣ) r⊩inr + return + (subst + r asmZ ._⊩_ r ([ f , g ] .map (inr y))) + (sym + ((λ* tracker r + ≡⟨ λ*ComputationRule tracker r + Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r)) + ≡⟨ cong r Id (pr₁ r) (f~ (pr₂ r)) (g~ (pr₂ r))) r≡pair⨾false⨾yᵣ + Id (pr₁ (pair false yᵣ)) (f~ (pr₂ (pair false yᵣ))) (g~ (pr₂ (pair false yᵣ))) + ≡⟨ cong x Id x (f~ (pr₂ (pair false yᵣ))) (g~ (pr₂ (pair false yᵣ)))) (pr₁pxy≡x _ _) + Id false (f~ (pr₂ (pair false yᵣ))) (g~ (pr₂ (pair false yᵣ))) + ≡⟨ ifFalseElse _ _ + g~ (pr₂ (pair false yᵣ)) + ≡⟨ cong x g~ x) (pr₂pxy≡y _ _) + g~ yᵣ + ))) + (g~tracks y yᵣ yᵣ⊩y))) }) -open BinCoproduct -BinCoproductsASM : BinCoproducts ASM -BinCoproductsASM (X , asmX) (Y , asmY) .binCoprodOb = X Y , asmX asmY -BinCoproductsASM (X , asmX) (Y , asmY) .binCoprodInj₁ = κ₁ -BinCoproductsASM (X , asmX) (Y , asmY) .binCoprodInj₂ = κ₂ -BinCoproductsASM (X , asmX) (Y , asmY) .univProp {Z , asmZ} f g = - uniqueExists - [ f , g ] - ((AssemblyMorphism≡ _ _ (funExt x refl))) , (AssemblyMorphism≡ _ _ (funExt y refl)))) - ! isProp× (isSetAssemblyMorphism _ _ _ _) (isSetAssemblyMorphism _ _ _ _)) - λ ! (κ₁⊚!≡f , κ₂⊚!≡g) AssemblyMorphism≡ _ _ (funExt λ { (inl x) i κ₁⊚!≡f (~ i) .map x ; (inr y) i κ₂⊚!≡g (~ i) .map y }) +open BinCoproduct +BinCoproductsASM : BinCoproducts ASM +BinCoproductsASM (X , asmX) (Y , asmY) .binCoprodOb = X Y , asmX asmY +BinCoproductsASM (X , asmX) (Y , asmY) .binCoprodInj₁ = κ₁ +BinCoproductsASM (X , asmX) (Y , asmY) .binCoprodInj₂ = κ₂ +BinCoproductsASM (X , asmX) (Y , asmY) .univProp {Z , asmZ} f g = + uniqueExists + [ f , g ] + ((AssemblyMorphism≡ _ _ (funExt x refl))) , (AssemblyMorphism≡ _ _ (funExt y refl)))) + ! isProp× (isSetAssemblyMorphism _ _ _ _) (isSetAssemblyMorphism _ _ _ _)) + λ ! (κ₁⊚!≡f , κ₂⊚!≡g) AssemblyMorphism≡ _ _ (funExt λ { (inl x) i κ₁⊚!≡f (~ i) .map x ; (inr y) i κ₂⊚!≡g (~ i) .map y }) -module _ - {X Y : Type } - (asmX : Assembly X) - (asmY : Assembly Y) - where +-- I have no idea why I did these since this can be derived from the universal property of the coproduct anyway? +module _ + {X Y : Type } + (asmX : Assembly X) + (asmY : Assembly Y) + where - asmX+Y = asmX asmY - asmY+X = asmY asmX + asmX+Y = asmX asmY + asmY+X = asmY asmX - X+Y→Y+X : AssemblyMorphism asmX+Y asmY+X - X+Y→Y+X .map (inl x) = inr x - X+Y→Y+X .map (inr y) = inl y - X+Y→Y+X .tracker = - do - let - tracker : Term as 1 - tracker = ` Id ̇ (` pr₁ ̇ # fzero) ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # fzero)) ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ # fzero)) - return - ((λ* tracker) , - { (inl x) r r⊩inl - transport - (propTruncIdempotent (asmY+X .⊩isPropValued (λ* tracker r) (inr x))) - (do - (rₓ , rₓ⊩x , r≡pair⨾true⨾rₓ) r⊩inl - let - λ*trackerEq : λ* tracker r pair false rₓ - λ*trackerEq = - λ* tracker r - ≡⟨ λ*ComputationRule tracker (r []) - Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r)) - ≡⟨ cong r Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r))) r≡pair⨾true⨾rₓ - Id (pr₁ (pair true rₓ)) (pair false (pr₂ (pair true rₓ))) (pair true (pr₂ (pair true rₓ))) - ≡⟨ cong r Id r (pair false (pr₂ (pair true rₓ))) (pair true (pr₂ (pair true rₓ)))) (pr₁pxy≡x _ _) - Id true (pair false (pr₂ (pair true rₓ))) (pair true (pr₂ (pair true rₓ))) - ≡⟨ ifTrueThen _ _ - pair false (pr₂ (pair true rₓ)) - ≡⟨ cong r pair false r) (pr₂pxy≡y _ _) - pair false rₓ - - return (return (rₓ , subst _ (sym λ*trackerEq) rₓ⊩x , λ*trackerEq))) - ; (inr y) r r⊩inr - transport - (propTruncIdempotent (asmY+X .⊩isPropValued (λ* tracker r) (inl y))) - (do - (yᵣ , yᵣ⊩y , r≡pair⨾false⨾yᵣ) r⊩inr - let - λ*trackerEq : λ* tracker r pair true yᵣ - λ*trackerEq = - λ* tracker r - ≡⟨ λ*ComputationRule tracker (r []) - Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r)) - ≡⟨ cong r Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r))) r≡pair⨾false⨾yᵣ - Id (pr₁ (pair false yᵣ)) (pair false (pr₂ (pair false yᵣ))) (pair true (pr₂ (pair false yᵣ))) - ≡⟨ cong r Id r (pair false (pr₂ (pair false yᵣ))) (pair true (pr₂ (pair false yᵣ)))) (pr₁pxy≡x _ _) - Id false (pair false (pr₂ (pair false yᵣ))) (pair true (pr₂ (pair false yᵣ))) - ≡⟨ ifFalseElse _ _ - pair true (pr₂ (pair false yᵣ)) - ≡⟨ cong r pair true r) (pr₂pxy≡y _ _) - pair true yᵣ - - return (return (yᵣ , subst _ (sym λ*trackerEq) yᵣ⊩y , λ*trackerEq))) })) + X+Y→Y+X : AssemblyMorphism asmX+Y asmY+X + X+Y→Y+X .map (inl x) = inr x + X+Y→Y+X .map (inr y) = inl y + X+Y→Y+X .tracker = + do + let + tracker : Term as 1 + tracker = ` Id ̇ (` pr₁ ̇ # zero) ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # zero)) ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ # zero)) + return + ((λ* tracker) , + { (inl x) r r⊩inl + transport + (propTruncIdempotent (asmY+X .⊩isPropValued (λ* tracker r) (inr x))) + (do + (rₓ , rₓ⊩x , r≡pair⨾true⨾rₓ) r⊩inl + let + λ*trackerEq : λ* tracker r pair false rₓ + λ*trackerEq = + λ* tracker r + ≡⟨ λ*ComputationRule tracker r + Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r)) + ≡⟨ cong r Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r))) r≡pair⨾true⨾rₓ + Id (pr₁ (pair true rₓ)) (pair false (pr₂ (pair true rₓ))) (pair true (pr₂ (pair true rₓ))) + ≡⟨ cong r Id r (pair false (pr₂ (pair true rₓ))) (pair true (pr₂ (pair true rₓ)))) (pr₁pxy≡x _ _) + Id true (pair false (pr₂ (pair true rₓ))) (pair true (pr₂ (pair true rₓ))) + ≡⟨ ifTrueThen _ _ + pair false (pr₂ (pair true rₓ)) + ≡⟨ cong r pair false r) (pr₂pxy≡y _ _) + pair false rₓ + + return (return (rₓ , subst _ (sym λ*trackerEq) rₓ⊩x , λ*trackerEq))) + ; (inr y) r r⊩inr + transport + (propTruncIdempotent (asmY+X .⊩isPropValued (λ* tracker r) (inl y))) + (do + (yᵣ , yᵣ⊩y , r≡pair⨾false⨾yᵣ) r⊩inr + let + λ*trackerEq : λ* tracker r pair true yᵣ + λ*trackerEq = + λ* tracker r + ≡⟨ λ*ComputationRule tracker r + Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r)) + ≡⟨ cong r Id (pr₁ r) (pair false (pr₂ r)) (pair true (pr₂ r))) r≡pair⨾false⨾yᵣ + Id (pr₁ (pair false yᵣ)) (pair false (pr₂ (pair false yᵣ))) (pair true (pr₂ (pair false yᵣ))) + ≡⟨ cong r Id r (pair false (pr₂ (pair false yᵣ))) (pair true (pr₂ (pair false yᵣ)))) (pr₁pxy≡x _ _) + Id false (pair false (pr₂ (pair false yᵣ))) (pair true (pr₂ (pair false yᵣ))) + ≡⟨ ifFalseElse _ _ + pair true (pr₂ (pair false yᵣ)) + ≡⟨ cong r pair true r) (pr₂pxy≡y _ _) + pair true yᵣ + + return (return (yᵣ , subst _ (sym λ*trackerEq) yᵣ⊩y , λ*trackerEq))) })) -CatIsoX+Y-Y+X : {X Y : Type } (asmX : Assembly X) (asmY : Assembly Y) CatIso ASM (X Y , asmX asmY) (Y X , asmY asmX) -CatIsoX+Y-Y+X asmX asmY = - (X+Y→Y+X asmX asmY) , - (isiso - (X+Y→Y+X asmY asmX) - (AssemblyMorphism≡ _ _ (funExt { (inl y) refl ; (inr x) refl }))) - (AssemblyMorphism≡ _ _ (funExt { (inl x) refl ; (inr y) refl })))) +CatIsoX+Y-Y+X : {X Y : Type } (asmX : Assembly X) (asmY : Assembly Y) CatIso ASM (X Y , asmX asmY) (Y X , asmY asmX) +CatIsoX+Y-Y+X asmX asmY = + (X+Y→Y+X asmX asmY) , + (isiso + (X+Y→Y+X asmY asmX) + (AssemblyMorphism≡ _ _ (funExt { (inl y) refl ; (inr x) refl }))) + (AssemblyMorphism≡ _ _ (funExt { (inl x) refl ; (inr y) refl })))) \ No newline at end of file diff --git a/docs/Realizability.Assembly.BinProducts.html b/docs/Realizability.Assembly.BinProducts.html index 3f768ee..9626ba2 100644 --- a/docs/Realizability.Assembly.BinProducts.html +++ b/docs/Realizability.Assembly.BinProducts.html @@ -1,201 +1,126 @@ -Realizability.Assembly.BinProducts
{-# OPTIONS --cubical --allow-unsolved-metas #-}
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation hiding (map)
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Categories.Limits.BinProduct
-open import Realizability.CombinatoryAlgebra
+Realizability.Assembly.BinProducts
{-# OPTIONS --cubical #-}
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.HITs.PropositionalTruncation hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Categories.Limits.BinProduct
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Assembly.BinProducts {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Assembly.BinProducts {} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open import Realizability.Assembly.Base ca
-open import Realizability.Assembly.Morphism ca
-open CombinatoryAlgebra ca
-open Assembly
-open AssemblyMorphism
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open CombinatoryAlgebra ca
+open Assembly
+open AssemblyMorphism
 
-infixl 23 _⊗_
-_⊗_ : {A B : Type }  Assembly A  Assembly B  Assembly (A × B)
-(as  bs) .isSetX = isSetΣ (as .isSetX)  _  bs .isSetX)
-(as  bs) ._⊩_ r (a , b) = (as ._⊩_ (pr₁  r) a) × (bs ._⊩_ (pr₂  r) b)
-(as  bs) .⊩isPropValued r (a , b) = isPropΣ (as .⊩isPropValued (pr₁  r) a)
-                                              _  bs .⊩isPropValued (pr₂  r) b)
-(as  bs) .⊩surjective (a , b) = do
-                                   (b~ , b~realizes)  bs .⊩surjective b
-                                   (a~ , a~realizes)  as .⊩surjective a
-                                   return
-                                     ( pair  a~  b~
-                                     , subst  x  as ._⊩_ x a) (sym (pr₁pxy≡x a~ b~)) a~realizes
-                                     , subst  x  bs ._⊩_ x b) (sym (pr₂pxy≡y a~ b~)) b~realizes
-                                     )
+infixl 23 _⊗_
+_⊗_ : {A B : Type }  Assembly A  Assembly B  Assembly (A × B)
+(as  bs) .isSetX = isSetΣ (as .isSetX)  _  bs .isSetX)
+(as  bs) ._⊩_ r (a , b) = (as ._⊩_ (pr₁  r) a) × (bs ._⊩_ (pr₂  r) b)
+(as  bs) .⊩isPropValued r (a , b) = isPropΣ (as .⊩isPropValued (pr₁  r) a)
+                                              _  bs .⊩isPropValued (pr₂  r) b)
+(as  bs) .⊩surjective (a , b) = do
+                                   (b~ , b~realizes)  bs .⊩surjective b
+                                   (a~ , a~realizes)  as .⊩surjective a
+                                   return
+                                     ( pair  a~  b~
+                                     , subst  x  as ._⊩_ x a) (sym (pr₁pxy≡x a~ b~)) a~realizes
+                                     , subst  x  bs ._⊩_ x b) (sym (pr₂pxy≡y a~ b~)) b~realizes
+                                     )
 
-⟪_,_⟫ : {X Y Z W : Type }
-        {xs : Assembly X}
-        {ys : Assembly Y}
-        {zs : Assembly Z}
-        {ws : Assembly W}
-        (f : AssemblyMorphism xs ys)
-        (g : AssemblyMorphism zs ws)
-         AssemblyMorphism (xs  zs) (ys  ws)
- f , g  .map (x , z) = f .map x , g .map z
-⟪_,_⟫ {ys = ys} {ws = ws} f g .tracker = (do
-                      (f~ , f~tracks)  f .tracker
-                      (g~ , g~tracks)  g .tracker
-                      return (s  (s  (k  pair)  (s  (k  f~)  (s  (k  pr₁)  Id)))  (s  (k  g~)  (s  (k  pr₂)  Id))
-                             , λ xz r r⊩xz 
-                               ( subst  y  ys ._⊩_ y (f .map (xz .fst)))
-                                 (sym (subst _
-                                             (sym (t⨾r≡pair_fg f~ g~ r))
-                                             (pr₁pxy≡x (f~  (pr₁  r)) (g~  (pr₂  r)))))
-                                 (f~tracks (xz .fst) (pr₁  r) (r⊩xz .fst)))
-                               , subst  y  ws ._⊩_ y (g .map (xz .snd)))
-                                 (sym (subst _
-                                             (sym (t⨾r≡pair_fg f~ g~ r))
-                                             (pr₂pxy≡y (f~  (pr₁  r)) (g~  (pr₂  r)))))
-                                 (g~tracks (xz .snd) (pr₂  r) (r⊩xz .snd))))
-                               where
-                      module _ (f~ g~ r : A) where
-                        subf≡fprr :  f pr  (s  (k  f)  (s  (k  pr)  Id)  r)  (f  (pr  r))
-                        subf≡fprr f pr =
-                                    s  (k  f)  (s  (k  pr)  Id)  r
-                                      ≡⟨ sabc≡ac_bc _ _ _ 
-                                    (k  f  r)  (s  (k  pr)  Id  r)
-                                      ≡⟨ cong  x  x  _) (kab≡a f r) 
-                                    f  (s  (k  pr)  Id  r)
-                                      ≡⟨ cong  x  f  x) (sabc≡ac_bc _ _ _) 
-                                    f  (k  pr  r  (Id  r))
-                                      ≡⟨ cong  x  f  (x  (Id  r))) (kab≡a _ _ ) 
-                                    f  (pr  (Id  r))
-                                      ≡⟨ cong  x  f  (pr  x)) (Ida≡a r) 
-                                    f  (pr  r)
-                                      
-                        t⨾r≡pair_fg :
-                          s  (s  (k  pair)  (s  (k  f~)  (s  (k  pr₁)  Id)))  (s  (k  g~)  (s  (k  pr₂)  Id))  r
-                           pair  (f~  (pr₁  r))  (g~  (pr₂  r))
-                        t⨾r≡pair_fg =
-                          s  (s  (k  pair)  (s  (k  f~)  (s  (k  pr₁)  Id)))  (s  (k  g~)  (s  (k  pr₂)  Id))  r
-                            ≡⟨ sabc≡ac_bc _ _ _ 
-                          s  (k  pair)  (s  (k  f~)  (s  (k  pr₁)  Id))  r  (s  (k  g~)  (s  (k  pr₂)  Id)  r)
-                            ≡⟨ cong  x  x  (s  (k  g~)  (s  (k  pr₂)  Id)  r)) (sabc≡ac_bc _ _ _) 
-                          k  pair  r  (s  (k  f~)  (s  (k  pr₁)  Id)  r)  (s  (k  g~)  (s  (k  pr₂)  Id)  r)
-                            ≡⟨ cong  x  x  (s  (k  f~)  (s  (k  pr₁)  Id)  r)  (s  (k  g~)  (s  (k  pr₂)  Id)  r))
-                              (kab≡a pair r) 
-                          pair  (s  (k  f~)  (s  (k  pr₁)  Id)  r)  (s  (k  g~)  (s  (k  pr₂)  Id)  r)
-                            ≡⟨ cong₂  x y  pair  x  y) (subf≡fprr f~ pr₁) (subf≡fprr g~ pr₂) 
-                          pair  (f~  (pr₁  r))  (g~  (pr₂  r))
-                            
-π₁ : {A B : Type } {as : Assembly A} {bs : Assembly B}  AssemblyMorphism (as  bs) as
-π₁ .map (a , b) = a
-π₁ .tracker =  pr₁ ,  (a , b) p (goal , _)  goal) ∣₁
+⟪_,_⟫ : {X Y Z W : Type }
+        {xs : Assembly X}
+        {ys : Assembly Y}
+        {zs : Assembly Z}
+        {ws : Assembly W}
+        (f : AssemblyMorphism xs ys)
+        (g : AssemblyMorphism zs ws)
+         AssemblyMorphism (xs  zs) (ys  ws)
+ f , g  .map (x , z) = f .map x , g .map z
+⟪_,_⟫ {ys = ys} {ws = ws} f g .tracker =
+  do
+    (f~ , f~⊩isTrackedF)  f .tracker
+    (g~ , g~⊩isTrackedG)  g .tracker
+    let
+      realizer : Term as 1
+      realizer = ` pair ̇ (` f~ ̇ (` pr₁ ̇ # zero)) ̇ (` g~ ̇ (` pr₂ ̇ # zero))
+    return
+      (λ* realizer ,
+       { (x , z) r (pr₁r⊩x , pr₂r⊩z) 
+        subst  r'  r' ⊩[ ys ] (f .map x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (f~⊩isTrackedF x (pr₁  r) pr₁r⊩x) ,
+        subst  r'  r' ⊩[ ws ] (g .map z)) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) (g~⊩isTrackedG z (pr₂  r) pr₂r⊩z) }))
+        
+π₁ : {A B : Type } {as : Assembly A} {bs : Assembly B}  AssemblyMorphism (as  bs) as
+π₁ .map (a , b) = a
+π₁ .tracker =  pr₁ ,  (a , b) p (goal , _)  goal) ∣₁
 
-π₂ : {A B : Type } {as : Assembly A} {bs : Assembly B}  AssemblyMorphism (as  bs) bs
-π₂ .map (a , b) = b
-π₂ .tracker =  pr₂ ,  (a , b) p (_ , goal)  goal) ∣₁
+π₂ : {A B : Type } {as : Assembly A} {bs : Assembly B}  AssemblyMorphism (as  bs) bs
+π₂ .map (a , b) = b
+π₂ .tracker =  pr₂ ,  (a , b) p (_ , goal)  goal) ∣₁
 
-⟨_,_⟩ : {X Y Z : Type }
-       {xs : Assembly X} {ys : Assembly Y} {zs : Assembly Z}
-       AssemblyMorphism zs xs
-       AssemblyMorphism zs ys
-       AssemblyMorphism zs (xs  ys)
- f , g  .map z = f .map z , g .map z
-⟨_,_⟩ {X} {Y} {Z} {xs} {ys} {zs} f g .tracker = map2 untruncated (f .tracker) (g .tracker) where
-  module _ 
-         ((f~ , f~tracks) : Σ[ f~  A ] tracks {xs = zs} {ys = xs}  f~ (f .map))
-         ((g~ , g~tracks) : Σ[ g~  A ] tracks {xs = zs} {ys = ys} g~ (g .map)) where
-           
-         _⊩X_ = xs ._⊩_
-         _⊩Y_ = ys ._⊩_
-         _⊩Z_ = zs ._⊩_
-             
-         t = s  (s  (k  pair)  (s  (k  f~)  Id))  (s  (k  g~)  Id)
-         untruncated : Σ[ t  A ] (∀ z zᵣ zᵣ⊩z  ((pr₁  (t  zᵣ)) ⊩X (f .map z)) × ((pr₂  (t  zᵣ)) ⊩Y (g .map z)))
-         untruncated = t , λ z zᵣ zᵣ⊩z  goal₁ z zᵣ zᵣ⊩z , goal₂ z zᵣ zᵣ⊩z where
-           module _ (z : Z) (zᵣ : A) (zᵣ⊩z : zᵣ ⊩Z z) where
-
-             pr₁⨾tracker⨾zᵣ≡f~⨾zᵣ : pr₁  (t  zᵣ)  f~  zᵣ
-             pr₁⨾tracker⨾zᵣ≡f~⨾zᵣ =
-               pr₁  (s  (s  (k  pair)  (s  (k  f~)  Id))  (s  (k  g~)  Id)  zᵣ)
-                          ≡⟨ cong  x  pr₁  x) (sabc≡ac_bc _ _ _) 
-               pr₁  (s  (k  pair)  (s  (k  f~)  Id)  zᵣ  (s  (k  g~)  Id  zᵣ))
-                          ≡⟨ cong  x  pr₁  (x  (s  (k  g~)  Id  zᵣ))) (sabc≡ac_bc _ _ _) 
-               pr₁  (k  pair  zᵣ  (s  (k  f~)  Id  zᵣ)  (s  (k  g~)  Id  zᵣ))
-                          ≡⟨ cong  x  pr₁  (x  (s  (k  f~)  Id  zᵣ)  (s  (k  g~)  Id  zᵣ))) (kab≡a _ _) 
-               pr₁  (pair  (s  (k  f~)  Id  zᵣ)  (s  (k  g~)  Id  zᵣ))
-                           ≡⟨ pr₁pxy≡x _ _ 
-               s  (k  f~)  Id  zᵣ
-                            ≡⟨ sabc≡ac_bc _ _ _ 
-               k  f~  zᵣ  (Id  zᵣ)
-                           ≡⟨ cong  x  x  (Id  zᵣ)) (kab≡a _ _) 
-               f~  (Id  zᵣ)
-                          ≡⟨ cong  x  f~  x) (Ida≡a _) 
-               f~  zᵣ
-                    
-
-             pr₂⨾tracker⨾zᵣ≡g~⨾zᵣ : pr₂  (t  zᵣ)  g~  zᵣ
-             pr₂⨾tracker⨾zᵣ≡g~⨾zᵣ =
-               pr₂  (s  (s  (k  pair)  (s  (k  f~)  Id))  (s  (k  g~)  Id)  zᵣ)
-                   ≡⟨ cong  x  pr₂  x) (sabc≡ac_bc _ _ _) 
-               pr₂  (s  (k  pair)  (s  (k  f~)  Id)  zᵣ  (s  (k  g~)  Id  zᵣ))
-                   ≡⟨ cong  x  pr₂  (x  (s  (k  g~)  Id  zᵣ))) (sabc≡ac_bc _ _ _) 
-               pr₂  (k  pair  zᵣ  (s  (k  f~)  Id  zᵣ)  (s  (k  g~)  Id  zᵣ))
-                   ≡⟨ cong  x  pr₂  (x  (s  (k  f~)  Id  zᵣ)  (s  (k  g~)  Id  zᵣ))) (kab≡a _ _) 
-               pr₂  (pair  (s  (k  f~)  Id  zᵣ)  (s  (k  g~)  Id  zᵣ))
-                   ≡⟨ pr₂pxy≡y _ _ 
-               s  (k  g~)  Id  zᵣ
-                   ≡⟨ sabc≡ac_bc _ _ _ 
-               k  g~  zᵣ  (Id  zᵣ)
-                   ≡⟨ cong  x  x  (Id  zᵣ)) (kab≡a _ _) 
-               g~  (Id  zᵣ)
-                  ≡⟨ cong  x  g~  x) (Ida≡a _) 
-               g~  zᵣ
-                    
-                  
-             goal₁ : (pr₁  (t  zᵣ)) ⊩X (f .map z)
-             goal₁ = subst  y  y ⊩X (f .map z)) (sym pr₁⨾tracker⨾zᵣ≡f~⨾zᵣ) (f~tracks z zᵣ zᵣ⊩z)
+⟨_,_⟩ : {X Y Z : Type }
+       {xs : Assembly X} {ys : Assembly Y} {zs : Assembly Z}
+       AssemblyMorphism zs xs
+       AssemblyMorphism zs ys
+       AssemblyMorphism zs (xs  ys)
+ f , g  .map z = f .map z , g .map z
+⟨_,_⟩ {X} {Y} {Z} {xs} {ys} {zs} f g .tracker =
+  do
+    (f~ , f~⊩isTrackedF)  f .tracker
+    (g~ , g~⊩isTrackedG)  g .tracker
+    let
+      realizer : Term as 1
+      realizer = ` pair ̇ (` f~ ̇ # zero) ̇ (` g~ ̇ # zero)
+    return
+      (λ* realizer ,
+       z r r⊩z 
+        subst  r'  r' ⊩[ xs ] (f .map z)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (f~⊩isTrackedF z r r⊩z) ,
+        subst  r'  r' ⊩[ ys ] (g .map z)) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) (g~⊩isTrackedG z r r⊩z)))
   
-             goal₂ : (pr₂  (t  zᵣ)) ⊩Y (g .map z)
-             goal₂ = subst  y  y ⊩Y (g .map z)) (sym pr₂⨾tracker⨾zᵣ≡g~⨾zᵣ) (g~tracks z zᵣ zᵣ⊩z)
-module _ {X Y : Type } (xs : Assembly X) (ys : Assembly Y) where
-    theπ₁ = π₁ {A = X} {B = Y} {as = xs} {bs = ys}
-    theπ₂ = π₂ {A = X} {B = Y} {as = xs} {bs = ys}
-    isBinProduct⊗ : ((Z , zs) : Σ[ Z  Type  ] Assembly Z)
-                    (f : AssemblyMorphism zs xs)
-                    (g : AssemblyMorphism zs ys)
-                    ∃![ fg  AssemblyMorphism zs (xs  ys) ] (fg  theπ₁  f) × (fg  theπ₂  g)
-    isBinProduct⊗ (Z , zs) f g =
-                  uniqueExists
-                    {B = λ fg  (fg  theπ₁  f) × (fg  theπ₂  g)}
-                     f , g 
-                    ( AssemblyMorphism≡ ( f , g   theπ₁) f (funExt  x  refl))
-                    , AssemblyMorphism≡ ( f , g   theπ₂) g (funExt  x  refl)))
-                     fg  isProp×
-                            (isSetAssemblyMorphism zs xs (fg  theπ₁) f)
-                            (isSetAssemblyMorphism zs ys (fg  theπ₂) g))
-                    -- TODO : Come up with a prettier proof
-                    λ fg (fgπ₁≡f , fgπ₂≡g)  sym ((lemma₂ fg fgπ₁≡f fgπ₂≡g)  (lemma₁ fg fgπ₁≡f fgπ₂≡g)) where
-                      module _ (fg : AssemblyMorphism zs (xs  ys))
-                               (fgπ₁≡f : fg  theπ₁  f)
-                               (fgπ₂≡g : fg  theπ₂  g) where
-                             lemma₁ :  fg  theπ₁ , fg  theπ₂    f , g 
-                             lemma₁ = AssemblyMorphism≡
-                                       fg  theπ₁ , fg  theπ₂ 
-                                       f , g 
-                                       i z  (fgπ₁≡f i .map z) , (fgπ₂≡g i .map z))
+module _ {X Y : Type } (xs : Assembly X) (ys : Assembly Y) where
+    theπ₁ = π₁ {A = X} {B = Y} {as = xs} {bs = ys}
+    theπ₂ = π₂ {A = X} {B = Y} {as = xs} {bs = ys}
+    isBinProduct⊗ : ((Z , zs) : Σ[ Z  Type  ] Assembly Z)
+                    (f : AssemblyMorphism zs xs)
+                    (g : AssemblyMorphism zs ys)
+                    ∃![ fg  AssemblyMorphism zs (xs  ys) ] (fg  theπ₁  f) × (fg  theπ₂  g)
+    isBinProduct⊗ (Z , zs) f g =
+                  uniqueExists
+                    {B = λ fg  (fg  theπ₁  f) × (fg  theπ₂  g)}
+                     f , g 
+                    ( AssemblyMorphism≡ ( f , g   theπ₁) f (funExt  x  refl))
+                    , AssemblyMorphism≡ ( f , g   theπ₂) g (funExt  x  refl)))
+                     fg  isProp×
+                            (isSetAssemblyMorphism zs xs (fg  theπ₁) f)
+                            (isSetAssemblyMorphism zs ys (fg  theπ₂) g))
+                    -- TODO : Come up with a prettier proof
+                    λ fg (fgπ₁≡f , fgπ₂≡g)  sym ((lemma₂ fg fgπ₁≡f fgπ₂≡g)  (lemma₁ fg fgπ₁≡f fgπ₂≡g)) where
+                      module _ (fg : AssemblyMorphism zs (xs  ys))
+                               (fgπ₁≡f : fg  theπ₁  f)
+                               (fgπ₂≡g : fg  theπ₂  g) where
+                             lemma₁ :  fg  theπ₁ , fg  theπ₂    f , g 
+                             lemma₁ = AssemblyMorphism≡
+                                       fg  theπ₁ , fg  theπ₂ 
+                                       f , g 
+                                       i z  (fgπ₁≡f i .map z) , (fgπ₂≡g i .map z))
 
-                             lemma₂ : fg   fg  theπ₁ , fg  theπ₂ 
-                             lemma₂ = AssemblyMorphism≡
-                                      fg
-                                       fg  theπ₁ , fg  theπ₂ 
-                                      (funExt λ x  ΣPathP (refl , refl))
+                             lemma₂ : fg   fg  theπ₁ , fg  theπ₂ 
+                             lemma₂ = AssemblyMorphism≡
+                                      fg
+                                       fg  theπ₁ , fg  theπ₂ 
+                                      (funExt λ x  ΣPathP (refl , refl))
 
-module _ where
-    open BinProduct
-    ASMBinProducts : BinProducts ASM
-    ASMBinProducts (X , xs) (Y , ys) .binProdOb = (X × Y) , (xs  ys)
-    ASMBinProducts (X , xs) (Y , ys) .binProdPr₁ = π₁ {as = xs} {bs = ys}
-    ASMBinProducts (X , xs) (Y , ys) .binProdPr₂ = π₂ {as = xs} {bs = ys}
-    ASMBinProducts (X , xs) (Y , ys) .univProp {z} f g = isBinProduct⊗ xs ys z f g
+module _ where
+    open BinProduct
+    ASMBinProducts : BinProducts ASM
+    ASMBinProducts (X , xs) (Y , ys) .binProdOb = (X × Y) , (xs  ys)
+    ASMBinProducts (X , xs) (Y , ys) .binProdPr₁ = π₁ {as = xs} {bs = ys}
+    ASMBinProducts (X , xs) (Y , ys) .binProdPr₂ = π₂ {as = xs} {bs = ys}
+    ASMBinProducts (X , xs) (Y , ys) .univProp {z} f g = isBinProduct⊗ xs ys z f g
 
\ No newline at end of file diff --git a/docs/Realizability.Assembly.Coequalizers.html b/docs/Realizability.Assembly.Coequalizers.html index b8faa75..4e7daf1 100644 --- a/docs/Realizability.Assembly.Coequalizers.html +++ b/docs/Realizability.Assembly.Coequalizers.html @@ -6,82 +6,87 @@ open import Cubical.HITs.PropositionalTruncation hiding (map) open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Data.Sigma -open import Cubical.Categories.Limits.Coequalizers -open import Realizability.CombinatoryAlgebra +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure -module Realizability.Assembly.Coequalizers {} {A : Type } (ca : CombinatoryAlgebra A) where +module Realizability.Assembly.Coequalizers {} {A : Type } (ca : CombinatoryAlgebra A) where -open import Realizability.Assembly.Base ca -open import Realizability.Assembly.Morphism ca -open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open CombinatoryAlgebra ca +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -module _ - {X Y : Type } - (xs : Assembly X) - (ys : Assembly Y) - (f g : AssemblyMorphism xs ys) - where - private - _⊩X_ = xs ._⊩_ - _⊩Y_ = ys ._⊩_ +module _ + {X Y : Type } + (xs : Assembly X) + (ys : Assembly Y) + (f g : AssemblyMorphism xs ys) + where + private + _⊩X_ = xs ._⊩_ + _⊩Y_ = ys ._⊩_ - _⊩coeq_ : (a : A) (x : SetCoequalizer (f .map) (g .map)) hProp - a ⊩coeq x = - setCoequalizerRec - isSetHProp - y (∃[ y' Y ] (inc {f = f .map} {g = g .map} y inc y') × (a ⊩Y y')) , squash₁) - x i (∃[ y' Y ] (coeq {f = f .map} {g = g .map} x i inc y') × (a ⊩Y y')) , squash₁) - x + _⊩coeq_ : (a : A) (x : SetCoequalizer (f .map) (g .map)) hProp + a ⊩coeq x = + setCoequalizerRec + isSetHProp + y (∃[ y' Y ] (inc {f = f .map} {g = g .map} y inc y') × (a ⊩Y y')) , squash₁) + x i (∃[ y' Y ] (coeq {f = f .map} {g = g .map} x i inc y') × (a ⊩Y y')) , squash₁) + x - coequalizer : Assembly (SetCoequalizer (f .map) (g .map)) - ⊩coeqSurjective : (x : SetCoequalizer (f .map) (g .map)) ∃[ a A ] ((a ⊩coeq x) .fst) + coequalizer : Assembly (SetCoequalizer (f .map) (g .map)) + ⊩coeqSurjective : (x : SetCoequalizer (f .map) (g .map)) ∃[ a A ] ((a ⊩coeq x) .fst) - coequalizer .isSetX = squash - coequalizer ._⊩_ a x = (a ⊩coeq x) .fst - coequalizer .⊩isPropValued a x = (a ⊩coeq x) .snd - coequalizer .⊩surjective x = ⊩coeqSurjective x + coequalizer .isSetX = squash + coequalizer ._⊩_ a x = (a ⊩coeq x) .fst + coequalizer .⊩isPropValued a x = (a ⊩coeq x) .snd + coequalizer .⊩surjective x = ⊩coeqSurjective x - ⊩coeqSurjective x = - setCoequalizerElimProp - {C = λ b ∃[ a A ] ((a ⊩coeq b) .fst)} - x squash₁) - b do - (b~ , b~realizes) ys .⊩surjective b - return (b~ , b~⊩coeq_inc_b b b~ b~realizes)) - x where - b~⊩coeq_inc_b : (b : Y) (b~ : A) (b~realizes : b~ ⊩Y b) (b~ ⊩coeq inc b) .fst - b~⊩coeq_inc_b b b~ b~realizes = b , refl , b~realizes ∣₁ - {- + ⊩coeqSurjective x = + setCoequalizerElimProp + {C = λ b ∃[ a A ] ((a ⊩coeq b) .fst)} + x squash₁) + b do + (b~ , b~realizes) ys .⊩surjective b + return (b~ , b~⊩coeq_inc_b b b~ b~realizes)) + x where + b~⊩coeq_inc_b : (b : Y) (b~ : A) (b~realizes : b~ ⊩Y b) (b~ ⊩coeq inc b) .fst + b~⊩coeq_inc_b b b~ b~realizes = b , refl , b~realizes ∣₁ + {- Coequalziers have a map E ← Y ⇇ X -} - ιcoequalizer : AssemblyMorphism ys coequalizer - ιcoequalizer .map = inc - ιcoequalizer .tracker = Id , y yᵣ yᵣ⊩y subst r (r ⊩coeq inc y) .fst) (sym (Ida≡a yᵣ)) y , refl , yᵣ⊩y ∣₁) ∣₁ + ιcoequalizer : AssemblyMorphism ys coequalizer + ιcoequalizer .map = inc + ιcoequalizer .tracker = Id , y yᵣ yᵣ⊩y subst r (r ⊩coeq inc y) .fst) (sym (Ida≡a yᵣ)) y , refl , yᵣ⊩y ∣₁) ∣₁ - coequalizerFactors : ((Z , zs) : Σ[ Z Type ] Assembly Z) - (ι' : AssemblyMorphism ys zs) - (f ι' g ι') - ∃![ ! AssemblyMorphism coequalizer zs ] (ιcoequalizer ! ι') - coequalizerFactors (Z , zs) ι' f⊚ι'≡g⊚ι' = - uniqueExists where - .map setCoequalizerRec (zs .isSetX) (ι' .map) λ x λ i f⊚ι'≡g⊚ι' i .map x - .tracker {!!}) - (AssemblyMorphism≡ _ _ (funExt λ x refl)) - ! isSetAssemblyMorphism ys zs (ιcoequalizer !) ι') - λ ! ιcoequalizer⊚!≡ι' AssemblyMorphism≡ _ _ - (funExt λ x - setCoequalizerElimProp - {C = λ x setCoequalizerRec (zs .isSetX) (ι' .map) x₁ i f⊚ι'≡g⊚ι' i .map x₁) x ! .map x} - x zs .isSetX _ _) y λ i ιcoequalizer⊚!≡ι' (~ i) .map y) x) - open Coequalizer - open IsCoequalizer - - ιIsCoequalizer : IsCoequalizer {C = ASM} f g ιcoequalizer - ιIsCoequalizer .glues = AssemblyMorphism≡ _ _ (funExt λ x SetCoequalizer.coeq x) - ιIsCoequalizer .univProp q qGlues = coequalizerFactors _ q qGlues - - ASMCoequalizer : Coequalizer {C = ASM} f g - ASMCoequalizer .coapex = (SetCoequalizer (f .map) (g .map)) , coequalizer - Coequalizer.coeq ASMCoequalizer = ιcoequalizer - ASMCoequalizer .isCoequalizer = ιIsCoequalizer + coequalizerFactors : ((Z , zs) : Σ[ Z Type ] Assembly Z) + (ι' : AssemblyMorphism ys zs) + (f ι' g ι') + ∃![ ! AssemblyMorphism coequalizer zs ] (ιcoequalizer ! ι') + coequalizerFactors (Z , zs) ι' f⊚ι'≡g⊚ι' = + uniqueExists + (let + map = x setCoequalizerRec (zs .isSetX) (ι' .map) x i f⊚ι'≡g⊚ι' i .map x) x) + in + makeAssemblyMorphism + map + (do + (ι'~ , ι'~⊩isTrackedι') ι' .tracker + return + (ι'~ , + x r r⊩x setCoequalizerElimProp {C = λ x (r : A) r ⊩[ coequalizer ] x (ι'~ r) ⊩[ zs ] (map x)} {!!} y r r⊩y {!!}) x r r⊩x)))) + {!!} + {!!} + {!!} + {- + uniqueExists (λ where + .map → setCoequalizerRec (zs .isSetX) (ι' .map) λ x → λ i → f⊚ι'≡g⊚ι' i .map x + .tracker → return ({!!} , (λ x r r⊩x → {!setCoequalizerElimProp {C = λ x → !}))) + (AssemblyMorphism≡ _ _ (funExt λ x → refl)) + (λ ! → isSetAssemblyMorphism ys zs (ιcoequalizer ⊚ !) ι') + λ ! ιcoequalizer⊚!≡ι' → AssemblyMorphism≡ _ _ + (funExt λ x → + setCoequalizerElimProp + {C = λ x → setCoequalizerRec (zs .isSetX) (ι' .map) (λ x₁ i → f⊚ι'≡g⊚ι' i .map x₁) x ≡ ! .map x} + (λ x → zs .isSetX _ _) (λ y → λ i → ιcoequalizer⊚!≡ι' (~ i) .map y) x) -}
\ No newline at end of file diff --git a/docs/Realizability.Assembly.Equalizers.html b/docs/Realizability.Assembly.Equalizers.html index 9f7c018..1f96bcd 100644 --- a/docs/Realizability.Assembly.Equalizers.html +++ b/docs/Realizability.Assembly.Equalizers.html @@ -1,42 +1,42 @@ -Realizability.Assembly.Equalizers
{-# OPTIONS --cubical --allow-unsolved-metas #-}
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation hiding (map)
-open import Realizability.CombinatoryAlgebra
+Realizability.Assembly.Equalizers
{-# OPTIONS --cubical #-}
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation hiding (map)
+open import Realizability.CombinatoryAlgebra
 
-module Realizability.Assembly.Equalizers {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Assembly.Equalizers {} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open import Realizability.Assembly.Base ca
-open import Realizability.Assembly.Morphism ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
 
-module _ {A B : Type } {as : Assembly A} {bs : Assembly B} (f g : AssemblyMorphism as bs) where
-  _⊩A_ = as ._⊩_
-  equalizer : Assembly (Σ[ a  A ] f .map a  g .map a)
-  equalizer .isSetX = isSetΣ (as .isSetX) λ x  isProp→isSet (bs .isSetX (f .map x) (g .map x))
-  equalizer ._⊩_ r (a , fa≡ga) = as ._⊩_ r a
-  equalizer .⊩isPropValued r (a , fa≡ga) = as .⊩isPropValued r a
-  equalizer .⊩surjective (a , fa≡ga) = as .⊩surjective a
+module _ {A B : Type } {as : Assembly A} {bs : Assembly B} (f g : AssemblyMorphism as bs) where
+  _⊩A_ = as ._⊩_
+  equalizer : Assembly (Σ[ a  A ] f .map a  g .map a)
+  equalizer .isSetX = isSetΣ (as .isSetX) λ x  isProp→isSet (bs .isSetX (f .map x) (g .map x))
+  equalizer ._⊩_ r (a , fa≡ga) = as ._⊩_ r a
+  equalizer .⊩isPropValued r (a , fa≡ga) = as .⊩isPropValued r a
+  equalizer .⊩surjective (a , fa≡ga) = as .⊩surjective a
 
-  ιequalizer : AssemblyMorphism equalizer as
-  ιequalizer .map (a , fa≡ga) = a
-  ιequalizer .tracker =  Id ,  x aₓ aₓ⊩x  subst  y  y ⊩A (x .fst)) (sym (Ida≡a aₓ)) aₓ⊩x) ∣₁
+  ιequalizer : AssemblyMorphism equalizer as
+  ιequalizer .map (a , fa≡ga) = a
+  ιequalizer .tracker =  Id ,  x aₓ aₓ⊩x  subst  y  y ⊩A (x .fst)) (sym (Ida≡a aₓ)) aₓ⊩x) ∣₁
                                                                                                  
-  equalizerFactors : ((Z , zs) : Σ[ Z  Type  ] (Assembly Z))
-                    (ι' : AssemblyMorphism zs as)
-                    (ι'  f  ι'  g)
-                    ∃![ !  AssemblyMorphism zs equalizer ] (!  ιequalizer  ι')
-  equalizerFactors (Z , zs) ι' ι'f≡ι'g =
-                   uniqueExists  where
-                                   .map z  ι' .map z , λ i  ι'f≡ι'g i .map z
-                                   .tracker  ι' .tracker)
-                                   (AssemblyMorphism≡ _ _ refl)
-                                    !  isSetAssemblyMorphism _ _ (!  ιequalizer) ι')
-                                   λ !' !'⊚ι≡ι'  AssemblyMorphism≡ _ _
-                                                  (funExt λ z  Σ≡Prop  x  bs .isSetX (f .map x) (g .map x))
-                                                           i  !'⊚ι≡ι' (~ i) .map z))
+  equalizerFactors : ((Z , zs) : Σ[ Z  Type  ] (Assembly Z))
+                    (ι' : AssemblyMorphism zs as)
+                    (ι'  f  ι'  g)
+                    ∃![ !  AssemblyMorphism zs equalizer ] (!  ιequalizer  ι')
+  equalizerFactors (Z , zs) ι' ι'f≡ι'g =
+                   uniqueExists  where
+                                   .map z  ι' .map z , λ i  ι'f≡ι'g i .map z
+                                   .tracker  ι' .tracker)
+                                   (AssemblyMorphism≡ _ _ refl)
+                                    !  isSetAssemblyMorphism _ _ (!  ιequalizer) ι')
+                                   λ !' !'⊚ι≡ι'  AssemblyMorphism≡ _ _
+                                                  (funExt λ z  Σ≡Prop  x  bs .isSetX (f .map x) (g .map x))
+                                                           i  !'⊚ι≡ι' (~ i) .map z))
 
 
\ No newline at end of file diff --git a/docs/Realizability.Assembly.Everything.html b/docs/Realizability.Assembly.Everything.html index e3bda71..f044745 100644 --- a/docs/Realizability.Assembly.Everything.html +++ b/docs/Realizability.Assembly.Everything.html @@ -8,5 +8,6 @@ open import Realizability.Assembly.Equalizers open import Realizability.Assembly.Exponentials open import Realizability.Assembly.Morphism -open import Realizability.Assembly.Regular.Everything +-- TODO : Fix regular structure modules +-- open import Realizability.Assembly.Regular.Everything
\ No newline at end of file diff --git a/docs/Realizability.Assembly.Exponentials.html b/docs/Realizability.Assembly.Exponentials.html index beee3f4..677cefa 100644 --- a/docs/Realizability.Assembly.Exponentials.html +++ b/docs/Realizability.Assembly.Exponentials.html @@ -2,128 +2,143 @@ Realizability.Assembly.Exponentials
{-# OPTIONS --cubical --allow-unsolved-metas #-}
 open import Cubical.Foundations.Prelude
 open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation hiding (map)
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Realizability.CombinatoryAlgebra
+open import Cubical.Data.FinData hiding (eq)
+open import Cubical.HITs.PropositionalTruncation hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Assembly.Exponentials {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Assembly.Exponentials {} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open import Realizability.Assembly.Base ca
-open import Realizability.Assembly.Morphism ca
-open import Realizability.Assembly.BinProducts ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.BinProducts ca
 
--- Exponential objects
-_⇒_ : {A B : Type }  (as : Assembly A)  (bs : Assembly B)  Assembly (AssemblyMorphism as bs)
-(as  bs) .isSetX = isSetAssemblyMorphism as bs
-(as  bs) ._⊩_ r f = tracks {xs = as} {ys = bs} r (f .map)
-_⇒_ {A} {B} as bs .⊩isPropValued r f = isPropTracks {X = A} {Y = B} {xs = as} {ys = bs}  r (f .map)
-(as  bs) .⊩surjective f = f .tracker
+-- Exponential objects
+_⇒_ : {A B : Type }  (as : Assembly A)  (bs : Assembly B)  Assembly (AssemblyMorphism as bs)
+(as  bs) .isSetX = isSetAssemblyMorphism as bs
+(as  bs) ._⊩_ r f = tracks {xs = as} {ys = bs} r (f .map)
+_⇒_ {A} {B} as bs .⊩isPropValued r f = isPropTracks {X = A} {Y = B} {xs = as} {ys = bs}  r (f .map)
+(as  bs) .⊩surjective f = f .tracker
 
--- What a distinguished gentleman
-eval : {X Y : Type }  (xs : Assembly X)  (ys : Assembly Y)  AssemblyMorphism ((xs  ys)  xs) ys
-eval xs ys .map (f , x) = f .map x
-eval {X} {Y} xs ys .tracker =
-        (s  (s  (k  pr₁)  Id)  (s  (k  pr₂)  Id))
-       ,  (f , x) r r⊩fx  subst
-                y  y ⊩Y (f .map x))
-               (sym (tracker⨾r≡pr₁r⨾pr₂r (f , x) r r⊩fx))
-               (pr₁r⨾pr₂rTracks (f , x) r r⊩fx))
-       ∣₁ where
-          _⊩Y_ = ys ._⊩_
-          module _ (fx : (AssemblyMorphism xs ys) × X)
-                   (r : A)
-                   (r⊩fx : ((xs  ys)  xs) ._⊩_ r (fx .fst , fx .snd)) where
-            f = fx .fst
-            x = fx .snd
+-- What a distinguished gentleman
+eval : {X Y : Type }  (xs : Assembly X)  (ys : Assembly Y)  AssemblyMorphism ((xs  ys)  xs) ys
+eval xs ys .map (f , x) = f .map x
+eval {X} {Y} xs ys .tracker =
+        (s  (s  (k  pr₁)  Id)  (s  (k  pr₂)  Id))
+       ,  (f , x) r r⊩fx  subst
+                y  y ⊩Y (f .map x))
+               (sym (tracker⨾r≡pr₁r⨾pr₂r (f , x) r r⊩fx))
+               (pr₁r⨾pr₂rTracks (f , x) r r⊩fx))
+       ∣₁ where
+          _⊩Y_ = ys ._⊩_
+          module _ (fx : (AssemblyMorphism xs ys) × X)
+                   (r : A)
+                   (r⊩fx : ((xs  ys)  xs) ._⊩_ r (fx .fst , fx .snd)) where
+            f = fx .fst
+            x = fx .snd
                           
-            pr₁r⨾pr₂rTracks : (pr₁  r  (pr₂  r)) ⊩Y (f .map x)
-            pr₁r⨾pr₂rTracks = r⊩fx .fst x (pr₂  r) (r⊩fx .snd)
+            pr₁r⨾pr₂rTracks : (pr₁  r  (pr₂  r)) ⊩Y (f .map x)
+            pr₁r⨾pr₂rTracks = r⊩fx .fst x (pr₂  r) (r⊩fx .snd)
                           
-            tracker⨾r≡pr₁r⨾pr₂r : s  (s  (k  pr₁)  Id)  (s  (k  pr₂)  Id)  r  (pr₁  r)  (pr₂  r)
-            tracker⨾r≡pr₁r⨾pr₂r =
-              s  (s  (k  pr₁)  Id)  (s  (k  pr₂)  Id)  r
-                ≡⟨ sabc≡ac_bc _ _ _  
-              (s  (k  pr₁)  Id  r)  (s  (k  pr₂)  Id  r)
-                ≡⟨ cong  x  x  (s  (k  pr₂)  Id  r)) (sabc≡ac_bc _ _ _)  
-              (k  pr₁  r  (Id  r))  (s  (k  pr₂)  Id  r)
-                ≡⟨ cong  x  (k  pr₁  r  (Id  r))  x) (sabc≡ac_bc _ _ _) 
-              (k  pr₁  r  (Id  r))  (k  pr₂  r  (Id  r))
-                ≡⟨ cong  x  (x  (Id  r))  (k  pr₂  r  (Id  r))) (kab≡a _ _) 
-              (pr₁  (Id  r))  (k  pr₂  r  (Id  r))
-                ≡⟨ cong  x  (pr₁  x)  (k  pr₂  r  (Id  r))) (Ida≡a r) 
-              (pr₁  r)  (k  pr₂  r  (Id  r))
-                ≡⟨ cong  x  (pr₁  r)  (x  (Id  r))) (kab≡a _ _)  
-              (pr₁  r)  (pr₂  (Id  r))
-                ≡⟨ cong  x  (pr₁  r)  (pr₂  x)) (Ida≡a r) 
-              (pr₁  r)  (pr₂  r)
-              
+            tracker⨾r≡pr₁r⨾pr₂r : s  (s  (k  pr₁)  Id)  (s  (k  pr₂)  Id)  r  (pr₁  r)  (pr₂  r)
+            tracker⨾r≡pr₁r⨾pr₂r =
+              s  (s  (k  pr₁)  Id)  (s  (k  pr₂)  Id)  r
+                ≡⟨ sabc≡ac_bc _ _ _  
+              (s  (k  pr₁)  Id  r)  (s  (k  pr₂)  Id  r)
+                ≡⟨ cong  x  x  (s  (k  pr₂)  Id  r)) (sabc≡ac_bc _ _ _)  
+              (k  pr₁  r  (Id  r))  (s  (k  pr₂)  Id  r)
+                ≡⟨ cong  x  (k  pr₁  r  (Id  r))  x) (sabc≡ac_bc _ _ _) 
+              (k  pr₁  r  (Id  r))  (k  pr₂  r  (Id  r))
+                ≡⟨ cong  x  (x  (Id  r))  (k  pr₂  r  (Id  r))) (kab≡a _ _) 
+              (pr₁  (Id  r))  (k  pr₂  r  (Id  r))
+                ≡⟨ cong  x  (pr₁  x)  (k  pr₂  r  (Id  r))) (Ida≡a r) 
+              (pr₁  r)  (k  pr₂  r  (Id  r))
+                ≡⟨ cong  x  (pr₁  r)  (x  (Id  r))) (kab≡a _ _)  
+              (pr₁  r)  (pr₂  (Id  r))
+                ≡⟨ cong  x  (pr₁  r)  (pr₂  x)) (Ida≡a r) 
+              (pr₁  r)  (pr₂  r)
+              
 
-module _ {X Y Z : Type }
-         {xs : Assembly X}
-         {ys : Assembly Y}
-         {zs : Assembly Z}
-         (f : AssemblyMorphism (zs  xs) ys) where
-         theEval = eval {X} {Y} xs ys
-         ⇒isExponential : ∃![ g  AssemblyMorphism zs (xs  ys) ]
-                           g , identityMorphism xs   theEval  f
-         ⇒isExponential = uniqueExists  where
-                                           .map z  λ where
-                                                        .map x  f .map (z , x)
-                                                        .tracker  do
-                                                                    (f~ , f~tracks)  f .tracker
-                                                                    (z~ , z~realizes)  zs .⊩surjective z
-                                                                    return ( (s  (k  f~)  (s  (k  (pair  z~))  Id)
-                                                                           , λ x aₓ aₓ⊩x
-                                                                            subst  k  k ⊩Y (f .map (z , x)))
-                                                                             (sym (eq f~ f~tracks z (z~ , z~realizes) x aₓ aₓ⊩x))
-                                                                             (pair⨾z~⨾aₓtracks f~ f~tracks z (z~ , z~realizes) x aₓ aₓ⊩x)))
-                                           .tracker  do
-                                                       (f~ , f~tracker)  f .tracker
-                                                       -- λ* x. λ* y. f~ ⨾ (pair ⨾ x ⨾ y)
-                                                       return ({!!} ,  z zᵣ zᵣ⊩z x xᵣ xᵣ⊩x  {!!})))
-                                        (AssemblyMorphism≡ _ _ (funExt  (z , x)  refl)))
-                                         g  isSetAssemblyMorphism _ _ ( g , identityMorphism xs   theEval) f)
-                                        λ g g×id⊚eval≡f  AssemblyMorphism≡ _ _
-                                                          (funExt  z  AssemblyMorphism≡ _ _
-                                                                         (funExt  x  λ i  g×id⊚eval≡f (~ i) .map (z , x))))) where
-                         _⊩X_ = xs ._⊩_
-                         _⊩Y_ = ys ._⊩_
-                         _⊩Z_ = zs ._⊩_
-                         _⊩Z×X_ = (zs  xs) ._⊩_
-                         Z×X = Z × X
-                         module _ (f~ : A)
-                                   (f~tracks : (∀ (zx : Z×X) (r : A) (rRealizes : (r ⊩Z×X zx))  ((f~  r) ⊩Y (f .map zx))))
-                                   (z : Z)
-                                   (zRealizer : Σ[ z~  A ] (z~ ⊩Z z))
-                                   (x : X)
-                                   (aₓ : A)
-                                   (aₓ⊩x : aₓ ⊩X x) where
-                            z~ : A
-                            z~ = zRealizer .fst
-                            z~realizes = zRealizer .snd
+module _ {X Y Z : Type }
+         {xs : Assembly X}
+         {ys : Assembly Y}
+         {zs : Assembly Z}
+         (f : AssemblyMorphism (zs  xs) ys) where
+         theEval = eval {X} {Y} xs ys
+         ⇒isExponential : ∃![ g  AssemblyMorphism zs (xs  ys) ]
+                           g , identityMorphism xs   theEval  f
+         ⇒isExponential = uniqueExists  where
+                                           .map z  λ where
+                                                        .map x  f .map (z , x)
+                                                        .tracker  do
+                                                                    (f~ , f~tracks)  f .tracker
+                                                                    (z~ , z~realizes)  zs .⊩surjective z
+                                                                    return ( (s  (k  f~)  (s  (k  (pair  z~))  Id)
+                                                                           , λ x aₓ aₓ⊩x
+                                                                            subst  k  k ⊩Y (f .map (z , x)))
+                                                                             (sym (eq f~ f~tracks z (z~ , z~realizes) x aₓ aₓ⊩x))
+                                                                             (pair⨾z~⨾aₓtracks f~ f~tracks z (z~ , z~realizes) x aₓ aₓ⊩x)))
+                                           .tracker  do
+                                                       (f~ , f~tracker)  f .tracker
+                                                       -- λ* x. λ* y. f~ ⨾ (pair ⨾ x ⨾ y)
+                                                       let
+                                                         realizer : Term as 2
+                                                         realizer = ` f~ ̇ (` pair ̇ # one ̇ # zero)
+                                                       return
+                                                         (λ*2 realizer ,
+                                                          z a a⊩z x b b⊩x 
+                                                           subst
+                                                              r'  r' ⊩[ ys ] (f .map (z , x)))
+                                                             (sym (λ*2ComputationRule realizer a b))
+                                                             (f~tracker
+                                                               (z , x)
+                                                               (pair  a  b)
+                                                               ((subst  r'  r' ⊩[ zs ] z) (sym (pr₁pxy≡x _ _)) a⊩z) ,
+                                                                (subst  r'  r' ⊩[ xs ] x) (sym (pr₂pxy≡y _ _)) b⊩x))))))
+                                        (AssemblyMorphism≡ _ _ (funExt  (z , x)  refl)))
+                                         g  isSetAssemblyMorphism _ _ ( g , identityMorphism xs   theEval) f)
+                                        λ g g×id⊚eval≡f  AssemblyMorphism≡ _ _
+                                                          (funExt  z  AssemblyMorphism≡ _ _
+                                                                         (funExt  x  λ i  g×id⊚eval≡f (~ i) .map (z , x))))) where
+                         _⊩X_ = xs ._⊩_
+                         _⊩Y_ = ys ._⊩_
+                         _⊩Z_ = zs ._⊩_
+                         _⊩Z×X_ = (zs  xs) ._⊩_
+                         Z×X = Z × X
+                         module _ (f~ : A)
+                                   (f~tracks : (∀ (zx : Z×X) (r : A) (rRealizes : (r ⊩Z×X zx))  ((f~  r) ⊩Y (f .map zx))))
+                                   (z : Z)
+                                   (zRealizer : Σ[ z~  A ] (z~ ⊩Z z))
+                                   (x : X)
+                                   (aₓ : A)
+                                   (aₓ⊩x : aₓ ⊩X x) where
+                            z~ : A
+                            z~ = zRealizer .fst
+                            z~realizes = zRealizer .snd
 
-                            eq : s  (k  f~)  (s  (k  (pair  z~))  Id)  aₓ  f~  (pair  z~  aₓ)
-                            eq =
-                              s  (k  f~)  (s  (k  (pair  z~))  Id)  aₓ
-                                ≡⟨ sabc≡ac_bc _ _ _ 
-                              (k  f~  aₓ)  (s  (k  (pair  z~))  Id  aₓ)
-                                ≡⟨ cong  x  x  (s  (k  (pair  z~))  Id  aₓ)) (kab≡a f~ aₓ) 
-                              f~  (s  (k  (pair  z~))  Id  aₓ)
-                                ≡⟨ cong  x  f~  x) (sabc≡ac_bc _ _ _) 
-                              f~  ((k  (pair  z~)  aₓ)  (Id  aₓ))
-                                ≡⟨ cong  x  f~  (x  (Id  aₓ))) (kab≡a (pair  z~) aₓ) 
-                              f~  (pair  z~  (Id  aₓ))
-                                ≡⟨ cong  x  f~  (pair  z~  x)) (Ida≡a aₓ) 
-                              f~  (pair  z~  aₓ)
-                                
+                            eq : s  (k  f~)  (s  (k  (pair  z~))  Id)  aₓ  f~  (pair  z~  aₓ)
+                            eq =
+                              s  (k  f~)  (s  (k  (pair  z~))  Id)  aₓ
+                                ≡⟨ sabc≡ac_bc _ _ _ 
+                              (k  f~  aₓ)  (s  (k  (pair  z~))  Id  aₓ)
+                                ≡⟨ cong  x  x  (s  (k  (pair  z~))  Id  aₓ)) (kab≡a f~ aₓ) 
+                              f~  (s  (k  (pair  z~))  Id  aₓ)
+                                ≡⟨ cong  x  f~  x) (sabc≡ac_bc _ _ _) 
+                              f~  ((k  (pair  z~)  aₓ)  (Id  aₓ))
+                                ≡⟨ cong  x  f~  (x  (Id  aₓ))) (kab≡a (pair  z~) aₓ) 
+                              f~  (pair  z~  (Id  aₓ))
+                                ≡⟨ cong  x  f~  (pair  z~  x)) (Ida≡a aₓ) 
+                              f~  (pair  z~  aₓ)
+                                
 
-                            pair⨾z~⨾aₓtracks : (f~  (pair  z~  aₓ)) ⊩Y (f .map (z , x))
-                            pair⨾z~⨾aₓtracks =
-                              f~tracks
-                                (z , x)
-                                (pair  z~  aₓ)
-                                ( (subst  y  y ⊩Z z) (sym (pr₁pxy≡x z~ aₓ)) z~realizes)
-                                , (subst  y  y ⊩X x) (sym (pr₂pxy≡y z~ aₓ)) aₓ⊩x))
+                            pair⨾z~⨾aₓtracks : (f~  (pair  z~  aₓ)) ⊩Y (f .map (z , x))
+                            pair⨾z~⨾aₓtracks =
+                              f~tracks
+                                (z , x)
+                                (pair  z~  aₓ)
+                                ( (subst  y  y ⊩Z z) (sym (pr₁pxy≡x z~ aₓ)) z~realizes)
+                                , (subst  y  y ⊩X x) (sym (pr₂pxy≡y z~ aₓ)) aₓ⊩x))
 
\ No newline at end of file diff --git a/docs/Realizability.Assembly.Morphism.html b/docs/Realizability.Assembly.Morphism.html index 52b2539..32739ba 100644 --- a/docs/Realizability.Assembly.Morphism.html +++ b/docs/Realizability.Assembly.Morphism.html @@ -4,152 +4,148 @@ open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism open import Cubical.Foundations.Function -open import Cubical.Data.Sigma -open import Cubical.HITs.PropositionalTruncation hiding (map) -open import Cubical.Reflection.RecordEquiv -open import Cubical.Categories.Category -open import Realizability.CombinatoryAlgebra +open import Cubical.Foundations.Equiv +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.HITs.PropositionalTruncation hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure -module Realizability.Assembly.Morphism {} {A : Type } (ca : CombinatoryAlgebra A) where +module Realizability.Assembly.Morphism {} {A : Type } (ca : CombinatoryAlgebra A) where -open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Base ca -open Assembly -open CombinatoryAlgebra ca -open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -module _ {X Y : Type } {xs : Assembly X} {ys : Assembly Y} (t : A) (f : X Y) where +module _ {X Y : Type } {xs : Assembly X} {ys : Assembly Y} (t : A) (f : X Y) where - tracks : Type - tracks = (x : X) (aₓ : A) (aₓ ⊩X x) (t aₓ) ⊩Y (f x) where - _⊩X_ = xs ._⊩_ - _⊩Y_ = ys ._⊩_ + tracks : Type + tracks = (x : X) (aₓ : A) (aₓ ⊩X x) (t aₓ) ⊩Y (f x) where + _⊩X_ = xs ._⊩_ + _⊩Y_ = ys ._⊩_ - isPropTracks : isProp tracks - isPropTracks = isPropΠ λ x - isPropΠ λ aₓ - isPropΠ λ aₓ⊩x - ys .⊩isPropValued (t aₓ) (f x) + isPropTracks : isProp tracks + isPropTracks = isPropΠ λ x + isPropΠ λ aₓ + isPropΠ λ aₓ⊩x + ys .⊩isPropValued (t aₓ) (f x) -record AssemblyMorphism {X Y : Type } (as : Assembly X) (bs : Assembly Y) : Type where - open Assembly as renaming (_⊩_ to _⊩X_) - open Assembly bs renaming (_⊩_ to _⊩Y_) - field - map : X Y - tracker : ∃[ t A ] ((x : X) (aₓ : A) (aₓ ⊩X x) (t aₓ) ⊩Y (map x)) -open AssemblyMorphism - -unquoteDecl AssemblyMorphismIsoΣ = declareRecordIsoΣ AssemblyMorphismIsoΣ (quote AssemblyMorphism) - -module _ {X Y : Type } (xs : Assembly X) (ys : Assembly Y) where +record AssemblyMorphism {X Y : Type } (as : Assembly X) (bs : Assembly Y) : Type where + no-eta-equality + constructor makeAssemblyMorphism + open Assembly as renaming (_⊩_ to _⊩X_) + open Assembly bs renaming (_⊩_ to _⊩Y_) + field + map : X Y + tracker : ∃[ t A ] ((x : X) (aₓ : A) (aₓ ⊩X x) (t aₓ) ⊩Y (map x)) +open AssemblyMorphism + +unquoteDecl AssemblyMorphismIsoΣ = declareRecordIsoΣ AssemblyMorphismIsoΣ (quote AssemblyMorphism) + +module _ {X Y : Type } (xs : Assembly X) (ys : Assembly Y) where - AssemblyMorphismΣ : Type - AssemblyMorphismΣ = Σ[ map (X Y) ] ∃[ t A ] ((x : X) (aₓ : A) (aₓ ⊩X x) (t aₓ) ⊩Y (map x)) where - _⊩X_ = xs ._⊩_ - _⊩Y_ = ys ._⊩_ + AssemblyMorphismΣ : Type + AssemblyMorphismΣ = Σ[ map (X Y) ] ∃[ t A ] ((x : X) (aₓ : A) (aₓ ⊩X x) (t aₓ) ⊩Y (map x)) where + _⊩X_ = xs ._⊩_ + _⊩Y_ = ys ._⊩_ - isSetAssemblyMorphismΣ : isSet AssemblyMorphismΣ - isSetAssemblyMorphismΣ = isSetΣ (isSet→ (ys .isSetX)) map isProp→isSet squash₁) + isSetAssemblyMorphismΣ : isSet AssemblyMorphismΣ + isSetAssemblyMorphismΣ = isSetΣ (isSet→ (ys .isSetX)) map isProp→isSet squash₁) - AssemblyMorphism≡Σ = isoToPath (AssemblyMorphismIsoΣ {as = xs} {bs = ys}) - - isSetAssemblyMorphism : isSet (AssemblyMorphism xs ys) - isSetAssemblyMorphism = subst t isSet t) (sym AssemblyMorphism≡Σ) isSetAssemblyMorphismΣ - -AssemblyMorphismΣ≡ : {X Y : Type } - {xs : Assembly X} - {ys : Assembly Y} - (f g : AssemblyMorphismΣ xs ys) - f .fst g .fst - --------------------------------- - f g -AssemblyMorphismΣ≡ f g = Σ≡Prop λ _ squash₁ - -module _ {X Y : Type } - {xs : Assembly X} - {ys : Assembly Y} - (f g : AssemblyMorphism xs ys) where - -- Necessary to please the constraint solver - theIso = AssemblyMorphismIsoΣ {X} {Y} {as = xs} {bs = ys} - thePath = AssemblyMorphismΣ≡ {X = X} {Y = Y} {xs = xs} {ys = ys} - open Iso - AssemblyMorphism≡ : (f .map g .map) f g - AssemblyMorphism≡ fmap≡gmap i = theIso .inv (thePath (theIso .fun f) (theIso .fun g) (fmap≡gmap) i) - -identityMorphism : {X : Type } (as : Assembly X) AssemblyMorphism as as -identityMorphism as .map x = x -identityMorphism as .tracker = Id , x aₓ aₓ⊩x subst y (as y) x) (sym (Ida≡a aₓ)) aₓ⊩x) ∣₁ - -compositeMorphism : {X Y Z : Type } {xs : Assembly X} {ys : Assembly Y} {zs : Assembly Z} - (f : AssemblyMorphism xs ys) - (g : AssemblyMorphism ys zs) - AssemblyMorphism xs zs -compositeMorphism f g .map x = g .map (f .map x) -compositeMorphism {X} {Y} {Z} {xs} {ys} {zs} f g .tracker = map2 untruncated (f .tracker) (g .tracker) where - open Assembly xs renaming (_⊩_ to _⊩X_) - open Assembly ys renaming (_⊩_ to _⊩Y_) - open Assembly zs renaming (_⊩_ to _⊩Z_) - module _ (fTracker : Σ[ f~ A ] tracks {xs = xs} {ys = ys} f~ (f .map)) - (gTracker : Σ[ g~ A ] tracks {xs = ys} {ys = zs} g~ (g .map)) where - - f~ = fTracker .fst - f~tracks = fTracker .snd - - g~ = gTracker .fst - g~tracks = gTracker .snd - - easierVariant : x aₓ aₓ⊩x (g~ (f~ aₓ)) ⊩Z g .map (f .map x) - easierVariant x aₓ aₓ⊩x = g~tracks (f .map x) (f~ aₓ) (f~tracks x aₓ aₓ⊩x) - - goal : (x : X) (aₓ : A) (aₓ⊩x : aₓ ⊩X x) - (B g~ f~ aₓ) ⊩Z (compositeMorphism f g .map x) - goal x aₓ aₓ⊩x = subst y y ⊩Z g .map (f .map x)) - (sym (Ba≡gfa g~ f~ aₓ)) - (easierVariant x aₓ aₓ⊩x) - - untruncated : Σ[ t A ] - ((x : X) (aₓ : A) - aₓ ⊩X x - (t aₓ) ⊩Z (compositeMorphism f g) .map x) - untruncated = B g~ f~ , goal - -infixl 23 _⊚_ -_⊚_ : {X Y Z : Type } {xs : Assembly X} {ys : Assembly Y} {zs : Assembly Z} - AssemblyMorphism xs ys - AssemblyMorphism ys zs - AssemblyMorphism xs zs -f g = compositeMorphism f g - -module _ {X Y : Type } (xs : Assembly X) (ys : Assembly Y) where - ⊚idL : (f : AssemblyMorphism xs ys) identityMorphism xs f f - ⊚idL f = AssemblyMorphism≡ (identityMorphism xs f) f (funExt λ x refl) - - ⊚idR : (f : AssemblyMorphism ys xs) f identityMorphism xs f - ⊚idR f = AssemblyMorphism≡ (f identityMorphism xs) f (funExt λ x refl) - -module _ {X Y Z W : Type } - (xs : Assembly X) - (ys : Assembly Y) - (zs : Assembly Z) - (ws : Assembly W) - (f : AssemblyMorphism xs ys) - (g : AssemblyMorphism ys zs) - (h : AssemblyMorphism zs ws) where - - ⊚assoc : (f g) h f (g h) - ⊚assoc = AssemblyMorphism≡ ((f g) h) (f (g h)) (∘-assoc (h .map) (g .map) (f .map)) - -open Category + AssemblyMorphism≡Σ = isoToPath (AssemblyMorphismIsoΣ {as = xs} {bs = ys}) + + isSetAssemblyMorphism : isSet (AssemblyMorphism xs ys) + isSetAssemblyMorphism = subst t isSet t) (sym AssemblyMorphism≡Σ) isSetAssemblyMorphismΣ + +AssemblyMorphism≡Equiv : {X Y : Type } {xs : Assembly X} {ys : Assembly Y} (f g : AssemblyMorphismΣ xs ys) (f .fst g .fst) (f g) +AssemblyMorphism≡Equiv {X} {Y} {xs} {ys} f g = Σ≡PropEquiv λ _ isPropPropTrunc + +AssemblyMorphismΣ≡ : {X Y : Type } + {xs : Assembly X} + {ys : Assembly Y} + (f g : AssemblyMorphismΣ xs ys) + f .fst g .fst + --------------------------------- + f g +AssemblyMorphismΣ≡ f g = Σ≡Prop λ _ squash₁ + +module _ {X Y : Type } + {xs : Assembly X} + {ys : Assembly Y} + (f g : AssemblyMorphism xs ys) where + -- Necessary to please the constraint solver + theIso = AssemblyMorphismIsoΣ {X} {Y} {as = xs} {bs = ys} + thePath = AssemblyMorphismΣ≡ {X = X} {Y = Y} {xs = xs} {ys = ys} + open Iso + AssemblyMorphism≡ : (f .map g .map) f g + map (AssemblyMorphism≡ fmap≡gmap i) x = fmap≡gmap i x + tracker (AssemblyMorphism≡ fmap≡gmap i) = + isProp→PathP + i isPropPropTrunc {A = Σ[ t A ] (∀ x aₓ aₓ ⊩[ xs ] x (t aₓ) ⊩[ ys ] (fmap≡gmap i x))}) + (f .tracker) (g .tracker) i + +identityMorphism : {X : Type } (as : Assembly X) AssemblyMorphism as as +identityMorphism as .map x = x +identityMorphism as .tracker = Id , x aₓ aₓ⊩x subst y (as y) x) (sym (Ida≡a aₓ)) aₓ⊩x) ∣₁ + +compositeMorphism : {X Y Z : Type } {xs : Assembly X} {ys : Assembly Y} {zs : Assembly Z} + (f : AssemblyMorphism xs ys) + (g : AssemblyMorphism ys zs) + AssemblyMorphism xs zs +compositeMorphism f g .map x = g .map (f .map x) +compositeMorphism {X} {Y} {Z} {xs} {ys} {zs} f g .tracker = + do + (f~ , isTrackedF) f .tracker + (g~ , isTrackedG) g .tracker + let + realizer : Term as 1 + realizer = ` g~ ̇ (` f~ ̇ # zero) + return + (λ* realizer , + x aₓ aₓ⊩x subst r' r' ⊩[ zs ] (g .map (f .map x))) (sym (λ*ComputationRule realizer aₓ)) (isTrackedG (f .map x) (f~ aₓ) (isTrackedF x aₓ aₓ⊩x)))) + +infixl 23 _⊚_ +_⊚_ : {X Y Z : Type } {xs : Assembly X} {ys : Assembly Y} {zs : Assembly Z} + AssemblyMorphism xs ys + AssemblyMorphism ys zs + AssemblyMorphism xs zs +f g = compositeMorphism f g + +module _ {X Y : Type } (xs : Assembly X) (ys : Assembly Y) where + ⊚idL : (f : AssemblyMorphism xs ys) identityMorphism xs f f + ⊚idL f = AssemblyMorphism≡ (identityMorphism xs f) f (funExt λ x refl) + + ⊚idR : (f : AssemblyMorphism ys xs) f identityMorphism xs f + ⊚idR f = AssemblyMorphism≡ (f identityMorphism xs) f (funExt λ x refl) + +module _ {X Y Z W : Type } + (xs : Assembly X) + (ys : Assembly Y) + (zs : Assembly Z) + (ws : Assembly W) + (f : AssemblyMorphism xs ys) + (g : AssemblyMorphism ys zs) + (h : AssemblyMorphism zs ws) where + + ⊚assoc : (f g) h f (g h) + ⊚assoc = AssemblyMorphism≡ ((f g) h) (f (g h)) (∘-assoc (h .map) (g .map) (f .map)) + +open Category -ASM : Category (ℓ-suc ) -ASM .ob = Σ[ X Type ] Assembly X -ASM .Hom[_,_] x y = AssemblyMorphism (x .snd) (y .snd) -ASM .id {x} = identityMorphism (x .snd) -ASM ._⋆_ f g = f g -ASM .⋆IdL {x} {y} f = ⊚idL (x .snd) (y .snd) f -ASM .⋆IdR {x} {y} f = ⊚idR (y .snd) (x .snd) f -ASM .⋆Assoc {x} {y} {z} {w} f g h = ⊚assoc (x .snd) (y .snd) (z .snd) (w .snd) f g h -ASM .isSetHom {x} {y} f g = isSetAssemblyMorphism (x .snd) (y .snd) f g - -open AssemblyMorphism public +ASM : Category (ℓ-suc ) +ASM .ob = Σ[ X Type ] Assembly X +ASM .Hom[_,_] x y = AssemblyMorphism (x .snd) (y .snd) +ASM .id {x} = identityMorphism (x .snd) +ASM ._⋆_ f g = f g +ASM .⋆IdL {x} {y} f = ⊚idL (x .snd) (y .snd) f +ASM .⋆IdR {x} {y} f = ⊚idR (y .snd) (x .snd) f +ASM .⋆Assoc {x} {y} {z} {w} f g h = ⊚assoc (x .snd) (y .snd) (z .snd) (w .snd) f g h +ASM .isSetHom {x} {y} f g = isSetAssemblyMorphism (x .snd) (y .snd) f g + +open AssemblyMorphism public \ No newline at end of file diff --git a/docs/Realizability.Assembly.SIP.html b/docs/Realizability.Assembly.SIP.html new file mode 100644 index 0000000..8b3edc0 --- /dev/null +++ b/docs/Realizability.Assembly.SIP.html @@ -0,0 +1,35 @@ + +Realizability.Assembly.SIP
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Realizability.CombinatoryAlgebra
+
+module Realizability.Assembly.SIP {} {A : Type } (ca : CombinatoryAlgebra A) where
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+
+module _ {X : Type } where
+
+  Assembly≡Iso :  (asmA asmB : Assembly X)  Iso (asmA  asmB) (∀ r x  r ⊩[ asmA ] x  r ⊩[ asmB ] x)
+  Iso.fun (Assembly≡Iso asmA asmB) path r x i = r ⊩[ path i ] x
+  Assembly._⊩_ (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) r x = pointwisePath r x i
+  Assembly.isSetX (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) = isPropIsSet {A = X} (asmA .isSetX) (asmB .isSetX) i
+  Assembly.⊩isPropValued (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) r x = isProp→PathP {B = λ j  isProp (pointwisePath r x j)}  j  isPropIsProp) (asmA .⊩isPropValued r x) (asmB .⊩isPropValued r x) i
+  Assembly.⊩surjective (Iso.inv (Assembly≡Iso asmA asmB) pointwisePath i) x = isProp→PathP {B = λ j  ∃[ a  A ] (pointwisePath a x j)}  j  isPropPropTrunc) (asmA .⊩surjective x) (asmB .⊩surjective x) i
+  Iso.rightInv (Assembly≡Iso asmA asmB) pointwise = funExt₂ λ r x  refl
+  Iso.leftInv (Assembly≡Iso asmA asmB) path = isSetAssembly X asmA asmB _ _
+
+  Assembly≡Equiv :  (asmA asmB : Assembly X)  (asmA  asmB)  (∀ r x  r ⊩[ asmA ] x  r ⊩[ asmB ] x)
+  Assembly≡Equiv asmA asmB = isoToEquiv (Assembly≡Iso asmA asmB)
+
+  Assembly≡ :  (asmA asmB : Assembly X)  (∀ r x  r ⊩[ asmA ] x  r ⊩[ asmB ] x)  (asmA  asmB)
+  Assembly≡ asmA asmB pointwise = Iso.inv (Assembly≡Iso asmA asmB) pointwise
+
\ No newline at end of file diff --git a/docs/Realizability.Assembly.SetsReflectiveSubcategory.html b/docs/Realizability.Assembly.SetsReflectiveSubcategory.html new file mode 100644 index 0000000..cf7c6ea --- /dev/null +++ b/docs/Realizability.Assembly.SetsReflectiveSubcategory.html @@ -0,0 +1,64 @@ + +Realizability.Assembly.SetsReflectiveSubcategory
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor
+open import Cubical.Categories.Instances.Sets
+open import Cubical.Categories.Adjoint
+open import Cubical.Categories.NaturalTransformation
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+
+module Realizability.Assembly.SetsReflectiveSubcategory {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+
+forgetfulFunctor : Functor ASM (SET )
+Functor.F-ob forgetfulFunctor (X , asmX) = X , asmX .isSetX
+Functor.F-hom forgetfulFunctor {X , asmX} {Y , asmY} f = f .map
+Functor.F-id forgetfulFunctor = refl
+Functor.F-seq forgetfulFunctor {X , asmX} {Y , asmY} {Z , asmZ} f g = refl
+
+ : Functor (SET ) ASM
+Functor.F-ob  (X , isSetX) = X , makeAssembly  a x  Unit*) isSetX  _ _  isPropUnit*) λ x   k , tt* ∣₁
+Functor.F-hom  {X , isSetX} {Y , isSetY} f = makeAssemblyMorphism f (return (k ,  _ _ _  tt*)))
+Functor.F-id  {X , isSetX} = AssemblyMorphism≡ _ _ refl
+Functor.F-seq  {X , isSetX} {Y , isSetY} {Z , isSetZ} f g = AssemblyMorphism≡ _ _ refl
+
+module _ where
+  open UnitCounit
+
+  adjointUnitCounit : forgetfulFunctor  
+  NatTrans.N-ob (_⊣_.η adjointUnitCounit) (X , asmX) = makeAssemblyMorphism  x  x) (return (k ,  _ _ _  tt*)))
+  NatTrans.N-hom (_⊣_.η adjointUnitCounit) {X , asmX} {Y , asmY} f = AssemblyMorphism≡ _ _ refl
+  NatTrans.N-ob (_⊣_.ε adjointUnitCounit) (X , isSetX) x = x
+  NatTrans.N-hom (_⊣_.ε adjointUnitCounit) {X , isSetX} {Y , isSetY} f = refl
+  TriangleIdentities.Δ₁ (_⊣_.triangleIdentities adjointUnitCounit) (X , asmX) = refl
+  TriangleIdentities.Δ₂ (_⊣_.triangleIdentities adjointUnitCounit) (X , isSetX) = AssemblyMorphism≡ _ _ refl
+
+module _ where
+  open NaturalBijection
+
+  adjointNaturalBijection : forgetfulFunctor  
+  Iso.fun (_⊣_.adjIso adjointNaturalBijection) f = makeAssemblyMorphism f (return (k ,  x r r⊩x  tt*)))
+  Iso.inv (_⊣_.adjIso adjointNaturalBijection) f = f .map
+  Iso.rightInv (_⊣_.adjIso adjointNaturalBijection) b = AssemblyMorphism≡ _ _ refl
+  Iso.leftInv (_⊣_.adjIso adjointNaturalBijection) a = refl
+  _⊣_.adjNatInD adjointNaturalBijection {X , isSetX} {Y , isSetY} f g = AssemblyMorphism≡ _ _ refl
+  _⊣_.adjNatInC adjointNaturalBijection {X , asmX} {Y , asmY} f g = refl
+
+
\ No newline at end of file diff --git a/docs/Realizability.Assembly.Terminal.html b/docs/Realizability.Assembly.Terminal.html new file mode 100644 index 0000000..b696816 --- /dev/null +++ b/docs/Realizability.Assembly.Terminal.html @@ -0,0 +1,51 @@ + +Realizability.Assembly.Terminal
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Categories.Limits.Terminal
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+
+module Realizability.Assembly.Terminal {} {A : Type } (ca : CombinatoryAlgebra A)  where
+
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open CombinatoryAlgebra ca
+open Assembly
+open AssemblyMorphism
+
+terminalAsm : Assembly Unit*
+(Assembly._⊩_ terminalAsm) a tt* = Unit*
+Assembly.isSetX terminalAsm = isSetUnit*
+(Assembly.⊩isPropValued terminalAsm) a tt* = isPropUnit*
+Assembly.⊩surjective terminalAsm tt* =  k , tt* ∣₁
+
+isTerminalTerminalAsm : isTerminal ASM (Unit* , terminalAsm)
+isTerminalTerminalAsm (X , asmX) =
+  inhProp→isContr
+    (makeAssemblyMorphism
+       x  tt*)
+      (return
+        (k ,  x r r⊩x  tt*))))
+     f g 
+      AssemblyMorphism≡ _ _ (funExt λ x  refl))
+
+TerminalASM : Terminal ASM
+fst TerminalASM = Unit* , terminalAsm
+snd TerminalASM = isTerminalTerminalAsm
+
+-- global element
+module _ {X : Type } (asmX : Assembly X) (x : X) (r : A) (r⊩x : r ⊩[ asmX ] x) where
+
+  globalElement : AssemblyMorphism terminalAsm asmX
+  AssemblyMorphism.map globalElement tt* = x
+  AssemblyMorphism.tracker globalElement =
+    return
+      ((k  r) ,
+       { tt* a tt*  subst  r'  r' ⊩[ asmX ] x) (sym (kab≡a _ _)) r⊩x }))
+
\ No newline at end of file diff --git a/docs/Realizability.CombinatoryAlgebra.html b/docs/Realizability.CombinatoryAlgebra.html index 1cb4a61..df9b188 100644 --- a/docs/Realizability.CombinatoryAlgebra.html +++ b/docs/Realizability.CombinatoryAlgebra.html @@ -12,63 +12,63 @@ record CombinatoryAlgebra {} (A : Type ) : Type where field - as : ApplicativeStructure A - fefermanStructure : Feferman as - open Feferman fefermanStructure public - open ApplicativeStructure as public - open ComputationRules as fefermanStructure public + as : ApplicativeStructure A + fefermanStructure : Feferman as + open Feferman fefermanStructure public + open ApplicativeStructure as public + open ComputationRules as fefermanStructure public module Combinators {} {A : Type } (ca : CombinatoryAlgebra A) where open CombinatoryAlgebra ca i : A - i = s k k + i = s k k k' : A - k' = k i + k' = k i - ia≡a : a i a a - ia≡a a = (cong x x a) refl) (sabc≡ac_bc k k a) (kab≡a a (k a)) + ia≡a : a i a a + ia≡a a = (cong x x a) refl) (sabc≡ac_bc k k a) (kab≡a a (k a)) - k'ab≡b : a b k' a b b - k'ab≡b a b = k' a b + k'ab≡b : a b k' a b b + k'ab≡b a b = k' a b ≡⟨ refl - (k i a b) - ≡⟨ cong x x b) (kab≡a i a) - (i b) + (k i a b) + ≡⟨ cong x x b) (kab≡a i a) + (i b) ≡⟨ ia≡a b b true : A - true = k + true = k false : A false = k' if_then_else_ : c t e A - if c then t else e = i c t e + if c then t else e = i c t e ifTrueThen : t e if true then t else e t ifTrueThen t e = if true then t else e ≡⟨ refl - i true t e - ≡⟨ cong x i x t e) refl - i k t e - ≡⟨ cong x x t e) (ia≡a k) - k t e - ≡⟨ kab≡a t e + i true t e + ≡⟨ cong x i x t e) refl + i k t e + ≡⟨ cong x x t e) (ia≡a k) + k t e + ≡⟨ kab≡a t e t ifFalseElse : t e if false then t else e e ifFalseElse t e = if false then t else e ≡⟨ refl - i false t e - ≡⟨ cong x i x t e) refl - i k' t e - ≡⟨ cong x x t e) (ia≡a k') - k' t e + i false t e + ≡⟨ cong x i x t e) refl + i k' t e + ≡⟨ cong x x t e) (ia≡a k') + k' t e ≡⟨ k'ab≡b t e e @@ -76,124 +76,124 @@ -- I used a Scheme script to generate this opaque pair : A - pair = s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) - (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) - (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) - (s (k (k)) (k (k))))) (s (k (k)) (k (k)))))) (s - (s (k (s)) (s (k (k)) (k (k)))) (s (k (k)) (s (k) (k))))))) - (s (s (k (s)) (s (k (k)) (k (k)))) (s (s (k (s)) (k (k))) (k (k)))) + pair = s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) + (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) + (s (s (k (s)) (s (s (k (s)) (s (k (k)) (k (s)))) + (s (k (k)) (k (k))))) (s (k (k)) (k (k)))))) (s + (s (k (s)) (s (k (k)) (k (k)))) (s (k (k)) (s (k) (k))))))) + (s (s (k (s)) (s (k (k)) (k (k)))) (s (s (k (s)) (k (k))) (k (k)))) opaque pr₁ : A - pr₁ = s (s k k) (k k) + pr₁ = s (s k k) (k k) pr₂ : A - pr₂ = s (s k k) (k k') + pr₂ = s (s k k) (k k') -- TODO : Prove computation rules - postulate pr₁pxy≡x : x y pr₁ (pair x y) x - postulate pr₂pxy≡y : x y pr₂ (pair x y) y + postulate pr₁pxy≡x : x y pr₁ (pair x y) x + postulate pr₂pxy≡y : x y pr₂ (pair x y) y -- Curry numbers ℕ→curry : A ℕ→curry zero = i - ℕ→curry (suc n) = pair k' (ℕ→curry n) + ℕ→curry (suc n) = pair k' (ℕ→curry n) Z : A Z = pr₁ opaque unfolding pr₁ - Zzero≡true : Z (ℕ→curry zero) true - Zzero≡true = Z (ℕ→curry zero) - ≡⟨ cong x Z x) refl - Z i - ≡⟨ cong x x i) refl - s (s k k) (k k) i - ≡⟨ sabc≡ac_bc (s k k) (k k) i - ((s k k) i) (k k i) - ≡⟨ cong x (x i) (k k i)) refl - (i i) (k k i) - ≡⟨ cong x x (k k i)) (ia≡a i) - i (k k i) - ≡⟨ cong x i x) (kab≡a k i) - i k - ≡⟨ ia≡a k - k + Zzero≡true : Z (ℕ→curry zero) true + Zzero≡true = Z (ℕ→curry zero) + ≡⟨ cong x Z x) refl + Z i + ≡⟨ cong x x i) refl + s (s k k) (k k) i + ≡⟨ sabc≡ac_bc (s k k) (k k) i + ((s k k) i) (k k i) + ≡⟨ cong x (x i) (k k i)) refl + (i i) (k k i) + ≡⟨ cong x x (k k i)) (ia≡a i) + i (k k i) + ≡⟨ cong x i x) (kab≡a k i) + i k + ≡⟨ ia≡a k + k - Zsuc≡false : n Z (ℕ→curry (suc n)) false - Zsuc≡false n = Z (ℕ→curry (suc n)) - ≡⟨ cong x Z x) refl - Z (pair k' (ℕ→curry n)) - ≡⟨ cong x x (pair k' (ℕ→curry n))) refl - pr₁ (pair k' (ℕ→curry n)) + Zsuc≡false : n Z (ℕ→curry (suc n)) false + Zsuc≡false n = Z (ℕ→curry (suc n)) + ≡⟨ cong x Z x) refl + Z (pair k' (ℕ→curry n)) + ≡⟨ cong x x (pair k' (ℕ→curry n))) refl + pr₁ (pair k' (ℕ→curry n)) ≡⟨ pr₁pxy≡x k' (ℕ→curry n) false S : A - S = pair k' + S = pair k' - Sn≡sucn : n S (ℕ→curry n) ℕ→curry (suc n) - Sn≡sucn n = S (ℕ→curry n) - ≡⟨ cong x x (ℕ→curry n)) refl - pair k' (ℕ→curry n) + Sn≡sucn : n S (ℕ→curry n) ℕ→curry (suc n) + Sn≡sucn n = S (ℕ→curry n) + ≡⟨ cong x x (ℕ→curry n)) refl + pair k' (ℕ→curry n) P : A - P = s (s (s (k pr₁) i) (k (ℕ→curry zero))) (s (k (pr₂)) i) + P = s (s (s (k pr₁) i) (k (ℕ→curry zero))) (s (k (pr₂)) i) - postulate Pzero≡zero : P (ℕ→curry zero) ℕ→curry zero - postulate Psucn≡n : n P (ℕ→curry (suc n)) ℕ→curry n + postulate Pzero≡zero : P (ℕ→curry zero) ℕ→curry zero + postulate Psucn≡n : n P (ℕ→curry (suc n)) ℕ→curry n B : g f A - B g f = s (k g) (s (k f) i) + B g f = s (k g) (s (k f) i) - Ba≡gfa : g f a B g f a g (f a) + Ba≡gfa : g f a B g f a g (f a) Ba≡gfa g f a = - s (k g) (s (k f) i) a - ≡⟨ sabc≡ac_bc (k g) (s (k f) i) a - (k g a) (s (k f) i a) - ≡⟨ cong x x (s (k f) i a)) (kab≡a g a) - g (s (k f) i a) - ≡⟨ cong x g x) (sabc≡ac_bc (k f) i a) - g ((k f a) (i a)) - ≡⟨ cong x g (x (i a))) (kab≡a f a) - g (f (i a)) - ≡⟨ cong x g (f x)) (ia≡a a) - g (f a) + s (k g) (s (k f) i) a + ≡⟨ sabc≡ac_bc (k g) (s (k f) i) a + (k g a) (s (k f) i a) + ≡⟨ cong x x (s (k f) i a)) (kab≡a g a) + g (s (k f) i a) + ≡⟨ cong x g x) (sabc≡ac_bc (k f) i a) + g ((k f a) (i a)) + ≡⟨ cong x g (x (i a))) (kab≡a f a) + g (f (i a)) + ≡⟨ cong x g (f x)) (ia≡a a) + g (f a) module Trivial {} {A : Type } (ca : CombinatoryAlgebra A) where open CombinatoryAlgebra ca open Combinators ca - module _ (isNonTrivial : s k ) where + module _ (isNonTrivial : s k ) where - k≠k' : k k' + k≠k' : k k' k≠k' k≡k' = isNonTrivial s≡k where - cond = if true then s else k - cond' = if false then s else k + cond = if true then s else k + cond' = if false then s else k condEq : cond cond' - condEq = cong x if x then s else k) k≡k' + condEq = cong x if x then s else k) k≡k' - cond≡s : cond s + cond≡s : cond s cond≡s = ifTrueThen _ _ - cond'≡k : cond' k + cond'≡k : cond' k cond'≡k = ifFalseElse _ _ - cond≡k : cond k - cond≡k = subst x x k) (sym condEq) cond'≡k + cond≡k : cond k + cond≡k = subst x x k) (sym condEq) cond'≡k - s≡k : s k + s≡k : s k s≡k = - s + s ≡⟨ sym cond≡s cond ≡⟨ cond≡k - k + k diff --git a/docs/Realizability.Modest.Base.html b/docs/Realizability.Modest.Base.html new file mode 100644 index 0000000..4065995 --- /dev/null +++ b/docs/Realizability.Modest.Base.html @@ -0,0 +1,87 @@ + +Realizability.Modest.Base
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.HITs.PropositionalTruncation hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.Base {} {A : Type } (ca : CombinatoryAlgebra A)  where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+
+module _ {X : Type } (asmX : Assembly X) where
+
+  isModest : Type _
+  isModest =  (x y : X) (a : A)  a ⊩[ asmX ] x  a ⊩[ asmX ] y  x  y
+
+  isPropIsModest : isProp isModest
+  isPropIsModest = isPropΠ3 λ x y a  isProp→ (isProp→ (asmX .isSetX x y))
+
+  isUniqueRealized : isModest   (a : A)  isProp (Σ[ x  X ] (a ⊩[ asmX ] x))
+  isUniqueRealized isMod a (x , a⊩x) (y , a⊩y) = Σ≡Prop  x'  asmX .⊩isPropValued a x') (isMod x y a a⊩x a⊩y)
+
+ModestSet : Type   Type (ℓ-suc )
+ModestSet X = Σ[ xs  Assembly X ] isModest xs
+
+MOD : Category (ℓ-suc ) 
+Category.ob MOD = Σ[ X  Type  ] Σ[ asmX  Assembly X ] isModest asmX
+Category.Hom[_,_] MOD (X , asmX , isModestAsmX) (Y , asmY , isModestAsmY) = AssemblyMorphism asmX asmY
+Category.id MOD {X , asmX , isModestAsmX} = identityMorphism asmX
+Category._⋆_ MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} {Z , asmZ , isModestAsmZ} f g = compositeMorphism f g
+Category.⋆IdL MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = ⊚idL asmX asmY f
+Category.⋆IdR MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = ⊚idR asmY asmX f
+Category.⋆Assoc MOD {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} {Z , asmZ , isModestAsmZ} {W , asmW , isModestAsmW} f g h = ⊚assoc asmX asmY asmZ asmW f g h
+Category.isSetHom MOD {X , asmX , idModestAsmX} {Y , asmY , isModestAsmY} = isSetAssemblyMorphism asmX asmY
+
+-- Every modest set is isomorphic to a canonically modest set
+module Canonical (X : Type ) (asmX : Assembly X) (isModestAsmX : isModest asmX) (resizing : hPropResizing ) where
+  open ResizedPowerset resizing
+  -- Replace every term of X by it's set of realisers
+  realisersOf : X   A
+  realisersOf x a = (a ⊩[ asmX ] x) , (asmX .⊩isPropValued a x)
+
+  resizedRealisersOf : X  𝓟 A
+  resizedRealisersOf x = ℙ→𝓟 A (realisersOf x)
+
+  realiserSet : Type 
+  realiserSet = Σ[ P  𝓟 A ] ∃[ x  X ] P  resizedRealisersOf x
+
+  canonicalModestSet : Assembly realiserSet
+  Assembly._⊩_ canonicalModestSet r (P , ∃x) = r ϵ P
+  Assembly.isSetX canonicalModestSet = isSetΣ (isSet𝓟 A)  P  isProp→isSet isPropPropTrunc)
+  Assembly.⊩isPropValued canonicalModestSet r (P , ∃x) = isPropϵ r P
+  Assembly.⊩surjective canonicalModestSet (P , ∃x) =
+    do
+      (x , P≡⊩x)  ∃x
+      (a , a⊩x)  asmX .⊩surjective x
+      return
+        (a ,
+        (subst
+           P  a ϵ P)
+          (sym P≡⊩x)
+          (subst  P  a  P) (sym (compIsIdFunc (realisersOf x))) a⊩x)))
+  {-
+  isModestCanonicalModestSet : isModest canonicalModestSet
+  isModestCanonicalModestSet x y a a⊩x a⊩y =
+    Σ≡Prop
+      (λ _ → isPropPropTrunc)
+      (𝓟≡ (x .fst) (y .fst) (⊆-extensionality (𝓟→ℙ A (x .fst)) (𝓟→ℙ A (y .fst)) ((λ b b⊩x → {!!}) , {!!}))) -}
+   
+  
+
\ No newline at end of file diff --git a/docs/Realizability.Modest.CanonicalPER.html b/docs/Realizability.Modest.CanonicalPER.html new file mode 100644 index 0000000..309af56 --- /dev/null +++ b/docs/Realizability.Modest.CanonicalPER.html @@ -0,0 +1,62 @@ + +Realizability.Modest.CanonicalPER
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Categories.Displayed.Reasoning
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Functor hiding (Id)
+open import Cubical.Categories.Constructions.Slice
+open import Categories.CartesianMorphism
+open import Categories.GenericObject
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.CanonicalPER {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.Terminal ca
+open import Realizability.Assembly.SetsReflectiveSubcategory ca
+open import Realizability.Modest.Base ca
+open import Realizability.Modest.UniformFamily ca
+open import Realizability.PERs.PER ca
+open import Realizability.PERs.SubQuotient ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Contravariant UNIMOD
+open UniformFamily
+open DisplayedUFamMap
+
+module _
+  {X : Type }
+  (asmX : Assembly X)
+  (isModestAsmX : isModest asmX) where
+
+  canonicalPER : PER
+  PER.relation canonicalPER a b = Σ[ x  X ] a ⊩[ asmX ] x × b ⊩[ asmX ] x
+  PER.isPropValued canonicalPER a b (x , a⊩x , b⊩x) (x' , a⊩x' , b⊩x') =
+    Σ≡Prop
+       x  isProp× (asmX .⊩isPropValued a x) (asmX .⊩isPropValued b x))
+      (isModestAsmX x x' a a⊩x a⊩x')
+  fst (PER.isPER canonicalPER) a b (x , a⊩x , b⊩x) = x , b⊩x , a⊩x
+  snd (PER.isPER canonicalPER) a b c (x , a⊩x , b⊩x) (x' , b⊩x' , c⊩x') =
+    x' , subst (a ⊩[ asmX ]_) (isModestAsmX x x' b b⊩x b⊩x') a⊩x , c⊩x'
+    
+  
+
\ No newline at end of file diff --git a/docs/Realizability.Modest.Everything.html b/docs/Realizability.Modest.Everything.html new file mode 100644 index 0000000..301899e --- /dev/null +++ b/docs/Realizability.Modest.Everything.html @@ -0,0 +1,11 @@ + +Realizability.Modest.Everything
module Realizability.Modest.Everything where
+
+open import Realizability.Modest.Base
+open import Realizability.Modest.CanonicalPER
+open import Realizability.Modest.UniformFamily
+open import Realizability.Modest.UniformFamilyCleavage
+open import Realizability.Modest.PartialSurjection
+-- open import Realizability.Modest.GenericUniformFamily
+open import Realizability.Modest.SubQuotientCanonicalPERIso
+
\ No newline at end of file diff --git a/docs/Realizability.Modest.PartialSurjection.html b/docs/Realizability.Modest.PartialSurjection.html new file mode 100644 index 0000000..1eb27cb --- /dev/null +++ b/docs/Realizability.Modest.PartialSurjection.html @@ -0,0 +1,390 @@ + +Realizability.Modest.PartialSurjection
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Foundations.Univalence
+open import Cubical.Functions.Surjection
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor.Base hiding (Id)
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.PartialSurjection {} {A : Type } (ca : CombinatoryAlgebra A) (resizing : hPropResizing ) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.SIP ca
+open import Realizability.Modest.Base ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open ResizedPowerset resizing
+
+record PartialSurjection (X : Type ) : Type (ℓ-suc ) where
+  no-eta-equality
+  constructor makePartialSurjection
+  field
+    support : A  Type 
+    enumeration : Σ[ a  A ] (support a)  X
+    isPropSupport :  a  isProp (support a)
+    isSurjectionEnumeration : isSurjection enumeration
+    isSetX : isSet X -- potentially redundant?
+open PartialSurjection
+
+module _ (X : Type ) (isCorrectHLevel : isSet X) where
+  -- first we need a Σ type equivalent to partial surjections
+  -- we could use RecordEquiv but this does not give hProps and hSets and
+  -- that causes problems when trying to compute the hlevel
+
+  PartialSurjectionΣ : Type (ℓ-suc )
+  PartialSurjectionΣ = Σ[ support  (A  hProp ) ] Σ[ enumeration  ((Σ[ a  A ]  support a )  X) ] isSurjection enumeration × isSet X
+
+  isSetPartialSurjectionΣ : isSet PartialSurjectionΣ
+  isSetPartialSurjectionΣ = isSetΣ (isSet→ isSetHProp)  support  isSetΣ (isSet→ isCorrectHLevel)  enum  isSet× (isProp→isSet isPropIsSurjection) (isProp→isSet isPropIsSet)))
+
+  PartialSurjectionIsoΣ : Iso (PartialSurjection X) PartialSurjectionΣ
+  Iso.fun PartialSurjectionIsoΣ surj =
+     a  (surj .support a) , (surj .isPropSupport a)) ,
+     { (a , suppA)  surj .enumeration (a , suppA) }) ,
+    surj .isSurjectionEnumeration ,
+    PartialSurjection.isSetX surj
+  Iso.inv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) =
+    makePartialSurjection  a   support a ) enumeration  a  str (support a)) isSurjectionEnumeration isSetX
+  Iso.rightInv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = refl
+  support (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .support a
+  enumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) (a , suppA) = surj .enumeration (a , suppA)
+  isPropSupport (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .isPropSupport a
+  isSurjectionEnumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSurjectionEnumeration
+  isSetX (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSetX
+
+  PartialSurjection≡Σ : PartialSurjection X  PartialSurjectionΣ
+  PartialSurjection≡Σ = isoToPath PartialSurjectionIsoΣ
+
+  isSetPartialSurjection : isSet (PartialSurjection X)
+  isSetPartialSurjection = subst isSet (sym PartialSurjection≡Σ) isSetPartialSurjectionΣ
+
+-- let us derive a structure of identity principle for partial surjections
+module SIP (X : Type ) (isCorrectHLevel : isSet X) where
+
+  PartialSurjection≡Iso :
+     (p q : PartialSurjection X)
+     Iso
+      (Σ[ suppPath  p .support  q .support ]
+      PathP  i  Σ[ a  A ] (suppPath i a)  X) (p .enumeration) (q .enumeration))
+      (p  q)
+  support (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z = suppPath i z
+  enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) (a , enum) = enumPath i (a , enum)
+  isPropSupport (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z =
+    isProp→PathP {B = λ j  isProp (suppPath j z)}  j  isPropIsProp) (p .isPropSupport z) (q .isPropSupport z) i
+  isSurjectionEnumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) b =
+    isProp→PathP
+      {B = λ j   fiber (enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) j)) b ∥₁}
+       j  isPropPropTrunc)
+      (p .isSurjectionEnumeration b) (q .isSurjectionEnumeration b) i
+  isSetX (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) = isPropIsSet (p .isSetX) (q .isSetX) i
+  Iso.inv (PartialSurjection≡Iso p q) p≡q =  i  p≡q i .support) ,  i  p≡q i .enumeration)
+  Iso.rightInv (PartialSurjection≡Iso p q) p≡q = isSetPartialSurjection X isCorrectHLevel _ _ _ _ 
+  Iso.leftInv (PartialSurjection≡Iso p q) (suppPath , enumPath) = ΣPathP (refl , refl)
+
+  PartialSurjection≡ :  (p q : PartialSurjection X)  Σ[ suppPath  p .support  q .support ] PathP  i  Σ[ a  A ] (suppPath i a)  X) (p .enumeration) (q .enumeration)  p  q
+  PartialSurjection≡ p q (suppPath , enumPath) = Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath)
+
+-- the type of partial surjections is equivalent to the type of modest assemblies on X
+module ModestSetIso (X : Type ) (isCorrectHLevel : isSet X) where
+
+  open SIP X isCorrectHLevel
+
+  {-# TERMINATING #-}
+  ModestSet→PartialSurjection : ModestSet X  PartialSurjection X
+  support (ModestSet→PartialSurjection (xs , isModestXs)) r = ∃[ x  X ] (r ⊩[ xs ] x)
+  enumeration (ModestSet→PartialSurjection (xs , isModestXs)) (r , ∃x) =
+    let
+      answer : Σ[ x  X ] (r ⊩[ xs ] x)
+      answer = PT.rec (isUniqueRealized xs isModestXs r)  t  t) ∃x
+    in fst answer
+  isPropSupport (ModestSet→PartialSurjection (xs , isModestXs)) r = isPropPropTrunc
+  isSurjectionEnumeration (ModestSet→PartialSurjection (xs , isModestXs)) x =
+    do
+      (a , a⊩x)  xs .⊩surjective x
+      return ((a ,  x , a⊩x ∣₁) , refl)
+  isSetX (ModestSet→PartialSurjection (xs , isModestXs)) = xs .isSetX
+
+  PartialSurjection→ModestSet : PartialSurjection X  ModestSet X
+  Assembly._⊩_ (fst (PartialSurjection→ModestSet surj)) r x =
+    Σ[ s  surj .support r ] surj .enumeration (r , s)  x
+  Assembly.isSetX (fst (PartialSurjection→ModestSet surj)) = surj .isSetX
+  Assembly.⊩isPropValued (fst (PartialSurjection→ModestSet surj)) a x (s , ≡x) (t , ≡x') =
+    Σ≡Prop  u  surj .isSetX (surj .enumeration (a , u)) x) (surj .isPropSupport a s t)
+  Assembly.⊩surjective (fst (PartialSurjection→ModestSet surj)) x =
+    do
+      ((a , s) , ≡x)  surj .isSurjectionEnumeration x
+      return (a , (s , ≡x))
+  snd (PartialSurjection→ModestSet surj) x y r (s , ≡x) (t , ≡x') =
+    x
+      ≡⟨ sym ≡x 
+    surj .enumeration (r , s)
+      ≡⟨ cong  s  surj .enumeration (r , s)) (surj .isPropSupport r s t) 
+    surj .enumeration (r , t)
+      ≡⟨ ≡x' 
+    y
+      
+
+  opaque
+    rightInv :  surj  ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)  surj
+    rightInv surj =
+      PartialSurjection≡
+        (ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)) surj
+        (funExt supportEq ,
+        funExtDep
+          {A = λ i  Σ-syntax A (funExt supportEq i)}
+          {B = λ _ _  X}
+          {f = ModestSet→PartialSurjection (PartialSurjection→ModestSet surj) .enumeration}
+          {g = surj .enumeration}
+          λ { {r , ∃x} {s , supp} p 
+            PT.elim
+              {P = λ ∃x  fst
+                             (PT.rec
+                              (isUniqueRealized (fst (PartialSurjection→ModestSet surj))
+                               (snd (PartialSurjection→ModestSet surj)) r)
+                               t  t) ∃x)
+                           surj .enumeration (s , supp)}
+              ∃x  surj .isSetX _ _)
+              { (x , suppR , ≡x) 
+               let
+                 ∃x' = transport (sym (supportEq s)) supp
+                 r≡s : r  s
+                 r≡s = PathPΣ p .fst
+               in
+               equivFun
+                 (propTruncIdempotent≃ (surj .isSetX x (surj .enumeration (s , supp))))
+                 (do
+                   (x' , suppS , ≡x')  ∃x'
+                   return
+                     (x
+                       ≡⟨ sym ≡x 
+                     surj .enumeration (r , suppR)
+                       ≡⟨ cong (surj .enumeration) (ΣPathP (r≡s , (isProp→PathP  i  surj .isPropSupport (PathPΣ p .fst i)) suppR supp))) 
+                     surj .enumeration (s , supp)
+                       )) })
+             ∃x }) where
+          supportEq :  r  (∃[ x  X ] (Σ[ supp  surj .support r ] (surj .enumeration (r , supp)  x)))  support surj r
+          supportEq =
+               r 
+                hPropExt
+                isPropPropTrunc
+                (surj .isPropSupport r)
+                 ∃x  PT.rec (surj .isPropSupport r)  { (x , supp , ≡x)  supp }) ∃x)
+                 supp  return (surj .enumeration (r , supp) , supp , refl)))
+
+  leftInv :  mod  PartialSurjection→ModestSet (ModestSet→PartialSurjection mod)  mod
+  leftInv (asmX , isModestAsmX) =
+    Σ≡Prop
+      isPropIsModest
+      (Assembly≡ _ _
+        λ r x 
+          hPropExt
+            (isPropΣ isPropPropTrunc  ∃x  asmX .isSetX _ _))
+            (asmX .⊩isPropValued r x)
+             { (∃x , ≡x) 
+              let
+                (x' , r⊩x') = PT.rec (isUniqueRealized asmX isModestAsmX r)  t  t) ∃x
+              in subst  x'  r ⊩[ asmX ] x') ≡x r⊩x'})
+            λ r⊩x   x , r⊩x ∣₁ , refl)
+
+  IsoModestSetPartialSurjection : Iso (ModestSet X) (PartialSurjection X)
+  Iso.fun IsoModestSetPartialSurjection = ModestSet→PartialSurjection
+  Iso.inv IsoModestSetPartialSurjection = PartialSurjection→ModestSet
+  Iso.rightInv IsoModestSetPartialSurjection = rightInv 
+  Iso.leftInv IsoModestSetPartialSurjection = leftInv
+
+  ModestSet≡PartialSurjection : ModestSet X  PartialSurjection X
+  ModestSet≡PartialSurjection = isoToPath IsoModestSetPartialSurjection
+
+record PartialSurjectionMorphism {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y) : Type  where
+  no-eta-equality
+  constructor makePartialSurjectionMorphism
+  field
+    map : X  Y
+    {-
+      The following "diagram" commutes
+                              
+      Xˢ -----------> X
+      |              |
+      |              |
+      |              |
+      |              |
+      |              |
+      ↓              ↓
+      Yˢ -----------> Y
+    -}
+    isTracked : ∃[ t  A ] (∀ (a : A) (sᵃ : psX .support a)  Σ[ sᵇ  (psY .support (t  a)) ] map (psX .enumeration (a , sᵃ))  psY .enumeration ((t  a) , sᵇ))
+open PartialSurjectionMorphism
+
+unquoteDecl PartialSurjectionMorphismIsoΣ = declareRecordIsoΣ PartialSurjectionMorphismIsoΣ (quote PartialSurjectionMorphism)
+
+PartialSurjectionMorphismΣ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  Type 
+PartialSurjectionMorphismΣ {X} {Y} psX psY =
+  Σ[ f  (X  Y) ] ∃[ t  A ] ((∀ (a : A) (sᵃ : psX .support a)  Σ[ sᵇ  (psY .support (t  a)) ] f (psX .enumeration (a , sᵃ))  psY .enumeration ((t  a) , sᵇ)))
+
+isSetPartialSurjectionMorphismΣ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  isSet (PartialSurjectionMorphismΣ psX psY)
+isSetPartialSurjectionMorphismΣ {X} {Y} psX psY = isSetΣ (isSet→ (psY .isSetX))  f  isProp→isSet isPropPropTrunc)
+
+PartialSurjectionMorphismΣ≡ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  PartialSurjectionMorphism psX psY  PartialSurjectionMorphismΣ psX psY
+PartialSurjectionMorphismΣ≡ {X} {Y} psX psY = isoToPath PartialSurjectionMorphismIsoΣ
+
+isSetPartialSurjectionMorphism : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  isSet (PartialSurjectionMorphism psX psY)
+isSetPartialSurjectionMorphism {X} {Y} psX psY = subst isSet (sym (PartialSurjectionMorphismΣ≡ psX psY)) (isSetPartialSurjectionMorphismΣ psX psY)
+
+-- SIP
+module MorphismSIP {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y) where
+  open PartialSurjectionMorphism
+  PartialSurjectionMorphism≡Iso :  (f g : PartialSurjectionMorphism psX psY)  Iso (f  g) (f .map  g .map)
+  Iso.fun (PartialSurjectionMorphism≡Iso f g) f≡g i = f≡g i .map
+  map (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) = fMap≡gMap i
+  isTracked (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) =
+    isProp→PathP
+      -- Agda can't infer the type B
+      {B = λ j  ∃-syntax A
+       t 
+         (a : A) (sᵃ : psX .support a) 
+         Σ-syntax (psY .support (t  a))
+          sᵇ 
+            fMap≡gMap j (psX .enumeration (a , sᵃ)) 
+            psY .enumeration (t  a , sᵇ)))}
+       j  isPropPropTrunc)
+      (f .isTracked) (g .isTracked) i
+  Iso.rightInv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap = refl
+  Iso.leftInv (PartialSurjectionMorphism≡Iso f g) f≡g = isSetPartialSurjectionMorphism psX psY f g _ _
+
+  PartialSurjectionMorphism≡ :  {f g : PartialSurjectionMorphism psX psY}  (f .map  g .map)  f  g
+  PartialSurjectionMorphism≡ {f} {g} fMap≡gMap = Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap
+
+-- morphisms between partial surjections are equivalent to assembly morphisms between corresponding modest assemblies
+module
+  _
+  {X Y : Type }
+  (psX : PartialSurjection X)
+  (psY : PartialSurjection Y) where
+  open ModestSetIso 
+  open MorphismSIP psX psY
+
+  asmX = PartialSurjection→ModestSet X (psX .isSetX) psX .fst
+  isModestAsmX = PartialSurjection→ModestSet X (psX .isSetX) psX .snd
+
+  asmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .fst
+  isModestAsmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .snd
+
+  PartialSurjectionHomModestSetHomIso : Iso (AssemblyMorphism asmX asmY) (PartialSurjectionMorphism psX psY)
+  map (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) = asmHom .map
+  isTracked (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) =
+    do
+      (map~ , isTrackedMap)  asmHom .tracker
+      return
+        (map~ ,
+         λ a aSuppX 
+           let
+             worker : (map~  a) ⊩[ asmY ] (asmHom .map (psX .enumeration (a , aSuppX)))
+             worker = isTrackedMap (psX .enumeration (a , aSuppX)) a (aSuppX , refl)
+           in
+           (worker .fst) ,
+           (sym (worker .snd)))
+  AssemblyMorphism.map (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) = surjHom .map
+  AssemblyMorphism.tracker (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) =
+    do
+      (t , isTrackedMap)  surjHom .isTracked
+      return
+        (t ,
+         { x a (aSuppX , ≡x) 
+          (isTrackedMap a aSuppX .fst) ,
+          (sym (cong (surjHom .map) (sym ≡x)  isTrackedMap a aSuppX .snd)) }))
+  Iso.rightInv PartialSurjectionHomModestSetHomIso surjHom = PartialSurjectionMorphism≡ refl
+  Iso.leftInv PartialSurjectionHomModestSetHomIso asmHom = AssemblyMorphism≡ _ _ refl
+
+  PartialSurjectionHom≡ModestSetHom : AssemblyMorphism asmX asmY  PartialSurjectionMorphism psX psY
+  PartialSurjectionHom≡ModestSetHom = isoToPath PartialSurjectionHomModestSetHomIso
+
+-- the category of partial surjections
+
+idPartSurjMorphism :  {X : Type }  (psX : PartialSurjection X)  PartialSurjectionMorphism psX psX
+map (idPartSurjMorphism {X} psX) x = x
+isTracked (idPartSurjMorphism {X} psX) =
+  return (Id ,  a aSuppX  (subst  r  psX .support r) (sym (Ida≡a a)) aSuppX) , (cong (psX .enumeration) (Σ≡Prop  b  psX .isPropSupport b) (sym (Ida≡a a))))))
+
+composePartSurjMorphism :
+   {X Y Z : Type } {psX : PartialSurjection X} {psY : PartialSurjection Y} {psZ : PartialSurjection Z}
+   (f : PartialSurjectionMorphism psX psY)
+   (g : PartialSurjectionMorphism psY psZ)
+   PartialSurjectionMorphism psX psZ
+map (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) x = g .map (f .map x)
+isTracked (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) =
+  do
+    (f~ , isTrackedF)  f .isTracked
+    (g~ , isTrackedG)  g .isTracked
+    let
+      realizer : Term as 1
+      realizer = ` g~ ̇ (` f~ ̇ # zero)
+    return
+      (λ* realizer ,
+       a aSuppX 
+        subst  r'  psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~  a) (isTrackedF a aSuppX .fst) .fst) ,
+       (g .map (f .map (psX .enumeration (a , aSuppX)))
+          ≡⟨ cong (g .map) (isTrackedF a aSuppX .snd) 
+        g .map (psY .enumeration (f~  a , fst (isTrackedF a aSuppX)))
+          ≡⟨ isTrackedG (f~  a) (fst (isTrackedF a aSuppX)) .snd 
+        psZ .enumeration (g~  (f~  a) , fst (isTrackedG (f~  a) (fst (isTrackedF a aSuppX))))
+          ≡⟨ cong (psZ .enumeration) (Σ≡Prop  z  psZ .isPropSupport z) (sym (λ*ComputationRule realizer a))) 
+        psZ .enumeration
+          (λ* realizer  a ,
+           subst  r'  psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~  a) (isTrackedF a aSuppX .fst) .fst))
+          )))
+
+idLPartSurjMorphism :
+   {X Y : Type }
+   {psX : PartialSurjection X}
+   {psY : PartialSurjection Y}
+   (f : PartialSurjectionMorphism psX psY)
+   composePartSurjMorphism (idPartSurjMorphism psX) f  f
+idLPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl
+
+idRPartSurjMorphism :
+   {X Y : Type }
+   {psX : PartialSurjection X}
+   {psY : PartialSurjection Y}
+   (f : PartialSurjectionMorphism psX psY)
+   composePartSurjMorphism f (idPartSurjMorphism psY)  f
+idRPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl
+
+assocComposePartSurjMorphism :
+   {X Y Z W : Type }
+   {psX : PartialSurjection X}
+   {psY : PartialSurjection Y}
+   {psZ : PartialSurjection Z}
+   {psW : PartialSurjection W}
+   (f : PartialSurjectionMorphism psX psY)
+   (g : PartialSurjectionMorphism psY psZ)
+   (h : PartialSurjectionMorphism psZ psW)
+   composePartSurjMorphism (composePartSurjMorphism f g) h  composePartSurjMorphism f (composePartSurjMorphism g h)
+assocComposePartSurjMorphism {X} {Y} {Z} {W} {psX} {psY} {psZ} {psW} f g h = MorphismSIP.PartialSurjectionMorphism≡ psX psW refl
+
+PARTSURJ : Category (ℓ-suc ) 
+Category.ob PARTSURJ = Σ[ X  Type  ] PartialSurjection X
+Category.Hom[_,_] PARTSURJ (X , surjX) (Y , surjY) = PartialSurjectionMorphism surjX surjY
+Category.id PARTSURJ {X , surjX} = idPartSurjMorphism surjX
+Category._⋆_ PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} f g = composePartSurjMorphism f g
+Category.⋆IdL PARTSURJ {X , surjX} {Y , surjY} f = idLPartSurjMorphism f
+Category.⋆IdR PARTSURJ {X , surjX} {Y , surjY} f = idRPartSurjMorphism f
+Category.⋆Assoc PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} {W , surjW} f g h = assocComposePartSurjMorphism f g h
+Category.isSetHom PARTSURJ {X , surjX} {Y , surjY} = isSetPartialSurjectionMorphism surjX surjY
+
\ No newline at end of file diff --git a/docs/Realizability.Modest.SubQuotientCanonicalPERIso.html b/docs/Realizability.Modest.SubQuotientCanonicalPERIso.html new file mode 100644 index 0000000..428ead4 --- /dev/null +++ b/docs/Realizability.Modest.SubQuotientCanonicalPERIso.html @@ -0,0 +1,149 @@ + +Realizability.Modest.SubQuotientCanonicalPERIso
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Categories.Displayed.Reasoning
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Functor hiding (Id)
+open import Cubical.Categories.Constructions.Slice
+open import Categories.CartesianMorphism
+open import Categories.GenericObject
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.SubQuotientCanonicalPERIso {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.Terminal ca
+open import Realizability.Assembly.SetsReflectiveSubcategory ca
+open import Realizability.Modest.Base ca
+open import Realizability.Modest.UniformFamily ca
+open import Realizability.Modest.CanonicalPER ca
+open import Realizability.PERs.PER ca
+open import Realizability.PERs.SubQuotient ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Contravariant UNIMOD
+open UniformFamily
+open DisplayedUFamMap
+
+module
+  _ {X : Type }
+  (asmX : Assembly X)
+  (isModestAsmX : isModest asmX) where
+
+  theCanonicalPER : PER
+  theCanonicalPER = canonicalPER asmX isModestAsmX
+
+  theSubQuotient : Assembly (subQuotient theCanonicalPER)
+  theSubQuotient = subQuotientAssembly theCanonicalPER
+
+  invert : AssemblyMorphism theSubQuotient asmX
+  AssemblyMorphism.map invert sq = SQ.rec (asmX .isSetX) reprAction reprActionCoh sq module Invert where
+
+    reprAction : Σ[ a  A ] (a ~[ theCanonicalPER ] a)  X
+    reprAction (a , x , a⊩x , _) = x
+
+    reprActionCoh :  a b a~b  reprAction a  reprAction b
+    reprActionCoh (a , x , a⊩x , _) (b , x' , b⊩x' , _) (x'' , a⊩x'' , b⊩x'') =
+      x
+        ≡⟨ isModestAsmX x x'' a a⊩x a⊩x'' 
+      x''
+        ≡⟨ isModestAsmX x'' x' b b⊩x'' b⊩x' 
+      x'
+        
+  AssemblyMorphism.tracker invert = return (Id ,  sq a a⊩sq  goal sq a a⊩sq)) where
+    realizability : (sq : subQuotient theCanonicalPER)  (a : A)  a ⊩[ theSubQuotient ] sq  a ⊩[ asmX ] (invert .map sq)
+    realizability sq a a⊩sq =
+      SQ.elimProp
+        {P = motive}
+        isPropMotive
+        elemMotive
+        sq a a⊩sq where
+
+      motive : (sq : subQuotient theCanonicalPER)  Type _
+      motive sq =  (a : A)  a ⊩[ theSubQuotient ] sq  a ⊩[ asmX ] (invert .map sq)
+
+      isPropMotive :  sq  isProp (motive sq)
+      isPropMotive sq = isPropΠ2 λ a a⊩sq  asmX .⊩isPropValued _ _
+
+      elemMotive : (x : domain theCanonicalPER)  motive [ x ]
+      elemMotive (r , x , r⊩x , _) a (x' , a⊩x' , r⊩x') = subst (a ⊩[ asmX ]_) (isModestAsmX x' x r r⊩x' r⊩x) a⊩x'
+
+    goal : (sq : subQuotient theCanonicalPER)  (a : A)  a ⊩[ theSubQuotient ] sq  (Id  a) ⊩[ asmX ] (invert .map sq)
+    goal sq a a⊩sq = subst (_⊩[ asmX ] _) (sym (Ida≡a a)) (realizability sq a a⊩sq)
+
+  forward : AssemblyMorphism asmX theSubQuotient
+  AssemblyMorphism.map forward x = subquot module Forward where
+    mainMap : Σ[ a  A ] (a ⊩[ asmX ] x)  subQuotient theCanonicalPER
+    mainMap (a , a⊩x) = [ a , x , a⊩x , a⊩x ]
+ 
+    mainMap2Constant : 2-Constant mainMap
+    mainMap2Constant (a , a⊩x) (b , b⊩x) = eq/ _ _ (x , a⊩x , b⊩x)
+
+    subquot : subQuotient theCanonicalPER
+    subquot = PT.rec→Set squash/ mainMap mainMap2Constant (asmX .⊩surjective x)
+  AssemblyMorphism.tracker forward =
+    return
+      (Id ,
+       x a a⊩x 
+        PT.elim
+          {P = λ surj  (Id  a) ⊩[ theSubQuotient ] (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)}
+           surj  theSubQuotient .⊩isPropValued (Id  a) (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj))
+           { (b , b⊩x)  x , subst (_⊩[ asmX ] x) (sym (Ida≡a a)) a⊩x , b⊩x })
+          (asmX .⊩surjective x)))
+
+  subQuotientCanonicalIso : CatIso MOD (X , asmX , isModestAsmX) (subQuotient theCanonicalPER , theSubQuotient , isModestSubQuotientAssembly theCanonicalPER)
+  fst subQuotientCanonicalIso = forward
+  isIso.inv (snd subQuotientCanonicalIso) = invert
+  isIso.sec (snd subQuotientCanonicalIso) = goal where
+    opaque
+      pointwise :  sq  (invert  forward) .map sq  sq
+      pointwise sq =
+        SQ.elimProp
+           sq  squash/ (forward .map (invert .map sq)) sq)
+           { d@(a , x , a⊩x , a⊩'x) 
+            PT.elim
+              {P = λ surj  PT.rec→Set squash/ (Forward.mainMap (Invert.reprAction [ d ] d)) (Forward.mainMap2Constant (Invert.reprAction [ d ] d)) surj  [ d ]}
+               surj  squash/ _ _)
+               { (b , b⊩x)  eq/ _ _ (x , b⊩x , a⊩x) })
+              (asmX .⊩surjective x) })
+          sq
+
+    goal : invert  forward  identityMorphism theSubQuotient
+    goal = AssemblyMorphism≡ _ _ (funExt pointwise)
+  isIso.ret (snd subQuotientCanonicalIso) = goal where
+    opaque
+      pointwise :  x  (forward  invert) .map x  x
+      pointwise x =
+        PT.elim
+          {P =
+            λ surj 
+              invert .map
+                (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)
+               x}
+           surj  asmX .isSetX _ _)
+           { (a , a⊩x)  refl })
+          (asmX .⊩surjective x)
+
+    goal : forward  invert  identityMorphism asmX
+    goal = AssemblyMorphism≡ _ _ (funExt pointwise)
+
\ No newline at end of file diff --git a/docs/Realizability.Modest.UniformFamily.html b/docs/Realizability.Modest.UniformFamily.html new file mode 100644 index 0000000..f91aa3a --- /dev/null +++ b/docs/Realizability.Modest.UniformFamily.html @@ -0,0 +1,219 @@ + +Realizability.Modest.UniformFamily
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Categories.Displayed.Reasoning
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Functor hiding (Id)
+open import Cubical.Categories.Constructions.Slice
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.UniformFamily {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.Terminal ca
+open import Realizability.Modest.Base ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+
+record UniformFamily {I : Type } (asmI : Assembly I) : Type (ℓ-suc ) where
+  no-eta-equality
+  field
+    carriers : I  Type 
+    assemblies :  i  Assembly (carriers i)
+    isModestFamily :  i  isModest (assemblies i)
+open UniformFamily
+record DisplayedUFamMap {I J : Type } (asmI : Assembly I) (asmJ : Assembly J) (u : AssemblyMorphism asmI asmJ) (X : UniformFamily asmI) (Y : UniformFamily asmJ) : Type  where
+  no-eta-equality
+  field
+    fibrewiseMap :  i  X .carriers i  Y .carriers (u .map i)
+    isTracked : ∃[ e  A ] (∀ (i : I) (a : A) (a⊩i : a ⊩[ asmI ] i) (x : X .carriers i) (b : A) (b⊩x : b ⊩[ X .assemblies i ] x)  (e  a  b) ⊩[ Y .assemblies (u .map i) ] (fibrewiseMap i x))
+
+open DisplayedUFamMap
+
+DisplayedUFamMapPathP :
+   {I J} (asmI : Assembly I) (asmJ : Assembly J) 
+   u v X Y
+   (uᴰ : DisplayedUFamMap asmI asmJ u X Y)
+   (vᴰ : DisplayedUFamMap asmI asmJ v X Y)
+   (p : u  v)
+   (∀ (i : I) (x : X .carriers i)  PathP  j  Y .carriers (p j .map i)) (uᴰ .fibrewiseMap i x) (vᴰ .fibrewiseMap i x))
+  -----------------------------------------------------------------------------------------------------------------------
+   PathP  i  DisplayedUFamMap asmI asmJ (p i) X Y) uᴰ vᴰ
+fibrewiseMap (DisplayedUFamMapPathP {I} {J} asmI asmJ u v X Y uᴰ vᴰ p pᴰ dimI) i x = pᴰ i x dimI
+isTracked (DisplayedUFamMapPathP {I} {J} asmI asmJ u v X Y uᴰ vᴰ p pᴰ dimI) =
+  isProp→PathP
+    {B = λ dimJ  ∃[ e  A ] ((i : I) (a : A)  a ⊩[ asmI ] i  (x : X .carriers i) (b : A)  b ⊩[ X .assemblies i ] x  (e  a  b) ⊩[ Y .assemblies (p dimJ .map i) ] pᴰ i x dimJ)}
+     dimJ  isPropPropTrunc)
+    (uᴰ .isTracked)
+    (vᴰ .isTracked)
+    dimI
+
+isSetDisplayedUFamMap :  {I J} (asmI : Assembly I) (asmJ : Assembly J)   u X Y  isSet (DisplayedUFamMap asmI asmJ u X Y)
+fibrewiseMap (isSetDisplayedUFamMap {I} {J} asmI asmJ u X Y f g p q dimI dimJ) i x =
+  isSet→isSet'
+    (Y .assemblies (u .map i) .isSetX)
+    {a₀₀ = fibrewiseMap f i x}
+    {a₀₁ = fibrewiseMap f i x}
+    refl
+    {a₁₀ = fibrewiseMap g i x}
+    {a₁₁ = fibrewiseMap g i x}
+    refl
+     dimK  fibrewiseMap (p dimK) i x)
+     dimK  fibrewiseMap (q dimK) i x)
+    dimJ dimI
+isTracked (isSetDisplayedUFamMap {I} {J} asmI asmJ u X Y f g p q dimI dimJ) =
+  isProp→SquareP
+    {B = λ dimI dimJ 
+      ∃[ e  A ]
+        ((i : I) (a : A) 
+         a ⊩[ asmI ] i 
+         (x : X .carriers i) (b : A) 
+         b ⊩[ X .assemblies i ] x 
+         (e  a  b) ⊩[ Y .assemblies (u .map i) ]
+         isSet→isSet'
+         (Y .assemblies
+          (u .map i)
+          .isSetX)
+          _  fibrewiseMap f i x)  _  fibrewiseMap g i x)
+          dimK  fibrewiseMap (p dimK) i x)
+          dimK  fibrewiseMap (q dimK) i x) dimJ dimI)}
+       dimI dimJ  isPropPropTrunc)
+      {a = isTracked f}
+      {b = isTracked g}
+      {c = isTracked f}
+      {d = isTracked g}
+      refl
+      refl
+       dimK  isTracked (p dimK))
+       dimK  isTracked (q dimK))
+      dimI dimJ
+
+DisplayedUFamMapPathPIso :
+   {I J} (asmI : Assembly I) (asmJ : Assembly J) 
+   u v X Y
+   (uᴰ : DisplayedUFamMap asmI asmJ u X Y)
+   (vᴰ : DisplayedUFamMap asmI asmJ v X Y)
+   (p : u  v)
+   Iso
+    (∀ (i : I) (x : X .carriers i)  PathP  dimI  Y .carriers (p dimI .map i)) (uᴰ .fibrewiseMap i x) (vᴰ .fibrewiseMap i x))
+    (PathP  dimI  DisplayedUFamMap asmI asmJ (p dimI) X Y) uᴰ vᴰ)
+Iso.fun (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) pᴰ = DisplayedUFamMapPathP asmI asmJ u v X Y uᴰ vᴰ p pᴰ
+Iso.inv (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) uᴰ≡vᴰ i x dimI = (uᴰ≡vᴰ dimI) .fibrewiseMap i x
+Iso.rightInv (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) uᴰ≡vᴰ dimI dimJ =
+  isSet→SquareP
+    {A = λ dimK dimL  DisplayedUFamMap asmI asmJ (p dimL) X Y}
+     dimI dimJ  isSetDisplayedUFamMap asmI asmJ (p dimJ) X Y)
+    {a₀₀ = uᴰ}
+    {a₀₁ = vᴰ}
+     dimK  DisplayedUFamMapPathP asmI asmJ u v X Y uᴰ vᴰ p  i x dimL  uᴰ≡vᴰ dimL .fibrewiseMap i x) dimK)
+    {a₁₀ = uᴰ}
+    {a₁₁ = vᴰ}
+    uᴰ≡vᴰ
+    refl
+    refl dimI dimJ
+Iso.leftInv (DisplayedUFamMapPathPIso {I} {J} asmI asmJ u v X Y uᴰ vᴰ p) pᴰ = refl
+
+idDisplayedUFamMap :  {I} (asmI : Assembly I) (p : UniformFamily asmI)  DisplayedUFamMap asmI asmI (identityMorphism asmI) p p
+DisplayedUFamMap.fibrewiseMap (idDisplayedUFamMap {I} asmI p) i pi = pi
+DisplayedUFamMap.isTracked (idDisplayedUFamMap {I} asmI p) =
+  return
+    (λ*2 realizer ,
+     λ i a a⊩i x b b⊩x 
+       subst
+          r  r ⊩[ p .assemblies i ] x)
+         (sym (λ*2ComputationRule realizer a b))
+         b⊩x) where
+  realizer : Term as 2
+  realizer = # zero
+
+module _
+  {I J K : Type }
+  (asmI : Assembly I)
+  (asmJ : Assembly J)
+  (asmK : Assembly K)
+  (f : AssemblyMorphism asmI asmJ)
+  (g : AssemblyMorphism asmJ asmK)
+  (X : UniformFamily asmI)
+  (Y : UniformFamily asmJ)
+  (Z : UniformFamily asmK)
+  (fᴰ : DisplayedUFamMap asmI asmJ f X Y)
+  (gᴰ : DisplayedUFamMap asmJ asmK g Y Z) where
+
+  composeDisplayedUFamMap : DisplayedUFamMap asmI asmK (f  g) X Z
+  DisplayedUFamMap.fibrewiseMap composeDisplayedUFamMap i Xi = gᴰ .fibrewiseMap (f .map i) (fᴰ .fibrewiseMap i Xi)
+  DisplayedUFamMap.isTracked composeDisplayedUFamMap =
+    do
+      (gᴰ~ , isTrackedGᴰ)  gᴰ .isTracked
+      (fᴰ~ , isTrackedFᴰ)  fᴰ .isTracked
+      (f~ , isTrackedF)  f .tracker
+      let
+        realizer : Term as 2
+        realizer = ` gᴰ~ ̇ (` f~ ̇ # one) ̇ (` fᴰ~ ̇ # one ̇ # zero)
+      return
+        (λ*2 realizer ,
+         i a a⊩i x b b⊩x 
+          subst
+            (_⊩[ Z .assemblies (g .map (f .map i)) ] _)
+            (sym (λ*2ComputationRule realizer a b))
+            (isTrackedGᴰ (f .map i) (f~  a) (isTrackedF i a a⊩i) (fᴰ .fibrewiseMap i x) (fᴰ~  a  b) (isTrackedFᴰ i a a⊩i x b b⊩x))))
+
+UNIMOD : Categoryᴰ ASM (ℓ-suc ) 
+Categoryᴰ.ob[ UNIMOD ] (I , asmI) = UniformFamily asmI
+Categoryᴰ.Hom[_][_,_] UNIMOD {I , asmI} {J , asmJ} u X Y = DisplayedUFamMap asmI asmJ u X Y
+Categoryᴰ.idᴰ UNIMOD {I , asmI} {X} = idDisplayedUFamMap asmI X
+Categoryᴰ._⋆ᴰ_ UNIMOD {I , asmI} {J , asmJ} {K , asmK} {f} {g} {X} {Y} {Z} fᴰ gᴰ = composeDisplayedUFamMap asmI asmJ asmK f g X Y Z fᴰ gᴰ
+Categoryᴰ.⋆IdLᴰ UNIMOD {I , asmI} {J , asmJ} {f} {X} {Y} fᴰ =
+  DisplayedUFamMapPathP
+    asmI asmJ
+    (identityMorphism asmI  f) f
+    X Y
+    (composeDisplayedUFamMap asmI asmI asmJ (Category.id ASM) f X X Y (idDisplayedUFamMap asmI X) fᴰ)
+    fᴰ
+    (Category.⋆IdL ASM f)
+     i x  refl)
+Categoryᴰ.⋆IdRᴰ UNIMOD {I , asmI} {J , asmJ} {f} {X} {Y} fᴰ =
+  DisplayedUFamMapPathP
+    asmI asmJ
+    (f  identityMorphism asmJ) f
+    X Y
+    (composeDisplayedUFamMap asmI asmJ asmJ f (Category.id ASM) X Y Y fᴰ (idDisplayedUFamMap asmJ Y))
+    fᴰ
+    (Category.⋆IdR ASM f)
+    λ i x  refl
+Categoryᴰ.⋆Assocᴰ UNIMOD {I , asmI} {J , asmJ} {K , asmK} {L , asmL} {f} {g} {h} {X} {Y} {Z} {W} fᴰ gᴰ hᴰ =
+  DisplayedUFamMapPathP
+    asmI asmL
+    ((f  g)  h) (f  (g  h))
+    X W
+    (composeDisplayedUFamMap
+      asmI asmK asmL
+      (f  g) h X Z W
+      (composeDisplayedUFamMap asmI asmJ asmK f g X Y Z fᴰ gᴰ)
+      hᴰ)
+    (composeDisplayedUFamMap
+      asmI asmJ asmL
+      f (g  h) X Y W
+      fᴰ (composeDisplayedUFamMap asmJ asmK asmL g h Y Z W gᴰ hᴰ))
+    (Category.⋆Assoc ASM f g h)
+    λ i x  refl
+Categoryᴰ.isSetHomᴰ UNIMOD {I , asmI} {J , asmJ} {f} {X} {Y} = isSetDisplayedUFamMap asmI asmJ f X Y
+
\ No newline at end of file diff --git a/docs/Realizability.Modest.UniformFamilyCleavage.html b/docs/Realizability.Modest.UniformFamilyCleavage.html new file mode 100644 index 0000000..5bd01ed --- /dev/null +++ b/docs/Realizability.Modest.UniformFamilyCleavage.html @@ -0,0 +1,102 @@ + +Realizability.Modest.UniformFamilyCleavage
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Categories.Displayed.Reasoning
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Functor hiding (Id)
+open import Cubical.Categories.Constructions.Slice
+open import Categories.CartesianMorphism
+open import Categories.GenericObject
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.UniformFamilyCleavage {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.Terminal ca
+open import Realizability.Assembly.SetsReflectiveSubcategory ca
+open import Realizability.Modest.Base ca
+open import Realizability.Modest.UniformFamily ca
+open import Realizability.Modest.CanonicalPER ca
+open import Realizability.Modest.UnresizedGeneric ca
+open import Realizability.PERs.PER ca
+open import Realizability.PERs.SubQuotient ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Contravariant UNIMOD
+open UniformFamily
+open DisplayedUFamMap
+
+uniformFamilyCleavage : cleavage
+uniformFamilyCleavage {X , asmX} {Y , asmY} f N =
+  N' , fᴰ , cartfᴰ where
+    N' : UniformFamily asmX
+    UniformFamily.carriers N' x = N .carriers (f .map x)
+    UniformFamily.assemblies N' x = N .assemblies (f .map x)
+    UniformFamily.isModestFamily N' x = N .isModestFamily (f .map x)
+
+    fᴰ : DisplayedUFamMap asmX asmY f N' N
+    DisplayedUFamMap.fibrewiseMap fᴰ x Nfx = Nfx
+    DisplayedUFamMap.isTracked fᴰ =
+      do
+        let
+          realizer : Term as 2
+          realizer = # zero
+        return
+          (λ*2 realizer ,
+           x a a⊩x Nfx b b⊩Nfx 
+            subst
+              (_⊩[ N .assemblies (f .map x) ] Nfx)
+              (sym (λ*2ComputationRule realizer a b))
+              b⊩Nfx))
+
+    opaque
+      unfolding isCartesian'
+      cart'fᴰ : isCartesian' f fᴰ
+      cart'fᴰ {Z , asmZ} {M} g hᴰ =
+        (! , !⋆fᴰ≡hᴰ) ,
+        λ { (!' , !'comm) 
+          Σ≡Prop
+             !  UNIMOD .Categoryᴰ.isSetHomᴰ _ _)
+            (DisplayedUFamMapPathP
+              _ _ _ _ _ _ _ _ _
+              λ z Mz 
+                sym
+                  (!' .fibrewiseMap z Mz
+                    ≡[ i ]⟨ !'comm i .fibrewiseMap z Mz 
+                  hᴰ .fibrewiseMap z Mz
+                    )) } where
+          ! : DisplayedUFamMap asmZ asmX g M N'
+          DisplayedUFamMap.fibrewiseMap ! z Mz = hᴰ .fibrewiseMap z Mz
+          DisplayedUFamMap.isTracked ! = hᴰ .isTracked
+
+          !⋆fᴰ≡hᴰ : composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ  hᴰ
+          !⋆fᴰ≡hᴰ =
+            DisplayedUFamMapPathP
+              asmZ asmY _ _
+              M N
+              (composeDisplayedUFamMap asmZ asmX asmY g f M N' N ! fᴰ) hᴰ refl
+              λ z Mz  refl
+
+    cartfᴰ : isCartesian f fᴰ
+    cartfᴰ = isCartesian'→isCartesian f fᴰ cart'fᴰ
+
\ No newline at end of file diff --git a/docs/Realizability.Modest.UnresizedGeneric.html b/docs/Realizability.Modest.UnresizedGeneric.html new file mode 100644 index 0000000..4fef522 --- /dev/null +++ b/docs/Realizability.Modest.UnresizedGeneric.html @@ -0,0 +1,94 @@ + +Realizability.Modest.UnresizedGeneric
open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Categories.Displayed.Reasoning
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Functor hiding (Id)
+open import Cubical.Categories.Constructions.Slice
+open import Categories.CartesianMorphism
+open import Categories.GenericObject
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.UnresizedGeneric {} {A : Type } (ca : CombinatoryAlgebra A) (resizing : hPropResizing ) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.Terminal ca
+open import Realizability.Assembly.SetsReflectiveSubcategory ca
+open import Realizability.Modest.Base ca
+open import Realizability.Modest.UniformFamily ca
+open import Realizability.Modest.CanonicalPER ca
+open import Realizability.PERs.PER ca
+open import Realizability.PERs.ResizedPER ca resizing
+open import Realizability.PERs.SubQuotient ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Contravariant UNIMOD
+open UniformFamily
+open DisplayedUFamMap
+
+module Unresized
+  {X : Type }
+  (asmX : Assembly X)
+  (M : UniformFamily asmX) where
+
+  theCanonicalPER :  x  PER
+  theCanonicalPER x = canonicalPER (M . assemblies x) (M .isModestFamily x)
+
+  elimRealizerForMx :  (x : X) (Mx : M .carriers x)  Σ[ a  A ] (a ⊩[ M .assemblies x ] Mx)  subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x))
+  elimRealizerForMx x Mx (a , a⊩Mx) = [ a , Mx , a⊩Mx , a⊩Mx ]
+
+  opaque
+    elimRealizerForMx2Constant :  x Mx  2-Constant (elimRealizerForMx x Mx)
+    elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx) =
+      eq/
+        (a , Mx , a⊩Mx , a⊩Mx)
+        (b , Mx , b⊩Mx , b⊩Mx)
+        (Mx , a⊩Mx , b⊩Mx)
+
+  mainMapType : Type _
+  mainMapType =
+     (x : X) (Mx : M .carriers x) 
+    Σ[ out  (subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x))) ]
+    (∀ (a : A)  a ⊩[ asmX ] x  (b : A)  b ⊩[ M .assemblies x ] Mx  (λ*2 (# zero)  a  b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] out)
+
+  opaque
+    mainMap : mainMapType
+    mainMap x Mx =
+      PT.rec→Set
+        (isSetΣ
+            squash/
+             out 
+              isSetΠ3
+                λ a a⊩x b 
+                  isSet→
+                    (isProp→isSet
+                      (str
+                        (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero)  a  b) out)))))
+        ((λ { (c , c⊩Mx) 
+          (elimRealizerForMx x Mx (c , c⊩Mx)) ,
+           a a⊩x b b⊩Mx 
+            subst (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx x Mx (c , c⊩Mx))) (sym (λ*2ComputationRule (# zero) a b)) (Mx , b⊩Mx , c⊩Mx)) }))
+         { (a , a⊩Mx) (b , b⊩Mx) 
+          Σ≡Prop  out  isPropΠ4 λ a a⊩x b b⊩Mx  str (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero)  a  b) out)) (elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx)) })
+        (M .assemblies x .⊩surjective Mx)
+
\ No newline at end of file diff --git a/docs/Realizability.PERs.Everything.html b/docs/Realizability.PERs.Everything.html new file mode 100644 index 0000000..72d56c1 --- /dev/null +++ b/docs/Realizability.PERs.Everything.html @@ -0,0 +1,7 @@ + +Realizability.PERs.Everything
module Realizability.PERs.Everything where
+
+open import Realizability.PERs.PER
+open import Realizability.PERs.ResizedPER
+open import Realizability.PERs.SubQuotient
+
\ No newline at end of file diff --git a/docs/Realizability.PERs.PER.html b/docs/Realizability.PERs.PER.html new file mode 100644 index 0000000..afe7afe --- /dev/null +++ b/docs/Realizability.PERs.PER.html @@ -0,0 +1,226 @@ + +Realizability.PERs.PER
open import Realizability.ApplicativeStructure
+open import Realizability.CombinatoryAlgebra
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Powerset
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Relation.Binary
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor hiding (Id)
+
+module Realizability.PERs.PER
+  {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+
+open CombinatoryAlgebra ca
+open Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+
+module BR = BinaryRelation
+
+isPartialEquivalenceRelation : (A  A  Type )  Type _
+isPartialEquivalenceRelation rel = BR.isSym rel × BR.isTrans rel
+
+isPropIsPartialEquivalenceRelation :  r  (∀ a b  isProp (r a b))  isProp (isPartialEquivalenceRelation r)
+isPropIsPartialEquivalenceRelation rel isPropValuedRel =
+  isProp×
+    (isPropΠ  x  isPropΠ λ y  isProp→ (isPropValuedRel y x)))
+    (isPropΠ λ x  isPropΠ λ y  isPropΠ λ z  isProp→ (isProp→ (isPropValuedRel x z)))
+
+record PER : Type (ℓ-suc ) where
+  no-eta-equality
+  constructor makePER
+  field
+    relation : A  A  Type 
+    isPropValued :  a b  isProp (relation a b)
+    isPER : isPartialEquivalenceRelation relation
+  isSymmetric = isPER .fst
+  isTransitive = isPER .snd
+
+open PER
+
+PERΣ : Type (ℓ-suc )
+PERΣ = Σ[ relation  (A  A  hProp ) ] isPartialEquivalenceRelation λ a b   relation a b 
+
+isSetPERΣ : isSet PERΣ
+isSetPERΣ =
+  isSetΣ
+    (isSet→ (isSet→ isSetHProp))
+     relation 
+      isProp→isSet
+        (isPropIsPartialEquivalenceRelation
+           a b   relation a b )
+           a b  str (relation a b))))
+
+PER≡ :  (R S : PER)  (R .relation  S .relation)  R  S
+relation (PER≡ R S rel≡ i) = rel≡ i
+isPropValued (PER≡ R S rel≡ i) a b =
+  isProp→PathP
+    {B = λ j  isProp (rel≡ j a b)}
+     j  isPropIsProp)
+    (R .isPropValued a b)
+    (S .isPropValued a b) i
+isPER (PER≡ R S rel≡ i) =
+  isProp→PathP
+    {B = λ j  isPartialEquivalenceRelation (rel≡ j)}
+     j  isPropIsPartialEquivalenceRelation (rel≡ j) λ a b  isPropRelJ a b j)
+    (R .isPER)
+    (S .isPER) i where
+      isPropRelJ :  a b j  isProp (rel≡ j a b)
+      isPropRelJ a b j = isProp→PathP {B = λ k  isProp (rel≡ k a b)}  k  isPropIsProp) (R .isPropValued a b) (S .isPropValued a b) j
+
+PERIsoΣ : Iso PER PERΣ
+Iso.fun PERIsoΣ per =  a b  per .relation a b , per .isPropValued a b) , per .isPER
+relation (Iso.inv PERIsoΣ perΣ) a b =  perΣ .fst a b 
+isPropValued (Iso.inv PERIsoΣ perΣ) a b = str (perΣ .fst a b)
+isPER (Iso.inv PERIsoΣ perΣ) = perΣ .snd
+Iso.rightInv PERIsoΣ perΣ = refl
+Iso.leftInv PERIsoΣ per = PER≡ _ _ refl
+
+isSetPER : isSet PER
+isSetPER = isOfHLevelRetractFromIso 2 PERIsoΣ isSetPERΣ
+
+PER≡Iso :  (R S : PER)  Iso (R  S) (R .relation  S .relation)
+Iso.fun (PER≡Iso R S) R≡S i = R≡S i .relation
+Iso.inv (PER≡Iso R S) rel≡ = PER≡ R S rel≡
+Iso.rightInv (PER≡Iso R S) rel≡ = refl
+Iso.leftInv (PER≡Iso R S) R≡S = isSetPER R S _ _
+
+_~[_]_ : A  PER  A  Type 
+a ~[ R ] b = R .relation a b
+
+isProp~ :  a R b  isProp (a ~[ R ] b)
+isProp~ a R b = R .isPropValued a b
+
+isTracker : (R S : PER)  A  Type 
+isTracker R S a =  r r'  r ~[ R ] r'  (a  r) ~[ S ] (a  r')
+
+perTracker : (R S : PER)  Type 
+perTracker R S = Σ[ a  A ] isTracker R S a
+
+isEquivTracker : (R S : PER)  perTracker R S  perTracker R S  Type 
+isEquivTracker R S (a , _) (b , _) = (∀ r  r ~[ R ] r  (a  r) ~[ S ] (b  r))
+
+isEquivRelIsEquivTracker : (R S : PER)  BR.isEquivRel (isEquivTracker R S)
+BinaryRelation.isEquivRel.reflexive (isEquivRelIsEquivTracker R S) (a , isTrackerA) r r~r = isTrackerA r r r~r
+BinaryRelation.isEquivRel.symmetric (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) a~b r r~r =
+  isSymmetric S (a  r) (b  r) (a~b r r~r)
+BinaryRelation.isEquivRel.transitive (isEquivRelIsEquivTracker R S) (a , isTrackerA) (b , isTrackerB) (c , isTrackerC) a~b b~c r r~r =
+  isTransitive S (a  r) (b  r) (c  r) (a~b r r~r) (b~c r r~r)
+
+isPropIsEquivTracker :  R S a b  isProp (isEquivTracker R S a b)
+isPropIsEquivTracker R S (a , isTrackerA) (b , isTrackerB) = isPropΠ2 λ r r~r  isPropValued S (a  r) (b  r)
+
+isEffectiveIsEquivTracker :  R S  BR.isEffective (isEquivTracker R S)
+isEffectiveIsEquivTracker R S = isEquivRel→isEffective (isPropIsEquivTracker R S) (isEquivRelIsEquivTracker R S)
+
+perMorphism : (R S : PER)  Type 
+perMorphism R S = perTracker R S / (isEquivTracker R S)
+
+effectiveIsEquivTracker :  R S a b  [ a ]  [ b ]  isEquivTracker R S a b
+effectiveIsEquivTracker R S a b eq' = effective (isPropIsEquivTracker R S) (isEquivRelIsEquivTracker R S) a b eq'
+
+isSetPerMorphism :  R S  isSet (perMorphism R S)
+isSetPerMorphism R S = squash/
+
+idPerMorphism : (R : PER)  perMorphism R R
+idPerMorphism R = [ Id ,  r r' r~r'  subst2  r r'  r ~[ R ] r') (sym (Ida≡a r)) (sym (Ida≡a r')) r~r') ]
+
+composePerTracker : (R S T : PER)  perTracker R S  perTracker S T  perTracker R T
+composePerTracker R S T (a , a⊩f) (b , b⊩g) =
+  let
+    realizer : Term as 1
+    realizer = ` b ̇ (` a ̇ # zero)
+  in
+  λ* realizer ,
+  λ r r' r~r' 
+    subst2
+      _~[ T ]_
+      (sym (λ*ComputationRule realizer r))
+      (sym (λ*ComputationRule realizer r'))
+      (b⊩g (a  r) (a  r') (a⊩f r r' r~r'))
+
+composePerMorphism : (R S T : PER)  perMorphism R S  perMorphism S T  perMorphism R T
+composePerMorphism R S T f g =
+  SQ.rec2
+    squash/
+     { (a , a⊩f) (b , b⊩g) 
+      [ composePerTracker R S T (a , a⊩f) (b , b⊩g) ] })
+     { (a , a⊩f) (b , b⊩f) (c , c⊩g) a~b 
+      eq/ _ _
+        λ r r~r 
+          subst2
+             car cbr  car ~[ T ] cbr)
+            (sym (λ*ComputationRule (` c ̇ (` a ̇ # zero)) r))
+            (sym (λ*ComputationRule (` c ̇ (` b ̇ # zero)) r))
+            (c⊩g (a  r) (b  r) (a~b r r~r)) })
+     { (a , a⊩f) (b , b⊩g) (c , c⊩g) b~c 
+      eq/ _ _
+        λ r r~r 
+          subst2
+             bar car  bar ~[ T ] car)
+            (sym (λ*ComputationRule (` b ̇ (` a ̇ # zero)) r))
+            (sym (λ*ComputationRule (` c ̇ (` a ̇ # zero)) r))
+            (b~c (a  r) (a⊩f r r r~r)) })
+    f g
+
+idLPerMorphism :  R S f  composePerMorphism R R S (idPerMorphism R) f  f
+idLPerMorphism R S f =
+  SQ.elimProp
+     f  squash/ (composePerMorphism R R S (idPerMorphism R) f) f)
+     { (a , a⊩f) 
+      eq/ _ _
+      λ r r~r 
+        subst
+           ar  ar ~[ S ] (a  r))
+          (sym (λ*ComputationRule (` a ̇ (` Id ̇ # zero)) r  cong  x  a  x) (Ida≡a r)))
+          (a⊩f r r r~r) })
+    f
+
+idRPerMorphism :  R S f  composePerMorphism R S S f (idPerMorphism S)  f
+idRPerMorphism R S f =
+  SQ.elimProp
+     f  squash/ (composePerMorphism R S S f (idPerMorphism S)) f)
+     { (a , a⊩f) 
+      eq/ _ _
+        λ r r~r 
+          subst  ar  ar ~[ S ] (a  r)) (sym (λ*ComputationRule (` Id ̇ (` a ̇ # zero)) r  Ida≡a (a  r))) (a⊩f r r r~r) })
+    f
+
+assocPerMorphism :  R S T U f g h  composePerMorphism R T U (composePerMorphism R S T f g) h  composePerMorphism R S U f (composePerMorphism S T U g h)
+assocPerMorphism R S T U f g h =
+  SQ.elimProp3
+     f g h  squash/ (composePerMorphism R T U (composePerMorphism R S T f g) h) (composePerMorphism R S U f (composePerMorphism S T U g h)))
+     { (a , a⊩f) (b , b⊩g) (c , c⊩h) 
+      eq/ _ _
+        λ r r~r 
+          subst2
+             cba cba'  cba ~[ U ] cba')
+            (sym (λ*ComputationRule (` c ̇ (`  as  (λ*abst (` b ̇ (` a ̇ # zero))) []  ̇ # zero)) r  cong  bar  c  bar) (λ*ComputationRule (` b ̇ (` a ̇ # zero)) r)))
+            (sym (λ*ComputationRule (`  as  (λ*abst (` c ̇ (` b ̇ # zero))) [] ̇ (` a ̇ # zero)) r  λ*ComputationRule (` c ̇ (` b ̇ # zero)) (a  r)))
+            (c⊩h (b  (a  r)) (b  (a  r)) (b⊩g (a  r) (a  r) (a⊩f r r r~r)) ) })
+    f g h
+
+PERCat : Category (ℓ-suc ) 
+Category.ob PERCat = PER
+Category.Hom[_,_] PERCat R S = perMorphism R S
+Category.id PERCat {R} = idPerMorphism R
+Category._⋆_ PERCat {R} {S} {T} f g = composePerMorphism R S T f g
+Category.⋆IdL PERCat {R} {S} f = idLPerMorphism R S f
+Category.⋆IdR PERCat {R} {S} f = idRPerMorphism R S f
+Category.⋆Assoc PERCat {R} {S} {T} {U} f g h = assocPerMorphism R S T U f g h
+Category.isSetHom PERCat {R} {S} = isSetPerMorphism R S
+
\ No newline at end of file diff --git a/docs/Realizability.PERs.ResizedPER.html b/docs/Realizability.PERs.ResizedPER.html new file mode 100644 index 0000000..7b8bf8c --- /dev/null +++ b/docs/Realizability.PERs.ResizedPER.html @@ -0,0 +1,197 @@ + +Realizability.PERs.ResizedPER
open import Realizability.ApplicativeStructure
+open import Realizability.CombinatoryAlgebra
+open import Realizability.PropResizing
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Path
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Relation.Binary
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor hiding (Id)
+
+module Realizability.PERs.ResizedPER
+  {} {A : Type } (ca : CombinatoryAlgebra A) (resizing : hPropResizing ) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.PERs.PER ca
+open import Realizability.Modest.Base ca
+
+open CombinatoryAlgebra ca
+open Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+
+smallHProp = resizing .fst
+hProp≃smallHProp = resizing .snd
+smallHProp≃hProp = invEquiv hProp≃smallHProp
+
+isSetSmallHProp : isSet smallHProp
+isSetSmallHProp = isOfHLevelRespectEquiv 2 hProp≃smallHProp isSetHProp
+
+hPropIsoSmallHProp : Iso (hProp ) smallHProp
+hPropIsoSmallHProp = equivToIso hProp≃smallHProp
+
+shrink : hProp   smallHProp
+shrink = Iso.fun hPropIsoSmallHProp
+
+enlarge : smallHProp  hProp 
+enlarge = Iso.inv hPropIsoSmallHProp
+
+enlarge⋆shrink≡id : section shrink enlarge
+enlarge⋆shrink≡id = Iso.rightInv hPropIsoSmallHProp
+
+shrink⋆enlarge≡id : retract shrink enlarge
+shrink⋆enlarge≡id = Iso.leftInv hPropIsoSmallHProp
+
+extractType : smallHProp  Type 
+extractType p =  enlarge p 
+
+isPropExtractType :  p  isProp (extractType p)
+isPropExtractType p = str (enlarge p)
+
+extractFromShrunk :  p isPropP  extractType (shrink (p , isPropP))  p
+extractFromShrunk p isPropP =
+  extractType (shrink (p , isPropP))
+    ≡⟨ refl 
+   enlarge (shrink (p , isPropP)) 
+    ≡⟨ cong ⟨_⟩ (shrink⋆enlarge≡id (p , isPropP)) 
+  p
+    
+
+shrinkFromExtracted :  p  shrink (extractType p , isPropExtractType p)  p
+shrinkFromExtracted p =
+  shrink (extractType p , isPropExtractType p)
+    ≡⟨ refl 
+  shrink (enlarge p)
+    ≡⟨ enlarge⋆shrink≡id p 
+  p
+    
+
+record ResizedPER : Type  where
+  no-eta-equality
+  constructor makeResizedPER
+  field
+    relation : A  A  smallHProp
+    isSymmetric :  a b  extractType (relation a b)  extractType (relation b a)
+    isTransitive :  a b c  extractType (relation a b)  extractType (relation b c)  extractType (relation a c)
+
+open ResizedPER
+
+unquoteDecl ResizedPERIsoΣ = declareRecordIsoΣ ResizedPERIsoΣ (quote ResizedPER)
+
+ResizedPERΣ : Type 
+ResizedPERΣ =
+  Σ[ relation  (A  A  smallHProp) ]
+  (∀ a b  extractType (relation a b)  extractType (relation b a)) ×
+  (∀ a b c  extractType (relation a b)  extractType (relation b c)  extractType (relation a c))
+
+isSetResizedPERΣ : isSet ResizedPERΣ
+isSetResizedPERΣ =
+  isSetΣ
+    (isSet→ (isSet→ isSetSmallHProp))
+     relation  isProp→isSet (isProp× (isPropΠ3 λ _ _ _  isPropExtractType _) (isPropΠ5 λ _ _ _ _ _  isPropExtractType _)))
+
+isSetResizedPER : isSet ResizedPER
+isSetResizedPER = isOfHLevelRetractFromIso 2 ResizedPERIsoΣ isSetResizedPERΣ
+
+ResizedPER≡Iso :  (R S : ResizedPER)  Iso (R  S) (∀ a b  R .relation a b  S .relation a b)
+Iso.fun (ResizedPER≡Iso R S) R≡S a b i = (R≡S i) .relation a b
+relation (Iso.inv (ResizedPER≡Iso R S) pointwise i) a b = pointwise a b i
+isSymmetric (Iso.inv (ResizedPER≡Iso R S) pointwise i) =
+  isProp→PathP
+    {B = λ j  (a b : A)  extractType (pointwise a b j)  extractType (pointwise b a j)}
+     j  isPropΠ3 λ _ _ _  isPropExtractType _)
+    (isSymmetric R)
+    (isSymmetric S) i
+isTransitive (Iso.inv (ResizedPER≡Iso R S) pointwise i) =
+  isProp→PathP
+    {B = λ j  (a b c : A)  extractType (pointwise a b j)  extractType (pointwise b c j)  extractType (pointwise a c j)}
+     j  isPropΠ5 λ _ _ _ _ _  isPropExtractType _)
+    (R .isTransitive)
+    (S .isTransitive)
+    i
+Iso.rightInv (ResizedPER≡Iso R S) pointwise = refl
+Iso.leftInv (ResizedPER≡Iso R S) R≡S = isSetResizedPER R S _ _
+
+ResizedPER≡ :  (R S : ResizedPER)  (∀ a b  R .relation a b  S .relation a b)  R  S
+ResizedPER≡ R S pointwise = Iso.inv (ResizedPER≡Iso R S) pointwise
+
+ResizedPERIsoPER : Iso ResizedPER PER
+PER.relation (Iso.fun ResizedPERIsoPER resized) a b = extractType (resized .relation a b)
+PER.isPropValued (Iso.fun ResizedPERIsoPER resized) a b = isPropExtractType _
+fst (PER.isPER (Iso.fun ResizedPERIsoPER resized)) a b a~b = resized .isSymmetric a b a~b
+snd (PER.isPER (Iso.fun ResizedPERIsoPER resized)) a b c a~b b~c = resized .isTransitive a b c a~b b~c
+relation (Iso.inv ResizedPERIsoPER per) a b = shrink (per .PER.relation a b , per .PER.isPropValued a b)
+isSymmetric (Iso.inv ResizedPERIsoPER per) a b a~[resized]b = b~[resized]a where
+  a~b : per .PER.relation a b
+  a~b = transport (extractFromShrunk _ _) a~[resized]b
+
+  b~a : per .PER.relation b a
+  b~a = per .PER.isPER .fst a b a~b
+
+  b~[resized]a : extractType (shrink (per .PER.relation b a , per .PER.isPropValued b a))
+  b~[resized]a = transport (sym (extractFromShrunk _ _)) b~a
+isTransitive (Iso.inv ResizedPERIsoPER per) a b c a~[resized]b b~[resized]c = a~[resized]c where
+  a~b : per .PER.relation a b
+  a~b = transport (extractFromShrunk _ _) a~[resized]b
+
+  b~c : per .PER.relation b c
+  b~c = transport (extractFromShrunk _ _) b~[resized]c
+
+  a~c : per .PER.relation a c
+  a~c = per .PER.isPER .snd a b c a~b b~c
+
+  a~[resized]c : extractType (shrink (per .PER.relation a c , per .PER.isPropValued a c))
+  a~[resized]c = transport (sym (extractFromShrunk _ _)) a~c
+Iso.rightInv ResizedPERIsoPER per =
+  PER≡ _ _
+    (funExt₂
+      λ a b 
+        extractFromShrunk (per .PER.relation a b) (per .PER.isPropValued a b))
+Iso.leftInv ResizedPERIsoPER resizedPer =
+  ResizedPER≡ _ _
+    λ a b  shrinkFromExtracted (resizedPer .relation a b)
+
+opaque
+  shrinkPER : PER  ResizedPER
+  shrinkPER = ResizedPERIsoPER .Iso.inv
+
+opaque
+  enlargePER : ResizedPER  PER
+  enlargePER = ResizedPERIsoPER .Iso.fun
+
+opaque
+  unfolding shrinkPER
+  unfolding enlargePER
+  shrinkPER⋆enlargePER≡id :  resized  shrinkPER (enlargePER resized)  resized
+  shrinkPER⋆enlargePER≡id resized = ResizedPERIsoPER .Iso.leftInv resized
+
+opaque
+  unfolding shrinkPER
+  unfolding enlargePER
+  enlargePER⋆shrinkPER≡id :  per  enlargePER (shrinkPER per)  per
+  enlargePER⋆shrinkPER≡id per = ResizedPERIsoPER .Iso.rightInv per
+
+ResizedPER≃PER : ResizedPER  PER
+ResizedPER≃PER = isoToEquiv ResizedPERIsoPER
+
+opaque
+  transportFromSmall :  {ℓ'} {P : ResizedPER  Type ℓ'}  (∀ per  P (shrinkPER per))   resized  P resized
+  transportFromSmall {ℓ'} {P} small resized = subst P (shrinkPER⋆enlargePER≡id resized) (small (enlargePER resized))
+
+opaque
+  transportFromLarge :  {ℓ'} {P : PER  Type ℓ'}  (∀ resized  P (enlargePER resized))   per  P per
+  transportFromLarge {ℓ'} {P} large per = subst P (enlargePER⋆shrinkPER≡id per) (large (shrinkPER per))
+
\ No newline at end of file diff --git a/docs/Realizability.PERs.SubQuotient.html b/docs/Realizability.PERs.SubQuotient.html new file mode 100644 index 0000000..e244cc6 --- /dev/null +++ b/docs/Realizability.PERs.SubQuotient.html @@ -0,0 +1,271 @@ + +Realizability.PERs.SubQuotient
open import Realizability.ApplicativeStructure
+open import Realizability.CombinatoryAlgebra
+open import Realizability.PropResizing
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Function
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Functions.Embedding
+open import Cubical.Functions.Surjection
+open import Cubical.Relation.Binary
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor hiding (Id)
+
+module Realizability.PERs.SubQuotient
+  {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.PERs.PER ca
+open import Realizability.Modest.Base ca
+
+open CombinatoryAlgebra ca
+open Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+
+module _
+  (per : PER) where
+
+  domain : Type 
+  domain = Σ[ a  A ] (per .PER.relation a a)
+
+  subQuotient : Type 
+  subQuotient = domain / λ { (a , _) (b , _)  per .PER.relation a b }
+
+  subQuotientRealizability : A  subQuotient  hProp 
+  subQuotientRealizability r [a] =
+    SQ.rec
+      isSetHProp
+       { (a , a~a)  r ~[ per ] a , isProp~ r per a })
+       { (a , a~a) (b , b~b) a~b 
+        Σ≡Prop
+           x  isPropIsProp)
+          (hPropExt
+            (isProp~ r per a)
+            (isProp~ r per b)
+             r~a  PER.isTransitive per r a b r~a a~b)
+             r~b  PER.isTransitive per r b a r~b (PER.isSymmetric per a b a~b))) })
+      [a]
+      
+  
+  subQuotientAssembly : Assembly subQuotient
+  Assembly._⊩_ subQuotientAssembly r [a] =  subQuotientRealizability r [a] 
+  Assembly.isSetX subQuotientAssembly = squash/
+  Assembly.⊩isPropValued subQuotientAssembly r [a] = str (subQuotientRealizability r [a])
+  Assembly.⊩surjective subQuotientAssembly [a] =
+    SQ.elimProp
+      {P = λ [a]  ∃[ r  A ]  subQuotientRealizability r [a] }
+       [a]  isPropPropTrunc)
+       { (a , a~a)  return (a , a~a) })
+      [a]
+      
+  isModestSubQuotientAssembly : isModest subQuotientAssembly
+  isModestSubQuotientAssembly x y a a⊩x a⊩y =
+    SQ.elimProp2
+      {P = λ x y  motive x y}
+      isPropMotive
+       { (x , x~x) (y , y~y) a a~x a~y 
+        eq/ (x , x~x) (y , y~y) (PER.isTransitive per x a y (PER.isSymmetric per a x a~x) a~y) })
+      x y
+      a a⊩x a⊩y where
+        motive :  (x y : subQuotient)  Type 
+        motive x y =  (a : A) (a⊩x : a ⊩[ subQuotientAssembly ] x) (a⊩y : a ⊩[ subQuotientAssembly ] y)  x  y
+
+        isPropMotive :  x y  isProp (motive x y)
+        isPropMotive x y = isPropΠ3 λ _ _ _  squash/ x y
+
+module _ (R S : PER) (f : perMorphism R S) where
+  
+  subQuotientAssemblyMorphism : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)
+  subQuotientAssemblyMorphism =
+    SQ.rec
+      (isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S))
+      mainMap
+      mainMapCoherence
+      f where
+
+      mainMap : perTracker R S  AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)
+      AssemblyMorphism.map (mainMap (f , fIsTracker)) sqR =
+        SQ.rec
+          squash/
+          mainMapRepr
+          mainMapReprCoherence
+          sqR module MainMapDefn where
+            mainMapRepr : domain R  subQuotient S
+            mainMapRepr (r , r~r) = [ f  r , fIsTracker r r r~r ]
+
+            mainMapReprCoherence : (a b : domain R)  R .PER.relation (a .fst) (b .fst)  mainMapRepr a  mainMapRepr b
+            mainMapReprCoherence (a , a~a) (b , b~b) a~b = eq/ _ _ (fIsTracker a b a~b)
+ 
+      AssemblyMorphism.tracker (mainMap (f , fIsTracker)) =
+        do
+          return
+            (f ,
+             sqR s s⊩sqR 
+              SQ.elimProp
+                {P =
+                  λ sqR
+                    (s : A)
+                   s ⊩[ subQuotientAssembly R ] sqR
+                    subQuotientRealizability S (f  s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR) }
+                 sqR  isPropΠ2 λ s s⊩sqR  str (subQuotientRealizability S (f  s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR)))
+                 { (a , a~a) s s~a  fIsTracker s a s~a })
+                sqR s s⊩sqR))
+
+      mainMapCoherence : (a b : perTracker R S)  isEquivTracker R S a b  mainMap a  mainMap b
+      mainMapCoherence (a , a~a) (b , b~b) a~b =
+        AssemblyMorphism≡ _ _
+          (funExt
+            λ sq 
+              SQ.elimProp
+                {P =
+                  λ sq 
+                    SQ.rec
+                      squash/
+                      (MainMapDefn.mainMapRepr a a~a sq)
+                      (MainMapDefn.mainMapReprCoherence a a~a sq) sq
+                    
+                    SQ.rec
+                      squash/
+                      (MainMapDefn.mainMapRepr b b~b sq)
+                      (MainMapDefn.mainMapReprCoherence b b~b sq) sq}
+                 sq  squash/ _ _)
+                 { (r , r~r)  eq/ _ _ (a~b r r~r) })
+                sq)
+    
+module _ (R S : PER) (f : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)) where
+  subQuotientAssemblyMorphism→perMorphism : perMorphism R S
+  subQuotientAssemblyMorphism→perMorphism =
+    PT.rec→Set squash/ mainMap mainMap2Constant (f .tracker) module InverseDefinition where
+      isSQTracker : A  Type 
+      isSQTracker t =  (q : subQuotient R) (a : A)  a ⊩[ subQuotientAssembly R ] q   subQuotientRealizability S (t  a) (f .AssemblyMorphism.map q) 
+      -- 🤢🤮
+      mainMap : Σ[ t  A ] (isSQTracker t)  perMorphism R S
+      mainMap (t , t⊩f) =
+        [ t ,
+           r r' r~r' 
+            let
+              r~r : r ~[ R ] r
+              r~r = PER.isTransitive R r r' r r~r' (PER.isSymmetric R r r' r~r')
+
+              r'~r' : r' ~[ R ] r'
+              r'~r' = PER.isTransitive R r' r r' (PER.isSymmetric R r r' r~r') r~r'
+            in
+            SQ.elimProp
+              {P = λ q   (t : A)   subQuotientRealizability S (t  r) q    subQuotientRealizability S (t  r') q   (t  r) ~[ S ] (t  r')}
+               q  isPropΠ3 λ t _ _  isProp~ (t  r) S (t  r'))
+               { (s , s~s) t tr~s tr'~s  PER.isTransitive S (t  r) s (t  r') tr~s (PER.isSymmetric S (t  r') s tr'~s) })
+              (f .AssemblyMorphism.map [ (r , r~r) ])
+              t
+              (t⊩f [ (r , r~r) ] r r~r)
+              (subst  eq   subQuotientRealizability S (t  r') (f .AssemblyMorphism.map eq) ) (eq/ _ _ (PER.isSymmetric R r r' r~r')) (t⊩f [ (r' , r'~r') ] r' r'~r'))) ]
+
+      mainMap2Constant : 2-Constant mainMap
+      mainMap2Constant (t , t⊩f) (t' , t'⊩f) =
+        eq/ _ _
+          λ r r~r 
+            SQ.elimProp
+              {P = λ q   subQuotientRealizability S (t  r) q    subQuotientRealizability S (t'  r) q   (t  r) ~[ S ] (t'  r)}
+               q  isPropΠ2 λ _ _  isProp~ (t  r) S (t'  r))
+               { (s , s~s) tr~s t'r~s  PER.isTransitive S (t  r) s (t'  r) tr~s (PER.isSymmetric S (t'  r) s t'r~s) })
+              (f .AssemblyMorphism.map [ (r , r~r) ])
+              (t⊩f [ (r , r~r) ] r r~r)
+              (t'⊩f [ (r , r~r) ] r r~r)
+
+subQuotientModestSet : PER  MOD .Category.ob
+subQuotientModestSet R = subQuotient R , subQuotientAssembly R , isModestSubQuotientAssembly R
+
+subQuotientFunctor : Functor PERCat MOD
+Functor.F-ob subQuotientFunctor R = subQuotientModestSet R
+Functor.F-hom subQuotientFunctor {R} {S} f = subQuotientAssemblyMorphism R S f
+Functor.F-id subQuotientFunctor {R} =
+  AssemblyMorphism≡ _ _
+    (funExt
+      λ sqR 
+        SQ.elimProp
+          {P = λ sqR  subQuotientAssemblyMorphism R R (PERCat .Category.id {R}) .AssemblyMorphism.map sqR  identityMorphism (subQuotientAssembly R) .AssemblyMorphism.map sqR}
+           sqR  squash/ _ _)
+           { (a , a~a) 
+            eq/ _ _
+              (subst (_~[ R ] a) (sym (Ida≡a a)) a~a) })
+          sqR)
+Functor.F-seq subQuotientFunctor {R} {S} {T} f g =
+  AssemblyMorphism≡ _ _
+    (funExt
+      λ sq 
+        SQ.elimProp3
+          {P = λ sqR f g 
+            subQuotientAssemblyMorphism R T (seq' PERCat {R} {S} {T} f g) .AssemblyMorphism.map sqR 
+            seq' MOD
+              {x = subQuotientModestSet R}
+              {y = subQuotientModestSet S}
+              {z = subQuotientModestSet T}
+              (subQuotientAssemblyMorphism R S f) (subQuotientAssemblyMorphism S T g) .AssemblyMorphism.map sqR}
+           sq f g  squash/ _ _)
+           { (a , a~a) (b , bIsTracker) (c , cIsTracker) 
+            eq/ _ _ (subst (_~[ T ] (c  (b  a))) (sym (λ*ComputationRule (` c ̇ (` b ̇ # zero)) a)) (cIsTracker (b  a) (b  a) (bIsTracker a a a~a))) })
+          sq f g)
+
+hasPropFibersSubQuotientFunctor :  R S  hasPropFibers (subQuotientAssemblyMorphism R S)
+hasPropFibersSubQuotientFunctor R S f (x , sqX≡f) (y , sqY≡f) =
+  Σ≡Prop
+       perMap  isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S) _ _)
+      (SQ.elimProp2
+        {P = λ x y  subQuotientAssemblyMorphism R S x  f  subQuotientAssemblyMorphism R S y  f  x  y}
+         x y  isPropΠ2 λ _ _  squash/ _ _)
+         { (x , x⊩f) (y , y⊩f) sqX≡f sqY≡f 
+          eq/ _ _
+            λ r r~r 
+              SQ.elimProp
+                {P = λ f[r]   subQuotientRealizability S (x  r) f[r]     subQuotientRealizability S (y  r) f[r]   (x  r) ~[ S ] (y  r)}
+                 f[r]  isPropΠ2 λ _ _  isProp~ (x  r) S (y  r))
+                 { (s , s~s) xr~s yr~s  PER.isTransitive S (x  r) s (y  r) xr~s (PER.isSymmetric S (y  r) s yr~s) })
+                (f .AssemblyMorphism.map [ (r , r~r) ])
+                (subst  f[r]   subQuotientRealizability S (x  r) f[r] ) (cong  m  m .AssemblyMorphism.map [ (r , r~r) ]) sqX≡f) (x⊩f r r r~r))
+                (subst  f[r]   subQuotientRealizability S (y  r) f[r] ) (cong  m  m .AssemblyMorphism.map [ (r , r~r) ]) sqY≡f) (y⊩f r r r~r)) })
+        x y sqX≡f sqY≡f)
+
+fiberSubQuotientFunctor :  R S f  fiber (subQuotientAssemblyMorphism R S) f
+fiberSubQuotientFunctor R S f =
+  (subQuotientAssemblyMorphism→perMorphism R S f) ,
+  (AssemblyMorphism≡ _ _
+      (funExt
+         qR 
+          SQ.elimProp
+            {P = λ qR  subQuotientAssemblyMorphism R S (subQuotientAssemblyMorphism→perMorphism R S f) .AssemblyMorphism.map qR  f .AssemblyMorphism.map qR}
+             qR  squash/ _ _)
+             { (r , r~r) 
+              PT.elim
+                {P =
+                  λ fTracker 
+                    subQuotientAssemblyMorphism R S (PT.rec→Set squash/ (InverseDefinition.mainMap R S f) (InverseDefinition.mainMap2Constant R S f) fTracker) .AssemblyMorphism.map [ (r , r~r) ]
+                     f .AssemblyMorphism.map [ (r , r~r) ]}
+                 fTracker  squash/ _ _)
+                 { (t , tIsTracker) 
+                  SQ.elimProp
+                    {P =
+                      λ fqR   subQuotientRealizability S (t  r) fqR  
+                        subQuotientAssemblyMorphism R S (InverseDefinition.mainMap R S f (t , tIsTracker)) .AssemblyMorphism.map [ (r , r~r) ]  fqR}
+                     fqR  isProp→ (squash/ _ _))
+                     { (s , s~s) tr~s  eq/ _ _ tr~s })
+                    (f .AssemblyMorphism.map [ (r , r~r) ])
+                    (tIsTracker [ (r , r~r) ] r r~r) })
+                (f .tracker) })
+            qR)))
+
+isFullyFaithfulSubQuotientFunctor : Functor.isFullyFaithful subQuotientFunctor
+equiv-proof (isFullyFaithfulSubQuotientFunctor R S) f = inhProp→isContr (fiberSubQuotientFunctor R S f) (hasPropFibersSubQuotientFunctor R S f)
+
\ No newline at end of file diff --git a/docs/Realizability.PropResizing.html b/docs/Realizability.PropResizing.html index 5c3006b..3582d77 100644 --- a/docs/Realizability.PropResizing.html +++ b/docs/Realizability.PropResizing.html @@ -1,25 +1,79 @@ Realizability.PropResizing
open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Structure
-open import Cubical.Data.Sigma
+open import Cubical.Foundations.Equiv.Properties
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Structure
+open import Cubical.Foundations.Powerset
+open import Cubical.Data.Sigma
 
-module Realizability.PropResizing where
+module Realizability.PropResizing where
 
--- Formulation of propositional resizing inspired by the corresponding formulation
--- in TypeTopology
--- https://www.cs.bham.ac.uk/~mhe/TypeTopology/UF.Size.html
+-- Formulation of propositional resizing inspired by the corresponding formulation
+-- in TypeTopology
+-- https://www.cs.bham.ac.uk/~mhe/TypeTopology/UF.Size.html
 
-copyOf :  {}  Type   (ℓ' : Level)  Type _
-copyOf {} X ℓ' = Σ[ copy  Type ℓ' ] X  copy
+copyOf :  {}  Type   (ℓ' : Level)  Type _
+copyOf {} X ℓ' = Σ[ copy  Type ℓ' ] X  copy
 
-copy = fst
-copyEquiv = snd
+copy = fst
+copyEquiv = snd
 
--- We need the principle that TypeTopology calls Ω resizing
--- that the universe of props in a universe 𝓤 has a copy in 𝓤
--- This we call hPropResizing
-hPropResizing :    Type _
-hPropResizing  = copyOf (hProp ) 
+-- We need the principle that TypeTopology calls Ω resizing
+-- that the universe of props in a universe 𝓤 has a copy in 𝓤
+-- This we call hPropResizing
+hPropResizing :    Type _
+hPropResizing  = copyOf (hProp ) 
+
+-- We obtain a copy of the powerset using hPropResizing
+module ResizedPowerset {} (resizing : hPropResizing ) where
+
+  smallHProp = resizing .fst
+  hProp≃smallHProp = resizing .snd
+  smallHProp≃hProp = invEquiv hProp≃smallHProp
+
+  𝓟 : Type   Type 
+  𝓟 X = X  smallHProp
+
+  ℙ≃𝓟 :  X   X  𝓟 X
+  ℙ≃𝓟 X =
+     X
+      ≃⟨ idEquiv ( X) 
+    (X  hProp )
+      ≃⟨ equiv→ (idEquiv X) hProp≃smallHProp 
+    (X  smallHProp)
+      ≃⟨ idEquiv (𝓟 X) 
+    𝓟 X
+      
+
+  𝓟≃ℙ :  X  𝓟 X   X
+  𝓟≃ℙ X = invEquiv (ℙ≃𝓟 X)
+
+  ℙ→𝓟 :  X   X  𝓟 X
+  ℙ→𝓟 X = equivFun (ℙ≃𝓟 X)
+
+  𝓟→ℙ :  X  𝓟 X   X
+  𝓟→ℙ X = equivFun (invEquiv (ℙ≃𝓟 X))
+
+  compIsIdEquiv :  X  compEquiv (ℙ≃𝓟 X) (invEquiv (ℙ≃𝓟 X))  idEquiv ( X)
+  compIsIdEquiv X = invEquiv-is-rinv (ℙ≃𝓟 X)
+
+  compIsIdFunc :  {X} (p :  X)  𝓟→ℙ X (ℙ→𝓟 X p)  p
+  compIsIdFunc {X} p i = equivFun (compIsIdEquiv X i) p
+
+  _ϵ_ :  {X}  X  𝓟 X  Type 
+  _ϵ_ {X} x P = x  𝓟→ℙ X P
+
+  isPropϵ :  {X} (x : X) P  isProp (x ϵ P)
+  isPropϵ {X} x P = ∈-isProp (𝓟→ℙ X P) x
+
+  isSet𝓟 :  X  isSet (𝓟 X)
+  isSet𝓟 X = isOfHLevelRespectEquiv 2 (ℙ≃𝓟 X) isSetℙ
+
+  𝓟≡Equiv :  {X} (P Q : 𝓟 X)  (P  Q)  (𝓟→ℙ X P  𝓟→ℙ X Q)
+  𝓟≡Equiv {X} P Q = congEquiv {x = P} {y = Q} (𝓟≃ℙ X)
+
+  𝓟≡ :  {X} (P Q : 𝓟 X)  𝓟→ℙ X P  𝓟→ℙ X Q  P  Q
+  𝓟≡ {X} P Q equ = equivFun (invEquiv (𝓟≡Equiv P Q)) equ
+  
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.BinProducts.html b/docs/Realizability.Topos.BinProducts.html index aa2db2e..fac2dcc 100644 --- a/docs/Realizability.Topos.BinProducts.html +++ b/docs/Realizability.Topos.BinProducts.html @@ -1,5 +1,5 @@ -Realizability.Topos.BinProducts
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+Realizability.Topos.BinProducts
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.HLevels
@@ -17,7 +17,7 @@
 module Realizability.Topos.BinProducts
   { ℓ' ℓ''} {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
 
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
 open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
@@ -49,7 +49,7 @@
     Predicate.isSetX (PartialEquivalenceRelation.equality binProdObRT) =
       isSet× (isSet× isSetX isSetY) (isSet× isSetX isSetY)
     Predicate.∣ PartialEquivalenceRelation.equality binProdObRT  ((x , y) , x' , y') r =
-      (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   perY .equality  (y , y')
+      (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   perY .equality  (y , y')
     Predicate.isPropValued (PartialEquivalenceRelation.equality binProdObRT) ((x , y) , x' , y') r =
       isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _)
     isPartialEquivalenceRelation.isSetX (PartialEquivalenceRelation.isPerEquality binProdObRT) = isSet× isSetX isSetY
@@ -59,36 +59,36 @@
         (sY , sY⊩isSymmetricY)  perY .isSymmetric
         let
           prover : ApplStrTerm as 1
-          prover = ` pair ̇ (` sX ̇ (` pr₁ ̇ # zero)) ̇ (` sY ̇ (` pr₂ ̇ # zero))
+          prover = ` pair ̇ (` sX ̇ (` pr₁ ̇ # zero)) ̇ (` sY ̇ (` pr₂ ̇ # zero))
         return
-          (λ* prover ,
+          (λ* prover ,
            { (x , y) (x' , y') r (pr₁r⊩x~x' , pr₂r⊩y~y') 
             subst
                r'  r'   perX .equality  (x' , x))
-              (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
-              (sX⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x') ,
+              (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+              (sX⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x') ,
             subst
                r'  r'   perY .equality  (y' , y))
-              (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
-              (sY⊩isSymmetricY y y' (pr₂  r) pr₂r⊩y~y') }))
+              (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+              (sY⊩isSymmetricY y y' (pr₂  r) pr₂r⊩y~y') }))
     isPartialEquivalenceRelation.isTransitive (PartialEquivalenceRelation.isPerEquality binProdObRT) =
       do
         (tX , tX⊩isTransitiveX)  perX .isTransitive
         (tY , tY⊩isTransitiveY)  perY .isTransitive
         let
           prover : ApplStrTerm as 2
-          prover = ` pair ̇ (` tX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` tY ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
+          prover = ` pair ̇ (` tX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` tY ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
         return
-          (λ*2 prover ,
+          (λ*2 prover ,
            { (x , y) (x' , y') (x'' , y'') a b (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x'~x'' , pr₂b⊩y'~y'') 
             subst
                r'  r'   perX .equality  (x , x''))
-              (sym (cong  x  pr₁  x) (λ*2ComputationRule prover a b)  pr₁pxy≡x _ _))
-              (tX⊩isTransitiveX x x' x'' (pr₁  a) (pr₁  b) pr₁a⊩x~x' pr₁b⊩x'~x'') ,
+              (sym (cong  x  pr₁  x) (λ*2ComputationRule prover a b)  pr₁pxy≡x _ _))
+              (tX⊩isTransitiveX x x' x'' (pr₁  a) (pr₁  b) pr₁a⊩x~x' pr₁b⊩x'~x'') ,
             subst
                r'  r'   perY .equality  (y , y''))
-              (sym (cong  x  pr₂  x) (λ*2ComputationRule prover a b)  pr₂pxy≡y _ _))
-              (tY⊩isTransitiveY y y' y'' (pr₂  a) (pr₂  b) pr₂a⊩y~y' pr₂b⊩y'~y'') }))
+              (sym (cong  x  pr₂  x) (λ*2ComputationRule prover a b)  pr₂pxy≡y _ _))
+              (tY⊩isTransitiveY y y' y'' (pr₂  a) (pr₂  b) pr₂a⊩y~y' pr₂b⊩y'~y'') }))
 
   opaque
     unfolding binProdObRT
@@ -97,7 +97,7 @@
     FunctionalRelation.relation binProdPr₁FuncRel =
       record
         { isSetX = isSet× (isSet× isSetX isSetY) isSetX
-        ; ∣_∣ = λ { ((x , y) , x') r  (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   perY .equality  (y , y) }
+        ; ∣_∣ = λ { ((x , y) , x') r  (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   perY .equality  (y , y) }
         ; isPropValued =  { ((x , y) , x') r  isProp× (perX .equality .isPropValued _ _) (perY .equality .isPropValued _ _) }) }
     FunctionalRelation.isFuncRel binProdPr₁FuncRel =
       record
@@ -106,31 +106,31 @@
            (stD , stD⊩isStrictDomainEqX)  idFuncRel perX .isStrictDomain
            let
              prover : ApplStrTerm as 1
-             prover = ` pair ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (# zero))
+             prover = ` pair ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (# zero))
            return
-             (λ* prover ,
+             (λ* prover ,
               { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) 
                subst
                   r'  r'   perX .equality  (x , x))
-                 (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
-                 (stD⊩isStrictDomainEqX x x' (pr₁  r) pr₁r⊩x~x') ,
+                 (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                 (stD⊩isStrictDomainEqX x x' (pr₁  r) pr₁r⊩x~x') ,
                subst
                   r'  r'   perY .equality  (y , y))
-                 (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                 (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
                  pr₂r⊩y~y }))
        ; isStrictCodomain =
          do
            (stC , stC⊩isStrictCodomainEqX)  idFuncRel perX .isStrictCodomain
            let
              prover : ApplStrTerm as 1
-             prover = ` stC ̇ (` pr₁ ̇ # zero)
+             prover = ` stC ̇ (` pr₁ ̇ # zero)
            return
-             (λ* prover ,
+             (λ* prover ,
               λ { (x , y) x' r (pr₁r⊩x~x' , pr₂r⊩y~y) 
                 subst
                    r'  r'   perX .equality  (x' , x'))
-                  (sym (λ*ComputationRule prover r))
-                  (stC⊩isStrictCodomainEqX x x' (pr₁  r) pr₁r⊩x~x') })
+                  (sym (λ*ComputationRule prover r))
+                  (stC⊩isStrictCodomainEqX x x' (pr₁  r) pr₁r⊩x~x') })
        ; isRelational =
          do
            (stC , stC⊩isStrictCodomainEqY)  idFuncRel perY .isStrictCodomain
@@ -138,36 +138,36 @@
            (s , s⊩isSymmetricX)  perX .isSymmetric
            let
              prover : ApplStrTerm as 3
-             prover = ` pair ̇ (` t ̇ (` s ̇ (` pr₁ ̇ # two)) ̇ (` t ̇ (` pr₁ ̇ # one) ̇ # zero)) ̇ (` stC ̇ (` pr₂ ̇ # two))
+             prover = ` pair ̇ (` t ̇ (` s ̇ (` pr₁ ̇ # two)) ̇ (` t ̇ (` pr₁ ̇ # one) ̇ # zero)) ̇ (` stC ̇ (` pr₂ ̇ # two))
            return
-             (λ*3 prover ,
+             (λ*3 prover ,
               ((λ { (x , y) (x' , y') x'' x''' a b c (pr₁a⊩x~x' , pr₂a⊩y~y') (pr₁b⊩x~x'' , pr₂b⊩y~y) c⊩x''~x''' 
                 subst
                    r'  r'   perX .equality  (x' , x'''))
-                  (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
+                  (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
                   (t⊩isTransitiveX
                     x' x x'''
-                    (s  (pr₁  a)) (t  (pr₁  b)  c)
-                    (s⊩isSymmetricX x x' (pr₁  a) pr₁a⊩x~x')
-                    (t⊩isTransitiveX x x'' x''' (pr₁  b) c pr₁b⊩x~x'' c⊩x''~x''')) ,
+                    (s  (pr₁  a)) (t  (pr₁  b)  c)
+                    (s⊩isSymmetricX x x' (pr₁  a) pr₁a⊩x~x')
+                    (t⊩isTransitiveX x x'' x''' (pr₁  b) c pr₁b⊩x~x'' c⊩x''~x''')) ,
                 subst
                    r'  r'   perY .equality  (y' , y'))
-                  (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _))
-                  (stC⊩isStrictCodomainEqY y y' (pr₂  a) pr₂a⊩y~y') })))
+                  (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _))
+                  (stC⊩isStrictCodomainEqY y y' (pr₂  a) pr₂a⊩y~y') })))
        ; isSingleValued =
          do
            (t , t⊩isTransitive)  perX .isTransitive
            (s , s⊩isSymmetric)  perX .isSymmetric
            let
              prover : ApplStrTerm as 2
-             prover = ` t ̇ (` s ̇ (` pr₁ ̇ # one)) ̇ (` pr₁ ̇ # zero)
+             prover = ` t ̇ (` s ̇ (` pr₁ ̇ # one)) ̇ (` pr₁ ̇ # zero)
            return
-             (λ*2 prover ,
+             (λ*2 prover ,
                { (x , y) x' x'' r₁ r₂ (pr₁r₁⊩x~x' , pr₂r₁⊩y~y) (pr₁r₂⊩x~x'' , pr₂r₂⊩y~y) 
                 subst
                    r'  r'   perX .equality  (x' , x''))
-                  (sym (λ*2ComputationRule prover r₁ r₂))
-                  (t⊩isTransitive x' x x'' (s  (pr₁  r₁)) (pr₁  r₂) (s⊩isSymmetric x x' (pr₁  r₁) pr₁r₁⊩x~x') pr₁r₂⊩x~x'')}))
+                  (sym (λ*2ComputationRule prover r₁ r₂))
+                  (t⊩isTransitive x' x x'' (s  (pr₁  r₁)) (pr₁  r₂) (s⊩isSymmetric x x' (pr₁  r₁) pr₁r₁⊩x~x') pr₁r₂⊩x~x'')}))
        ; isTotal =
          do
            return
@@ -175,8 +175,8 @@
                { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) 
                 return
                   (x ,
-                  ((subst  r'  r'   perX .equality  (x , x)) (cong  x  pr₁  x) (sym (Ida≡a _))) pr₁r⊩x~x) ,
-                   (subst  r'  r'   perY .equality  (y , y)) (cong  x  pr₂  x) (sym (Ida≡a _))) pr₂r⊩y~y))) }))
+                  ((subst  r'  r'   perX .equality  (x , x)) (cong  x  pr₁  x) (sym (Ida≡a _))) pr₁r⊩x~x) ,
+                   (subst  r'  r'   perY .equality  (y , y)) (cong  x  pr₂  x) (sym (Ida≡a _))) pr₂r⊩y~y))) }))
        }
 
   opaque
@@ -192,7 +192,7 @@
     FunctionalRelation.relation binProdPr₂FuncRel =
       record
         { isSetX = isSet× (isSet× isSetX isSetY) isSetY
-        ; ∣_∣ = λ { ((x , y) , y') r  (pr₁  r)   perY .equality  (y , y') × (pr₂  r)   perX .equality  (x , x) }
+        ; ∣_∣ = λ { ((x , y) , y') r  (pr₁  r)   perY .equality  (y , y') × (pr₂  r)   perX .equality  (x , x) }
         ; isPropValued = λ { ((x , y) , y') r  isProp× (perY .equality .isPropValued _ _) (perX .equality .isPropValued _ _) } }
     FunctionalRelation.isFuncRel binProdPr₂FuncRel =
       record
@@ -201,79 +201,79 @@
            (stD , stD⊩isStrictDomainEqY)  idFuncRel perY .isStrictDomain
            let
              prover : ApplStrTerm as 1
-             prover = ` pair ̇ (` pr₂ ̇ (# zero)) ̇ (` stD ̇ (` pr₁ ̇ # zero))
+             prover = ` pair ̇ (` pr₂ ̇ (# zero)) ̇ (` stD ̇ (` pr₁ ̇ # zero))
            return
-             (λ* prover ,
+             (λ* prover ,
               { (x , y) y' r (pr₁r⊩y~y' , pr₂r⊩x~x) 
                 subst
                    r'  r'   perX .equality  (x , x))
-                  (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                  (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
                   pr₂r⊩x~x ,
                 subst
                    r'  r'   perY .equality  (y , y))
-                  (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
-                  (stD⊩isStrictDomainEqY y y' (pr₁  r) pr₁r⊩y~y') }))
+                  (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                  (stD⊩isStrictDomainEqY y y' (pr₁  r) pr₁r⊩y~y') }))
        ; isStrictCodomain =
          do
            (stC , stC⊩isStrictCodomainEqY)  idFuncRel perY .isStrictCodomain
            let
              prover : ApplStrTerm as 1
-             prover = ` stC ̇ (` pr₁ ̇ # zero)
+             prover = ` stC ̇ (` pr₁ ̇ # zero)
            return
-             (λ* prover ,
+             (λ* prover ,
               { (x , y) y' r (pr₁r⊩y~y' , pr₂r⊩x~x) 
                subst
                   r'  r'   perY .equality  (y' , y'))
-                 (sym (λ*ComputationRule prover r))
-                 (stC⊩isStrictCodomainEqY y y' (pr₁  r) pr₁r⊩y~y') }))
+                 (sym (λ*ComputationRule prover r))
+                 (stC⊩isStrictCodomainEqY y y' (pr₁  r) pr₁r⊩y~y') }))
        ; isRelational =
          do
            (stC , stC⊩isStrictCodomainEqX)  idFuncRel perX .isStrictCodomain
            (relY , relY⊩isRelationalEqY)  idFuncRel perY .isRelational
            let
              prover : ApplStrTerm as 3
-             prover = ` pair ̇ (` relY ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇ (` stC ̇ (` pr₁ ̇ # two))
+             prover = ` pair ̇ (` relY ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇ (` stC ̇ (` pr₁ ̇ # two))
            return
-             (λ*3 prover ,
+             (λ*3 prover ,
               { (x , y₁) (x' , y₂) y₃ y₄ a b c (pr₁a⊩x~x' , pr₂a⊩y₁~y₂) (pr₁b⊩y₁~y₃ , pr₂b⊩x~x) c⊩y₃~y₄ 
                subst
                   r'  r'   perY .equality  (y₂ , y₄))
-                 (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
-                 (relY⊩isRelationalEqY y₁ y₂ y₃ y₄ (pr₂  a) (pr₁  b) c pr₂a⊩y₁~y₂ pr₁b⊩y₁~y₃ c⊩y₃~y₄) ,
+                 (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
+                 (relY⊩isRelationalEqY y₁ y₂ y₃ y₄ (pr₂  a) (pr₁  b) c pr₂a⊩y₁~y₂ pr₁b⊩y₁~y₃ c⊩y₃~y₄) ,
                subst
                   r'  r'   perX .equality  (x' , x'))
-                 (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _))
-                 (stC⊩isStrictCodomainEqX x x' (pr₁  a) pr₁a⊩x~x') }))
+                 (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _))
+                 (stC⊩isStrictCodomainEqX x x' (pr₁  a) pr₁a⊩x~x') }))
        ; isSingleValued =
          do
            (svY , svY⊩isSingleValuedY)  idFuncRel perY .isSingleValued
            let
              prover : ApplStrTerm as 2
-             prover = ` svY ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
+             prover = ` svY ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
            return
-             (λ*2 prover ,
+             (λ*2 prover ,
               { (x , y) y' y'' r₁ r₂ (pr₁r₁⊩y~y' , pr₂r₁⊩x~x) (pr₁r₂⊩y~y'' , pr₂r₂⊩) 
                subst
                   r'  r'   perY .equality  (y' , y''))
-                 (sym (λ*2ComputationRule prover r₁ r₂))
-                 (svY⊩isSingleValuedY y y' y'' (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩y~y' pr₁r₂⊩y~y'') }))
+                 (sym (λ*2ComputationRule prover r₁ r₂))
+                 (svY⊩isSingleValuedY y y' y'' (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩y~y' pr₁r₂⊩y~y'') }))
        ; isTotal =
          do
            let
              prover : ApplStrTerm as 1
-             prover = ` pair ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero)
+             prover = ` pair ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero)
            return
-             (λ* prover ,
+             (λ* prover ,
               { (x , y) r (pr₁r⊩x~x , pr₂r⊩y~y) 
                return
                  (y ,
                    (subst
                       r'  r'   perY .equality  (y , y))
-                     (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                     (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
                      pr₂r⊩y~y ,
                     subst
                       r'  r'   perX .equality  (x , x))
-                     (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                     (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
                      pr₁r⊩x~x)) }))
        }
 
@@ -296,7 +296,7 @@
               { relation =
                 record
                   { isSetX = isSet× isSetZ (isSet× isSetX isSetY)
-                  ; ∣_∣ = λ { (z , x , y) r  (pr₁  r)   F .relation  (z , x) × (pr₂  r)   G .relation  (z , y) }
+                  ; ∣_∣ = λ { (z , x , y) r  (pr₁  r)   F .relation  (z , x) × (pr₂  r)   G .relation  (z , y) }
                 ; isPropValued = λ { (z , x , y) r  isProp× (F .relation .isPropValued _ _) (G .relation .isPropValued _ _) } }
               ; isFuncRel =
                 record
@@ -305,71 +305,71 @@
                      (stFD , stFD⊩isStrictDomain)  F .isStrictDomain
                      let
                        prover : ApplStrTerm as 1
-                       prover = ` stFD ̇ (` pr₁ ̇ # zero)
+                       prover = ` stFD ̇ (` pr₁ ̇ # zero)
                      return
-                       (λ* prover ,
+                       (λ* prover ,
                          { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
                           subst
                              r'  r'   perZ .equality  (z , z))
-                            (sym (λ*ComputationRule prover r))
-                            (stFD⊩isStrictDomain z x (pr₁  r) pr₁r⊩Fzx) }))
+                            (sym (λ*ComputationRule prover r))
+                            (stFD⊩isStrictDomain z x (pr₁  r) pr₁r⊩Fzx) }))
                  ; isStrictCodomain =
                    do
                      (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
                      (stGC , stGC⊩isStrictCodomainG)  G .isStrictCodomain
                      let
                        prover : ApplStrTerm as 1
-                       prover = ` pair ̇ (` stFC ̇ (` pr₁ ̇ # zero)) ̇ (` stGC ̇ (` pr₂ ̇ # zero))
+                       prover = ` pair ̇ (` stFC ̇ (` pr₁ ̇ # zero)) ̇ (` stGC ̇ (` pr₂ ̇ # zero))
                      return
-                       (λ* prover ,
+                       (λ* prover ,
                         { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
                          subst
                             r'  r'   perX .equality  (x , x))
-                           (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
-                           (stFC⊩isStrictCodomainF z x (pr₁  r) pr₁r⊩Fzx) ,
+                           (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                           (stFC⊩isStrictCodomainF z x (pr₁  r) pr₁r⊩Fzx) ,
                          subst
                             r'  r'   perY .equality  (y , y))
-                           (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
-                           (stGC⊩isStrictCodomainG z y (pr₂  r) pr₂r⊩Gzy) }))
+                           (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                           (stGC⊩isStrictCodomainG z y (pr₂  r) pr₂r⊩Gzy) }))
                  ; isRelational =
                    do
                      (relF , relF⊩isRelationalF)  F .isRelational
                      (relG , relG⊩isRelationalG)  G .isRelational
                      let
                        prover : ApplStrTerm as 3
-                       prover = ` pair ̇ (` relF ̇ # two ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` relG ̇ # two ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
+                       prover = ` pair ̇ (` relF ̇ # two ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` relG ̇ # two ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
                      return
-                       (λ*3 prover ,
+                       (λ*3 prover ,
                         { z z' (x , y) (x' , y') a b c a⊩z~z' (pr₁b⊩Fzx , pr₂b⊩Gzy) (pr₁c⊩x~x' , pr₂c⊩y~y') 
-                         (subst  r'  r'   F .relation  (z' , x')) (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _)) (relF⊩isRelationalF z z' x x' _ _ _ a⊩z~z' pr₁b⊩Fzx pr₁c⊩x~x')) ,
-                         subst  r'  r'   G .relation  (z' , y')) (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _)) (relG⊩isRelationalG z z' y y' _ _ _ a⊩z~z' pr₂b⊩Gzy pr₂c⊩y~y') }))
+                         (subst  r'  r'   F .relation  (z' , x')) (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _)) (relF⊩isRelationalF z z' x x' _ _ _ a⊩z~z' pr₁b⊩Fzx pr₁c⊩x~x')) ,
+                         subst  r'  r'   G .relation  (z' , y')) (sym (cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _)) (relG⊩isRelationalG z z' y y' _ _ _ a⊩z~z' pr₂b⊩Gzy pr₂c⊩y~y') }))
                  ; isSingleValued =
                    do
                      (svF , svF⊩isSingleValuedF)  F .isSingleValued
                      (svG , svG⊩isSingleValuedG)  G .isSingleValued
                      let
                        prover : ApplStrTerm as 2
-                       prover = ` pair ̇ (` svF ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` svG ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
+                       prover = ` pair ̇ (` svF ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` svG ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ # zero))
                      return
-                       (λ*2 prover ,
+                       (λ*2 prover ,
                         { z (x , y) (x' , y') r₁ r₂ (pr₁r₁⊩Fzx , pr₂r₁⊩Gzy) (pr₁r₂⊩Fzx' , pr₂r₂⊩Gzy') 
                          subst
                             r'  r'   perX .equality  (x , x'))
-                           (sym (cong  x  pr₁  x) (λ*2ComputationRule prover r₁ r₂)  pr₁pxy≡x _ _))
-                           (svF⊩isSingleValuedF z x x' (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩Fzx pr₁r₂⊩Fzx') ,
+                           (sym (cong  x  pr₁  x) (λ*2ComputationRule prover r₁ r₂)  pr₁pxy≡x _ _))
+                           (svF⊩isSingleValuedF z x x' (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩Fzx pr₁r₂⊩Fzx') ,
                          subst
                             r'  r'   perY .equality  (y , y'))
-                           (sym (cong  x  pr₂  x) (λ*2ComputationRule prover r₁ r₂)  pr₂pxy≡y _ _))
-                           (svG⊩isSingleValuedG z y y' (pr₂  r₁) (pr₂  r₂) pr₂r₁⊩Gzy pr₂r₂⊩Gzy') }))
+                           (sym (cong  x  pr₂  x) (λ*2ComputationRule prover r₁ r₂)  pr₂pxy≡y _ _))
+                           (svG⊩isSingleValuedG z y y' (pr₂  r₁) (pr₂  r₂) pr₂r₁⊩Gzy pr₂r₂⊩Gzy') }))
                  ; isTotal =
                    do
                      (tlF , tlF⊩isTotalF)  F .isTotal
                      (tlG , tlG⊩isTotalG)  G .isTotal
                      let
                        prover : ApplStrTerm as 1
-                       prover = ` pair ̇ (` tlF ̇ # zero) ̇ (` tlG ̇ # zero)
+                       prover = ` pair ̇ (` tlF ̇ # zero) ̇ (` tlG ̇ # zero)
                      return
-                       (λ* prover ,
+                       (λ* prover ,
                         { z r r⊩z~z 
                          do
                            (x , tlFr⊩Fzx)  tlF⊩isTotalF z r r⊩z~z
@@ -378,11 +378,11 @@
                              ((x , y) ,
                               (subst
                                  r'  r'   F .relation  (z , x))
-                                (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                                (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
                                 tlFr⊩Fzx ,
                                subst
                                  r'  r'   G .relation  (z , y))
-                                (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                                (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
                                 tlGr⊩Gzy)) }))
                  }}
 
@@ -402,17 +402,17 @@
                   (s , s⊩F≤F')  F≤F'
                   let
                     prover : ApplStrTerm as 1
-                    prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+                    prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
                   return
-                    (λ* prover ,
+                    (λ* prover ,
                       { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
                        subst
                           r'  r'   F' .relation  (z , x))
-                         (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
-                         (s⊩F≤F' z x (pr₁  r) pr₁r⊩Fzx) ,
+                         (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                         (s⊩F≤F' z x (pr₁  r) pr₁r⊩Fzx) ,
                        subst
                           r'  r'   G .relation  (z , y))
-                         (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                         (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
                          pr₂r⊩Gzy }))
             in
             eq/ _ _ (answer , F≤G→G≤F perZ binProdObRT (theFuncRel F G) (theFuncRel F' G) answer) })
@@ -423,18 +423,18 @@
                   (s , s⊩G≤G')  G≤G'
                   let
                     prover : ApplStrTerm as 1
-                    prover = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero))
+                    prover = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero))
                   return
-                    (λ* prover ,
+                    (λ* prover ,
                      { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
                       (subst
                          r'  r'   F .relation  (z , x))
-                        (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                        (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
                         pr₁r⊩Fzx) ,
                       (subst
                          r'  r'   G' .relation  (z , y))
-                        (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
-                        (s⊩G≤G' z y (pr₂  r) pr₂r⊩Gzy)) }))
+                        (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                        (s⊩G≤G' z y (pr₂  r) pr₂r⊩Gzy)) }))
             in eq/ _ _ (answer , (F≤G→G≤F perZ binProdObRT (theFuncRel F G) (theFuncRel F G') answer)) })
           f g
   opaque
@@ -473,9 +473,9 @@
                   (stD , stD⊩isStrictDomain)  theFuncRel' .isStrictDomain
                   let
                     prover : ApplStrTerm as 1
-                    prover = ` relF ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ (` p ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+                    prover = ` relF ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ (` p ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
                   return
-                    (λ* prover ,
+                    (λ* prover ,
                     λ z x r r⊩∃ 
                       transport
                         (propTruncIdempotent (F .relation .isPropValued _ _))
@@ -484,12 +484,12 @@
                           return
                             (subst
                                r'  r'   F .relation  (z , x))
-                              (sym (λ*ComputationRule prover r))
+                              (sym (λ*ComputationRule prover r))
                               (relF⊩isRelationalF
                                 z z x' x
-                                (stD  (pr₁  r)) (pr₁  (p  (pr₁  r))) (pr₁  (pr₂  r))
-                                (stD⊩isStrictDomain z (x' , y) (pr₁  r) pr₁r⊩theFuncRel'zx'y )
-                                (p⊩theFuncRel'≤theFuncRel z (x' , y) (pr₁  r) pr₁r⊩theFuncRel'zx'y .fst)
+                                (stD  (pr₁  r)) (pr₁  (p  (pr₁  r))) (pr₁  (pr₂  r))
+                                (stD⊩isStrictDomain z (x' , y) (pr₁  r) pr₁r⊩theFuncRel'zx'y )
+                                (p⊩theFuncRel'≤theFuncRel z (x' , y) (pr₁  r) pr₁r⊩theFuncRel'zx'y .fst)
                                  pr₁pr₂r⊩x~x'))))
             in
             eq/ _ _ (answer , F≤G→G≤F perZ perX (composeFuncRel _ _ _ theFuncRel' binProdPr₁FuncRel) F answer))
@@ -517,9 +517,9 @@
                   (st , st⊩isStrictDomainTheFuncRel')  theFuncRel' .isStrictDomain
                   let
                     prover : ApplStrTerm as 1
-                    prover = ` relG ̇ (` st ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (` p ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+                    prover = ` relG ̇ (` st ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (` p ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
                   return
-                    (λ* prover ,
+                    (λ* prover ,
                      z y r r⊩∃ 
                       transport
                         (propTruncIdempotent (G .relation .isPropValued _ _))
@@ -528,12 +528,12 @@
                           return
                             (subst
                                r'  r'   G .relation  (z , y))
-                              (sym (λ*ComputationRule prover r)) 
+                              (sym (λ*ComputationRule prover r)) 
                               (relG⊩isRelationalG
                                 z z y' y
-                                (st  (pr₁  r)) (pr₂  (p  (pr₁  r))) (pr₁  (pr₂  r))
-                                (st⊩isStrictDomainTheFuncRel' z (x , y') (pr₁  r) pr₁r⊩theFuncRel'zxy')
-                                (p⊩theFuncRel'≤theFuncRel z (x , y') (pr₁  r) pr₁r⊩theFuncRel'zxy' .snd)
+                                (st  (pr₁  r)) (pr₂  (p  (pr₁  r))) (pr₁  (pr₂  r))
+                                (st⊩isStrictDomainTheFuncRel' z (x , y') (pr₁  r) pr₁r⊩theFuncRel'zxy')
+                                (p⊩theFuncRel'≤theFuncRel z (x , y') (pr₁  r) pr₁r⊩theFuncRel'zxy' .snd)
                                 pr₁pr₂r⊩y'~y)))))
             in
             eq/ _ _ (answer , F≤G→G≤F perZ perY (composeFuncRel _ _ _ theFuncRel' binProdPr₂FuncRel) G answer))
@@ -562,17 +562,17 @@
                     let
                       realizer : ApplStrTerm as 1 -- cursed
                       realizer =
-                        ` rel!' ̇ (` stD!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # zero)))) ̇
-                          (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # zero))) ̇
-                          (` pair ̇
-                           (` tX ̇
-                            (` sX ̇
-                             (` pr₁ ̇
-                              (` sv!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # zero)))))) ̇
-                            (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pr₁ ̇ # zero))))) ̇
-                           (` pr₁ ̇ (` pr₂ ̇ (` q' ̇ (` pr₂ ̇ # zero)))))
+                        ` rel!' ̇ (` stD!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # zero)))) ̇
+                          (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # zero))) ̇
+                          (` pair ̇
+                           (` tX ̇
+                            (` sX ̇
+                             (` pr₁ ̇
+                              (` sv!' ̇ (` pr₁ ̇ (` q ̇ (` pr₁ ̇ # zero))) ̇ (` pr₁ ̇ (` q' ̇ (` pr₂ ̇ # zero)))))) ̇
+                            (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pr₁ ̇ # zero))))) ̇
+                           (` pr₁ ̇ (` pr₂ ̇ (` q' ̇ (` pr₂ ̇ # zero)))))
                     return
-                      (λ* realizer ,
+                      (λ* realizer ,
                        { z (x , y) r (pr₁r⊩Fzx , pr₂r⊩Gzy) 
                         transport
                           (propTruncIdempotent (!' .relation .isPropValued _ _))
@@ -585,7 +585,7 @@
                             return
                               (subst
                                  r'  r'   !' .relation  (z , x , y))
-                                (sym (λ*ComputationRule realizer r))
+                                (sym (λ*ComputationRule realizer r))
                                 (rel!'⊩isRelational!'
                                   z z
                                   (x'' , y'')
diff --git a/docs/Realizability.Topos.Equalizer.html b/docs/Realizability.Topos.Equalizer.html
index 916a1e2..f5ada32 100644
--- a/docs/Realizability.Topos.Equalizer.html
+++ b/docs/Realizability.Topos.Equalizer.html
@@ -36,7 +36,7 @@
 There is additional bureacracy because we have to deal with eliminators of set quotients. This makes things a little more complicated.
 
 -}
-open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.HLevels
@@ -56,7 +56,7 @@
 module Realizability.Topos.Equalizer
   { ℓ' ℓ''} {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
 
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
 open import Realizability.Topos.Object { = } {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial 
@@ -101,8 +101,8 @@
                 record
                   { isSetX = isSet× (perX .isSetX) (perX .isSetX)
                   ; ∣_∣ = λ { (x , x') r 
-                    ((pr₁  r)   perX .equality  (x , x')) ×
-                    (∃[ y  Y ] (pr₁  (pr₂  r))   F .relation  (x , y) × (pr₂  (pr₂  r))   G .relation  (x , y)) }
+                    ((pr₁  r)   perX .equality  (x , x')) ×
+                    (∃[ y  Y ] (pr₁  (pr₂  r))   F .relation  (x , y) × (pr₂  (pr₂  r))   G .relation  (x , y)) }
                   ; isPropValued = λ { (x , x') r  isProp× (perX .equality .isPropValued _ _) isPropPropTrunc } }
               ; isPerEquality =
                 record
@@ -116,40 +116,40 @@
                       let
                         prover : ApplStrTerm as 1
                         prover =
-                          ` pair ̇
-                            (` s ̇ (` pr₁ ̇ # zero)) ̇
-                            (` pair ̇
-                              (` relF ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)))) ̇
-                              (` relG ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)))))
+                          ` pair ̇
+                            (` s ̇ (` pr₁ ̇ # zero)) ̇
+                            (` pair ̇
+                              (` relF ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)))) ̇
+                              (` relG ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) ̇ (` stFC ̇ (` pr₁ ̇ (` pr₂ ̇ # zero)))))
                       return
-                        (λ* prover ,
+                        (λ* prover ,
                          { x x' r (pr₁r⊩x~x' , pr₂r⊩∃) 
                           subst
                              r'  r'   perX .equality  (x' , x))
-                            (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
-                            (s⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x') ,
+                            (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                            (s⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x') ,
                           do
                             (y , pr₁pr₂r⊩Fxy , pr₂pr₂r⊩Gxy)  pr₂r⊩∃
                             return
                               (y ,
                               subst
                                  r'  r'   F .relation  (x' , y))
-                                (sym (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                                (sym (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
                                 (relF⊩isRelationalF
                                   x x' y y
-                                  (pr₁  r) (pr₁  (pr₂  r)) (stFC  (pr₁  (pr₂  r)))
+                                  (pr₁  r) (pr₁  (pr₂  r)) (stFC  (pr₁  (pr₂  r)))
                                   pr₁r⊩x~x'
                                   pr₁pr₂r⊩Fxy
-                                  (stFC⊩isStrictCodomainF x y (pr₁  (pr₂  r)) pr₁pr₂r⊩Fxy)) ,
+                                  (stFC⊩isStrictCodomainF x y (pr₁  (pr₂  r)) pr₁pr₂r⊩Fxy)) ,
                               subst
                                  r'  r'   G .relation  (x' , y))
-                                (sym (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                                (sym (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
                                 (relG⊩isRelationalG
                                   x x' y y
-                                  (pr₁  r) (pr₂  (pr₂  r)) (stFC  (pr₁  (pr₂  r)))
+                                  (pr₁  r) (pr₂  (pr₂  r)) (stFC  (pr₁  (pr₂  r)))
                                   pr₁r⊩x~x'
                                   pr₂pr₂r⊩Gxy
-                                  (stFC⊩isStrictCodomainF x y (pr₁  (pr₂  r)) pr₁pr₂r⊩Fxy))) }))
+                                  (stFC⊩isStrictCodomainF x y (pr₁  (pr₂  r)) pr₁pr₂r⊩Fxy))) }))
                   ; isTransitive =
                     do
                       (t , t⊩isTransitiveX)  perX .isTransitive
@@ -157,14 +157,14 @@
                       (relG , relG⊩isRelationalG)  G .isRelational
                       let
                         prover : ApplStrTerm as 2
-                        prover = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ # one)) ̇ (` pr₂ ̇ (` pr₂ ̇ # one)))
+                        prover = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ # one)) ̇ (` pr₂ ̇ (` pr₂ ̇ # one)))
                       return
-                        (λ*2 prover ,
+                        (λ*2 prover ,
                         λ { x₁ x₂ x₃ a b (pr₁a⊩x₁~x₂ , pr₂a⊩∃) (pr₁b⊩x₂~x₃ , pr₂b⊩∃) 
                           subst
                              r'  r'   perX .equality  (x₁ , x₃))
-                            (sym (cong  x  pr₁  x) (λ*2ComputationRule prover a b)  pr₁pxy≡x _ _))
-                            (t⊩isTransitiveX x₁ x₂ x₃ (pr₁  a) (pr₁  b) pr₁a⊩x₁~x₂ pr₁b⊩x₂~x₃) ,
+                            (sym (cong  x  pr₁  x) (λ*2ComputationRule prover a b)  pr₁pxy≡x _ _))
+                            (t⊩isTransitiveX x₁ x₂ x₃ (pr₁  a) (pr₁  b) pr₁a⊩x₁~x₂ pr₁b⊩x₂~x₃) ,
                           do
                             (y , (pr₁pr₂a⊩Fx₁y , pr₂pr₂a⊩Gx₁y))  pr₂a⊩∃
                             (y' , (pr₁pr₂a⊩Fx₂y' , pr₂pr₂a⊩Gx₂y'))  pr₂b⊩∃
@@ -172,11 +172,11 @@
                               (y ,
                               subst
                                  r'  r'   F .relation  (x₁ , y))
-                                (sym (cong  x  pr₁  (pr₂  x)) (λ*2ComputationRule prover a b)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                                (sym (cong  x  pr₁  (pr₂  x)) (λ*2ComputationRule prover a b)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
                                 pr₁pr₂a⊩Fx₁y ,
                               subst
                                  r'  r'   G .relation  (x₁ , y))
-                                (sym (cong  x  pr₂  (pr₂  x)) (λ*2ComputationRule prover a b)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                                (sym (cong  x  pr₂  (pr₂  x)) (λ*2ComputationRule prover a b)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
                                 pr₂pr₂a⊩Gx₁y) }) } }
 
   opaque
@@ -194,11 +194,11 @@
                (stC , stC⊩isStrictCodomain)  idFuncRel perX .isStrictCodomain
                let
                  prover : ApplStrTerm as 1
-                 prover = ` stC ̇ (` pr₁ ̇ # zero)
+                 prover = ` stC ̇ (` pr₁ ̇ # zero)
                return
-                 (λ* prover ,
+                 (λ* prover ,
                   { x x' r (pr₁r⊩x~x' , pr₂r⊩∃) 
-                   subst  r'  r'   perX .equality  (x' , x')) (sym (λ*ComputationRule prover r)) (stC⊩isStrictCodomain x x' (pr₁  r) pr₁r⊩x~x') }))
+                   subst  r'  r'   perX .equality  (x' , x')) (sym (λ*ComputationRule prover r)) (stC⊩isStrictCodomain x x' (pr₁  r) pr₁r⊩x~x') }))
            ; isRelational =
              do
                (relEqX , relEqX⊩isRelationalEqX)  idFuncRel perX .isRelational
@@ -208,52 +208,52 @@
                let
                  prover : ApplStrTerm as 3
                  prover =
-                   ` pair ̇
-                     (` relEqX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇
-                     (` pair ̇
-                       (` relF ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` pr₁ ̇ (` pr₂ ̇ # one)))) ̇
-                       (` relG ̇ (` pr₁ ̇ # two) ̇ (` pr₂ ̇ (` pr₂ ̇ # two)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` pr₁ ̇ (` pr₂ ̇ # one)))))
+                   ` pair ̇
+                     (` relEqX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇
+                     (` pair ̇
+                       (` relF ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` pr₁ ̇ (` pr₂ ̇ # one)))) ̇
+                       (` relG ̇ (` pr₁ ̇ # two) ̇ (` pr₂ ̇ (` pr₂ ̇ # two)) ̇ (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ # two)) ̇ (` pr₁ ̇ (` pr₂ ̇ # one)))))
                return
-                 (λ*3 prover ,
+                 (λ*3 prover ,
                   x₁ x₂ x₃ x₄ a b c (pr₁a⊩x₁~x₂ , pr₂a⊩) (pr₁b⊩x₁~x₃ , pr₂b⊩) c⊩x₃~x₄ 
                    subst
                       r'  r'   perX .equality  (x₂ , x₄))
-                     (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
-                     (relEqX⊩isRelationalEqX x₁ x₂ x₃ x₄ (pr₁  a) (pr₁  b) c pr₁a⊩x₁~x₂ pr₁b⊩x₁~x₃ c⊩x₃~x₄) ,
+                     (sym (cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _))
+                     (relEqX⊩isRelationalEqX x₁ x₂ x₃ x₄ (pr₁  a) (pr₁  b) c pr₁a⊩x₁~x₂ pr₁b⊩x₁~x₃ c⊩x₃~x₄) ,
                    do
                      (y , pr₁pr₂a⊩Fx₁y , pr₂pr₂a⊩Gx₁y)  pr₂a⊩
                      (y' , pr₁pr₂b⊩Fx₁y' , pr₂pr₂b⊩Gx₁y')  pr₂b⊩
                      let
-                       y~y' = svF⊩isSingleValuedF x₁ y y' (pr₁  (pr₂  a)) (pr₁  (pr₂  b)) pr₁pr₂a⊩Fx₁y pr₁pr₂b⊩Fx₁y'
+                       y~y' = svF⊩isSingleValuedF x₁ y y' (pr₁  (pr₂  a)) (pr₁  (pr₂  b)) pr₁pr₂a⊩Fx₁y pr₁pr₂b⊩Fx₁y'
                      return
                        (y' ,
                        subst
                           r'  r'   F .relation  (x₂ , y'))
-                         (sym (cong  x  pr₁  (pr₂  x)) (λ*3ComputationRule prover a b c)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                         (sym (cong  x  pr₁  (pr₂  x)) (λ*3ComputationRule prover a b c)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
                          (relF⊩isRelationalF
                            x₁ x₂ y y'
-                           (pr₁  a) (pr₁  (pr₂  a)) (svF  (pr₁  (pr₂  a))  (pr₁  (pr₂  b)))
+                           (pr₁  a) (pr₁  (pr₂  a)) (svF  (pr₁  (pr₂  a))  (pr₁  (pr₂  b)))
                            pr₁a⊩x₁~x₂ pr₁pr₂a⊩Fx₁y y~y') ,
                        subst
                           r'  r'   G .relation  (x₂ , y'))
-                         (sym (cong  x  pr₂  (pr₂  x)) (λ*3ComputationRule prover a b c)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                         (sym (cong  x  pr₂  (pr₂  x)) (λ*3ComputationRule prover a b c)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
                          (relG⊩isRelationalG
                            x₁ x₂ y y'
-                           (pr₁  a) (pr₂  (pr₂  a)) (svF  (pr₁  (pr₂  a))  (pr₁  (pr₂  b)))
+                           (pr₁  a) (pr₂  (pr₂  a)) (svF  (pr₁  (pr₂  a))  (pr₁  (pr₂  b)))
                            pr₁a⊩x₁~x₂ pr₂pr₂a⊩Gx₁y y~y'))))
            ; isSingleValued =
              do
                (svEqX , svEqX⊩isSingleValuedEqX)  idFuncRel perX .isSingleValued
                let
                  prover : ApplStrTerm as 2
-                 prover = ` svEqX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
+                 prover = ` svEqX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
                return
-                 (λ*2 prover ,
+                 (λ*2 prover ,
                   x₁ x₂ x₃ r₁ r₂ (pr₁r₁⊩x₁~x₂ , pr₁r₁⊩) (pr₁r₂⊩x₁~x₃ , pr₂r₂⊩) 
                    subst
                       r'  r'   perX .equality  (x₂ , x₃))
-                     (sym (λ*2ComputationRule prover r₁ r₂))
-                     (svEqX⊩isSingleValuedEqX x₁ x₂ x₃ (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩x₁~x₂ pr₁r₂⊩x₁~x₃)))
+                     (sym (λ*2ComputationRule prover r₁ r₂))
+                     (svEqX⊩isSingleValuedEqX x₁ x₂ x₃ (pr₁  r₁) (pr₁  r₂) pr₁r₁⊩x₁~x₂ pr₁r₂⊩x₁~x₃)))
            ; isTotal = idFuncRel (equalizerPer F G) .isTotal
            } }
 
@@ -278,13 +278,13 @@
             let
               realizer : ApplStrTerm as 1
               realizer =
-                ` pair ̇
-                  (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))))) ̇
-                  (` relG ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
-                     (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
-                     (` relF ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ # zero) ̇ (` stCF ̇ (` pr₂ ̇ # zero)))))
+                ` pair ̇
+                  (` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))))) ̇
+                  (` relG ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
+                     (` svF ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
+                     (` relF ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ # zero) ̇ (` stCF ̇ (` pr₂ ̇ # zero)))))
             return
-              (λ* realizer ,
+              (λ* realizer ,
               -- unfold everything and bring it back in together
                x y r r⊩∃ 
                 do
@@ -297,7 +297,7 @@
                     (x' ,
                     (subst
                        r'  r'   perX .equality  (x , x'))
-                      (sym (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _))
+                      (sym (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _))
                       ⊩x~x' ,
                     do
                       return
@@ -305,22 +305,22 @@
                         subst
                            r'  r'   F .relation  (x , y'))
                           (sym
-                            (cong  x  pr₁  (pr₂  (pr₁  x))) (λ*ComputationRule realizer r) 
-                             cong  x  pr₁  (pr₂  x)) (pr₁pxy≡x _ _) 
-                             cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                            (cong  x  pr₁  (pr₂  (pr₁  x))) (λ*ComputationRule realizer r) 
+                             cong  x  pr₁  (pr₂  x)) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  x) (pr₂pxy≡y _ _) 
                              pr₁pxy≡x _ _))
                           ⊩Fxy' ,
                         subst
                            r'  r'   G .relation  (x , y'))
                           (sym
-                            (cong  x  pr₂  (pr₂  (pr₁  x))) (λ*ComputationRule realizer r) 
-                             cong  x  pr₂  (pr₂  x)) (pr₁pxy≡x _ _) 
-                             cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                            (cong  x  pr₂  (pr₂  (pr₁  x))) (λ*ComputationRule realizer r) 
+                             cong  x  pr₂  (pr₂  x)) (pr₁pxy≡x _ _) 
+                             cong  x  pr₂  x) (pr₂pxy≡y _ _) 
                              pr₂pxy≡y _ _))
                           ⊩Gxy')) ,
                     subst
                        r'  r'   G .relation  (x' , y))
-                      (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                      (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
                       (relG⊩isRelationalG x x' y' y _ _ _ ⊩x~x' ⊩Gxy' y'~y))))
       in
       eq/ _ _
@@ -367,26 +367,26 @@
                        -- possibly the ugliest realizer out there
                        prover : ApplStrTerm as 1
                        prover =
-                         ` pair ̇
-                           (` stCH ̇ # zero) ̇
-                           (` pair ̇
-                             (` tlF ̇ (` stCH ̇ # zero)) ̇
-                             (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCH ̇ # zero))))) ̇ # zero) ̇
-                             (` pr₂ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCH ̇ # zero))))) ̇
-                              (` stCF ̇ (` tlF ̇ (` stCH ̇ # zero)))))
+                         ` pair ̇
+                           (` stCH ̇ # zero) ̇
+                           (` pair ̇
+                             (` tlF ̇ (` stCH ̇ # zero)) ̇
+                             (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCH ̇ # zero))))) ̇ # zero) ̇
+                             (` pr₂ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCH ̇ # zero))))) ̇
+                              (` stCF ̇ (` tlF ̇ (` stCH ̇ # zero)))))
                      return
-                       (λ* prover ,
+                       (λ* prover ,
                         z x r r⊩Hzx 
                          let
                              x~x = stCH⊩isStrictCodomainH z x r r⊩Hzx
                          in
-                         subst  r  r   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) x~x ,
+                         subst  r  r   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) x~x ,
                          (do
-                           (y , ⊩Fxy)  tlF⊩isTotalF x (stCH  r) x~x
+                           (y , ⊩Fxy)  tlF⊩isTotalF x (stCH  r) x~x
                            let
                              hope =
                                p⊩H⋆F≤H⋆G
-                                 z y (pair  r  (tlF  (stCH  r)))
+                                 z y (pair  r  (tlF  (stCH  r)))
                                  (return
                                    (x ,
                                     subst  r'  r'   H .relation  (z , x)) (sym (pr₁pxy≡x _ _)) r⊩Hzx ,
@@ -396,8 +396,8 @@
                              subst
                                 r'  r'   F .relation  (x , y))
                                (sym
-                                 (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r) 
-                                  cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                                 (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r) 
+                                  cong  x  pr₁  x) (pr₂pxy≡y _ _) 
                                   pr₁pxy≡x _ _))
                                ⊩Fxy ,
                              -- god I wish there was a better way to do this :(
@@ -409,8 +409,8 @@
                                    (subst
                                       r'  r'   G .relation  (x , y))
                                      (sym
-                                       (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r) 
-                                        cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                                       (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r) 
+                                        cong  x  pr₂  x) (pr₂pxy≡y _ _) 
                                         pr₂pxy≡y _ _))
                                      (relG⊩isRelationalG x' x y y _ _ _ (svH⊩isSingleValuedH z x' x _ _ ⊩Hzx' r⊩Hzx) ⊩Gx'y (stCF⊩isStrictCodomainF x y _ ⊩Fxy))))))))
                  ; isRelational =
@@ -418,14 +418,14 @@
                      (relH , relH⊩isRelationalH)  H .isRelational
                      let
                        prover : ApplStrTerm as 3
-                       prover = ` relH ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
+                       prover = ` relH ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
                      return
-                       (λ*3 prover ,
+                       (λ*3 prover ,
                         z z' x x' a b c a⊩z~z' b⊩Hzx (pr₁c⊩x~x' , pr₂c⊩∃) 
                          subst
                             r'  r'   H .relation  (z' , x'))
-                           (sym (λ*3ComputationRule prover a b c))
-                           (relH⊩isRelationalH z z' x x' a b (pr₁  c) a⊩z~z' b⊩Hzx pr₁c⊩x~x')))
+                           (sym (λ*3ComputationRule prover a b c))
+                           (relH⊩isRelationalH z z' x x' a b (pr₁  c) a⊩z~z' b⊩Hzx pr₁c⊩x~x')))
                  ; isSingleValued =
                    do
                      (svH , svH⊩isSingleValuedH)  H .isSingleValued
@@ -438,15 +438,15 @@
                      let
                        prover : ApplStrTerm as 2
                        prover =
-                         ` pair ̇
-                           (` svH ̇ # one ̇ # zero) ̇
-                           (` pair ̇
-                             (` tlF ̇ (` stCH ̇ # one)) ̇
-                             (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # one ̇ (` tlF ̇ (` stCH ̇ # one))))) ̇ # one) ̇
-                               (` pr₂ ̇ (` p ̇ (` pair ̇ # one ̇ (` tlF ̇ (` stCH ̇ # one))))) ̇
-                               (` stCF ̇(` tlF ̇ (` stCH ̇ # one)))))
+                         ` pair ̇
+                           (` svH ̇ # one ̇ # zero) ̇
+                           (` pair ̇
+                             (` tlF ̇ (` stCH ̇ # one)) ̇
+                             (` relG ̇ (` svH ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # one ̇ (` tlF ̇ (` stCH ̇ # one))))) ̇ # one) ̇
+                               (` pr₂ ̇ (` p ̇ (` pair ̇ # one ̇ (` tlF ̇ (` stCH ̇ # one))))) ̇
+                               (` stCF ̇(` tlF ̇ (` stCH ̇ # one)))))
                      return
-                       (λ*2 prover ,
+                       (λ*2 prover ,
                         z x x' r₁ r₂ r₁⊩Hzx r₂⊩Hzx' 
                          let
                            x~x' = svH⊩isSingleValuedH z x x' r₁ r₂ r₁⊩Hzx r₂⊩Hzx'
@@ -454,7 +454,7 @@
                          in
                          subst
                             r'  r'   perX .equality  (x , x'))
-                           (sym (cong  x  pr₁  x) (λ*2ComputationRule prover r₁ r₂)  pr₁pxy≡x _ _))
+                           (sym (cong  x  pr₁  x) (λ*2ComputationRule prover r₁ r₂)  pr₁pxy≡x _ _))
                            x~x' ,
                          do
                            (y , ⊩Fxy)  tlF⊩isTotalF x _ x~x
@@ -462,7 +462,7 @@
                              y~y = stCF⊩isStrictCodomainF x y _ ⊩Fxy
                              hope =
                                p⊩H⋆F≤H⋆G z y
-                               (pair  r₁  (tlF  (stCH  r₁)))
+                               (pair  r₁  (tlF  (stCH  r₁)))
                                (do
                                  return
                                    (x ,
@@ -477,14 +477,14 @@
                              subst
                                 r'  r'   F .relation  (x , y))
                                (sym
-                                 (cong  x  pr₁  (pr₂  x)) (λ*2ComputationRule prover r₁ r₂) 
-                                  cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                                 (cong  x  pr₁  (pr₂  x)) (λ*2ComputationRule prover r₁ r₂) 
+                                  cong  x  pr₁  x) (pr₂pxy≡y _ _) 
                                   pr₁pxy≡x _ _)) ⊩Fxy ,
                              subst
                                 r'  r'   G .relation  (x , y))
                                (sym
-                                 (cong  x  pr₂  (pr₂  x)) (λ*2ComputationRule prover r₁ r₂) 
-                                  cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                                 (cong  x  pr₂  (pr₂  x)) (λ*2ComputationRule prover r₁ r₂) 
+                                  cong  x  pr₂  x) (pr₂pxy≡y _ _) 
                                   pr₂pxy≡y _ _))
                                (relG⊩isRelationalG x'' x y y _ _ _ x''~x ⊩Gx''y y~y))))
                  ; isTotal = H .isTotal
@@ -511,9 +511,9 @@
                 (stDH , stDH⊩isStrictDomainH)  H .isStrictDomain
                 let
                   prover : ApplStrTerm as 1
-                  prover = ` relH ̇ (` stDH ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+                  prover = ` relH ̇ (` stDH ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
                 return
-                  (λ* prover ,
+                  (λ* prover ,
                    z x r r⊩∃x' 
                     transport
                       (propTruncIdempotent (H .relation .isPropValued _ _))
@@ -522,8 +522,8 @@
                         return
                           (subst
                              r'  r'   H .relation  (z , x))
-                            (sym (λ*ComputationRule prover r))
-                            (relH⊩isRelationalH z z x' x _ _ _ (stDH⊩isStrictDomainH z x' (pr₁  r) pr₁r⊩Hzx') pr₁r⊩Hzx' pr₁pr₂r⊩x'~x)))))
+                            (sym (λ*ComputationRule prover r))
+                            (relH⊩isRelationalH z z x' x _ _ _ (stDH⊩isStrictDomainH z x' (pr₁  r) pr₁r⊩Hzx') pr₁r⊩Hzx' pr₁pr₂r⊩x'~x)))))
             !funcRel⋆inc≡H = eq/ _ _ (answer , F≤G→G≤F _ _ (composeFuncRel _ _ _ (!funcRel H H⋆F≡H⋆G) (equalizerFuncRel F G)) H answer)
           in !funcRel⋆inc≡H ,
           λ !' !'⋆inc≡H 
@@ -544,9 +544,9 @@
                       (stDH , stDH⊩isStrictDomainH)  H .isStrictDomain
                       let
                         prover : ApplStrTerm as 1
-                        prover = ` rel!' ̇ (` stDH ̇ # zero) ̇ (` pr₁ ̇ (` q ̇ # zero)) ̇ (` pr₂ ̇ (` q ̇ # zero))
+                        prover = ` rel!' ̇ (` stDH ̇ # zero) ̇ (` pr₁ ̇ (` q ̇ # zero)) ̇ (` pr₂ ̇ (` q ̇ # zero))
                       return
-                        (λ* prover ,
+                        (λ* prover ,
                          z x r r⊩Hzx 
                           transport
                             (propTruncIdempotent (!'funcRel .relation .isPropValued _ _))
@@ -555,7 +555,7 @@
                               return
                                 (subst
                                    r'  r'   !'funcRel .relation  (z , x))
-                                  (sym (λ*ComputationRule prover r))
+                                  (sym (λ*ComputationRule prover r))
                                   (rel!'⊩isRelational!'FuncRel
                                     z z x' x _ _ _
                                     (stDH⊩isStrictDomainH z x r r⊩Hzx)
diff --git a/docs/Realizability.Topos.Everything.html b/docs/Realizability.Topos.Everything.html
index 4ec7be1..0e760ab 100644
--- a/docs/Realizability.Topos.Everything.html
+++ b/docs/Realizability.Topos.Everything.html
@@ -1,13 +1,14 @@
 
-Realizability.Topos.Everything
module Realizability.Topos.Everything where
+Realizability.Topos.Everything
{-# OPTIONS --allow-unsolved-metas #-}
+module Realizability.Topos.Everything where
 
-open import Realizability.Topos.Object
-open import Realizability.Topos.FunctionalRelation
-open import Realizability.Topos.TerminalObject
-open import Realizability.Topos.BinProducts
-open import Realizability.Topos.Equalizer
-open import Realizability.Topos.MonicReprFuncRel
-open import Realizability.Topos.StrictRelation
-open import Realizability.Topos.ResizedPredicate
-open import Realizability.Topos.SubobjectClassifier
+open import Realizability.Topos.Object
+open import Realizability.Topos.FunctionalRelation
+open import Realizability.Topos.TerminalObject
+open import Realizability.Topos.BinProducts
+open import Realizability.Topos.Equalizer
+open import Realizability.Topos.MonicReprFuncRel
+open import Realizability.Topos.StrictRelation
+open import Realizability.Topos.ResizedPredicate
+open import Realizability.Topos.SubobjectClassifier
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.FunctionalRelation.html b/docs/Realizability.Topos.FunctionalRelation.html index dcc7b73..bdd8871 100644 --- a/docs/Realizability.Topos.FunctionalRelation.html +++ b/docs/Realizability.Topos.FunctionalRelation.html @@ -1,5 +1,5 @@ -Realizability.Topos.FunctionalRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+Realizability.Topos.FunctionalRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.Structure
@@ -21,7 +21,7 @@
   { ℓ' ℓ''}
   {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
   where
 
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
@@ -45,10 +45,10 @@
   equalityY = perY .equality
   
   realizesStrictDomain : A  Type _
-  realizesStrictDomain stD = (∀ x y r  r   relation  (x , y)  (stD  r)   equalityX  (x , x))
+  realizesStrictDomain stD = (∀ x y r  r   relation  (x , y)  (stD  r)   equalityX  (x , x))
 
   realizesStrictCodomain : A  Type _
-  realizesStrictCodomain stC = (∀ x y r  r   relation  (x , y)  (stC  r)   equalityY  (y , y))
+  realizesStrictCodomain stC = (∀ x y r  r   relation  (x , y)  (stC  r)   equalityY  (y , y))
 
   realizesRelational : A  Type _
   realizesRelational rel =
@@ -57,7 +57,7 @@
          b   relation  (x , y)
          c   equalityY  (y , y')
         ------------------------------------------
-         (rel  a  b  c)   relation  (x' , y'))
+         (rel  a  b  c)   relation  (x' , y'))
 
   realizesSingleValued : A  Type _
   realizesSingleValued sv =
@@ -65,11 +65,11 @@
          r₁   relation  (x , y)
          r₂   relation  (x , y')
         -----------------------------------
-         (sv  r₁  r₂)   equalityY  (y , y'))
+         (sv  r₁  r₂)   equalityY  (y , y'))
 
   realizesTotal : A  Type _
   realizesTotal tl =
-        (∀ x r  r   equalityX  (x , x)  ∃[ y  Y ] (tl  r)   relation  (x , y))
+        (∀ x r  r   equalityX  (x , x)  ∃[ y  Y ] (tl  r)   relation  (x , y))
     
   record isFunctionalRelation : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
     constructor makeIsFunctionalRelation
@@ -90,7 +90,7 @@
 open FunctionalRelation
 
 pointwiseEntailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  (F G : FunctionalRelation perX perY)  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
-pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe  A ] (∀ x y r  r   F .relation  (x , y)  (pe  r)   G .relation  (x , y))
+pointwiseEntailment {X} {Y} perX perY F G = ∃[ pe  A ] (∀ x y r  r   F .relation  (x , y)  (pe  r)   G .relation  (x , y))
 
 -- Directly taken from "Realizability with Scott's Graph Model" by Tom de Jong
 -- Lemma 4.3.5
@@ -111,24 +111,24 @@
       (stGD , stGD⊩isStrictDomainG)  G .isStrictDomain
       let
         prover : ApplStrTerm as 1
-        prover = ` rlF ̇ (` stGD ̇ # zero) ̇ (` tlF ̇ (` stGD ̇ # zero)) ̇ (` svG ̇ (` r ̇ (` tlF ̇ (` stGD ̇ # zero))) ̇ # zero)
+        prover = ` rlF ̇ (` stGD ̇ # zero) ̇ (` tlF ̇ (` stGD ̇ # zero)) ̇ (` svG ̇ (` r ̇ (` tlF ̇ (` stGD ̇ # zero))) ̇ # zero)
       return
-        (λ* prover ,
+        (λ* prover ,
          x y s s⊩Gxy 
           subst
              r'  r'   F .relation  (x , y))
-            (sym (λ*ComputationRule prover s))
+            (sym (λ*ComputationRule prover s))
             (transport
               (propTruncIdempotent (F .relation .isPropValued _ _))
               (do
-                (y' , tlF⨾stGDs⊩Fxy')  tlF⊩isTotalF x (stGD  s) (stGD⊩isStrictDomainG x y s s⊩Gxy)
+                (y' , tlF⨾stGDs⊩Fxy')  tlF⊩isTotalF x (stGD  s) (stGD⊩isStrictDomainG x y s s⊩Gxy)
                 return
                   (rlF⊩isRelationalF
                     x x y' y
-                    (stGD  s) (tlF  (stGD  s)) (svG  (r  (tlF  (stGD  s)))  s)
+                    (stGD  s) (tlF  (stGD  s)) (svG  (r  (tlF  (stGD  s)))  s)
                     (stGD⊩isStrictDomainG x y s s⊩Gxy)
                     tlF⨾stGDs⊩Fxy'
-                    (svG⊩isSingleValuedG x y' y (r  (tlF  (stGD  s))) s (r⊩F≤G x y' (tlF  (stGD  s)) tlF⨾stGDs⊩Fxy') s⊩Gxy))))))
+                    (svG⊩isSingleValuedG x y' y (r  (tlF  (stGD  s))) s (r⊩F≤G x y' (tlF  (stGD  s)) tlF⨾stGDs⊩Fxy') s⊩Gxy))))))
 
 bientailment :  {X Y : Type ℓ'}  (perX : PartialEquivalenceRelation X)  (perY : PartialEquivalenceRelation Y)  FunctionalRelation perX perY  FunctionalRelation perX perY  Type _
 bientailment {X} {Y} perX perY F G = pointwiseEntailment perX perY F G × pointwiseEntailment perX perY G F
@@ -153,10 +153,10 @@
         (p , p⊩G≤H)  G≤H
         let
           prover : ApplStrTerm as 1
-          prover = ` p ̇ (` s ̇ # zero)
+          prover = ` p ̇ (` s ̇ # zero)
         return
-          (λ* prover ,
-           x y r r⊩Fxy  subst  r'  r'   H .relation  (x , y)) (sym (λ*ComputationRule prover r)) (p⊩G≤H x y (s  r) (s⊩F≤G x y r r⊩Fxy))))
+          (λ* prover ,
+           x y r r⊩Fxy  subst  r'  r'   H .relation  (x , y)) (sym (λ*ComputationRule prover r)) (p⊩G≤H x y (s  r) (s⊩F≤G x y r r⊩Fxy))))
   in
   answer , F≤G→G≤F perX perY F H answer
 
@@ -169,56 +169,56 @@
       (t , t⊩isTransitive)  perX .isTransitive
       let
         prover : ApplStrTerm as 1
-        prover = ` t ̇ # zero ̇ (` s ̇ # zero)
+        prover = ` t ̇ # zero ̇ (` s ̇ # zero)
       return
-        (λ* prover ,
+        (λ* prover ,
          λ x x' r r⊩x~x' 
            subst
               r'  r'   perX .equality  (x , x))
-             (sym (λ*ComputationRule prover r))
-             (t⊩isTransitive x x' x r (s  r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')))
+             (sym (λ*ComputationRule prover r))
+             (t⊩isTransitive x x' x r (s  r) r⊩x~x' (s⊩isSymmetric x x' r r⊩x~x')))
   isFunctionalRelation.isStrictCodomain (isFuncRel (idFuncRel {X} perX)) =
     do
       (s , s⊩isSymmetric)  perX .isSymmetric
       (t , t⊩isTransitive)  perX .isTransitive
       let
         prover : ApplStrTerm as 1
-        prover = ` t ̇ (` s ̇ # zero) ̇ # zero
+        prover = ` t ̇ (` s ̇ # zero) ̇ # zero
       return
-        (λ* prover ,
+        (λ* prover ,
          x x' r r⊩x~x' 
           subst
              r'  r'   perX .equality  (x' , x'))
-            (sym (λ*ComputationRule prover r))
-            (t⊩isTransitive x' x x' (s  r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x')))
+            (sym (λ*ComputationRule prover r))
+            (t⊩isTransitive x' x x' (s  r) r (s⊩isSymmetric x x' r r⊩x~x') r⊩x~x')))
   isFunctionalRelation.isRelational (isFuncRel (idFuncRel {X} perX)) =
     do
       (s , s⊩isSymmetric)  perX .isSymmetric
       (t , t⊩isTransitive)  perX .isTransitive
       let
         prover : ApplStrTerm as 3
-        prover = ` t ̇ (` t ̇ (` s ̇ # two) ̇ # one) ̇ # zero
+        prover = ` t ̇ (` t ̇ (` s ̇ # two) ̇ # one) ̇ # zero
       return
-        (λ*3 prover ,
+        (λ*3 prover ,
          x₁ x₂ x₃ x₄ a b c a⊩x₁~x₂ b⊩x₁~x₃ c⊩x₃~x₄ 
           subst
              r'  r'   perX .equality  (x₂ , x₄))
-            (sym (λ*3ComputationRule prover a b c))
-            (t⊩isTransitive x₂ x₃ x₄ (t  (s  a)  b) c (t⊩isTransitive x₂ x₁ x₃ (s  a) b (s⊩isSymmetric x₁ x₂ a a⊩x₁~x₂) b⊩x₁~x₃) c⊩x₃~x₄)))
+            (sym (λ*3ComputationRule prover a b c))
+            (t⊩isTransitive x₂ x₃ x₄ (t  (s  a)  b) c (t⊩isTransitive x₂ x₁ x₃ (s  a) b (s⊩isSymmetric x₁ x₂ a a⊩x₁~x₂) b⊩x₁~x₃) c⊩x₃~x₄)))
   isFunctionalRelation.isSingleValued (isFuncRel (idFuncRel {X} perX)) =
     do
       (s , s⊩isSymmetric)  perX .isSymmetric
       (t , t⊩isTransitive)  perX .isTransitive
       let
         prover : ApplStrTerm as 2
-        prover = ` t ̇ (` s ̇ # one) ̇ # zero
+        prover = ` t ̇ (` s ̇ # one) ̇ # zero
       return
-        (λ*2 prover ,
+        (λ*2 prover ,
          x₁ x₂ x₃ r₁ r₂ r₁⊩x₁~x₂ r₂⊩x₁~x₃ 
           subst
              r'  r'   perX .equality  (x₂ , x₃))
-            (sym (λ*2ComputationRule prover r₁ r₂))
-            (t⊩isTransitive x₂ x₁ x₃ (s  r₁) r₂ (s⊩isSymmetric x₁ x₂ r₁ r₁⊩x₁~x₂) r₂⊩x₁~x₃)))
+            (sym (λ*2ComputationRule prover r₁ r₂))
+            (t⊩isTransitive x₂ x₁ x₃ (s  r₁) r₂ (s⊩isSymmetric x₁ x₂ r₁ r₁⊩x₁~x₂) r₂⊩x₁~x₃)))
   isFunctionalRelation.isTotal (isFuncRel (idFuncRel {X} perX)) =
     do
       (s , s⊩isSymmetric)  perX .isSymmetric
@@ -242,42 +242,42 @@
      FunctionalRelation perX perZ
   isSetPredicateBase (relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) = isSet× (perX .isSetX) (perZ .isSetX)
    relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)  (x , z) r =
-    ∃[ y  Y ] (pr₁  r)   F .relation  (x , y) × (pr₂  r)   G .relation  (y , z)
+    ∃[ y  Y ] (pr₁  r)   F .relation  (x , y) × (pr₂  r)   G .relation  (y , z)
   isPropValued (relation (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) (x , z) r = isPropPropTrunc
   isFunctionalRelation.isStrictDomain (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
     do
       (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
       let
         prover : ApplStrTerm as 1
-        prover = ` stFD ̇ (` pr₁ ̇ # zero)
+        prover = ` stFD ̇ (` pr₁ ̇ # zero)
       return
-        (λ* prover ,
+        (λ* prover ,
          x z r r⊩∃y 
           subst
              r'  r'   perX .equality  (x , x))
-            (sym (λ*ComputationRule prover r))
+            (sym (λ*ComputationRule prover r))
             (transport
               (propTruncIdempotent (perX .equality .isPropValued _ _))
               (do
                 (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
-                return (stFD⊩isStrictDomainF x y (pr₁  r) pr₁r⊩Fxy)))))
+                return (stFD⊩isStrictDomainF x y (pr₁  r) pr₁r⊩Fxy)))))
   isFunctionalRelation.isStrictCodomain (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
     do
       (stGC , stGC⊩isStrictCodomainG)  G .isStrictCodomain
       let
         prover : ApplStrTerm as 1
-        prover = ` stGC ̇ (` pr₂ ̇ # zero)
+        prover = ` stGC ̇ (` pr₂ ̇ # zero)
       return
-        (λ* prover ,
+        (λ* prover ,
          λ x z r r⊩∃y 
            subst
               r'  r'   perZ .equality  (z , z))
-             (sym (λ*ComputationRule prover r))
+             (sym (λ*ComputationRule prover r))
              (transport
                (propTruncIdempotent (perZ .equality .isPropValued _ _))
                (do
                  (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
-                 return (stGC⊩isStrictCodomainG y z (pr₂  r) pr₂r⊩Gyz))))
+                 return (stGC⊩isStrictCodomainG y z (pr₂  r) pr₂r⊩Gyz))))
   isFunctionalRelation.isRelational (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
     do
       (rlF , rlF⊩isRelationalF)  F .isRelational
@@ -285,28 +285,28 @@
       (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
       let
         prover : ApplStrTerm as 3
-        prover = ` pair ̇ (` rlF ̇ # two ̇ (` pr₁ ̇ # one) ̇ (` stFC ̇ (` pr₁ ̇ # one))) ̇ (` rlG ̇ (` stFC ̇ (` pr₁ ̇ # one)) ̇ (` pr₂ ̇ # one) ̇ # zero)
+        prover = ` pair ̇ (` rlF ̇ # two ̇ (` pr₁ ̇ # one) ̇ (` stFC ̇ (` pr₁ ̇ # one))) ̇ (` rlG ̇ (` stFC ̇ (` pr₁ ̇ # one)) ̇ (` pr₂ ̇ # one) ̇ # zero)
       return
-        (λ*3 prover ,
+        (λ*3 prover ,
          x x' z z' a b c a⊩x~x' b⊩∃y c⊩z~z' 
           do
             (y , pr₁b⊩Fxy , pr₂b⊩Gyz)  b⊩∃y
             let
-              pr₁proofEq : pr₁  (λ*3 prover  a  b  c)  rlF  a  (pr₁  b)  (stFC  (pr₁  b))
-              pr₁proofEq = cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _
+              pr₁proofEq : pr₁  (λ*3 prover  a  b  c)  rlF  a  (pr₁  b)  (stFC  (pr₁  b))
+              pr₁proofEq = cong  x  pr₁  x) (λ*3ComputationRule prover a b c)  pr₁pxy≡x _ _
 
-              pr₂proofEq : pr₂  (λ*3 prover  a  b  c)  rlG  (stFC  (pr₁  b))  (pr₂  b)  c
-              pr₂proofEq = cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _
+              pr₂proofEq : pr₂  (λ*3 prover  a  b  c)  rlG  (stFC  (pr₁  b))  (pr₂  b)  c
+              pr₂proofEq = cong  x  pr₂  x) (λ*3ComputationRule prover a b c)  pr₂pxy≡y _ _
             return
               (y ,
                subst
                   r'  r'   F .relation  (x' , y))
                  (sym pr₁proofEq)
-                 (rlF⊩isRelationalF x x' y y a (pr₁  b) (stFC  (pr₁  b)) a⊩x~x' pr₁b⊩Fxy (stFC⊩isStrictCodomainF x y (pr₁  b) pr₁b⊩Fxy)) ,
+                 (rlF⊩isRelationalF x x' y y a (pr₁  b) (stFC  (pr₁  b)) a⊩x~x' pr₁b⊩Fxy (stFC⊩isStrictCodomainF x y (pr₁  b) pr₁b⊩Fxy)) ,
                subst
                   r'  r'   G .relation  (y , z'))
                  (sym pr₂proofEq)
-                 (rlG⊩isRelationalG y y z z' (stFC  (pr₁  b)) (pr₂  b) c (stFC⊩isStrictCodomainF x y (pr₁  b) pr₁b⊩Fxy) pr₂b⊩Gyz c⊩z~z'))))
+                 (rlG⊩isRelationalG y y z z' (stFC  (pr₁  b)) (pr₂  b) c (stFC⊩isStrictCodomainF x y (pr₁  b) pr₁b⊩Fxy) pr₂b⊩Gyz c⊩z~z'))))
   isFunctionalRelation.isSingleValued (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
     do
       (svF , svF⊩isSingleValuedF)  F .isSingleValued
@@ -315,9 +315,9 @@
       (stGC , stGC⊩isStrictCodomainG)  G .isStrictCodomain
       let
         prover : ApplStrTerm as 2
-        prover = ` svG ̇ (` pr₂ ̇ # one) ̇ (` relG ̇ (` svF ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ # one)) ̇ (` pr₂ ̇ # zero) ̇ (` stGC ̇ (` pr₂ ̇ # zero)))
+        prover = ` svG ̇ (` pr₂ ̇ # one) ̇ (` relG ̇ (` svF ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ # one)) ̇ (` pr₂ ̇ # zero) ̇ (` stGC ̇ (` pr₂ ̇ # zero)))
       return
-        (λ*2 prover ,
+        (λ*2 prover ,
          x z z' r₁ r₂ r₁⊩∃y r₂⊩∃y 
           transport
             (propTruncIdempotent (perZ .equality .isPropValued _ _))
@@ -327,20 +327,20 @@
               return
                 (subst
                    r'  r'   perZ .equality  (z , z'))
-                  (sym (λ*2ComputationRule prover r₁ r₂))
+                  (sym (λ*2ComputationRule prover r₁ r₂))
                   (svG⊩isSingleValuedG
                     y z z'
-                    (pr₂  r₁)
-                    (relG  (svF  (pr₁  r₂)  (pr₁  r₁))  (pr₂  r₂)  (stGC  (pr₂  r₂)))
+                    (pr₂  r₁)
+                    (relG  (svF  (pr₁  r₂)  (pr₁  r₁))  (pr₂  r₂)  (stGC  (pr₂  r₂)))
                     pr₂r₁⊩Gyz
                     (relG⊩isRelationalG
                       y' y z' z'
-                      (svF  (pr₁  r₂)  (pr₁  r₁))
-                      (pr₂  r₂)
-                      (stGC  (pr₂  r₂))
-                      (svF⊩isSingleValuedF x y' y (pr₁  r₂) (pr₁  r₁) pr₁r₂⊩Fxy' pr₁r₁⊩Fxy)
+                      (svF  (pr₁  r₂)  (pr₁  r₁))
+                      (pr₂  r₂)
+                      (stGC  (pr₂  r₂))
+                      (svF⊩isSingleValuedF x y' y (pr₁  r₂) (pr₁  r₁) pr₁r₂⊩Fxy' pr₁r₁⊩Fxy)
                       pr₂r₂⊩Gy'z'
-                      (stGC⊩isStrictCodomainG y' z' (pr₂  r₂) pr₂r₂⊩Gy'z')))))))
+                      (stGC⊩isStrictCodomainG y' z' (pr₂  r₂) pr₂r₂⊩Gy'z')))))))
   isFunctionalRelation.isTotal (isFuncRel (composeFuncRel {X} {Y} {Z} perX perY perZ F G)) =
     do
       (tlF , tlF⊩isTotalF)  F .isTotal
@@ -348,19 +348,19 @@
       (stFC , stFC⊩isStrictCodomainF)  F .isStrictCodomain
       let
         prover : ApplStrTerm as 1
-        prover = ` pair ̇ (` tlF ̇ # zero) ̇ (` tlG ̇ (` stFC ̇ (` tlF ̇ # zero)))
+        prover = ` pair ̇ (` tlF ̇ # zero) ̇ (` tlG ̇ (` stFC ̇ (` tlF ̇ # zero)))
       return
-        (λ* prover ,
+        (λ* prover ,
          x r r⊩x~x 
           do
             (y , ⊩Fxy)  tlF⊩isTotalF x r r⊩x~x
-            (z , ⊩Gyz)  tlG⊩isTotalG y (stFC  (tlF  r)) (stFC⊩isStrictCodomainF x y (tlF  r) ⊩Fxy)
+            (z , ⊩Gyz)  tlG⊩isTotalG y (stFC  (tlF  r)) (stFC⊩isStrictCodomainF x y (tlF  r) ⊩Fxy)
             return
               (z ,
               return
                 (y ,
-                ((subst  r'  r'   F .relation  (x , y)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) ⊩Fxy) ,
-                 (subst  r'  r'   G .relation  (y , z)) (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _)) ⊩Gyz))))))
+                ((subst  r'  r'   F .relation  (x , y)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) ⊩Fxy) ,
+                 (subst  r'  r'   G .relation  (y , z)) (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _)) ⊩Gyz))))))
 
 opaque
   unfolding composeFuncRel
@@ -383,9 +383,9 @@
               (s , s⊩F≤F')  F≤F'
               let
                 prover : ApplStrTerm as 1
-                prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+                prover = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
               return
-                (λ* prover ,
+                (λ* prover ,
                  x z r r⊩∃y 
                   do
                     (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
@@ -393,11 +393,11 @@
                       (y ,
                        subst
                           r'  r'   F' .relation  (x , y))
-                         (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
-                         (s⊩F≤F' x y (pr₁  r) pr₁r⊩Fxy) ,
+                         (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                         (s⊩F≤F' x y (pr₁  r) pr₁r⊩Fxy) ,
                        subst
                           r'  r'   G .relation  (y , z))
-                         (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
+                         (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _))
                          pr₂r⊩Gyz))))
           in
         (answer , F≤G→G≤F perX perZ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F' G) answer) })
@@ -407,17 +407,17 @@
             (s , s⊩G≤G')  G≤G'
             let
               prover : ApplStrTerm as 1
-              prover = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero))
+              prover = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero))
             return
-              (λ* prover ,
+              (λ* prover ,
                x z r r⊩∃y 
                  do
                    (y , pr₁r⊩Fxy , pr₂r⊩Gyz)  r⊩∃y
 
                    return
                      (y ,
-                      subst  r'  r'   F .relation  (x , y)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) pr₁r⊩Fxy ,
-                      subst  r'  r'   G' .relation  (y , z)) (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _)) (s⊩G≤G' y z (pr₂  r) pr₂r⊩Gyz)))))
+                      subst  r'  r'   F .relation  (x , y)) (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _)) pr₁r⊩Fxy ,
+                      subst  r'  r'   G' .relation  (y , z)) (sym (cong  x  pr₂  x) (λ*ComputationRule prover r)  pr₂pxy≡y _ _)) (s⊩G≤G' y z (pr₂  r) pr₂r⊩Gyz)))))
           in
         (answer , F≤G→G≤F perX perZ (composeFuncRel perX perY perZ F G) (composeFuncRel perX perY perZ F G') answer) })
       f g
@@ -444,9 +444,9 @@
               (sX , sX⊩isSymmetricX)  perX .isSymmetric
               let
                 prover : ApplStrTerm as 1
-                prover = ` relF ̇ (` sX ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero) ̇ (` stFC ̇ (` pr₂ ̇ # zero))
+                prover = ` relF ̇ (` sX ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero) ̇ (` stFC ̇ (` pr₂ ̇ # zero))
               return
-                (λ* prover ,
+                (λ* prover ,
                   x y r r⊩∃x' 
                    transport
                      (propTruncIdempotent (F .relation .isPropValued _ _))
@@ -455,13 +455,13 @@
                        return
                          (subst
                             r'  r'   F .relation  (x , y))
-                           (sym (λ*ComputationRule prover r))
+                           (sym (λ*ComputationRule prover r))
                            (relF⊩isRelationalF
                              x' x y y
-                             (sX  (pr₁  r)) (pr₂  r) (stFC  (pr₂  r))
-                             (sX⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x')
+                             (sX  (pr₁  r)) (pr₂  r) (stFC  (pr₂  r))
+                             (sX⊩isSymmetricX x x' (pr₁  r) pr₁r⊩x~x')
                              pr₂r⊩Fx'y
-                             (stFC⊩isStrictCodomainF x' y (pr₂  r) pr₂r⊩Fx'y))))))
+                             (stFC⊩isStrictCodomainF x' y (pr₂  r) pr₂r⊩Fx'y))))))
         in
         eq/ _ _ (answer , F≤G→G≤F perX perY (composeFuncRel perX perX perY (idFuncRel perX) F) F answer))
       f
@@ -487,9 +487,9 @@
               (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
               let
                 prover : ApplStrTerm as 1
-                prover = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ # zero)
+                prover = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ # zero)
               return
-                (λ* prover ,
+                (λ* prover ,
                  x y r r⊩∃y' 
                   transport
                     (propTruncIdempotent (F .relation .isPropValued _ _))
@@ -498,8 +498,8 @@
                       return
                         (subst
                            r'  r'   F .relation  (x , y))
-                          (sym (λ*ComputationRule prover r))
-                          (relF⊩isRelationalF x x y' y (stFD  (pr₁  r)) (pr₁  r) (pr₂  r) (stFD⊩isStrictDomainF x y' (pr₁  r) pr₁r⊩Fxy') pr₁r⊩Fxy' pr₂r⊩y'~y)))))
+                          (sym (λ*ComputationRule prover r))
+                          (relF⊩isRelationalF x x y' y (stFD  (pr₁  r)) (pr₁  r) (pr₂  r) (stFD⊩isStrictDomainF x y' (pr₁  r) pr₁r⊩Fxy') pr₁r⊩Fxy' pr₂r⊩y'~y)))))
         in
         eq/ _ _ (answer , F≤G→G≤F perX perY (composeFuncRel perX perY perY F (idFuncRel perY)) F answer))
       f
@@ -528,9 +528,9 @@
             do
               let
                 prover : ApplStrTerm as 1
-                prover = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero))
+                prover = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero))
               return
-                (λ* prover ,
+                (λ* prover ,
                  x w r r⊩∃z 
                   transport
                     (propTruncIdempotent isPropPropTrunc)
@@ -542,21 +542,21 @@
                           (y ,
                             (subst
                                r'  r'   F .relation  (x , y))
-                              (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
+                              (sym (cong  x  pr₁  x) (λ*ComputationRule prover r)  pr₁pxy≡x _ _))
                               pr₁pr₁r⊩Fxy ,
                             return
                               (z ,
                                 ((subst
                                    r'  r'   G .relation  (y , z))
                                   (sym
-                                    (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r) 
-                                     cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                                    (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule prover r) 
+                                     cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
                                   pr₂pr₁r⊩Gyz) ,
                                  (subst
                                    r'  r'   H .relation  (z , w))
                                   (sym
-                                    (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r) 
-                                     cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                                    (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule prover r) 
+                                     cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
                                   pr₂r⊩Hzw)))))))))
         in
         eq/ _ _
diff --git a/docs/Realizability.Topos.MonicReprFuncRel.html b/docs/Realizability.Topos.MonicReprFuncRel.html
index 5c0622c..6768c8a 100644
--- a/docs/Realizability.Topos.MonicReprFuncRel.html
+++ b/docs/Realizability.Topos.MonicReprFuncRel.html
@@ -1,5 +1,5 @@
 
-Realizability.Topos.MonicReprFuncRel
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+Realizability.Topos.MonicReprFuncRel
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.Structure
@@ -23,7 +23,7 @@
   { ℓ' ℓ''}
   {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
   where
 
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
@@ -47,7 +47,7 @@
   opaque
     isInjectiveFuncRel : Type (ℓ-max  (ℓ-max ℓ' ℓ''))
     isInjectiveFuncRel =
-      ∃[ inj  A ] (∀ x x' y r₁ r₂  r₁   F .relation  (x , y)  r₂   F .relation  (x' , y)  (inj  r₁  r₂)   perX .equality  (x , x'))
+      ∃[ inj  A ] (∀ x x' y r₁ r₂  r₁   F .relation  (x , y)  r₂   F .relation  (x' , y)  (inj  r₁  r₂)   perX .equality  (x , x'))
 
   opaque
     unfolding isInjectiveFuncRel
@@ -79,11 +79,11 @@
                 let
                   realizer : ApplStrTerm as 1
                   realizer =
-                    ` relB ̇ (` stDA ̇ # zero) ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCA ̇ # zero))))) ̇
-                      (` injF ̇ (` pr₂ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCA ̇ # zero))))) ̇
-                      (` tlF ̇ (` stCA ̇ # zero)))
+                    ` relB ̇ (` stDA ̇ # zero) ̇ (` pr₁ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCA ̇ # zero))))) ̇
+                      (` injF ̇ (` pr₂ ̇ (` p ̇ (` pair ̇ # zero ̇ (` tlF ̇ (` stCA ̇ # zero))))) ̇
+                      (` tlF ̇ (` stCA ̇ # zero)))
                 return
-                  (λ* realizer ,
+                  (λ* realizer ,
                    z x r r⊩Azx 
                     transport
                       (propTruncIdempotent (B .relation .isPropValued _ _))
@@ -95,7 +95,7 @@
                         (x' , ⊩Bzx' , ⊩Fx'y)  
                           p⊩A⋆F≤B⋆F
                             z y
-                            (pair  r  (tlF  (stCA  r)))
+                            (pair  r  (tlF  (stCA  r)))
                             (return
                               (x ,
                               subst  r'  r'   A .relation  (z , x)) (sym (pr₁pxy≡x _ _)) r⊩Azx ,
@@ -105,7 +105,7 @@
                         return
                           (subst
                              r'  r'   B .relation  (z , x))
-                            (sym (λ*ComputationRule realizer r))
+                            (sym (λ*ComputationRule realizer r))
                             (relB⊩isRelationalB z z x' x _ _ _ z~z ⊩Bzx' x'~x)))))
           in
           eq/ A B (answer , F≤G→G≤F perZ perX A B answer))
@@ -185,83 +185,83 @@
         let
           cursed : ApplStrTerm as 2
           cursed =
-             (` pair ̇
-                  (` pair ̇
-                    (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)) ̇
-                    (` pair ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)) ̇ # one) ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # zero) ̇ (` stDF ̇ # one)) ̇ # zero))) ̇
-                  (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)))
+             (` pair ̇
+                  (` pair ̇
+                    (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)) ̇
+                    (` pair ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)) ̇ # one) ̇ (` pair ̇ (` pair ̇ (` stDF ̇ # zero) ̇ (` stDF ̇ # one)) ̇ # zero))) ̇
+                  (` pair ̇ (` stDF ̇ # one) ̇ (` stDF ̇ # zero)))
           realizer : ApplStrTerm as 2
-          realizer = ` t ̇ (` s ̇ (` pr₁ ̇ (` pr₂ ̇ (` p ̇ cursed)))) ̇ (` s ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₁ ̇ (` p ̇ cursed)))))
+          realizer = ` t ̇ (` s ̇ (` pr₁ ̇ (` pr₂ ̇ (` p ̇ cursed)))) ̇ (` s ̇ (` pr₂ ̇ (` pr₁ ̇ (` pr₁ ̇ (` p ̇ cursed)))))
         return
-          (λ*2 realizer ,
+          (λ*2 realizer ,
            x₁ x₂ y r₁ r₂ r₁⊩Fx₁y r₂⊩Fx₂y 
             let
               x₁~x₁ = stDF⊩isStrictDomainF x₁ y r₁ r₁⊩Fx₁y
               x₂~x₂ = stDF⊩isStrictDomainF x₂ y r₂ r₂⊩Fx₂y
               foo =
                 p⊩kπ₁≤kπ₂ (x₁ , x₂) x₁
-                (pair 
-                  (pair 
-                    (pair  (stDF  r₁)  (stDF  r₂)) 
-                    (pair  (pair  (pair  (stDF  r₁)  (stDF  r₂))  r₁)  (pair  (pair  (stDF  r₂)  (stDF  r₁))  r₂))) 
-                  (pair  (stDF  r₁)  (stDF  r₂)))
+                (pair 
+                  (pair 
+                    (pair  (stDF  r₁)  (stDF  r₂)) 
+                    (pair  (pair  (pair  (stDF  r₁)  (stDF  r₂))  r₁)  (pair  (pair  (stDF  r₂)  (stDF  r₁))  r₂))) 
+                  (pair  (stDF  r₁)  (stDF  r₂)))
                 (return
                   (((x₁ , x₂)) ,
-                  ((subst  r'  r'   perX .equality  (x₁ , x₁)) (sym (cong  x  pr₁  (pr₁  x)) (pr₁pxy≡x _ _)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _)) x₁~x₁ ,
-                    subst  r'  r'   perX .equality  (x₂ , x₂)) (sym (cong  x  pr₂  (pr₁  x)) (pr₁pxy≡x _ _)  cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _)) x₂~x₂) ,
+                  ((subst  r'  r'   perX .equality  (x₁ , x₁)) (sym (cong  x  pr₁  (pr₁  x)) (pr₁pxy≡x _ _)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _)) x₁~x₁ ,
+                    subst  r'  r'   perX .equality  (x₂ , x₂)) (sym (cong  x  pr₂  (pr₁  x)) (pr₁pxy≡x _ _)  cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _)) x₂~x₂) ,
                  return
                   (y ,
                     return
                       (x₁ ,
                         (subst  r'  r'   perX .equality  (x₁ , x₁))
                           (sym
-                            (cong  x  pr₁  (pr₁  (pr₁  (pr₂  x)))) (pr₁pxy≡x _ _) 
-                             cong  x  pr₁  (pr₁  (pr₁  x))) (pr₂pxy≡y _ _) 
-                             cong  x  pr₁  (pr₁  x)) (pr₁pxy≡x _ _) 
-                             cong  x  pr₁  x) (pr₁pxy≡x _ _) 
+                            (cong  x  pr₁  (pr₁  (pr₁  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  (pr₁  (pr₁  x))) (pr₂pxy≡y _ _) 
+                             cong  x  pr₁  (pr₁  x)) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  x) (pr₁pxy≡x _ _) 
                              pr₁pxy≡x _ _))
                           x₁~x₁ ,
                          subst  r'  r'   perX .equality  (x₂ , x₂))
                            (sym
-                             (cong  x  pr₂  (pr₁  (pr₁  (pr₂  x)))) (pr₁pxy≡x _ _) 
-                              cong  x  pr₂  (pr₁  (pr₁  x))) (pr₂pxy≡y _ _) 
-                              cong  x  pr₂  (pr₁  x)) (pr₁pxy≡x _ _) 
-                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                             (cong  x  pr₂  (pr₁  (pr₁  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₁  (pr₁  x))) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  (pr₁  x)) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
                               pr₂pxy≡y _ _))
                            x₂~x₂) ,
                          subst  r'  r'   F .relation  (x₁ , y))
                            (sym
-                             (cong  x  pr₂  (pr₁  (pr₂  x))) (pr₁pxy≡x _ _) 
-                              cong  x  pr₂  (pr₁  x)) (pr₂pxy≡y _ _) 
-                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                             (cong  x  pr₂  (pr₁  (pr₂  x))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₁  x)) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
                               pr₂pxy≡y _ _))
                            r₁⊩Fx₁y) ,
                     return
                       (x₂ ,
                         (subst  r'  r'   perX .equality  (x₂ , x₂))
                           (sym
-                            (cong  x  pr₁  (pr₁  (pr₂  (pr₂  x)))) (pr₁pxy≡x _ _) 
-                             cong  x  pr₁  (pr₁  (pr₂  x))) (pr₂pxy≡y _ _) 
-                             cong  x  pr₁  (pr₁  x)) (pr₂pxy≡y _ _) 
-                             cong  x  pr₁  x) (pr₁pxy≡x _ _) 
+                            (cong  x  pr₁  (pr₁  (pr₂  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                             cong  x  pr₁  (pr₁  (pr₂  x))) (pr₂pxy≡y _ _) 
+                             cong  x  pr₁  (pr₁  x)) (pr₂pxy≡y _ _) 
+                             cong  x  pr₁  x) (pr₁pxy≡x _ _) 
                              pr₁pxy≡x _ _))
                           x₂~x₂ ,
                          subst  r'  r'   perX .equality  (x₁ , x₁))
                            (sym
-                             (cong  x  pr₂  (pr₁  (pr₂  (pr₂  x)))) (pr₁pxy≡x _ _) 
-                              cong  x  pr₂  (pr₁  (pr₂  x))) (pr₂pxy≡y _ _) 
-                              cong  x  pr₂  (pr₁  x)) (pr₂pxy≡y _ _) 
-                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                             (cong  x  pr₂  (pr₁  (pr₂  (pr₂  x)))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₁  (pr₂  x))) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  (pr₁  x)) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  x) (pr₁pxy≡x _ _) 
                               pr₂pxy≡y _ _))
                            x₁~x₁) ,
                          subst  r'  r'   F .relation  (x₂ , y))
                            (sym
-                             (cong  x  pr₂  (pr₂  (pr₂  x))) (pr₁pxy≡x _ _) 
-                              cong  x  pr₂  (pr₂  x)) (pr₂pxy≡y _ _) 
-                              cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                             (cong  x  pr₂  (pr₂  (pr₂  x))) (pr₁pxy≡x _ _) 
+                              cong  x  pr₂  (pr₂  x)) (pr₂pxy≡y _ _) 
+                              cong  x  pr₂  x) (pr₂pxy≡y _ _) 
                               pr₂pxy≡y _ _)) r₂⊩Fx₂y))) ,
-                         subst  r'  r'   perX .equality  (x₁ , x₁)) (sym (cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _)) x₁~x₁ ,
-                         subst  r'  r'   perX .equality  (x₂ , x₂)) (sym (cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _)) x₂~x₂))
+                         subst  r'  r'   perX .equality  (x₁ , x₁)) (sym (cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _)) x₁~x₁ ,
+                         subst  r'  r'   perX .equality  (x₂ , x₂)) (sym (cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _)) x₂~x₂))
             in
             transport
               (propTruncIdempotent (perX .equality .isPropValued _ _))
@@ -275,6 +275,6 @@
                 return
                   (subst
                      r'  r'   perX .equality  (x₁ , x₂))
-                    (sym (λ*2ComputationRule realizer r₁ r₂))
+                    (sym (λ*2ComputationRule realizer r₁ r₂))
                     (t⊩isTransitiveEqX x₁ x₂' x₂ _ _ (s⊩isSymmetricEqX x₂' x₁ _ x₂'~x₁) (s⊩isSymmetricEqX x₂ x₂' _ x₂~x₂'))))))
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.Object.html b/docs/Realizability.Topos.Object.html index 1737fdc..c8af074 100644 --- a/docs/Realizability.Topos.Object.html +++ b/docs/Realizability.Topos.Object.html @@ -1,5 +1,5 @@ -Realizability.Topos.Object
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+Realizability.Topos.Object
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.Structure
@@ -18,7 +18,7 @@
   { ℓ' ℓ''}
   {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
   where
   
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
@@ -31,8 +31,8 @@
 record isPartialEquivalenceRelation (X : Type ℓ') (equality : Predicate (X × X)) : Type (ℓ-max (ℓ-max (ℓ-suc ) (ℓ-suc ℓ')) (ℓ-suc ℓ'')) where
   field
     isSetX : isSet X
-    isSymmetric : ∃[ s  A ] (∀ x y r  r   equality  (x , y)  (s  r)   equality  (y , x))
-    isTransitive : ∃[ t  A ] (∀ x y z a b  a   equality  (x , y)  b   equality  (y , z)  (t  a  b)   equality  (x , z))
+    isSymmetric : ∃[ s  A ] (∀ x y r  r   equality  (x , y)  (s  r)   equality  (y , x))
+    isTransitive : ∃[ t  A ] (∀ x y z a b  a   equality  (x , y)  b   equality  (y , z)  (t  a  b)   equality  (x , z))
 
 open isPartialEquivalenceRelation
 isPropIsPartialEquivalenceRelation :  {X : Type ℓ'}  (equality : Predicate (X × X))  isProp (isPartialEquivalenceRelation X equality)
diff --git a/docs/Realizability.Topos.ResizedPredicate.html b/docs/Realizability.Topos.ResizedPredicate.html
index 0a56df2..c7deb1b 100644
--- a/docs/Realizability.Topos.ResizedPredicate.html
+++ b/docs/Realizability.Topos.ResizedPredicate.html
@@ -15,8 +15,8 @@
   {}
   {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
-  (resizing : hPropResizing )
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (resizing : hPropResizing )
   where
 
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = } {ℓ'' = } ca
@@ -71,14 +71,14 @@
   open PartialEquivalenceRelation
 
   entailmentResizedPredicate :  (ϕ ψ : ResizedPredicate X)  A  Type 
-  entailmentResizedPredicate ϕ ψ r =  (x : X) (a : A) (⊩ϕx : a   toPredicate ϕ  x)  (r  a)   toPredicate ψ  x
+  entailmentResizedPredicate ϕ ψ r =  (x : X) (a : A) (⊩ϕx : a   toPredicate ϕ  x)  (r  a)   toPredicate ψ  x
 
   isPropEntailmentResizedPredicate :  ϕ ψ a  isProp (entailmentResizedPredicate ϕ ψ a)
   isPropEntailmentResizedPredicate ϕ ψ a =
     isPropΠ λ x  isPropΠ λ b  isPropΠ λ _  (toPredicate ψ) .isPropValued _ _
 
   isStrictResizedPredicate :  (ϕ : ResizedPredicate X)  A  Type 
-  isStrictResizedPredicate ϕ r =  (x : X) (a : A) (⊩ϕx : a   toPredicate ϕ  x)  (r  a)   perX .equality  (x , x)
+  isStrictResizedPredicate ϕ r =  (x : X) (a : A) (⊩ϕx : a   toPredicate ϕ  x)  (r  a)   perX .equality  (x , x)
 
   isPropIsStrictResizedPredicate :  ϕ r  isProp (isStrictResizedPredicate ϕ r)
   isPropIsStrictResizedPredicate ϕ r =
@@ -86,7 +86,7 @@
 
   isRelationalResizedPredicate :  (ϕ : ResizedPredicate X)  A  Type 
   isRelationalResizedPredicate ϕ r =
-     (x x' : X) (a b : A) (⊩x~x' : a   perX .equality  (x , x')) (⊩ϕx : b   toPredicate ϕ  x)  (r  a  b)   toPredicate ϕ  x'
+     (x x' : X) (a b : A) (⊩x~x' : a   perX .equality  (x , x')) (⊩ϕx : b   toPredicate ϕ  x)  (r  a  b)   toPredicate ϕ  x'
 
   isPropIsRelationalResizedPredicate :  ϕ r  isProp (isRelationalResizedPredicate ϕ r)
   isPropIsRelationalResizedPredicate ϕ r =
diff --git a/docs/Realizability.Topos.StrictRelation.html b/docs/Realizability.Topos.StrictRelation.html
index 5594652..d509d03 100644
--- a/docs/Realizability.Topos.StrictRelation.html
+++ b/docs/Realizability.Topos.StrictRelation.html
@@ -1,5 +1,5 @@
 
-Realizability.Topos.StrictRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+Realizability.Topos.StrictRelation
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.Structure
@@ -26,7 +26,7 @@
   { ℓ' ℓ''}
   {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
   where
 
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
@@ -47,8 +47,8 @@
 
 record isStrictRelation {X : Type ℓ'} (perX : PartialEquivalenceRelation X) (ϕ : Predicate X) : Type (ℓ-max  (ℓ-max ℓ' ℓ'')) where
   field
-    isStrict : ∃[ st  A ] (∀ x r  r   ϕ  x  (st  r)   perX .equality  (x , x))
-    isRelational : ∃[ rel  A ] (∀ x x' r s  r   ϕ  x  s   perX .equality  (x , x')  (rel  r  s)   ϕ  x')
+    isStrict : ∃[ st  A ] (∀ x r  r   ϕ  x  (st  r)   perX .equality  (x , x))
+    isRelational : ∃[ rel  A ] (∀ x x' r s  r   ϕ  x  s   perX .equality  (x , x')  (rel  r  s)   ϕ  x')
 
 record StrictRelation {X : Type ℓ'} (perX : PartialEquivalenceRelation X) : Type (ℓ-max (ℓ-suc ) (ℓ-max (ℓ-suc ℓ') (ℓ-suc ℓ''))) where
   field
@@ -63,7 +63,7 @@
   {-# TERMINATING #-}
   subPer : PartialEquivalenceRelation X
   Predicate.isSetX (equality subPer) = isSet× (perX .isSetX) (perX .isSetX)
-   equality subPer  (x , x') r = (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   ϕ .predicate  x
+   equality subPer  (x , x') r = (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   ϕ .predicate  x
   isPropValued (equality subPer) (x , x') r = isProp× (perX .equality .isPropValued _ _) (ϕ .predicate .isPropValued _ _)
   isPartialEquivalenceRelation.isSetX (isPerEquality subPer) = perX .isSetX
   isPartialEquivalenceRelation.isSymmetric (isPerEquality subPer) =
@@ -73,17 +73,17 @@
       (relϕ , relϕ⊩isRelationalϕ)  ϕ .isRelational
       let
         realizer : ApplStrTerm as 1
-        realizer = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` relϕ ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero))
+        realizer = ` pair ̇ (` s ̇ (` pr₁ ̇ # zero)) ̇ (` relϕ ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero))
       return
-        (λ* realizer ,
+        (λ* realizer ,
          { x x' r (pr₁r⊩x~x' , pr₂r⊩ϕx) 
           subst
              r'  r'   perX .equality  (x' , x))
-            (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+            (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
             (s⊩isSymmetricX x x' _ pr₁r⊩x~x') ,
           subst
              r'  r'   ϕ .predicate  x')
-            (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+            (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
             (relϕ⊩isRelationalϕ x x' _ _ pr₂r⊩ϕx pr₁r⊩x~x') }))
   isPartialEquivalenceRelation.isTransitive (isPerEquality subPer) =
     do
@@ -91,17 +91,17 @@
       (relϕ , relϕ⊩isRelationalϕ)  ϕ .isRelational
       let
         realizer : ApplStrTerm as 2
-        realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # one)
+        realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # one)
       return
-        (λ*2 realizer ,
+        (λ*2 realizer ,
          { x₁ x₂ x₃ a b (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₂~x₃ , ⊩ϕx₂) 
           subst
              r'  r'   perX .equality  (x₁ , x₃))
-            (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer a b)  pr₁pxy≡x _ _))
+            (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer a b)  pr₁pxy≡x _ _))
             (t⊩isTransitiveX x₁ x₂ x₃ _ _ ⊩x₁~x₂ ⊩x₂~x₃) ,
           subst
              r'  r'   ϕ .predicate  x₁)
-            (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer a b)  pr₂pxy≡y _ _))
+            (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer a b)  pr₂pxy≡y _ _))
             ⊩ϕx₁ }))
 
   opaque
@@ -109,56 +109,56 @@
     {-# TERMINATING #-}
     incFuncRel : FunctionalRelation subPer perX
     Predicate.isSetX (relation incFuncRel) = isSet× (perX .isSetX) (perX .isSetX)
-    Predicate.∣ relation incFuncRel  (x , x') r = (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   ϕ .predicate  x
+    Predicate.∣ relation incFuncRel  (x , x') r = (pr₁  r)   perX .equality  (x , x') × (pr₂  r)   ϕ .predicate  x
     Predicate.isPropValued (relation incFuncRel) (x , x') r = isProp× (perX .equality .isPropValued _ _) (ϕ .predicate .isPropValued _ _)
     isFunctionalRelation.isStrictDomain (isFuncRel incFuncRel) =
       do
         (stD , stD⊩isStrictDomain)  idFuncRel perX .isStrictDomain
         let
           realizer : ApplStrTerm as 1
-          realizer = ` pair ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+          realizer = ` pair ̇ (` stD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
         return
-          (λ* realizer ,
+          (λ* realizer ,
            { x x' r (⊩x~x' , ⊩ϕx) 
-            (subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (stD⊩isStrictDomain x x' _ ⊩x~x')) ,
-            (subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx) }))
+            (subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (stD⊩isStrictDomain x x' _ ⊩x~x')) ,
+            (subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx) }))
     isFunctionalRelation.isStrictCodomain (isFuncRel incFuncRel) =
       do
         (stC , stC⊩isStrictCodomain)  idFuncRel perX .isStrictCodomain
         let
           realizer : ApplStrTerm as 1
-          realizer = ` stC ̇ (` pr₁ ̇ # zero)
+          realizer = ` stC ̇ (` pr₁ ̇ # zero)
         return
-          (λ* realizer ,
-           { x x' r (⊩x~x' , ⊩ϕx)  subst  r'  r'   perX .equality  (x' , x')) (sym (λ*ComputationRule realizer r)) (stC⊩isStrictCodomain x x' _ ⊩x~x')}))
+          (λ* realizer ,
+           { x x' r (⊩x~x' , ⊩ϕx)  subst  r'  r'   perX .equality  (x' , x')) (sym (λ*ComputationRule realizer r)) (stC⊩isStrictCodomain x x' _ ⊩x~x')}))
     isFunctionalRelation.isRelational (isFuncRel incFuncRel) =
       do
         (relX , relX⊩isRelationalX)  idFuncRel perX .isRelational
         (relϕ , relϕ⊩isRelationalϕ)  ϕ .isRelational
         let
           realizer : ApplStrTerm as 3
-          realizer = ` pair ̇ (` relX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇ (` relϕ ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # two))
+          realizer = ` pair ̇ (` relX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ # zero) ̇ (` relϕ ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # two))
         return
-          (λ*3 realizer ,
+          (λ*3 realizer ,
            { x₁ x₂ x₃ x₄ a b c (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₁~x₃ , ⊩ϕx₁') c⊩x₃~x₄ 
             subst
                r'  r'   perX .equality  (x₂ , x₄))
-              (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
+              (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
               (relX⊩isRelationalX x₁ x₂ x₃ x₄ _ _ _ ⊩x₁~x₂ ⊩x₁~x₃ c⊩x₃~x₄) ,
             subst
                r'  r'   ϕ .predicate  x₂)
-              (sym (cong  x  pr₂  x) (λ*3ComputationRule realizer a b c)  pr₂pxy≡y _ _))
+              (sym (cong  x  pr₂  x) (λ*3ComputationRule realizer a b c)  pr₂pxy≡y _ _))
               (relϕ⊩isRelationalϕ x₁ x₂ _ _ ⊩ϕx₁ ⊩x₁~x₂) }))
     isFunctionalRelation.isSingleValued (isFuncRel incFuncRel) =
       do
         (sv , sv⊩isSingleValuedX)  idFuncRel perX .isSingleValued
         let
           realizer : ApplStrTerm as 2
-          realizer = ` sv ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
+          realizer = ` sv ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)
         return
-          (λ*2 realizer ,
+          (λ*2 realizer ,
            { x x' x'' r₁ r₂ (⊩x~x' , ⊩ϕx) (⊩x~x'' , ⊩ϕx') 
-            subst  r'  r'   perX .equality  (x' , x'')) (sym (λ*2ComputationRule realizer r₁ r₂)) (sv⊩isSingleValuedX x x' x'' _ _ ⊩x~x' ⊩x~x'') }))
+            subst  r'  r'   perX .equality  (x' , x'')) (sym (λ*2ComputationRule realizer r₁ r₂)) (sv⊩isSingleValuedX x x' x'' _ _ ⊩x~x' ⊩x~x'') }))
     isFunctionalRelation.isTotal (isFuncRel incFuncRel) =
       do
         return
@@ -166,8 +166,8 @@
            { x r (pr₁r⊩x~x , pr₂r⊩ϕx) 
             return
               (x ,
-              subst  r'  r'   perX .equality  (x , x)) (cong  x  pr₁  x) (sym (Ida≡a _))) pr₁r⊩x~x ,
-              subst  r'  r'   ϕ .predicate  x) (cong  x  pr₂  x) (sym (Ida≡a _))) pr₂r⊩ϕx) }))
+              subst  r'  r'   perX .equality  (x , x)) (cong  x  pr₁  x) (sym (Ida≡a _))) pr₁r⊩x~x ,
+              subst  r'  r'   ϕ .predicate  x) (cong  x  pr₂  x) (sym (Ida≡a _))) pr₂r⊩ϕx) }))
 
   opaque
     unfolding isInjectiveFuncRel
@@ -179,17 +179,17 @@
         (s , s⊩isSymmetricX)  perX .isSymmetric
         let
           realizer : ApplStrTerm as 2
-          realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` s ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ # one)
+          realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ # one) ̇ (` s ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ # one)
         return
-          (λ*2 realizer ,
+          (λ*2 realizer ,
            x₁ x₂ x₃ r₁ r₂ (⊩x₁~x₃ , ⊩ϕx₁) (⊩x₂~x₃ , ⊩ϕx₂) 
             subst
                r'  r'   perX .equality  (x₁ , x₂))
-              (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
+              (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
               (t⊩isTransitiveX x₁ x₃ x₂ _ _ ⊩x₁~x₃ (s⊩isSymmetricX x₂ x₃ _ ⊩x₂~x₃)) ,
             subst
                r'  r'   ϕ .predicate  x₁)
-              (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁  r₂)  pr₂pxy≡y _ _))
+              (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁  r₂)  pr₂pxy≡y _ _))
               ⊩ϕx₁))
 
   isMonicInc : isMonic RT [ incFuncRel ]
@@ -225,9 +225,9 @@
       (stDF , stDF⊩isStrictDomainF)  F .isStrictDomain
       let
         realizer : ApplStrTerm as 2
-        realizer = ` relF ̇ (` stDF ̇ # one) ̇ # one ̇ # zero
+        realizer = ` relF ̇ (` stDF ̇ # one) ̇ # one ̇ # zero
       return
-        (λ*2 realizer ,
+        (λ*2 realizer ,
          x x' r s r⊩∃y s⊩x~x' 
           do
             (y , ⊩Fyx)  r⊩∃y
@@ -235,7 +235,7 @@
               (y ,
               subst
                  r'  r'   F .relation  (y , x'))
-                (sym (λ*2ComputationRule realizer r s))
+                (sym (λ*2ComputationRule realizer r s))
                 (relF⊩isRelationalF y y x x' _ _ _ (stDF⊩isStrictDomainF y x _ ⊩Fyx) ⊩Fyx s⊩x~x'))))
 
   perψ : PartialEquivalenceRelation X
@@ -255,46 +255,46 @@
       (stCF , stCF⊩isStrictCodomain)  F .isStrictCodomain
       let
         realizer : ApplStrTerm as 1
-        realizer = ` pair ̇ (` stCF ̇ # zero) ̇ # zero
+        realizer = ` pair ̇ (` stCF ̇ # zero) ̇ # zero
       return
-        (λ* realizer ,
+        (λ* realizer ,
          y x r ⊩Fyx 
           subst
              r'  r'   perX .equality  (x , x))
-            (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+            (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
             (stCF⊩isStrictCodomain y x _ ⊩Fyx) ,
            y ,
             subst
                r'  r'   F .relation  (y , x))
-              (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+              (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
               ⊩Fyx ∣₁))
   isFunctionalRelation.isRelational (isFuncRel perY≤perψFuncRel) =
     do
       (relF , relF⊩isRelationalF)  F .isRelational
       let
         realizer : ApplStrTerm as 3
-        realizer = ` relF ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
+        realizer = ` relF ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
       return
-        (λ*3 realizer ,
+        (λ*3 realizer ,
          { y y' x x' a b c ⊩y~y' ⊩Fyx (⊩x~x' , ⊩Fy''x) 
-          subst  r'  r'   F .relation  (y' , x')) (sym (λ*3ComputationRule realizer a b c)) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') }))
+          subst  r'  r'   F .relation  (y' , x')) (sym (λ*3ComputationRule realizer a b c)) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') }))
   isFunctionalRelation.isSingleValued (isFuncRel perY≤perψFuncRel) =
     do
       (svF , svF⊩isSingleValuedF)  F .isSingleValued
       let
         realizer : ApplStrTerm as 2
-        realizer = ` pair ̇ (` svF ̇ # one ̇ # zero) ̇ # one
+        realizer = ` pair ̇ (` svF ̇ # one ̇ # zero) ̇ # one
       return
-        (λ*2 realizer ,
+        (λ*2 realizer ,
          y x x' r₁ r₂ ⊩Fyx ⊩Fyx' 
           subst
              r'  r'   perX .equality  (x , x'))
-            (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
+            (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
             (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') ,
            y ,
             (subst
                r'  r'   F .relation  (y , x))
-              (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _))
+              (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _))
               ⊩Fyx) ∣₁))
   isFunctionalRelation.isTotal (isFuncRel perY≤perψFuncRel) =
     do
@@ -319,9 +319,9 @@
             (relF , relF⊩isRelationalF)  F .isRelational
             let
               realizer : ApplStrTerm as 1
-              realizer = ` relF ̇ (` stDF ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+              realizer = ` relF ̇ (` stDF ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
             return
-              (λ* realizer ,
+              (λ* realizer ,
                y x r r⊩∃x' 
                 transport
                   (propTruncIdempotent (F .relation .isPropValued _ _))
@@ -330,7 +330,7 @@
                     return
                       (subst
                          r  r   F .relation  (y , x))
-                        (sym (λ*ComputationRule realizer r))
+                        (sym (λ*ComputationRule realizer r))
                         (relF⊩isRelationalF y y x' x _ _ _ (stDF⊩isStrictDomainF y x' _ ⊩Fyx') ⊩Fyx' ⊩x'~x)))))
       in
       eq/ _ _ (answer , F≤G→G≤F perY perX (composeFuncRel _ _ _ perY≤perψFuncRel (InducedSubobject.incFuncRel perX ψ)) F answer)
@@ -347,19 +347,19 @@
         (stCF , stCF⊩isStrictCodomainF)  F .isStrictCodomain
         let
           realizer : ApplStrTerm as 1
-          realizer = ` pair ̇ (` stCF ̇ # zero) ̇ # zero
+          realizer = ` pair ̇ (` stCF ̇ # zero) ̇ # zero
         return
-          (λ* realizer ,
+          (λ* realizer ,
            x y r ⊩Fyx 
             (subst
                r'  r'   perX .equality  (x , x))
-              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
               (stCF⊩isStrictCodomainF y x _ ⊩Fyx)) ,
             (return
               (y ,
               (subst
                  r'  r'   F .relation  (y , x))
-                (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
                 ⊩Fyx)))))
     isFunctionalRelation.isStrictCodomain (isFuncRel perψ≤perYFuncRel) =
       do
@@ -372,11 +372,11 @@
         (relF , relF⊩isRelationalF)  F .isRelational
         let
           realizer : ApplStrTerm as 3
-          realizer = ` relF ̇ # zero ̇ # one ̇ (` pr₁ ̇ # two)
+          realizer = ` relF ̇ # zero ̇ # one ̇ (` pr₁ ̇ # two)
         return
-          (λ*3 realizer ,
+          (λ*3 realizer ,
            { x x' y y' a b c (⊩x~x' , ⊩ψx) ⊩Fyx ⊩y~y' 
-            subst  r'  r'   F .relation  (y' , x')) (sym (λ*3ComputationRule realizer a b c)) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') }))
+            subst  r'  r'   F .relation  (y' , x')) (sym (λ*3ComputationRule realizer a b c)) (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x') }))
     isFunctionalRelation.isSingleValued (isFuncRel perψ≤perYFuncRel) =
       let
         isInjectiveFuncRelF = isMonic→isInjectiveFuncRel perY perX F isMonicF
@@ -405,9 +405,9 @@
             (svF , svF⊩isSingleValuedF)  F .isSingleValued
             let
               realizer : ApplStrTerm as 1
-              realizer = ` pair ̇ (` svF ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero)
+              realizer = ` pair ̇ (` svF ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero)
             return
-              (λ* realizer ,
+              (λ* realizer ,
                x x' r r⊩∃y 
                 transport
                   (propTruncIdempotent (isProp× (perX .equality .isPropValued _ _) isPropPropTrunc))
@@ -416,13 +416,13 @@
                     return
                       (subst
                          r'  r'   perX .equality  (x , x'))
-                        (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                        (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
                         (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') ,
                       return
                         (y ,
                         (subst
                            r'  r'   F .relation  (y , x))
-                          (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                          (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
                           ⊩Fyx))))))
       in
       eq/ _ _ (answer , F≤G→G≤F perψ perX (composeFuncRel _ _ _ perψ≤perYFuncRel F) (InducedSubobject.incFuncRel perX ψ) answer)
@@ -474,9 +474,9 @@
             (q , q⊩incϕ≤F⋆incψ)  q
             let
               realizer : ApplStrTerm as 1
-              realizer = ` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # zero) ̇ # zero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # zero) ̇ # zero))))
+              realizer = ` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # zero) ̇ # zero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` q ̇ (` pair ̇ (` stϕ ̇ # zero) ̇ # zero))))
             return
-              (λ* realizer ,
+              (λ* realizer ,
                x a a⊩ϕx 
                 transport
                   (propTruncIdempotent (ψ .predicate .isPropValued _ _))
@@ -484,10 +484,10 @@
                     (x' , ⊩Fxx' , ⊩x'~x , ⊩ψx') 
                       q⊩incϕ≤F⋆incψ
                         x x
-                        (pair  (stϕ  a)  a)
+                        (pair  (stϕ  a)  a)
                         ((subst  r'  r'   perX .equality  (x , x)) (sym (pr₁pxy≡x _ _)) (stϕ⊩isStrictϕ x a a⊩ϕx)) ,
                          (subst  r'  r'   ϕ .predicate  x) (sym (pr₂pxy≡y _ _)) a⊩ϕx))
-                    return (subst  r'  r'   ψ .predicate  x) (sym (λ*ComputationRule realizer a)) (relψ⊩isRelationalψ x' x _ _ ⊩ψx' ⊩x'~x))))))
+                    return (subst  r'  r'   ψ .predicate  x) (sym (λ*ComputationRule realizer a)) (relψ⊩isRelationalψ x' x _ _ ⊩ψx' ⊩x'~x))))))
         f
         f⋆incψ≡incϕ
 
@@ -500,36 +500,36 @@
     {-# TERMINATING #-}
     funcRel : FunctionalRelation perϕ perψ
     Predicate.isSetX (relation funcRel) = isSet× (perX .isSetX) (perX .isSetX)
-    Predicate.∣ relation funcRel  (x , x') r = (pr₁  r)   perX .equality  (x , x') × ((pr₁  (pr₂  r))   ϕ .predicate  x) × ((pr₂  (pr₂  r))   ψ .predicate  x)
+    Predicate.∣ relation funcRel  (x , x') r = (pr₁  r)   perX .equality  (x , x') × ((pr₁  (pr₂  r))   ϕ .predicate  x) × ((pr₂  (pr₂  r))   ψ .predicate  x)
     Predicate.isPropValued (relation funcRel) (x , x') r = isProp× (perX .equality .isPropValued _ _) (isProp× (ϕ .predicate .isPropValued _ _) (ψ .predicate .isPropValued _ _))
     isFunctionalRelation.isStrictDomain (isFuncRel funcRel) =
       do
         (stϕ , stϕ⊩isStrictϕ)  ϕ .isStrict
         let
           realizer : ApplStrTerm as 1
-          realizer = ` pair ̇ (` stϕ ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+          realizer = ` pair ̇ (` stϕ ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
         return
-          (λ* realizer ,
+          (λ* realizer ,
            { x x' r (⊩x~x' , ⊩ϕx , ⊩ψx) 
-            subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (stϕ⊩isStrictϕ x _ ⊩ϕx) ,
-            subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx}))
+            subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _)) (stϕ⊩isStrictϕ x _ ⊩ϕx) ,
+            subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx}))
     isFunctionalRelation.isStrictCodomain (isFuncRel funcRel) =
       do
         (stCX , stCX⊩isStrictCodomainX)  idFuncRel perX .isStrictCodomain
         (relψ , relψ⊩isRelationalψ)  ψ .isRelational
         let
           realizer : ApplStrTerm as 1
-          realizer = ` pair ̇ (` stCX ̇ (` pr₁ ̇ # zero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero))
+          realizer = ` pair ̇ (` stCX ̇ (` pr₁ ̇ # zero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) ̇ (` pr₁ ̇ # zero))
         return
-          (λ* realizer ,
+          (λ* realizer ,
            { x x' r (⊩x~x' , ⊩ϕx , ⊩ψx) 
             subst
                r'  r'   perX .equality  (x' , x'))
-              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
               (stCX⊩isStrictCodomainX x x' _ ⊩x~x') ,
             subst
                r'  r'   ψ .predicate  x')
-              (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+              (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
               (relψ⊩isRelationalψ x x' _ _ ⊩ψx ⊩x~x')}))
     isFunctionalRelation.isRelational (isFuncRel funcRel) =
       do
@@ -539,21 +539,21 @@
         let
           realizer : ApplStrTerm as 3
           realizer =
-            ` pair ̇ (` relX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` relϕ ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # two)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` pr₁ ̇ # two)))
+            ` pair ̇ (` relX ̇ (` pr₁ ̇ # two) ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` relϕ ̇ (` pr₂ ̇ # two) ̇ (` pr₁ ̇ # two)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` pr₁ ̇ # two)))
         return
-          (λ*3 realizer ,
+          (λ*3 realizer ,
           λ { x₁ x₂ x₃ x₄ a b c (⊩x₁~x₂ , ⊩ϕx₁) (⊩x₁~x₃ , ⊩'ϕx₁ , ⊩ψx₁) (⊩x₃~x₄ , ⊩ψx₃) 
             subst
                r'  r'   perX .equality  (x₂ , x₄))
-              (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
+              (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
               (relX⊩isRelationalX x₁ x₂ x₃ x₄ _ _ _ ⊩x₁~x₂ ⊩x₁~x₃ ⊩x₃~x₄) ,
             subst
                r'  r'   ϕ .predicate  x₂)
-              (sym (cong  x  pr₁  (pr₂  x)) (λ*3ComputationRule realizer a b c)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+              (sym (cong  x  pr₁  (pr₂  x)) (λ*3ComputationRule realizer a b c)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
               (relϕ⊩isRelationalϕ x₁ x₂ _ _ ⊩ϕx₁ ⊩x₁~x₂) ,
             subst
                r'  r'   ψ .predicate  x₂)
-              (sym (cong  x  pr₂  (pr₂  x)) (λ*3ComputationRule realizer a b c)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+              (sym (cong  x  pr₂  (pr₂  x)) (λ*3ComputationRule realizer a b c)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
               (relψ⊩isRelationalψ x₁ x₂ _ _ ⊩ψx₁ ⊩x₁~x₂)})
     isFunctionalRelation.isSingleValued (isFuncRel funcRel) =
       do
@@ -561,35 +561,35 @@
         (relψ , relψ⊩isRelationalψ)  ψ .isRelational
         let
           realizer : ApplStrTerm as 2
-          realizer = ` pair ̇ (` svX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` pr₁ ̇ # one))
+          realizer = ` pair ̇ (` svX ̇ (` pr₁ ̇ # one) ̇ (` pr₁ ̇ # zero)) ̇ (` relψ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` pr₁ ̇ # one))
         return
-          (λ*2 realizer ,
+          (λ*2 realizer ,
            { x₁ x₂ x₃ r₁ r₂ (⊩x₁~x₂ , ⊩ϕx , ⊩ψx) (⊩x₁~x₃ , ⊩'ϕx , ⊩'ψx) 
-            (subst  r'  r'   perX .equality  (x₂ , x₃)) (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _)) (svX⊩isSingleValuedX x₁ x₂ x₃ _ _ ⊩x₁~x₂ ⊩x₁~x₃)) ,
-             subst  r'  r'   ψ .predicate  x₂) (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _)) (relψ⊩isRelationalψ x₁ x₂ _ _ ⊩ψx ⊩x₁~x₂)}))
+            (subst  r'  r'   perX .equality  (x₂ , x₃)) (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _)) (svX⊩isSingleValuedX x₁ x₂ x₃ _ _ ⊩x₁~x₂ ⊩x₁~x₃)) ,
+             subst  r'  r'   ψ .predicate  x₂) (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _)) (relψ⊩isRelationalψ x₁ x₂ _ _ ⊩ψx ⊩x₁~x₂)}))
     isFunctionalRelation.isTotal (isFuncRel funcRel) =
       do
         (tl , tl⊩isTotalIncψ)  incψ .isTotal
         (s , s⊩ϕ≤ψ)  ϕ≤ψ
         let
           realizer : ApplStrTerm as 1
-          realizer = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` pair ̇ (` pr₂ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero)))
+          realizer = ` pair ̇ (` pr₁ ̇ # zero) ̇ (` pair ̇ (` pr₂ ̇ # zero) ̇ (` s ̇ (` pr₂ ̇ # zero)))
         return
-          (λ* realizer ,
+          (λ* realizer ,
            { x r (⊩x~x , ⊩ϕx) 
             return
               (x ,
               subst
                  r'  r'   perX .equality  (x , x))
-                (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
                 ⊩x~x ,
               subst
                  r'  r'   ϕ .predicate  x)
-                (sym (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
+                (sym (cong  x  pr₁  (pr₂  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₂pxy≡y _ _)  pr₁pxy≡x _ _))
                 ⊩ϕx ,
               subst
                  r'  r'   ψ .predicate  x)
-                (sym (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule realizer r)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
+                (sym (cong  x  pr₂  (pr₂  x)) (λ*ComputationRule realizer r)  cong  x  pr₂  x) (pr₂pxy≡y _ _)  pr₂pxy≡y _ _))
                 (s⊩ϕ≤ψ x _ ⊩ϕx))}))
     
     funcRel⋆incψ≡incϕ : [ funcRel ]  [ incψ ]  [ incϕ ]
@@ -600,9 +600,9 @@
             (t , t⊩isTransitiveX)  perX .isTransitive
             let
               realizer : ApplStrTerm as 1
-              realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)))
+              realizer = ` pair ̇ (` t ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)))
             return
-              (λ* realizer ,
+              (λ* realizer ,
                { x x' r ⊩∃x'' 
                 transport
                   (propTruncIdempotent (isPropΣ (perX .equality .isPropValued _ _) λ _  ϕ .predicate .isPropValued _ _))
@@ -611,11 +611,11 @@
                     return
                       ((subst
                          r'  r'   perX .equality  (x , x'))
-                        (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                        (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
                         (t⊩isTransitiveX x x'' x' _ _ ⊩x~x'' ⊩x''~x')) ,
                        (subst
                           r'  r'   ϕ .predicate  x)
-                         (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
+                         (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _))
                          ⊩ϕx)))}))
       in
       eq/ _ _ (answer , F≤G→G≤F perϕ perX (composeFuncRel _ _ _ funcRel incψ) incϕ answer)
diff --git a/docs/Realizability.Topos.SubobjectClassifier.html b/docs/Realizability.Topos.SubobjectClassifier.html
index 4d86d16..3d028fc 100644
--- a/docs/Realizability.Topos.SubobjectClassifier.html
+++ b/docs/Realizability.Topos.SubobjectClassifier.html
@@ -1,5 +1,5 @@
 
-Realizability.Topos.SubobjectClassifier
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; ⟦_⟧ to pre⟦_⟧)
+Realizability.Topos.SubobjectClassifier
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; ⟦_⟧ to pre⟦_⟧)
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.HLevels
 open import Cubical.Foundations.Equiv
@@ -21,8 +21,8 @@
   {}
   {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
-  (resizing : hPropResizing )
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  )
+  (resizing : hPropResizing )
   where
   
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = } {ℓ'' = } ca
@@ -49,8 +49,8 @@
 Ωper : PartialEquivalenceRelation (ResizedPredicate Unit*)
 Predicate.isSetX (equality Ωper) = isSet× isSetResizedPredicate isSetResizedPredicate
 Predicate.∣ equality Ωper  (α , β) r =
-  (∀ (a : A) (⊩α : a   toPredicate α  tt*)  ((pr₁  r)  a)   toPredicate β  tt*) ×
-  (∀ (a : A) (⊩β : a   toPredicate β  tt*)  ((pr₂  r)  a)   toPredicate α  tt*)
+  (∀ (a : A) (⊩α : a   toPredicate α  tt*)  ((pr₁  r)  a)   toPredicate β  tt*) ×
+  (∀ (a : A) (⊩β : a   toPredicate β  tt*)  ((pr₂  r)  a)   toPredicate α  tt*)
 Predicate.isPropValued (equality Ωper) (α , β) r =
   isProp×
     (isPropΠ  _  isPropΠ λ _  (toPredicate β) .isPropValued _ _))
@@ -60,41 +60,41 @@
   do
     let
       ent₁ : ApplStrTerm as 2
-      ent₁ = ` pr₂ ̇ # one ̇ # zero
+      ent₁ = ` pr₂ ̇ # one ̇ # zero
 
       ent₂ : ApplStrTerm as 2
-      ent₂ = ` pr₁ ̇ # one ̇ # zero
+      ent₂ = ` pr₁ ̇ # one ̇ # zero
 
       realizer : ApplStrTerm as 1
-      realizer = ` pair ̇ (λ*abst ent₁) ̇ (λ*abst ent₂)
+      realizer = ` pair ̇ (λ*abst ent₁) ̇ (λ*abst ent₂)
     return
-      (λ* realizer ,
+      (λ* realizer ,
       λ { α β r (pr₁r⊩α≤β , pr₂r⊩β≤α) 
          a a⊩β 
           let
-            eq : pr₁  (λ* realizer  r)  a  pr₂  r  a
+            eq : pr₁  (λ* realizer  r)  a  pr₂  r  a
             eq =
-              pr₁  (λ* realizer  r)  a
-                ≡⟨ cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
-              pr₁  (pair  _  _)  a
-                ≡⟨ cong  x  x  a) (pr₁pxy≡x _ _) 
-               (λ*abst ent₁)  (r  [])  a
-                ≡⟨ βreduction ent₁ a (r  []) 
-              pr₂  r  a
+              pr₁  (λ* realizer  r)  a
+                ≡⟨ cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
+              pr₁  (pair  _  _)  a
+                ≡⟨ cong  x  x  a) (pr₁pxy≡x _ _) 
+               (λ*abst ent₁)  (r  [])  a
+                ≡⟨ βreduction ent₁ a (r  []) 
+              pr₂  r  a
                 
           in
           subst  r'  r'   toPredicate α  tt*) (sym eq) (pr₂r⊩β≤α a a⊩β)) ,
          a a⊩α 
           let
-            eq : pr₂  (λ* realizer  r)  a  pr₁  r  a
+            eq : pr₂  (λ* realizer  r)  a  pr₁  r  a
             eq =
-              pr₂  (λ* realizer  r)  a
-                ≡⟨ cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
-              pr₂  (pair  _  _)  a
-                ≡⟨ cong  x  x  a) (pr₂pxy≡y _ _) 
-               λ*abst ent₂  (r  [])  a
-                ≡⟨ βreduction ent₂ a (r  []) 
-              pr₁  r  a
+              pr₂  (λ* realizer  r)  a
+                ≡⟨ cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+              pr₂  (pair  _  _)  a
+                ≡⟨ cong  x  x  a) (pr₂pxy≡y _ _) 
+               λ*abst ent₂  (r  [])  a
+                ≡⟨ βreduction ent₂ a (r  []) 
+              pr₁  r  a
                 
           in
           subst  r'  r'   toPredicate β  tt*) (sym eq) (pr₁r⊩α≤β a a⊩α)) })
@@ -102,115 +102,115 @@
   do
     let
       closure1 : ApplStrTerm as 3
-      closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ # two ̇ # zero)
+      closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ # two ̇ # zero)
 
       closure2 : ApplStrTerm as 3
-      closure2 = ` pr₂ ̇ # two ̇ (` pr₂ ̇ # one ̇ # zero)
+      closure2 = ` pr₂ ̇ # two ̇ (` pr₂ ̇ # one ̇ # zero)
 
-      realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2)
+      realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2)
     return
-      (λ*2 realizer ,
+      (λ*2 realizer ,
        { x y z a b (⊩x≤y , ⊩y≤x) (⊩y≤z , ⊩z≤y) 
          r r⊩x 
           subst
              r'  r'   toPredicate z  tt*)
             (sym
-              (cong  x  pr₁  x  r) (λ*2ComputationRule realizer a b) 
-               cong  x  x  r) (pr₁pxy≡x _ _) 
-               βreduction closure1 r (b  a  [])))
+              (cong  x  pr₁  x  r) (λ*2ComputationRule realizer a b) 
+               cong  x  x  r) (pr₁pxy≡x _ _) 
+               βreduction closure1 r (b  a  [])))
             (⊩y≤z _ (⊩x≤y r r⊩x))) ,
          r r⊩z 
           subst
              r'  r'   toPredicate x  tt*)
             (sym
-              (cong  x  pr₂  x  r) (λ*2ComputationRule realizer a b) 
-               cong  x  x  r) (pr₂pxy≡y _ _) 
-               βreduction closure2 r (b  a  [])))
+              (cong  x  pr₂  x  r) (λ*2ComputationRule realizer a b) 
+               cong  x  x  r) (pr₂pxy≡y _ _) 
+               βreduction closure2 r (b  a  [])))
             (⊩y≤x _ (⊩z≤y r r⊩z))) }))
 
 opaque
   unfolding terminalPer
   trueFuncRel : FunctionalRelation terminalPer Ωper
   Predicate.isSetX (relation trueFuncRel) = isSet× isSetUnit* isSetResizedPredicate
-  Predicate.∣ relation trueFuncRel  (tt* , p) r =  (a : A)  (r  a)   toPredicate p  tt*
+  Predicate.∣ relation trueFuncRel  (tt* , p) r =  (a : A)  (r  a)   toPredicate p  tt*
   Predicate.isPropValued (relation trueFuncRel) (tt* , p) r = isPropΠ λ a  (toPredicate p) .isPropValued _ _
   isFunctionalRelation.isStrictDomain (isFuncRel trueFuncRel) =
     do
       return
-        (k ,
+        (k ,
          { tt* y r r⊩⊤≤y  tt*}))
   isFunctionalRelation.isStrictCodomain (isFuncRel trueFuncRel) =
     do
       let
         idClosure : ApplStrTerm as 2
-        idClosure = # zero
+        idClosure = # zero
         realizer : ApplStrTerm as 1
-        realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure)
+        realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure)
       return
-        (λ* realizer ,
+        (λ* realizer ,
          { tt* y r r⊩⊤≤y 
            a a⊩y 
             subst
                r'  r'   toPredicate y  tt*)
               (sym
-                (cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
-                 cong  x  x  a) (pr₁pxy≡x _ _) 
-                 βreduction idClosure a (r  [])))
+                (cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₁pxy≡x _ _) 
+                 βreduction idClosure a (r  [])))
               a⊩y) ,
            a a⊩y 
             subst
                r'  r'   toPredicate y  tt*)
               (sym
-                (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
-                 cong  x  x  a) (pr₂pxy≡y _ _) 
-                 βreduction idClosure a (r  [])))
+                (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₂pxy≡y _ _) 
+                 βreduction idClosure a (r  [])))
               a⊩y)}))
   isFunctionalRelation.isRelational (isFuncRel trueFuncRel) =
     do
       let
         realizer : ApplStrTerm as 4
-        realizer = ` pr₁ ̇ # one ̇ (# two  ̇ ` k)
+        realizer = ` pr₁ ̇ # one ̇ (# two  ̇ ` k)
       return
-        (λ*4 realizer ,
+        (λ*4 realizer ,
          { tt* tt* x y a b c tt* b⊩⊤≤x (pr₁c⊩x≤y , pr₂c⊩y≤x) r 
           subst
              r'  r'   toPredicate y  tt*)
-            (sym (λ*4ComputationRule realizer a b c r))
-            (pr₁c⊩x≤y (b  k) (b⊩⊤≤x k))}))
+            (sym (λ*4ComputationRule realizer a b c r))
+            (pr₁c⊩x≤y (b  k) (b⊩⊤≤x k))}))
   isFunctionalRelation.isSingleValued (isFuncRel trueFuncRel) =
     do
       let
         closure1 : ApplStrTerm as 3
-        closure1 = # one ̇ ` k
+        closure1 = # one ̇ ` k
 
         closure2 : ApplStrTerm as 3
-        closure2 = # two ̇ ` k
+        closure2 = # two ̇ ` k
         
         realizer : ApplStrTerm as 2
-        realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2)
+        realizer = ` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2)
       return
-        (λ*2 realizer ,
+        (λ*2 realizer ,
          { tt* x y r₁ r₂ r₁⊩⊤≤x r₂⊩⊤≤y 
            a a⊩x 
             subst
                r'  r'   toPredicate y  tt*)
               (sym
-                (cong  x  pr₁  x  a) (λ*2ComputationRule realizer r₁ r₂) 
-                 cong  x  x  a) (pr₁pxy≡x _ _) 
-                 βreduction closure1 a (r₂  r₁  [])))
-              (r₂⊩⊤≤y k)) ,
+                (cong  x  pr₁  x  a) (λ*2ComputationRule realizer r₁ r₂) 
+                 cong  x  x  a) (pr₁pxy≡x _ _) 
+                 βreduction closure1 a (r₂  r₁  [])))
+              (r₂⊩⊤≤y k)) ,
            a a⊩y 
             subst
                r'  r'   toPredicate x  tt*)
               (sym
-                (cong  x  pr₂  x  a) (λ*2ComputationRule realizer r₁ r₂) 
-                 cong  x  x  a) (pr₂pxy≡y _ _) 
-                 βreduction closure2 a (r₂  r₁  [])))
-              (r₁⊩⊤≤x k))}))
+                (cong  x  pr₂  x  a) (λ*2ComputationRule realizer r₁ r₂) 
+                 cong  x  x  a) (pr₂pxy≡y _ _) 
+                 βreduction closure2 a (r₂  r₁  [])))
+              (r₁⊩⊤≤x k))}))
   isFunctionalRelation.isTotal (isFuncRel trueFuncRel) =
     do
       return
-        (k ,
+        (k ,
          { tt* r tt* 
           let
              = pre1 Unit* isSetUnit* isNonTrivial
@@ -218,7 +218,7 @@
           
             fromPredicate  ,
              a 
-              subst  p  (k  r  a)   p  tt*) (sym (compIsIdFunc )) tt*)
+              subst  p  (k  r  a)   p  tt*) (sym (compIsIdFunc )) tt*)
           ∣₁ }))
 
 opaque
@@ -228,7 +228,7 @@
   isInjectiveTrueFuncRel =
     do
       return
-        (k ,
+        (k ,
          { tt* tt* p r₁ r₂ r₁⊩⊤≤p r₂⊩⊤≤p  tt* }))
 
 truePredicate : Predicate Unit*
@@ -253,15 +253,15 @@
   charFuncRel : FunctionalRelation perX Ωper
   Predicate.isSetX (relation charFuncRel) = isSet× (perX .isSetX) isSetResizedPredicate
   Predicate.∣ relation charFuncRel  (x , p) r =
-    (pr₁  r)   perX .equality  (x , x) ×
-    (∀ (b : A) (b⊩ϕx : b   ϕ .predicate  x)  (pr₁  (pr₂  r)  b)   toPredicate p  tt*) ×
-    (∀ (b : A) (b⊩px : b   toPredicate p  tt*)  (pr₂  (pr₂  r)  b)   ϕ .predicate  x)
+    (pr₁  r)   perX .equality  (x , x) ×
+    (∀ (b : A) (b⊩ϕx : b   ϕ .predicate  x)  (pr₁  (pr₂  r)  b)   toPredicate p  tt*) ×
+    (∀ (b : A) (b⊩px : b   toPredicate p  tt*)  (pr₂  (pr₂  r)  b)   ϕ .predicate  x)
   Predicate.isPropValued (relation charFuncRel) (x , p) r =
     isProp×
       (perX .equality .isPropValued _ _)
       (isProp×
-        (isPropΠ  _  isPropΠ λ _  (toPredicate p) .isPropValued _ _))
-        (isPropΠ λ _  isPropΠ λ _  ϕ .predicate .isPropValued _ _))
+        (isPropΠ  _  isPropΠ λ _  (toPredicate p) .isPropValued _ _))
+        (isPropΠ λ _  isPropΠ λ _  ϕ .predicate .isPropValued _ _))
   isFunctionalRelation.isStrictDomain (isFuncRel charFuncRel) =
     do
       return
@@ -271,27 +271,27 @@
     do
       let
         idClosure : ApplStrTerm as 2
-        idClosure = # zero
+        idClosure = # zero
         realizer : ApplStrTerm as 1
-        realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure)
+        realizer = ` pair ̇ (λ*abst idClosure) ̇ (λ*abst idClosure)
       return
-        (λ* realizer ,
+        (λ* realizer ,
          x y r x₁ 
            a a⊩y 
             subst
                r'  r'   toPredicate y  tt*)
               (sym
-                (cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
-                 cong  x  x  a) (pr₁pxy≡x _ _) 
-                 βreduction idClosure a (r  [])))
+                (cong  x  pr₁  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₁pxy≡x _ _) 
+                 βreduction idClosure a (r  [])))
               a⊩y) ,
            a a⊩y 
             subst
                r'  r'   toPredicate y  tt*)
               (sym
-                (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
-                 cong  x  x  a) (pr₂pxy≡y _ _) 
-                 βreduction idClosure a (r  [])))
+                (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+                 cong  x  x  a) (pr₂pxy≡y _ _) 
+                 βreduction idClosure a (r  [])))
               a⊩y)))
   isFunctionalRelation.isRelational (isFuncRel charFuncRel) =
     do
@@ -300,15 +300,15 @@
       (relϕ , relϕ⊩isRelationalϕ)  isStrictRelation.isRelational (ϕ .isStrictRelationPredicate)
       let
         closure1 : ApplStrTerm as 4
-        closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` relϕ ̇ # zero ̇ (` sX ̇ # three)))
+        closure1 = ` pr₁ ̇ # one ̇ (` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` relϕ ̇ # zero ̇ (` sX ̇ # three)))
 
         closure2 : ApplStrTerm as 4
-        closure2 = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ # one ̇ # zero)) ̇ # three
+        closure2 = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ # one ̇ # zero)) ̇ # three
 
         realizer : ApplStrTerm as 3
-        realizer = ` pair ̇ (` tX ̇ (` sX ̇ # two) ̇ # two) ̇ (` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2))
+        realizer = ` pair ̇ (` tX ̇ (` sX ̇ # two) ̇ # two) ̇ (` pair ̇ (λ*abst closure1) ̇ (λ*abst closure2))
       return
-        (λ*3 realizer ,
+        (λ*3 realizer ,
          { x x' p p' a b c a⊩x~x' (⊩x~x , ⊩ϕx≤p , ⊩p≤ϕx) (⊩p≤p' , ⊩p'≤p) 
           let
             ⊩x'~x = sX⊩isSymmetricX x x' a a⊩x~x'
@@ -316,60 +316,60 @@
           in
           subst
              r'  r'   perX .equality  (x' , x'))
-            (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
+            (sym (cong  x  pr₁  x) (λ*3ComputationRule realizer a b c)  pr₁pxy≡x _ _))
             ⊩x'~x' ,
            r r⊩ϕx' 
             subst
                r'  r'   toPredicate p'  tt*)
               (sym
-                (cong  x  pr₁  (pr₂  x)  r) (λ*3ComputationRule realizer a b c) 
-                 cong  x  pr₁  x  r) (pr₂pxy≡y _ _) 
-                 cong  x  x  r) (pr₁pxy≡x _ _) 
-                 βreduction closure1 r (c  b  a  [])))
+                (cong  x  pr₁  (pr₂  x)  r) (λ*3ComputationRule realizer a b c) 
+                 cong  x  pr₁  x  r) (pr₂pxy≡y _ _) 
+                 cong  x  x  r) (pr₁pxy≡x _ _) 
+                 βreduction closure1 r (c  b  a  [])))
               (⊩p≤p' _ (⊩ϕx≤p _ (relϕ⊩isRelationalϕ x' x _ _ r⊩ϕx' ⊩x'~x)))) ,
           λ r r⊩p' 
             subst
                r'  r'   ϕ .predicate  x')
               (sym
-                (cong  x  pr₂  (pr₂  x)  r) (λ*3ComputationRule realizer a b c) 
-                 cong  x  pr₂  x  r) (pr₂pxy≡y _ _) 
-                 cong  x  x  r) (pr₂pxy≡y _ _) 
-                 βreduction closure2 r (c  b  a  [])))
+                (cong  x  pr₂  (pr₂  x)  r) (λ*3ComputationRule realizer a b c) 
+                 cong  x  pr₂  x  r) (pr₂pxy≡y _ _) 
+                 cong  x  x  r) (pr₂pxy≡y _ _) 
+                 βreduction closure2 r (c  b  a  [])))
               (relϕ⊩isRelationalϕ x x' _ _ (⊩p≤ϕx _ (⊩p'≤p r r⊩p')) a⊩x~x') }))
   isFunctionalRelation.isSingleValued (isFuncRel charFuncRel) =
     do
       let
         closure1 : ApplStrTerm as 3
-        closure1 = ` pr₁ ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ # zero)
+        closure1 = ` pr₁ ̇ (` pr₂ ̇ # one) ̇ (` pr₂ ̇ (` pr₂ ̇ # two) ̇ # zero)
 
         closure2 : ApplStrTerm as 3
-        closure2 = ` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ (` pr₂ ̇ # one) ̇ # zero)
+        closure2 = ` pr₁ ̇ (` pr₂ ̇ # two) ̇ (` pr₂ ̇ (` pr₂ ̇ # one) ̇ # zero)
 
         realizer : ApplStrTerm as 2
-        realizer = ` pair ̇ λ*abst closure1 ̇ λ*abst closure2
+        realizer = ` pair ̇ λ*abst closure1 ̇ λ*abst closure2
       return
-        (λ*2 realizer ,
+        (λ*2 realizer ,
          { x y y' r₁ r₂ (⊩x~x , ⊩ϕx≤y , ⊩y≤ϕx) (⊩'x~x , ⊩ϕx≤y' , ⊩y'≤ϕx) 
            a a⊩y 
             subst
                r'  r'   toPredicate y'  tt*)
-              (sym (cong  x  pr₁  x  a) (λ*2ComputationRule realizer r₁ r₂)  cong  x  x  a) (pr₁pxy≡x _ _)  βreduction closure1 a (r₂  r₁  [])))
+              (sym (cong  x  pr₁  x  a) (λ*2ComputationRule realizer r₁ r₂)  cong  x  x  a) (pr₁pxy≡x _ _)  βreduction closure1 a (r₂  r₁  [])))
               (⊩ϕx≤y' _ (⊩y≤ϕx a a⊩y))) ,
            a a⊩y' 
             subst
                r'  r'   toPredicate y  tt*)
-              (sym (cong  x  pr₂  x  a) (λ*2ComputationRule realizer r₁ r₂)  cong  x  x  a) (pr₂pxy≡y _ _)  βreduction closure2 a (r₂  r₁  [])))
+              (sym (cong  x  pr₂  x  a) (λ*2ComputationRule realizer r₁ r₂)  cong  x  x  a) (pr₂pxy≡y _ _)  βreduction closure2 a (r₂  r₁  [])))
               (⊩ϕx≤y _ (⊩y'≤ϕx a a⊩y'))) }))
   isFunctionalRelation.isTotal (isFuncRel charFuncRel) =
     do
       let
         idClosure : ApplStrTerm as 2
-        idClosure = # zero
+        idClosure = # zero
 
         realizer : ApplStrTerm as 1
-        realizer = ` pair ̇ # zero ̇ (` pair ̇ λ*abst idClosure ̇ λ*abst idClosure)
+        realizer = ` pair ̇ # zero ̇ (` pair ̇ λ*abst idClosure ̇ λ*abst idClosure)
       return
-        (λ* realizer ,
+        (λ* realizer ,
          x r r⊩x~x 
           let
             resultPredicate : Predicate Unit*
@@ -383,25 +383,25 @@
             (fromPredicate resultPredicate ,
             subst
                r'  r'   perX .equality  (x , x))
-              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+              (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
               r⊩x~x ,
              b b⊩ϕx 
               subst
                  r  r   toPredicate (fromPredicate resultPredicate)  tt*)
                 (sym
-                  (cong  x  pr₁  (pr₂  x)  b) (λ*ComputationRule realizer r) 
-                   cong  x  pr₁  x  b) (pr₂pxy≡y _ _) 
-                   cong  x  x  b) (pr₁pxy≡x _ _) 
-                   βreduction idClosure b (r  [])))
+                  (cong  x  pr₁  (pr₂  x)  b) (λ*ComputationRule realizer r) 
+                   cong  x  pr₁  x  b) (pr₂pxy≡y _ _) 
+                   cong  x  x  b) (pr₁pxy≡x _ _) 
+                   βreduction idClosure b (r  [])))
                 (subst  p  b   p  tt*) (sym (compIsIdFunc resultPredicate)) b⊩ϕx)) ,
              b b⊩'ϕx 
               subst
                  r  r   ϕ .predicate  x)
                 (sym
-                  (cong  x  pr₂  (pr₂  x)  b) (λ*ComputationRule realizer r) 
-                   cong  x  pr₂  x  b) (pr₂pxy≡y _ _) 
-                   cong  x  x  b) (pr₂pxy≡y _ _) 
-                   βreduction idClosure b (r  [])))
+                  (cong  x  pr₂  (pr₂  x)  b) (λ*ComputationRule realizer r) 
+                   cong  x  pr₂  x  b) (pr₂pxy≡y _ _) 
+                   cong  x  x  b) (pr₂pxy≡y _ _) 
+                   βreduction idClosure b (r  [])))
                 let foo = subst  p  b   p  tt*) (compIsIdFunc resultPredicate) b⊩'ϕx in foo))))
 
   subobjectCospan :  char  Cospan RT
@@ -426,14 +426,14 @@
             (relϕ , relϕ⊩isRelationalϕ)  StrictRelation.isRelational ϕ
             let
               closure : ApplStrTerm as 2
-              closure = (` pr₁ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` relϕ ̇ (` pr₂ ̇ (` pr₁ ̇ # one)) ̇ (` pr₁ ̇ (` pr₁ ̇ # one))))
+              closure = (` pr₁ ̇ (` pr₂ ̇ (` pr₂ ̇ # one)) ̇ (` relϕ ̇ (` pr₂ ̇ (` pr₁ ̇ # one)) ̇ (` pr₁ ̇ (` pr₁ ̇ # one))))
               realizer : ApplStrTerm as 1
               realizer =
-                ` pair ̇
-                  (` pair ̇ (` stX ̇ (` pr₁ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
-                  λ*abst closure
+                ` pair ̇
+                  (` pair ̇ (` stX ̇ (` pr₁ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇
+                  λ*abst closure
             return
-              (λ* realizer ,
+              (λ* realizer ,
                { x p r r⊩∃x' 
                 do
                   (x' , (⊩x~x' , ⊩ϕx) , ⊩x'~x' , ⊩ϕx'≤p , ⊩p≤ϕx')  r⊩∃x'
@@ -441,19 +441,19 @@
                     (tt* ,
                     ((subst
                        r'  r'   perX .equality  (x , x))
-                      (sym (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _))
+                      (sym (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _))
                       (stX⊩isStrictDomainX x x' _ ⊩x~x')) ,
                      (subst
                         r'  r'   ϕ .predicate  x)
-                       (sym (cong  x  pr₂  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _))
+                       (sym (cong  x  pr₂  (pr₁  x)) (λ*ComputationRule realizer r)  cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _))
                        ⊩ϕx)) ,
                     λ r' 
                       let
-                        eq : pr₂  (λ* realizer  r)  r'  pr₁  (pr₂  (pr₂  r))  (relϕ  (pr₂  (pr₁  r))  (pr₁  (pr₁  r)))
+                        eq : pr₂  (λ* realizer  r)  r'  pr₁  (pr₂  (pr₂  r))  (relϕ  (pr₂  (pr₁  r))  (pr₁  (pr₁  r)))
                         eq =
-                          cong  x  pr₂  x  r') (λ*ComputationRule realizer r) 
-                          cong  x  x  r') (pr₂pxy≡y _ _) 
-                          βreduction closure r' (r  [])
+                          cong  x  pr₂  x  r') (λ*ComputationRule realizer r) 
+                          cong  x  x  r') (pr₂pxy≡y _ _) 
+                          βreduction closure r' (r  [])
                       in
                       subst
                          r'  r'   toPredicate p  tt*)
@@ -487,7 +487,7 @@
     opaque
       unfolding trueFuncRel
       trueFuncRelTruePredicate :  a  (a   trueFuncRel .relation  (tt* , fromPredicate truePredicate))
-      trueFuncRelTruePredicate a = λ b  subst  p  (a  b)   p  tt*) (sym (compIsIdFunc truePredicate)) tt*
+      trueFuncRelTruePredicate a = λ b  subst  p  (a  b)   p  tt*) (sym (compIsIdFunc truePredicate)) tt*
 
     opaque
       unfolding composeFuncRel
@@ -514,17 +514,17 @@
           let
             realizer : ApplStrTerm as 1
             realizer =
-              ` pair ̇
-                (` stFC ̇ # zero) ̇
-                (` relϕ ̇
-                  (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k)))) ̇ ` k) ̇
-                  (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k))) ̇ # zero))
+              ` pair ̇
+                (` stFC ̇ # zero) ̇
+                (` relϕ ̇
+                  (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k)))) ̇ ` k) ̇
+                  (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # zero)) ̇ ` k))) ̇ # zero))
           return
-            (λ* realizer ,
+            (λ* realizer ,
               y x r r⊩Hyx 
                subst
                   r'  r'   perX .equality  (x , x))
-                 (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
+                 (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x _ _))
                  (stFC⊩isStrictCodomainF y x _ r⊩Hyx) ,
                (equivFun
                  (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _))
@@ -533,30 +533,30 @@
                      ent⊩entailment
                      y
                      (fromPredicate truePredicate)
-                     (pair  (a  (stFD  r))  k)
+                     (pair  (a  (stFD  r))  k)
                      (return
                        (tt* ,
                         subst
                            r  r   G .relation  (y , tt*))
                           (sym (pr₁pxy≡x _ _))
-                          (a⊩idY≤G y tt* (stFD  r) (stFD⊩isStrictDomainF y x _ r⊩Hyx))  ,
+                          (a⊩idY≤G y tt* (stFD  r) (stFD⊩isStrictDomainF y x _ r⊩Hyx))  ,
                         trueFuncRelTruePredicate _))
                    let
                      ⊩x'~x = svF⊩isSingleValuedF y x' x _ _ ⊩Fyx' r⊩Hyx
-                     ⊩ϕx = relϕ⊩isRelationalϕ x' x _ _ (⊩⊤≤ϕx' k (subst  p  k   p  tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x'~x
-                   return (subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx)))))
+                     ⊩ϕx = relϕ⊩isRelationalϕ x' x _ _ (⊩⊤≤ϕx' k (subst  p  k   p  tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x'~x
+                   return (subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (λ*ComputationRule realizer r)  pr₂pxy≡y _ _)) ⊩ϕx)))))
       isFunctionalRelation.isRelational (isFuncRel H) =
         do
           (relF , relF⊩isRelationalF)  isFunctionalRelation.isRelational (F .isFuncRel)
           let
             realizer : ApplStrTerm as 3
-            realizer = ` relF ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
+            realizer = ` relF ̇ # two ̇ # one ̇ (` pr₁ ̇ # zero)
           return
-            (λ*3 realizer ,
+            (λ*3 realizer ,
               y y' x x' a b c ⊩y~y' ⊩Fyx (⊩x~x' , ⊩ϕx) 
                subst
                   r'  r'   F .relation  (y' , x'))
-                 (sym (λ*3ComputationRule realizer a b c))
+                 (sym (λ*3ComputationRule realizer a b c))
                  (relF⊩isRelationalF y y' x x' _ _ _ ⊩y~y' ⊩Fyx ⊩x~x')))
       isFunctionalRelation.isSingleValued (isFuncRel H) =
         do
@@ -568,15 +568,15 @@
           let
             realizer : ApplStrTerm as 2
             realizer =
-              ` pair ̇
-                (` svF ̇ # one ̇ # zero) ̇
-                (` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD  ̇ # one)) ̇ ` k)))) ̇ ` k) ̇ (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # one)) ̇ ` k))) ̇ # one))
+              ` pair ̇
+                (` svF ̇ # one ̇ # zero) ̇
+                (` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` pr₂ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD  ̇ # one)) ̇ ` k)))) ̇ ` k) ̇ (` svF ̇ (` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ (` stFD ̇ # one)) ̇ ` k))) ̇ # one))
           return
-            (λ*2 realizer ,
+            (λ*2 realizer ,
               y x x' r₁ r₂ ⊩Fyx ⊩Fyx' 
                subst
                   r'  r'   perX .equality  (x , x'))
-                 (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
+                 (sym (cong  x  pr₁  x) (λ*2ComputationRule realizer r₁ r₂)  pr₁pxy≡x _ _))
                  (svF⊩isSingleValuedF y x x' _ _ ⊩Fyx ⊩Fyx') ,
                (equivFun
                  (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _))
@@ -585,18 +585,18 @@
                      ent⊩entailment
                      y
                      (fromPredicate truePredicate)
-                     (pair  (a  (stFD  r₁))  k)
+                     (pair  (a  (stFD  r₁))  k)
                      (return
                        (tt* ,
                         subst  r  r   G .relation  (y , tt*)) (sym (pr₁pxy≡x _ _)) (a⊩idY≤G y tt* _ (stFD⊩isStrictDomainF y x _ ⊩Fyx))  ,
                         trueFuncRelTruePredicate _))
                    let
                      ⊩x''~x = svF⊩isSingleValuedF y x'' x _ _ ⊩Fyx'' ⊩Fyx
-                     ⊩ϕx = relϕ⊩isRelationalϕ x'' x _ _ (⊩⊤≤ϕx'' k (subst  p  k   p  tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x''~x
+                     ⊩ϕx = relϕ⊩isRelationalϕ x'' x _ _ (⊩⊤≤ϕx'' k (subst  p  k   p  tt*) (sym (compIsIdFunc truePredicate)) tt*)) ⊩x''~x
                    return
                      (subst
                         r'  r'   ϕ .predicate  x)
-                       (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _))
+                       (sym (cong  x  pr₂  x) (λ*2ComputationRule realizer r₁ r₂)  pr₂pxy≡y _ _))
                        ⊩ϕx)))))
       isFunctionalRelation.isTotal (isFuncRel H) =
         do
@@ -604,21 +604,21 @@
           (a , a⊩idY≤G)  idY≤G
           let
             realizer : ApplStrTerm as 1
-            realizer = ` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ # zero) ̇ ` k))
+            realizer = ` pr₁ ̇ (` ent ̇ (` pair ̇ (` a ̇ # zero) ̇ ` k))
           return
-            (λ* realizer ,
+            (λ* realizer ,
              { y r r⊩y~y 
               do
                 (x , ⊩Fyx , ⊩x~x , ⊩ϕx≤⊤ , ⊩⊤≤ϕx) 
                   ent⊩entailment
                     y
                     (fromPredicate truePredicate)
-                    (pair  (a  r)  k)
+                    (pair  (a  r)  k)
                     (return
                       (tt* ,
                        subst  r  r   G .relation  (y , tt*)) (sym (pr₁pxy≡x _ _)) (a⊩idY≤G y tt* r r⊩y~y)  ,
                        trueFuncRelTruePredicate _))
-                return (x , subst  r'  r'   F .relation  (y , x)) (sym (λ*ComputationRule realizer r)) ⊩Fyx) }))
+                return (x , subst  r'  r'   F .relation  (y , x)) (sym (λ*ComputationRule realizer r)) ⊩Fyx) }))
 
     opaque
       unfolding composeRTMorphism
@@ -633,9 +633,9 @@
               (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
               let
                 realizer : ApplStrTerm as 1
-                realizer = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
+                realizer = ` relF ̇ (` stFD ̇ (` pr₁ ̇ # zero)) ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))
               return
-                 (λ* realizer ,
+                 (λ* realizer ,
                   y x r ⊩∃x' 
                    equivFun
                      (propTruncIdempotent≃ (F .relation .isPropValued _ _))
@@ -644,7 +644,7 @@
                        return
                          (subst
                             r'  r'   F .relation  (y , x))
-                           (sym (λ*ComputationRule realizer r))
+                           (sym (λ*ComputationRule realizer r))
                            (relF⊩isRelationalF y y x' x _ _ _ (stFD⊩isStrictDomainF y x' _ ⊩Hyx') ⊩Hyx' ⊩x'~x)))))
         in eq/ _ _ (F≤G→G≤F perY perX (composeFuncRel _ _ _ H incFuncRel) F answer , answer)
 
@@ -660,15 +660,15 @@
               (a , a⊩idY≤G)  idY≤G
               let
                 realizer : ApplStrTerm as 1
-                realizer = ` a ̇ (` stHD ̇ (` pr₁ ̇ # zero))
+                realizer = ` a ̇ (` stHD ̇ (` pr₁ ̇ # zero))
               return
-                (λ* realizer ,
+                (λ* realizer ,
                   { y tt* r r⊩∃x 
                    equivFun
                      (propTruncIdempotent≃ (G .relation .isPropValued _ _))
                      (do
                        (x , ⊩Hyx , ⊩x~x , ⊩ϕx)  r⊩∃x
-                       return (subst  r'  r'   G .relation  (y , tt*)) (sym (λ*ComputationRule realizer r)) (a⊩idY≤G y tt* _ (stHD⊩isStrictDomainH y x _ ⊩Hyx)))) }))
+                       return (subst  r'  r'   G .relation  (y , tt*)) (sym (λ*ComputationRule realizer r)) (a⊩idY≤G y tt* _ (stHD⊩isStrictDomainH y x _ ⊩Hyx)))) }))
         in eq/ _ _ (F≤G→G≤F perY terminalPer (composeFuncRel _ _ _ H (terminalFuncRel subPer)) G answer , answer)
 
     opaque
@@ -687,9 +687,9 @@
               (stDH , stDH⊩isStrictDomainH)  H .isStrictDomain
               let
                 realizer : ApplStrTerm as 1
-                realizer = ` relH' ̇ (` stDH ̇ # zero) ̇ (` pr₁ ̇ (` a ̇ # zero)) ̇ (` pr₂ ̇ (` a ̇ # zero))
+                realizer = ` relH' ̇ (` stDH ̇ # zero) ̇ (` pr₁ ̇ (` a ̇ # zero)) ̇ (` pr₂ ̇ (` a ̇ # zero))
               return
-                (λ* realizer ,
+                (λ* realizer ,
                   y x r r⊩Hyx 
                    equivFun
                      (propTruncIdempotent≃ (H' .relation .isPropValued _ _))
@@ -698,7 +698,7 @@
                        return
                          (subst
                             r'  r'   H' .relation  (y , x))
-                           (sym (λ*ComputationRule realizer r))
+                           (sym (λ*ComputationRule realizer r))
                            (relH'⊩isRelationalH y y x' x _ _ _ (stDH⊩isStrictDomainH y x r r⊩Hyx) ⊩H'yx' (⊩x'~x , ⊩ϕx'))))))
         in
         eq/ _ _ (answer , (F≤G→G≤F _ _ H H' answer))
@@ -708,7 +708,7 @@
     classifies {Y , perY} f g f⋆char≡g⋆true =
       SQ.elimProp2
         {P = λ f g   (commutes : f  [ charFuncRel ]  g  [ trueFuncRel ])  ∃![ hk  RTMorphism perY subPer ] (f  hk  [ incFuncRel ]) × (g  hk  [ terminalFuncRel subPer ])}
-         f g  isPropΠ λ _  isPropIsContr)
+         f g  isPropΠ λ _  isPropIsContr)
          F G F⋆char≡G⋆true 
            let
              entailment = [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G F charFuncRel G trueFuncRel F⋆char≡G⋆true
@@ -722,7 +722,7 @@
              λ { h' (f≡h'⋆inc , g≡h'⋆term) 
                SQ.elimProp
                  {P = λ h'   (comm1 : [ F ]  h'  [ incFuncRel ]) (comm2 : [ G ]  h'  [ terminalFuncRel subPer ])  [ UnivPropWithRepr.H perY F G entailment ]  h'}
-                  h'  isPropΠ λ _  isPropΠ λ _  squash/ _ _)
+                  h'  isPropΠ λ _  isPropΠ λ _  squash/ _ _)
                   H' F≡H'⋆inc G≡H'⋆term 
                    UnivPropWithRepr.isUniqueH perY F G entailment H' F≡H'⋆inc G≡H'⋆term)
                  h'
@@ -753,11 +753,11 @@
         (stCC , stCC⊩isStrictCodomainC)  C .isStrictCodomain
         let
           realizer : ApplStrTerm as 2
-          realizer = ` relC ̇ # zero ̇ # one ̇ (` stCC ̇ # one)
+          realizer = ` relC ̇ # zero ̇ # one ̇ (` stCC ̇ # one)
         return
-          (λ*2 realizer ,
+          (λ*2 realizer ,
            λ x x' a b a⊩Cx⊤ b⊩x~x' 
-             subst  r'  r'   C .relation  (x' , )) (sym (λ*2ComputationRule realizer a b)) (relC⊩isRelationalC x x'   _ _ _ b⊩x~x' a⊩Cx⊤ (stCC⊩isStrictCodomainC x  a a⊩Cx⊤)))
+             subst  r'  r'   C .relation  (x' , )) (sym (λ*2ComputationRule realizer a b)) (relC⊩isRelationalC x x'   _ _ _ b⊩x~x' a⊩Cx⊤ (stCC⊩isStrictCodomainC x  a a⊩Cx⊤)))
 
     perψ = InducedSubobject.subPer perX ψ
     incFuncRelψ = InducedSubobject.incFuncRel perX ψ
@@ -779,12 +779,12 @@
               (sX , sX⊩isSymmetricX)  perX .isSymmetric
               let
                 closure : ApplStrTerm as 2
-                closure = ` pr₁ ̇ (` svC ̇ (` pr₂ ̇ (` pr₁ ̇ # one)) ̇ (` relC ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # one))) ̇ (` pr₂ ̇ # one) ̇ (` stCC ̇ (` pr₂ ̇ # one)))) ̇ ` k
+                closure = ` pr₁ ̇ (` svC ̇ (` pr₂ ̇ (` pr₁ ̇ # one)) ̇ (` relC ̇ (` sX ̇ (` pr₁ ̇ (` pr₁ ̇ # one))) ̇ (` pr₂ ̇ # one) ̇ (` stCC ̇ (` pr₂ ̇ # one)))) ̇ ` k
 
                 realizer : ApplStrTerm as 1
-                realizer = ` pair ̇ (` pair ̇ (` stDC ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (λ*abst closure)
+                realizer = ` pair ̇ (` pair ̇ (` stDC ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₁ ̇ # zero))) ̇ (λ*abst closure)
               return
-                (λ* realizer ,
+                (λ* realizer ,
                  λ { x p r r⊩∃x' 
                    do
                      (x' , (⊩x~x' , ⊩Cx⊤) , ⊩Cx'p)  r⊩∃x'
@@ -796,37 +796,37 @@
                        (subst
                           r'  r'   perX .equality  (x , x))
                          (sym
-                           (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r) 
-                            cong  x  pr₁  x) (pr₁pxy≡x _ _) 
+                           (cong  x  pr₁  (pr₁  x)) (λ*ComputationRule realizer r) 
+                            cong  x  pr₁  x) (pr₁pxy≡x _ _) 
                             pr₁pxy≡x _ _ ))
                          (stDC⊩isStrictDomainC x  _ ⊩Cx⊤) ,
                         subst
                            r'  r'   C .relation  (x , ))
                           (sym
-                            (cong  x  pr₂  (pr₁  x)) (λ*ComputationRule realizer r) 
-                             cong  x  pr₂  x) (pr₁pxy≡x _ _) 
+                            (cong  x  pr₂  (pr₁  x)) (λ*ComputationRule realizer r) 
+                             cong  x  pr₂  x) (pr₁pxy≡x _ _) 
                              pr₂pxy≡y _ _))
                           ⊩Cx⊤) ,
                         λ a 
                           subst
                              r'  r'   toPredicate p  tt*)
                             (sym
-                              (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
-                               cong  x  x  a) (pr₂pxy≡y _ _) 
-                               βreduction closure a (r  [])))
-                            (⊩⊤≤p k (subst  q  k   q  tt*) (sym (compIsIdFunc truePredicate)) tt*))) })
+                              (cong  x  pr₂  x  a) (λ*ComputationRule realizer r) 
+                               cong  x  x  a) (pr₂pxy≡y _ _) 
+                               βreduction closure a (r  [])))
+                            (⊩⊤≤p k (subst  q  k   q  tt*) (sym (compIsIdFunc truePredicate)) tt*))) })
         in eq/ _ _ (answer , F≤G→G≤F _ _ (composeFuncRel _ _ _ incFuncRelψ C) (composeFuncRel _ _ _ (terminalFuncRel perψ) trueFuncRel) answer)
 
     opaque
       unfolding InducedSubobject.incFuncRel
       unfolding composeFuncRel
-      ⊩Cx⊤≤ϕx : ∃[ ent  A ] (∀ (x : X) (r : A)  r   C .relation  (x , )  (ent  r)   ϕ .predicate  x)
+      ⊩Cx⊤≤ϕx : ∃[ ent  A ] (∀ (x : X) (r : A)  r   C .relation  (x , )  (ent  r)   ϕ .predicate  x)
       ⊩Cx⊤≤ϕx =
         let
           ((h , incψ≡h⋆incϕ , termψ≡h⋆termϕ) , isUniqueH) = classifies [ incFuncRelψ ] [ terminalFuncRel perψ ] pbSqCommutes
         in
         SQ.elimProp
-          {P = λ h   (incψ≡h⋆incϕ : [ incFuncRelψ ]  h  [ incFuncRel ])  ∃[ ent  A ] (∀ (x : X) (r : A)  r   C .relation  (x , )  (ent  r)   ϕ .predicate  x)}
+          {P = λ h   (incψ≡h⋆incϕ : [ incFuncRelψ ]  h  [ incFuncRel ])  ∃[ ent  A ] (∀ (x : X) (r : A)  r   C .relation  (x , )  (ent  r)   ϕ .predicate  x)}
            h  isPropΠ λ _  isPropPropTrunc)
            H incψ≡H⋆incϕ 
             do
@@ -834,9 +834,9 @@
               (stDC , stDC⊩isStrictDomainC)  C .isStrictDomain
               (relϕ , relϕ⊩isRelationalϕ)  isStrictRelation.isRelational (ϕ .isStrictRelationPredicate)
               let
-                realizer = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) 
+                realizer = ` relϕ ̇ (` pr₂ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) ̇ (` pr₁ ̇ (` pr₂ ̇ (` a ̇ (` pair ̇ (` stDC ̇ # zero) ̇ # zero)))) 
               return
-                (λ* realizer ,
+                (λ* realizer ,
                   x r r⊩Cx⊤ 
                    equivFun
                      (propTruncIdempotent≃ (ϕ .predicate .isPropValued _ _))
@@ -844,11 +844,11 @@
                        (x' , ⊩Hxx' , ⊩x'~x , ⊩ϕx') 
                            a⊩incψ≤H⋆incϕ
                            x x
-                           (pair  (stDC  r)  r)
+                           (pair  (stDC  r)  r)
                            (subst  r'  r'   perX .equality  (x , x)) (sym (pr₁pxy≡x _ _)) (stDC⊩isStrictDomainC x  r r⊩Cx⊤) ,
                             subst  r'  r'   C .relation  (x , )) (sym (pr₂pxy≡y _ _)) r⊩Cx⊤)
                        return
-                         (subst  r'  r'   ϕ .predicate  x) (sym (λ*ComputationRule realizer r)) (relϕ⊩isRelationalϕ x' x _ _ ⊩ϕx' ⊩x'~x))))))
+                         (subst  r'  r'   ϕ .predicate  x) (sym (λ*ComputationRule realizer r)) (relϕ⊩isRelationalϕ x' x _ _ ⊩ϕx' ⊩x'~x))))))
           h
           incψ≡h⋆incϕ
 
@@ -881,53 +881,53 @@
                 (d , d⊩X'x⊤≤ϕx)  PullbackHelper.⊩Cx⊤≤ϕx charFuncRel' commutes classifies
                 let
                   closure1 : ApplStrTerm as 2
-                  closure1 = ` pr₂ ̇ (` a ̇ (` pair ̇ (` pair ̇ (` stDX'  ̇ # one) ̇ # zero) ̇ # one)) ̇ ` k
+                  closure1 = ` pr₂ ̇ (` a ̇ (` pair ̇ (` pair ̇ (` stDX'  ̇ # one) ̇ # zero) ̇ # one)) ̇ ` k
                   closure2 : ApplStrTerm as 2
-                  closure2 = ` d ̇ (` relX' ̇ (` stDX' ̇ # one) ̇ # one ̇ (` pair ̇ ` k ̇ (` k ̇ # zero)))
+                  closure2 = ` d ̇ (` relX' ̇ (` stDX' ̇ # one) ̇ # one ̇ (` pair ̇ ` k ̇ (` k ̇ # zero)))
                   realizer : ApplStrTerm as 1
-                  realizer = ` pair ̇ (` stDX' ̇ # zero) ̇ (` pair ̇ λ*abst closure1 ̇ λ*abst closure2)
+                  realizer = ` pair ̇ (` stDX' ̇ # zero) ̇ (` pair ̇ λ*abst closure1 ̇ λ*abst closure2)
                 return
-                  (λ* realizer ,
+                  (λ* realizer ,
                     { x p r r⊩X'xp 
                      let
                        ⊩x~x = stDX'⊩isStrictDomainX' x p r r⊩X'xp
                      in
-                     subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x  _ _)) ⊩x~x ,
+                     subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (λ*ComputationRule realizer r)  pr₁pxy≡x  _ _)) ⊩x~x ,
                       b b⊩ϕx 
                        let
                          goal =
                            a⊩inc⋆X'≤term⋆true
-                             x p (pair  (pair  (stDX'  r)  b)  r)
+                             x p (pair  (pair  (stDX'  r)  b)  r)
                              (return
-                               (x , (subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _)) ⊩x~x ,
-                               subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _)) b⊩ϕx) ,
+                               (x , (subst  r'  r'   perX .equality  (x , x)) (sym (cong  x  pr₁  x) (pr₁pxy≡x _ _)  pr₁pxy≡x _ _)) ⊩x~x ,
+                               subst  r'  r'   ϕ .predicate  x) (sym (cong  x  pr₂  x) (pr₁pxy≡x _ _)  pr₂pxy≡y _ _)) b⊩ϕx) ,
                                subst  r'  r'   charFuncRel' .relation  (x , p)) (sym (pr₂pxy≡y _ _)) r⊩X'xp))
 
-                         eq : pr₁  (pr₂  (λ* realizer  r))  b  pr₂  (a  (pair  (pair  (stDX'  r)  b)  r))  k
+                         eq : pr₁  (pr₂  (λ* realizer  r))  b  pr₂  (a  (pair  (pair  (stDX'  r)  b)  r))  k
                          eq =
-                           cong  x  pr₁  (pr₂  x)  b) (λ*ComputationRule realizer r)  cong  x  pr₁  x  b) (pr₂pxy≡y _ _)  cong  x  x  b) (pr₁pxy≡x _ _)  βreduction closure1 b (r  [])
+                           cong  x  pr₁  (pr₂  x)  b) (λ*ComputationRule realizer r)  cong  x  pr₁  x  b) (pr₂pxy≡y _ _)  cong  x  x  b) (pr₁pxy≡x _ _)  βreduction closure1 b (r  [])
                        in
                        equivFun
                          (propTruncIdempotent≃ (toPredicate p .isPropValued _ _))
                          (do
                            (tt* , ⊩'x~x , ⊩⊤≤p)  goal
-                           return (subst  r'  r'   toPredicate p  tt*) (sym eq) (⊩⊤≤p k)))) ,
+                           return (subst  r'  r'   toPredicate p  tt*) (sym eq) (⊩⊤≤p k)))) ,
                       c c⊩p 
                        let
                          ⊩X'x⊤ =
                            relX'⊩isRelationalX'
                              x x p  _ _
-                             (pair  k  (k  c))
+                             (pair  k  (k  c))
                              ⊩x~x r⊩X'xp
-                             ((λ b b⊩p  subst  q  (pr₁  (pair  k  (k  c)))   q  tt*) (sym (compIsIdFunc truePredicate)) tt*) ,
-                               b b⊩⊤  subst  r'  r'   toPredicate p  tt*) (sym (cong  x  x  b) (pr₂pxy≡y _ _)  kab≡a _ _)) c⊩p))
+                             ((λ b b⊩p  subst  q  (pr₁  (pair  k  (k  c)))   q  tt*) (sym (compIsIdFunc truePredicate)) tt*) ,
+                               b b⊩⊤  subst  r'  r'   toPredicate p  tt*) (sym (cong  x  x  b) (pr₂pxy≡y _ _)  kab≡a _ _)) c⊩p))
 
-                         eq : pr₂  (pr₂  (λ* realizer  r))  c  d  (relX'  (stDX'  r)  r  (pair  k  (k  c)))
+                         eq : pr₂  (pr₂  (λ* realizer  r))  c  d  (relX'  (stDX'  r)  r  (pair  k  (k  c)))
                          eq =
-                           cong  x  pr₂  (pr₂  x)  c) (λ*ComputationRule realizer r) 
-                           cong  x  pr₂  x  c) (pr₂pxy≡y _ _) 
-                           cong  x  x  c) (pr₂pxy≡y _ _) 
-                           βreduction closure2 c (r  [])
+                           cong  x  pr₂  (pr₂  x)  c) (λ*ComputationRule realizer r) 
+                           cong  x  pr₂  x  c) (pr₂pxy≡y _ _) 
+                           cong  x  x  c) (pr₂pxy≡y _ _) 
+                           βreduction closure2 c (r  [])
                        in
                        subst
                           r'  r'   ϕ .predicate  x)
diff --git a/docs/Realizability.Topos.TerminalObject.html b/docs/Realizability.Topos.TerminalObject.html
index 0eb2b9f..77abc3d 100644
--- a/docs/Realizability.Topos.TerminalObject.html
+++ b/docs/Realizability.Topos.TerminalObject.html
@@ -1,5 +1,5 @@
 
-Realizability.Topos.TerminalObject
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+Realizability.Topos.TerminalObject
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Realizability.CombinatoryAlgebra
 open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.HLevels
@@ -17,7 +17,7 @@
   { ℓ' ℓ''}
   {A : Type }
   (ca : CombinatoryAlgebra A)
-  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
+  (isNonTrivial : CombinatoryAlgebra.s ca  CombinatoryAlgebra.k ca  ) where
 
 open CombinatoryAlgebra ca
 open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
@@ -35,9 +35,9 @@
   isPropValued (equality terminalPer) _ _ = isPropUnit*
   isPartialEquivalenceRelation.isSetX (isPerEquality terminalPer) = isSetUnit*
   isPartialEquivalenceRelation.isSymmetric (isPerEquality terminalPer) =
-    return (k ,  { tt* tt* r tt*  tt* }))
+    return (k ,  { tt* tt* r tt*  tt* }))
   isPartialEquivalenceRelation.isTransitive (isPerEquality terminalPer) =
-    return (k ,  { tt* tt* tt* _ _ tt* tt*  tt* }))
+    return (k ,  { tt* tt* tt* _ _ tt* tt*  tt* }))
 
 open FunctionalRelation
 
@@ -54,22 +54,22 @@
       ; isFuncRel =
         record
           { isStrictDomain = return (Id ,  { y tt* r r⊩y~y  subst  r'  r'   perY .equality  (y , y)) (sym (Ida≡a _)) r⊩y~y }))
-          ; isStrictCodomain = return (k ,  { y tt* r r⊩y~y  tt* }))
+          ; isStrictCodomain = return (k ,  { y tt* r r⊩y~y  tt* }))
           ; isRelational =
             (do
             (t , t⊩isTransitive)  perY .isTransitive
             (s , s⊩isSymmetric)  perY .isSymmetric
             let
               prover : ApplStrTerm as 3
-              prover = ` t ̇ (` s ̇ # two) ̇ # two
+              prover = ` t ̇ (` s ̇ # two) ̇ # two
             return
-              (λ*3 prover ,
+              (λ*3 prover ,
                { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt* 
                 subst
                    r'  r'   perY .equality  (y' , y'))
-                  (sym (λ*3ComputationRule prover a b c))
-                  (t⊩isTransitive y' y y' (s  a) a (s⊩isSymmetric y y' a a⊩y~y') a⊩y~y') })))
-          ; isSingleValued = (return (k ,  { y tt* tt* r₁ r₂ r₁⊩y~y r₂⊩y~y  tt* })))
+                  (sym (λ*3ComputationRule prover a b c))
+                  (t⊩isTransitive y' y y' (s  a) a (s⊩isSymmetric y y' a a⊩y~y') a⊩y~y') })))
+          ; isSingleValued = (return (k ,  { y tt* tt* r₁ r₂ r₁⊩y~y r₂⊩y~y  tt* })))
           ; isTotal = return
                       (Id ,
                        y r r⊩y~y 
@@ -93,15 +93,15 @@
                   (stFD , stFD⊩isStrictDomainF)  F .isStrictDomain
                   let
                     prover : ApplStrTerm as 1
-                    prover = ` tlG ̇ (` stFD ̇ # zero)
+                    prover = ` tlG ̇ (` stFD ̇ # zero)
                   return
-                    (λ* prover ,
+                    (λ* prover ,
                      { y tt* r r⊩Fy 
                       transport
                         (propTruncIdempotent (G .relation .isPropValued _ _))
                         (do
-                          (tt* , tlGstFD⊩Gy)  tlG⊩isTotalG y (stFD  r) (stFD⊩isStrictDomainF y tt* r r⊩Fy)
-                          return (subst  r'  r'   G .relation  (y , tt*)) (sym (λ*ComputationRule prover r)) tlGstFD⊩Gy)) }))
+                          (tt* , tlGstFD⊩Gy)  tlG⊩isTotalG y (stFD  r) (stFD⊩isStrictDomainF y tt* r r⊩Fy)
+                          return (subst  r'  r'   G .relation  (y , tt*)) (sym (λ*ComputationRule prover r)) tlGstFD⊩Gy)) }))
             in
             eq/ _ _ (answer , F≤G→G≤F perY terminalPer F G answer))
           f g
diff --git a/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html b/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html
index f96528d..49a3f10 100644
--- a/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html
+++ b/docs/Realizability.Tripos.Prealgebra.Meets.Commutativity.html
@@ -28,14 +28,14 @@
   x⊓y≤y⊓x x y =
     do
       let
-        proof : Term as 1
-        proof = ` pair ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero)
+        proof : Term as 1
+        proof = ` pair ̇ (` pr₂ ̇ # zero) ̇ (` pr₁ ̇ # zero)
       return
-        (λ* proof ,
+        (λ* proof ,
            x' a a⊩x⊓y 
             subst
                r  r   y  x  x')
-              (sym (λ*ComputationRule proof a))
+              (sym (λ*ComputationRule proof a))
               ((subst  r  r   y  x') (sym (pr₁pxy≡x _ _)) (a⊩x⊓y .snd)) ,
                (subst  r  r   x  x') (sym (pr₂pxy≡y _ _)) (a⊩x⊓y .fst)))))
 
@@ -45,15 +45,15 @@
       (f , f⊩x≤y)  x≤y
       (g , g⊩y≤x)  y≤x
       let
-        proof : Term as 1
-        proof = ` pair ̇ (` f ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
+        proof : Term as 1
+        proof = ` pair ̇ (` f ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)
       return
-        ((λ* proof) ,
+        ((λ* proof) ,
            x' a a⊩x⊓z 
             subst
                r  r   y  z  x')
-              (sym (λ*ComputationRule proof a))
-              ((subst  r  r   y  x') (sym (pr₁pxy≡x _ _)) (f⊩x≤y x' (pr₁  a) (a⊩x⊓z .fst))) ,
+              (sym (λ*ComputationRule proof a))
+              ((subst  r  r   y  x') (sym (pr₁pxy≡x _ _)) (f⊩x≤y x' (pr₁  a) (a⊩x⊓z .fst))) ,
                (subst  r  r   z  x') (sym (pr₂pxy≡y _ _)) (a⊩x⊓z .snd)))))
 
 
@@ -63,14 +63,14 @@
       (f , f⊩y≤z)  y≤z
       (g , g⊩z≤y)  z≤y
       let
-        proof : Term as 1
-        proof = ` pair ̇ (`  pr₁ ̇ # zero) ̇ (` f ̇ (` pr₂ ̇ # zero))
+        proof : Term as 1
+        proof = ` pair ̇ (`  pr₁ ̇ # zero) ̇ (` f ̇ (` pr₂ ̇ # zero))
       return
-        ((λ* proof) ,
+        ((λ* proof) ,
            { x' a (pr₁a⊩x , pr₂a⊩y) 
             subst
                r  r   x  z  x')
-              (sym (λ*ComputationRule proof a))
+              (sym (λ*ComputationRule proof a))
               ((subst  r  r   x  x') (sym (pr₁pxy≡x _ _)) pr₁a⊩x) ,
-               (subst  r  r   z  x') (sym (pr₂pxy≡y _ _)) (f⊩y≤z x' (pr₂  a) pr₂a⊩y))) }))
+               (subst  r  r   z  x') (sym (pr₂pxy≡y _ _)) (f⊩y≤z x' (pr₂  a) pr₂a⊩y))) }))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html b/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html index 90b9965..08a4f58 100644 --- a/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html +++ b/docs/Realizability.Tripos.Prealgebra.Meets.Identity.html @@ -18,7 +18,7 @@ open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s k ) where +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s k ) where private PredicateX = Predicate X open Predicate open PredicateProperties X @@ -36,14 +36,14 @@ x≤x⊓1 x = do let - proof : Term as 1 - proof = ` pair ̇ # zero ̇ ` true + proof : Term as 1 + proof = ` pair ̇ # zero ̇ ` true return - ((λ* proof) , + ((λ* proof) , x' a a⊩x subst r x pre1 x' r) - (sym (λ*ComputationRule proof a)) + (sym (λ*ComputationRule proof a)) (subst r r x x') (sym (pr₁pxy≡x _ _)) @@ -56,14 +56,14 @@ x≤1⊓x x = do let - proof : Term as 1 - proof = ` pair ̇ ` false ̇ # zero + proof : Term as 1 + proof = ` pair ̇ ` false ̇ # zero return - ((λ* proof) , + ((λ* proof) , x' a a⊩x subst r r pre1 x x') - (sym (λ*ComputationRule proof a)) + (sym (λ*ComputationRule proof a)) (tt* , (subst r r x x') diff --git a/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html b/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html index 50a1a61..939904d 100644 --- a/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html +++ b/docs/Realizability.Tripos.Prealgebra.Predicate.Properties.html @@ -1,6 +1,6 @@ Realizability.Tripos.Prealgebra.Predicate.Properties
open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
 open import Cubical.Foundations.Prelude as P
 open import Cubical.Foundations.HLevels
 open import Cubical.Foundations.Equiv
@@ -31,7 +31,7 @@
   private PredicateX = Predicate X
   open Predicate
   _≤_ : Predicate  X  Predicate  X  Type (ℓ-max (ℓ-max  ℓ') ℓ'')
-  ϕ  ψ = ∃[ b  A ] (∀ (x : X) (a : A)  a  ( ϕ  x)  (b  a)   ψ  x)
+  ϕ  ψ = ∃[ b  A ] (∀ (x : X) (a : A)  a  ( ϕ  x)  (b  a)   ψ  x)
 
   isProp≤ :  ϕ ψ  isProp (ϕ  ψ)
   isProp≤ ϕ ψ = isPropPropTrunc
@@ -49,7 +49,7 @@
                                subst
                                   r  r   ξ  x)
                                  (sym (Ba≡gfa b a a'))
-                                 (ψ≤[b]ξ x (a  a')
+                                 (ψ≤[b]ξ x (a  a')
                                  (ϕ≤[a]ψ x a' a'⊩ϕx))))
     
 
@@ -71,19 +71,19 @@
   infix 25 _⊔_
   _⊔_ : PredicateX  PredicateX  PredicateX
   (ϕ  ψ) .isSetX = ϕ .isSetX
-   ϕ  ψ  x a =  ((pr₁  a  k) × ((pr₂  a)   ϕ  x))  ((pr₁  a  k') × ((pr₂  a)   ψ  x)) ∥₁
+   ϕ  ψ  x a =  ((pr₁  a  k) × ((pr₂  a)   ϕ  x))  ((pr₁  a  k') × ((pr₂  a)   ψ  x)) ∥₁
   (ϕ  ψ) .isPropValued x a = isPropPropTrunc
 
   infix 25 _⊓_
   _⊓_ : PredicateX  PredicateX  PredicateX
   (ϕ  ψ) .isSetX = ϕ .isSetX
-   ϕ  ψ  x a = ((pr₁  a)   ϕ  x) × ((pr₂  a)   ψ  x)
-  (ϕ  ψ) .isPropValued x a = isProp× (ϕ .isPropValued x (pr₁  a)) (ψ .isPropValued x (pr₂  a))
+   ϕ  ψ  x a = ((pr₁  a)   ϕ  x) × ((pr₂  a)   ψ  x)
+  (ϕ  ψ) .isPropValued x a = isProp× (ϕ .isPropValued x (pr₁  a)) (ψ .isPropValued x (pr₂  a))
 
   infix 25 _⇒_
   _⇒_ : PredicateX  PredicateX  PredicateX
   (ϕ  ψ) .isSetX = ϕ .isSetX
-   ϕ  ψ  x a =  b  (b   ϕ  x)  (a  b)   ψ  x
+   ϕ  ψ  x a =  b  (b   ϕ  x)  (a  b)   ψ  x
   (ϕ  ψ) .isPropValued x a = isPropΠ λ a  isPropΠ λ a⊩ϕx  ψ .isPropValued _ _
 
 module _ where
@@ -94,37 +94,37 @@
     Lift' = Lift {i = } {j = (ℓ-max ℓ' ℓ'')}
 
     kRealized : Predicate' Unit*
-    kRealized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a  Lift' (a  k) ; isPropValued = λ x a  isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k) }
+    kRealized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a  Lift' (a  k) ; isPropValued = λ x a  isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k) }
 
     k'Realized : Predicate' Unit*
-    k'Realized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a  Lift' (a  k') ; isPropValued = λ x a  isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k') }
+    k'Realized = record { isSetX = isSetUnit* ; ∣_∣ = λ x a  Lift' (a  k') ; isPropValued = λ x a  isOfHLevelRespectEquiv 1 LiftEquiv (isSetA a k') }
 
     kRealized≤k'Realized : kRealized  k'Realized
     kRealized≤k'Realized =
       do
         let
           prover : ApplStrTerm as 1
-          prover = ` k'
-        return (λ* prover , λ { x a (lift a≡k)  lift (λ*ComputationRule prover a) })
+          prover = ` k'
+        return (λ* prover , λ { x a (lift a≡k)  lift (λ*ComputationRule prover a) })
 
     k'Realized≤kRealized : k'Realized  kRealized
     k'Realized≤kRealized =
       do
         let
           prover : ApplStrTerm as 1
-          prover = ` k
-        return (λ* prover , λ { x a (lift a≡k')  lift (λ*ComputationRule prover a) })
+          prover = ` k
+        return (λ* prover , λ { x a (lift a≡k')  lift (λ*ComputationRule prover a) })
 
     kRealized≡k'Realized : kRealized  k'Realized
     kRealized≡k'Realized = antiSym kRealized k'Realized kRealized≤k'Realized k'Realized≤kRealized
 
-    Lift≡ : Lift' (k  k)  Lift' (k  k')
-    Lift≡ i =  kRealized≡k'Realized i  tt* k
+    Lift≡ : Lift' (k  k)  Lift' (k  k')
+    Lift≡ i =  kRealized≡k'Realized i  tt* k
 
-    Liftk≡k' : Lift' (k  k')
+    Liftk≡k' : Lift' (k  k')
     Liftk≡k' = transport Lift≡ (lift refl)
 
-    k≡k' : k  k'
+    k≡k' : k  k'
     k≡k' = Liftk≡k' .lower
 
 module Morphism {X Y : Type ℓ'} (isSetX : isSet X) (isSetY : isSet Y)  where
@@ -149,8 +149,8 @@
     λ ϕ 
       record
         { isSetX = isSetY
-        ; ∣_∣ = λ y a  (∀ b x  f x  y  (a  b)   ϕ  x)
-        ; isPropValued = λ y a  isPropΠ λ a'  isPropΠ λ x  isPropΠ λ fx≡y  ϕ .isPropValued x (a  a') }
+        ; ∣_∣ = λ y a  (∀ b x  f x  y  (a  b)   ϕ  x)
+        ; isPropValued = λ y a  isPropΠ λ a'  isPropΠ λ x  isPropΠ λ fx≡y  ϕ .isPropValued x (a  a') }
 
   `∃[_] : (X  Y)  (PredicateX  PredicateY)
   `∃[ f ] =
@@ -178,10 +178,10 @@
          y b b⊩∃fϕ 
           equivFun
             (propTruncIdempotent≃
-              (ψ .isPropValued y (a~  b)))
+              (ψ .isPropValued y (a~  b)))
               (do
                 (x , fx≡y , b⊩ϕx)  b⊩∃fϕ
-                return (subst  y'  (a~  b)   ψ  y') fx≡y (a~proves x b b⊩ϕx)))))
+                return (subst  y'  (a~  b)   ψ  y') fx≡y (a~proves x b b⊩ϕx)))))
 
   `∃isLeftAdjoint :  ϕ ψ f  `∃[ f ] ϕ ≤Y ψ  ϕ ≤X ( f) ψ
   `∃isLeftAdjoint ϕ ψ f =
@@ -195,73 +195,73 @@
   `∀isRightAdjoint→ ϕ ψ f p =
     do
       (a~ , a~proves)  p
-      let realizer = (s  (s  (k  a~)  Id)  (k  k))
+      let realizer = (s  (s  (k  a~)  Id)  (k  k))
       return
         (realizer ,
          x a a⊩ψfx 
           equivFun
             (propTruncIdempotent≃
-              (ϕ .isPropValued x (realizer  a) ))
+              (ϕ .isPropValued x (realizer  a) ))
               (do
                 let ∀prover = a~proves (f x) a a⊩ψfx
                 return
                   (subst
                      ϕ~  ϕ~   ϕ  x)
                     (sym
-                      (realizer  a
+                      (realizer  a
                         ≡⟨ refl 
-                       s  (s  (k  a~)  Id)  (k  k)  a
-                        ≡⟨ sabc≡ac_bc _ _ _ 
-                       s  (k  a~)  Id  a  (k  k  a)
-                        ≡⟨ cong  x  x  (k  k  a)) (sabc≡ac_bc _ _ _) 
-                       k  a~  a  (Id  a)  (k  k  a)
-                        ≡⟨ cong  x  k  a~  a  x  (k  k  a)) (Ida≡a a) 
-                       k  a~  a  a  (k  k  a)
-                        ≡⟨ cong  x  k  a~  a  a  x) (kab≡a _ _) 
-                       (k  a~  a)  a  k
-                        ≡⟨ cong  x  x  a  k) (kab≡a _ _) 
-                       a~  a  k
+                       s  (s  (k  a~)  Id)  (k  k)  a
+                        ≡⟨ sabc≡ac_bc _ _ _ 
+                       s  (k  a~)  Id  a  (k  k  a)
+                        ≡⟨ cong  x  x  (k  k  a)) (sabc≡ac_bc _ _ _) 
+                       k  a~  a  (Id  a)  (k  k  a)
+                        ≡⟨ cong  x  k  a~  a  x  (k  k  a)) (Ida≡a a) 
+                       k  a~  a  a  (k  k  a)
+                        ≡⟨ cong  x  k  a~  a  a  x) (kab≡a _ _) 
+                       (k  a~  a)  a  k
+                        ≡⟨ cong  x  x  a  k) (kab≡a _ _) 
+                       a~  a  k
                          ))
-                    (∀prover k x refl)))))
+                    (∀prover k x refl)))))
 
   `∀isRightAdjoint← :  ϕ ψ f  ( f) ψ ≤X ϕ  ψ ≤Y `∀[ f ] ϕ
   `∀isRightAdjoint← ϕ ψ f p =
     do
       (a~ , a~proves)  p
-      let realizer = (s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id))
+      let realizer = (s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id))
       return
         (realizer ,
          y b b⊩ψy a x fx≡y 
           subst
              r  r   ϕ  x)
             (sym
-              (realizer  b  a
+              (realizer  b  a
                  ≡⟨ refl 
-               s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id)  b  a
-                 ≡⟨ cong  x  x  a) (sabc≡ac_bc _ _ _) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  (s  (k  k)  Id  b)  a
-                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  x  a) (sabc≡ac_bc (k  k) Id b) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  ((k  k  b)  (Id  b))  a
-                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (x  (Id  b))  a) (kab≡a _ _) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  (Id  b))  a
-                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (k  x)  a) (Ida≡a b) 
-               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  b)  a
-                 ≡⟨ cong  x  x  (k  b)  a) (sabc≡ac_bc _ _ _) 
-               k  s  b  (s  (k  k)  (k  a~)  b)  (k  b)  a
-                 ≡⟨ cong  x  x  (s  (k  k)  (k  a~)  b)  (k  b)  a) (kab≡a _ _) 
-               s  (s  (k  k)  (k  a~)  b)  (k  b)  a
-                 ≡⟨ sabc≡ac_bc _ _ _ 
-               s  (k  k)  (k  a~)  b  a  (k  b  a)
-                 ≡⟨ cong  x  s  (k  k)  (k  a~)  b  a  x) (kab≡a b a) 
-               s  (k  k)  (k  a~)  b  a  b
-                 ≡⟨ cong  x  x  a  b) (sabc≡ac_bc (k  k) (k  a~) b) 
-               k  k  b  (k  a~  b)  a  b
-                 ≡⟨ cong  x  x  (k  a~  b)  a  b) (kab≡a _ _) 
-               k  (k  a~  b)  a  b
-                 ≡⟨ cong  x  k  x  a  b) (kab≡a _ _) 
-               k  a~  a  b
-                 ≡⟨ cong  x  x  b) (kab≡a _ _) 
-               a~  b
+               s  (s  (k  s)  (s  (k  k)  (k  a~)))  (s  (k  k)  Id)  b  a
+                 ≡⟨ cong  x  x  a) (sabc≡ac_bc _ _ _) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  (s  (k  k)  Id  b)  a
+                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  x  a) (sabc≡ac_bc (k  k) Id b) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  ((k  k  b)  (Id  b))  a
+                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (x  (Id  b))  a) (kab≡a _ _) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  (Id  b))  a
+                 ≡⟨ cong  x  s  (k  s)  (s  (k  k)  (k  a~))  b  (k  x)  a) (Ida≡a b) 
+               s  (k  s)  (s  (k  k)  (k  a~))  b  (k  b)  a
+                 ≡⟨ cong  x  x  (k  b)  a) (sabc≡ac_bc _ _ _) 
+               k  s  b  (s  (k  k)  (k  a~)  b)  (k  b)  a
+                 ≡⟨ cong  x  x  (s  (k  k)  (k  a~)  b)  (k  b)  a) (kab≡a _ _) 
+               s  (s  (k  k)  (k  a~)  b)  (k  b)  a
+                 ≡⟨ sabc≡ac_bc _ _ _ 
+               s  (k  k)  (k  a~)  b  a  (k  b  a)
+                 ≡⟨ cong  x  s  (k  k)  (k  a~)  b  a  x) (kab≡a b a) 
+               s  (k  k)  (k  a~)  b  a  b
+                 ≡⟨ cong  x  x  a  b) (sabc≡ac_bc (k  k) (k  a~) b) 
+               k  k  b  (k  a~  b)  a  b
+                 ≡⟨ cong  x  x  (k  a~  b)  a  b) (kab≡a _ _) 
+               k  (k  a~  b)  a  b
+                 ≡⟨ cong  x  k  x  a  b) (kab≡a _ _) 
+               k  a~  a  b
+                 ≡⟨ cong  x  x  b) (kab≡a _ _) 
+               a~  b
                    ))
             (a~proves x b (subst  x'  b   ψ  x') (sym fx≡y) b⊩ψy))))
 
diff --git a/docs/index.html b/docs/index.html
index b05e2a9..b173d19 100644
--- a/docs/index.html
+++ b/docs/index.html
@@ -5,5 +5,8 @@
 open import Realizability.CombinatoryAlgebra
 open import Realizability.ApplicativeStructure
 open import Realizability.Topos.Everything
-open import Realizability.Choice
+open import Realizability.Assembly.Everything
+open import Realizability.PERs.Everything
+open import Realizability.Modest.Everything
+open import Realizability.Choice
 
\ No newline at end of file diff --git a/src/Realizability/Modest/Everything.agda b/src/Realizability/Modest/Everything.agda new file mode 100644 index 0000000..44d8ce4 --- /dev/null +++ b/src/Realizability/Modest/Everything.agda @@ -0,0 +1,9 @@ +module Realizability.Modest.Everything where + +open import Realizability.Modest.Base +open import Realizability.Modest.CanonicalPER +open import Realizability.Modest.UniformFamily +open import Realizability.Modest.UniformFamilyCleavage +open import Realizability.Modest.PartialSurjection +-- open import Realizability.Modest.GenericUniformFamily +open import Realizability.Modest.SubQuotientCanonicalPERIso diff --git a/src/Realizability/Modest/PartialSurjection.agda b/src/Realizability/Modest/PartialSurjection.agda index a475d78..155b73b 100644 --- a/src/Realizability/Modest/PartialSurjection.agda +++ b/src/Realizability/Modest/PartialSurjection.agda @@ -386,12 +386,3 @@ Category.⋆IdL PARTSURJ {X , surjX} {Y , surjY} f = idLPartSurjMorphism f Category.⋆IdR PARTSURJ {X , surjX} {Y , surjY} f = idRPartSurjMorphism f Category.⋆Assoc PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} {W , surjW} f g h = assocComposePartSurjMorphism f g h Category.isSetHom PARTSURJ {X , surjX} {Y , surjY} = isSetPartialSurjectionMorphism surjX surjY - -open Category -open ModestSetIso - -L : Functor MOD PARTSURJ -Functor.F-ob L (X , modX) = X , (ModestSet→PartialSurjection X (modX .fst .isSetX) modX) -Functor.F-hom L {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = {!!} -Functor.F-id L = {!!} -Functor.F-seq L = {!!} diff --git a/src/Realizability/PERs/Everything.agda b/src/Realizability/PERs/Everything.agda new file mode 100644 index 0000000..31455f2 --- /dev/null +++ b/src/Realizability/PERs/Everything.agda @@ -0,0 +1,5 @@ +module Realizability.PERs.Everything where + +open import Realizability.PERs.PER +open import Realizability.PERs.ResizedPER +open import Realizability.PERs.SubQuotient diff --git a/src/Realizability/PERs/SubQuotient.agda b/src/Realizability/PERs/SubQuotient.agda index ae9720f..c515dd8 100644 --- a/src/Realizability/PERs/SubQuotient.agda +++ b/src/Realizability/PERs/SubQuotient.agda @@ -9,7 +9,10 @@ open import Cubical.Foundations.Univalence open import Cubical.Foundations.Powerset open import Cubical.Foundations.HLevels open import Cubical.Foundations.Path +open import Cubical.Foundations.Function open import Cubical.Functions.FunExtEquiv +open import Cubical.Functions.Embedding +open import Cubical.Functions.Surjection open import Cubical.Relation.Binary open import Cubical.Data.Sigma open import Cubical.Data.FinData @@ -84,7 +87,7 @@ module _ isPropMotive x y = isPropΠ3 λ _ _ _ → squash/ x y module _ (R S : PER) (f : perMorphism R S) where - + subQuotientAssemblyMorphism : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S) subQuotientAssemblyMorphism = SQ.rec @@ -142,6 +145,44 @@ module _ (R S : PER) (f : perMorphism R S) where (λ { (r , r~r) → eq/ _ _ (a~b r r~r) }) sq) +module _ (R S : PER) (f : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)) where + subQuotientAssemblyMorphism→perMorphism : perMorphism R S + subQuotientAssemblyMorphism→perMorphism = + PT.rec→Set squash/ mainMap mainMap2Constant (f .tracker) module InverseDefinition where + isSQTracker : A → Type ℓ + isSQTracker t = ∀ (q : subQuotient R) (a : A) → a ⊩[ subQuotientAssembly R ] q → ⟨ subQuotientRealizability S (t ⨾ a) (f .AssemblyMorphism.map q) ⟩ + -- 🤢🤮 + mainMap : Σ[ t ∈ A ] (isSQTracker t) → perMorphism R S + mainMap (t , t⊩f) = + [ t , + (λ r r' r~r' → + let + r~r : r ~[ R ] r + r~r = PER.isTransitive R r r' r r~r' (PER.isSymmetric R r r' r~r') + + r'~r' : r' ~[ R ] r' + r'~r' = PER.isTransitive R r' r r' (PER.isSymmetric R r r' r~r') r~r' + in + SQ.elimProp + {P = λ q → ∀ (t : A) → ⟨ subQuotientRealizability S (t ⨾ r) q ⟩ → ⟨ subQuotientRealizability S (t ⨾ r') q ⟩ → (t ⨾ r) ~[ S ] (t ⨾ r')} + (λ q → isPropΠ3 λ t _ _ → isProp~ (t ⨾ r) S (t ⨾ r')) + (λ { (s , s~s) t tr~s tr'~s → PER.isTransitive S (t ⨾ r) s (t ⨾ r') tr~s (PER.isSymmetric S (t ⨾ r') s tr'~s) }) + (f .AssemblyMorphism.map [ (r , r~r) ]) + t + (t⊩f [ (r , r~r) ] r r~r) + (subst (λ eq → ⟨ subQuotientRealizability S (t ⨾ r') (f .AssemblyMorphism.map eq) ⟩) (eq/ _ _ (PER.isSymmetric R r r' r~r')) (t⊩f [ (r' , r'~r') ] r' r'~r'))) ] + + mainMap2Constant : 2-Constant mainMap + mainMap2Constant (t , t⊩f) (t' , t'⊩f) = + eq/ _ _ + λ r r~r → + SQ.elimProp + {P = λ q → ⟨ subQuotientRealizability S (t ⨾ r) q ⟩ → ⟨ subQuotientRealizability S (t' ⨾ r) q ⟩ → (t ⨾ r) ~[ S ] (t' ⨾ r)} + (λ q → isPropΠ2 λ _ _ → isProp~ (t ⨾ r) S (t' ⨾ r)) + (λ { (s , s~s) tr~s t'r~s → PER.isTransitive S (t ⨾ r) s (t' ⨾ r) tr~s (PER.isSymmetric S (t' ⨾ r) s t'r~s) }) + (f .AssemblyMorphism.map [ (r , r~r) ]) + (t⊩f [ (r , r~r) ] r r~r) + (t'⊩f [ (r , r~r) ] r r~r) subQuotientModestSet : PER → MOD .Category.ob subQuotientModestSet R = subQuotient R , subQuotientAssembly R , isModestSubQuotientAssembly R @@ -176,3 +217,53 @@ Functor.F-seq subQuotientFunctor {R} {S} {T} f g = (λ { (a , a~a) (b , bIsTracker) (c , cIsTracker) → eq/ _ _ (subst (_~[ T ] (c ⨾ (b ⨾ a))) (sym (λ*ComputationRule (` c ̇ (` b ̇ # zero)) a)) (cIsTracker (b ⨾ a) (b ⨾ a) (bIsTracker a a a~a))) }) sq f g) + +hasPropFibersSubQuotientFunctor : ∀ R S → hasPropFibers (subQuotientAssemblyMorphism R S) +hasPropFibersSubQuotientFunctor R S f (x , sqX≡f) (y , sqY≡f) = + Σ≡Prop + (λ perMap → isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S) _ _) + (SQ.elimProp2 + {P = λ x y → subQuotientAssemblyMorphism R S x ≡ f → subQuotientAssemblyMorphism R S y ≡ f → x ≡ y} + (λ x y → isPropΠ2 λ _ _ → squash/ _ _) + (λ { (x , x⊩f) (y , y⊩f) sqX≡f sqY≡f → + eq/ _ _ + λ r r~r → + SQ.elimProp + {P = λ f[r] → ⟨ subQuotientRealizability S (x ⨾ r) f[r] ⟩ → ⟨ subQuotientRealizability S (y ⨾ r) f[r] ⟩ → (x ⨾ r) ~[ S ] (y ⨾ r)} + (λ f[r] → isPropΠ2 λ _ _ → isProp~ (x ⨾ r) S (y ⨾ r)) + (λ { (s , s~s) xr~s yr~s → PER.isTransitive S (x ⨾ r) s (y ⨾ r) xr~s (PER.isSymmetric S (y ⨾ r) s yr~s) }) + (f .AssemblyMorphism.map [ (r , r~r) ]) + (subst (λ f[r] → ⟨ subQuotientRealizability S (x ⨾ r) f[r] ⟩) (cong (λ m → m .AssemblyMorphism.map [ (r , r~r) ]) sqX≡f) (x⊩f r r r~r)) + (subst (λ f[r] → ⟨ subQuotientRealizability S (y ⨾ r) f[r] ⟩) (cong (λ m → m .AssemblyMorphism.map [ (r , r~r) ]) sqY≡f) (y⊩f r r r~r)) }) + x y sqX≡f sqY≡f) + +fiberSubQuotientFunctor : ∀ R S f → fiber (subQuotientAssemblyMorphism R S) f +fiberSubQuotientFunctor R S f = + (subQuotientAssemblyMorphism→perMorphism R S f) , + (AssemblyMorphism≡ _ _ + (funExt + (λ qR → + SQ.elimProp + {P = λ qR → subQuotientAssemblyMorphism R S (subQuotientAssemblyMorphism→perMorphism R S f) .AssemblyMorphism.map qR ≡ f .AssemblyMorphism.map qR} + (λ qR → squash/ _ _) + (λ { (r , r~r) → + PT.elim + {P = + λ fTracker → + subQuotientAssemblyMorphism R S (PT.rec→Set squash/ (InverseDefinition.mainMap R S f) (InverseDefinition.mainMap2Constant R S f) fTracker) .AssemblyMorphism.map [ (r , r~r) ] + ≡ f .AssemblyMorphism.map [ (r , r~r) ]} + (λ fTracker → squash/ _ _) + (λ { (t , tIsTracker) → + SQ.elimProp + {P = + λ fqR → ⟨ subQuotientRealizability S (t ⨾ r) fqR ⟩ → + subQuotientAssemblyMorphism R S (InverseDefinition.mainMap R S f (t , tIsTracker)) .AssemblyMorphism.map [ (r , r~r) ] ≡ fqR} + (λ fqR → isProp→ (squash/ _ _)) + (λ { (s , s~s) tr~s → eq/ _ _ tr~s }) + (f .AssemblyMorphism.map [ (r , r~r) ]) + (tIsTracker [ (r , r~r) ] r r~r) }) + (f .tracker) }) + qR))) + +isFullyFaithfulSubQuotientFunctor : Functor.isFullyFaithful subQuotientFunctor +equiv-proof (isFullyFaithfulSubQuotientFunctor R S) f = inhProp→isContr (fiberSubQuotientFunctor R S f) (hasPropFibersSubQuotientFunctor R S f) diff --git a/src/Realizability/Topos/SubobjectClassifier.agda b/src/Realizability/Topos/SubobjectClassifier.agda index 5cc8c95..d1d9a56 100644 --- a/src/Realizability/Topos/SubobjectClassifier.agda +++ b/src/Realizability/Topos/SubobjectClassifier.agda @@ -936,33 +936,3 @@ module ClassifiesStrictRelations char commutes classifies - -module ClassifiesSubobjects where - - subobjectClassifierUnivProp : Type _ - subobjectClassifierUnivProp = - ∀ {X Y : Type ℓ} - {perX : PartialEquivalenceRelation X} - {perY : PartialEquivalenceRelation Y} - → (f : RTMorphism perX perY) - → isMonic RT f - → ∃![ char ∈ RTMorphism perY Ωper ] - Σ[ commutes ∈ f ⋆ char ≡ [ terminalFuncRel perX ] ⋆ [ trueFuncRel ] ] - isPullback RT (cospan (Y , perY) ((ResizedPredicate Unit*) , Ωper) (Unit* , terminalPer) char [ trueFuncRel ]) f [ terminalFuncRel perX ] commutes - - isSubobjectClassifier : subobjectClassifierUnivProp - isSubobjectClassifier {X} {Y} {perX} {perY} f isMonicF = - SQ.elimProp - {P = λ f → ∀ (isMonic : isMonic RT f) → ∃![ char ∈ RTMorphism perY Ωper ] (Σ[ commutes ∈ f ⋆ char ≡ [ terminalFuncRel perX ] ⋆ [ trueFuncRel ] ] isPullback RT (cospan (Y , perY) ((ResizedPredicate Unit*) , Ωper) (Unit* , terminalPer) char [ trueFuncRel ]) f [ terminalFuncRel perX ] commutes) } - (λ f → isPropΠ λ isMonicF → isPropIsContr) - (λ F isMonicF → - let - ϕ = SubobjectIsoMonicFuncRel.ψ perY perX F isMonicF - in - uniqueExists - [ ClassifiesStrictRelations.charFuncRel Y perY ϕ ] - ({!ClassifiesStrictRelations.subobjectSquareCommutes Y perY ϕ!} , {!!}) - {!!} - {!!}) - f - isMonicF diff --git a/src/index.agda b/src/index.agda index c30a8cb..0ad3713 100644 --- a/src/index.agda +++ b/src/index.agda @@ -4,4 +4,7 @@ module index where open import Realizability.CombinatoryAlgebra open import Realizability.ApplicativeStructure open import Realizability.Topos.Everything +open import Realizability.Assembly.Everything +open import Realizability.PERs.Everything +open import Realizability.Modest.Everything open import Realizability.Choice From 93cc09800c920d01b6e6f956f0d1bcff5471c404 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sat, 12 Oct 2024 20:05:27 +0530 Subject: [PATCH 60/61] Clean and refactor tripos modules --- docs/Realizability.Assembly.Everything.html | 19 +- ...ty.Assembly.SetsReflectiveSubcategory.html | 2 +- docs/Realizability.Modest.CanonicalPER.html | 115 +- ...ealizability.Modest.PartialSurjection.html | 756 +++++------ ...ity.Modest.SubQuotientCanonicalPERIso.html | 297 ++-- ...Realizability.Modest.UnresizedGeneric.html | 14 +- docs/Realizability.PERs.SubQuotient.html | 491 +++---- docs/Realizability.Topos.Everything.html | 22 +- ...alizability.Topos.SubobjectClassifier.html | 8 +- docs/Realizability.Tripos.Everything.html | 16 +- .../Realizability.Tripos.Logic.Semantics.html | 1190 ++++++++--------- ...ability.Tripos.Prealgebra.Implication.html | 158 ++- ...Tripos.Prealgebra.Joins.Associativity.html | 892 ++++++------ ...zability.Tripos.Prealgebra.Joins.Base.html | 177 ++- ...Tripos.Prealgebra.Joins.Commutativity.html | 297 ++-- ...y.Tripos.Prealgebra.Joins.Idempotency.html | 85 +- ...lity.Tripos.Prealgebra.Joins.Identity.html | 203 ++- ...Tripos.Prealgebra.Meets.Associativity.html | 232 ++-- ...y.Tripos.Prealgebra.Meets.Idempotency.html | 70 +- docs/index.html | 12 +- realizability.agda-lib | 2 +- src/Realizability/Assembly/Everything.agda | 5 +- src/Realizability/Modest/CanonicalPER.agda | 9 +- src/Realizability/Modest/EquivToPERs.agda | 73 + .../Modest/PartialSurjection.agda | 2 + .../Modest/SubQuotientCanonicalPERIso.agda | 5 + src/Realizability/PERs/SubQuotient.agda | 5 +- src/Realizability/Topos/Everything.agda | 2 +- src/Realizability/Tripos/Everything.agda | 8 +- src/Realizability/Tripos/Logic/Semantics.agda | 78 +- .../Tripos/Prealgebra/Implication.agda | 32 +- .../Tripos/Prealgebra/Joins/#Base.agda# | 27 - .../Prealgebra/Joins/Associativity.agda | 58 +- .../Tripos/Prealgebra/Joins/Base.agda | 27 +- .../Prealgebra/Joins/Commutativity.agda | 21 +- .../Tripos/Prealgebra/Joins/Idempotency.agda | 21 +- .../Tripos/Prealgebra/Joins/Identity.agda | 29 +- .../Prealgebra/Meets/Associativity.agda | 26 +- .../Tripos/Prealgebra/Meets/Idempotency.agda | 20 +- src/index.agda | 4 +- 40 files changed, 2750 insertions(+), 2760 deletions(-) create mode 100644 src/Realizability/Modest/EquivToPERs.agda delete mode 100644 src/Realizability/Tripos/Prealgebra/Joins/#Base.agda# diff --git a/docs/Realizability.Assembly.Everything.html b/docs/Realizability.Assembly.Everything.html index f044745..852dd84 100644 --- a/docs/Realizability.Assembly.Everything.html +++ b/docs/Realizability.Assembly.Everything.html @@ -1,13 +1,12 @@ Realizability.Assembly.Everything
{-# OPTIONS --cubical #-}
-module Realizability.Assembly.Everything where
-open import Realizability.Assembly.Base
-open import Realizability.Assembly.BinCoproducts
-open import Realizability.Assembly.BinProducts
-open import Realizability.Assembly.Coequalizers
-open import Realizability.Assembly.Equalizers
-open import Realizability.Assembly.Exponentials
-open import Realizability.Assembly.Morphism
--- TODO : Fix regular structure modules
--- open import Realizability.Assembly.Regular.Everything
+-- These modules contain some elementary results about the category of assemblies
+-- In particular, finite limits and exponentials
+module Realizability.Assembly.Everything where
+open import Realizability.Assembly.Base
+open import Realizability.Assembly.BinCoproducts
+open import Realizability.Assembly.BinProducts
+open import Realizability.Assembly.Equalizers
+open import Realizability.Assembly.Exponentials
+open import Realizability.Assembly.Morphism
 
\ No newline at end of file diff --git a/docs/Realizability.Assembly.SetsReflectiveSubcategory.html b/docs/Realizability.Assembly.SetsReflectiveSubcategory.html index cf7c6ea..0065c5c 100644 --- a/docs/Realizability.Assembly.SetsReflectiveSubcategory.html +++ b/docs/Realizability.Assembly.SetsReflectiveSubcategory.html @@ -43,7 +43,7 @@ open UnitCounit adjointUnitCounit : forgetfulFunctor - NatTrans.N-ob (_⊣_.η adjointUnitCounit) (X , asmX) = makeAssemblyMorphism x x) (return (k , _ _ _ tt*))) + NatTrans.N-ob (_⊣_.η adjointUnitCounit) (X , asmX) = makeAssemblyMorphism x x) (return (k , _ _ _ tt*))) NatTrans.N-hom (_⊣_.η adjointUnitCounit) {X , asmX} {Y , asmY} f = AssemblyMorphism≡ _ _ refl NatTrans.N-ob (_⊣_.ε adjointUnitCounit) (X , isSetX) x = x NatTrans.N-hom (_⊣_.ε adjointUnitCounit) {X , isSetX} {Y , isSetY} f = refl diff --git a/docs/Realizability.Modest.CanonicalPER.html b/docs/Realizability.Modest.CanonicalPER.html index 309af56..733b398 100644 --- a/docs/Realizability.Modest.CanonicalPER.html +++ b/docs/Realizability.Modest.CanonicalPER.html @@ -1,62 +1,67 @@ -Realizability.Modest.CanonicalPER
open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Function
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Powerset
-open import Cubical.Foundations.Path
-open import Cubical.Foundations.Structure using (⟨_⟩; str)
-open import Cubical.Data.Sigma
-open import Cubical.Data.FinData
-open import Cubical.Data.Unit
-open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Reflection.RecordEquiv
-open import Cubical.Categories.Category
-open import Cubical.Categories.Displayed.Base
-open import Cubical.Categories.Displayed.Reasoning
-open import Cubical.Categories.Limits.Pullback
-open import Cubical.Categories.Functor hiding (Id)
-open import Cubical.Categories.Constructions.Slice
-open import Categories.CartesianMorphism
-open import Categories.GenericObject
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure
-open import Realizability.PropResizing
+Realizability.Modest.CanonicalPER
{-# OPTIONS --allow-unsolved-metas #-}
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Categories.Displayed.Reasoning
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Functor hiding (Id)
+open import Cubical.Categories.Constructions.Slice
+open import Categories.CartesianMorphism
+open import Categories.GenericObject
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
 
-module Realizability.Modest.CanonicalPER {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Modest.CanonicalPER {} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open import Realizability.Assembly.Base ca
-open import Realizability.Assembly.Morphism ca
-open import Realizability.Assembly.Terminal ca
-open import Realizability.Assembly.SetsReflectiveSubcategory ca
-open import Realizability.Modest.Base ca
-open import Realizability.Modest.UniformFamily ca
-open import Realizability.PERs.PER ca
-open import Realizability.PERs.SubQuotient ca
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.Terminal ca
+open import Realizability.Assembly.SetsReflectiveSubcategory ca
+open import Realizability.Modest.Base ca
+open import Realizability.Modest.UniformFamily ca
+open import Realizability.PERs.PER ca
+open import Realizability.PERs.SubQuotient ca
 
-open Assembly
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Contravariant UNIMOD
-open UniformFamily
-open DisplayedUFamMap
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Contravariant UNIMOD
+open UniformFamily
+open DisplayedUFamMap
 
-module _
-  {X : Type }
-  (asmX : Assembly X)
-  (isModestAsmX : isModest asmX) where
+module _
+  {X : Type }
+  (asmX : Assembly X)
+  (isModestAsmX : isModest asmX) where
 
-  canonicalPER : PER
-  PER.relation canonicalPER a b = Σ[ x  X ] a ⊩[ asmX ] x × b ⊩[ asmX ] x
-  PER.isPropValued canonicalPER a b (x , a⊩x , b⊩x) (x' , a⊩x' , b⊩x') =
-    Σ≡Prop
-       x  isProp× (asmX .⊩isPropValued a x) (asmX .⊩isPropValued b x))
-      (isModestAsmX x x' a a⊩x a⊩x')
-  fst (PER.isPER canonicalPER) a b (x , a⊩x , b⊩x) = x , b⊩x , a⊩x
-  snd (PER.isPER canonicalPER) a b c (x , a⊩x , b⊩x) (x' , b⊩x' , c⊩x') =
-    x' , subst (a ⊩[ asmX ]_) (isModestAsmX x x' b b⊩x b⊩x') a⊩x , c⊩x'
-    
-  
+  canonicalPER : PER
+  PER.relation canonicalPER a b = Σ[ x  X ] a ⊩[ asmX ] x × b ⊩[ asmX ] x
+  PER.isPropValued canonicalPER a b (x , a⊩x , b⊩x) (x' , a⊩x' , b⊩x') =
+    Σ≡Prop
+       x  isProp× (asmX .⊩isPropValued a x) (asmX .⊩isPropValued b x))
+      (isModestAsmX x x' a a⊩x a⊩x')
+  fst (PER.isPER canonicalPER) a b (x , a⊩x , b⊩x) = x , b⊩x , a⊩x
+  snd (PER.isPER canonicalPER) a b c (x , a⊩x , b⊩x) (x' , b⊩x' , c⊩x') =
+    x' , subst (a ⊩[ asmX ]_) (isModestAsmX x x' b b⊩x b⊩x') a⊩x , c⊩x'
+
+CanonicalPERFunctor : Functor MOD PERCat
+Functor.F-ob CanonicalPERFunctor (X , asmX , isModestAsmX) = canonicalPER asmX isModestAsmX
+Functor.F-hom CanonicalPERFunctor {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = {!!}
+Functor.F-id CanonicalPERFunctor = {!!}
+Functor.F-seq CanonicalPERFunctor = {!!}
 
\ No newline at end of file diff --git a/docs/Realizability.Modest.PartialSurjection.html b/docs/Realizability.Modest.PartialSurjection.html index 1eb27cb..ce4f6b7 100644 --- a/docs/Realizability.Modest.PartialSurjection.html +++ b/docs/Realizability.Modest.PartialSurjection.html @@ -1,223 +1,225 @@ -Realizability.Modest.PartialSurjection
open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Function
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Powerset
-open import Cubical.Foundations.Structure using (⟨_⟩; str)
-open import Cubical.Foundations.Univalence
-open import Cubical.Functions.Surjection
-open import Cubical.Functions.FunExtEquiv
-open import Cubical.Data.Sigma
-open import Cubical.Data.FinData
-open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Reflection.RecordEquiv
-open import Cubical.Categories.Category
-open import Cubical.Categories.Functor.Base hiding (Id)
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure
-open import Realizability.PropResizing
-
-module Realizability.Modest.PartialSurjection {} {A : Type } (ca : CombinatoryAlgebra A) (resizing : hPropResizing ) where
-
-open import Realizability.Assembly.Base ca
-open import Realizability.Assembly.Morphism ca
-open import Realizability.Assembly.SIP ca
-open import Realizability.Modest.Base ca
-
-open Assembly
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open ResizedPowerset resizing
-
-record PartialSurjection (X : Type ) : Type (ℓ-suc ) where
-  no-eta-equality
-  constructor makePartialSurjection
-  field
-    support : A  Type 
-    enumeration : Σ[ a  A ] (support a)  X
-    isPropSupport :  a  isProp (support a)
-    isSurjectionEnumeration : isSurjection enumeration
-    isSetX : isSet X -- potentially redundant?
-open PartialSurjection
-
-module _ (X : Type ) (isCorrectHLevel : isSet X) where
-  -- first we need a Σ type equivalent to partial surjections
-  -- we could use RecordEquiv but this does not give hProps and hSets and
-  -- that causes problems when trying to compute the hlevel
-
-  PartialSurjectionΣ : Type (ℓ-suc )
-  PartialSurjectionΣ = Σ[ support  (A  hProp ) ] Σ[ enumeration  ((Σ[ a  A ]  support a )  X) ] isSurjection enumeration × isSet X
-
-  isSetPartialSurjectionΣ : isSet PartialSurjectionΣ
-  isSetPartialSurjectionΣ = isSetΣ (isSet→ isSetHProp)  support  isSetΣ (isSet→ isCorrectHLevel)  enum  isSet× (isProp→isSet isPropIsSurjection) (isProp→isSet isPropIsSet)))
-
-  PartialSurjectionIsoΣ : Iso (PartialSurjection X) PartialSurjectionΣ
-  Iso.fun PartialSurjectionIsoΣ surj =
-     a  (surj .support a) , (surj .isPropSupport a)) ,
-     { (a , suppA)  surj .enumeration (a , suppA) }) ,
-    surj .isSurjectionEnumeration ,
-    PartialSurjection.isSetX surj
-  Iso.inv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) =
-    makePartialSurjection  a   support a ) enumeration  a  str (support a)) isSurjectionEnumeration isSetX
-  Iso.rightInv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = refl
-  support (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .support a
-  enumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) (a , suppA) = surj .enumeration (a , suppA)
-  isPropSupport (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .isPropSupport a
-  isSurjectionEnumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSurjectionEnumeration
-  isSetX (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSetX
-
-  PartialSurjection≡Σ : PartialSurjection X  PartialSurjectionΣ
-  PartialSurjection≡Σ = isoToPath PartialSurjectionIsoΣ
-
-  isSetPartialSurjection : isSet (PartialSurjection X)
-  isSetPartialSurjection = subst isSet (sym PartialSurjection≡Σ) isSetPartialSurjectionΣ
-
--- let us derive a structure of identity principle for partial surjections
-module SIP (X : Type ) (isCorrectHLevel : isSet X) where
-
-  PartialSurjection≡Iso :
-     (p q : PartialSurjection X)
-     Iso
-      (Σ[ suppPath  p .support  q .support ]
-      PathP  i  Σ[ a  A ] (suppPath i a)  X) (p .enumeration) (q .enumeration))
-      (p  q)
-  support (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z = suppPath i z
-  enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) (a , enum) = enumPath i (a , enum)
-  isPropSupport (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z =
-    isProp→PathP {B = λ j  isProp (suppPath j z)}  j  isPropIsProp) (p .isPropSupport z) (q .isPropSupport z) i
-  isSurjectionEnumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) b =
-    isProp→PathP
-      {B = λ j   fiber (enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) j)) b ∥₁}
-       j  isPropPropTrunc)
-      (p .isSurjectionEnumeration b) (q .isSurjectionEnumeration b) i
-  isSetX (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) = isPropIsSet (p .isSetX) (q .isSetX) i
-  Iso.inv (PartialSurjection≡Iso p q) p≡q =  i  p≡q i .support) ,  i  p≡q i .enumeration)
-  Iso.rightInv (PartialSurjection≡Iso p q) p≡q = isSetPartialSurjection X isCorrectHLevel _ _ _ _ 
-  Iso.leftInv (PartialSurjection≡Iso p q) (suppPath , enumPath) = ΣPathP (refl , refl)
-
-  PartialSurjection≡ :  (p q : PartialSurjection X)  Σ[ suppPath  p .support  q .support ] PathP  i  Σ[ a  A ] (suppPath i a)  X) (p .enumeration) (q .enumeration)  p  q
-  PartialSurjection≡ p q (suppPath , enumPath) = Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath)
-
--- the type of partial surjections is equivalent to the type of modest assemblies on X
-module ModestSetIso (X : Type ) (isCorrectHLevel : isSet X) where
-
-  open SIP X isCorrectHLevel
-
-  {-# TERMINATING #-}
-  ModestSet→PartialSurjection : ModestSet X  PartialSurjection X
-  support (ModestSet→PartialSurjection (xs , isModestXs)) r = ∃[ x  X ] (r ⊩[ xs ] x)
-  enumeration (ModestSet→PartialSurjection (xs , isModestXs)) (r , ∃x) =
-    let
-      answer : Σ[ x  X ] (r ⊩[ xs ] x)
-      answer = PT.rec (isUniqueRealized xs isModestXs r)  t  t) ∃x
-    in fst answer
-  isPropSupport (ModestSet→PartialSurjection (xs , isModestXs)) r = isPropPropTrunc
-  isSurjectionEnumeration (ModestSet→PartialSurjection (xs , isModestXs)) x =
-    do
-      (a , a⊩x)  xs .⊩surjective x
-      return ((a ,  x , a⊩x ∣₁) , refl)
-  isSetX (ModestSet→PartialSurjection (xs , isModestXs)) = xs .isSetX
-
-  PartialSurjection→ModestSet : PartialSurjection X  ModestSet X
-  Assembly._⊩_ (fst (PartialSurjection→ModestSet surj)) r x =
-    Σ[ s  surj .support r ] surj .enumeration (r , s)  x
-  Assembly.isSetX (fst (PartialSurjection→ModestSet surj)) = surj .isSetX
-  Assembly.⊩isPropValued (fst (PartialSurjection→ModestSet surj)) a x (s , ≡x) (t , ≡x') =
-    Σ≡Prop  u  surj .isSetX (surj .enumeration (a , u)) x) (surj .isPropSupport a s t)
-  Assembly.⊩surjective (fst (PartialSurjection→ModestSet surj)) x =
-    do
-      ((a , s) , ≡x)  surj .isSurjectionEnumeration x
-      return (a , (s , ≡x))
-  snd (PartialSurjection→ModestSet surj) x y r (s , ≡x) (t , ≡x') =
-    x
-      ≡⟨ sym ≡x 
-    surj .enumeration (r , s)
-      ≡⟨ cong  s  surj .enumeration (r , s)) (surj .isPropSupport r s t) 
-    surj .enumeration (r , t)
-      ≡⟨ ≡x' 
-    y
-      
-
-  opaque
-    rightInv :  surj  ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)  surj
-    rightInv surj =
-      PartialSurjection≡
-        (ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)) surj
-        (funExt supportEq ,
-        funExtDep
-          {A = λ i  Σ-syntax A (funExt supportEq i)}
-          {B = λ _ _  X}
-          {f = ModestSet→PartialSurjection (PartialSurjection→ModestSet surj) .enumeration}
-          {g = surj .enumeration}
-          λ { {r , ∃x} {s , supp} p 
-            PT.elim
-              {P = λ ∃x  fst
-                             (PT.rec
-                              (isUniqueRealized (fst (PartialSurjection→ModestSet surj))
-                               (snd (PartialSurjection→ModestSet surj)) r)
-                               t  t) ∃x)
-                           surj .enumeration (s , supp)}
-              ∃x  surj .isSetX _ _)
-              { (x , suppR , ≡x) 
-               let
-                 ∃x' = transport (sym (supportEq s)) supp
-                 r≡s : r  s
-                 r≡s = PathPΣ p .fst
-               in
-               equivFun
-                 (propTruncIdempotent≃ (surj .isSetX x (surj .enumeration (s , supp))))
-                 (do
-                   (x' , suppS , ≡x')  ∃x'
-                   return
-                     (x
-                       ≡⟨ sym ≡x 
-                     surj .enumeration (r , suppR)
-                       ≡⟨ cong (surj .enumeration) (ΣPathP (r≡s , (isProp→PathP  i  surj .isPropSupport (PathPΣ p .fst i)) suppR supp))) 
-                     surj .enumeration (s , supp)
-                       )) })
-             ∃x }) where
-          supportEq :  r  (∃[ x  X ] (Σ[ supp  surj .support r ] (surj .enumeration (r , supp)  x)))  support surj r
-          supportEq =
-               r 
-                hPropExt
-                isPropPropTrunc
-                (surj .isPropSupport r)
-                 ∃x  PT.rec (surj .isPropSupport r)  { (x , supp , ≡x)  supp }) ∃x)
-                 supp  return (surj .enumeration (r , supp) , supp , refl)))
-
-  leftInv :  mod  PartialSurjection→ModestSet (ModestSet→PartialSurjection mod)  mod
-  leftInv (asmX , isModestAsmX) =
-    Σ≡Prop
-      isPropIsModest
-      (Assembly≡ _ _
-        λ r x 
-          hPropExt
-            (isPropΣ isPropPropTrunc  ∃x  asmX .isSetX _ _))
-            (asmX .⊩isPropValued r x)
-             { (∃x , ≡x) 
-              let
-                (x' , r⊩x') = PT.rec (isUniqueRealized asmX isModestAsmX r)  t  t) ∃x
-              in subst  x'  r ⊩[ asmX ] x') ≡x r⊩x'})
-            λ r⊩x   x , r⊩x ∣₁ , refl)
-
-  IsoModestSetPartialSurjection : Iso (ModestSet X) (PartialSurjection X)
-  Iso.fun IsoModestSetPartialSurjection = ModestSet→PartialSurjection
-  Iso.inv IsoModestSetPartialSurjection = PartialSurjection→ModestSet
-  Iso.rightInv IsoModestSetPartialSurjection = rightInv 
-  Iso.leftInv IsoModestSetPartialSurjection = leftInv
-
-  ModestSet≡PartialSurjection : ModestSet X  PartialSurjection X
-  ModestSet≡PartialSurjection = isoToPath IsoModestSetPartialSurjection
-
-record PartialSurjectionMorphism {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y) : Type  where
-  no-eta-equality
-  constructor makePartialSurjectionMorphism
-  field
-    map : X  Y
-    {-
+Realizability.Modest.PartialSurjection
-- This module defines a **partial surjection** and shows that it is equivalent to a modest set.
+-- A partial surjection on a set X is a surjection from the combinatory algebra 𝔸 ↠ X that is only defined for a certain subset of 𝔸
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Foundations.Univalence
+open import Cubical.Functions.Surjection
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor.Base hiding (Id)
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.PartialSurjection {} {A : Type } (ca : CombinatoryAlgebra A) (resizing : hPropResizing ) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.SIP ca
+open import Realizability.Modest.Base ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open ResizedPowerset resizing
+
+record PartialSurjection (X : Type ) : Type (ℓ-suc ) where
+  no-eta-equality
+  constructor makePartialSurjection
+  field
+    support : A  Type 
+    enumeration : Σ[ a  A ] (support a)  X
+    isPropSupport :  a  isProp (support a)
+    isSurjectionEnumeration : isSurjection enumeration
+    isSetX : isSet X -- potentially redundant?
+open PartialSurjection
+
+module _ (X : Type ) (isCorrectHLevel : isSet X) where
+  -- first we need a Σ type equivalent to partial surjections
+  -- we could use RecordEquiv but this does not give hProps and hSets and
+  -- that causes problems when trying to compute the hlevel
+
+  PartialSurjectionΣ : Type (ℓ-suc )
+  PartialSurjectionΣ = Σ[ support  (A  hProp ) ] Σ[ enumeration  ((Σ[ a  A ]  support a )  X) ] isSurjection enumeration × isSet X
+
+  isSetPartialSurjectionΣ : isSet PartialSurjectionΣ
+  isSetPartialSurjectionΣ = isSetΣ (isSet→ isSetHProp)  support  isSetΣ (isSet→ isCorrectHLevel)  enum  isSet× (isProp→isSet isPropIsSurjection) (isProp→isSet isPropIsSet)))
+
+  PartialSurjectionIsoΣ : Iso (PartialSurjection X) PartialSurjectionΣ
+  Iso.fun PartialSurjectionIsoΣ surj =
+     a  (surj .support a) , (surj .isPropSupport a)) ,
+     { (a , suppA)  surj .enumeration (a , suppA) }) ,
+    surj .isSurjectionEnumeration ,
+    PartialSurjection.isSetX surj
+  Iso.inv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) =
+    makePartialSurjection  a   support a ) enumeration  a  str (support a)) isSurjectionEnumeration isSetX
+  Iso.rightInv PartialSurjectionIsoΣ (support , enumeration , isSurjectionEnumeration , isSetX) = refl
+  support (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .support a
+  enumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) (a , suppA) = surj .enumeration (a , suppA)
+  isPropSupport (Iso.leftInv PartialSurjectionIsoΣ surj i) a = surj .isPropSupport a
+  isSurjectionEnumeration (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSurjectionEnumeration
+  isSetX (Iso.leftInv PartialSurjectionIsoΣ surj i) = surj .isSetX
+
+  PartialSurjection≡Σ : PartialSurjection X  PartialSurjectionΣ
+  PartialSurjection≡Σ = isoToPath PartialSurjectionIsoΣ
+
+  isSetPartialSurjection : isSet (PartialSurjection X)
+  isSetPartialSurjection = subst isSet (sym PartialSurjection≡Σ) isSetPartialSurjectionΣ
+
+-- let us derive a structure of identity principle for partial surjections
+module SIP (X : Type ) (isCorrectHLevel : isSet X) where
+
+  PartialSurjection≡Iso :
+     (p q : PartialSurjection X)
+     Iso
+      (Σ[ suppPath  p .support  q .support ]
+      PathP  i  Σ[ a  A ] (suppPath i a)  X) (p .enumeration) (q .enumeration))
+      (p  q)
+  support (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z = suppPath i z
+  enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) (a , enum) = enumPath i (a , enum)
+  isPropSupport (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) z =
+    isProp→PathP {B = λ j  isProp (suppPath j z)}  j  isPropIsProp) (p .isPropSupport z) (q .isPropSupport z) i
+  isSurjectionEnumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) b =
+    isProp→PathP
+      {B = λ j   fiber (enumeration (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) j)) b ∥₁}
+       j  isPropPropTrunc)
+      (p .isSurjectionEnumeration b) (q .isSurjectionEnumeration b) i
+  isSetX (Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath) i) = isPropIsSet (p .isSetX) (q .isSetX) i
+  Iso.inv (PartialSurjection≡Iso p q) p≡q =  i  p≡q i .support) ,  i  p≡q i .enumeration)
+  Iso.rightInv (PartialSurjection≡Iso p q) p≡q = isSetPartialSurjection X isCorrectHLevel _ _ _ _ 
+  Iso.leftInv (PartialSurjection≡Iso p q) (suppPath , enumPath) = ΣPathP (refl , refl)
+
+  PartialSurjection≡ :  (p q : PartialSurjection X)  Σ[ suppPath  p .support  q .support ] PathP  i  Σ[ a  A ] (suppPath i a)  X) (p .enumeration) (q .enumeration)  p  q
+  PartialSurjection≡ p q (suppPath , enumPath) = Iso.fun (PartialSurjection≡Iso p q) (suppPath , enumPath)
+
+-- the type of partial surjections is equivalent to the type of modest assemblies on X
+module ModestSetIso (X : Type ) (isCorrectHLevel : isSet X) where
+
+  open SIP X isCorrectHLevel
+
+  {-# TERMINATING #-}
+  ModestSet→PartialSurjection : ModestSet X  PartialSurjection X
+  support (ModestSet→PartialSurjection (xs , isModestXs)) r = ∃[ x  X ] (r ⊩[ xs ] x)
+  enumeration (ModestSet→PartialSurjection (xs , isModestXs)) (r , ∃x) =
+    let
+      answer : Σ[ x  X ] (r ⊩[ xs ] x)
+      answer = PT.rec (isUniqueRealized xs isModestXs r)  t  t) ∃x
+    in fst answer
+  isPropSupport (ModestSet→PartialSurjection (xs , isModestXs)) r = isPropPropTrunc
+  isSurjectionEnumeration (ModestSet→PartialSurjection (xs , isModestXs)) x =
+    do
+      (a , a⊩x)  xs .⊩surjective x
+      return ((a ,  x , a⊩x ∣₁) , refl)
+  isSetX (ModestSet→PartialSurjection (xs , isModestXs)) = xs .isSetX
+
+  PartialSurjection→ModestSet : PartialSurjection X  ModestSet X
+  Assembly._⊩_ (fst (PartialSurjection→ModestSet surj)) r x =
+    Σ[ s  surj .support r ] surj .enumeration (r , s)  x
+  Assembly.isSetX (fst (PartialSurjection→ModestSet surj)) = surj .isSetX
+  Assembly.⊩isPropValued (fst (PartialSurjection→ModestSet surj)) a x (s , ≡x) (t , ≡x') =
+    Σ≡Prop  u  surj .isSetX (surj .enumeration (a , u)) x) (surj .isPropSupport a s t)
+  Assembly.⊩surjective (fst (PartialSurjection→ModestSet surj)) x =
+    do
+      ((a , s) , ≡x)  surj .isSurjectionEnumeration x
+      return (a , (s , ≡x))
+  snd (PartialSurjection→ModestSet surj) x y r (s , ≡x) (t , ≡x') =
+    x
+      ≡⟨ sym ≡x 
+    surj .enumeration (r , s)
+      ≡⟨ cong  s  surj .enumeration (r , s)) (surj .isPropSupport r s t) 
+    surj .enumeration (r , t)
+      ≡⟨ ≡x' 
+    y
+      
+
+  opaque
+    rightInv :  surj  ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)  surj
+    rightInv surj =
+      PartialSurjection≡
+        (ModestSet→PartialSurjection (PartialSurjection→ModestSet surj)) surj
+        (funExt supportEq ,
+        funExtDep
+          {A = λ i  Σ-syntax A (funExt supportEq i)}
+          {B = λ _ _  X}
+          {f = ModestSet→PartialSurjection (PartialSurjection→ModestSet surj) .enumeration}
+          {g = surj .enumeration}
+          λ { {r , ∃x} {s , supp} p 
+            PT.elim
+              {P = λ ∃x  fst
+                             (PT.rec
+                              (isUniqueRealized (fst (PartialSurjection→ModestSet surj))
+                               (snd (PartialSurjection→ModestSet surj)) r)
+                               t  t) ∃x)
+                           surj .enumeration (s , supp)}
+              ∃x  surj .isSetX _ _)
+              { (x , suppR , ≡x) 
+               let
+                 ∃x' = transport (sym (supportEq s)) supp
+                 r≡s : r  s
+                 r≡s = PathPΣ p .fst
+               in
+               equivFun
+                 (propTruncIdempotent≃ (surj .isSetX x (surj .enumeration (s , supp))))
+                 (do
+                   (x' , suppS , ≡x')  ∃x'
+                   return
+                     (x
+                       ≡⟨ sym ≡x 
+                     surj .enumeration (r , suppR)
+                       ≡⟨ cong (surj .enumeration) (ΣPathP (r≡s , (isProp→PathP  i  surj .isPropSupport (PathPΣ p .fst i)) suppR supp))) 
+                     surj .enumeration (s , supp)
+                       )) })
+             ∃x }) where
+          supportEq :  r  (∃[ x  X ] (Σ[ supp  surj .support r ] (surj .enumeration (r , supp)  x)))  support surj r
+          supportEq =
+               r 
+                hPropExt
+                isPropPropTrunc
+                (surj .isPropSupport r)
+                 ∃x  PT.rec (surj .isPropSupport r)  { (x , supp , ≡x)  supp }) ∃x)
+                 supp  return (surj .enumeration (r , supp) , supp , refl)))
+
+  leftInv :  mod  PartialSurjection→ModestSet (ModestSet→PartialSurjection mod)  mod
+  leftInv (asmX , isModestAsmX) =
+    Σ≡Prop
+      isPropIsModest
+      (Assembly≡ _ _
+        λ r x 
+          hPropExt
+            (isPropΣ isPropPropTrunc  ∃x  asmX .isSetX _ _))
+            (asmX .⊩isPropValued r x)
+             { (∃x , ≡x) 
+              let
+                (x' , r⊩x') = PT.rec (isUniqueRealized asmX isModestAsmX r)  t  t) ∃x
+              in subst  x'  r ⊩[ asmX ] x') ≡x r⊩x'})
+            λ r⊩x   x , r⊩x ∣₁ , refl)
+
+  IsoModestSetPartialSurjection : Iso (ModestSet X) (PartialSurjection X)
+  Iso.fun IsoModestSetPartialSurjection = ModestSet→PartialSurjection
+  Iso.inv IsoModestSetPartialSurjection = PartialSurjection→ModestSet
+  Iso.rightInv IsoModestSetPartialSurjection = rightInv 
+  Iso.leftInv IsoModestSetPartialSurjection = leftInv
+
+  ModestSet≡PartialSurjection : ModestSet X  PartialSurjection X
+  ModestSet≡PartialSurjection = isoToPath IsoModestSetPartialSurjection
+
+record PartialSurjectionMorphism {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y) : Type  where
+  no-eta-equality
+  constructor makePartialSurjectionMorphism
+  field
+    map : X  Y
+    {-
       The following "diagram" commutes
                               
       Xˢ -----------> X
@@ -229,162 +231,162 @@
       ↓              ↓
       Yˢ -----------> Y
     -}
-    isTracked : ∃[ t  A ] (∀ (a : A) (sᵃ : psX .support a)  Σ[ sᵇ  (psY .support (t  a)) ] map (psX .enumeration (a , sᵃ))  psY .enumeration ((t  a) , sᵇ))
-open PartialSurjectionMorphism
-
-unquoteDecl PartialSurjectionMorphismIsoΣ = declareRecordIsoΣ PartialSurjectionMorphismIsoΣ (quote PartialSurjectionMorphism)
-
-PartialSurjectionMorphismΣ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  Type 
-PartialSurjectionMorphismΣ {X} {Y} psX psY =
-  Σ[ f  (X  Y) ] ∃[ t  A ] ((∀ (a : A) (sᵃ : psX .support a)  Σ[ sᵇ  (psY .support (t  a)) ] f (psX .enumeration (a , sᵃ))  psY .enumeration ((t  a) , sᵇ)))
-
-isSetPartialSurjectionMorphismΣ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  isSet (PartialSurjectionMorphismΣ psX psY)
-isSetPartialSurjectionMorphismΣ {X} {Y} psX psY = isSetΣ (isSet→ (psY .isSetX))  f  isProp→isSet isPropPropTrunc)
-
-PartialSurjectionMorphismΣ≡ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  PartialSurjectionMorphism psX psY  PartialSurjectionMorphismΣ psX psY
-PartialSurjectionMorphismΣ≡ {X} {Y} psX psY = isoToPath PartialSurjectionMorphismIsoΣ
-
-isSetPartialSurjectionMorphism : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  isSet (PartialSurjectionMorphism psX psY)
-isSetPartialSurjectionMorphism {X} {Y} psX psY = subst isSet (sym (PartialSurjectionMorphismΣ≡ psX psY)) (isSetPartialSurjectionMorphismΣ psX psY)
-
--- SIP
-module MorphismSIP {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y) where
-  open PartialSurjectionMorphism
-  PartialSurjectionMorphism≡Iso :  (f g : PartialSurjectionMorphism psX psY)  Iso (f  g) (f .map  g .map)
-  Iso.fun (PartialSurjectionMorphism≡Iso f g) f≡g i = f≡g i .map
-  map (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) = fMap≡gMap i
-  isTracked (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) =
-    isProp→PathP
-      -- Agda can't infer the type B
-      {B = λ j  ∃-syntax A
-       t 
-         (a : A) (sᵃ : psX .support a) 
-         Σ-syntax (psY .support (t  a))
-          sᵇ 
-            fMap≡gMap j (psX .enumeration (a , sᵃ)) 
-            psY .enumeration (t  a , sᵇ)))}
-       j  isPropPropTrunc)
-      (f .isTracked) (g .isTracked) i
-  Iso.rightInv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap = refl
-  Iso.leftInv (PartialSurjectionMorphism≡Iso f g) f≡g = isSetPartialSurjectionMorphism psX psY f g _ _
-
-  PartialSurjectionMorphism≡ :  {f g : PartialSurjectionMorphism psX psY}  (f .map  g .map)  f  g
-  PartialSurjectionMorphism≡ {f} {g} fMap≡gMap = Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap
-
--- morphisms between partial surjections are equivalent to assembly morphisms between corresponding modest assemblies
-module
-  _
-  {X Y : Type }
-  (psX : PartialSurjection X)
-  (psY : PartialSurjection Y) where
-  open ModestSetIso 
-  open MorphismSIP psX psY
-
-  asmX = PartialSurjection→ModestSet X (psX .isSetX) psX .fst
-  isModestAsmX = PartialSurjection→ModestSet X (psX .isSetX) psX .snd
-
-  asmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .fst
-  isModestAsmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .snd
-
-  PartialSurjectionHomModestSetHomIso : Iso (AssemblyMorphism asmX asmY) (PartialSurjectionMorphism psX psY)
-  map (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) = asmHom .map
-  isTracked (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) =
-    do
-      (map~ , isTrackedMap)  asmHom .tracker
-      return
-        (map~ ,
-         λ a aSuppX 
-           let
-             worker : (map~  a) ⊩[ asmY ] (asmHom .map (psX .enumeration (a , aSuppX)))
-             worker = isTrackedMap (psX .enumeration (a , aSuppX)) a (aSuppX , refl)
-           in
-           (worker .fst) ,
-           (sym (worker .snd)))
-  AssemblyMorphism.map (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) = surjHom .map
-  AssemblyMorphism.tracker (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) =
-    do
-      (t , isTrackedMap)  surjHom .isTracked
-      return
-        (t ,
-         { x a (aSuppX , ≡x) 
-          (isTrackedMap a aSuppX .fst) ,
-          (sym (cong (surjHom .map) (sym ≡x)  isTrackedMap a aSuppX .snd)) }))
-  Iso.rightInv PartialSurjectionHomModestSetHomIso surjHom = PartialSurjectionMorphism≡ refl
-  Iso.leftInv PartialSurjectionHomModestSetHomIso asmHom = AssemblyMorphism≡ _ _ refl
-
-  PartialSurjectionHom≡ModestSetHom : AssemblyMorphism asmX asmY  PartialSurjectionMorphism psX psY
-  PartialSurjectionHom≡ModestSetHom = isoToPath PartialSurjectionHomModestSetHomIso
-
--- the category of partial surjections
-
-idPartSurjMorphism :  {X : Type }  (psX : PartialSurjection X)  PartialSurjectionMorphism psX psX
-map (idPartSurjMorphism {X} psX) x = x
-isTracked (idPartSurjMorphism {X} psX) =
-  return (Id ,  a aSuppX  (subst  r  psX .support r) (sym (Ida≡a a)) aSuppX) , (cong (psX .enumeration) (Σ≡Prop  b  psX .isPropSupport b) (sym (Ida≡a a))))))
-
-composePartSurjMorphism :
-   {X Y Z : Type } {psX : PartialSurjection X} {psY : PartialSurjection Y} {psZ : PartialSurjection Z}
-   (f : PartialSurjectionMorphism psX psY)
-   (g : PartialSurjectionMorphism psY psZ)
-   PartialSurjectionMorphism psX psZ
-map (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) x = g .map (f .map x)
-isTracked (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) =
-  do
-    (f~ , isTrackedF)  f .isTracked
-    (g~ , isTrackedG)  g .isTracked
-    let
-      realizer : Term as 1
-      realizer = ` g~ ̇ (` f~ ̇ # zero)
-    return
-      (λ* realizer ,
-       a aSuppX 
-        subst  r'  psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~  a) (isTrackedF a aSuppX .fst) .fst) ,
-       (g .map (f .map (psX .enumeration (a , aSuppX)))
-          ≡⟨ cong (g .map) (isTrackedF a aSuppX .snd) 
-        g .map (psY .enumeration (f~  a , fst (isTrackedF a aSuppX)))
-          ≡⟨ isTrackedG (f~  a) (fst (isTrackedF a aSuppX)) .snd 
-        psZ .enumeration (g~  (f~  a) , fst (isTrackedG (f~  a) (fst (isTrackedF a aSuppX))))
-          ≡⟨ cong (psZ .enumeration) (Σ≡Prop  z  psZ .isPropSupport z) (sym (λ*ComputationRule realizer a))) 
-        psZ .enumeration
-          (λ* realizer  a ,
-           subst  r'  psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~  a) (isTrackedF a aSuppX .fst) .fst))
-          )))
-
-idLPartSurjMorphism :
-   {X Y : Type }
-   {psX : PartialSurjection X}
-   {psY : PartialSurjection Y}
-   (f : PartialSurjectionMorphism psX psY)
-   composePartSurjMorphism (idPartSurjMorphism psX) f  f
-idLPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl
-
-idRPartSurjMorphism :
-   {X Y : Type }
-   {psX : PartialSurjection X}
-   {psY : PartialSurjection Y}
-   (f : PartialSurjectionMorphism psX psY)
-   composePartSurjMorphism f (idPartSurjMorphism psY)  f
-idRPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl
-
-assocComposePartSurjMorphism :
-   {X Y Z W : Type }
-   {psX : PartialSurjection X}
-   {psY : PartialSurjection Y}
-   {psZ : PartialSurjection Z}
-   {psW : PartialSurjection W}
-   (f : PartialSurjectionMorphism psX psY)
-   (g : PartialSurjectionMorphism psY psZ)
-   (h : PartialSurjectionMorphism psZ psW)
-   composePartSurjMorphism (composePartSurjMorphism f g) h  composePartSurjMorphism f (composePartSurjMorphism g h)
-assocComposePartSurjMorphism {X} {Y} {Z} {W} {psX} {psY} {psZ} {psW} f g h = MorphismSIP.PartialSurjectionMorphism≡ psX psW refl
-
-PARTSURJ : Category (ℓ-suc ) 
-Category.ob PARTSURJ = Σ[ X  Type  ] PartialSurjection X
-Category.Hom[_,_] PARTSURJ (X , surjX) (Y , surjY) = PartialSurjectionMorphism surjX surjY
-Category.id PARTSURJ {X , surjX} = idPartSurjMorphism surjX
-Category._⋆_ PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} f g = composePartSurjMorphism f g
-Category.⋆IdL PARTSURJ {X , surjX} {Y , surjY} f = idLPartSurjMorphism f
-Category.⋆IdR PARTSURJ {X , surjX} {Y , surjY} f = idRPartSurjMorphism f
-Category.⋆Assoc PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} {W , surjW} f g h = assocComposePartSurjMorphism f g h
-Category.isSetHom PARTSURJ {X , surjX} {Y , surjY} = isSetPartialSurjectionMorphism surjX surjY
+    isTracked : ∃[ t  A ] (∀ (a : A) (sᵃ : psX .support a)  Σ[ sᵇ  (psY .support (t  a)) ] map (psX .enumeration (a , sᵃ))  psY .enumeration ((t  a) , sᵇ))
+open PartialSurjectionMorphism
+
+unquoteDecl PartialSurjectionMorphismIsoΣ = declareRecordIsoΣ PartialSurjectionMorphismIsoΣ (quote PartialSurjectionMorphism)
+
+PartialSurjectionMorphismΣ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  Type 
+PartialSurjectionMorphismΣ {X} {Y} psX psY =
+  Σ[ f  (X  Y) ] ∃[ t  A ] ((∀ (a : A) (sᵃ : psX .support a)  Σ[ sᵇ  (psY .support (t  a)) ] f (psX .enumeration (a , sᵃ))  psY .enumeration ((t  a) , sᵇ)))
+
+isSetPartialSurjectionMorphismΣ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  isSet (PartialSurjectionMorphismΣ psX psY)
+isSetPartialSurjectionMorphismΣ {X} {Y} psX psY = isSetΣ (isSet→ (psY .isSetX))  f  isProp→isSet isPropPropTrunc)
+
+PartialSurjectionMorphismΣ≡ : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  PartialSurjectionMorphism psX psY  PartialSurjectionMorphismΣ psX psY
+PartialSurjectionMorphismΣ≡ {X} {Y} psX psY = isoToPath PartialSurjectionMorphismIsoΣ
+
+isSetPartialSurjectionMorphism : {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y)  isSet (PartialSurjectionMorphism psX psY)
+isSetPartialSurjectionMorphism {X} {Y} psX psY = subst isSet (sym (PartialSurjectionMorphismΣ≡ psX psY)) (isSetPartialSurjectionMorphismΣ psX psY)
+
+-- SIP
+module MorphismSIP {X Y : Type } (psX : PartialSurjection X) (psY : PartialSurjection Y) where
+  open PartialSurjectionMorphism
+  PartialSurjectionMorphism≡Iso :  (f g : PartialSurjectionMorphism psX psY)  Iso (f  g) (f .map  g .map)
+  Iso.fun (PartialSurjectionMorphism≡Iso f g) f≡g i = f≡g i .map
+  map (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) = fMap≡gMap i
+  isTracked (Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap i) =
+    isProp→PathP
+      -- Agda can't infer the type B
+      {B = λ j  ∃-syntax A
+       t 
+         (a : A) (sᵃ : psX .support a) 
+         Σ-syntax (psY .support (t  a))
+          sᵇ 
+            fMap≡gMap j (psX .enumeration (a , sᵃ)) 
+            psY .enumeration (t  a , sᵇ)))}
+       j  isPropPropTrunc)
+      (f .isTracked) (g .isTracked) i
+  Iso.rightInv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap = refl
+  Iso.leftInv (PartialSurjectionMorphism≡Iso f g) f≡g = isSetPartialSurjectionMorphism psX psY f g _ _
+
+  PartialSurjectionMorphism≡ :  {f g : PartialSurjectionMorphism psX psY}  (f .map  g .map)  f  g
+  PartialSurjectionMorphism≡ {f} {g} fMap≡gMap = Iso.inv (PartialSurjectionMorphism≡Iso f g) fMap≡gMap
+
+-- morphisms between partial surjections are equivalent to assembly morphisms between corresponding modest assemblies
+module
+  _
+  {X Y : Type }
+  (psX : PartialSurjection X)
+  (psY : PartialSurjection Y) where
+  open ModestSetIso 
+  open MorphismSIP psX psY
+
+  asmX = PartialSurjection→ModestSet X (psX .isSetX) psX .fst
+  isModestAsmX = PartialSurjection→ModestSet X (psX .isSetX) psX .snd
+
+  asmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .fst
+  isModestAsmY = PartialSurjection→ModestSet Y (psY .isSetX) psY .snd
+
+  PartialSurjectionHomModestSetHomIso : Iso (AssemblyMorphism asmX asmY) (PartialSurjectionMorphism psX psY)
+  map (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) = asmHom .map
+  isTracked (Iso.fun PartialSurjectionHomModestSetHomIso asmHom) =
+    do
+      (map~ , isTrackedMap)  asmHom .tracker
+      return
+        (map~ ,
+         λ a aSuppX 
+           let
+             worker : (map~  a) ⊩[ asmY ] (asmHom .map (psX .enumeration (a , aSuppX)))
+             worker = isTrackedMap (psX .enumeration (a , aSuppX)) a (aSuppX , refl)
+           in
+           (worker .fst) ,
+           (sym (worker .snd)))
+  AssemblyMorphism.map (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) = surjHom .map
+  AssemblyMorphism.tracker (Iso.inv PartialSurjectionHomModestSetHomIso surjHom) =
+    do
+      (t , isTrackedMap)  surjHom .isTracked
+      return
+        (t ,
+         { x a (aSuppX , ≡x) 
+          (isTrackedMap a aSuppX .fst) ,
+          (sym (cong (surjHom .map) (sym ≡x)  isTrackedMap a aSuppX .snd)) }))
+  Iso.rightInv PartialSurjectionHomModestSetHomIso surjHom = PartialSurjectionMorphism≡ refl
+  Iso.leftInv PartialSurjectionHomModestSetHomIso asmHom = AssemblyMorphism≡ _ _ refl
+
+  PartialSurjectionHom≡ModestSetHom : AssemblyMorphism asmX asmY  PartialSurjectionMorphism psX psY
+  PartialSurjectionHom≡ModestSetHom = isoToPath PartialSurjectionHomModestSetHomIso
+
+-- the category of partial surjections
+
+idPartSurjMorphism :  {X : Type }  (psX : PartialSurjection X)  PartialSurjectionMorphism psX psX
+map (idPartSurjMorphism {X} psX) x = x
+isTracked (idPartSurjMorphism {X} psX) =
+  return (Id ,  a aSuppX  (subst  r  psX .support r) (sym (Ida≡a a)) aSuppX) , (cong (psX .enumeration) (Σ≡Prop  b  psX .isPropSupport b) (sym (Ida≡a a))))))
+
+composePartSurjMorphism :
+   {X Y Z : Type } {psX : PartialSurjection X} {psY : PartialSurjection Y} {psZ : PartialSurjection Z}
+   (f : PartialSurjectionMorphism psX psY)
+   (g : PartialSurjectionMorphism psY psZ)
+   PartialSurjectionMorphism psX psZ
+map (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) x = g .map (f .map x)
+isTracked (composePartSurjMorphism {X} {Y} {Z} {psX} {psY} {psZ} f g) =
+  do
+    (f~ , isTrackedF)  f .isTracked
+    (g~ , isTrackedG)  g .isTracked
+    let
+      realizer : Term as 1
+      realizer = ` g~ ̇ (` f~ ̇ # zero)
+    return
+      (λ* realizer ,
+       a aSuppX 
+        subst  r'  psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~  a) (isTrackedF a aSuppX .fst) .fst) ,
+       (g .map (f .map (psX .enumeration (a , aSuppX)))
+          ≡⟨ cong (g .map) (isTrackedF a aSuppX .snd) 
+        g .map (psY .enumeration (f~  a , fst (isTrackedF a aSuppX)))
+          ≡⟨ isTrackedG (f~  a) (fst (isTrackedF a aSuppX)) .snd 
+        psZ .enumeration (g~  (f~  a) , fst (isTrackedG (f~  a) (fst (isTrackedF a aSuppX))))
+          ≡⟨ cong (psZ .enumeration) (Σ≡Prop  z  psZ .isPropSupport z) (sym (λ*ComputationRule realizer a))) 
+        psZ .enumeration
+          (λ* realizer  a ,
+           subst  r'  psZ .support r') (sym (λ*ComputationRule realizer a)) (isTrackedG (f~  a) (isTrackedF a aSuppX .fst) .fst))
+          )))
+
+idLPartSurjMorphism :
+   {X Y : Type }
+   {psX : PartialSurjection X}
+   {psY : PartialSurjection Y}
+   (f : PartialSurjectionMorphism psX psY)
+   composePartSurjMorphism (idPartSurjMorphism psX) f  f
+idLPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl
+
+idRPartSurjMorphism :
+   {X Y : Type }
+   {psX : PartialSurjection X}
+   {psY : PartialSurjection Y}
+   (f : PartialSurjectionMorphism psX psY)
+   composePartSurjMorphism f (idPartSurjMorphism psY)  f
+idRPartSurjMorphism {X} {Y} {psX} {psY} f = MorphismSIP.PartialSurjectionMorphism≡ psX psY refl
+
+assocComposePartSurjMorphism :
+   {X Y Z W : Type }
+   {psX : PartialSurjection X}
+   {psY : PartialSurjection Y}
+   {psZ : PartialSurjection Z}
+   {psW : PartialSurjection W}
+   (f : PartialSurjectionMorphism psX psY)
+   (g : PartialSurjectionMorphism psY psZ)
+   (h : PartialSurjectionMorphism psZ psW)
+   composePartSurjMorphism (composePartSurjMorphism f g) h  composePartSurjMorphism f (composePartSurjMorphism g h)
+assocComposePartSurjMorphism {X} {Y} {Z} {W} {psX} {psY} {psZ} {psW} f g h = MorphismSIP.PartialSurjectionMorphism≡ psX psW refl
+
+PARTSURJ : Category (ℓ-suc ) 
+Category.ob PARTSURJ = Σ[ X  Type  ] PartialSurjection X
+Category.Hom[_,_] PARTSURJ (X , surjX) (Y , surjY) = PartialSurjectionMorphism surjX surjY
+Category.id PARTSURJ {X , surjX} = idPartSurjMorphism surjX
+Category._⋆_ PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} f g = composePartSurjMorphism f g
+Category.⋆IdL PARTSURJ {X , surjX} {Y , surjY} f = idLPartSurjMorphism f
+Category.⋆IdR PARTSURJ {X , surjX} {Y , surjY} f = idRPartSurjMorphism f
+Category.⋆Assoc PARTSURJ {X , surjX} {Y , surjY} {Z , surjZ} {W , surjW} f g h = assocComposePartSurjMorphism f g h
+Category.isSetHom PARTSURJ {X , surjX} {Y , surjY} = isSetPartialSurjectionMorphism surjX surjY
 
\ No newline at end of file diff --git a/docs/Realizability.Modest.SubQuotientCanonicalPERIso.html b/docs/Realizability.Modest.SubQuotientCanonicalPERIso.html index 428ead4..c07c4c1 100644 --- a/docs/Realizability.Modest.SubQuotientCanonicalPERIso.html +++ b/docs/Realizability.Modest.SubQuotientCanonicalPERIso.html @@ -1,149 +1,154 @@ -Realizability.Modest.SubQuotientCanonicalPERIso
open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Function
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Powerset
-open import Cubical.Foundations.Path
-open import Cubical.Foundations.Structure using (⟨_⟩; str)
-open import Cubical.Data.Sigma
-open import Cubical.Data.FinData
-open import Cubical.Data.Unit
-open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.HITs.SetQuotients as SQ
-open import Cubical.Reflection.RecordEquiv
-open import Cubical.Categories.Category
-open import Cubical.Categories.Displayed.Base
-open import Cubical.Categories.Displayed.Reasoning
-open import Cubical.Categories.Limits.Pullback
-open import Cubical.Categories.Functor hiding (Id)
-open import Cubical.Categories.Constructions.Slice
-open import Categories.CartesianMorphism
-open import Categories.GenericObject
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure
-open import Realizability.PropResizing
-
-module Realizability.Modest.SubQuotientCanonicalPERIso {} {A : Type } (ca : CombinatoryAlgebra A) where
-
-open import Realizability.Assembly.Base ca
-open import Realizability.Assembly.Morphism ca
-open import Realizability.Assembly.Terminal ca
-open import Realizability.Assembly.SetsReflectiveSubcategory ca
-open import Realizability.Modest.Base ca
-open import Realizability.Modest.UniformFamily ca
-open import Realizability.Modest.CanonicalPER ca
-open import Realizability.PERs.PER ca
-open import Realizability.PERs.SubQuotient ca
-
-open Assembly
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Contravariant UNIMOD
-open UniformFamily
-open DisplayedUFamMap
-
-module
-  _ {X : Type }
-  (asmX : Assembly X)
-  (isModestAsmX : isModest asmX) where
-
-  theCanonicalPER : PER
-  theCanonicalPER = canonicalPER asmX isModestAsmX
-
-  theSubQuotient : Assembly (subQuotient theCanonicalPER)
-  theSubQuotient = subQuotientAssembly theCanonicalPER
-
-  invert : AssemblyMorphism theSubQuotient asmX
-  AssemblyMorphism.map invert sq = SQ.rec (asmX .isSetX) reprAction reprActionCoh sq module Invert where
-
-    reprAction : Σ[ a  A ] (a ~[ theCanonicalPER ] a)  X
-    reprAction (a , x , a⊩x , _) = x
-
-    reprActionCoh :  a b a~b  reprAction a  reprAction b
-    reprActionCoh (a , x , a⊩x , _) (b , x' , b⊩x' , _) (x'' , a⊩x'' , b⊩x'') =
-      x
-        ≡⟨ isModestAsmX x x'' a a⊩x a⊩x'' 
-      x''
-        ≡⟨ isModestAsmX x'' x' b b⊩x'' b⊩x' 
-      x'
-        
-  AssemblyMorphism.tracker invert = return (Id ,  sq a a⊩sq  goal sq a a⊩sq)) where
-    realizability : (sq : subQuotient theCanonicalPER)  (a : A)  a ⊩[ theSubQuotient ] sq  a ⊩[ asmX ] (invert .map sq)
-    realizability sq a a⊩sq =
-      SQ.elimProp
-        {P = motive}
-        isPropMotive
-        elemMotive
-        sq a a⊩sq where
-
-      motive : (sq : subQuotient theCanonicalPER)  Type _
-      motive sq =  (a : A)  a ⊩[ theSubQuotient ] sq  a ⊩[ asmX ] (invert .map sq)
-
-      isPropMotive :  sq  isProp (motive sq)
-      isPropMotive sq = isPropΠ2 λ a a⊩sq  asmX .⊩isPropValued _ _
-
-      elemMotive : (x : domain theCanonicalPER)  motive [ x ]
-      elemMotive (r , x , r⊩x , _) a (x' , a⊩x' , r⊩x') = subst (a ⊩[ asmX ]_) (isModestAsmX x' x r r⊩x' r⊩x) a⊩x'
-
-    goal : (sq : subQuotient theCanonicalPER)  (a : A)  a ⊩[ theSubQuotient ] sq  (Id  a) ⊩[ asmX ] (invert .map sq)
-    goal sq a a⊩sq = subst (_⊩[ asmX ] _) (sym (Ida≡a a)) (realizability sq a a⊩sq)
-
-  forward : AssemblyMorphism asmX theSubQuotient
-  AssemblyMorphism.map forward x = subquot module Forward where
-    mainMap : Σ[ a  A ] (a ⊩[ asmX ] x)  subQuotient theCanonicalPER
-    mainMap (a , a⊩x) = [ a , x , a⊩x , a⊩x ]
+Realizability.Modest.SubQuotientCanonicalPERIso
--  This module shows that any modest set M is isomorphic to the subquotient of the canonical PER of M.
+--  Effectively, this shows that the subquotient functor is **split essentially surjective** on objects.
+--  Since the subquotient functor is fully faithful, this implies that it is an equivalence of categories.
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Unit
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.Categories.Category
+open import Cubical.Categories.Displayed.Base
+open import Cubical.Categories.Displayed.Reasoning
+open import Cubical.Categories.Limits.Pullback
+open import Cubical.Categories.Functor hiding (Id)
+open import Cubical.Categories.Constructions.Slice
+open import Cubical.Categories.Equivalence
+open import Cubical.Categories.NaturalTransformation
+open import Categories.CartesianMorphism
+open import Categories.GenericObject
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
+open import Realizability.PropResizing
+
+module Realizability.Modest.SubQuotientCanonicalPERIso {} {A : Type } (ca : CombinatoryAlgebra A) where
+
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.Assembly.Terminal ca
+open import Realizability.Assembly.SetsReflectiveSubcategory ca
+open import Realizability.Modest.Base ca
+open import Realizability.Modest.UniformFamily ca
+open import Realizability.Modest.CanonicalPER ca
+open import Realizability.PERs.PER ca
+open import Realizability.PERs.SubQuotient ca
+
+open Assembly
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Contravariant UNIMOD
+open UniformFamily
+open DisplayedUFamMap
+
+module
+  _ {X : Type }
+  (asmX : Assembly X)
+  (isModestAsmX : isModest asmX) where
+
+  theCanonicalPER : PER
+  theCanonicalPER = canonicalPER asmX isModestAsmX
+
+  theSubQuotient : Assembly (subQuotient theCanonicalPER)
+  theSubQuotient = subQuotientAssembly theCanonicalPER
+
+  invert : AssemblyMorphism theSubQuotient asmX
+  AssemblyMorphism.map invert sq = SQ.rec (asmX .isSetX) reprAction reprActionCoh sq module Invert where
+
+    reprAction : Σ[ a  A ] (a ~[ theCanonicalPER ] a)  X
+    reprAction (a , x , a⊩x , _) = x
+
+    reprActionCoh :  a b a~b  reprAction a  reprAction b
+    reprActionCoh (a , x , a⊩x , _) (b , x' , b⊩x' , _) (x'' , a⊩x'' , b⊩x'') =
+      x
+        ≡⟨ isModestAsmX x x'' a a⊩x a⊩x'' 
+      x''
+        ≡⟨ isModestAsmX x'' x' b b⊩x'' b⊩x' 
+      x'
+        
+  AssemblyMorphism.tracker invert = return (Id ,  sq a a⊩sq  goal sq a a⊩sq)) where
+    realizability : (sq : subQuotient theCanonicalPER)  (a : A)  a ⊩[ theSubQuotient ] sq  a ⊩[ asmX ] (invert .map sq)
+    realizability sq a a⊩sq =
+      SQ.elimProp
+        {P = motive}
+        isPropMotive
+        elemMotive
+        sq a a⊩sq where
+
+      motive : (sq : subQuotient theCanonicalPER)  Type _
+      motive sq =  (a : A)  a ⊩[ theSubQuotient ] sq  a ⊩[ asmX ] (invert .map sq)
+
+      isPropMotive :  sq  isProp (motive sq)
+      isPropMotive sq = isPropΠ2 λ a a⊩sq  asmX .⊩isPropValued _ _
+
+      elemMotive : (x : domain theCanonicalPER)  motive [ x ]
+      elemMotive (r , x , r⊩x , _) a (x' , a⊩x' , r⊩x') = subst (a ⊩[ asmX ]_) (isModestAsmX x' x r r⊩x' r⊩x) a⊩x'
+
+    goal : (sq : subQuotient theCanonicalPER)  (a : A)  a ⊩[ theSubQuotient ] sq  (Id  a) ⊩[ asmX ] (invert .map sq)
+    goal sq a a⊩sq = subst (_⊩[ asmX ] _) (sym (Ida≡a a)) (realizability sq a a⊩sq)
+
+  forward : AssemblyMorphism asmX theSubQuotient
+  AssemblyMorphism.map forward x = subquot module Forward where
+    mainMap : Σ[ a  A ] (a ⊩[ asmX ] x)  subQuotient theCanonicalPER
+    mainMap (a , a⊩x) = [ a , x , a⊩x , a⊩x ]
  
-    mainMap2Constant : 2-Constant mainMap
-    mainMap2Constant (a , a⊩x) (b , b⊩x) = eq/ _ _ (x , a⊩x , b⊩x)
-
-    subquot : subQuotient theCanonicalPER
-    subquot = PT.rec→Set squash/ mainMap mainMap2Constant (asmX .⊩surjective x)
-  AssemblyMorphism.tracker forward =
-    return
-      (Id ,
-       x a a⊩x 
-        PT.elim
-          {P = λ surj  (Id  a) ⊩[ theSubQuotient ] (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)}
-           surj  theSubQuotient .⊩isPropValued (Id  a) (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj))
-           { (b , b⊩x)  x , subst (_⊩[ asmX ] x) (sym (Ida≡a a)) a⊩x , b⊩x })
-          (asmX .⊩surjective x)))
-
-  subQuotientCanonicalIso : CatIso MOD (X , asmX , isModestAsmX) (subQuotient theCanonicalPER , theSubQuotient , isModestSubQuotientAssembly theCanonicalPER)
-  fst subQuotientCanonicalIso = forward
-  isIso.inv (snd subQuotientCanonicalIso) = invert
-  isIso.sec (snd subQuotientCanonicalIso) = goal where
-    opaque
-      pointwise :  sq  (invert  forward) .map sq  sq
-      pointwise sq =
-        SQ.elimProp
-           sq  squash/ (forward .map (invert .map sq)) sq)
-           { d@(a , x , a⊩x , a⊩'x) 
-            PT.elim
-              {P = λ surj  PT.rec→Set squash/ (Forward.mainMap (Invert.reprAction [ d ] d)) (Forward.mainMap2Constant (Invert.reprAction [ d ] d)) surj  [ d ]}
-               surj  squash/ _ _)
-               { (b , b⊩x)  eq/ _ _ (x , b⊩x , a⊩x) })
-              (asmX .⊩surjective x) })
-          sq
-
-    goal : invert  forward  identityMorphism theSubQuotient
-    goal = AssemblyMorphism≡ _ _ (funExt pointwise)
-  isIso.ret (snd subQuotientCanonicalIso) = goal where
-    opaque
-      pointwise :  x  (forward  invert) .map x  x
-      pointwise x =
-        PT.elim
-          {P =
-            λ surj 
-              invert .map
-                (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)
-               x}
-           surj  asmX .isSetX _ _)
-           { (a , a⊩x)  refl })
-          (asmX .⊩surjective x)
-
-    goal : forward  invert  identityMorphism asmX
-    goal = AssemblyMorphism≡ _ _ (funExt pointwise)
+    mainMap2Constant : 2-Constant mainMap
+    mainMap2Constant (a , a⊩x) (b , b⊩x) = eq/ _ _ (x , a⊩x , b⊩x)
+
+    subquot : subQuotient theCanonicalPER
+    subquot = PT.rec→Set squash/ mainMap mainMap2Constant (asmX .⊩surjective x)
+  AssemblyMorphism.tracker forward =
+    return
+      (Id ,
+       x a a⊩x 
+        PT.elim
+          {P = λ surj  (Id  a) ⊩[ theSubQuotient ] (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)}
+           surj  theSubQuotient .⊩isPropValued (Id  a) (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj))
+           { (b , b⊩x)  x , subst (_⊩[ asmX ] x) (sym (Ida≡a a)) a⊩x , b⊩x })
+          (asmX .⊩surjective x)))
+
+  subQuotientCanonicalIso : CatIso MOD (X , asmX , isModestAsmX) (subQuotient theCanonicalPER , theSubQuotient , isModestSubQuotientAssembly theCanonicalPER)
+  fst subQuotientCanonicalIso = forward
+  isIso.inv (snd subQuotientCanonicalIso) = invert
+  isIso.sec (snd subQuotientCanonicalIso) = goal where
+    opaque
+      pointwise :  sq  (invert  forward) .map sq  sq
+      pointwise sq =
+        SQ.elimProp
+           sq  squash/ (forward .map (invert .map sq)) sq)
+           { d@(a , x , a⊩x , a⊩'x) 
+            PT.elim
+              {P = λ surj  PT.rec→Set squash/ (Forward.mainMap (Invert.reprAction [ d ] d)) (Forward.mainMap2Constant (Invert.reprAction [ d ] d)) surj  [ d ]}
+               surj  squash/ _ _)
+               { (b , b⊩x)  eq/ _ _ (x , b⊩x , a⊩x) })
+              (asmX .⊩surjective x) })
+          sq
+
+    goal : invert  forward  identityMorphism theSubQuotient
+    goal = AssemblyMorphism≡ _ _ (funExt pointwise)
+  isIso.ret (snd subQuotientCanonicalIso) = goal where
+    opaque
+      pointwise :  x  (forward  invert) .map x  x
+      pointwise x =
+        PT.elim
+          {P =
+            λ surj 
+              invert .map
+                (PT.rec→Set squash/ (Forward.mainMap x) (Forward.mainMap2Constant x) surj)
+               x}
+           surj  asmX .isSetX _ _)
+           { (a , a⊩x)  refl })
+          (asmX .⊩surjective x)
+
+    goal : forward  invert  identityMorphism asmX
+    goal = AssemblyMorphism≡ _ _ (funExt pointwise)
 
\ No newline at end of file diff --git a/docs/Realizability.Modest.UnresizedGeneric.html b/docs/Realizability.Modest.UnresizedGeneric.html index 4fef522..4e23a82 100644 --- a/docs/Realizability.Modest.UnresizedGeneric.html +++ b/docs/Realizability.Modest.UnresizedGeneric.html @@ -52,9 +52,9 @@ (M : UniformFamily asmX) where theCanonicalPER : x PER - theCanonicalPER x = canonicalPER (M . assemblies x) (M .isModestFamily x) + theCanonicalPER x = canonicalPER (M . assemblies x) (M .isModestFamily x) - elimRealizerForMx : (x : X) (Mx : M .carriers x) Σ[ a A ] (a ⊩[ M .assemblies x ] Mx) subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) + elimRealizerForMx : (x : X) (Mx : M .carriers x) Σ[ a A ] (a ⊩[ M .assemblies x ] Mx) subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x)) elimRealizerForMx x Mx (a , a⊩Mx) = [ a , Mx , a⊩Mx , a⊩Mx ] opaque @@ -68,8 +68,8 @@ mainMapType : Type _ mainMapType = (x : X) (Mx : M .carriers x) - Σ[ out (subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x))) ] - (∀ (a : A) a ⊩[ asmX ] x (b : A) b ⊩[ M .assemblies x ] Mx (λ*2 (# zero) a b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] out) + Σ[ out (subQuotient (canonicalPER (M .assemblies x) (M .isModestFamily x))) ] + (∀ (a : A) a ⊩[ asmX ] x (b : A) b ⊩[ M .assemblies x ] Mx (λ*2 (# zero) a b) ⊩[ subQuotientAssembly (theCanonicalPER x) ] out) opaque mainMap : mainMapType @@ -83,12 +83,12 @@ isSet→ (isProp→isSet (str - (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero) a b) out))))) + (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero) a b) out))))) ((λ { (c , c⊩Mx) (elimRealizerForMx x Mx (c , c⊩Mx)) , a a⊩x b b⊩Mx - subst (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx x Mx (c , c⊩Mx))) (sym (λ*2ComputationRule (# zero) a b)) (Mx , b⊩Mx , c⊩Mx)) })) + subst (_⊩[ subQuotientAssembly (theCanonicalPER x) ] (elimRealizerForMx x Mx (c , c⊩Mx))) (sym (λ*2ComputationRule (# zero) a b)) (Mx , b⊩Mx , c⊩Mx)) })) { (a , a⊩Mx) (b , b⊩Mx) - Σ≡Prop out isPropΠ4 λ a a⊩x b b⊩Mx str (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero) a b) out)) (elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx)) }) + Σ≡Prop out isPropΠ4 λ a a⊩x b b⊩Mx str (subQuotientRealizability (theCanonicalPER x) (λ*2 (# zero) a b) out)) (elimRealizerForMx2Constant x Mx (a , a⊩Mx) (b , b⊩Mx)) }) (M .assemblies x .⊩surjective Mx)
\ No newline at end of file diff --git a/docs/Realizability.PERs.SubQuotient.html b/docs/Realizability.PERs.SubQuotient.html index e244cc6..1b81bf4 100644 --- a/docs/Realizability.PERs.SubQuotient.html +++ b/docs/Realizability.PERs.SubQuotient.html @@ -1,271 +1,274 @@ -Realizability.PERs.SubQuotient
open import Realizability.ApplicativeStructure
-open import Realizability.CombinatoryAlgebra
-open import Realizability.PropResizing
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Structure using (⟨_⟩; str)
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Univalence
-open import Cubical.Foundations.Powerset
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Path
-open import Cubical.Foundations.Function
-open import Cubical.Functions.FunExtEquiv
-open import Cubical.Functions.Embedding
-open import Cubical.Functions.Surjection
-open import Cubical.Relation.Binary
-open import Cubical.Data.Sigma
-open import Cubical.Data.FinData
-open import Cubical.Data.Vec
-open import Cubical.Reflection.RecordEquiv
-open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.HITs.SetQuotients as SQ
-open import Cubical.Categories.Category
-open import Cubical.Categories.Functor hiding (Id)
+Realizability.PERs.SubQuotient
--  The subquotient construction embeds the category PER into the category MOD of modest sets.
+--  It turns out to be an equivalence of categories! In this module, we merely define it and show
+--  that it is an embedding of categories --- is full and faithful.
+open import Realizability.ApplicativeStructure
+open import Realizability.CombinatoryAlgebra
+open import Realizability.PropResizing
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Structure using (⟨_⟩; str)
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Powerset
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Path
+open import Cubical.Foundations.Function
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Functions.Embedding
+open import Cubical.Functions.Surjection
+open import Cubical.Relation.Binary
+open import Cubical.Data.Sigma
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Reflection.RecordEquiv
+open import Cubical.HITs.PropositionalTruncation as PT hiding (map)
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.HITs.SetQuotients as SQ
+open import Cubical.Categories.Category
+open import Cubical.Categories.Functor hiding (Id)
 
-module Realizability.PERs.SubQuotient
-  {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.PERs.SubQuotient
+  {} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open import Realizability.Assembly.Base ca
-open import Realizability.Assembly.Morphism ca
-open import Realizability.PERs.PER ca
-open import Realizability.Modest.Base ca
+open import Realizability.Assembly.Base ca
+open import Realizability.Assembly.Morphism ca
+open import Realizability.PERs.PER ca
+open import Realizability.Modest.Base ca
 
-open CombinatoryAlgebra ca
-open Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open CombinatoryAlgebra ca
+open Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-module _
-  (per : PER) where
+module _
+  (per : PER) where
 
-  domain : Type 
-  domain = Σ[ a  A ] (per .PER.relation a a)
+  domain : Type 
+  domain = Σ[ a  A ] (per .PER.relation a a)
 
-  subQuotient : Type 
-  subQuotient = domain / λ { (a , _) (b , _)  per .PER.relation a b }
+  subQuotient : Type 
+  subQuotient = domain / λ { (a , _) (b , _)  per .PER.relation a b }
 
-  subQuotientRealizability : A  subQuotient  hProp 
-  subQuotientRealizability r [a] =
-    SQ.rec
-      isSetHProp
-       { (a , a~a)  r ~[ per ] a , isProp~ r per a })
-       { (a , a~a) (b , b~b) a~b 
-        Σ≡Prop
-           x  isPropIsProp)
-          (hPropExt
-            (isProp~ r per a)
-            (isProp~ r per b)
-             r~a  PER.isTransitive per r a b r~a a~b)
-             r~b  PER.isTransitive per r b a r~b (PER.isSymmetric per a b a~b))) })
-      [a]
+  subQuotientRealizability : A  subQuotient  hProp 
+  subQuotientRealizability r [a] =
+    SQ.rec
+      isSetHProp
+       { (a , a~a)  r ~[ per ] a , isProp~ r per a })
+       { (a , a~a) (b , b~b) a~b 
+        Σ≡Prop
+           x  isPropIsProp)
+          (hPropExt
+            (isProp~ r per a)
+            (isProp~ r per b)
+             r~a  PER.isTransitive per r a b r~a a~b)
+             r~b  PER.isTransitive per r b a r~b (PER.isSymmetric per a b a~b))) })
+      [a]
       
   
-  subQuotientAssembly : Assembly subQuotient
-  Assembly._⊩_ subQuotientAssembly r [a] =  subQuotientRealizability r [a] 
-  Assembly.isSetX subQuotientAssembly = squash/
-  Assembly.⊩isPropValued subQuotientAssembly r [a] = str (subQuotientRealizability r [a])
-  Assembly.⊩surjective subQuotientAssembly [a] =
-    SQ.elimProp
-      {P = λ [a]  ∃[ r  A ]  subQuotientRealizability r [a] }
-       [a]  isPropPropTrunc)
-       { (a , a~a)  return (a , a~a) })
-      [a]
+  subQuotientAssembly : Assembly subQuotient
+  Assembly._⊩_ subQuotientAssembly r [a] =  subQuotientRealizability r [a] 
+  Assembly.isSetX subQuotientAssembly = squash/
+  Assembly.⊩isPropValued subQuotientAssembly r [a] = str (subQuotientRealizability r [a])
+  Assembly.⊩surjective subQuotientAssembly [a] =
+    SQ.elimProp
+      {P = λ [a]  ∃[ r  A ]  subQuotientRealizability r [a] }
+       [a]  isPropPropTrunc)
+       { (a , a~a)  return (a , a~a) })
+      [a]
       
-  isModestSubQuotientAssembly : isModest subQuotientAssembly
-  isModestSubQuotientAssembly x y a a⊩x a⊩y =
-    SQ.elimProp2
-      {P = λ x y  motive x y}
-      isPropMotive
-       { (x , x~x) (y , y~y) a a~x a~y 
-        eq/ (x , x~x) (y , y~y) (PER.isTransitive per x a y (PER.isSymmetric per a x a~x) a~y) })
-      x y
-      a a⊩x a⊩y where
-        motive :  (x y : subQuotient)  Type 
-        motive x y =  (a : A) (a⊩x : a ⊩[ subQuotientAssembly ] x) (a⊩y : a ⊩[ subQuotientAssembly ] y)  x  y
+  isModestSubQuotientAssembly : isModest subQuotientAssembly
+  isModestSubQuotientAssembly x y a a⊩x a⊩y =
+    SQ.elimProp2
+      {P = λ x y  motive x y}
+      isPropMotive
+       { (x , x~x) (y , y~y) a a~x a~y 
+        eq/ (x , x~x) (y , y~y) (PER.isTransitive per x a y (PER.isSymmetric per a x a~x) a~y) })
+      x y
+      a a⊩x a⊩y where
+        motive :  (x y : subQuotient)  Type 
+        motive x y =  (a : A) (a⊩x : a ⊩[ subQuotientAssembly ] x) (a⊩y : a ⊩[ subQuotientAssembly ] y)  x  y
 
-        isPropMotive :  x y  isProp (motive x y)
-        isPropMotive x y = isPropΠ3 λ _ _ _  squash/ x y
+        isPropMotive :  x y  isProp (motive x y)
+        isPropMotive x y = isPropΠ3 λ _ _ _  squash/ x y
 
-module _ (R S : PER) (f : perMorphism R S) where
+module _ (R S : PER) (f : perMorphism R S) where
   
-  subQuotientAssemblyMorphism : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)
-  subQuotientAssemblyMorphism =
-    SQ.rec
-      (isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S))
-      mainMap
-      mainMapCoherence
-      f where
+  subQuotientAssemblyMorphism : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)
+  subQuotientAssemblyMorphism =
+    SQ.rec
+      (isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S))
+      mainMap
+      mainMapCoherence
+      f where
 
-      mainMap : perTracker R S  AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)
-      AssemblyMorphism.map (mainMap (f , fIsTracker)) sqR =
-        SQ.rec
-          squash/
-          mainMapRepr
-          mainMapReprCoherence
-          sqR module MainMapDefn where
-            mainMapRepr : domain R  subQuotient S
-            mainMapRepr (r , r~r) = [ f  r , fIsTracker r r r~r ]
+      mainMap : perTracker R S  AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)
+      AssemblyMorphism.map (mainMap (f , fIsTracker)) sqR =
+        SQ.rec
+          squash/
+          mainMapRepr
+          mainMapReprCoherence
+          sqR module MainMapDefn where
+            mainMapRepr : domain R  subQuotient S
+            mainMapRepr (r , r~r) = [ f  r , fIsTracker r r r~r ]
 
-            mainMapReprCoherence : (a b : domain R)  R .PER.relation (a .fst) (b .fst)  mainMapRepr a  mainMapRepr b
-            mainMapReprCoherence (a , a~a) (b , b~b) a~b = eq/ _ _ (fIsTracker a b a~b)
+            mainMapReprCoherence : (a b : domain R)  R .PER.relation (a .fst) (b .fst)  mainMapRepr a  mainMapRepr b
+            mainMapReprCoherence (a , a~a) (b , b~b) a~b = eq/ _ _ (fIsTracker a b a~b)
  
-      AssemblyMorphism.tracker (mainMap (f , fIsTracker)) =
-        do
-          return
-            (f ,
-             sqR s s⊩sqR 
-              SQ.elimProp
-                {P =
-                  λ sqR
-                    (s : A)
-                   s ⊩[ subQuotientAssembly R ] sqR
-                    subQuotientRealizability S (f  s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR) }
-                 sqR  isPropΠ2 λ s s⊩sqR  str (subQuotientRealizability S (f  s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR)))
-                 { (a , a~a) s s~a  fIsTracker s a s~a })
-                sqR s s⊩sqR))
+      AssemblyMorphism.tracker (mainMap (f , fIsTracker)) =
+        do
+          return
+            (f ,
+             sqR s s⊩sqR 
+              SQ.elimProp
+                {P =
+                  λ sqR
+                    (s : A)
+                   s ⊩[ subQuotientAssembly R ] sqR
+                    subQuotientRealizability S (f  s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR) }
+                 sqR  isPropΠ2 λ s s⊩sqR  str (subQuotientRealizability S (f  s) (SQ.rec squash/ (MainMapDefn.mainMapRepr f fIsTracker sqR) (MainMapDefn.mainMapReprCoherence f fIsTracker sqR) sqR)))
+                 { (a , a~a) s s~a  fIsTracker s a s~a })
+                sqR s s⊩sqR))
 
-      mainMapCoherence : (a b : perTracker R S)  isEquivTracker R S a b  mainMap a  mainMap b
-      mainMapCoherence (a , a~a) (b , b~b) a~b =
-        AssemblyMorphism≡ _ _
-          (funExt
-            λ sq 
-              SQ.elimProp
-                {P =
-                  λ sq 
-                    SQ.rec
-                      squash/
-                      (MainMapDefn.mainMapRepr a a~a sq)
-                      (MainMapDefn.mainMapReprCoherence a a~a sq) sq
-                    
-                    SQ.rec
-                      squash/
-                      (MainMapDefn.mainMapRepr b b~b sq)
-                      (MainMapDefn.mainMapReprCoherence b b~b sq) sq}
-                 sq  squash/ _ _)
-                 { (r , r~r)  eq/ _ _ (a~b r r~r) })
-                sq)
+      mainMapCoherence : (a b : perTracker R S)  isEquivTracker R S a b  mainMap a  mainMap b
+      mainMapCoherence (a , a~a) (b , b~b) a~b =
+        AssemblyMorphism≡ _ _
+          (funExt
+            λ sq 
+              SQ.elimProp
+                {P =
+                  λ sq 
+                    SQ.rec
+                      squash/
+                      (MainMapDefn.mainMapRepr a a~a sq)
+                      (MainMapDefn.mainMapReprCoherence a a~a sq) sq
+                    
+                    SQ.rec
+                      squash/
+                      (MainMapDefn.mainMapRepr b b~b sq)
+                      (MainMapDefn.mainMapReprCoherence b b~b sq) sq}
+                 sq  squash/ _ _)
+                 { (r , r~r)  eq/ _ _ (a~b r r~r) })
+                sq)
     
-module _ (R S : PER) (f : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)) where
-  subQuotientAssemblyMorphism→perMorphism : perMorphism R S
-  subQuotientAssemblyMorphism→perMorphism =
-    PT.rec→Set squash/ mainMap mainMap2Constant (f .tracker) module InverseDefinition where
-      isSQTracker : A  Type 
-      isSQTracker t =  (q : subQuotient R) (a : A)  a ⊩[ subQuotientAssembly R ] q   subQuotientRealizability S (t  a) (f .AssemblyMorphism.map q) 
-      -- 🤢🤮
-      mainMap : Σ[ t  A ] (isSQTracker t)  perMorphism R S
-      mainMap (t , t⊩f) =
-        [ t ,
-           r r' r~r' 
-            let
-              r~r : r ~[ R ] r
-              r~r = PER.isTransitive R r r' r r~r' (PER.isSymmetric R r r' r~r')
+module _ (R S : PER) (f : AssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S)) where
+  subQuotientAssemblyMorphism→perMorphism : perMorphism R S
+  subQuotientAssemblyMorphism→perMorphism =
+    PT.rec→Set squash/ mainMap mainMap2Constant (f .tracker) module InverseDefinition where
+      isSQTracker : A  Type 
+      isSQTracker t =  (q : subQuotient R) (a : A)  a ⊩[ subQuotientAssembly R ] q   subQuotientRealizability S (t  a) (f .AssemblyMorphism.map q) 
 
-              r'~r' : r' ~[ R ] r'
-              r'~r' = PER.isTransitive R r' r r' (PER.isSymmetric R r r' r~r') r~r'
-            in
-            SQ.elimProp
-              {P = λ q   (t : A)   subQuotientRealizability S (t  r) q    subQuotientRealizability S (t  r') q   (t  r) ~[ S ] (t  r')}
-               q  isPropΠ3 λ t _ _  isProp~ (t  r) S (t  r'))
-               { (s , s~s) t tr~s tr'~s  PER.isTransitive S (t  r) s (t  r') tr~s (PER.isSymmetric S (t  r') s tr'~s) })
-              (f .AssemblyMorphism.map [ (r , r~r) ])
-              t
-              (t⊩f [ (r , r~r) ] r r~r)
-              (subst  eq   subQuotientRealizability S (t  r') (f .AssemblyMorphism.map eq) ) (eq/ _ _ (PER.isSymmetric R r r' r~r')) (t⊩f [ (r' , r'~r') ] r' r'~r'))) ]
+      mainMap : Σ[ t  A ] (isSQTracker t)  perMorphism R S
+      mainMap (t , t⊩f) =
+        [ t ,
+           r r' r~r' 
+            let
+              r~r : r ~[ R ] r
+              r~r = PER.isTransitive R r r' r r~r' (PER.isSymmetric R r r' r~r')
 
-      mainMap2Constant : 2-Constant mainMap
-      mainMap2Constant (t , t⊩f) (t' , t'⊩f) =
-        eq/ _ _
-          λ r r~r 
-            SQ.elimProp
-              {P = λ q   subQuotientRealizability S (t  r) q    subQuotientRealizability S (t'  r) q   (t  r) ~[ S ] (t'  r)}
-               q  isPropΠ2 λ _ _  isProp~ (t  r) S (t'  r))
-               { (s , s~s) tr~s t'r~s  PER.isTransitive S (t  r) s (t'  r) tr~s (PER.isSymmetric S (t'  r) s t'r~s) })
-              (f .AssemblyMorphism.map [ (r , r~r) ])
-              (t⊩f [ (r , r~r) ] r r~r)
-              (t'⊩f [ (r , r~r) ] r r~r)
+              r'~r' : r' ~[ R ] r'
+              r'~r' = PER.isTransitive R r' r r' (PER.isSymmetric R r r' r~r') r~r'
+            in
+            SQ.elimProp
+              {P = λ q   (t : A)   subQuotientRealizability S (t  r) q    subQuotientRealizability S (t  r') q   (t  r) ~[ S ] (t  r')}
+               q  isPropΠ3 λ t _ _  isProp~ (t  r) S (t  r'))
+               { (s , s~s) t tr~s tr'~s  PER.isTransitive S (t  r) s (t  r') tr~s (PER.isSymmetric S (t  r') s tr'~s) })
+              (f .AssemblyMorphism.map [ (r , r~r) ])
+              t
+              (t⊩f [ (r , r~r) ] r r~r)
+              (subst  eq   subQuotientRealizability S (t  r') (f .AssemblyMorphism.map eq) ) (eq/ _ _ (PER.isSymmetric R r r' r~r')) (t⊩f [ (r' , r'~r') ] r' r'~r'))) ]
 
-subQuotientModestSet : PER  MOD .Category.ob
-subQuotientModestSet R = subQuotient R , subQuotientAssembly R , isModestSubQuotientAssembly R
+      mainMap2Constant : 2-Constant mainMap
+      mainMap2Constant (t , t⊩f) (t' , t'⊩f) =
+        eq/ _ _
+          λ r r~r 
+            SQ.elimProp
+              {P = λ q   subQuotientRealizability S (t  r) q    subQuotientRealizability S (t'  r) q   (t  r) ~[ S ] (t'  r)}
+               q  isPropΠ2 λ _ _  isProp~ (t  r) S (t'  r))
+               { (s , s~s) tr~s t'r~s  PER.isTransitive S (t  r) s (t'  r) tr~s (PER.isSymmetric S (t'  r) s t'r~s) })
+              (f .AssemblyMorphism.map [ (r , r~r) ])
+              (t⊩f [ (r , r~r) ] r r~r)
+              (t'⊩f [ (r , r~r) ] r r~r)
 
-subQuotientFunctor : Functor PERCat MOD
-Functor.F-ob subQuotientFunctor R = subQuotientModestSet R
-Functor.F-hom subQuotientFunctor {R} {S} f = subQuotientAssemblyMorphism R S f
-Functor.F-id subQuotientFunctor {R} =
-  AssemblyMorphism≡ _ _
-    (funExt
-      λ sqR 
-        SQ.elimProp
-          {P = λ sqR  subQuotientAssemblyMorphism R R (PERCat .Category.id {R}) .AssemblyMorphism.map sqR  identityMorphism (subQuotientAssembly R) .AssemblyMorphism.map sqR}
-           sqR  squash/ _ _)
-           { (a , a~a) 
-            eq/ _ _
-              (subst (_~[ R ] a) (sym (Ida≡a a)) a~a) })
-          sqR)
-Functor.F-seq subQuotientFunctor {R} {S} {T} f g =
-  AssemblyMorphism≡ _ _
-    (funExt
-      λ sq 
-        SQ.elimProp3
-          {P = λ sqR f g 
-            subQuotientAssemblyMorphism R T (seq' PERCat {R} {S} {T} f g) .AssemblyMorphism.map sqR 
-            seq' MOD
-              {x = subQuotientModestSet R}
-              {y = subQuotientModestSet S}
-              {z = subQuotientModestSet T}
-              (subQuotientAssemblyMorphism R S f) (subQuotientAssemblyMorphism S T g) .AssemblyMorphism.map sqR}
-           sq f g  squash/ _ _)
-           { (a , a~a) (b , bIsTracker) (c , cIsTracker) 
-            eq/ _ _ (subst (_~[ T ] (c  (b  a))) (sym (λ*ComputationRule (` c ̇ (` b ̇ # zero)) a)) (cIsTracker (b  a) (b  a) (bIsTracker a a a~a))) })
-          sq f g)
+subQuotientModestSet : PER  MOD .Category.ob
+subQuotientModestSet R = subQuotient R , subQuotientAssembly R , isModestSubQuotientAssembly R
 
-hasPropFibersSubQuotientFunctor :  R S  hasPropFibers (subQuotientAssemblyMorphism R S)
-hasPropFibersSubQuotientFunctor R S f (x , sqX≡f) (y , sqY≡f) =
-  Σ≡Prop
-       perMap  isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S) _ _)
-      (SQ.elimProp2
-        {P = λ x y  subQuotientAssemblyMorphism R S x  f  subQuotientAssemblyMorphism R S y  f  x  y}
-         x y  isPropΠ2 λ _ _  squash/ _ _)
-         { (x , x⊩f) (y , y⊩f) sqX≡f sqY≡f 
-          eq/ _ _
-            λ r r~r 
-              SQ.elimProp
-                {P = λ f[r]   subQuotientRealizability S (x  r) f[r]     subQuotientRealizability S (y  r) f[r]   (x  r) ~[ S ] (y  r)}
-                 f[r]  isPropΠ2 λ _ _  isProp~ (x  r) S (y  r))
-                 { (s , s~s) xr~s yr~s  PER.isTransitive S (x  r) s (y  r) xr~s (PER.isSymmetric S (y  r) s yr~s) })
-                (f .AssemblyMorphism.map [ (r , r~r) ])
-                (subst  f[r]   subQuotientRealizability S (x  r) f[r] ) (cong  m  m .AssemblyMorphism.map [ (r , r~r) ]) sqX≡f) (x⊩f r r r~r))
-                (subst  f[r]   subQuotientRealizability S (y  r) f[r] ) (cong  m  m .AssemblyMorphism.map [ (r , r~r) ]) sqY≡f) (y⊩f r r r~r)) })
-        x y sqX≡f sqY≡f)
+subQuotientFunctor : Functor PERCat MOD
+Functor.F-ob subQuotientFunctor R = subQuotientModestSet R
+Functor.F-hom subQuotientFunctor {R} {S} f = subQuotientAssemblyMorphism R S f
+Functor.F-id subQuotientFunctor {R} =
+  AssemblyMorphism≡ _ _
+    (funExt
+      λ sqR 
+        SQ.elimProp
+          {P = λ sqR  subQuotientAssemblyMorphism R R (PERCat .Category.id {R}) .AssemblyMorphism.map sqR  identityMorphism (subQuotientAssembly R) .AssemblyMorphism.map sqR}
+           sqR  squash/ _ _)
+           { (a , a~a) 
+            eq/ _ _
+              (subst (_~[ R ] a) (sym (Ida≡a a)) a~a) })
+          sqR)
+Functor.F-seq subQuotientFunctor {R} {S} {T} f g =
+  AssemblyMorphism≡ _ _
+    (funExt
+      λ sq 
+        SQ.elimProp3
+          {P = λ sqR f g 
+            subQuotientAssemblyMorphism R T (seq' PERCat {R} {S} {T} f g) .AssemblyMorphism.map sqR 
+            seq' MOD
+              {x = subQuotientModestSet R}
+              {y = subQuotientModestSet S}
+              {z = subQuotientModestSet T}
+              (subQuotientAssemblyMorphism R S f) (subQuotientAssemblyMorphism S T g) .AssemblyMorphism.map sqR}
+           sq f g  squash/ _ _)
+           { (a , a~a) (b , bIsTracker) (c , cIsTracker) 
+            eq/ _ _ (subst (_~[ T ] (c  (b  a))) (sym (λ*ComputationRule (` c ̇ (` b ̇ # zero)) a)) (cIsTracker (b  a) (b  a) (bIsTracker a a a~a))) })
+          sq f g)
 
-fiberSubQuotientFunctor :  R S f  fiber (subQuotientAssemblyMorphism R S) f
-fiberSubQuotientFunctor R S f =
-  (subQuotientAssemblyMorphism→perMorphism R S f) ,
-  (AssemblyMorphism≡ _ _
-      (funExt
-         qR 
-          SQ.elimProp
-            {P = λ qR  subQuotientAssemblyMorphism R S (subQuotientAssemblyMorphism→perMorphism R S f) .AssemblyMorphism.map qR  f .AssemblyMorphism.map qR}
-             qR  squash/ _ _)
-             { (r , r~r) 
-              PT.elim
-                {P =
-                  λ fTracker 
-                    subQuotientAssemblyMorphism R S (PT.rec→Set squash/ (InverseDefinition.mainMap R S f) (InverseDefinition.mainMap2Constant R S f) fTracker) .AssemblyMorphism.map [ (r , r~r) ]
-                     f .AssemblyMorphism.map [ (r , r~r) ]}
-                 fTracker  squash/ _ _)
-                 { (t , tIsTracker) 
-                  SQ.elimProp
-                    {P =
-                      λ fqR   subQuotientRealizability S (t  r) fqR  
-                        subQuotientAssemblyMorphism R S (InverseDefinition.mainMap R S f (t , tIsTracker)) .AssemblyMorphism.map [ (r , r~r) ]  fqR}
-                     fqR  isProp→ (squash/ _ _))
-                     { (s , s~s) tr~s  eq/ _ _ tr~s })
-                    (f .AssemblyMorphism.map [ (r , r~r) ])
-                    (tIsTracker [ (r , r~r) ] r r~r) })
-                (f .tracker) })
-            qR)))
+hasPropFibersSubQuotientFunctor :  R S  hasPropFibers (subQuotientAssemblyMorphism R S)
+hasPropFibersSubQuotientFunctor R S f (x , sqX≡f) (y , sqY≡f) =
+  Σ≡Prop
+       perMap  isSetAssemblyMorphism (subQuotientAssembly R) (subQuotientAssembly S) _ _)
+      (SQ.elimProp2
+        {P = λ x y  subQuotientAssemblyMorphism R S x  f  subQuotientAssemblyMorphism R S y  f  x  y}
+         x y  isPropΠ2 λ _ _  squash/ _ _)
+         { (x , x⊩f) (y , y⊩f) sqX≡f sqY≡f 
+          eq/ _ _
+            λ r r~r 
+              SQ.elimProp
+                {P = λ f[r]   subQuotientRealizability S (x  r) f[r]     subQuotientRealizability S (y  r) f[r]   (x  r) ~[ S ] (y  r)}
+                 f[r]  isPropΠ2 λ _ _  isProp~ (x  r) S (y  r))
+                 { (s , s~s) xr~s yr~s  PER.isTransitive S (x  r) s (y  r) xr~s (PER.isSymmetric S (y  r) s yr~s) })
+                (f .AssemblyMorphism.map [ (r , r~r) ])
+                (subst  f[r]   subQuotientRealizability S (x  r) f[r] ) (cong  m  m .AssemblyMorphism.map [ (r , r~r) ]) sqX≡f) (x⊩f r r r~r))
+                (subst  f[r]   subQuotientRealizability S (y  r) f[r] ) (cong  m  m .AssemblyMorphism.map [ (r , r~r) ]) sqY≡f) (y⊩f r r r~r)) })
+        x y sqX≡f sqY≡f)
 
-isFullyFaithfulSubQuotientFunctor : Functor.isFullyFaithful subQuotientFunctor
-equiv-proof (isFullyFaithfulSubQuotientFunctor R S) f = inhProp→isContr (fiberSubQuotientFunctor R S f) (hasPropFibersSubQuotientFunctor R S f)
+fiberSubQuotientFunctor :  R S f  fiber (subQuotientAssemblyMorphism R S) f
+fiberSubQuotientFunctor R S f =
+  (subQuotientAssemblyMorphism→perMorphism R S f) ,
+  (AssemblyMorphism≡ _ _
+      (funExt
+         qR 
+          SQ.elimProp
+            {P = λ qR  subQuotientAssemblyMorphism R S (subQuotientAssemblyMorphism→perMorphism R S f) .AssemblyMorphism.map qR  f .AssemblyMorphism.map qR}
+             qR  squash/ _ _)
+             { (r , r~r) 
+              PT.elim
+                {P =
+                  λ fTracker 
+                    subQuotientAssemblyMorphism R S (PT.rec→Set squash/ (InverseDefinition.mainMap R S f) (InverseDefinition.mainMap2Constant R S f) fTracker) .AssemblyMorphism.map [ (r , r~r) ]
+                     f .AssemblyMorphism.map [ (r , r~r) ]}
+                 fTracker  squash/ _ _)
+                 { (t , tIsTracker) 
+                  SQ.elimProp
+                    {P =
+                      λ fqR   subQuotientRealizability S (t  r) fqR  
+                        subQuotientAssemblyMorphism R S (InverseDefinition.mainMap R S f (t , tIsTracker)) .AssemblyMorphism.map [ (r , r~r) ]  fqR}
+                     fqR  isProp→ (squash/ _ _))
+                     { (s , s~s) tr~s  eq/ _ _ tr~s })
+                    (f .AssemblyMorphism.map [ (r , r~r) ])
+                    (tIsTracker [ (r , r~r) ] r r~r) })
+                (f .tracker) })
+            qR)))
+
+isFullyFaithfulSubQuotientFunctor : Functor.isFullyFaithful subQuotientFunctor
+equiv-proof (isFullyFaithfulSubQuotientFunctor R S) f = inhProp→isContr (fiberSubQuotientFunctor R S f) (hasPropFibersSubQuotientFunctor R S f)
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.Everything.html b/docs/Realizability.Topos.Everything.html index 0e760ab..5831390 100644 --- a/docs/Realizability.Topos.Everything.html +++ b/docs/Realizability.Topos.Everything.html @@ -1,14 +1,14 @@ -Realizability.Topos.Everything
{-# OPTIONS --allow-unsolved-metas #-}
-module Realizability.Topos.Everything where
+Realizability.Topos.Everything
-- Modules on the realizability topos over 𝔸
+module Realizability.Topos.Everything where
 
-open import Realizability.Topos.Object
-open import Realizability.Topos.FunctionalRelation
-open import Realizability.Topos.TerminalObject
-open import Realizability.Topos.BinProducts
-open import Realizability.Topos.Equalizer
-open import Realizability.Topos.MonicReprFuncRel
-open import Realizability.Topos.StrictRelation
-open import Realizability.Topos.ResizedPredicate
-open import Realizability.Topos.SubobjectClassifier
+open import Realizability.Topos.Object
+open import Realizability.Topos.FunctionalRelation
+open import Realizability.Topos.TerminalObject
+open import Realizability.Topos.BinProducts
+open import Realizability.Topos.Equalizer
+open import Realizability.Topos.MonicReprFuncRel
+open import Realizability.Topos.StrictRelation
+open import Realizability.Topos.ResizedPredicate
+open import Realizability.Topos.SubobjectClassifier
 
\ No newline at end of file diff --git a/docs/Realizability.Topos.SubobjectClassifier.html b/docs/Realizability.Topos.SubobjectClassifier.html index 3d028fc..c120f09 100644 --- a/docs/Realizability.Topos.SubobjectClassifier.html +++ b/docs/Realizability.Topos.SubobjectClassifier.html @@ -260,8 +260,8 @@ isProp× (perX .equality .isPropValued _ _) (isProp× - (isPropΠ _ isPropΠ λ _ (toPredicate p) .isPropValued _ _)) - (isPropΠ λ _ isPropΠ λ _ ϕ .predicate .isPropValued _ _)) + (isPropΠ _ isPropΠ λ _ (toPredicate p) .isPropValued _ _)) + (isPropΠ λ _ isPropΠ λ _ ϕ .predicate .isPropValued _ _)) isFunctionalRelation.isStrictDomain (isFuncRel charFuncRel) = do return @@ -708,7 +708,7 @@ classifies {Y , perY} f g f⋆char≡g⋆true = SQ.elimProp2 {P = λ f g (commutes : f [ charFuncRel ] g [ trueFuncRel ]) ∃![ hk RTMorphism perY subPer ] (f hk [ incFuncRel ]) × (g hk [ terminalFuncRel subPer ])} - f g isPropΠ λ _ isPropIsContr) + f g isPropΠ λ _ isPropIsContr) F G F⋆char≡G⋆true let entailment = [F]⋆[G]≡[H]⋆[I]→H⋆I≤F⋆G F charFuncRel G trueFuncRel F⋆char≡G⋆true @@ -722,7 +722,7 @@ λ { h' (f≡h'⋆inc , g≡h'⋆term) SQ.elimProp {P = λ h' (comm1 : [ F ] h' [ incFuncRel ]) (comm2 : [ G ] h' [ terminalFuncRel subPer ]) [ UnivPropWithRepr.H perY F G entailment ] h'} - h' isPropΠ λ _ isPropΠ λ _ squash/ _ _) + h' isPropΠ λ _ isPropΠ λ _ squash/ _ _) H' F≡H'⋆inc G≡H'⋆term UnivPropWithRepr.isUniqueH perY F G entailment H' F≡H'⋆inc G≡H'⋆term) h' diff --git a/docs/Realizability.Tripos.Everything.html b/docs/Realizability.Tripos.Everything.html index df4532f..480edd6 100644 --- a/docs/Realizability.Tripos.Everything.html +++ b/docs/Realizability.Tripos.Everything.html @@ -1,9 +1,11 @@ -Realizability.Tripos.Everything
module Realizability.Tripos.Everything where
-open import Realizability.Tripos.Predicate
-open import Realizability.Tripos.Algebra
-open import Realizability.Tripos.Algebra.Base
-open import Realizability.Tripos.Prealgebra.Everything
-open import Realizability.Tripos.Logic.Syntax
-open import Realizability.Tripos.Logic.Semantics
+Realizability.Tripos.Everything
--  Modules on the realizability tripos.
+
+--  The "Prealgebra" modules establish that 𝔸-valued predicates on a set X form a Heyting prealgebra.
+--  The "Syntax" and "Semantics" modules deal with the internal logic of the tripos.
+module Realizability.Tripos.Everything where
+
+open import Realizability.Tripos.Prealgebra.Everything
+open import Realizability.Tripos.Logic.Syntax
+open import Realizability.Tripos.Logic.Semantics
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Logic.Semantics.html b/docs/Realizability.Tripos.Logic.Semantics.html index af8380b..e1f4f8a 100644 --- a/docs/Realizability.Tripos.Logic.Semantics.html +++ b/docs/Realizability.Tripos.Logic.Semantics.html @@ -1,605 +1,599 @@ -Realizability.Tripos.Logic.Semantics
{-# OPTIONS --allow-unsolved-metas #-}
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.HLevels
-open import Cubical.Foundations.Equiv
-open import Cubical.Foundations.Univalence
-open import Cubical.Foundations.Isomorphism
-open import Cubical.Foundations.Function
-open import Cubical.Foundations.Structure
-open import Cubical.Functions.FunExtEquiv
-open import Cubical.Data.Sigma
-open import Cubical.Data.Empty renaming (elim to ⊥elim ; rec* to ⊥rec*)
-open import Cubical.Data.Unit
-open import Cubical.Data.Sum
-open import Cubical.Data.Vec
-open import Cubical.Data.Nat
-open import Cubical.Data.Fin
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-
-module
-  Realizability.Tripos.Logic.Semantics
-  { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A)  where
-open CombinatoryAlgebra ca
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
-
-open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate')
-open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
-open import Realizability.Tripos.Prealgebra.Meets.Identity ca
-open import Realizability.Tripos.Prealgebra.Joins.Identity ca
-open import Realizability.Tripos.Prealgebra.Implication ca
-open import Realizability.Tripos.Logic.Syntax { = ℓ'}
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
-open Predicate'
-open PredicateProperties hiding (_≤_ ; isTrans≤)
-open Morphism {ℓ' = ℓ'} {ℓ'' = ℓ''}
-private
-  Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''}
-RelationInterpretation :  {n : }  (Vec Sort n)  Type _
-RelationInterpretation {n} relSym = (∀ i   Predicate   lookup i relSym ⟧ˢ )
-
-⟦_⟧ᶜ : Context  hSet ℓ'
- [] ⟧ᶜ = Unit* , isSetUnit* 
- c  x ⟧ᶜ = ( c ⟧ᶜ .fst ×  x ⟧ˢ .fst) , isSet× ( c ⟧ᶜ .snd) ( x ⟧ˢ .snd)
-
-⟦_⟧ⁿ :  {Γ s}  s  Γ    Γ ⟧ᶜ     s ⟧ˢ 
-⟦_⟧ⁿ {.(_  s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧
-⟦_⟧ⁿ {.(_  _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) =  s∈Γ ⟧ⁿ ⟦Γ⟧
-
-⟦_⟧ᵗ :  {Γ s}  Term Γ s    Γ ⟧ᶜ     s ⟧ˢ 
-⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ =  x ⟧ⁿ ⟦Γ⟧
-⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = ( t ⟧ᵗ ⟦Γ⟧) , ( t₁ ⟧ᵗ ⟦Γ⟧)
-⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst ( t ⟧ᵗ ⟦Γ⟧)
-⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd ( t ⟧ᵗ ⟦Γ⟧)
-⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x ( t ⟧ᵗ ⟦Γ⟧)
-
--- R for renamings and r for relations
-⟦_⟧ᴿ :  {Γ Δ}  Renaming Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
- id ⟧ᴿ ctx = ctx
- drop ren ⟧ᴿ (ctx , _) =  ren ⟧ᴿ ctx
- keep ren ⟧ᴿ (ctx , s) = ( ren ⟧ᴿ ctx) , s
-
--- B for suBstitution and s for sorts
-⟦_⟧ᴮ :  {Γ Δ}  Substitution Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
- id ⟧ᴮ ctx = ctx
- t , sub ⟧ᴮ ctx = ( sub ⟧ᴮ ctx) , ( t ⟧ᵗ ctx)
- drop sub ⟧ᴮ (ctx , s) =  sub ⟧ᴮ ctx
-
-renamingVarSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (v : s  Δ)   renamingVar ren v ⟧ⁿ   v ⟧ⁿ   ren ⟧ᴿ
-renamingVarSound {Γ} {.Γ} {s} id v = refl
-renamingVarSound {.(_  _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
-renamingVarSound {.(_  s)} {.(_  s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  ⟦s⟧ }
-renamingVarSound {.(_  _)} {.(_  _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
-
-renamingTermSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (t : Term Δ s)   renamingTerm ren t ⟧ᵗ   t ⟧ᵗ   ren ⟧ᴿ
-renamingTermSound {Γ} {.Γ} {s} id t = refl
-renamingTermSound {.(_  _)} {Δ} {s} (drop ren) (var x) =
- funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren x i ⟦Γ⟧ }
-renamingTermSound {.(_  _)} {Δ} {.(_  _)} r@(drop ren) (t `, t₁) =
- funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) }
-renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₁ t) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
-renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₂ t) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
-renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (fun f t) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
-renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (var v) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingVarSound r v i x }
-renamingTermSound {.(_  _)} {.(_  _)} {.(_  _)} r@(keep ren) (t `, t₁) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) }
-renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₁ t) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
-renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₂ t) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
-renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (fun f t) =
- funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
-
-
-substitutionVarSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (v : s  Δ)   substitutionVar subs v ⟧ᵗ   v ⟧ⁿ   subs ⟧ᴮ
-substitutionVarSound {Γ} {.Γ} {s} id t = refl
-substitutionVarSound {Γ} {.(_  s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧  refl
-substitutionVarSound {Γ} {.(_  _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i  substitutionVarSound subs t i ⟦Γ⟧
-substitutionVarSound {.(_  _)} {Δ} {s} r@(drop subs) t =
- -- TODO : Fix unsolved constraints
- funExt
-   λ { x@(⟦Γ⟧ , ⟦s⟧) 
-      substitutionVar (drop subs) t ⟧ᵗ x
-       ≡[ i ]⟨  renamingTermSound (drop id) (substitutionVar subs t) i x  
-      substitutionVar subs t ⟧ᵗ ( drop id ⟧ᴿ x)
-       ≡⟨ refl 
-      substitutionVar subs t ⟧ᵗ ⟦Γ⟧
-       ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ 
-      t ⟧ⁿ ( subs ⟧ᴮ ⟦Γ⟧)
-       }
-
-substitutionTermSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (t : Term Δ s)   substitutionTerm subs t ⟧ᵗ   t ⟧ᵗ   subs ⟧ᴮ
-substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x
-substitutionTermSound {Γ} {Δ} {.(_  _)} subs (t `, t₁) = funExt λ x i  (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x)
-substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i  substitutionTermSound subs t i x .fst
-substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i  substitutionTermSound subs t i x .snd
-substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i  f (substitutionTermSound subs t i x)
-
-semanticSubstitution :  {Γ Δ}  (subs : Substitution Γ Δ)  Predicate   Δ ⟧ᶜ   Predicate   Γ ⟧ᶜ 
-semanticSubstitution {Γ} {Δ} subs = ⋆_ (str  Γ ⟧ᶜ) (str  Δ ⟧ᶜ)  subs ⟧ᴮ
-
-module Interpretation
-  {n : }
-  (relSym : Vec Sort n)
-  (⟦_⟧ʳ : RelationInterpretation relSym) (isNonTrivial : s  k  ) where
-  open Relational relSym
-  ⟦_⟧ᶠ :  {Γ}  Formula Γ  Predicate   Γ ⟧ᶜ 
-  ⟦_⟧ᶠ {Γ} ⊤ᵗ = pre1   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial
-  ⟦_⟧ᶠ {Γ} ⊥ᵗ = pre0   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial
-  ⟦_⟧ᶠ {Γ} (form `∨ form₁) = _⊔_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
-  ⟦_⟧ᶠ {Γ} (form `∧ form₁) = _⊓_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
-  ⟦_⟧ᶠ {Γ} (form `→ form₁) = _⇒_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
-  ⟦_⟧ᶠ {Γ} (`∃ {B = B} form) = `∃[_] (isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ)) (str  Γ ⟧ᶜ)  { (⟦Γ⟧ , ⟦B⟧)  ⟦Γ⟧ })  form ⟧ᶠ
-  ⟦_⟧ᶠ {Γ} (`∀ {B = B} form) = `∀[_] (isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ)) (str  Γ ⟧ᶜ)  { (⟦Γ⟧ , ⟦B⟧)  ⟦Γ⟧ })  form ⟧ᶠ
-  ⟦_⟧ᶠ {Γ} (rel R t) = ⋆_ (str  Γ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ
-
-  -- Due to a shortcut in the soundness of negation termination checking fails
-  -- TODO : Fix
-  {-# TERMINATING #-}
-  substitutionFormulaSound :  {Γ Δ}  (subs : Substitution Γ Δ)  (f : Formula Δ)   substitutionFormula subs f ⟧ᶠ  semanticSubstitution subs  f ⟧ᶠ
-  substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (pre1   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial)
-      (semanticSubstitution subs (pre1   Δ ⟧ᶜ  (str  Δ ⟧ᶜ) isNonTrivial))
-       γ a a⊩1γ  tt*)
-      λ γ a a⊩1subsγ  tt*
-  substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (pre0   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial)
-      (semanticSubstitution subs (pre0   Δ ⟧ᶜ  (str  Δ ⟧ᶜ) isNonTrivial))
-       _ _ bot  ⊥rec* bot)
-      λ _ _ bot  bot
-  substitutionFormulaSound {Γ} {Δ} subs (f `∨ f₁) =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (_⊔_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
-      (semanticSubstitution subs (_⊔_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
-       γ a a⊩substFormFs 
-        a⊩substFormFs >>=
-          λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) 
-                    inl (pr₁a≡k , subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁
-            ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) 
-                    inr (pr₁a≡k' , subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ })
-      λ γ a a⊩semanticSubsFs 
-        a⊩semanticSubsFs >>=
-          λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) 
-                    inl (pr₁a≡k , (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁
-            ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) 
-                    inr (pr₁a≡k' , (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ }
-  substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (_⊓_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
-      (semanticSubstitution subs (_⊓_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
-       γ a a⊩substFormulaFs 
-        (subst  form  (pr₁  a)   form  γ) (substitutionFormulaSound subs f) (a⊩substFormulaFs .fst)) ,
-        (subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f₁) (a⊩substFormulaFs .snd)))
-      λ γ a a⊩semanticSubstFs 
-        (subst  form  (pr₁  a)   form  γ) (sym (substitutionFormulaSound subs f)) (a⊩semanticSubstFs .fst)) ,
-        (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f₁)) (a⊩semanticSubstFs .snd))
-  substitutionFormulaSound {Γ} {Δ} subs (f `→ f₁) =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (_⇒_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
-      (semanticSubstitution subs (_⇒_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
-       γ a a⊩substFormulaFs 
-        λ b b⊩semanticSubstFs 
-          subst
-             form  (a  b)   form  γ)
-            (substitutionFormulaSound subs f₁)
-            (a⊩substFormulaFs
-              b
-              (subst
-                 form  b   form  γ)
-                (sym (substitutionFormulaSound subs f))
-                b⊩semanticSubstFs)))
-      λ γ a a⊩semanticSubstFs 
-        λ b b⊩substFormulaFs 
-          subst
-             form  (a  b)   form  γ)
-            (sym (substitutionFormulaSound subs f₁))
-            (a⊩semanticSubstFs
-              b
-              (subst
-                 form  b   form  γ)
-                (substitutionFormulaSound subs f)
-                b⊩substFormulaFs))
-  substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (`∃[ isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Γ ⟧ᶜ)  { (f , s)  f })  substitutionFormula (var here , drop subs) f ⟧ᶠ)
-      (semanticSubstitution subs (`∃[ isSet× (str  Δ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Δ ⟧ᶜ)  { (γ , b)  γ })  f ⟧ᶠ))
-       γ a a⊩πSubstFormulaF 
-        a⊩πSubstFormulaF >>=
-          λ { ((γ' , b) , γ'≡γ , a⊩substFormFγ') 
-             (( subs ⟧ᴮ γ') , b) ,
-              ((cong  subs ⟧ᴮ γ'≡γ) ,
-                (subst
-                   form  a   form  (γ' , b))
-                  (substitutionFormulaSound (var here , drop subs) f)
-                  a⊩substFormFγ' )) ∣₁ })
-      λ γ a a⊩semanticSubstF 
-        a⊩semanticSubstF >>=
-          λ (x@(δ , b) , δ≡subsγ , a⊩fx) 
-             (γ , b) ,
-              (refl ,
-                (subst
-                   form  a   form  (γ , b))
-                  (sym (substitutionFormulaSound (var here , drop subs) f))
-                  (subst  x  a    f ⟧ᶠ  (x , b)) δ≡subsγ a⊩fx))) ∣₁
-  substitutionFormulaSound {Γ} {Δ} subs (`∀ {B = B} f) =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (`∀[ isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Γ ⟧ᶜ)  { (f , s)  f })  substitutionFormula (var here , drop subs) f ⟧ᶠ)
-      (semanticSubstitution subs (`∀[ isSet× (str  Δ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Δ ⟧ᶜ)  { (f , s)  f })  f ⟧ᶠ))
-       γ a a⊩substFormF 
-        λ { r x@(δ , b) δ≡subsγ 
-          subst
-             g  (a  r)    f ⟧ᶠ  (g , b))
-            (sym δ≡subsγ)
-            (subst
-               form  (a  r)   form  (γ , b))
-              (substitutionFormulaSound (var here , drop subs) f)
-              (a⊩substFormF r (γ , b) refl)) })
-      λ γ a a⊩semanticSubsF 
-        λ { r x@(γ' , b) γ'≡γ 
-          subst
-             form  (a  r)   form  (γ' , b))
-            (sym (substitutionFormulaSound (var here , drop subs) f))
-            (subst
-               g  (a  r)    f ⟧ᶠ  (g , b))
-              (cong  subs ⟧ᴮ (sym γ'≡γ))
-              (a⊩semanticSubsF r ( subs ⟧ᴮ γ , b) refl)) }
-  substitutionFormulaSound {Γ} {Δ} subs (rel R t) =
-    Predicate≡
-        Γ ⟧ᶜ 
-      (⋆_ (str  Γ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  substitutionTerm subs t ⟧ᵗ  R ⟧ʳ)
-      (semanticSubstitution subs (⋆_ (str  Δ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ))
-       γ a a⊩substTR 
-        subst  transform  a    R ⟧ʳ  (transform γ)) (substitutionTermSound subs t) a⊩substTR)
-      λ γ a a⊩semSubst 
-        subst  transform  a    R ⟧ʳ  (transform γ)) (sym (substitutionTermSound subs t)) a⊩semSubst
-
-  weakenFormulaMonotonic :  {Γ B}  (γ :   Γ ⟧ᶜ )  (ϕ : Formula Γ)  (a : A)  (b :   B ⟧ˢ )  a    ϕ ⟧ᶠ  γ  a    weakenFormula {S = B} ϕ ⟧ᶠ  (γ , b)
-  weakenFormulaMonotonic {Γ} {B} γ ϕ a b =
-    hPropExt
-      ( ϕ ⟧ᶠ .isPropValued γ a)
-      ( weakenFormula ϕ ⟧ᶠ .isPropValued (γ , b) a)
-       a⊩ϕγ  subst  form  a   form  (γ , b)) (sym (substitutionFormulaSound (drop id) ϕ)) a⊩ϕγ)
-      λ a⊩weakenϕγb  subst  form  a   form  (γ , b)) (substitutionFormulaSound (drop id) ϕ) a⊩weakenϕγb
-module Soundness
-  {n}
-  {relSym : Vec Sort n}
-  (isNonTrivial : s  k  )
-  (⟦_⟧ʳ : RelationInterpretation relSym) where
-  open Relational relSym
-  open Interpretation relSym ⟦_⟧ʳ isNonTrivial
-  -- Acknowledgements : 1lab's "the internal logic of a regular hyperdoctrine"
-  infix 35 _⊨_
+Realizability.Tripos.Logic.Semantics
open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm)
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.HLevels
+open import Cubical.Foundations.Equiv
+open import Cubical.Foundations.Univalence
+open import Cubical.Foundations.Isomorphism
+open import Cubical.Foundations.Function
+open import Cubical.Foundations.Structure
+open import Cubical.Functions.FunExtEquiv
+open import Cubical.Data.Sigma
+open import Cubical.Data.Empty renaming (elim to ⊥elim ; rec* to ⊥rec*)
+open import Cubical.Data.Unit
+open import Cubical.Data.Sum
+open import Cubical.Data.Vec
+open import Cubical.Data.Nat
+open import Cubical.Data.FinData
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+
+module
+  Realizability.Tripos.Logic.Semantics
+  { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A)  where
+open CombinatoryAlgebra ca
+
+open import Realizability.Tripos.Prealgebra.Predicate.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Predicate.Properties {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Meets.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Joins.Identity {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Implication {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Logic.Syntax { = ℓ'}
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open Predicate
+open PredicateProperties hiding (_≤_ ; isTrans≤)
+open Morphism
+RelationInterpretation :  {n : }  (Vec Sort n)  Type _
+RelationInterpretation {n} relSym = (∀ i   Predicate   lookup i relSym ⟧ˢ )
+
+⟦_⟧ᶜ : Context  hSet ℓ'
+ [] ⟧ᶜ = Unit* , isSetUnit* 
+ c  x ⟧ᶜ = ( c ⟧ᶜ .fst ×  x ⟧ˢ .fst) , isSet× ( c ⟧ᶜ .snd) ( x ⟧ˢ .snd)
+
+⟦_⟧ⁿ :  {Γ s}  s  Γ    Γ ⟧ᶜ     s ⟧ˢ 
+⟦_⟧ⁿ {.(_  s)} {s} _∈_.here (⟦Γ⟧ , ⟦s⟧) = ⟦s⟧
+⟦_⟧ⁿ {.(_  _)} {s} (_∈_.there s∈Γ) (⟦Γ⟧ , ⟦s⟧) =  s∈Γ ⟧ⁿ ⟦Γ⟧
+
+⟦_⟧ᵗ :  {Γ s}  Term Γ s    Γ ⟧ᶜ     s ⟧ˢ 
+⟦_⟧ᵗ {Γ} {s} (var x) ⟦Γ⟧ =  x ⟧ⁿ ⟦Γ⟧
+⟦_⟧ᵗ {Γ} {s} (t `, t₁) ⟦Γ⟧ = ( t ⟧ᵗ ⟦Γ⟧) , ( t₁ ⟧ᵗ ⟦Γ⟧)
+⟦_⟧ᵗ {Γ} {s} (π₁ t) ⟦Γ⟧ = fst ( t ⟧ᵗ ⟦Γ⟧)
+⟦_⟧ᵗ {Γ} {s} (π₂ t) ⟦Γ⟧ = snd ( t ⟧ᵗ ⟦Γ⟧)
+⟦_⟧ᵗ {Γ} {s} (fun x t) ⟦Γ⟧ = x ( t ⟧ᵗ ⟦Γ⟧)
+
+-- R for renamings and r for relations
+⟦_⟧ᴿ :  {Γ Δ}  Renaming Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
+ id ⟧ᴿ ctx = ctx
+ drop ren ⟧ᴿ (ctx , _) =  ren ⟧ᴿ ctx
+ keep ren ⟧ᴿ (ctx , s) = ( ren ⟧ᴿ ctx) , s
+
+-- B for suBstitution and s for sorts
+⟦_⟧ᴮ :  {Γ Δ}  Substitution Γ Δ    Γ ⟧ᶜ     Δ ⟧ᶜ 
+ id ⟧ᴮ ctx = ctx
+ t , sub ⟧ᴮ ctx = ( sub ⟧ᴮ ctx) , ( t ⟧ᵗ ctx)
+ drop sub ⟧ᴮ (ctx , s) =  sub ⟧ᴮ ctx
+
+renamingVarSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (v : s  Δ)   renamingVar ren v ⟧ⁿ   v ⟧ⁿ   ren ⟧ᴿ
+renamingVarSound {Γ} {.Γ} {s} id v = refl
+renamingVarSound {.(_  _)} {Δ} {s} (drop ren) v = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
+renamingVarSound {.(_  s)} {.(_  s)} {s} (keep ren) here = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  ⟦s⟧ }
+renamingVarSound {.(_  _)} {.(_  _)} {s} (keep ren) (there v) = funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren v i ⟦Γ⟧ }
+
+renamingTermSound :  {Γ Δ s}  (ren : Renaming Γ Δ)  (t : Term Δ s)   renamingTerm ren t ⟧ᵗ   t ⟧ᵗ   ren ⟧ᴿ
+renamingTermSound {Γ} {.Γ} {s} id t = refl
+renamingTermSound {.(_  _)} {Δ} {s} (drop ren) (var x) =
+ funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingVarSound ren x i ⟦Γ⟧ }
+renamingTermSound {.(_  _)} {Δ} {.(_  _)} r@(drop ren) (t `, t₁) =
+ funExt λ { (⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i (⟦Γ⟧ , ⟦s⟧) , renamingTermSound r t₁ i (⟦Γ⟧ , ⟦s⟧) }
+renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₁ t) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
+renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (π₂ t) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
+renamingTermSound {.(_  _)} {Δ} {s} r@(drop ren) (fun f t) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
+renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (var v) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingVarSound r v i x }
+renamingTermSound {.(_  _)} {.(_  _)} {.(_  _)} r@(keep ren) (t `, t₁) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  (renamingTermSound r t i x) , (renamingTermSound r t₁ i x) }
+renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₁ t) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .fst }
+renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (π₂ t) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  renamingTermSound r t i x .snd }
+renamingTermSound {.(_  _)} {.(_  _)} {s} r@(keep ren) (fun f t) =
+ funExt λ { x@(⟦Γ⟧ , ⟦s⟧) i  f (renamingTermSound r t i x) }
+
+
+substitutionVarSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (v : s  Δ)   substitutionVar subs v ⟧ᵗ   v ⟧ⁿ   subs ⟧ᴮ
+substitutionVarSound {Γ} {.Γ} {s} id t = refl
+substitutionVarSound {Γ} {.(_  s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧  refl
+substitutionVarSound {Γ} {.(_  _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i  substitutionVarSound subs t i ⟦Γ⟧
+substitutionVarSound {Γ'  s'} {Δ} {s} r@(drop subs) t = funExt pointwise where
+  pointwise : (x :  Γ' ⟧ᶜ .fst ×  s' ⟧ˢ .fst)   substitutionVar r t ⟧ᵗ x  ( t ⟧ⁿ   r ⟧ᴮ) x
+  pointwise x@(⟦Γ'⟧ , ⟦s'⟧) =
+     substitutionVar r t ⟧ᵗ x
+      ≡[ i ]⟨ renamingTermSound (drop {s = s'} id) (substitutionVar subs t) i x 
+     substitutionVar subs t ⟧ᵗ ( drop {Δ = Γ'} {s = s'} id ⟧ᴿ x)
+      ≡⟨ refl 
+     substitutionVar subs t ⟧ᵗ ⟦Γ'⟧
+      ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ'⟧ 
+     t ⟧ⁿ ( subs ⟧ᴮ ⟦Γ'⟧)
+      
+
+substitutionTermSound :  {Γ Δ s}  (subs : Substitution Γ Δ)  (t : Term Δ s)   substitutionTerm subs t ⟧ᵗ   t ⟧ᵗ   subs ⟧ᴮ
+substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x
+substitutionTermSound {Γ} {Δ} {.(_  _)} subs (t `, t₁) = funExt λ x i  (substitutionTermSound subs t i x) , (substitutionTermSound subs t₁ i x)
+substitutionTermSound {Γ} {Δ} {s} subs (π₁ t) = funExt λ x i  substitutionTermSound subs t i x .fst
+substitutionTermSound {Γ} {Δ} {s} subs (π₂ t) = funExt λ x i  substitutionTermSound subs t i x .snd
+substitutionTermSound {Γ} {Δ} {s} subs (fun f t) = funExt λ x i  f (substitutionTermSound subs t i x)
+
+semanticSubstitution :  {Γ Δ}  (subs : Substitution Γ Δ)  Predicate   Δ ⟧ᶜ   Predicate   Γ ⟧ᶜ 
+semanticSubstitution {Γ} {Δ} subs = ⋆_ (str  Γ ⟧ᶜ) (str  Δ ⟧ᶜ)  subs ⟧ᴮ
+
+module Interpretation
+  {n : }
+  (relSym : Vec Sort n)
+  (⟦_⟧ʳ : RelationInterpretation relSym) (isNonTrivial : s  k  ) where
+  open Relational relSym
+  ⟦_⟧ᶠ :  {Γ}  Formula Γ  Predicate   Γ ⟧ᶜ 
+  ⟦_⟧ᶠ {Γ} ⊤ᵗ = pre1   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial
+  ⟦_⟧ᶠ {Γ} ⊥ᵗ = pre0   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial
+  ⟦_⟧ᶠ {Γ} (form `∨ form₁) = _⊔_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (form `∧ form₁) = _⊓_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (form `→ form₁) = _⇒_   Γ ⟧ᶜ   form ⟧ᶠ  form₁ ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (`∃ {B = B} form) = `∃[_] (isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ)) (str  Γ ⟧ᶜ)  { (⟦Γ⟧ , ⟦B⟧)  ⟦Γ⟧ })  form ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (`∀ {B = B} form) = `∀[_] (isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ)) (str  Γ ⟧ᶜ)  { (⟦Γ⟧ , ⟦B⟧)  ⟦Γ⟧ })  form ⟧ᶠ
+  ⟦_⟧ᶠ {Γ} (rel R t) = ⋆_ (str  Γ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ
+
+  -- Due to a shortcut in the soundness of negation termination checking fails
+  -- TODO : Fix
+  {-# TERMINATING #-}
+  substitutionFormulaSound :  {Γ Δ}  (subs : Substitution Γ Δ)  (f : Formula Δ)   substitutionFormula subs f ⟧ᶠ  semanticSubstitution subs  f ⟧ᶠ
+  substitutionFormulaSound {Γ} {Δ} subs ⊤ᵗ =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (pre1   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial)
+      (semanticSubstitution subs (pre1   Δ ⟧ᶜ  (str  Δ ⟧ᶜ) isNonTrivial))
+       γ a a⊩1γ  tt*)
+      λ γ a a⊩1subsγ  tt*
+  substitutionFormulaSound {Γ} {Δ} subs ⊥ᵗ =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (pre0   Γ ⟧ᶜ  (str  Γ ⟧ᶜ) isNonTrivial)
+      (semanticSubstitution subs (pre0   Δ ⟧ᶜ  (str  Δ ⟧ᶜ) isNonTrivial))
+       _ _ bot  ⊥rec* bot)
+      λ _ _ bot  bot
+  substitutionFormulaSound {Γ} {Δ} subs (f `∨ f₁) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (_⊔_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
+      (semanticSubstitution subs (_⊔_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
+       γ a a⊩substFormFs 
+        a⊩substFormFs >>=
+          λ { (inl (pr₁a≡k , pr₂a⊩substFormF)) 
+                    inl (pr₁a≡k , subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f) pr₂a⊩substFormF) ∣₁
+            ; (inr (pr₁a≡k' , pr₂a⊩substFormF₁)) 
+                    inr (pr₁a≡k' , subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f₁) pr₂a⊩substFormF₁) ∣₁ })
+      λ γ a a⊩semanticSubsFs 
+        a⊩semanticSubsFs >>=
+          λ { (inl (pr₁a≡k , pr₂a⊩semanticSubsF)) 
+                    inl (pr₁a≡k , (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f)) pr₂a⊩semanticSubsF)) ∣₁
+            ; (inr (pr₁a≡k' , pr₂a⊩semanticSubsF₁)) 
+                    inr (pr₁a≡k' , (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f₁)) pr₂a⊩semanticSubsF₁)) ∣₁ }
+  substitutionFormulaSound {Γ} {Δ} subs (f `∧ f₁) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (_⊓_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
+      (semanticSubstitution subs (_⊓_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
+       γ a a⊩substFormulaFs 
+        (subst  form  (pr₁  a)   form  γ) (substitutionFormulaSound subs f) (a⊩substFormulaFs .fst)) ,
+        (subst  form  (pr₂  a)   form  γ) (substitutionFormulaSound subs f₁) (a⊩substFormulaFs .snd)))
+      λ γ a a⊩semanticSubstFs 
+        (subst  form  (pr₁  a)   form  γ) (sym (substitutionFormulaSound subs f)) (a⊩semanticSubstFs .fst)) ,
+        (subst  form  (pr₂  a)   form  γ) (sym (substitutionFormulaSound subs f₁)) (a⊩semanticSubstFs .snd))
+  substitutionFormulaSound {Γ} {Δ} subs (f `→ f₁) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (_⇒_   Γ ⟧ᶜ   substitutionFormula subs f ⟧ᶠ  substitutionFormula subs f₁ ⟧ᶠ)
+      (semanticSubstitution subs (_⇒_   Δ ⟧ᶜ   f ⟧ᶠ  f₁ ⟧ᶠ))
+       γ a a⊩substFormulaFs 
+        λ b b⊩semanticSubstFs 
+          subst
+             form  (a  b)   form  γ)
+            (substitutionFormulaSound subs f₁)
+            (a⊩substFormulaFs
+              b
+              (subst
+                 form  b   form  γ)
+                (sym (substitutionFormulaSound subs f))
+                b⊩semanticSubstFs)))
+      λ γ a a⊩semanticSubstFs 
+        λ b b⊩substFormulaFs 
+          subst
+             form  (a  b)   form  γ)
+            (sym (substitutionFormulaSound subs f₁))
+            (a⊩semanticSubstFs
+              b
+              (subst
+                 form  b   form  γ)
+                (substitutionFormulaSound subs f)
+                b⊩substFormulaFs))
+  substitutionFormulaSound {Γ} {Δ} subs (`∃ {B = B} f) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (`∃[ isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Γ ⟧ᶜ)  { (f , s)  f })  substitutionFormula (var here , drop subs) f ⟧ᶠ)
+      (semanticSubstitution subs (`∃[ isSet× (str  Δ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Δ ⟧ᶜ)  { (γ , b)  γ })  f ⟧ᶠ))
+       γ a a⊩πSubstFormulaF 
+        a⊩πSubstFormulaF >>=
+          λ { ((γ' , b) , γ'≡γ , a⊩substFormFγ') 
+             (( subs ⟧ᴮ γ') , b) ,
+              ((cong  subs ⟧ᴮ γ'≡γ) ,
+                (subst
+                   form  a   form  (γ' , b))
+                  (substitutionFormulaSound (var here , drop subs) f)
+                  a⊩substFormFγ' )) ∣₁ })
+      λ γ a a⊩semanticSubstF 
+        a⊩semanticSubstF >>=
+          λ (x@(δ , b) , δ≡subsγ , a⊩fx) 
+             (γ , b) ,
+              (refl ,
+                (subst
+                   form  a   form  (γ , b))
+                  (sym (substitutionFormulaSound (var here , drop subs) f))
+                  (subst  x  a    f ⟧ᶠ  (x , b)) δ≡subsγ a⊩fx))) ∣₁
+  substitutionFormulaSound {Γ} {Δ} subs (`∀ {B = B} f) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (`∀[ isSet× (str  Γ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Γ ⟧ᶜ)  { (f , s)  f })  substitutionFormula (var here , drop subs) f ⟧ᶠ)
+      (semanticSubstitution subs (`∀[ isSet× (str  Δ ⟧ᶜ) (str  B ⟧ˢ) ] (str  Δ ⟧ᶜ)  { (f , s)  f })  f ⟧ᶠ))
+       γ a a⊩substFormF 
+        λ { r x@(δ , b) δ≡subsγ 
+          subst
+             g  (a  r)    f ⟧ᶠ  (g , b))
+            (sym δ≡subsγ)
+            (subst
+               form  (a  r)   form  (γ , b))
+              (substitutionFormulaSound (var here , drop subs) f)
+              (a⊩substFormF r (γ , b) refl)) })
+      λ γ a a⊩semanticSubsF 
+        λ { r x@(γ' , b) γ'≡γ 
+          subst
+             form  (a  r)   form  (γ' , b))
+            (sym (substitutionFormulaSound (var here , drop subs) f))
+            (subst
+               g  (a  r)    f ⟧ᶠ  (g , b))
+              (cong  subs ⟧ᴮ (sym γ'≡γ))
+              (a⊩semanticSubsF r ( subs ⟧ᴮ γ , b) refl)) }
+  substitutionFormulaSound {Γ} {Δ} subs (rel R t) =
+    Predicate≡
+        Γ ⟧ᶜ 
+      (⋆_ (str  Γ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  substitutionTerm subs t ⟧ᵗ  R ⟧ʳ)
+      (semanticSubstitution subs (⋆_ (str  Δ ⟧ᶜ) (str  lookup R relSym ⟧ˢ)  t ⟧ᵗ  R ⟧ʳ))
+       γ a a⊩substTR 
+        subst  transform  a    R ⟧ʳ  (transform γ)) (substitutionTermSound subs t) a⊩substTR)
+      λ γ a a⊩semSubst 
+        subst  transform  a    R ⟧ʳ  (transform γ)) (sym (substitutionTermSound subs t)) a⊩semSubst
+
+  weakenFormulaMonotonic :  {Γ B}  (γ :   Γ ⟧ᶜ )  (ϕ : Formula Γ)  (a : A)  (b :   B ⟧ˢ )  a    ϕ ⟧ᶠ  γ  a    weakenFormula {S = B} ϕ ⟧ᶠ  (γ , b)
+  weakenFormulaMonotonic {Γ} {B} γ ϕ a b =
+    hPropExt
+      ( ϕ ⟧ᶠ .isPropValued γ a)
+      ( weakenFormula ϕ ⟧ᶠ .isPropValued (γ , b) a)
+       a⊩ϕγ  subst  form  a   form  (γ , b)) (sym (substitutionFormulaSound (drop id) ϕ)) a⊩ϕγ)
+      λ a⊩weakenϕγb  subst  form  a   form  (γ , b)) (substitutionFormulaSound (drop id) ϕ) a⊩weakenϕγb
+module Soundness
+  {n}
+  {relSym : Vec Sort n}
+  (isNonTrivial : s  k  )
+  (⟦_⟧ʳ : RelationInterpretation relSym) where
+  open Relational relSym
+  open Interpretation relSym ⟦_⟧ʳ isNonTrivial
+  -- Acknowledgements : 1lab's "the internal logic of a regular hyperdoctrine"
+  infix 35 _⊨_
   
-  module PredProps = PredicateProperties
+  module PredProps = PredicateProperties
   
-  _⊨_ :  {Γ}  Formula Γ  Formula Γ  Type (ℓ-max (ℓ-max  ℓ'') ℓ')
-  _⊨_ {Γ} ϕ ψ =  ϕ ⟧ᶠ   ψ ⟧ᶠ where open PredProps   Γ ⟧ᶜ 
-
-  entails = _⊨_
-
-  holdsInTripos :  {Γ}  Formula Γ  Type (ℓ-max (ℓ-max  ℓ'') ℓ')
-  holdsInTripos {Γ} form = ⊤ᵗ  form
-
-  private
-    variable
-      Γ Δ : Context
-      ϕ ψ θ : Formula Γ
-      χ μ ν : Formula Δ
-
-  cut :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ψ  θ  ϕ  θ
-  cut {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤  ϕ ⟧ᶠ  ψ ⟧ᶠ  θ ⟧ᶠ ϕ⊨ψ ψ⊨θ where open PredProps   Γ ⟧ᶜ 
-
-  substitutionEntailment :  {Γ Δ} (subs : Substitution Γ Δ)  {ϕ ψ : Formula Δ}  ϕ  ψ  substitutionFormula subs ϕ  substitutionFormula subs ψ
-  substitutionEntailment {Γ} {Δ} subs {ϕ} {ψ} ϕ⊨ψ =
-    subst2
-       ϕ' ψ'  ϕ' ≤Γ ψ')
-      (sym (substitutionFormulaSound subs ϕ))
-      (sym (substitutionFormulaSound subs ψ))
-      (ϕ⊨ψ >>=
-        λ { (a , a⊩ϕ≤ψ) 
-           a ,  γ b b⊩ϕsubsγ  a⊩ϕ≤ψ ( subs ⟧ᴮ γ) b b⊩ϕsubsγ) ∣₁ }) where
-      open PredProps {ℓ'' = ℓ''}   Γ ⟧ᶜ  renaming (_≤_ to _≤Γ_)
-      open PredProps {ℓ'' = ℓ''}   Δ ⟧ᶜ  renaming (_≤_ to _≤Δ_)
-
-  `∧intro :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  entails ϕ θ  entails ϕ (ψ `∧ θ)
-  `∧intro {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ϕ⊨θ =
-    do
-      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
-      (b , b⊩ϕ⊨θ)  ϕ⊨θ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` pair ̇ (` a ̇ # fzero) ̇ (` b ̇ # fzero)
-      return
-        (λ* prover ,
-          λ γ r r⊩ϕγ 
-            let
-              proofEq : λ* prover  r  pair  (a  r)  (b  r)
-              proofEq = λ*ComputationRule prover (r  [])
-
-              pr₁proofEq : pr₁  (λ* prover  r)  a  r
-              pr₁proofEq =
-                pr₁  (λ* prover  r)
-                  ≡⟨ cong  x  pr₁  x) proofEq 
-                pr₁  (pair  (a  r)  (b  r))
-                  ≡⟨ pr₁pxy≡x _ _ 
-                a  r
-                  
-
-              pr₂proofEq : pr₂  (λ* prover  r)  b  r
-              pr₂proofEq =
-                pr₂  (λ* prover  r)
-                  ≡⟨ cong  x  pr₂  x) proofEq 
-                pr₂  (pair  (a  r)  (b  r))
-                  ≡⟨ pr₂pxy≡y _ _ 
-                b  r
-                  
-            in
-            subst  r  r    ψ ⟧ᶠ  γ) (sym pr₁proofEq) (a⊩ϕ⊨ψ γ r r⊩ϕγ) ,
-            subst  r  r    θ ⟧ᶠ  γ) (sym pr₂proofEq) (b⊩ϕ⊨θ γ r r⊩ϕγ))
-
-  `∧elim1 :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `∧ θ)  ϕ  ψ
-  `∧elim1 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ =
-    do
-      (a , a⊩ϕ⊨ψ∧θ)  ϕ⊨ψ∧θ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` pr₁ ̇ (` a ̇ # fzero)
-      return
-        (λ* prover ,
-          λ γ b b⊩ϕγ  subst  r  r    ψ ⟧ᶠ  γ) (sym (λ*ComputationRule prover (b  []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst))
+  _⊨_ :  {Γ}  Formula Γ  Formula Γ  Type (ℓ-max (ℓ-max  ℓ'') ℓ')
+  _⊨_ {Γ} ϕ ψ =  ϕ ⟧ᶠ   ψ ⟧ᶠ where open PredProps   Γ ⟧ᶜ 
+
+  entails = _⊨_
+
+  holdsInTripos :  {Γ}  Formula Γ  Type (ℓ-max (ℓ-max  ℓ'') ℓ')
+  holdsInTripos {Γ} form = ⊤ᵗ  form
+
+  private
+    variable
+      Γ Δ : Context
+      ϕ ψ θ : Formula Γ
+      χ μ ν : Formula Δ
+
+  cut :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ψ  θ  ϕ  θ
+  cut {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ψ⊨θ = isTrans≤  ϕ ⟧ᶠ  ψ ⟧ᶠ  θ ⟧ᶠ ϕ⊨ψ ψ⊨θ where open PredProps   Γ ⟧ᶜ 
+
+  substitutionEntailment :  {Γ Δ} (subs : Substitution Γ Δ)  {ϕ ψ : Formula Δ}  ϕ  ψ  substitutionFormula subs ϕ  substitutionFormula subs ψ
+  substitutionEntailment {Γ} {Δ} subs {ϕ} {ψ} ϕ⊨ψ =
+    subst2
+       ϕ' ψ'  ϕ' ≤Γ ψ')
+      (sym (substitutionFormulaSound subs ϕ))
+      (sym (substitutionFormulaSound subs ψ))
+      (ϕ⊨ψ >>=
+        λ { (a , a⊩ϕ≤ψ) 
+           a ,  γ b b⊩ϕsubsγ  a⊩ϕ≤ψ ( subs ⟧ᴮ γ) b b⊩ϕsubsγ) ∣₁ }) where
+      open PredProps   Γ ⟧ᶜ  renaming (_≤_ to _≤Γ_)
+      open PredProps   Δ ⟧ᶜ  renaming (_≤_ to _≤Δ_)
+
+  `∧intro :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  entails ϕ θ  entails ϕ (ψ `∧ θ)
+  `∧intro {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ ϕ⊨θ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      (b , b⊩ϕ⊨θ)  ϕ⊨θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pair ̇ (` a ̇ # zero) ̇ (` b ̇ # zero)
+      return
+        (λ* prover ,
+          λ γ r r⊩ϕγ 
+            let
+              proofEq : λ* prover  r  pair  (a  r)  (b  r)
+              proofEq = λ*ComputationRule prover r
+
+              pr₁proofEq : pr₁  (λ* prover  r)  a  r
+              pr₁proofEq =
+                pr₁  (λ* prover  r)
+                  ≡⟨ cong  x  pr₁  x) proofEq 
+                pr₁  (pair  (a  r)  (b  r))
+                  ≡⟨ pr₁pxy≡x _ _ 
+                a  r
+                  
+
+              pr₂proofEq : pr₂  (λ* prover  r)  b  r
+              pr₂proofEq =
+                pr₂  (λ* prover  r)
+                  ≡⟨ cong  x  pr₂  x) proofEq 
+                pr₂  (pair  (a  r)  (b  r))
+                  ≡⟨ pr₂pxy≡y _ _ 
+                b  r
+                  
+            in
+            subst  r  r    ψ ⟧ᶠ  γ) (sym pr₁proofEq) (a⊩ϕ⊨ψ γ r r⊩ϕγ) ,
+            subst  r  r    θ ⟧ᶠ  γ) (sym pr₂proofEq) (b⊩ϕ⊨θ γ r r⊩ϕγ))
+
+  `∧elim1 :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `∧ θ)  ϕ  ψ
+  `∧elim1 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ =
+    do
+      (a , a⊩ϕ⊨ψ∧θ)  ϕ⊨ψ∧θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pr₁ ̇ (` a ̇ # zero)
+      return
+        (λ* prover ,
+          λ γ b b⊩ϕγ  subst  r  r    ψ ⟧ᶠ  γ) (sym (λ*ComputationRule prover b)) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst))
           
-  `∧elim2 :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `∧ θ)  ϕ  θ
-  `∧elim2 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ =
-    do
-      (a , a⊩ϕ⊨ψ∧θ)  ϕ⊨ψ∧θ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` pr₂ ̇ (` a ̇ # fzero)
-      return
-        (λ* prover ,
-          λ γ b b⊩ϕγ  subst  r  r    θ ⟧ᶠ  γ) (sym (λ*ComputationRule prover (b  []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd))
-
-  `∃intro :  {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ  B)} {t : Term Γ B}  ϕ  substitutionFormula (t , id) ψ  ϕ  `∃ ψ
-  `∃intro {Γ} {ϕ} {B} {ψ} {t} ϕ⊨ψ[t/x] =
-    do
-      (a , a⊩ϕ⊨ψ[t/x])  ϕ⊨ψ[t/x]
-      return
-        (a ,  γ b b⊩ϕγ   (γ , ( t ⟧ᵗ γ)) ,
-        (refl , (subst  form  (a  b)   form  γ) (substitutionFormulaSound (t , id) ψ) (a⊩ϕ⊨ψ[t/x] γ b b⊩ϕγ))) ∣₁))
-
-  `∃elim :  {Γ} {ϕ θ : Formula Γ} {B} {ψ : Formula (Γ  B)}  ϕ  `∃ ψ  (weakenFormula ϕ `∧ ψ)  weakenFormula θ  ϕ  θ
-  `∃elim {Γ} {ϕ} {θ} {B} {ψ} ϕ⊨∃ψ ϕ∧ψ⊨θ =
-    do
-      (a , a⊩ϕ⊨∃ψ)  ϕ⊨∃ψ
-      (b , b⊩ϕ∧ψ⊨θ)  ϕ∧ψ⊨θ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` b ̇ (` pair ̇ # fzero ̇ (` a ̇ # fzero))
-      return
-        (λ* prover ,
-         γ c c⊩ϕγ 
-          subst
-             r  r    θ ⟧ᶠ  γ)
-            (sym (λ*ComputationRule prover (c  [])))
-            (transport
-              (propTruncIdempotent ( θ ⟧ᶠ .isPropValued γ (b  (pair  c  (a  c)))))
-              (a⊩ϕ⊨∃ψ γ c c⊩ϕγ >>=
-                λ { (x@(γ' , b') , (γ'≡γ , a⨾c⊩ψx)) 
-                   transport
-                    (sym
-                      (weakenFormulaMonotonic γ θ (b  (pair  c  (a  c))) b'))
-                    (b⊩ϕ∧ψ⊨θ
-                      (γ , b')
-                      (pair  c  (a  c))
-                      (subst
-                         r  r    weakenFormula ϕ ⟧ᶠ  (γ , b'))
-                        (sym (pr₁pxy≡x _ _))
-                        (transport
-                          (weakenFormulaMonotonic γ ϕ c b') c⊩ϕγ) ,
-                      subst  r  r    ψ ⟧ᶠ  (γ , b')) (sym (pr₂pxy≡y _ _)) (subst  g  (a  c)    ψ ⟧ᶠ  (g , b')) γ'≡γ a⨾c⊩ψx)) ) ∣₁ }))))
-
-  `∀intro :  {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ  B)}  weakenFormula ϕ  ψ  ϕ  `∀ ψ
-  `∀intro {Γ} {ϕ} {B} {ψ} ϕ⊨ψ =
-    do
-      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
-      let
-        prover : ApplStrTerm as 2
-        prover = ` a ̇ # fzero
-      return
-        (λ* prover ,
-         γ b b⊩ϕ  λ { c x@(γ' , b') γ'≡γ 
-          subst
-             r  r    ψ ⟧ᶠ  (γ' , b'))
-            (sym (λ*ComputationRule prover (b  c  [])))
-            (a⊩ϕ⊨ψ
-              (γ' , b')
-              b
-              (transport (weakenFormulaMonotonic γ' ϕ b b') (subst  g  b    ϕ ⟧ᶠ  g) (sym γ'≡γ) b⊩ϕ))) }))
-
-  `∀elim :  {Γ} {B} {ϕ : Formula Γ} {ψ : Formula (Γ  B)}  ϕ  `∀ ψ  (t : Term Γ B)  ϕ  substitutionFormula (t , id) ψ
-  `∀elim {Γ} {B} {ϕ} {ψ} ϕ⊨∀ψ t =
-    do
-      (a , a⊩ϕ⊨∀ψ)  ϕ⊨∀ψ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` a ̇ # fzero ̇ ` k
-      return
-        (λ* prover ,
-         γ b b⊩ϕγ 
-          subst
-           form  (λ* prover  b)   form  γ)
-          (sym (substitutionFormulaSound (t , id) ψ))
-          (subst
-             r  r    ψ ⟧ᶠ  (γ ,  t ⟧ᵗ γ))
-            (sym (λ*ComputationRule prover (b  [])))
-            (a⊩ϕ⊨∀ψ γ b b⊩ϕγ k (γ ,  t ⟧ᵗ γ) refl))))
-
-  `→intro :  {Γ} {ϕ ψ θ : Formula Γ}  (ϕ `∧ ψ)  θ  ϕ  (ψ `→ θ)
-  `→intro {Γ} {ϕ} {ψ} {θ} ϕ∧ψ⊨θ = a⊓b≤c→a≤b⇒c   Γ ⟧ᶜ  (str  Γ ⟧ᶜ)  ϕ ⟧ᶠ  ψ ⟧ᶠ  θ ⟧ᶠ ϕ∧ψ⊨θ
-
-  `→elim :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `→ θ)  ϕ  ψ  ϕ  θ
-  `→elim {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ→θ ϕ⊨ψ =
-    do
-      (a , a⊩ϕ⊨ψ→θ)  ϕ⊨ψ→θ
-      (b , b⊩ϕ⊨ψ)  ϕ⊨ψ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` a ̇ (# fzero) ̇ (` b ̇ # fzero)
-      return
-        (λ* prover ,
-         γ c c⊩ϕγ 
-          subst
-             r  r    θ ⟧ᶠ  γ)
-            (sym (λ*ComputationRule prover (c  [])))
-            (a⊩ϕ⊨ψ→θ γ c c⊩ϕγ (b  c) (b⊩ϕ⊨ψ γ c c⊩ϕγ))))
-
-  `∨introR :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ϕ  (ψ `∨ θ)
-  `∨introR {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ =
-    do
-      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` pair ̇ ` true ̇ (` a ̇ # fzero)
-      return
-        ((λ* prover) ,
-         γ b b⊩ϕγ 
-          let
-            pr₁proofEq : pr₁  (λ* prover  b)  true
-            pr₁proofEq =
-              pr₁  (λ* prover  b)
-                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (b  [])) 
-              pr₁  (pair  true  (a  b))
-                ≡⟨ pr₁pxy≡x _ _ 
-              true
-                
-
-            pr₂proofEq : pr₂  (λ* prover  b)  a  b
-            pr₂proofEq =
-              pr₂  (λ* prover  b)
-                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (b  [])) 
-              pr₂  (pair  true  (a  b))
-                ≡⟨ pr₂pxy≡y _ _ 
-              a  b
-                
-          in  inl (pr₁proofEq , subst  r  r    ψ ⟧ᶠ  γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁))
-
-  `∨introL :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ϕ  (θ `∨ ψ)
-  `∨introL {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ =
-    do
-      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
-      let
-        prover : ApplStrTerm as 1
-        prover = ` pair ̇ ` false ̇ (` a ̇ # fzero)
-      return
-        ((λ* prover) ,
-         γ b b⊩ϕγ 
-          let
-            pr₁proofEq : pr₁  (λ* prover  b)  false
-            pr₁proofEq =
-              pr₁  (λ* prover  b)
-                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (b  [])) 
-              pr₁  (pair  false  (a  b))
-                ≡⟨ pr₁pxy≡x _ _ 
-              false
-                
-
-            pr₂proofEq : pr₂  (λ* prover  b)  a  b
-            pr₂proofEq =
-              pr₂  (λ* prover  b)
-                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (b  [])) 
-              pr₂  (pair  false  (a  b))
-                ≡⟨ pr₂pxy≡y _ _ 
-              a  b
-                
-          in  inr (pr₁proofEq , subst  r  r    ψ ⟧ᶠ  γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁))
-
-  -- Pretty sure this is code duplication
-  `if_then_else_ :  {as n}  ApplStrTerm as n  ApplStrTerm as n  ApplStrTerm as n  ApplStrTerm as n
-  `if a then b else c = ` Id ̇ a ̇ b ̇ c
-
-  `∨elim :  {Γ} {ϕ ψ θ χ : Formula Γ}  (ϕ `∧ ψ)  χ  (ϕ `∧ θ)  χ  (ϕ `∧ (ψ `∨ θ))  χ
-  `∨elim {Γ} {ϕ} {ψ} {θ} {χ} ϕ∧ψ⊨χ ϕ∧θ⊨χ =
-    do
-      (a , a⊩ϕ∧ψ⊨χ)  ϕ∧ψ⊨χ
-      (b , b⊩ϕ∧θ⊨χ)  ϕ∧θ⊨χ
-      let
-        prover : ApplStrTerm as 1
-        prover =
-          (`if ` pr₁ ̇ (` pr₂ ̇ # fzero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))
-      return
-        (λ* prover ,
-        
-          { γ c (pr₁⨾c⊩ϕγ , pr₂⨾c⊩ψ∨θ) 
-            transport
-            (propTruncIdempotent ( χ ⟧ᶠ .isPropValued γ (λ* prover  c)))
-            (pr₂⨾c⊩ψ∨θ >>=
-              λ { (inl (pr₁⨾pr₂⨾c≡true , pr₂⨾pr₂⨾c⊩ψ)) 
-                  let
-                    proofEq : λ* prover  c  a  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                    proofEq =
-                      λ* prover  c
-                        ≡⟨ λ*ComputationRule prover (c  []) 
-                      (if (pr₁  (pr₂  c)) then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                        ≡⟨ cong  x  (if x then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) pr₁⨾pr₂⨾c≡true 
-                      (if true then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                        ≡⟨ cong  x  x  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) (ifTrueThen a b) 
-                      a  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                        
-                  in
-                   subst
-                     r  r    χ ⟧ᶠ  γ)
-                    (sym proofEq)
-                    (a⊩ϕ∧ψ⊨χ
-                      γ
-                      (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                      ((
-                      subst
-                         r  r    ϕ ⟧ᶠ  γ)
-                        (sym (pr₁pxy≡x _ _))
-                        pr₁⨾c⊩ϕγ) ,
-                      subst
-                         r  r    ψ ⟧ᶠ  γ)
-                        (sym (pr₂pxy≡y _ _))
-                        pr₂⨾pr₂⨾c⊩ψ)) ∣₁
-                ; (inr (pr₁pr₂⨾c≡false , pr₂⨾pr₂⨾c⊩θ)) 
-                  let
-                    proofEq : λ* prover  c  b  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                    proofEq =
-                      λ* prover  c
-                        ≡⟨ λ*ComputationRule prover (c  []) 
-                      (if (pr₁  (pr₂  c)) then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                        ≡⟨ cong  x  (if x then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) pr₁pr₂⨾c≡false 
-                      (if false then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                        ≡⟨ cong  x  x  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) (ifFalseElse a b) 
-                      b  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                        
-                  in
-                   subst
-                     r  r    χ ⟧ᶠ  γ)
-                    (sym proofEq)
-                    (b⊩ϕ∧θ⊨χ
-                      γ
-                      (pair  (pr₁  c)  (pr₂  (pr₂  c)))
-                      ((subst  r  r    ϕ ⟧ᶠ  γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) ,
-                       (subst  r  r    θ ⟧ᶠ  γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) }))
+  `∧elim2 :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `∧ θ)  ϕ  θ
+  `∧elim2 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ =
+    do
+      (a , a⊩ϕ⊨ψ∧θ)  ϕ⊨ψ∧θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pr₂ ̇ (` a ̇ # zero)
+      return
+        (λ* prover ,
+          λ γ b b⊩ϕγ  subst  r  r    θ ⟧ᶠ  γ) (sym (λ*ComputationRule prover b)) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd))
+
+  `∃intro :  {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ  B)} {t : Term Γ B}  ϕ  substitutionFormula (t , id) ψ  ϕ  `∃ ψ
+  `∃intro {Γ} {ϕ} {B} {ψ} {t} ϕ⊨ψ[t/x] =
+    do
+      (a , a⊩ϕ⊨ψ[t/x])  ϕ⊨ψ[t/x]
+      return
+        (a ,  γ b b⊩ϕγ   (γ , ( t ⟧ᵗ γ)) ,
+        (refl , (subst  form  (a  b)   form  γ) (substitutionFormulaSound (t , id) ψ) (a⊩ϕ⊨ψ[t/x] γ b b⊩ϕγ))) ∣₁))
+
+  `∃elim :  {Γ} {ϕ θ : Formula Γ} {B} {ψ : Formula (Γ  B)}  ϕ  `∃ ψ  (weakenFormula ϕ `∧ ψ)  weakenFormula θ  ϕ  θ
+  `∃elim {Γ} {ϕ} {θ} {B} {ψ} ϕ⊨∃ψ ϕ∧ψ⊨θ =
+    do
+      (a , a⊩ϕ⊨∃ψ)  ϕ⊨∃ψ
+      (b , b⊩ϕ∧ψ⊨θ)  ϕ∧ψ⊨θ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` b ̇ (` pair ̇ # zero ̇ (` a ̇ # zero))
+      return
+        (λ* prover ,
+         γ c c⊩ϕγ 
+          subst
+             r  r    θ ⟧ᶠ  γ)
+            (sym (λ*ComputationRule prover c))
+            (transport
+              (propTruncIdempotent ( θ ⟧ᶠ .isPropValued γ (b  (pair  c  (a  c)))))
+              (a⊩ϕ⊨∃ψ γ c c⊩ϕγ >>=
+                λ { (x@(γ' , b') , (γ'≡γ , a⨾c⊩ψx)) 
+                   transport
+                    (sym
+                      (weakenFormulaMonotonic γ θ (b  (pair  c  (a  c))) b'))
+                    (b⊩ϕ∧ψ⊨θ
+                      (γ , b')
+                      (pair  c  (a  c))
+                      (subst
+                         r  r    weakenFormula ϕ ⟧ᶠ  (γ , b'))
+                        (sym (pr₁pxy≡x _ _))
+                        (transport
+                          (weakenFormulaMonotonic γ ϕ c b') c⊩ϕγ) ,
+                      subst  r  r    ψ ⟧ᶠ  (γ , b')) (sym (pr₂pxy≡y _ _)) (subst  g  (a  c)    ψ ⟧ᶠ  (g , b')) γ'≡γ a⨾c⊩ψx)) ) ∣₁ }))))
+
+  `∀intro :  {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ  B)}  weakenFormula ϕ  ψ  ϕ  `∀ ψ
+  `∀intro {Γ} {ϕ} {B} {ψ} ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 2
+        prover = ` a ̇ # one
+      return
+        (λ*2 prover ,
+         γ b b⊩ϕ  λ { c x@(γ' , b') γ'≡γ 
+          subst
+             r  r    ψ ⟧ᶠ  (γ' , b'))
+            (sym (λ*2ComputationRule prover b c))
+            (a⊩ϕ⊨ψ
+              (γ' , b')
+              b
+              (transport (weakenFormulaMonotonic γ' ϕ b b') (subst  g  b    ϕ ⟧ᶠ  g) (sym γ'≡γ) b⊩ϕ))) }))
+
+  `∀elim :  {Γ} {B} {ϕ : Formula Γ} {ψ : Formula (Γ  B)}  ϕ  `∀ ψ  (t : Term Γ B)  ϕ  substitutionFormula (t , id) ψ
+  `∀elim {Γ} {B} {ϕ} {ψ} ϕ⊨∀ψ t =
+    do
+      (a , a⊩ϕ⊨∀ψ)  ϕ⊨∀ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` a ̇ # zero ̇ ` k
+      return
+        (λ* prover ,
+         γ b b⊩ϕγ 
+          subst
+           form  (λ* prover  b)   form  γ)
+          (sym (substitutionFormulaSound (t , id) ψ))
+          (subst
+             r  r    ψ ⟧ᶠ  (γ ,  t ⟧ᵗ γ))
+            (sym (λ*ComputationRule prover b))
+            (a⊩ϕ⊨∀ψ γ b b⊩ϕγ k (γ ,  t ⟧ᵗ γ) refl))))
+
+  `→intro :  {Γ} {ϕ ψ θ : Formula Γ}  (ϕ `∧ ψ)  θ  ϕ  (ψ `→ θ)
+  `→intro {Γ} {ϕ} {ψ} {θ} ϕ∧ψ⊨θ = a⊓b≤c→a≤b⇒c   Γ ⟧ᶜ  (str  Γ ⟧ᶜ)  ϕ ⟧ᶠ  ψ ⟧ᶠ  θ ⟧ᶠ ϕ∧ψ⊨θ
+
+  `→elim :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  (ψ `→ θ)  ϕ  ψ  ϕ  θ
+  `→elim {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ→θ ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ→θ)  ϕ⊨ψ→θ
+      (b , b⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` a ̇ (# zero) ̇ (` b ̇ # zero)
+      return
+        (λ* prover ,
+         γ c c⊩ϕγ 
+          subst
+             r  r    θ ⟧ᶠ  γ)
+            (sym (λ*ComputationRule prover c))
+            (a⊩ϕ⊨ψ→θ γ c c⊩ϕγ (b  c) (b⊩ϕ⊨ψ γ c c⊩ϕγ))))
+
+  `∨introR :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ϕ  (ψ `∨ θ)
+  `∨introR {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pair ̇ ` true ̇ (` a ̇ # zero)
+      return
+        ((λ* prover) ,
+         γ b b⊩ϕγ 
+          let
+            pr₁proofEq : pr₁  (λ* prover  b)  true
+            pr₁proofEq =
+              pr₁  (λ* prover  b)
+                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover b) 
+              pr₁  (pair  true  (a  b))
+                ≡⟨ pr₁pxy≡x _ _ 
+              true
+                
+
+            pr₂proofEq : pr₂  (λ* prover  b)  a  b
+            pr₂proofEq =
+              pr₂  (λ* prover  b)
+                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover b) 
+              pr₂  (pair  true  (a  b))
+                ≡⟨ pr₂pxy≡y _ _ 
+              a  b
+                
+          in  inl (pr₁proofEq , subst  r  r    ψ ⟧ᶠ  γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁))
+
+  `∨introL :  {Γ} {ϕ ψ θ : Formula Γ}  ϕ  ψ  ϕ  (θ `∨ ψ)
+  `∨introL {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ =
+    do
+      (a , a⊩ϕ⊨ψ)  ϕ⊨ψ
+      let
+        prover : ApplStrTerm as 1
+        prover = ` pair ̇ ` false ̇ (` a ̇ # zero)
+      return
+        ((λ* prover) ,
+         γ b b⊩ϕγ 
+          let
+            pr₁proofEq : pr₁  (λ* prover  b)  false
+            pr₁proofEq =
+              pr₁  (λ* prover  b)
+                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover b) 
+              pr₁  (pair  false  (a  b))
+                ≡⟨ pr₁pxy≡x _ _ 
+              false
+                
+
+            pr₂proofEq : pr₂  (λ* prover  b)  a  b
+            pr₂proofEq =
+              pr₂  (λ* prover  b)
+                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover b) 
+              pr₂  (pair  false  (a  b))
+                ≡⟨ pr₂pxy≡y _ _ 
+              a  b
+                
+          in  inr (pr₁proofEq , subst  r  r    ψ ⟧ᶠ  γ) (sym pr₂proofEq) (a⊩ϕ⊨ψ γ b b⊩ϕγ)) ∣₁))
+
+  -- Pretty sure this is code duplication
+  `if_then_else_ :  {as n}  ApplStrTerm as n  ApplStrTerm as n  ApplStrTerm as n  ApplStrTerm as n
+  `if a then b else c = ` Id ̇ a ̇ b ̇ c
+
+  `∨elim :  {Γ} {ϕ ψ θ χ : Formula Γ}  (ϕ `∧ ψ)  χ  (ϕ `∧ θ)  χ  (ϕ `∧ (ψ `∨ θ))  χ
+  `∨elim {Γ} {ϕ} {ψ} {θ} {χ} ϕ∧ψ⊨χ ϕ∧θ⊨χ =
+    do
+      (a , a⊩ϕ∧ψ⊨χ)  ϕ∧ψ⊨χ
+      (b , b⊩ϕ∧θ⊨χ)  ϕ∧θ⊨χ
+      let
+        prover : ApplStrTerm as 1
+        prover =
+          (`if ` pr₁ ̇ (` pr₂ ̇ # zero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))
+      return
+        (λ* prover ,
+        
+          { γ c (pr₁⨾c⊩ϕγ , pr₂⨾c⊩ψ∨θ) 
+            transport
+            (propTruncIdempotent ( χ ⟧ᶠ .isPropValued γ (λ* prover  c)))
+            (pr₂⨾c⊩ψ∨θ >>=
+              λ { (inl (pr₁⨾pr₂⨾c≡true , pr₂⨾pr₂⨾c⊩ψ)) 
+                  let
+                    proofEq : λ* prover  c  a  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                    proofEq =
+                      λ* prover  c
+                        ≡⟨ λ*ComputationRule prover c 
+                      (if (pr₁  (pr₂  c)) then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  (if x then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) pr₁⨾pr₂⨾c≡true 
+                      (if true then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  x  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) (ifTrueThen a b) 
+                      a  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        
+                  in
+                   subst
+                     r  r    χ ⟧ᶠ  γ)
+                    (sym proofEq)
+                    (a⊩ϕ∧ψ⊨χ
+                      γ
+                      (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                      ((
+                      subst
+                         r  r    ϕ ⟧ᶠ  γ)
+                        (sym (pr₁pxy≡x _ _))
+                        pr₁⨾c⊩ϕγ) ,
+                      subst
+                         r  r    ψ ⟧ᶠ  γ)
+                        (sym (pr₂pxy≡y _ _))
+                        pr₂⨾pr₂⨾c⊩ψ)) ∣₁
+                ; (inr (pr₁pr₂⨾c≡false , pr₂⨾pr₂⨾c⊩θ)) 
+                  let
+                    proofEq : λ* prover  c  b  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                    proofEq =
+                      λ* prover  c
+                        ≡⟨ λ*ComputationRule prover c 
+                      (if (pr₁  (pr₂  c)) then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  (if x then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) pr₁pr₂⨾c≡false 
+                      (if false then a else b)  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        ≡⟨ cong  x  x  (pair  (pr₁  c)  (pr₂  (pr₂  c)))) (ifFalseElse a b) 
+                      b  (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                        
+                  in
+                   subst
+                     r  r    χ ⟧ᶠ  γ)
+                    (sym proofEq)
+                    (b⊩ϕ∧θ⊨χ
+                      γ
+                      (pair  (pr₁  c)  (pr₂  (pr₂  c)))
+                      ((subst  r  r    ϕ ⟧ᶠ  γ) (sym (pr₁pxy≡x _ _)) pr₁⨾c⊩ϕγ) ,
+                       (subst  r  r    θ ⟧ᶠ  γ) (sym (pr₂pxy≡y _ _)) pr₂⨾pr₂⨾c⊩θ))) ∣₁ }) }))
 
   
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Implication.html b/docs/Realizability.Tripos.Prealgebra.Implication.html index 968a8e7..ef5ab24 100644 --- a/docs/Realizability.Tripos.Prealgebra.Implication.html +++ b/docs/Realizability.Tripos.Prealgebra.Implication.html @@ -1,91 +1,89 @@ Realizability.Tripos.Prealgebra.Implication
open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Univalence
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
+open import Realizability.ApplicativeStructure
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Univalence
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
 
-module Realizability.Tripos.Prealgebra.Implication {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Tripos.Prealgebra.Implication { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open import Realizability.Tripos.Prealgebra.Predicate ca
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
 
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private
-  λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-  λ* = `λ* as fefermanStructure
+module _ (X : Type ℓ') (isSetX' : isSet X) where
+  PredicateX = Predicate  X
+  open Predicate
+  open PredicateProperties  X
+  -- ⇒ is Heyting implication
+  a⊓b≤c→a≤b⇒c :  a b c  (a  b  c)  a  (b  c)
+  a⊓b≤c→a≤b⇒c a b c a⊓b≤c =
+    do
+      (a~ , a~proves)  a⊓b≤c
+      let
+        prover : Term as 2
+        prover = (` a~ ̇ (` pair ̇ (# one)  ̇ (# zero)))
+      return
+        (λ*2 prover ,
+          λ x aₓ aₓ⊩ax bₓ bₓ⊩bx 
+            subst
+               r  r   c  x)
+              (sym (λ*2ComputationRule prover aₓ bₓ))
+              (a~proves
+                x
+                (pair  aₓ  bₓ)
+                ((subst  r  r   a  x) (sym (pr₁pxy≡x _ _)) aₓ⊩ax) ,
+                 (subst  r  r   b  x) (sym (pr₂pxy≡y _ _)) bₓ⊩bx))))
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where
-  PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  -- ⇒ is Heyting implication
-  a⊓b≤c→a≤b⇒c :  a b c  (a  b  c)  a  (b  c)
-  a⊓b≤c→a≤b⇒c a b c a⊓b≤c =
-    do
-      (a~ , a~proves)  a⊓b≤c
-      let prover = (` a~ ̇ (` pair ̇ (# fzero)  ̇ (# fone)))
-      return
-        (λ* prover ,
-          λ x aₓ aₓ⊩ax bₓ bₓ⊩bx 
-            subst
-               r  r   c  x)
-              (sym (λ*ComputationRule prover (aₓ  bₓ  [])))
-              (a~proves
-                x
-                (pair  aₓ  bₓ)
-                ((subst  r  r   a  x) (sym (pr₁pxy≡x _ _)) aₓ⊩ax) ,
-                 (subst  r  r   b  x) (sym (pr₂pxy≡y _ _)) bₓ⊩bx))))
+  a≤b⇒c→a⊓b≤c :  a b c  a  (b  c)  (a  b  c)
+  a≤b⇒c→a⊓b≤c a b c a≤b⇒c =
+    do
+      (a~ , a~proves)  a≤b⇒c
+      let prover = ` a~ ̇ (` pr₁ ̇ (# zero)) ̇ (` pr₂ ̇ (# zero))
+      return
+        (λ* prover ,
+          λ { x abₓ (pr₁abₓ⊩ax , pr₂abₓ⊩bx) 
+            subst
+               r  r   c  x)
+              (sym (λ*ComputationRule prover abₓ))
+              (a~proves x (pr₁  abₓ) pr₁abₓ⊩ax (pr₂  abₓ) pr₂abₓ⊩bx) })
 
-  a≤b⇒c→a⊓b≤c :  a b c  a  (b  c)  (a  b  c)
-  a≤b⇒c→a⊓b≤c a b c a≤b⇒c =
-    do
-      (a~ , a~proves)  a≤b⇒c
-      let prover = ` a~ ̇ (` pr₁ ̇ (# fzero)) ̇ (` pr₂ ̇ (# fzero))
-      return
-        (λ* prover ,
-          λ { x abₓ (pr₁abₓ⊩ax , pr₂abₓ⊩bx) 
-            subst
-               r  r   c  x)
-              (sym (λ*ComputationRule prover (abₓ  [])))
-              (a~proves x (pr₁  abₓ) pr₁abₓ⊩ax (pr₂  abₓ) pr₂abₓ⊩bx) })
+  ⇒isRightAdjointOf⊓ :  a b c  (a  b  c)  (a  b  c)
+  ⇒isRightAdjointOf⊓ a b c = hPropExt (isProp≤ (a  b) c) (isProp≤ a (b  c)) (a⊓b≤c→a≤b⇒c a b c) (a≤b⇒c→a⊓b≤c a b c)
 
-  ⇒isRightAdjointOf⊓ :  a b c  (a  b  c)  (a  b  c)
-  ⇒isRightAdjointOf⊓ a b c = hPropExt (isProp≤ (a  b) c) (isProp≤ a (b  c)) (a⊓b≤c→a≤b⇒c a b c) (a≤b⇒c→a⊓b≤c a b c)
+  antiSym→a⇒c≤b⇒c :  a b c  a  b  b  a  (a  c)  (b  c)
+  antiSym→a⇒c≤b⇒c a b c a≤b b≤a =
+    do
+      (α , αProves)  a≤b
+      (β , βProves)  b≤a
+      let
+        prover : Term as 2
+        prover = (# one) ̇ (` β ̇ # zero)
+      return
+        (λ*2 prover ,
+           x r r⊩a⇒c r' r'⊩b 
+            subst
+               witness  witness   c  x)
+              (sym (λ*2ComputationRule prover r r'))
+              (r⊩a⇒c (β  r') (βProves x r' r'⊩b))))
 
-  antiSym→a⇒c≤b⇒c :  a b c  a  b  b  a  (a  c)  (b  c)
-  antiSym→a⇒c≤b⇒c a b c a≤b b≤a =
-    do
-      (α , αProves)  a≤b
-      (β , βProves)  b≤a
-      let
-        prover : Term as 2
-        prover = (# fzero) ̇ (` β ̇ # fone)
-      return
-        (λ* prover ,
-           x r r⊩a⇒c r' r'⊩b 
-            subst
-               witness  witness   c  x)
-              (sym (λ*ComputationRule prover (r  r'  [])))
-              (r⊩a⇒c (β  r') (βProves x r' r'⊩b))))
-
-  antiSym→a⇒b≤a⇒c :  a b c  b  c  c  b  (a  b)  (a  c)
-  antiSym→a⇒b≤a⇒c a b c b≤c c≤b =
-    do
-      (β , βProves)  b≤c
-      (γ , γProves)  c≤b
-      let
-        prover : Term as 2
-        prover = ` β ̇ ((# fzero) ̇ (# fone))
-      return
-        (λ* prover ,
-           x α α⊩a⇒b a' a'⊩a 
-            subst
-               r  r   c  x)
-              (sym (λ*ComputationRule prover (α  a'  [])))
-              (βProves x (α  a') (α⊩a⇒b a' a'⊩a))))
+  antiSym→a⇒b≤a⇒c :  a b c  b  c  c  b  (a  b)  (a  c)
+  antiSym→a⇒b≤a⇒c a b c b≤c c≤b =
+    do
+      (β , βProves)  b≤c
+      (γ , γProves)  c≤b
+      let
+        prover : Term as 2
+        prover = ` β ̇ ((# one) ̇ (# zero))
+      return
+        (λ*2 prover ,
+           x α α⊩a⇒b a' a'⊩a 
+            subst
+               r  r   c  x)
+              (sym (λ*2ComputationRule prover α a'))
+              (βProves x (α  a') (α⊩a⇒b a' a'⊩a))))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Joins.Associativity.html b/docs/Realizability.Tripos.Prealgebra.Joins.Associativity.html index ecd2f05..7d812b5 100644 --- a/docs/Realizability.Tripos.Prealgebra.Joins.Associativity.html +++ b/docs/Realizability.Tripos.Prealgebra.Joins.Associativity.html @@ -2,246 +2,242 @@ Realizability.Tripos.Prealgebra.Joins.Associativity
{-# OPTIONS --allow-unsolved-metas #-}
 open import Cubical.Foundations.Prelude
 open import Cubical.Data.Empty renaming (rec to ⊥rec)
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Tripos.Prealgebra.Joins.Associativity {} {A : Type } (ca : CombinatoryAlgebra A) where
-open import Realizability.Tripos.Prealgebra.Predicate ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module Realizability.Tripos.Prealgebra.Joins.Associativity { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private
-  λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-  λ* = `λ* as fefermanStructure
+module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
+  private PredicateX = Predicate X
+  open Predicate
+  open PredicateProperties X
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
+  module →⊔Assoc (x y z : PredicateX) where
+    →proverTerm : Term as 1
+    →proverTerm =
+            ` Id ̇
+            (` pr₁ ̇ (# zero)) ̇
+            (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# zero)))) ̇
+            (` Id ̇
+              (` pr₁ ̇ (` pr₂ ̇ (# zero))) ̇
+              (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))) ̇
+              (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero))))
 
-  module →⊔Assoc (x y z : Predicate {ℓ'' = ℓ''} X) where
-    →proverTerm : Term as 1
-    →proverTerm =
-            ` Id ̇
-            (` pr₁ ̇ (# fzero)) ̇
-            (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# fzero)))) ̇
-            (` Id ̇
-              (` pr₁ ̇ (` pr₂ ̇ (# fzero))) ̇
-              (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))) ̇
-              (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))))
+    →prover = λ* →proverTerm
 
-    →prover = λ* →proverTerm
+    module Pr₁a≡true
+      (a : A)
+      (pr₁a≡true : pr₁  a  true) where
 
-    module Pr₁a≡true
-      (a : A)
-      (pr₁a≡true : pr₁  a  true) where
+      private proof = →prover  a
 
-      private proof = →prover  a
+      proof≡pair⨾true⨾pair⨾true⨾pr₂a : proof  pair  true  (pair  true  (pr₂  a))
+      proof≡pair⨾true⨾pair⨾true⨾pr₂a =
+        proof
+          ≡⟨ λ*ComputationRule →proverTerm a 
+        (Id 
+         (pr₁  a) 
+         (pair  true  (pair  true  (pr₂  a))) 
+         (Id 
+           (pr₁  (pr₂  a)) 
+           (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+           (pair  false  (pr₂  (pr₂  a)))))
+           ≡⟨  cong  x 
+                     (Id 
+                       x 
+                       (pair  true  (pair  true  (pr₂  a))) 
+                       (Id 
+                         (pr₁  (pr₂  a)) 
+                         (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+                         (pair  false  (pr₂  (pr₂  a))))))
+                pr₁a≡true 
+         (Id 
+           k 
+           (pair  true  (pair  true  (pr₂  a))) 
+           (Id 
+             (pr₁  (pr₂  a)) 
+             (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+             (pair  false  (pr₂  (pr₂  a)))))
+            ≡⟨ ifTrueThen _ _ 
+          pair  true  (pair  true  (pr₂  a))
+            
 
-      proof≡pair⨾true⨾pair⨾true⨾pr₂a : proof  pair  true  (pair  true  (pr₂  a))
-      proof≡pair⨾true⨾pair⨾true⨾pr₂a =
-        proof
-          ≡⟨ λ*ComputationRule →proverTerm (a  []) 
-        (Id 
-         (pr₁  a) 
-         (pair  true  (pair  true  (pr₂  a))) 
-         (Id 
-           (pr₁  (pr₂  a)) 
-           (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-           (pair  false  (pr₂  (pr₂  a)))))
-           ≡⟨  cong  x 
-                     (Id 
-                       x 
-                       (pair  true  (pair  true  (pr₂  a))) 
-                       (Id 
-                         (pr₁  (pr₂  a)) 
-                         (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-                         (pair  false  (pr₂  (pr₂  a))))))
-                pr₁a≡true 
-         (Id 
-           k 
-           (pair  true  (pair  true  (pr₂  a))) 
-           (Id 
-             (pr₁  (pr₂  a)) 
-             (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-             (pair  false  (pr₂  (pr₂  a)))))
-            ≡⟨ ifTrueThen _ _ 
-          pair  true  (pair  true  (pr₂  a))
-            
+      pr₁⨾proof≡true : pr₁  (→prover  a)  true
+      pr₁⨾proof≡true =
+                     (pr₁  (→prover  a))
+                       ≡⟨ cong  x  pr₁  x) proof≡pair⨾true⨾pair⨾true⨾pr₂a 
+                     pr₁  (pair  true  (pair  true  (pr₂  a)))
+                       ≡⟨ pr₁pxy≡x _ _ 
+                     true
+                        
 
-      pr₁⨾proof≡true : pr₁  (→prover  a)  true
-      pr₁⨾proof≡true =
-                     (pr₁  (→prover  a))
-                       ≡⟨ cong  x  pr₁  x) proof≡pair⨾true⨾pair⨾true⨾pr₂a 
-                     pr₁  (pair  true  (pair  true  (pr₂  a)))
-                       ≡⟨ pr₁pxy≡x _ _ 
-                     true
-                        
+      pr₁pr₂proof≡true : pr₁  (pr₂  (→prover  a))  true
+      pr₁pr₂proof≡true =
+        pr₁  (pr₂  proof)
+          ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair⨾true⨾pair⨾true⨾pr₂a 
+        pr₁  (pr₂  (pair  true  (pair  true  (pr₂  a))))
+          ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+        pr₁  (pair  true  (pr₂  a))
+          ≡⟨ pr₁pxy≡x _ _ 
+        true
+           
 
-      pr₁pr₂proof≡true : pr₁  (pr₂  (→prover  a))  true
-      pr₁pr₂proof≡true =
-        pr₁  (pr₂  proof)
-          ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair⨾true⨾pair⨾true⨾pr₂a 
-        pr₁  (pr₂  (pair  true  (pair  true  (pr₂  a))))
-          ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
-        pr₁  (pair  true  (pr₂  a))
-          ≡⟨ pr₁pxy≡x _ _ 
-        true
-           
+      pr₂pr₂proof≡pr₂a : pr₂  (pr₂  proof)  pr₂  a
+      pr₂pr₂proof≡pr₂a = 
+        pr₂  (pr₂  proof)
+          ≡⟨ cong  x  pr₂  (pr₂  x)) proof≡pair⨾true⨾pair⨾true⨾pr₂a 
+        pr₂  (pr₂  (pair  true  (pair  true  (pr₂  a))))
+          ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+        pr₂  (pair  true  (pr₂  a))
+          ≡⟨ pr₂pxy≡y _ _ 
+        pr₂  a
+          
 
-      pr₂pr₂proof≡pr₂a : pr₂  (pr₂  proof)  pr₂  a
-      pr₂pr₂proof≡pr₂a = 
-        pr₂  (pr₂  proof)
-          ≡⟨ cong  x  pr₂  (pr₂  x)) proof≡pair⨾true⨾pair⨾true⨾pr₂a 
-        pr₂  (pr₂  (pair  true  (pair  true  (pr₂  a))))
-          ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
-        pr₂  (pair  true  (pr₂  a))
-          ≡⟨ pr₂pxy≡y _ _ 
-        pr₂  a
-          
+    module Pr₁a≡false
+      (a : A)
+      (pr₁a≡false : pr₁  a  false) where
+      private proof = →prover  a
+      proof≡y⊔z : proof  (Id 
+                    (pr₁  (pr₂  a)) 
+                    (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+                    (pair  false  (pr₂  (pr₂  a))))
+      proof≡y⊔z =
+        proof
+          ≡⟨ λ*ComputationRule →proverTerm a 
+        (Id 
+         (pr₁  a) 
+         (pair  true  (pair  true  (pr₂  a))) 
+         (Id 
+           (pr₁  (pr₂  a)) 
+           (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+           (pair  false  (pr₂  (pr₂  a)))))
+           ≡⟨  cong  x 
+                     (Id 
+                       x 
+                       (pair  true  (pair  true  (pr₂  a))) 
+                       (Id 
+                         (pr₁  (pr₂  a)) 
+                         (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+                         (pair  false  (pr₂  (pr₂  a))))))
+                pr₁a≡false 
+         (Id 
+           k' 
+           (pair  true  (pair  true  (pr₂  a))) 
+           (Id 
+             (pr₁  (pr₂  a)) 
+             (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+             (pair  false  (pr₂  (pr₂  a)))))
+            ≡⟨ ifFalseElse _ _ 
+         (Id 
+             (pr₁  (pr₂  a)) 
+             (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+             (pair  false  (pr₂  (pr₂  a))))
+            
 
-    module Pr₁a≡false
-      (a : A)
-      (pr₁a≡false : pr₁  a  false) where
-      private proof = →prover  a
-      proof≡y⊔z : proof  (Id 
-                    (pr₁  (pr₂  a)) 
-                    (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-                    (pair  false  (pr₂  (pr₂  a))))
-      proof≡y⊔z =
-        proof
-          ≡⟨ λ*ComputationRule →proverTerm (a  []) 
-        (Id 
-         (pr₁  a) 
-         (pair  true  (pair  true  (pr₂  a))) 
-         (Id 
-           (pr₁  (pr₂  a)) 
-           (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-           (pair  false  (pr₂  (pr₂  a)))))
-           ≡⟨  cong  x 
-                     (Id 
-                       x 
-                       (pair  true  (pair  true  (pr₂  a))) 
-                       (Id 
-                         (pr₁  (pr₂  a)) 
-                         (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-                         (pair  false  (pr₂  (pr₂  a))))))
-                pr₁a≡false 
-         (Id 
-           k' 
-           (pair  true  (pair  true  (pr₂  a))) 
-           (Id 
-             (pr₁  (pr₂  a)) 
-             (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-             (pair  false  (pr₂  (pr₂  a)))))
-            ≡⟨ ifFalseElse _ _ 
-         (Id 
-             (pr₁  (pr₂  a)) 
-             (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-             (pair  false  (pr₂  (pr₂  a))))
-            
-
-      module _ (pr₁pr₂a≡true : pr₁  (pr₂  a)  true) where
-        proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a : proof  pair  true  (pair  false  (pr₂  (pr₂  a)))
-        proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a =
-          proof
-            ≡⟨ proof≡y⊔z 
-          Id 
-              (pr₁  (pr₂  a)) 
-              (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-              (pair  false  (pr₂  (pr₂  a)))
-            ≡⟨ cong
-                x  (Id 
-                        x 
-                        (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-                        (pair  false  (pr₂  (pr₂  a)))))
-               pr₁pr₂a≡true 
-          Id  true  (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-            (pair  false  (pr₂  (pr₂  a)))
-            ≡⟨ ifTrueThen _ _ 
-          (pair  true  (pair  false  (pr₂  (pr₂  a))))
-            
+      module _ (pr₁pr₂a≡true : pr₁  (pr₂  a)  true) where
+        proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a : proof  pair  true  (pair  false  (pr₂  (pr₂  a)))
+        proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a =
+          proof
+            ≡⟨ proof≡y⊔z 
+          Id 
+              (pr₁  (pr₂  a)) 
+              (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+              (pair  false  (pr₂  (pr₂  a)))
+            ≡⟨ cong
+                x  (Id 
+                        x 
+                        (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+                        (pair  false  (pr₂  (pr₂  a)))))
+               pr₁pr₂a≡true 
+          Id  true  (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+            (pair  false  (pr₂  (pr₂  a)))
+            ≡⟨ ifTrueThen _ _ 
+          (pair  true  (pair  false  (pr₂  (pr₂  a))))
+            
         
-        pr₁proof≡true : pr₁  proof  true
-        pr₁proof≡true =
-          pr₁  proof
-            ≡⟨ cong  x  pr₁  x) proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a 
-         pr₁  (pair  true  (pair  false  (pr₂  (pr₂  a))))
-            ≡⟨ pr₁pxy≡x _ _ 
-         true
-            
+        pr₁proof≡true : pr₁  proof  true
+        pr₁proof≡true =
+          pr₁  proof
+            ≡⟨ cong  x  pr₁  x) proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a 
+         pr₁  (pair  true  (pair  false  (pr₂  (pr₂  a))))
+            ≡⟨ pr₁pxy≡x _ _ 
+         true
+            
 
-        pr₁pr₂proof≡k' : pr₁  (pr₂  proof)  k'
-        pr₁pr₂proof≡k' =
-          pr₁  (pr₂  proof)
-            ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a 
-          pr₁  (pr₂  (pair  true  (pair  false  (pr₂  (pr₂  a)))))
-            ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
-          pr₁  (pair  false  (pr₂  (pr₂  a)))
-            ≡⟨ pr₁pxy≡x _ _ 
-          false
-            
+        pr₁pr₂proof≡k' : pr₁  (pr₂  proof)  k'
+        pr₁pr₂proof≡k' =
+          pr₁  (pr₂  proof)
+            ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a 
+          pr₁  (pr₂  (pair  true  (pair  false  (pr₂  (pr₂  a)))))
+            ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+          pr₁  (pair  false  (pr₂  (pr₂  a)))
+            ≡⟨ pr₁pxy≡x _ _ 
+          false
+            
 
         
-        pr₂pr₂proof≡pr₂pr₂a : pr₂  (pr₂  proof)  pr₂  (pr₂  a)
-        pr₂pr₂proof≡pr₂pr₂a =
-          pr₂  (pr₂  proof)
-            ≡⟨ cong  x  pr₂  (pr₂  x)) proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a 
-          pr₂  (pr₂  (pair  true  (pair  false  (pr₂  (pr₂  a)))))
-            ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
-          pr₂  (pair  false  (pr₂  (pr₂  a)))
-            ≡⟨ pr₂pxy≡y _ _ 
-          pr₂  (pr₂  a)
-            
+        pr₂pr₂proof≡pr₂pr₂a : pr₂  (pr₂  proof)  pr₂  (pr₂  a)
+        pr₂pr₂proof≡pr₂pr₂a =
+          pr₂  (pr₂  proof)
+            ≡⟨ cong  x  pr₂  (pr₂  x)) proof≡pair⨾true⨾pair⨾false⨾pr₂⨾pr₂⨾a 
+          pr₂  (pr₂  (pair  true  (pair  false  (pr₂  (pr₂  a)))))
+            ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+          pr₂  (pair  false  (pr₂  (pr₂  a)))
+            ≡⟨ pr₂pxy≡y _ _ 
+          pr₂  (pr₂  a)
+            
 
-      module _ (pr₁pr₂a≡false : pr₁  (pr₂  a)  false) where
+      module _ (pr₁pr₂a≡false : pr₁  (pr₂  a)  false) where
 
-        proof≡pair⨾false⨾pr₂pr₂a : proof  pair  false  (pr₂  (pr₂  a))
-        proof≡pair⨾false⨾pr₂pr₂a =
-          proof
-            ≡⟨ proof≡y⊔z 
-          Id 
-            (pr₁  (pr₂  a)) 
-            (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-            (pair  false  (pr₂  (pr₂  a)))
-            ≡⟨ cong
-                x 
-                 Id 
-                 x 
-                 (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
-                 (pair  false  (pr₂  (pr₂  a))))
-                 pr₁pr₂a≡false 
-          ifFalseElse _ _
+        proof≡pair⨾false⨾pr₂pr₂a : proof  pair  false  (pr₂  (pr₂  a))
+        proof≡pair⨾false⨾pr₂pr₂a =
+          proof
+            ≡⟨ proof≡y⊔z 
+          Id 
+            (pr₁  (pr₂  a)) 
+            (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+            (pair  false  (pr₂  (pr₂  a)))
+            ≡⟨ cong
+                x 
+                 Id 
+                 x 
+                 (pair  true  (pair  false  (pr₂  (pr₂  a)))) 
+                 (pair  false  (pr₂  (pr₂  a))))
+                 pr₁pr₂a≡false 
+          ifFalseElse _ _
 
-        pr₁proof≡false : pr₁  proof  false
-        pr₁proof≡false =
-          pr₁  proof
-            ≡⟨ cong  x  pr₁  x) proof≡pair⨾false⨾pr₂pr₂a 
-          pr₁  (pair  false  (pr₂  (pr₂  a)))
-            ≡⟨ pr₁pxy≡x _ _ 
-          false
-            
+        pr₁proof≡false : pr₁  proof  false
+        pr₁proof≡false =
+          pr₁  proof
+            ≡⟨ cong  x  pr₁  x) proof≡pair⨾false⨾pr₂pr₂a 
+          pr₁  (pair  false  (pr₂  (pr₂  a)))
+            ≡⟨ pr₁pxy≡x _ _ 
+          false
+            
 
-        pr₂proof≡pr₂pr₂a : pr₂  proof  pr₂  (pr₂  a)
-        pr₂proof≡pr₂pr₂a =
-          pr₂  proof
-            ≡⟨ cong  x  pr₂  x) proof≡pair⨾false⨾pr₂pr₂a 
-          pr₂  (pair  false  (pr₂  (pr₂  a)))
-            ≡⟨ pr₂pxy≡y _ _ 
-          pr₂  (pr₂  a)
-            
+        pr₂proof≡pr₂pr₂a : pr₂  proof  pr₂  (pr₂  a)
+        pr₂proof≡pr₂pr₂a =
+          pr₂  proof
+            ≡⟨ cong  x  pr₂  x) proof≡pair⨾false⨾pr₂pr₂a 
+          pr₂  (pair  false  (pr₂  (pr₂  a)))
+            ≡⟨ pr₂pxy≡y _ _ 
+          pr₂  (pr₂  a)
+            
             
 
-  x⊔_y⊔z≤x⊔y_⊔z :  x y z  (x  (y  z))  ((x  y)  z)
-  x⊔_y⊔z≤x⊔y_⊔z x y z =
-    (do
-      let
-        {-
+  x⊔_y⊔z≤x⊔y_⊔z :  x y z  (x  (y  z))  ((x  y)  z)
+  x⊔_y⊔z≤x⊔y_⊔z x y z =
+    (do
+      let
+        {-
         if pr₁ a then
           ⟨ true , ⟨ true , pr₂ a ⟩⟩
         else
@@ -250,247 +246,247 @@
           else
             ⟨ false , pr₂ (pr₂ a) ⟩
         -}
-        prover : Term as 1
-        prover =
-          ` Id ̇
-            (` pr₁ ̇ (# fzero)) ̇
-            (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# fzero)))) ̇
-            (` Id ̇
-              (` pr₁ ̇ (` pr₂ ̇ (# fzero))) ̇
-              (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))) ̇
-              (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))))
-      return
-        (λ* prover ,
-          λ x' a a⊩x⊔_y⊔z 
-            transport
-              (propTruncIdempotent (((x  y)  z) .isPropValued x' (λ* prover  a)))
-              (a⊩x⊔_y⊔z
-                >>=  { (inl (pr₁a≡true , pr₁a⊩x))
-                        inl
-                           (Pr₁a≡true.pr₁⨾proof≡true a pr₁a≡true ,
-                           transport
-                             (propTruncIdempotent isPropPropTrunc)
-                              a⊩x⊔_y⊔z
-                               >>=  { (inl (pr₁a≡k , pr₂a⊩x)) 
-                                          inl
-                                           (Pr₁a≡true.pr₁pr₂proof≡true a pr₁a≡true ,
-                                            subst
-                                               r  r   x  x')
-                                              (sym (Pr₁a≡true.pr₂pr₂proof≡pr₂a a pr₁a≡true))
-                                              pr₂a⊩x) ∣₁
-                                      ; (inr (pr₁a≡k' , pr₂a⊩y⊔z))  ⊥rec (Trivial.k≠k' ca isNonTrivial (k ≡⟨ sym pr₁a≡true  pr₁  a ≡⟨ pr₁a≡k'  k' )) }) ∣₁) ∣₁ ∣₁
-                       ; (inr (pr₁a≡false , pr₂a⊩y⊔z))
-                       pr₂a⊩y⊔z >>=
-                        { (inl (pr₁pr₂a≡k  , pr₂pr₂a⊩y)) 
-                             inl (Pr₁a≡false.pr₁proof≡true a pr₁a≡false pr₁pr₂a≡k ,
-                               inr
-                              (Pr₁a≡false.pr₁pr₂proof≡k' a pr₁a≡false pr₁pr₂a≡k ,
-                              subst
-                                 r  r   y  x')
-                                (sym (Pr₁a≡false.pr₂pr₂proof≡pr₂pr₂a a pr₁a≡false pr₁pr₂a≡k))
-                                pr₂pr₂a⊩y) ∣₁) ∣₁
-                          ; (inr (pr₁pr₂a≡k' , pr₂pr₂a⊩z)) 
-                             inr (
-                              Pr₁a≡false.pr₁proof≡false a pr₁a≡false pr₁pr₂a≡k' ,
-                              subst
-                                 r  r   z  x')
-                                (sym (Pr₁a≡false.pr₂proof≡pr₂pr₂a a pr₁a≡false pr₁pr₂a≡k'))
-                                pr₂pr₂a⊩z) ∣₁ }) ∣₁
-                                })))) where open →⊔Assoc x y z
+        prover : Term as 1
+        prover =
+          ` Id ̇
+            (` pr₁ ̇ (# zero)) ̇
+            (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# zero)))) ̇
+            (` Id ̇
+              (` pr₁ ̇ (` pr₂ ̇ (# zero))) ̇
+              (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))) ̇
+              (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero))))
+      return
+        (λ* prover ,
+          λ x' a a⊩x⊔_y⊔z 
+            transport
+              (propTruncIdempotent (((x  y)  z) .isPropValued x' (λ* prover  a)))
+              (a⊩x⊔_y⊔z
+                >>=  { (inl (pr₁a≡true , pr₁a⊩x))
+                        inl
+                           (Pr₁a≡true.pr₁⨾proof≡true a pr₁a≡true ,
+                           transport
+                             (propTruncIdempotent isPropPropTrunc)
+                              a⊩x⊔_y⊔z
+                               >>=  { (inl (pr₁a≡k , pr₂a⊩x)) 
+                                          inl
+                                           (Pr₁a≡true.pr₁pr₂proof≡true a pr₁a≡true ,
+                                            subst
+                                               r  r   x  x')
+                                              (sym (Pr₁a≡true.pr₂pr₂proof≡pr₂a a pr₁a≡true))
+                                              pr₂a⊩x) ∣₁
+                                      ; (inr (pr₁a≡k' , pr₂a⊩y⊔z))  ⊥rec (Trivial.k≠k' ca isNonTrivial (k ≡⟨ sym pr₁a≡true  pr₁  a ≡⟨ pr₁a≡k'  k' )) }) ∣₁) ∣₁ ∣₁
+                       ; (inr (pr₁a≡false , pr₂a⊩y⊔z))
+                       pr₂a⊩y⊔z >>=
+                        { (inl (pr₁pr₂a≡k  , pr₂pr₂a⊩y)) 
+                             inl (Pr₁a≡false.pr₁proof≡true a pr₁a≡false pr₁pr₂a≡k ,
+                               inr
+                              (Pr₁a≡false.pr₁pr₂proof≡k' a pr₁a≡false pr₁pr₂a≡k ,
+                              subst
+                                 r  r   y  x')
+                                (sym (Pr₁a≡false.pr₂pr₂proof≡pr₂pr₂a a pr₁a≡false pr₁pr₂a≡k))
+                                pr₂pr₂a⊩y) ∣₁) ∣₁
+                          ; (inr (pr₁pr₂a≡k' , pr₂pr₂a⊩z)) 
+                             inr (
+                              Pr₁a≡false.pr₁proof≡false a pr₁a≡false pr₁pr₂a≡k' ,
+                              subst
+                                 r  r   z  x')
+                                (sym (Pr₁a≡false.pr₂proof≡pr₂pr₂a a pr₁a≡false pr₁pr₂a≡k'))
+                                pr₂pr₂a⊩z) ∣₁ }) ∣₁
+                                })))) where open →⊔Assoc x y z
 
-  `if_then_else_ :  {n}  Term as n  Term as n  Term as n  Term as n
-  `if c then t else e = ` Id ̇ c ̇ t ̇ e
+  `if_then_else_ :  {n}  Term as n  Term as n  Term as n  Term as n
+  `if c then t else e = ` Id ̇ c ̇ t ̇ e
 
-  x⊔y_⊔z≤x⊔_y⊔z :  x y z  (x  y)  z  x  (y  z)
-  x⊔y_⊔z≤x⊔_y⊔z x y z =
-    do
-      let
-        proof : Term as 1
-        proof =
-          `if ` pr₁ ̇ # fzero then
-            `if ` pr₁ ̇ (` pr₂ ̇ (# fzero)) then
-              ` pair ̇ ` true  ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))
-            else
-              (` pair ̇ ` false ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))))
-          else
-            (` pair ̇ ` false ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # fzero)))
-      return
-        (λ* proof ,
-           x' a a⊩x⊔y_⊔z 
-            a⊩x⊔y_⊔z >>=
-              let
-                mainThen =
-                  if (pr₁  (pr₂  a)) then
-                    (pair  true  (pr₂  (pr₂  a)))
-                  else
-                    (pair  false  (pair  true  (pr₂  (pr₂  a))))
+  x⊔y_⊔z≤x⊔_y⊔z :  x y z  (x  y)  z  x  (y  z)
+  x⊔y_⊔z≤x⊔_y⊔z x y z =
+    do
+      let
+        proof : Term as 1
+        proof =
+          `if ` pr₁ ̇ # zero then
+            `if ` pr₁ ̇ (` pr₂ ̇ (# zero)) then
+              ` pair ̇ ` true  ̇ (` pr₂ ̇ (` pr₂ ̇ # zero))
+            else
+              (` pair ̇ ` false ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (` pr₂ ̇ # zero))))
+          else
+            (` pair ̇ ` false ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # zero)))
+      return
+        (λ* proof ,
+           x' a a⊩x⊔y_⊔z 
+            a⊩x⊔y_⊔z >>=
+              let
+                mainThen =
+                  if (pr₁  (pr₂  a)) then
+                    (pair  true  (pr₂  (pr₂  a)))
+                  else
+                    (pair  false  (pair  true  (pr₂  (pr₂  a))))
 
-                mainElse =
-                  (pair  false  (pair  false  (pr₂  a)))
-              in
-              λ { (inl (pr₁a≡k , pr₂a⊩x⊔y)) 
+                mainElse =
+                  (pair  false  (pair  false  (pr₂  a)))
+              in
+              λ { (inl (pr₁a≡k , pr₂a⊩x⊔y)) 
                            
-                           let
-                             proofEq :
-                               (λ* proof  a)  if (pr₁  (pr₂  a)) then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a))))
-                             proofEq =
-                               λ* proof  a
-                                 ≡⟨ λ*ComputationRule proof (a  []) 
-                               if (pr₁  a) then
-                                 if (pr₁  (pr₂  a)) then
-                                   (pair  true  (pr₂  (pr₂  a)))
-                                 else
-                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
-                               else
-                                 (pair  false  (pair  false  (pr₂  a)))
-                                 ≡⟨ cong  x  if x then mainThen else mainElse) pr₁a≡k 
-                               if true then
-                                 if (pr₁  (pr₂  a)) then
-                                   (pair  true  (pr₂  (pr₂  a)))
-                                 else
-                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
-                               else
-                                 (pair  false  (pair  false  (pr₂  a)))
-                                 ≡⟨ ifTrueThen _ _ 
-                               refl
-                           in
-                           pr₂a⊩x⊔y >>=
-                             λ { (inl (pr₁pr₂a≡k , pr₂pr₂a⊩x)) 
-                               let
-                                 proof≡pair : λ* proof  a  pair  true  (pr₂  (pr₂  a))
-                                 proof≡pair =
-                                   λ* proof  a
-                                     ≡⟨ proofEq 
-                                   (if (pr₁  (pr₂  a)) then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
-                                     ≡⟨ cong  x  if x then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a))))) pr₁pr₂a≡k 
-                                   (if true then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
-                                     ≡⟨ ifTrueThen _ _ 
-                                   pair  true  (pr₂  (pr₂  a))
-                                     
+                           let
+                             proofEq :
+                               (λ* proof  a)  if (pr₁  (pr₂  a)) then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a))))
+                             proofEq =
+                               λ* proof  a
+                                 ≡⟨ λ*ComputationRule proof a 
+                               if (pr₁  a) then
+                                 if (pr₁  (pr₂  a)) then
+                                   (pair  true  (pr₂  (pr₂  a)))
+                                 else
+                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
+                               else
+                                 (pair  false  (pair  false  (pr₂  a)))
+                                 ≡⟨ cong  x  if x then mainThen else mainElse) pr₁a≡k 
+                               if true then
+                                 if (pr₁  (pr₂  a)) then
+                                   (pair  true  (pr₂  (pr₂  a)))
+                                 else
+                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
+                               else
+                                 (pair  false  (pair  false  (pr₂  a)))
+                                 ≡⟨ ifTrueThen _ _ 
+                               refl
+                           in
+                           pr₂a⊩x⊔y >>=
+                             λ { (inl (pr₁pr₂a≡k , pr₂pr₂a⊩x)) 
+                               let
+                                 proof≡pair : λ* proof  a  pair  true  (pr₂  (pr₂  a))
+                                 proof≡pair =
+                                   λ* proof  a
+                                     ≡⟨ proofEq 
+                                   (if (pr₁  (pr₂  a)) then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
+                                     ≡⟨ cong  x  if x then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a))))) pr₁pr₂a≡k 
+                                   (if true then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
+                                     ≡⟨ ifTrueThen _ _ 
+                                   pair  true  (pr₂  (pr₂  a))
+                                     
                                  
-                                 pr₁proofEq : pr₁  (λ* proof  a)  k
-                                 pr₁proofEq =
-                                   pr₁  (λ* proof  a)
-                                     ≡⟨ cong  x  pr₁  x) proof≡pair 
-                                   pr₁  (pair  true  (pr₂  (pr₂  a)))
-                                     ≡⟨ pr₁pxy≡x _ _ 
-                                   true
-                                     
+                                 pr₁proofEq : pr₁  (λ* proof  a)  k
+                                 pr₁proofEq =
+                                   pr₁  (λ* proof  a)
+                                     ≡⟨ cong  x  pr₁  x) proof≡pair 
+                                   pr₁  (pair  true  (pr₂  (pr₂  a)))
+                                     ≡⟨ pr₁pxy≡x _ _ 
+                                   true
+                                     
 
-                                 pr₂proofEq : pr₂  (λ* proof  a)  pr₂  (pr₂  a)
-                                 pr₂proofEq =
-                                   pr₂  (λ* proof  a)
-                                     ≡⟨ cong  x  pr₂  x) proof≡pair 
-                                   pr₂  (pair  true  (pr₂  (pr₂  a)))
-                                     ≡⟨ pr₂pxy≡y _ _ 
-                                   pr₂  (pr₂  a)
-                                     
-                               in  inl (pr₁proofEq , subst  r  r   x  x') (sym pr₂proofEq) pr₂pr₂a⊩x) ∣₁
-                               ; (inr (pr₁pr₂a≡k' , pr₂pr₂a⊩y)) 
-                               let
-                                 proof≡pair : λ* proof  a  pair  false  (pair  true  (pr₂  (pr₂  a)))
-                                 proof≡pair =
-                                   λ* proof  a
-                                     ≡⟨ proofEq 
-                                   (if (pr₁  (pr₂  a)) then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
-                                     ≡⟨ cong  x  if x then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a))))) pr₁pr₂a≡k' 
-                                   (if false then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
-                                     ≡⟨ ifFalseElse _ _ 
-                                   pair  false  (pair  true  (pr₂  (pr₂  a)))
-                                     
+                                 pr₂proofEq : pr₂  (λ* proof  a)  pr₂  (pr₂  a)
+                                 pr₂proofEq =
+                                   pr₂  (λ* proof  a)
+                                     ≡⟨ cong  x  pr₂  x) proof≡pair 
+                                   pr₂  (pair  true  (pr₂  (pr₂  a)))
+                                     ≡⟨ pr₂pxy≡y _ _ 
+                                   pr₂  (pr₂  a)
+                                     
+                               in  inl (pr₁proofEq , subst  r  r   x  x') (sym pr₂proofEq) pr₂pr₂a⊩x) ∣₁
+                               ; (inr (pr₁pr₂a≡k' , pr₂pr₂a⊩y)) 
+                               let
+                                 proof≡pair : λ* proof  a  pair  false  (pair  true  (pr₂  (pr₂  a)))
+                                 proof≡pair =
+                                   λ* proof  a
+                                     ≡⟨ proofEq 
+                                   (if (pr₁  (pr₂  a)) then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
+                                     ≡⟨ cong  x  if x then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a))))) pr₁pr₂a≡k' 
+                                   (if false then (pair  true  (pr₂  (pr₂  a))) else (pair  false  (pair  true  (pr₂  (pr₂  a)))))
+                                     ≡⟨ ifFalseElse _ _ 
+                                   pair  false  (pair  true  (pr₂  (pr₂  a)))
+                                     
 
-                                 pr₁proofEq : pr₁  (λ* proof  a)  k'
-                                 pr₁proofEq =
-                                   pr₁  (λ* proof  a)
-                                     ≡⟨ cong  x  pr₁  x) proof≡pair 
-                                   pr₁  (pair  false  (pair  true  (pr₂  (pr₂  a))))
-                                     ≡⟨ pr₁pxy≡x _ _ 
-                                   false
-                                     
+                                 pr₁proofEq : pr₁  (λ* proof  a)  k'
+                                 pr₁proofEq =
+                                   pr₁  (λ* proof  a)
+                                     ≡⟨ cong  x  pr₁  x) proof≡pair 
+                                   pr₁  (pair  false  (pair  true  (pr₂  (pr₂  a))))
+                                     ≡⟨ pr₁pxy≡x _ _ 
+                                   false
+                                     
 
-                                 pr₁pr₂proofEq : pr₁  (pr₂  (λ* proof  a))  true
-                                 pr₁pr₂proofEq =
-                                   pr₁  (pr₂  (λ* proof  a))
-                                     ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair 
-                                   pr₁  (pr₂  (pair  false  (pair  true  (pr₂  (pr₂  a)))))
-                                     ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
-                                   pr₁  (pair  true  (pr₂  (pr₂  a)))
-                                     ≡⟨ pr₁pxy≡x _ _ 
-                                   true
-                                     
+                                 pr₁pr₂proofEq : pr₁  (pr₂  (λ* proof  a))  true
+                                 pr₁pr₂proofEq =
+                                   pr₁  (pr₂  (λ* proof  a))
+                                     ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair 
+                                   pr₁  (pr₂  (pair  false  (pair  true  (pr₂  (pr₂  a)))))
+                                     ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                                   pr₁  (pair  true  (pr₂  (pr₂  a)))
+                                     ≡⟨ pr₁pxy≡x _ _ 
+                                   true
+                                     
 
-                                 pr₂pr₂proofEq : pr₂  (pr₂  (λ* proof  a))  pr₂  (pr₂  a)
-                                 pr₂pr₂proofEq =
-                                   pr₂  (pr₂  (λ* proof  a))
-                                     ≡⟨ cong  x  pr₂  (pr₂  x)) proof≡pair 
-                                   pr₂  (pr₂  (pair  false  (pair  true  (pr₂  (pr₂  a)))))
-                                     ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
-                                   pr₂  (pair  true  (pr₂  (pr₂  a)))
-                                     ≡⟨ pr₂pxy≡y _ _ 
-                                   pr₂  (pr₂  a)
-                                     
-                               in  inr (pr₁proofEq ,  inl (pr₁pr₂proofEq , subst  r  r   y  x') (sym pr₂pr₂proofEq) pr₂pr₂a⊩y) ∣₁) ∣₁ }
-                           ; (inr (pr₁a≡k' , pr₂a⊩z)) 
-                           let
-                             proof≡pair : λ* proof  a  pair  false  (pair  false  (pr₂  a))
-                             proof≡pair =
-                               λ* proof  a
-                                 ≡⟨ λ*ComputationRule proof (a  []) 
-                               if (pr₁  a) then
-                                 if (pr₁  (pr₂  a)) then
-                                   (pair  true  (pr₂  (pr₂  a)))
-                                 else
-                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
-                               else
-                                 (pair  false  (pair  false  (pr₂  a)))
-                                 ≡⟨ cong  x  if x then mainThen else mainElse) pr₁a≡k' 
-                               if false then
-                                 if (pr₁  (pr₂  a)) then
-                                   (pair  true  (pr₂  (pr₂  a)))
-                                 else
-                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
-                               else
-                                 (pair  false  (pair  false  (pr₂  a)))
-                                 ≡⟨ ifFalseElse _ _ 
-                               (pair  false  (pair  false  (pr₂  a)))
-                                 
+                                 pr₂pr₂proofEq : pr₂  (pr₂  (λ* proof  a))  pr₂  (pr₂  a)
+                                 pr₂pr₂proofEq =
+                                   pr₂  (pr₂  (λ* proof  a))
+                                     ≡⟨ cong  x  pr₂  (pr₂  x)) proof≡pair 
+                                   pr₂  (pr₂  (pair  false  (pair  true  (pr₂  (pr₂  a)))))
+                                     ≡⟨ cong  x  pr₂  x) (pr₂pxy≡y _ _) 
+                                   pr₂  (pair  true  (pr₂  (pr₂  a)))
+                                     ≡⟨ pr₂pxy≡y _ _ 
+                                   pr₂  (pr₂  a)
+                                     
+                               in  inr (pr₁proofEq ,  inl (pr₁pr₂proofEq , subst  r  r   y  x') (sym pr₂pr₂proofEq) pr₂pr₂a⊩y) ∣₁) ∣₁ }
+                           ; (inr (pr₁a≡k' , pr₂a⊩z)) 
+                           let
+                             proof≡pair : λ* proof  a  pair  false  (pair  false  (pr₂  a))
+                             proof≡pair =
+                               λ* proof  a
+                                 ≡⟨ λ*ComputationRule proof a 
+                               if (pr₁  a) then
+                                 if (pr₁  (pr₂  a)) then
+                                   (pair  true  (pr₂  (pr₂  a)))
+                                 else
+                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
+                               else
+                                 (pair  false  (pair  false  (pr₂  a)))
+                                 ≡⟨ cong  x  if x then mainThen else mainElse) pr₁a≡k' 
+                               if false then
+                                 if (pr₁  (pr₂  a)) then
+                                   (pair  true  (pr₂  (pr₂  a)))
+                                 else
+                                   (pair  false  (pair  true  (pr₂  (pr₂  a))))
+                               else
+                                 (pair  false  (pair  false  (pr₂  a)))
+                                 ≡⟨ ifFalseElse _ _ 
+                               (pair  false  (pair  false  (pr₂  a)))
+                                 
 
-                             pr₁proof≡false : pr₁  (λ* proof  a)  false
-                             pr₁proof≡false =
-                               pr₁  (λ* proof  a)
-                                 ≡⟨ cong  x  pr₁  x) proof≡pair 
-                               pr₁  (pair  false  (pair  false  (pr₂  a)))
-                                 ≡⟨ pr₁pxy≡x _ _ 
-                               false
-                                 
+                             pr₁proof≡false : pr₁  (λ* proof  a)  false
+                             pr₁proof≡false =
+                               pr₁  (λ* proof  a)
+                                 ≡⟨ cong  x  pr₁  x) proof≡pair 
+                               pr₁  (pair  false  (pair  false  (pr₂  a)))
+                                 ≡⟨ pr₁pxy≡x _ _ 
+                               false
+                                 
 
-                             pr₂proof≡pair : pr₂  (λ* proof  a)  (pair  false  (pr₂  a))
-                             pr₂proof≡pair =
-                               (pr₂  (λ* proof  a))
-                                 ≡⟨ cong  x  (pr₂  x)) proof≡pair 
-                               (pr₂  (pair  false  (pair  false  (pr₂  a))))
-                                 ≡⟨ (pr₂pxy≡y _ _) 
-                               (pair  false  (pr₂  a))
-                                 
+                             pr₂proof≡pair : pr₂  (λ* proof  a)  (pair  false  (pr₂  a))
+                             pr₂proof≡pair =
+                               (pr₂  (λ* proof  a))
+                                 ≡⟨ cong  x  (pr₂  x)) proof≡pair 
+                               (pr₂  (pair  false  (pair  false  (pr₂  a))))
+                                 ≡⟨ (pr₂pxy≡y _ _) 
+                               (pair  false  (pr₂  a))
+                                 
 
-                             pr₁pr₂proof≡false : pr₁  (pr₂  (λ* proof  a))  false
-                             pr₁pr₂proof≡false =
-                               pr₁  (pr₂  (λ* proof  a))
-                                 ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair 
-                               pr₁  (pr₂  (pair  false  (pair  false  (pr₂  a))))
-                                 ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
-                               pr₁  (pair  false  (pr₂  a))
-                                 ≡⟨ pr₁pxy≡x _ _ 
-                               false
-                                 
+                             pr₁pr₂proof≡false : pr₁  (pr₂  (λ* proof  a))  false
+                             pr₁pr₂proof≡false =
+                               pr₁  (pr₂  (λ* proof  a))
+                                 ≡⟨ cong  x  pr₁  (pr₂  x)) proof≡pair 
+                               pr₁  (pr₂  (pair  false  (pair  false  (pr₂  a))))
+                                 ≡⟨ cong  x  pr₁  x) (pr₂pxy≡y _ _) 
+                               pr₁  (pair  false  (pr₂  a))
+                                 ≡⟨ pr₁pxy≡x _ _ 
+                               false
+                                 
 
-                             pr₂pr₂proof≡pr₂a : pr₂  (pr₂  (λ* proof  a))  pr₂  a
-                             pr₂pr₂proof≡pr₂a =
-                               pr₂  (pr₂  (λ* proof  a))
-                                 ≡⟨ cong  x  pr₂  x ) pr₂proof≡pair 
-                               pr₂  (pair  false  (pr₂  a))
-                                 ≡⟨ pr₂pxy≡y _ _ 
-                               pr₂  a
-                                 
-                           in  inr (pr₁proof≡false ,  inr (pr₁pr₂proof≡false , subst  r  r   z  x') (sym pr₂pr₂proof≡pr₂a) pr₂a⊩z) ∣₁) ∣₁ } ))
+                             pr₂pr₂proof≡pr₂a : pr₂  (pr₂  (λ* proof  a))  pr₂  a
+                             pr₂pr₂proof≡pr₂a =
+                               pr₂  (pr₂  (λ* proof  a))
+                                 ≡⟨ cong  x  pr₂  x ) pr₂proof≡pair 
+                               pr₂  (pair  false  (pr₂  a))
+                                 ≡⟨ pr₂pxy≡y _ _ 
+                               pr₂  a
+                                 
+                           in  inr (pr₁proof≡false ,  inr (pr₁pr₂proof≡false , subst  r  r   z  x') (sym pr₂pr₂proof≡pr₂a) pr₂a⊩z) ∣₁) ∣₁ } ))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Joins.Base.html b/docs/Realizability.Tripos.Prealgebra.Joins.Base.html index 310ecf7..16c7598 100644 --- a/docs/Realizability.Tripos.Prealgebra.Joins.Base.html +++ b/docs/Realizability.Tripos.Prealgebra.Joins.Base.html @@ -2,96 +2,93 @@ Realizability.Tripos.Prealgebra.Joins.Base
open import Cubical.Foundations.Prelude
 open import Cubical.Foundations.Equiv
 open import Cubical.Data.Unit
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum renaming (rec to sumRec)
-open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum renaming (rec to sumRec)
+open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Tripos.Prealgebra.Joins.Base {} {A : Type } (ca : CombinatoryAlgebra A) where
-open import Realizability.Tripos.Prealgebra.Predicate ca
-open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module Realizability.Tripos.Prealgebra.Joins.Base { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
-
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
-  PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  antiSym→x⊔z≤y⊔z' :  x y z  x  y  y  x  (x  z)  (y  z)
-  antiSym→x⊔z≤y⊔z' x y z x≤y y≤x =
-    do
-      (x⊨y , x⊨yProves)  x≤y
-      let prover = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# fzero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# fzero)))
-      return
-        (λ* prover ,
-          λ x' a a⊩x⊔zx' 
-            equivFun
-              (propTruncIdempotent≃
-                ((y  z) .isPropValued x' (λ* prover  a)))
-                (do
-                  a⊩x⊔z  a⊩x⊔zx'
-                  return
-                     sumRec
-                       { (pr₁⨾a≡k , pr₂⨾a⊩xx') 
-                        inl
-                          (((pr₁  (λ* prover  a))
-                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (a  [])) 
-                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
-                            (pr₁  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  pr₁  x) (ifTrueThen _ _) 
-                            (pr₁  (pair  k  (x⊨y  (pr₂  a))))
-                             ≡⟨ pr₁pxy≡x _ _ 
-                            (k)
-                             ) ,
-                             subst
-                                r  r   y  x')
-                               (sym
-                                 (pr₂  (λ* prover  a)
-                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (a  [])) 
-                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
-                                  pr₂  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  pr₂  x) (ifTrueThen _ _) 
-                                  pr₂  (pair  k  (x⊨y  (pr₂  a)))
-                                   ≡⟨ pr₂pxy≡y _ _ 
-                                  x⊨y  (pr₂  a)
-                                   ))
-                               (x⊨yProves x' (pr₂  a) pr₂⨾a⊩xx')) })
-                       { (pr₁⨾a≡k' , pr₂a⊩zx') 
-                        inr
-                         ((((pr₁  (λ* prover  a))
-                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (a  [])) 
-                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
-                            (pr₁  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  pr₁  x) (ifFalseElse _ _) 
-                            (pr₁  (pair  k'  (pr₂  a)))
-                             ≡⟨ pr₁pxy≡x _ _ 
-                            (k')
-                             )) ,
-                               subst
-                                  r  r   z  x')
-                                 (sym
-                                  ((pr₂  (λ* prover  a)
-                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (a  [])) 
-                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
-                                  pr₂  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  pr₂  x) (ifFalseElse _ _) 
-                                  pr₂  (pair  k'  (pr₂  a))
-                                   ≡⟨ pr₂pxy≡y _ _ 
-                                  pr₂  a
-                                    )))
-                                 pr₂a⊩zx') })
-                      a⊩x⊔z ∣₁))
+module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
+  PredicateX = Predicate X
+  open Predicate
+  open PredicateProperties X
+  antiSym→x⊔z≤y⊔z' :  x y z  x  y  y  x  (x  z)  (y  z)
+  antiSym→x⊔z≤y⊔z' x y z x≤y y≤x =
+    do
+      (x⊨y , x⊨yProves)  x≤y
+      let prover = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# zero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# zero)))
+      return
+        (λ* prover ,
+          λ x' a a⊩x⊔zx' 
+            equivFun
+              (propTruncIdempotent≃
+                ((y  z) .isPropValued x' (λ* prover  a)))
+                (do
+                  a⊩x⊔z  a⊩x⊔zx'
+                  return
+                     sumRec
+                       { (pr₁⨾a≡k , pr₂⨾a⊩xx') 
+                        inl
+                          (((pr₁  (λ* prover  a))
+                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover a) 
+                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
+                            (pr₁  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  pr₁  x) (ifTrueThen _ _) 
+                            (pr₁  (pair  k  (x⊨y  (pr₂  a))))
+                             ≡⟨ pr₁pxy≡x _ _ 
+                            (k)
+                             ) ,
+                             subst
+                                r  r   y  x')
+                               (sym
+                                 (pr₂  (λ* prover  a)
+                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover a) 
+                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
+                                  pr₂  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  pr₂  x) (ifTrueThen _ _) 
+                                  pr₂  (pair  k  (x⊨y  (pr₂  a)))
+                                   ≡⟨ pr₂pxy≡y _ _ 
+                                  x⊨y  (pr₂  a)
+                                   ))
+                               (x⊨yProves x' (pr₂  a) pr₂⨾a⊩xx')) })
+                       { (pr₁⨾a≡k' , pr₂a⊩zx') 
+                        inr
+                         ((((pr₁  (λ* prover  a))
+                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover a) 
+                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
+                            (pr₁  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  pr₁  x) (ifFalseElse _ _) 
+                            (pr₁  (pair  k'  (pr₂  a)))
+                             ≡⟨ pr₁pxy≡x _ _ 
+                            (k')
+                             )) ,
+                               subst
+                                  r  r   z  x')
+                                 (sym
+                                  ((pr₂  (λ* prover  a)
+                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover a) 
+                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
+                                  pr₂  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  pr₂  x) (ifFalseElse _ _) 
+                                  pr₂  (pair  k'  (pr₂  a))
+                                   ≡⟨ pr₂pxy≡y _ _ 
+                                  pr₂  a
+                                    )))
+                                 pr₂a⊩zx') })
+                      a⊩x⊔z ∣₁))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Joins.Commutativity.html b/docs/Realizability.Tripos.Prealgebra.Joins.Commutativity.html index 43ec9b1..65b075e 100644 --- a/docs/Realizability.Tripos.Prealgebra.Joins.Commutativity.html +++ b/docs/Realizability.Tripos.Prealgebra.Joins.Commutativity.html @@ -1,159 +1,154 @@ Realizability.Tripos.Prealgebra.Joins.Commutativity
open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
-open import Cubical.Foundations.Prelude
-open import Cubical.Foundations.Equiv
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum renaming (rec to sumRec)
-open import Cubical.Relation.Binary.Order.Preorder
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
+open import Realizability.ApplicativeStructure
+open import Cubical.Foundations.Prelude
+open import Cubical.Foundations.Equiv
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum renaming (rec to sumRec)
+open import Cubical.Relation.Binary.Order.Preorder
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
 
-module Realizability.Tripos.Prealgebra.Joins.Commutativity {} {A : Type } (ca : CombinatoryAlgebra A) where
+module Realizability.Tripos.Prealgebra.Joins.Commutativity {} {A : Type } (ca : CombinatoryAlgebra A) where
 
-open import Realizability.Tripos.Prealgebra.Predicate ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where
+  open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+  private PredicateX = Predicate X
+  open Predicate
+  open PredicateProperties  X
+  open PreorderReasoning preorder≤
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
+  -- ⊔ is commutative upto anti-symmetry
+  a⊔b→b⊔a :  a b  a  b  b  a
+  a⊔b→b⊔a a b =
+    do
+      let prover = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# zero))) ̇ (` pair ̇ ` k ̇ (` pr₂ ̇ (# zero)))
+      let λ*eq = λ (r : A)  λ*ComputationRule prover r
+      return
+        (λ* prover ,
+          λ x r r⊩a⊔b 
+            r⊩a⊔b >>=
+              λ { (inl (pr₁r≡k  , pr₂r⊩a)) 
+                   inr ((pr₁  (λ* prover  r)
+                           ≡⟨ cong  x  pr₁  x) (λ*eq r) 
+                          pr₁  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                           ≡⟨ cong  x  pr₁  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) (pr₁r≡k) 
+                          pr₁  (Id  k  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                           ≡⟨ cong  x  pr₁  x) (ifTrueThen _ _) 
+                          pr₁  (pair  k'  (pr₂  r))
+                           ≡⟨ pr₁pxy≡x _ _ 
+                          k'
+                           ) ,
+                        subst
+                           r  r   a  x)
+                          (sym
+                            ((pr₂  (λ* prover  r)
+                                ≡⟨ cong  x  pr₂  x) (λ*eq r) 
+                              pr₂  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                                ≡⟨ cong  x  pr₂  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) (pr₁r≡k) 
+                              pr₂  (Id  k  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                                ≡⟨ cong  x  pr₂  x) (ifTrueThen _ _) 
+                              pr₂  (pair  k'  (pr₂  r))
+                                ≡⟨ pr₂pxy≡y _ _ 
+                              pr₂  r
+                                )))
+                          pr₂r⊩a) ∣₁
+                ; (inr (pr₁r≡k' , pr₂r⊩b)) 
+                   inl (((pr₁  (λ* prover  r)
+                           ≡⟨ cong  x  pr₁  x) (λ*eq r) 
+                          pr₁  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                           ≡⟨ cong  x  pr₁  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) pr₁r≡k' 
+                          pr₁  (Id  k'  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                           ≡⟨ cong  x  pr₁  x) (ifFalseElse _ _) 
+                          pr₁  (pair  k  (pr₂  r))
+                           ≡⟨ pr₁pxy≡x _ _ 
+                          k
+                           ) ,
+                        subst
+                           r  r   b  x)
+                          (sym
+                            ((pr₂  (λ* prover  r)
+                                ≡⟨ cong  x  pr₂  x) (λ*eq r) 
+                              pr₂  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                                ≡⟨ cong  x  pr₂  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) pr₁r≡k' 
+                              pr₂  (Id  k'  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
+                                ≡⟨ cong  x  pr₂  x) (ifFalseElse _ _) 
+                              pr₂  (pair  k  (pr₂  r))
+                                ≡⟨ pr₂pxy≡y _ _ 
+                              pr₂  r
+                                )))
+                          pr₂r⊩b)) ∣₁ })
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where
-  
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  open PreorderReasoning preorder≤
-
-  -- ⊔ is commutative upto anti-symmetry
-  a⊔b→b⊔a :  a b  a  b  b  a
-  a⊔b→b⊔a a b =
-    do
-      let prover = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# fzero))) ̇ (` pair ̇ ` k ̇ (` pr₂ ̇ (# fzero)))
-      let λ*eq = λ (r : A)  λ*ComputationRule prover (r  [])
-      return
-        (λ* prover ,
-          λ x r r⊩a⊔b 
-            r⊩a⊔b >>=
-              λ { (inl (pr₁r≡k  , pr₂r⊩a)) 
-                   inr ((pr₁  (λ* prover  r)
-                           ≡⟨ cong  x  pr₁  x) (λ*eq r) 
-                          pr₁  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                           ≡⟨ cong  x  pr₁  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) (pr₁r≡k) 
-                          pr₁  (Id  k  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                           ≡⟨ cong  x  pr₁  x) (ifTrueThen _ _) 
-                          pr₁  (pair  k'  (pr₂  r))
-                           ≡⟨ pr₁pxy≡x _ _ 
-                          k'
-                           ) ,
-                        subst
-                           r  r   a  x)
-                          (sym
-                            ((pr₂  (λ* prover  r)
-                                ≡⟨ cong  x  pr₂  x) (λ*eq r) 
-                              pr₂  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                                ≡⟨ cong  x  pr₂  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) (pr₁r≡k) 
-                              pr₂  (Id  k  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                                ≡⟨ cong  x  pr₂  x) (ifTrueThen _ _) 
-                              pr₂  (pair  k'  (pr₂  r))
-                                ≡⟨ pr₂pxy≡y _ _ 
-                              pr₂  r
-                                )))
-                          pr₂r⊩a) ∣₁
-                ; (inr (pr₁r≡k' , pr₂r⊩b)) 
-                   inl (((pr₁  (λ* prover  r)
-                           ≡⟨ cong  x  pr₁  x) (λ*eq r) 
-                          pr₁  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                           ≡⟨ cong  x  pr₁  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) pr₁r≡k' 
-                          pr₁  (Id  k'  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                           ≡⟨ cong  x  pr₁  x) (ifFalseElse _ _) 
-                          pr₁  (pair  k  (pr₂  r))
-                           ≡⟨ pr₁pxy≡x _ _ 
-                          k
-                           ) ,
-                        subst
-                           r  r   b  x)
-                          (sym
-                            ((pr₂  (λ* prover  r)
-                                ≡⟨ cong  x  pr₂  x) (λ*eq r) 
-                              pr₂  (Id  (pr₁  r)  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                                ≡⟨ cong  x  pr₂  (Id  x  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))) pr₁r≡k' 
-                              pr₂  (Id  k'  (pair  k'  (pr₂  r))  (pair  k  (pr₂  r)))
-                                ≡⟨ cong  x  pr₂  x) (ifFalseElse _ _) 
-                              pr₂  (pair  k  (pr₂  r))
-                                ≡⟨ pr₂pxy≡y _ _ 
-                              pr₂  r
-                                )))
-                          pr₂r⊩b)) ∣₁ })
-
-  antiSym→x⊔z≤y⊔z :  x y z  x  y  y  x  (x  z)  (y  z)
-  antiSym→x⊔z≤y⊔z x y z x≤y y≤x =
-    do
-      (x⊨y , x⊨yProves)  x≤y
-      let prover = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# fzero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# fzero)))
-      return
-        (λ* prover ,
-          λ x' a a⊩x⊔zx' 
-            equivFun
-              (propTruncIdempotent≃
-                ((y  z) .isPropValued x' (λ* prover  a)))
-                (do
-                  a⊩x⊔z  a⊩x⊔zx'
-                  return
-                     sumRec
-                       { (pr₁⨾a≡k , pr₂⨾a⊩xx') 
-                        inl
-                          (((pr₁  (λ* prover  a))
-                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (a  [])) 
-                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
-                            (pr₁  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  pr₁  x) (ifTrueThen _ _) 
-                            (pr₁  (pair  k  (x⊨y  (pr₂  a))))
-                             ≡⟨ pr₁pxy≡x _ _ 
-                            (k)
-                             ) ,
-                             subst
-                                r  r   y  x')
-                               (sym
-                                 (pr₂  (λ* prover  a)
-                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (a  [])) 
-                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
-                                  pr₂  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  pr₂  x) (ifTrueThen _ _) 
-                                  pr₂  (pair  k  (x⊨y  (pr₂  a)))
-                                   ≡⟨ pr₂pxy≡y _ _ 
-                                  x⊨y  (pr₂  a)
-                                   ))
-                               (x⊨yProves x' (pr₂  a) pr₂⨾a⊩xx')) })
-                       { (pr₁⨾a≡k' , pr₂a⊩zx') 
-                        inr
-                         ((((pr₁  (λ* prover  a))
-                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover (a  [])) 
-                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
-                            (pr₁  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
-                             ≡⟨ cong  x  pr₁  x) (ifFalseElse _ _) 
-                            (pr₁  (pair  k'  (pr₂  a)))
-                             ≡⟨ pr₁pxy≡x _ _ 
-                            (k')
-                             )) ,
-                               subst
-                                  r  r   z  x')
-                                 (sym
-                                  ((pr₂  (λ* prover  a)
-                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover (a  [])) 
-                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
-                                  pr₂  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
-                                   ≡⟨ cong  x  pr₂  x) (ifFalseElse _ _) 
-                                  pr₂  (pair  k'  (pr₂  a))
-                                   ≡⟨ pr₂pxy≡y _ _ 
-                                  pr₂  a
-                                    )))
-                                 pr₂a⊩zx') })
-                      a⊩x⊔z ∣₁))
+  antiSym→x⊔z≤y⊔z :  x y z  x  y  y  x  (x  z)  (y  z)
+  antiSym→x⊔z≤y⊔z x y z x≤y y≤x =
+    do
+      (x⊨y , x⊨yProves)  x≤y
+      let prover = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# zero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# zero)))
+      return
+        (λ* prover ,
+          λ x' a a⊩x⊔zx' 
+            equivFun
+              (propTruncIdempotent≃
+                ((y  z) .isPropValued x' (λ* prover  a)))
+                (do
+                  a⊩x⊔z  a⊩x⊔zx'
+                  return
+                     sumRec
+                       { (pr₁⨾a≡k , pr₂⨾a⊩xx') 
+                        inl
+                          (((pr₁  (λ* prover  a))
+                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover a) 
+                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
+                            (pr₁  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  pr₁  x) (ifTrueThen _ _) 
+                            (pr₁  (pair  k  (x⊨y  (pr₂  a))))
+                             ≡⟨ pr₁pxy≡x _ _ 
+                            (k)
+                             ) ,
+                             subst
+                                r  r   y  x')
+                               (sym
+                                 (pr₂  (λ* prover  a)
+                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover a) 
+                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k 
+                                  pr₂  (Id  k  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  pr₂  x) (ifTrueThen _ _) 
+                                  pr₂  (pair  k  (x⊨y  (pr₂  a)))
+                                   ≡⟨ pr₂pxy≡y _ _ 
+                                  x⊨y  (pr₂  a)
+                                   ))
+                               (x⊨yProves x' (pr₂  a) pr₂⨾a⊩xx')) })
+                       { (pr₁⨾a≡k' , pr₂a⊩zx') 
+                        inr
+                         ((((pr₁  (λ* prover  a))
+                             ≡⟨ cong  x  pr₁  x) (λ*ComputationRule prover a) 
+                            (pr₁  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  (pr₁  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
+                            (pr₁  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))
+                             ≡⟨ cong  x  pr₁  x) (ifFalseElse _ _) 
+                            (pr₁  (pair  k'  (pr₂  a)))
+                             ≡⟨ pr₁pxy≡x _ _ 
+                            (k')
+                             )) ,
+                               subst
+                                  r  r   z  x')
+                                 (sym
+                                  ((pr₂  (λ* prover  a)
+                                   ≡⟨ cong  x  pr₂  x) (λ*ComputationRule prover a) 
+                                  pr₂  (Id  (pr₁  a)  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  (pr₂  (Id  x  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a))))) pr₁⨾a≡k' 
+                                  pr₂  (Id  k'  (pair  k  (x⊨y  (pr₂  a)))  (pair  k'  (pr₂  a)))
+                                   ≡⟨ cong  x  pr₂  x) (ifFalseElse _ _) 
+                                  pr₂  (pair  k'  (pr₂  a))
+                                   ≡⟨ pr₂pxy≡y _ _ 
+                                  pr₂  a
+                                    )))
+                                 pr₂a⊩zx') })
+                      a⊩x⊔z ∣₁))
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Joins.Idempotency.html b/docs/Realizability.Tripos.Prealgebra.Joins.Idempotency.html index 97ff8be..396191a 100644 --- a/docs/Realizability.Tripos.Prealgebra.Joins.Idempotency.html +++ b/docs/Realizability.Tripos.Prealgebra.Joins.Idempotency.html @@ -1,52 +1,49 @@ Realizability.Tripos.Prealgebra.Joins.Idempotency
open import Cubical.Foundations.Prelude
 open import Cubical.Data.Unit
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum
-open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum
+open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Tripos.Prealgebra.Joins.Idempotency {} {A : Type } (ca : CombinatoryAlgebra A) where
-open import Realizability.Tripos.Prealgebra.Predicate ca
-open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module Realizability.Tripos.Prealgebra.Joins.Idempotency { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
+module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
+  private PredicateX = Predicate X
+  open Predicate
+  open PredicateProperties X
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
+  x⊔x≤x :  x  x  x  x
+  x⊔x≤x x =
+      return
+        (pr₂ ,
+           x' a a⊩x⊔x 
+            transport
+              (propTruncIdempotent (x .isPropValued x' (pr₂  a)))
+              (do
+                a⊩x⊔x  a⊩x⊔x
+                let
+                  goal : ((pr₁  a  k) × (pr₂  a)   x  x')  ((pr₁  a  k') × (pr₂  a)   x  x')  (pr₂  a)   x  x'
+                  goal = λ { (inl (_ , pr₂a⊩x))  pr₂a⊩x ; (inr (_ , pr₂a⊩x))  pr₂a⊩x }
+                return (goal a⊩x⊔x))))
 
-  x⊔x≤x :  x  x  x  x
-  x⊔x≤x x =
-      return
-        (pr₂ ,
-           x' a a⊩x⊔x 
-            transport
-              (propTruncIdempotent (x .isPropValued x' (pr₂  a)))
-              (do
-                a⊩x⊔x  a⊩x⊔x
-                let
-                  goal : ((pr₁  a  k) × (pr₂  a)   x  x')  ((pr₁  a  k') × (pr₂  a)   x  x')  (pr₂  a)   x  x'
-                  goal = λ { (inl (_ , pr₂a⊩x))  pr₂a⊩x ; (inr (_ , pr₂a⊩x))  pr₂a⊩x }
-                return (goal a⊩x⊔x))))
-
-  x≤x⊔x :  x  x  x  x
-  x≤x⊔x x =
-    let prover = ` pair ̇ ` true ̇ # fzero in
-     λ* prover ,
-       x' a a⊩x 
-        subst
-           r  r   x  x  x')
-          (sym (λ*ComputationRule prover (a  [])))
-           inl (pr₁pxy≡x _ _ , subst  r  r   x  x') (sym (pr₂pxy≡y _ _)) a⊩x) ∣₁) ∣₁
+  x≤x⊔x :  x  x  x  x
+  x≤x⊔x x =
+    let prover = ` pair ̇ ` true ̇ # zero in
+     λ* prover ,
+       x' a a⊩x 
+        subst
+           r  r   x  x  x')
+          (sym (λ*ComputationRule prover a))
+           inl (pr₁pxy≡x _ _ , subst  r  r   x  x') (sym (pr₂pxy≡y _ _)) a⊩x) ∣₁) ∣₁
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Joins.Identity.html b/docs/Realizability.Tripos.Prealgebra.Joins.Identity.html index 7aa8ddf..d2aba48 100644 --- a/docs/Realizability.Tripos.Prealgebra.Joins.Identity.html +++ b/docs/Realizability.Tripos.Prealgebra.Joins.Identity.html @@ -1,120 +1,107 @@ -Realizability.Tripos.Prealgebra.Joins.Identity
{-# OPTIONS --allow-unsolved-metas #-}
-open import Cubical.Foundations.Prelude
-open import Cubical.Data.Unit
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum
-open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+Realizability.Tripos.Prealgebra.Joins.Identity
open import Cubical.Foundations.Prelude
+open import Cubical.Data.Unit
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum
+open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Tripos.Prealgebra.Joins.Identity {} {A : Type } (ca : CombinatoryAlgebra A) where
-open import Realizability.Tripos.Prealgebra.Predicate ca
-open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module Realizability.Tripos.Prealgebra.Joins.Identity { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
+module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
+  private PredicateX = Predicate  X
+  open Predicate
+  open PredicateProperties  X
+  open PreorderReasoning preorder≤
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  open PreorderReasoning preorder≤
+  pre0 : PredicateX
+  Predicate.isSetX pre0 = isSetX'
+   pre0  = λ x a  ⊥*
+  Predicate.isPropValued pre0 = λ x a  isProp⊥*
 
+  x⊔0≤x :  x  x  pre0  x
+  x⊔0≤x x =
+    do
+      return
+        (pr₂ ,
+           x' a a⊩x⊔0 
+            transport
+              (propTruncIdempotent (x .isPropValued x' (pr₂  a)))
+              (do
+                a⊩x⊔0  a⊩x⊔0
+                let
+                  goal : ((pr₁  a  k) × (pr₂  a)   x  x')  ((pr₁  a  k') × ⊥*)  (pr₂  a)   x  x'
+                  goal = λ { (inl (pr₁a≡k , pr₂a⊩x))  pr₂a⊩x ; (inr (_ , bottom))  ⊥*rec bottom }
+                return (goal a⊩x⊔0))))
 
-  pre0 : PredicateX
-  Predicate.isSetX pre0 = isSetX'
-   pre0  = λ x a  ⊥*
-  Predicate.isPropValued pre0 = λ x a  isProp⊥*
+  x≤0⊔x :  x  x  (pre0  x)
+  x≤0⊔x x =
+    let
+      proof : Term as 1
+      proof = ` pair ̇ ` false ̇ # zero
+    in
+      return
+        ((λ* proof) ,
+           x' a a⊩x 
+            let
+              pr₁proofEq : pr₁  (λ* proof  a)  false
+              pr₁proofEq =
+                pr₁  (λ* proof  a)
+                  ≡⟨ cong  x  pr₁  x) (λ*ComputationRule proof a) 
+                pr₁  (pair  false  a)
+                  ≡⟨ pr₁pxy≡x _ _ 
+                false
+                  
 
-  x⊔0≤x :  x  x  pre0  x
-  x⊔0≤x x =
-    do
-      return
-        (pr₂ ,
-           x' a a⊩x⊔0 
-            transport
-              (propTruncIdempotent (x .isPropValued x' (pr₂  a)))
-              (do
-                a⊩x⊔0  a⊩x⊔0
-                let
-                  goal : ((pr₁  a  k) × (pr₂  a)   x  x')  ((pr₁  a  k') × ⊥*)  (pr₂  a)   x  x'
-                  goal = λ { (inl (pr₁a≡k , pr₂a⊩x))  pr₂a⊩x ; (inr (_ , bottom))  ⊥*rec bottom }
-                return (goal a⊩x⊔0))))
+              pr₂proofEq : pr₂  (λ* proof  a)  a
+              pr₂proofEq =
+                pr₂  (λ* proof  a)
+                  ≡⟨ cong  x  pr₂  x) (λ*ComputationRule proof a) 
+                pr₂  (pair  false  a)
+                  ≡⟨ pr₂pxy≡y _ _ 
+                a
+                  
+            in
+             inr (pr₁proofEq , subst  r  r   x  x') (sym pr₂proofEq) a⊩x) ∣₁))
 
-  x≤0⊔x :  x  x  (pre0  x)
-  x≤0⊔x x =
-    let
-      proof : Term as 1
-      proof = ` pair ̇ ` false ̇ # fzero
-    in
-      return
-        ((λ* proof) ,
-           x' a a⊩x 
-            let
-              pr₁proofEq : pr₁  (λ* proof  a)  false
-              pr₁proofEq =
-                pr₁  (λ* proof  a)
-                  ≡⟨ cong  x  pr₁  x) (λ*ComputationRule proof (a  [])) 
-                pr₁  (pair  false  a)
-                  ≡⟨ pr₁pxy≡x _ _ 
-                false
-                  
+  x≤x⊔0 :  x  x  x  pre0
+  x≤x⊔0 x =
+    let
+      proof : Term as 1
+      proof = ` pair ̇ ` true ̇ # zero
+    in return
+      ((λ* proof) ,
+         x' a a⊩x 
+          let
+            pr₁proofEq : pr₁  (λ* proof  a)  true
+            pr₁proofEq =
+              pr₁  (λ* proof  a)
+                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule proof a) 
+              pr₁  (pair  true  a)
+                ≡⟨ pr₁pxy≡x _ _ 
+              true
+                
 
-              pr₂proofEq : pr₂  (λ* proof  a)  a
-              pr₂proofEq =
-                pr₂  (λ* proof  a)
-                  ≡⟨ cong  x  pr₂  x) (λ*ComputationRule proof (a  [])) 
-                pr₂  (pair  false  a)
-                  ≡⟨ pr₂pxy≡y _ _ 
-                a
-                  
-            in
-             inr (pr₁proofEq , subst  r  r   x  x') (sym pr₂proofEq) a⊩x) ∣₁))
-
-  x≤x⊔0 :  x  x  x  pre0
-  x≤x⊔0 x =
-    let
-      proof : Term as 1
-      proof = ` pair ̇ ` true ̇ # fzero
-    in return
-      ((λ* proof) ,
-         x' a a⊩x 
-          let
-            pr₁proofEq : pr₁  (λ* proof  a)  true
-            pr₁proofEq =
-              pr₁  (λ* proof  a)
-                ≡⟨ cong  x  pr₁  x) (λ*ComputationRule proof (a  [])) 
-              pr₁  (pair  true  a)
-                ≡⟨ pr₁pxy≡x _ _ 
-              true
-                
-
-            pr₂proofEq : pr₂  (λ* proof  a)  a
-            pr₂proofEq =
-              pr₂  (λ* proof  a)
-                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule proof (a  [])) 
-              pr₂  (pair  true  a)
-                ≡⟨ pr₂pxy≡y _ _ 
-              a
-                
-          in
-           inl (pr₁proofEq , subst  r  r   x  x') (sym pr₂proofEq) a⊩x) ∣₁))
-  0⊔x≤x :  x  pre0  x  x
-  0⊔x≤x x =
-    pre0  x
-      ≲⟨ a⊔b→b⊔a X isSetX' pre0 x 
-    x  pre0
-      ≲⟨ x⊔0≤x x 
-    x
-      
+            pr₂proofEq : pr₂  (λ* proof  a)  a
+            pr₂proofEq =
+              pr₂  (λ* proof  a)
+                ≡⟨ cong  x  pr₂  x) (λ*ComputationRule proof a) 
+              pr₂  (pair  true  a)
+                ≡⟨ pr₂pxy≡y _ _ 
+              a
+                
+          in
+           inl (pr₁proofEq , subst  r  r   x  x') (sym pr₂proofEq) a⊩x) ∣₁))
 
   
   
diff --git a/docs/Realizability.Tripos.Prealgebra.Meets.Associativity.html b/docs/Realizability.Tripos.Prealgebra.Meets.Associativity.html
index 2c2d682..c560ab1 100644
--- a/docs/Realizability.Tripos.Prealgebra.Meets.Associativity.html
+++ b/docs/Realizability.Tripos.Prealgebra.Meets.Associativity.html
@@ -1,136 +1,132 @@
 
 Realizability.Tripos.Prealgebra.Meets.Associativity
open import Cubical.Foundations.Prelude
 open import Cubical.Data.Unit
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum
-open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum
+open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Tripos.Prealgebra.Meets.Associativity {} {A : Type } (ca : CombinatoryAlgebra A) where
-open import Realizability.Tripos.Prealgebra.Predicate ca
-open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module Realizability.Tripos.Prealgebra.Meets.Associativity { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
+module _  (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
+  open Predicate
+  open PredicateProperties X
+  open PreorderReasoning preorder≤
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  open PreorderReasoning preorder≤
+  x⊓_y⊓z≤x⊓y_⊓z :  x y z  x  (y  z)  (x  y)  z
+  x⊓_y⊓z≤x⊓y_⊓z x y z =
+    let
+      proof : Term as 1
+      proof = ` pair ̇ (` pair ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero))
+    in
+      return
+        (λ* proof ,
+          λ x' a a⊩x⊓_y⊓z 
+            let
+              λ*eq = λ*ComputationRule proof a
+              
+              pr₁proofEq : pr₁  (λ* proof  a)  pair  (pr₁  a)  (pr₁  (pr₂  a))
+              pr₁proofEq =
+                pr₁  (λ* proof  a)
+                  ≡⟨ cong  x  pr₁  x) λ*eq 
+                pr₁  (pair  (pair  (pr₁  a)  (pr₁  (pr₂  a)))  (pr₂  (pr₂  a)))
+                  ≡⟨ pr₁pxy≡x _ _ 
+                pair  (pr₁  a)  (pr₁  (pr₂  a))
+                  
 
-  x⊓_y⊓z≤x⊓y_⊓z :  x y z  x  (y  z)  (x  y)  z
-  x⊓_y⊓z≤x⊓y_⊓z x y z =
-    let
-      proof : Term as 1
-      proof = ` pair ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))
-    in
-      return
-        (λ* proof ,
-          λ x' a a⊩x⊓_y⊓z 
-            let
-              λ*eq = λ*ComputationRule proof (a  [])
-              -- Unfortunately, proof assistants
-              pr₁proofEq : pr₁  (λ* proof  a)  pair  (pr₁  a)  (pr₁  (pr₂  a))
-              pr₁proofEq =
-                pr₁  (λ* proof  a)
-                  ≡⟨ cong  x  pr₁  x) λ*eq 
-                pr₁  (pair  (pair  (pr₁  a)  (pr₁  (pr₂  a)))  (pr₂  (pr₂  a)))
-                  ≡⟨ pr₁pxy≡x _ _ 
-                pair  (pr₁  a)  (pr₁  (pr₂  a))
-                  
+              pr₁pr₁proofEq : pr₁  (pr₁  (λ* proof  a))  pr₁  a
+              pr₁pr₁proofEq =
+                pr₁  (pr₁  (λ* proof  a))
+                  ≡⟨ cong  x  pr₁  x) pr₁proofEq 
+                pr₁  (pair  (pr₁  a)  (pr₁  (pr₂  a)))
+                  ≡⟨ pr₁pxy≡x _ _ 
+                pr₁  a
+                  
 
-              pr₁pr₁proofEq : pr₁  (pr₁  (λ* proof  a))  pr₁  a
-              pr₁pr₁proofEq =
-                pr₁  (pr₁  (λ* proof  a))
-                  ≡⟨ cong  x  pr₁  x) pr₁proofEq 
-                pr₁  (pair  (pr₁  a)  (pr₁  (pr₂  a)))
-                  ≡⟨ pr₁pxy≡x _ _ 
-                pr₁  a
-                  
+              pr₂pr₁proofEq : pr₂  (pr₁  (λ* proof  a))  pr₁  (pr₂  a)
+              pr₂pr₁proofEq =
+                pr₂  (pr₁  (λ* proof  a))
+                  ≡⟨ cong  x  pr₂  x) pr₁proofEq 
+                pr₂  (pair  (pr₁  a)  (pr₁  (pr₂  a)))
+                  ≡⟨ pr₂pxy≡y _ _ 
+                pr₁  (pr₂  a)
+                  
 
-              pr₂pr₁proofEq : pr₂  (pr₁  (λ* proof  a))  pr₁  (pr₂  a)
-              pr₂pr₁proofEq =
-                pr₂  (pr₁  (λ* proof  a))
-                  ≡⟨ cong  x  pr₂  x) pr₁proofEq 
-                pr₂  (pair  (pr₁  a)  (pr₁  (pr₂  a)))
-                  ≡⟨ pr₂pxy≡y _ _ 
-                pr₁  (pr₂  a)
-                  
+              pr₂proofEq : pr₂  (λ* proof  a)  pr₂  (pr₂  a)
+              pr₂proofEq =
+                pr₂  (λ* proof  a)
+                  ≡⟨ cong  x  pr₂  x) λ*eq 
+                pr₂  (pair  (pair  (pr₁  a)  (pr₁  (pr₂  a)))  (pr₂  (pr₂  a)))
+                  ≡⟨ pr₂pxy≡y _ _ 
+                pr₂  (pr₂  a)
+                  
+            in
+              (subst  r  r   x  x') (sym pr₁pr₁proofEq) (a⊩x⊓_y⊓z .fst) ,
+               subst  r  r   y  x') (sym pr₂pr₁proofEq) (a⊩x⊓_y⊓z .snd .fst)) ,
+               subst  r  r   z  x') (sym pr₂proofEq) (a⊩x⊓_y⊓z .snd .snd))
 
-              pr₂proofEq : pr₂  (λ* proof  a)  pr₂  (pr₂  a)
-              pr₂proofEq =
-                pr₂  (λ* proof  a)
-                  ≡⟨ cong  x  pr₂  x) λ*eq 
-                pr₂  (pair  (pair  (pr₁  a)  (pr₁  (pr₂  a)))  (pr₂  (pr₂  a)))
-                  ≡⟨ pr₂pxy≡y _ _ 
-                pr₂  (pr₂  a)
-                  
-            in
-              (subst  r  r   x  x') (sym pr₁pr₁proofEq) (a⊩x⊓_y⊓z .fst) ,
-               subst  r  r   y  x') (sym pr₂pr₁proofEq) (a⊩x⊓_y⊓z .snd .fst)) ,
-               subst  r  r   z  x') (sym pr₂proofEq) (a⊩x⊓_y⊓z .snd .snd))
+  x⊓y_⊓z≤x⊓_y⊓z :  x y z  (x  y)  z  x  (y  z)
+  x⊓y_⊓z≤x⊓_y⊓z x y z =
+    let
+      proof : Term as 1
+      proof = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero))
+    in
+      return
+        (λ* proof ,
+          λ { x' a ((pr₁pr₁a⊩x , pr₂pr₁a⊩y) , pr₂a⊩z) 
+            let
+              λ*eq = λ*ComputationRule proof a
 
-  x⊓y_⊓z≤x⊓_y⊓z :  x y z  (x  y)  z  x  (y  z)
-  x⊓y_⊓z≤x⊓_y⊓z x y z =
-    let
-      proof : Term as 1
-      proof = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero))
-    in
-      return
-        (λ* proof ,
-          λ { x' a ((pr₁pr₁a⊩x , pr₂pr₁a⊩y) , pr₂a⊩z) 
-            let
-              λ*eq = λ*ComputationRule proof (a  [])
+              pr₂proofEq : pr₂  (λ* proof  a)  pair  (pr₂  (pr₁  a))  (pr₂  a)
+              pr₂proofEq =
+                pr₂  (λ* proof  a)
+                  ≡⟨ cong  x  pr₂  x) λ*eq 
+                pr₂  (pair  (pr₁  (pr₁  a))  (pair  (pr₂  (pr₁  a))  (pr₂  a)))
+                  ≡⟨ pr₂pxy≡y _ _ 
+                pair  (pr₂  (pr₁  a))  (pr₂  a)
+                  
 
-              pr₂proofEq : pr₂  (λ* proof  a)  pair  (pr₂  (pr₁  a))  (pr₂  a)
-              pr₂proofEq =
-                pr₂  (λ* proof  a)
-                  ≡⟨ cong  x  pr₂  x) λ*eq 
-                pr₂  (pair  (pr₁  (pr₁  a))  (pair  (pr₂  (pr₁  a))  (pr₂  a)))
-                  ≡⟨ pr₂pxy≡y _ _ 
-                pair  (pr₂  (pr₁  a))  (pr₂  a)
-                  
+              pr₁proofEq : pr₁  (λ* proof  a)  pr₁  (pr₁  a)
+              pr₁proofEq =
+                pr₁  (λ* proof  a)
+                  ≡⟨ cong  x  pr₁  x) λ*eq 
+                pr₁  (pair  (pr₁  (pr₁  a))  (pair  (pr₂  (pr₁  a))  (pr₂  a)))
+                  ≡⟨ pr₁pxy≡x _ _ 
+                pr₁  (pr₁  a)
+                  
 
-              pr₁proofEq : pr₁  (λ* proof  a)  pr₁  (pr₁  a)
-              pr₁proofEq =
-                pr₁  (λ* proof  a)
-                  ≡⟨ cong  x  pr₁  x) λ*eq 
-                pr₁  (pair  (pr₁  (pr₁  a))  (pair  (pr₂  (pr₁  a))  (pr₂  a)))
-                  ≡⟨ pr₁pxy≡x _ _ 
-                pr₁  (pr₁  a)
-                  
+              pr₁pr₂proofEq : pr₁  (pr₂  (λ* proof  a))  pr₂  (pr₁  a)
+              pr₁pr₂proofEq =
+                pr₁  (pr₂  (λ* proof  a))
+                  ≡⟨ cong  x  pr₁  x) pr₂proofEq 
+                pr₁  (pair  (pr₂  (pr₁  a))  (pr₂  a))
+                  ≡⟨ pr₁pxy≡x _ _ 
+                pr₂  (pr₁  a)
+                  
 
-              pr₁pr₂proofEq : pr₁  (pr₂  (λ* proof  a))  pr₂  (pr₁  a)
-              pr₁pr₂proofEq =
-                pr₁  (pr₂  (λ* proof  a))
-                  ≡⟨ cong  x  pr₁  x) pr₂proofEq 
-                pr₁  (pair  (pr₂  (pr₁  a))  (pr₂  a))
-                  ≡⟨ pr₁pxy≡x _ _ 
-                pr₂  (pr₁  a)
-                  
-
-              pr₂pr₂proofEq : pr₂  (pr₂  (λ* proof  a))  pr₂  a
-              pr₂pr₂proofEq =
-                pr₂  (pr₂  (λ* proof  a))
-                  ≡⟨ cong  x  pr₂  x) pr₂proofEq 
-                pr₂  (pair  (pr₂  (pr₁  a))  (pr₂  a))
-                  ≡⟨ pr₂pxy≡y _ _ 
-                pr₂  a
-                  
+              pr₂pr₂proofEq : pr₂  (pr₂  (λ* proof  a))  pr₂  a
+              pr₂pr₂proofEq =
+                pr₂  (pr₂  (λ* proof  a))
+                  ≡⟨ cong  x  pr₂  x) pr₂proofEq 
+                pr₂  (pair  (pr₂  (pr₁  a))  (pr₂  a))
+                  ≡⟨ pr₂pxy≡y _ _ 
+                pr₂  a
+                  
                   
-            in
-              subst  r  r   x  x') (sym pr₁proofEq) pr₁pr₁a⊩x ,
-              subst  r  r   y  x') (sym pr₁pr₂proofEq) pr₂pr₁a⊩y ,
-              subst  r  r   z  x') (sym pr₂pr₂proofEq) pr₂a⊩z })
+            in
+              subst  r  r   x  x') (sym pr₁proofEq) pr₁pr₁a⊩x ,
+              subst  r  r   y  x') (sym pr₁pr₂proofEq) pr₂pr₁a⊩y ,
+              subst  r  r   z  x') (sym pr₂pr₂proofEq) pr₂a⊩z })
 
   
 
\ No newline at end of file diff --git a/docs/Realizability.Tripos.Prealgebra.Meets.Idempotency.html b/docs/Realizability.Tripos.Prealgebra.Meets.Idempotency.html index 253f724..fc7b730 100644 --- a/docs/Realizability.Tripos.Prealgebra.Meets.Idempotency.html +++ b/docs/Realizability.Tripos.Prealgebra.Meets.Idempotency.html @@ -1,45 +1,41 @@ Realizability.Tripos.Prealgebra.Meets.Idempotency
open import Cubical.Foundations.Prelude
 open import Cubical.Data.Unit
-open import Cubical.Data.Fin
-open import Cubical.Data.Vec
-open import Cubical.Data.Sum
-open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
-open import Cubical.Data.Sigma
-open import Cubical.HITs.PropositionalTruncation
-open import Cubical.HITs.PropositionalTruncation.Monad
-open import Cubical.Relation.Binary.Order.Preorder
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
+open import Cubical.Data.FinData
+open import Cubical.Data.Vec
+open import Cubical.Data.Sum
+open import Cubical.Data.Empty renaming (rec* to ⊥*rec)
+open import Cubical.Data.Sigma
+open import Cubical.HITs.PropositionalTruncation
+open import Cubical.HITs.PropositionalTruncation.Monad
+open import Cubical.Relation.Binary.Order.Preorder
+open import Realizability.CombinatoryAlgebra
+open import Realizability.ApplicativeStructure
 
-module Realizability.Tripos.Prealgebra.Meets.Idempotency {} {A : Type } (ca : CombinatoryAlgebra A) where
-open import Realizability.Tripos.Prealgebra.Predicate ca
-open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
-open CombinatoryAlgebra ca
-open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
+module Realizability.Tripos.Prealgebra.Meets.Idempotency { ℓ' ℓ''} {A : Type } (ca : CombinatoryAlgebra A) where
+open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca
+open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca
+open CombinatoryAlgebra ca
+open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a)
 
-private λ*ComputationRule = `λ*ComputationRule as fefermanStructure
-private λ* = `λ* as fefermanStructure
+module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
+  open Predicate
+  open PredicateProperties X
+  open PreorderReasoning preorder≤
 
-module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s  k  ) where
-  private PredicateX = Predicate {ℓ'' = ℓ''} X
-  open Predicate
-  open PredicateProperties {ℓ'' = ℓ''} X
-  open PreorderReasoning preorder≤
+  x⊓x≤x :  x  x  x  x
+  x⊓x≤x x = return (pr₁ ,  x' a a⊩x⊓x  a⊩x⊓x .fst))
 
-  x⊓x≤x :  x  x  x  x
-  x⊓x≤x x = return (pr₁ ,  x' a a⊩x⊓x  a⊩x⊓x .fst))
-
-  x≤x⊓x :  x  x  x  x
-  x≤x⊓x x =
-    let
-      proof : Term as 1
-      proof = ` pair ̇ # fzero ̇ # fzero
-    in
-      return
-        ((λ* proof) ,
-           x' a a⊩x 
-            let λ*eq = λ*ComputationRule proof (a  []) in
-            (subst  r  r   x  x') (sym (cong  x  pr₁  x) λ*eq  pr₁pxy≡x _ _)) a⊩x) ,
-             subst  r  r   x  x') (sym (cong  x  pr₂  x) λ*eq  pr₂pxy≡y _ _)) a⊩x))
+  x≤x⊓x :  x  x  x  x
+  x≤x⊓x x =
+    let
+      proof : Term as 1
+      proof = ` pair ̇ # zero ̇ # zero
+    in
+      return
+        ((λ* proof) ,
+           x' a a⊩x 
+            let λ*eq = λ*ComputationRule proof a in
+            (subst  r  r   x  x') (sym (cong  x  pr₁  x) λ*eq  pr₁pxy≡x _ _)) a⊩x) ,
+             subst  r  r   x  x') (sym (cong  x  pr₂  x) λ*eq  pr₂pxy≡y _ _)) a⊩x))
 
\ No newline at end of file diff --git a/docs/index.html b/docs/index.html index b173d19..020e0aa 100644 --- a/docs/index.html +++ b/docs/index.html @@ -2,11 +2,9 @@ index
{-# OPTIONS --cubical #-}
 module index where
 
-open import Realizability.CombinatoryAlgebra
-open import Realizability.ApplicativeStructure
-open import Realizability.Topos.Everything
-open import Realizability.Assembly.Everything
-open import Realizability.PERs.Everything
-open import Realizability.Modest.Everything
-open import Realizability.Choice
+open import Realizability.Topos.Everything
+open import Realizability.Tripos.Everything
+open import Realizability.Assembly.Everything
+open import Realizability.PERs.Everything
+open import Realizability.Modest.Everything
 
\ No newline at end of file diff --git a/realizability.agda-lib b/realizability.agda-lib index 3986295..b5faaa2 100644 --- a/realizability.agda-lib +++ b/realizability.agda-lib @@ -1,4 +1,4 @@ name: realizability depend: cubical include: src -flags: --cubical -WnoUnsupportedIndexedMatch \ No newline at end of file +flags: --cubical -WnoUnsupportedIndexedMatch --inversion-max-depth=5000 \ No newline at end of file diff --git a/src/Realizability/Assembly/Everything.agda b/src/Realizability/Assembly/Everything.agda index 549c4a2..41a68e1 100644 --- a/src/Realizability/Assembly/Everything.agda +++ b/src/Realizability/Assembly/Everything.agda @@ -1,11 +1,10 @@ {-# OPTIONS --cubical #-} +-- These modules contain some elementary results about the category of assemblies +-- In particular, finite limits and exponentials module Realizability.Assembly.Everything where open import Realizability.Assembly.Base open import Realizability.Assembly.BinCoproducts open import Realizability.Assembly.BinProducts -open import Realizability.Assembly.Coequalizers open import Realizability.Assembly.Equalizers open import Realizability.Assembly.Exponentials open import Realizability.Assembly.Morphism --- TODO : Fix regular structure modules --- open import Realizability.Assembly.Regular.Everything diff --git a/src/Realizability/Modest/CanonicalPER.agda b/src/Realizability/Modest/CanonicalPER.agda index f0d0682..5b4a82e 100644 --- a/src/Realizability/Modest/CanonicalPER.agda +++ b/src/Realizability/Modest/CanonicalPER.agda @@ -1,3 +1,4 @@ +{-# OPTIONS --allow-unsolved-metas #-} open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism @@ -56,5 +57,9 @@ module _ fst (PER.isPER canonicalPER) a b (x , a⊩x , b⊩x) = x , b⊩x , a⊩x snd (PER.isPER canonicalPER) a b c (x , a⊩x , b⊩x) (x' , b⊩x' , c⊩x') = x' , subst (a ⊩[ asmX ]_) (isModestAsmX x x' b b⊩x b⊩x') a⊩x , c⊩x' - - + +CanonicalPERFunctor : Functor MOD PERCat +Functor.F-ob CanonicalPERFunctor (X , asmX , isModestAsmX) = canonicalPER asmX isModestAsmX +Functor.F-hom CanonicalPERFunctor {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = {!!} +Functor.F-id CanonicalPERFunctor = {!!} +Functor.F-seq CanonicalPERFunctor = {!!} diff --git a/src/Realizability/Modest/EquivToPERs.agda b/src/Realizability/Modest/EquivToPERs.agda new file mode 100644 index 0000000..a3bf517 --- /dev/null +++ b/src/Realizability/Modest/EquivToPERs.agda @@ -0,0 +1,73 @@ +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.HLevels +open import Cubical.Foundations.Isomorphism +open import Cubical.Foundations.Function +open import Cubical.Foundations.Equiv +open import Cubical.Foundations.Powerset +open import Cubical.Foundations.Path +open import Cubical.Foundations.Structure using (⟨_⟩; str) +open import Cubical.Data.Sigma +open import Cubical.Data.FinData +open import Cubical.Data.Unit +open import Cubical.HITs.PropositionalTruncation as PT hiding (map) +open import Cubical.HITs.PropositionalTruncation.Monad +open import Cubical.HITs.SetQuotients as SQ +open import Cubical.Reflection.RecordEquiv +open import Cubical.Categories.Category +open import Cubical.Categories.Displayed.Base +open import Cubical.Categories.Displayed.Reasoning +open import Cubical.Categories.Limits.Pullback +open import Cubical.Categories.Functor hiding (Id) +open import Cubical.Categories.Constructions.Slice +open import Cubical.Categories.Equivalence +open import Cubical.Categories.NaturalTransformation +open import Categories.CartesianMorphism +open import Categories.GenericObject +open import Realizability.CombinatoryAlgebra +open import Realizability.ApplicativeStructure +open import Realizability.PropResizing + +module Realizability.Modest.EquivToPERs {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where + +open import Realizability.Assembly.Base ca +open import Realizability.Assembly.Morphism ca +open import Realizability.Assembly.Terminal ca +open import Realizability.Assembly.SetsReflectiveSubcategory ca +open import Realizability.Modest.Base ca +open import Realizability.Modest.UniformFamily ca +open import Realizability.Modest.CanonicalPER ca +open import Realizability.Modest.SubQuotientCanonicalPERIso ca +open import Realizability.PERs.PER ca +open import Realizability.PERs.SubQuotient ca + +open Assembly +open CombinatoryAlgebra ca +open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) +open Contravariant UNIMOD +open UniformFamily +open DisplayedUFamMap + +CanonicalPERWeakInverseOfSubQuotient : WeakInverse subQuotientFunctor +WeakInverse.invFunc CanonicalPERWeakInverseOfSubQuotient = CanonicalPERFunctor +NatTrans.N-ob (NatIso.trans (WeakInverse.η CanonicalPERWeakInverseOfSubQuotient)) R = [ Id , isTrackerId ] where + isTrackerId : isTracker R (canonicalPER (subQuotientAssembly R) (isModestSubQuotientAssembly R)) Id + isTrackerId r r' r~r' = + subst2 + _~[ canonicalPER (subQuotientAssembly R) (isModestSubQuotientAssembly R) ]_ + (sym (Ida≡a r)) + (sym (Ida≡a r')) + ([ r , r~r ] , (r~r , PER.isSymmetric R r r' r~r')) where + r~r : r ~[ R ] r + r~r = PER.isTransitive R r r' r r~r' (PER.isSymmetric R r r' r~r') +NatTrans.N-hom (NatIso.trans (WeakInverse.η CanonicalPERWeakInverseOfSubQuotient)) {R} {S} f = {!!} +isIso.inv (NatIso.nIso (WeakInverse.η CanonicalPERWeakInverseOfSubQuotient) R) = {!!} +isIso.sec (NatIso.nIso (WeakInverse.η CanonicalPERWeakInverseOfSubQuotient) R) = {!!} +isIso.ret (NatIso.nIso (WeakInverse.η CanonicalPERWeakInverseOfSubQuotient) R) = {!!} +NatTrans.N-ob (NatIso.trans (WeakInverse.ε CanonicalPERWeakInverseOfSubQuotient)) (X , asmX , isModestAsmX)= {!invert asmX isModestAsmX!} +NatTrans.N-hom (NatIso.trans (WeakInverse.ε CanonicalPERWeakInverseOfSubQuotient)) {X , asmX , isModestAsmX} {Y , asmY , isModestAsmY} f = {!forward!} +isIso.inv (NatIso.nIso (WeakInverse.ε CanonicalPERWeakInverseOfSubQuotient) (X , asmX , isModestAsmX)) = {!!} +isIso.sec (NatIso.nIso (WeakInverse.ε CanonicalPERWeakInverseOfSubQuotient) (X , asmX , isModestAsmX)) = {!!} +isIso.ret (NatIso.nIso (WeakInverse.ε CanonicalPERWeakInverseOfSubQuotient) (X , asmX , isModestAsmX)) = {!!} + +isEquivalenceSubQuotient : isEquivalence subQuotientFunctor +isEquivalenceSubQuotient = ∣ CanonicalPERWeakInverseOfSubQuotient ∣₁ diff --git a/src/Realizability/Modest/PartialSurjection.agda b/src/Realizability/Modest/PartialSurjection.agda index 155b73b..01d6358 100644 --- a/src/Realizability/Modest/PartialSurjection.agda +++ b/src/Realizability/Modest/PartialSurjection.agda @@ -1,3 +1,5 @@ +-- This module defines a **partial surjection** and shows that it is equivalent to a modest set. +-- A partial surjection on a set X is a surjection from the combinatory algebra 𝔸 ↠ X that is only defined for a certain subset of 𝔸 open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism diff --git a/src/Realizability/Modest/SubQuotientCanonicalPERIso.agda b/src/Realizability/Modest/SubQuotientCanonicalPERIso.agda index 82cb630..2a5a288 100644 --- a/src/Realizability/Modest/SubQuotientCanonicalPERIso.agda +++ b/src/Realizability/Modest/SubQuotientCanonicalPERIso.agda @@ -1,3 +1,6 @@ +-- This module shows that any modest set M is isomorphic to the subquotient of the canonical PER of M. +-- Effectively, this shows that the subquotient functor is **split essentially surjective** on objects. +-- Since the subquotient functor is fully faithful, this implies that it is an equivalence of categories. open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Isomorphism @@ -19,6 +22,8 @@ open import Cubical.Categories.Displayed.Reasoning open import Cubical.Categories.Limits.Pullback open import Cubical.Categories.Functor hiding (Id) open import Cubical.Categories.Constructions.Slice +open import Cubical.Categories.Equivalence +open import Cubical.Categories.NaturalTransformation open import Categories.CartesianMorphism open import Categories.GenericObject open import Realizability.CombinatoryAlgebra diff --git a/src/Realizability/PERs/SubQuotient.agda b/src/Realizability/PERs/SubQuotient.agda index c515dd8..d412ee5 100644 --- a/src/Realizability/PERs/SubQuotient.agda +++ b/src/Realizability/PERs/SubQuotient.agda @@ -1,3 +1,6 @@ +-- The subquotient construction embeds the category PER into the category MOD of modest sets. +-- It turns out to be an equivalence of categories! In this module, we merely define it and show +-- that it is an embedding of categories --- is full and faithful. open import Realizability.ApplicativeStructure open import Realizability.CombinatoryAlgebra open import Realizability.PropResizing @@ -151,7 +154,7 @@ module _ (R S : PER) (f : AssemblyMorphism (subQuotientAssembly R) (subQuotientA PT.rec→Set squash/ mainMap mainMap2Constant (f .tracker) module InverseDefinition where isSQTracker : A → Type ℓ isSQTracker t = ∀ (q : subQuotient R) (a : A) → a ⊩[ subQuotientAssembly R ] q → ⟨ subQuotientRealizability S (t ⨾ a) (f .AssemblyMorphism.map q) ⟩ - -- 🤢🤮 + mainMap : Σ[ t ∈ A ] (isSQTracker t) → perMorphism R S mainMap (t , t⊩f) = [ t , diff --git a/src/Realizability/Topos/Everything.agda b/src/Realizability/Topos/Everything.agda index 7783f54..b760b9d 100644 --- a/src/Realizability/Topos/Everything.agda +++ b/src/Realizability/Topos/Everything.agda @@ -1,4 +1,4 @@ -{-# OPTIONS --allow-unsolved-metas #-} +-- Modules on the realizability topos over 𝔸 module Realizability.Topos.Everything where open import Realizability.Topos.Object diff --git a/src/Realizability/Tripos/Everything.agda b/src/Realizability/Tripos/Everything.agda index 987f8f7..979756a 100644 --- a/src/Realizability/Tripos/Everything.agda +++ b/src/Realizability/Tripos/Everything.agda @@ -1,7 +1,9 @@ +-- Modules on the realizability tripos. + +-- The "Prealgebra" modules establish that 𝔸-valued predicates on a set X form a Heyting prealgebra. +-- The "Syntax" and "Semantics" modules deal with the internal logic of the tripos. module Realizability.Tripos.Everything where -open import Realizability.Tripos.Predicate -open import Realizability.Tripos.Algebra -open import Realizability.Tripos.Algebra.Base + open import Realizability.Tripos.Prealgebra.Everything open import Realizability.Tripos.Logic.Syntax open import Realizability.Tripos.Logic.Semantics diff --git a/src/Realizability/Tripos/Logic/Semantics.agda b/src/Realizability/Tripos/Logic/Semantics.agda index d19a65a..916afec 100644 --- a/src/Realizability/Tripos/Logic/Semantics.agda +++ b/src/Realizability/Tripos/Logic/Semantics.agda @@ -1,6 +1,5 @@ -{-# OPTIONS --allow-unsolved-metas #-} open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm) open import Cubical.Foundations.Prelude open import Cubical.Foundations.HLevels open import Cubical.Foundations.Equiv @@ -15,7 +14,7 @@ open import Cubical.Data.Unit open import Cubical.Data.Sum open import Cubical.Data.Vec open import Cubical.Data.Nat -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Relation.Binary.Order.Preorder @@ -24,8 +23,6 @@ module Realizability.Tripos.Logic.Semantics {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where open CombinatoryAlgebra ca -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure open import Realizability.Tripos.Prealgebra.Predicate.Base {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Predicate.Properties {ℓ' = ℓ'} {ℓ'' = ℓ''} ca @@ -101,18 +98,17 @@ substitutionVarSound : ∀ {Γ Δ s} → (subs : Substitution Γ Δ) → (v : s substitutionVarSound {Γ} {.Γ} {s} id t = refl substitutionVarSound {Γ} {.(_ ′ s)} {s} (t' , subs) here = funExt λ ⟦Γ⟧ → refl substitutionVarSound {Γ} {.(_ ′ _)} {s} (t' , subs) (there t) = funExt λ ⟦Γ⟧ i → substitutionVarSound subs t i ⟦Γ⟧ -substitutionVarSound {.(_ ′ _)} {Δ} {s} r@(drop subs) t = - -- TODO : Fix unsolved constraints - funExt - λ { x@(⟦Γ⟧ , ⟦s⟧) → - ⟦ substitutionVar (drop subs) t ⟧ᵗ x - ≡[ i ]⟨ renamingTermSound (drop id) (substitutionVar subs t) i x ⟩ - ⟦ substitutionVar subs t ⟧ᵗ (⟦ drop id ⟧ᴿ x) - ≡⟨ refl ⟩ - ⟦ substitutionVar subs t ⟧ᵗ ⟦Γ⟧ - ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ⟧ ⟩ - ⟦ t ⟧ⁿ (⟦ subs ⟧ᴮ ⟦Γ⟧) - ∎} +substitutionVarSound {Γ' ′ s'} {Δ} {s} r@(drop subs) t = funExt pointwise where + pointwise : (x : ⟦ Γ' ⟧ᶜ .fst × ⟦ s' ⟧ˢ .fst) → ⟦ substitutionVar r t ⟧ᵗ x ≡ (⟦ t ⟧ⁿ ∘ ⟦ r ⟧ᴮ) x + pointwise x@(⟦Γ'⟧ , ⟦s'⟧) = + ⟦ substitutionVar r t ⟧ᵗ x + ≡[ i ]⟨ renamingTermSound (drop {s = s'} id) (substitutionVar subs t) i x ⟩ + ⟦ substitutionVar subs t ⟧ᵗ (⟦ drop {Δ = Γ'} {s = s'} id ⟧ᴿ x) + ≡⟨ refl ⟩ + ⟦ substitutionVar subs t ⟧ᵗ ⟦Γ'⟧ + ≡[ i ]⟨ substitutionVarSound subs t i ⟦Γ'⟧ ⟩ + ⟦ t ⟧ⁿ (⟦ subs ⟧ᴮ ⟦Γ'⟧) + ∎ substitutionTermSound : ∀ {Γ Δ s} → (subs : Substitution Γ Δ) → (t : Term Δ s) → ⟦ substitutionTerm subs t ⟧ᵗ ≡ ⟦ t ⟧ᵗ ∘ ⟦ subs ⟧ᴮ substitutionTermSound {Γ} {Δ} {s} subs (var x) = substitutionVarSound subs x @@ -323,13 +319,13 @@ module Soundness (b , b⊩ϕ⊨θ) ← ϕ⊨θ let prover : ApplStrTerm as 1 - prover = ` pair ̇ (` a ̇ # fzero) ̇ (` b ̇ # fzero) + prover = ` pair ̇ (` a ̇ # zero) ̇ (` b ̇ # zero) return (λ* prover , λ γ r r⊩ϕγ → let proofEq : λ* prover ⨾ r ≡ pair ⨾ (a ⨾ r) ⨾ (b ⨾ r) - proofEq = λ*ComputationRule prover (r ∷ []) + proofEq = λ*ComputationRule prover r pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ r) ≡ a ⨾ r pr₁proofEq = @@ -358,10 +354,10 @@ module Soundness (a , a⊩ϕ⊨ψ∧θ) ← ϕ⊨ψ∧θ let prover : ApplStrTerm as 1 - prover = ` pr₁ ̇ (` a ̇ # fzero) + prover = ` pr₁ ̇ (` a ̇ # zero) return (λ* prover , - λ γ b b⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) (sym (λ*ComputationRule prover (b ∷ []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst)) + λ γ b b⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ γ) (sym (λ*ComputationRule prover b)) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .fst)) `∧elim2 : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ (ψ `∧ θ) → ϕ ⊨ θ `∧elim2 {Γ} {ϕ} {ψ} {θ} ϕ⊨ψ∧θ = @@ -369,10 +365,10 @@ module Soundness (a , a⊩ϕ⊨ψ∧θ) ← ϕ⊨ψ∧θ let prover : ApplStrTerm as 1 - prover = ` pr₂ ̇ (` a ̇ # fzero) + prover = ` pr₂ ̇ (` a ̇ # zero) return (λ* prover , - λ γ b b⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym (λ*ComputationRule prover (b ∷ []))) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd)) + λ γ b b⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) (sym (λ*ComputationRule prover b)) (a⊩ϕ⊨ψ∧θ γ b b⊩ϕγ .snd)) `∃intro : ∀ {Γ} {ϕ : Formula Γ} {B} {ψ : Formula (Γ ′ B)} {t : Term Γ B} → ϕ ⊨ substitutionFormula (t , id) ψ → ϕ ⊨ `∃ ψ `∃intro {Γ} {ϕ} {B} {ψ} {t} ϕ⊨ψ[t/x] = @@ -389,13 +385,13 @@ module Soundness (b , b⊩ϕ∧ψ⊨θ) ← ϕ∧ψ⊨θ let prover : ApplStrTerm as 1 - prover = ` b ̇ (` pair ̇ # fzero ̇ (` a ̇ # fzero)) + prover = ` b ̇ (` pair ̇ # zero ̇ (` a ̇ # zero)) return (λ* prover , (λ γ c c⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) - (sym (λ*ComputationRule prover (c ∷ []))) + (sym (λ*ComputationRule prover c)) (transport (propTruncIdempotent (⟦ θ ⟧ᶠ .isPropValued γ (b ⨾ (pair ⨾ c ⨾ (a ⨾ c))))) (a⊩ϕ⊨∃ψ γ c c⊩ϕγ >>= @@ -419,13 +415,13 @@ module Soundness (a , a⊩ϕ⊨ψ) ← ϕ⊨ψ let prover : ApplStrTerm as 2 - prover = ` a ̇ # fzero + prover = ` a ̇ # one return - (λ* prover , + (λ*2 prover , (λ γ b b⊩ϕ → λ { c x@(γ' , b') γ'≡γ → subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ (γ' , b')) - (sym (λ*ComputationRule prover (b ∷ c ∷ []))) + (sym (λ*2ComputationRule prover b c)) (a⊩ϕ⊨ψ (γ' , b') b @@ -437,7 +433,7 @@ module Soundness (a , a⊩ϕ⊨∀ψ) ← ϕ⊨∀ψ let prover : ApplStrTerm as 1 - prover = ` a ̇ # fzero ̇ ` k + prover = ` a ̇ # zero ̇ ` k return (λ* prover , (λ γ b b⊩ϕγ → @@ -446,7 +442,7 @@ module Soundness (sym (substitutionFormulaSound (t , id) ψ)) (subst (λ r → r ⊩ ∣ ⟦ ψ ⟧ᶠ ∣ (γ , ⟦ t ⟧ᵗ γ)) - (sym (λ*ComputationRule prover (b ∷ []))) + (sym (λ*ComputationRule prover b)) (a⊩ϕ⊨∀ψ γ b b⊩ϕγ k (γ , ⟦ t ⟧ᵗ γ) refl)))) `→intro : ∀ {Γ} {ϕ ψ θ : Formula Γ} → (ϕ `∧ ψ) ⊨ θ → ϕ ⊨ (ψ `→ θ) @@ -459,13 +455,13 @@ module Soundness (b , b⊩ϕ⊨ψ) ← ϕ⊨ψ let prover : ApplStrTerm as 1 - prover = ` a ̇ (# fzero) ̇ (` b ̇ # fzero) + prover = ` a ̇ (# zero) ̇ (` b ̇ # zero) return (λ* prover , (λ γ c c⊩ϕγ → subst (λ r → r ⊩ ∣ ⟦ θ ⟧ᶠ ∣ γ) - (sym (λ*ComputationRule prover (c ∷ []))) + (sym (λ*ComputationRule prover c)) (a⊩ϕ⊨ψ→θ γ c c⊩ϕγ (b ⨾ c) (b⊩ϕ⊨ψ γ c c⊩ϕγ)))) `∨introR : ∀ {Γ} {ϕ ψ θ : Formula Γ} → ϕ ⊨ ψ → ϕ ⊨ (ψ `∨ θ) @@ -474,7 +470,7 @@ module Soundness (a , a⊩ϕ⊨ψ) ← ϕ⊨ψ let prover : ApplStrTerm as 1 - prover = ` pair ̇ ` true ̇ (` a ̇ # fzero) + prover = ` pair ̇ ` true ̇ (` a ̇ # zero) return ((λ* prover) , (λ γ b b⊩ϕγ → @@ -482,7 +478,7 @@ module Soundness pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ b) ≡ true pr₁proofEq = pr₁ ⨾ (λ* prover ⨾ b) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover b) ⟩ pr₁ ⨾ (pair ⨾ true ⨾ (a ⨾ b)) ≡⟨ pr₁pxy≡x _ _ ⟩ true @@ -491,7 +487,7 @@ module Soundness pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ b) ≡ a ⨾ b pr₂proofEq = pr₂ ⨾ (λ* prover ⨾ b) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover b) ⟩ pr₂ ⨾ (pair ⨾ true ⨾ (a ⨾ b)) ≡⟨ pr₂pxy≡y _ _ ⟩ a ⨾ b @@ -504,7 +500,7 @@ module Soundness (a , a⊩ϕ⊨ψ) ← ϕ⊨ψ let prover : ApplStrTerm as 1 - prover = ` pair ̇ ` false ̇ (` a ̇ # fzero) + prover = ` pair ̇ ` false ̇ (` a ̇ # zero) return ((λ* prover) , (λ γ b b⊩ϕγ → @@ -512,7 +508,7 @@ module Soundness pr₁proofEq : pr₁ ⨾ (λ* prover ⨾ b) ≡ false pr₁proofEq = pr₁ ⨾ (λ* prover ⨾ b) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover b) ⟩ pr₁ ⨾ (pair ⨾ false ⨾ (a ⨾ b)) ≡⟨ pr₁pxy≡x _ _ ⟩ false @@ -521,7 +517,7 @@ module Soundness pr₂proofEq : pr₂ ⨾ (λ* prover ⨾ b) ≡ a ⨾ b pr₂proofEq = pr₂ ⨾ (λ* prover ⨾ b) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (b ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover b) ⟩ pr₂ ⨾ (pair ⨾ false ⨾ (a ⨾ b)) ≡⟨ pr₂pxy≡y _ _ ⟩ a ⨾ b @@ -540,7 +536,7 @@ module Soundness let prover : ApplStrTerm as 1 prover = - (`if ` pr₁ ̇ (` pr₂ ̇ # fzero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero))) + (`if ` pr₁ ̇ (` pr₂ ̇ # zero) then ` a else (` b)) ̇ (` pair ̇ (` pr₁ ̇ # zero) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero))) return (λ* prover , (λ @@ -553,7 +549,7 @@ module Soundness proofEq : λ* prover ⨾ c ≡ a ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) proofEq = λ* prover ⨾ c - ≡⟨ λ*ComputationRule prover (c ∷ []) ⟩ + ≡⟨ λ*ComputationRule prover c ⟩ (if (pr₁ ⨾ (pr₂ ⨾ c)) then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) ≡⟨ cong (λ x → (if x then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c)))) pr₁⨾pr₂⨾c≡true ⟩ (if true then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) @@ -581,7 +577,7 @@ module Soundness proofEq : λ* prover ⨾ c ≡ b ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) proofEq = λ* prover ⨾ c - ≡⟨ λ*ComputationRule prover (c ∷ []) ⟩ + ≡⟨ λ*ComputationRule prover c ⟩ (if (pr₁ ⨾ (pr₂ ⨾ c)) then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) ≡⟨ cong (λ x → (if x then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c)))) pr₁pr₂⨾c≡false ⟩ (if false then a else b) ⨾ (pair ⨾ (pr₁ ⨾ c) ⨾ (pr₂ ⨾ (pr₂ ⨾ c))) diff --git a/src/Realizability/Tripos/Prealgebra/Implication.agda b/src/Realizability/Tripos/Prealgebra/Implication.agda index f3d53d7..191ad7f 100644 --- a/src/Realizability/Tripos/Prealgebra/Implication.agda +++ b/src/Realizability/Tripos/Prealgebra/Implication.agda @@ -1,10 +1,10 @@ open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure open import Cubical.Foundations.Prelude open import Cubical.Foundations.Univalence open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec module Realizability.Tripos.Prealgebra.Implication {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where @@ -14,10 +14,6 @@ open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private - λ*ComputationRule = `λ*ComputationRule as fefermanStructure - λ* = `λ* as fefermanStructure - module _ (X : Type ℓ') (isSetX' : isSet X) where PredicateX = Predicate X open Predicate @@ -27,13 +23,15 @@ module _ (X : Type ℓ') (isSetX' : isSet X) where a⊓b≤c→a≤b⇒c a b c a⊓b≤c = do (a~ , a~proves) ← a⊓b≤c - let prover = (` a~ ̇ (` pair ̇ (# fzero) ̇ (# fone))) + let + prover : Term as 2 + prover = (` a~ ̇ (` pair ̇ (# one) ̇ (# zero))) return - (λ* prover , + (λ*2 prover , λ x aₓ aₓ⊩ax bₓ bₓ⊩bx → subst (λ r → r ⊩ ∣ c ∣ x) - (sym (λ*ComputationRule prover (aₓ ∷ bₓ ∷ []))) + (sym (λ*2ComputationRule prover aₓ bₓ)) (a~proves x (pair ⨾ aₓ ⨾ bₓ) @@ -44,13 +42,13 @@ module _ (X : Type ℓ') (isSetX' : isSet X) where a≤b⇒c→a⊓b≤c a b c a≤b⇒c = do (a~ , a~proves) ← a≤b⇒c - let prover = ` a~ ̇ (` pr₁ ̇ (# fzero)) ̇ (` pr₂ ̇ (# fzero)) + let prover = ` a~ ̇ (` pr₁ ̇ (# zero)) ̇ (` pr₂ ̇ (# zero)) return (λ* prover , λ { x abₓ (pr₁abₓ⊩ax , pr₂abₓ⊩bx) → subst (λ r → r ⊩ ∣ c ∣ x) - (sym (λ*ComputationRule prover (abₓ ∷ []))) + (sym (λ*ComputationRule prover abₓ)) (a~proves x (pr₁ ⨾ abₓ) pr₁abₓ⊩ax (pr₂ ⨾ abₓ) pr₂abₓ⊩bx) }) ⇒isRightAdjointOf⊓ : ∀ a b c → (a ⊓ b ≤ c) ≡ (a ≤ b ⇒ c) @@ -63,13 +61,13 @@ module _ (X : Type ℓ') (isSetX' : isSet X) where (β , βProves) ← b≤a let prover : Term as 2 - prover = (# fzero) ̇ (` β ̇ # fone) + prover = (# one) ̇ (` β ̇ # zero) return - (λ* prover , + (λ*2 prover , (λ x r r⊩a⇒c r' r'⊩b → subst (λ witness → witness ⊩ ∣ c ∣ x) - (sym (λ*ComputationRule prover (r ∷ r' ∷ []))) + (sym (λ*2ComputationRule prover r r')) (r⊩a⇒c (β ⨾ r') (βProves x r' r'⊩b)))) antiSym→a⇒b≤a⇒c : ∀ a b c → b ≤ c → c ≤ b → (a ⇒ b) ≤ (a ⇒ c) @@ -79,11 +77,11 @@ module _ (X : Type ℓ') (isSetX' : isSet X) where (γ , γProves) ← c≤b let prover : Term as 2 - prover = ` β ̇ ((# fzero) ̇ (# fone)) + prover = ` β ̇ ((# one) ̇ (# zero)) return - (λ* prover , + (λ*2 prover , (λ x α α⊩a⇒b a' a'⊩a → subst (λ r → r ⊩ ∣ c ∣ x) - (sym (λ*ComputationRule prover (α ∷ a' ∷ []))) + (sym (λ*2ComputationRule prover α a')) (βProves x (α ⨾ a') (α⊩a⇒b a' a'⊩a)))) diff --git a/src/Realizability/Tripos/Prealgebra/Joins/#Base.agda# b/src/Realizability/Tripos/Prealgebra/Joins/#Base.agda# deleted file mode 100644 index 7892fae..0000000 --- a/src/Realizability/Tripos/Prealgebra/Joins/#Base.agda# +++ /dev/null @@ -1,27 +0,0 @@ -open import Cubical.Foundations.Prelude -open import Cubical.Data.Unit -open import Cubical.Data.Fin -open import Cubical.Data.Vec -open import Cubical.Data.Sum -open import Cubical.Data.Empty renaming (rec* to ⊥*rec) -open import Cubical.Data.Sigma -open import Cubical.HITs.PropositionalTruncation -open import Cubical.HITs.PropositionalTruncation.Monad -open import Cubical.Relation.Binary.Order.Preorder -open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) - -module Realizability.Tripos.Prealgebra.Joins.Base {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca -open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca -open CombinatoryAlgebra ca -open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) - -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure - -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - PredicateX = Predicate {ℓ'' = ℓ''} X - open Predicate - open PredicateProperties {ℓ'' = ℓ''} X - \ No newline at end of file diff --git a/src/Realizability/Tripos/Prealgebra/Joins/Associativity.agda b/src/Realizability/Tripos/Prealgebra/Joins/Associativity.agda index 3f2dbe0..a18680e 100644 --- a/src/Realizability/Tripos/Prealgebra/Joins/Associativity.agda +++ b/src/Realizability/Tripos/Prealgebra/Joins/Associativity.agda @@ -1,38 +1,34 @@ {-# OPTIONS --allow-unsolved-metas #-} open import Cubical.Foundations.Prelude open import Cubical.Data.Empty renaming (rec to ⊥rec) -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Sum open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure -module Realizability.Tripos.Prealgebra.Joins.Associativity {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +module Realizability.Tripos.Prealgebra.Joins.Associativity {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private - λ*ComputationRule = `λ*ComputationRule as fefermanStructure - λ* = `λ* as fefermanStructure - -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - private PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where + private PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X - module →⊔Assoc (x y z : Predicate {ℓ'' = ℓ''} X) where + module →⊔Assoc (x y z : PredicateX) where →proverTerm : Term as 1 →proverTerm = ` Id ̇ - (` pr₁ ̇ (# fzero)) ̇ - (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# fzero)))) ̇ + (` pr₁ ̇ (# zero)) ̇ + (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# zero)))) ̇ (` Id ̇ - (` pr₁ ̇ (` pr₂ ̇ (# fzero))) ̇ - (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))) ̇ - (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))) + (` pr₁ ̇ (` pr₂ ̇ (# zero))) ̇ + (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))) ̇ + (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))) →prover = λ* →proverTerm @@ -45,7 +41,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ proof≡pair⨾true⨾pair⨾true⨾pr₂a : proof ≡ pair ⨾ true ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ a)) proof≡pair⨾true⨾pair⨾true⨾pr₂a = proof - ≡⟨ λ*ComputationRule →proverTerm (a ∷ []) ⟩ + ≡⟨ λ*ComputationRule →proverTerm a ⟩ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ true ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ a))) ⨾ @@ -114,7 +110,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ (pair ⨾ false ⨾ (pr₂ ⨾ (pr₂ ⨾ a)))) proof≡y⊔z = proof - ≡⟨ λ*ComputationRule →proverTerm (a ∷ []) ⟩ + ≡⟨ λ*ComputationRule →proverTerm a ⟩ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ true ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ a))) ⨾ @@ -252,12 +248,12 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ prover : Term as 1 prover = ` Id ̇ - (` pr₁ ̇ (# fzero)) ̇ - (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# fzero)))) ̇ + (` pr₁ ̇ (# zero)) ̇ + (` pair ̇ ` true ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (# zero)))) ̇ (` Id ̇ - (` pr₁ ̇ (` pr₂ ̇ (# fzero))) ̇ - (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))) ̇ - (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))) + (` pr₁ ̇ (` pr₂ ̇ (# zero))) ̇ + (` pair ̇ ` true ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))) ̇ + (` pair ̇ ` false ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))) return (λ* prover , λ x' a a⊩x⊔_y⊔z → @@ -306,13 +302,13 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ let proof : Term as 1 proof = - `if ` pr₁ ̇ # fzero then - `if ` pr₁ ̇ (` pr₂ ̇ (# fzero)) then - ` pair ̇ ` true ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)) + `if ` pr₁ ̇ # zero then + `if ` pr₁ ̇ (` pr₂ ̇ (# zero)) then + ` pair ̇ ` true ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) else - (` pair ̇ ` false ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)))) + (` pair ̇ ` false ̇ (` pair ̇ ` true ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)))) else - (` pair ̇ ` false ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # fzero))) + (` pair ̇ ` false ̇ (` pair ̇ ` false ̇ (` pr₂ ̇ # zero))) return (λ* proof , (λ x' a a⊩x⊔y_⊔z → @@ -334,7 +330,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ (λ* proof ⨾ a) ≡ if (pr₁ ⨾ (pr₂ ⨾ a)) then (pair ⨾ true ⨾ (pr₂ ⨾ (pr₂ ⨾ a))) else (pair ⨾ false ⨾ (pair ⨾ true ⨾ (pr₂ ⨾ (pr₂ ⨾ a)))) proofEq = λ* proof ⨾ a - ≡⟨ λ*ComputationRule proof (a ∷ []) ⟩ + ≡⟨ λ*ComputationRule proof a ⟩ if (pr₁ ⨾ a) then if (pr₁ ⨾ (pr₂ ⨾ a)) then (pair ⨾ true ⨾ (pr₂ ⨾ (pr₂ ⨾ a))) @@ -434,7 +430,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ proof≡pair : λ* proof ⨾ a ≡ pair ⨾ false ⨾ (pair ⨾ false ⨾ (pr₂ ⨾ a)) proof≡pair = λ* proof ⨾ a - ≡⟨ λ*ComputationRule proof (a ∷ []) ⟩ + ≡⟨ λ*ComputationRule proof a ⟩ if (pr₁ ⨾ a) then if (pr₁ ⨾ (pr₂ ⨾ a)) then (pair ⨾ true ⨾ (pr₂ ⨾ (pr₂ ⨾ a))) diff --git a/src/Realizability/Tripos/Prealgebra/Joins/Base.agda b/src/Realizability/Tripos/Prealgebra/Joins/Base.agda index e3e6c78..007e663 100644 --- a/src/Realizability/Tripos/Prealgebra/Joins/Base.agda +++ b/src/Realizability/Tripos/Prealgebra/Joins/Base.agda @@ -1,7 +1,7 @@ open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv open import Cubical.Data.Unit -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Sum renaming (rec to sumRec) open import Cubical.Data.Empty renaming (rec* to ⊥*rec) @@ -10,26 +10,23 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Relation.Binary.Order.Preorder open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure -module Realizability.Tripos.Prealgebra.Joins.Base {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +module Realizability.Tripos.Prealgebra.Joins.Base {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure - -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where + PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X antiSym→x⊔z≤y⊔z' : ∀ x y z → x ≤ y → y ≤ x → (x ⊔ z) ≤ (y ⊔ z) antiSym→x⊔z≤y⊔z' x y z x≤y y≤x = do (x⊨y , x⊨yProves) ← x≤y - let prover = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# fzero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# fzero))) + let prover = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# zero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# zero))) return (λ* prover , λ x' a a⊩x⊔zx' → @@ -43,7 +40,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ (λ { (pr₁⨾a≡k , pr₂⨾a⊩xx') → inl (((pr₁ ⨾ (λ* prover ⨾ a)) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover a) ⟩ (pr₁ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) ≡⟨ cong (λ x → (pr₁ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k ⟩ (pr₁ ⨾ (Id ⨾ k ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) @@ -56,7 +53,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ (λ r → r ⊩ ∣ y ∣ x') (sym (pr₂ ⨾ (λ* prover ⨾ a) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover a) ⟩ pr₂ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) ≡⟨ cong (λ x → (pr₂ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k ⟩ pr₂ ⨾ (Id ⨾ k ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) @@ -69,7 +66,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ (λ { (pr₁⨾a≡k' , pr₂a⊩zx') → inr ((((pr₁ ⨾ (λ* prover ⨾ a)) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover a) ⟩ (pr₁ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) ≡⟨ cong (λ x → (pr₁ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k' ⟩ (pr₁ ⨾ (Id ⨾ k' ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) @@ -82,7 +79,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ (λ r → r ⊩ ∣ z ∣ x') (sym ((pr₂ ⨾ (λ* prover ⨾ a) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover a) ⟩ pr₂ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) ≡⟨ cong (λ x → (pr₂ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k' ⟩ pr₂ ⨾ (Id ⨾ k' ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) diff --git a/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda b/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda index 5affaf6..1bf9605 100644 --- a/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda +++ b/src/Realizability/Tripos/Prealgebra/Joins/Commutativity.agda @@ -1,8 +1,8 @@ open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure open import Cubical.Foundations.Prelude open import Cubical.Foundations.Equiv -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Sum renaming (rec to sumRec) open import Cubical.Relation.Binary.Order.Preorder @@ -14,9 +14,6 @@ module Realizability.Tripos.Prealgebra.Joins.Commutativity {ℓ} {A : Type ℓ} open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure - module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca private PredicateX = Predicate X @@ -28,8 +25,8 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where a⊔b→b⊔a : ∀ a b → a ⊔ b ≤ b ⊔ a a⊔b→b⊔a a b = do - let prover = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# fzero))) ̇ (` pair ̇ ` k ̇ (` pr₂ ̇ (# fzero))) - let λ*eq = λ (r : A) → λ*ComputationRule prover (r ∷ []) + let prover = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# zero))) ̇ (` pair ̇ ` k ̇ (` pr₂ ̇ (# zero))) + let λ*eq = λ (r : A) → λ*ComputationRule prover r return (λ* prover , λ x r r⊩a⊔b → @@ -89,7 +86,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where antiSym→x⊔z≤y⊔z x y z x≤y y≤x = do (x⊨y , x⊨yProves) ← x≤y - let prover = ` Id ̇ (` pr₁ ̇ (# fzero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# fzero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# fzero))) + let prover = ` Id ̇ (` pr₁ ̇ (# zero)) ̇ (` pair ̇ ` k ̇ (` x⊨y ̇ (` pr₂ ̇ (# zero)))) ̇ (` pair ̇ ` k' ̇ (` pr₂ ̇ (# zero))) return (λ* prover , λ x' a a⊩x⊔zx' → @@ -103,7 +100,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where (λ { (pr₁⨾a≡k , pr₂⨾a⊩xx') → inl (((pr₁ ⨾ (λ* prover ⨾ a)) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover a) ⟩ (pr₁ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) ≡⟨ cong (λ x → (pr₁ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k ⟩ (pr₁ ⨾ (Id ⨾ k ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) @@ -116,7 +113,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where (λ r → r ⊩ ∣ y ∣ x') (sym (pr₂ ⨾ (λ* prover ⨾ a) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover a) ⟩ pr₂ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) ≡⟨ cong (λ x → (pr₂ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k ⟩ pr₂ ⨾ (Id ⨾ k ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) @@ -129,7 +126,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where (λ { (pr₁⨾a≡k' , pr₂a⊩zx') → inr ((((pr₁ ⨾ (λ* prover ⨾ a)) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule prover a) ⟩ (pr₁ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) ≡⟨ cong (λ x → (pr₁ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k' ⟩ (pr₁ ⨾ (Id ⨾ k' ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a)))) @@ -142,7 +139,7 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) where (λ r → r ⊩ ∣ z ∣ x') (sym ((pr₂ ⨾ (λ* prover ⨾ a) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule prover a) ⟩ pr₂ ⨾ (Id ⨾ (pr₁ ⨾ a) ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) ≡⟨ cong (λ x → (pr₂ ⨾ (Id ⨾ x ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))))) pr₁⨾a≡k' ⟩ pr₂ ⨾ (Id ⨾ k' ⨾ (pair ⨾ k ⨾ (x⊨y ⨾ (pr₂ ⨾ a))) ⨾ (pair ⨾ k' ⨾ (pr₂ ⨾ a))) diff --git a/src/Realizability/Tripos/Prealgebra/Joins/Idempotency.agda b/src/Realizability/Tripos/Prealgebra/Joins/Idempotency.agda index 32f42ba..c4a1c95 100644 --- a/src/Realizability/Tripos/Prealgebra/Joins/Idempotency.agda +++ b/src/Realizability/Tripos/Prealgebra/Joins/Idempotency.agda @@ -1,6 +1,6 @@ open import Cubical.Foundations.Prelude open import Cubical.Data.Unit -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Sum open import Cubical.Data.Empty renaming (rec* to ⊥*rec) @@ -9,21 +9,18 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Relation.Binary.Order.Preorder open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure -module Realizability.Tripos.Prealgebra.Joins.Idempotency {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +module Realizability.Tripos.Prealgebra.Joins.Idempotency {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure - -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - private PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where + private PredicateX = Predicate X open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X x⊔x≤x : ∀ x → x ⊔ x ≤ x x⊔x≤x x = @@ -41,10 +38,10 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ x≤x⊔x : ∀ x → x ≤ x ⊔ x x≤x⊔x x = - let prover = ` pair ̇ ` true ̇ # fzero in + let prover = ` pair ̇ ` true ̇ # zero in ∣ λ* prover , (λ x' a a⊩x → subst (λ r → r ⊩ ∣ x ⊔ x ∣ x') - (sym (λ*ComputationRule prover (a ∷ []))) + (sym (λ*ComputationRule prover a)) ∣ inl (pr₁pxy≡x _ _ , subst (λ r → r ⊩ ∣ x ∣ x') (sym (pr₂pxy≡y _ _)) a⊩x) ∣₁) ∣₁ diff --git a/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda b/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda index f8cb0f1..1554c3d 100644 --- a/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda +++ b/src/Realizability/Tripos/Prealgebra/Joins/Identity.agda @@ -1,7 +1,6 @@ -{-# OPTIONS --allow-unsolved-metas #-} open import Cubical.Foundations.Prelude open import Cubical.Data.Unit -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Sum open import Cubical.Data.Empty renaming (rec* to ⊥*rec) @@ -10,7 +9,7 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Relation.Binary.Order.Preorder open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure module Realizability.Tripos.Prealgebra.Joins.Identity {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca @@ -18,16 +17,12 @@ open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure - module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where private PredicateX = Predicate X open Predicate open PredicateProperties X open PreorderReasoning preorder≤ - pre0 : PredicateX Predicate.isSetX pre0 = isSetX' ∣ pre0 ∣ = λ x a → ⊥* @@ -52,7 +47,7 @@ module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) wh x≤0⊔x x = let proof : Term as 1 - proof = ` pair ̇ ` false ̇ # fzero + proof = ` pair ̇ ` false ̇ # zero in return ((λ* proof) , @@ -61,7 +56,7 @@ module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) wh pr₁proofEq : pr₁ ⨾ (λ* proof ⨾ a) ≡ false pr₁proofEq = pr₁ ⨾ (λ* proof ⨾ a) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule proof (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule proof a) ⟩ pr₁ ⨾ (pair ⨾ false ⨾ a) ≡⟨ pr₁pxy≡x _ _ ⟩ false @@ -70,7 +65,7 @@ module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) wh pr₂proofEq : pr₂ ⨾ (λ* proof ⨾ a) ≡ a pr₂proofEq = pr₂ ⨾ (λ* proof ⨾ a) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule proof (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule proof a) ⟩ pr₂ ⨾ (pair ⨾ false ⨾ a) ≡⟨ pr₂pxy≡y _ _ ⟩ a @@ -82,7 +77,7 @@ module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) wh x≤x⊔0 x = let proof : Term as 1 - proof = ` pair ̇ ` true ̇ # fzero + proof = ` pair ̇ ` true ̇ # zero in return ((λ* proof) , (λ x' a a⊩x → @@ -90,7 +85,7 @@ module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) wh pr₁proofEq : pr₁ ⨾ (λ* proof ⨾ a) ≡ true pr₁proofEq = pr₁ ⨾ (λ* proof ⨾ a) - ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule proof (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₁ ⨾ x) (λ*ComputationRule proof a) ⟩ pr₁ ⨾ (pair ⨾ true ⨾ a) ≡⟨ pr₁pxy≡x _ _ ⟩ true @@ -99,21 +94,13 @@ module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) wh pr₂proofEq : pr₂ ⨾ (λ* proof ⨾ a) ≡ a pr₂proofEq = pr₂ ⨾ (λ* proof ⨾ a) - ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule proof (a ∷ [])) ⟩ + ≡⟨ cong (λ x → pr₂ ⨾ x) (λ*ComputationRule proof a) ⟩ pr₂ ⨾ (pair ⨾ true ⨾ a) ≡⟨ pr₂pxy≡y _ _ ⟩ a ∎ in ∣ inl (pr₁proofEq , subst (λ r → r ⊩ ∣ x ∣ x') (sym pr₂proofEq) a⊩x) ∣₁)) - 0⊔x≤x : ∀ x → pre0 ⊔ x ≤ x - 0⊔x≤x x = - pre0 ⊔ x - ≲⟨ a⊔b→b⊔a X isSetX' pre0 x ⟩ - x ⊔ pre0 - ≲⟨ x⊔0≤x x ⟩ - x - ◾ diff --git a/src/Realizability/Tripos/Prealgebra/Meets/Associativity.agda b/src/Realizability/Tripos/Prealgebra/Meets/Associativity.agda index 150644c..5fa6569 100644 --- a/src/Realizability/Tripos/Prealgebra/Meets/Associativity.agda +++ b/src/Realizability/Tripos/Prealgebra/Meets/Associativity.agda @@ -1,6 +1,6 @@ open import Cubical.Foundations.Prelude open import Cubical.Data.Unit -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Sum open import Cubical.Data.Empty renaming (rec* to ⊥*rec) @@ -9,35 +9,31 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Relation.Binary.Order.Preorder open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure -module Realizability.Tripos.Prealgebra.Meets.Associativity {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +module Realizability.Tripos.Prealgebra.Meets.Associativity {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure - -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - private PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X open PreorderReasoning preorder≤ x⊓_y⊓z≤x⊓y_⊓z : ∀ x y z → x ⊓ (y ⊓ z) ≤ (x ⊓ y) ⊓ z x⊓_y⊓z≤x⊓y_⊓z x y z = let proof : Term as 1 - proof = ` pair ̇ (` pair ̇ (` pr₁ ̇ # fzero) ̇ (` pr₁ ̇ (` pr₂ ̇ # fzero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # fzero)) + proof = ` pair ̇ (` pair ̇ (` pr₁ ̇ # zero) ̇ (` pr₁ ̇ (` pr₂ ̇ # zero))) ̇ (` pr₂ ̇ (` pr₂ ̇ # zero)) in return (λ* proof , λ x' a a⊩x⊓_y⊓z → let - λ*eq = λ*ComputationRule proof (a ∷ []) - -- Unfortunately, proof assistants + λ*eq = λ*ComputationRule proof a + pr₁proofEq : pr₁ ⨾ (λ* proof ⨾ a) ≡ pair ⨾ (pr₁ ⨾ a) ⨾ (pr₁ ⨾ (pr₂ ⨾ a)) pr₁proofEq = pr₁ ⨾ (λ* proof ⨾ a) @@ -82,13 +78,13 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ x⊓y_⊓z≤x⊓_y⊓z x y z = let proof : Term as 1 - proof = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # fzero)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₂ ̇ # fzero)) + proof = ` pair ̇ (` pr₁ ̇ (` pr₁ ̇ # zero)) ̇ (` pair ̇ (` pr₂ ̇ (` pr₁ ̇ # zero)) ̇ (` pr₂ ̇ # zero)) in return (λ* proof , λ { x' a ((pr₁pr₁a⊩x , pr₂pr₁a⊩y) , pr₂a⊩z) → let - λ*eq = λ*ComputationRule proof (a ∷ []) + λ*eq = λ*ComputationRule proof a pr₂proofEq : pr₂ ⨾ (λ* proof ⨾ a) ≡ pair ⨾ (pr₂ ⨾ (pr₁ ⨾ a)) ⨾ (pr₂ ⨾ a) pr₂proofEq = diff --git a/src/Realizability/Tripos/Prealgebra/Meets/Idempotency.agda b/src/Realizability/Tripos/Prealgebra/Meets/Idempotency.agda index a1c0bad..74844db 100644 --- a/src/Realizability/Tripos/Prealgebra/Meets/Idempotency.agda +++ b/src/Realizability/Tripos/Prealgebra/Meets/Idempotency.agda @@ -1,6 +1,6 @@ open import Cubical.Foundations.Prelude open import Cubical.Data.Unit -open import Cubical.Data.Fin +open import Cubical.Data.FinData open import Cubical.Data.Vec open import Cubical.Data.Sum open import Cubical.Data.Empty renaming (rec* to ⊥*rec) @@ -9,21 +9,17 @@ open import Cubical.HITs.PropositionalTruncation open import Cubical.HITs.PropositionalTruncation.Monad open import Cubical.Relation.Binary.Order.Preorder open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure renaming (λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*) +open import Realizability.ApplicativeStructure -module Realizability.Tripos.Prealgebra.Meets.Idempotency {ℓ} {A : Type ℓ} (ca : CombinatoryAlgebra A) where -open import Realizability.Tripos.Prealgebra.Predicate ca +module Realizability.Tripos.Prealgebra.Meets.Idempotency {ℓ ℓ' ℓ''} {A : Type ℓ} (ca : CombinatoryAlgebra A) where +open import Realizability.Tripos.Prealgebra.Predicate {ℓ' = ℓ'} {ℓ'' = ℓ''} ca open import Realizability.Tripos.Prealgebra.Joins.Commutativity ca open CombinatoryAlgebra ca open Realizability.CombinatoryAlgebra.Combinators ca renaming (i to Id; ia≡a to Ida≡a) -private λ*ComputationRule = `λ*ComputationRule as fefermanStructure -private λ* = `λ* as fefermanStructure - -module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where - private PredicateX = Predicate {ℓ'' = ℓ''} X +module _ (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ k → ⊥) where open Predicate - open PredicateProperties {ℓ'' = ℓ''} X + open PredicateProperties X open PreorderReasoning preorder≤ x⊓x≤x : ∀ x → x ⊓ x ≤ x @@ -33,11 +29,11 @@ module _ {ℓ' ℓ''} (X : Type ℓ') (isSetX' : isSet X) (isNonTrivial : s ≡ x≤x⊓x x = let proof : Term as 1 - proof = ` pair ̇ # fzero ̇ # fzero + proof = ` pair ̇ # zero ̇ # zero in return ((λ* proof) , (λ x' a a⊩x → - let λ*eq = λ*ComputationRule proof (a ∷ []) in + let λ*eq = λ*ComputationRule proof a in (subst (λ r → r ⊩ ∣ x ∣ x') (sym (cong (λ x → pr₁ ⨾ x) λ*eq ∙ pr₁pxy≡x _ _)) a⊩x) , subst (λ r → r ⊩ ∣ x ∣ x') (sym (cong (λ x → pr₂ ⨾ x) λ*eq ∙ pr₂pxy≡y _ _)) a⊩x)) diff --git a/src/index.agda b/src/index.agda index 0ad3713..cab8cd7 100644 --- a/src/index.agda +++ b/src/index.agda @@ -1,10 +1,8 @@ {-# OPTIONS --cubical #-} module index where -open import Realizability.CombinatoryAlgebra -open import Realizability.ApplicativeStructure open import Realizability.Topos.Everything +open import Realizability.Tripos.Everything open import Realizability.Assembly.Everything open import Realizability.PERs.Everything open import Realizability.Modest.Everything -open import Realizability.Choice From 29db2fb9865cba18c1b0efffe30edea8737cca24 Mon Sep 17 00:00:00 2001 From: Rahul Chhabra Date: Sat, 12 Oct 2024 20:08:22 +0530 Subject: [PATCH 61/61] Update README --- README.md | 77 ++++++++----------------------------------------------- 1 file changed, 10 insertions(+), 67 deletions(-) diff --git a/README.md b/README.md index 8f2c425..c63b5f3 100644 --- a/README.md +++ b/README.md @@ -1,73 +1,16 @@ # Experiments with Realizability in Univalent Type Theory -This repository contains accompanying code for my (upcoming) Bachelor's thesis. I am studying categorical realisability through the lens of univalent type theory. +This project formalises categorical realizability from the ground up in Cubical Agda. -This formalisation loosely follows my study (see the timeline below). Much of the development has been done inside Agda first, so it is convenient -for me to think using a cubical viewpoint, but most of the results should directly translate over to their Book HoTT analogues. +Notable results include : -This project is very much a work-in-progress, so the code is not very pretty - it is littered with `{!!}` and `--allow-unsolved-metas`. -Most of the holes have comments alongside them that explain how to fill that hole. +* The category of assemblies having all finite limits and being cartesian closed +* Modest sets are equivalent to partial surjections from the combinatory algebra +* The construction of the subquotient embedding from PERs to modest sets and that it is split essentially surjective on objects +* The internal logic of the realizability tripos +* The construction of the realizability topos -I periodically clean and reorganise the code. +Current formalisation targets include : - -# Code Organisation - -- At the highest level, we have combinatory algebra machinery -- Most modules are parameterised by a combinatory algebra `ca : CombinatoryAlgebra A` -- There are modules on the category of assemblies in `Realizability.Assembly` -- Lemmas and stuff relating to the regularity of $\mathsf{Asm}$ is found in `Realizability.Assembly.Regular` -- Tripos modules are to be found in `Realizability.Tripos` - -# Writing - -There are some notes relating to the project on my [abstract non-sense](https://github.com/rahulc29/abstract-nonsense) repository. - -# Build Instructions - -You will need Agda >= 2.6.4 and a [custom fork](https://github.com/rahulc29/cubical/tree/rahulc29/realizability) of the Cubical library to build the code. - -The custom fork has a few additional definitions in the category theory modules. I will hopefully integrate them into the Cubical library. - -# Timeline - -Weak timeline : - -Upto November 2023: - -- [x] Combinatory Algebras - - [x] Applicative Structures - - [x] Feferman structure on an AS - - [x] Combinatorial completeness - - [ ] Computation rule for $\lambda*$ - - Combinators - - [x] Identity, booleans, if-then-else, pairs, projections, B combinator, some Curry numerals - - [x] Computation rule for pairs - - [ ] Fixpoint combinators and primitive recursion combinator -- [x] Category of Assemblies - - [x] Define assemblies - - [x] Define the category $\mathsf{Asm}$ - - [x] Cartesian closure and similar structure - - [x] Binary products - - [x] Binary coproducts - - [x] Universal property - - [x] Equalisers - - [x] Exponentials - - [x] Initial and terminal objects - - [x] Coequalisers (December 2023) - -December 2023: - -- [ ] $\mathsf{Asm}$ is regular - - [x] Kernel pairs of morphisms exist - - [x] Kernel pairs have coequalisers - - [ ] Regular epics stable under pullback -- [ ] Exact completion - - [x] Internal equivalence relations of a category - - [ ] Functional relations -- [ ] Tripos - - [x] Heyting-valued Predicates - - [x] $\forall$ and $\exists$ are adjoints - - [x] Beck-Chevalley condition - - [ ] Heyting prealgebra structure - - [ ] Interpret intuitionistic logic +* Uniform families of PERs indexed by assemblies form a small and complete fibration +* The category of assemblies is exactly the category of double negation stable objects of the realizability topos