-
Notifications
You must be signed in to change notification settings - Fork 157
New notion of equivalence: Bijective relations #1252
New issue
Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.
By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.
Already on GitHub? Sign in to your account
base: master
Are you sure you want to change the base?
Changes from 8 commits
bf1e4fc
648993b
643c0a3
5600413
896d912
3f16baa
7fcb3e0
bf41a31
5479bcf
861451e
File filter
Filter by extension
Conversations
Jump to
Diff view
Diff view
There are no files selected for viewing
| Original file line number | Diff line number | Diff line change |
|---|---|---|
| @@ -0,0 +1,150 @@ | ||
| {- | ||
| Bijective Relations ([BijectiveRel]) | ||
| - Path to BijectiveRel ([pathToBijectiveRel]) | ||
| - BijectiveRel is equivalent to Equiv ([BijectiveRel≃Equiv]) | ||
| -} | ||
| module Cubical.Foundations.Equiv.BijectiveRel where | ||
|
|
||
| open import Cubical.Foundations.Prelude | ||
| open import Cubical.Foundations.Isomorphism | ||
| open import Cubical.Foundations.Equiv | ||
| open import Cubical.Foundations.Univalence | ||
| open import Cubical.Foundations.Univalence.Dependent | ||
| open import Cubical.Foundations.GroupoidLaws | ||
| open import Cubical.Foundations.Function | ||
| open import Cubical.Foundations.HLevels | ||
| open import Cubical.Functions.FunExtEquiv | ||
| open import Cubical.Relation.Binary | ||
| open import Cubical.Reflection.RecordEquiv | ||
| open import Cubical.Reflection.StrictEquiv | ||
| open import Cubical.Data.Sigma | ||
|
|
||
| private variable | ||
| ℓ ℓ' ℓ'' : Level | ||
| A B C : Type ℓ | ||
| R S : Rel A B ℓ | ||
|
|
||
| open HeterogenousRelation | ||
|
|
||
| record isBijectiveRel {A : Type ℓ} {B : Type ℓ'} (R : Rel A B ℓ'') : Type (ℓ-max ℓ (ℓ-max ℓ' ℓ'')) where | ||
| field | ||
| rContr : isFunctionalRel R | ||
| lContr : isFunctionalRel (flip R) | ||
|
|
||
| trr : A → B | ||
| trr a = rContr a .fst .fst | ||
|
|
||
| trl : B → A | ||
| trl b = lContr b .fst .fst | ||
|
|
||
| liftr : ∀ a → R a (trr a) | ||
| liftr a = rContr a .fst .snd | ||
|
|
||
| liftl : ∀ b → R (trl b) b | ||
| liftl b = lContr b .fst .snd | ||
|
|
||
| rightIsId : ∀ a → isIdentitySystem (trr a) (R a) (liftr a) | ||
| rightIsId a = isContrTotal→isIdentitySystem (rContr a) | ||
|
|
||
| module _ (a : A) where | ||
| open isIdentitySystem (rightIsId a) using () | ||
| renaming (isoPath to rightIsoPath; equivPath to rightEquivPath) public | ||
|
|
||
| leftIsId : ∀ b → isIdentitySystem (trl b) (flip R b) (liftl b) | ||
| leftIsId b = isContrTotal→isIdentitySystem (lContr b) | ||
|
|
||
| module _ (b : B) where | ||
| open isIdentitySystem (leftIsId b) using () | ||
| renaming (isoPath to leftIsoPath; equivPath to leftEquivPath) public | ||
|
|
||
| isEquivTrr : isEquiv trr | ||
| isEquivTrr .equiv-proof b = isOfHLevelRetractFromIso 0 (Σ-cong-iso-snd (λ a → rightIsoPath a b)) (lContr b) | ||
|
|
||
| isEquivTrl : isEquiv trl | ||
| isEquivTrl .equiv-proof a = isOfHLevelRetractFromIso 0 (Σ-cong-iso-snd (λ b → leftIsoPath b a)) (rContr a) | ||
|
|
||
| open isBijectiveRel | ||
|
|
||
| unquoteDecl isBijectiveRelIsoΣ = declareRecordIsoΣ isBijectiveRelIsoΣ (quote isBijectiveRel) | ||
|
|
||
| isPropIsBijectiveRel : {R : Rel A B ℓ''} → isProp (isBijectiveRel R) | ||
| isPropIsBijectiveRel x y i .rContr a = isPropIsContr (x .rContr a) (y .rContr a) i | ||
| isPropIsBijectiveRel x y i .lContr a = isPropIsContr (x .lContr a) (y .lContr a) i | ||
|
|
||
| BijectiveRel : ∀ (A : Type ℓ) (B : Type ℓ') ℓ'' → Type (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓ'')) | ||
| BijectiveRel A B ℓ'' = Σ[ R ∈ Rel A B ℓ'' ] isBijectiveRel R | ||
|
|
||
| BijectiveRelIsoΣ : Iso (BijectiveRel A B ℓ'') (Σ[ R ∈ Rel A B ℓ'' ] isFunctionalRel R × isFunctionalRel (flip R)) | ||
| BijectiveRelIsoΣ = Σ-cong-iso-snd λ _ → isBijectiveRelIsoΣ | ||
|
|
||
| BijectiveRelPathP : {A : I → Type ℓ} {B : I → Type ℓ'} {R₀ : BijectiveRel (A i0) (B i0) ℓ''} {R₁ : BijectiveRel (A i1) (B i1) ℓ''} | ||
| → PathP (λ i → Rel (A i) (B i) ℓ'') (R₀ .fst) (R₁ .fst) | ||
| → PathP (λ i → BijectiveRel (A i) (B i) ℓ'') R₀ R₁ | ||
| BijectiveRelPathP = ΣPathPProp λ _ → isPropIsBijectiveRel | ||
|
|
||
| BijectiveRelEq : {R₀ R₁ : BijectiveRel A B ℓ''} → (∀ a b → R₀ .fst a b ≃ R₁ .fst a b) → R₀ ≡ R₁ | ||
| BijectiveRelEq h = BijectiveRelPathP (funExt₂ λ a b → ua (h a b)) | ||
|
|
||
| BijectiveRel→Equiv : BijectiveRel A B ℓ → A ≃ B | ||
| BijectiveRel→Equiv (R , Rbij) .fst = trr Rbij | ||
| BijectiveRel→Equiv (R , Rbij) .snd = isEquivTrr Rbij | ||
|
|
||
| Equiv→BijectiveRel : A ≃ B → BijectiveRel A B _ | ||
| Equiv→BijectiveRel e .fst = graphRel (e .fst) | ||
| Equiv→BijectiveRel e .snd .rContr a = isContrSingl (e .fst a) | ||
| Equiv→BijectiveRel e .snd .lContr = e .snd .equiv-proof | ||
|
|
||
| EquivIsoBijectiveRel : (A B : Type ℓ) → Iso (A ≃ B) (BijectiveRel A B ℓ) | ||
| EquivIsoBijectiveRel A B .Iso.fun = Equiv→BijectiveRel | ||
| EquivIsoBijectiveRel A B .Iso.inv = BijectiveRel→Equiv | ||
| EquivIsoBijectiveRel A B .Iso.rightInv (R , Rbij) = BijectiveRelEq $ rightEquivPath Rbij | ||
| EquivIsoBijectiveRel A B .Iso.leftInv e = equivEq refl | ||
|
|
||
| Equiv≃BijectiveRel : (A B : Type ℓ) → (A ≃ B) ≃ (BijectiveRel A B ℓ) | ||
| Equiv≃BijectiveRel A B = isoToEquiv (EquivIsoBijectiveRel A B) | ||
|
|
||
| isBijectiveIdRel : isBijectiveRel (idRel A) | ||
| isBijectiveIdRel .rContr = isContrSingl | ||
| isBijectiveIdRel .lContr = isContrSingl' | ||
|
|
||
| idBijectiveRel : BijectiveRel A A _ | ||
| idBijectiveRel = _ , isBijectiveIdRel | ||
|
|
||
| isBijectiveInvRel : isBijectiveRel R → isBijectiveRel (invRel R) | ||
| isBijectiveInvRel Rbij .rContr = Rbij .lContr | ||
| isBijectiveInvRel Rbij .lContr = Rbij .rContr | ||
|
|
||
| invBijectiveRel : BijectiveRel A B ℓ' → BijectiveRel B A ℓ' | ||
| invBijectiveRel (_ , Rbij) = _ , isBijectiveInvRel Rbij | ||
|
|
||
| isBijectiveCompRel : isBijectiveRel R → isBijectiveRel S → isBijectiveRel (compRel R S) | ||
| isBijectiveCompRel Rbij Sbij .rContr = isFunctionalCompRel (Rbij .rContr) (Sbij .rContr) | ||
| isBijectiveCompRel Rbij Sbij .lContr = isCofunctionalCompRel (Rbij .lContr) (Sbij .lContr) | ||
|
|
||
| compBijectiveRel : BijectiveRel A B ℓ → BijectiveRel B C ℓ' → BijectiveRel A C _ | ||
| compBijectiveRel (_ , Rbij) (_ , Sbij) = _ , isBijectiveCompRel Rbij Sbij | ||
|
|
||
| isBijectivePathP : (A : I → Type ℓ) → isBijectiveRel (PathP A) | ||
| isBijectivePathP A .rContr = isContrSinglP A | ||
| isBijectivePathP A .lContr = isContrSinglP' A | ||
|
|
||
| pathToBijectiveRel : A ≡ B → BijectiveRel A B _ | ||
| pathToBijectiveRel P = _ , isBijectivePathP λ i → P i | ||
|
|
||
| BijectiveRelToPath : BijectiveRel A B ℓ → A ≡ B | ||
| BijectiveRelToPath R = ua (BijectiveRel→Equiv R) | ||
|
|
||
| path→BijectiveRel→Equiv : (P : A ≡ B) → BijectiveRel→Equiv (pathToBijectiveRel P) ≡ pathToEquiv P | ||
| path→BijectiveRel→Equiv P = equivEq refl | ||
|
|
||
| pathIsoBijectiveRel : Iso (A ≡ B) (BijectiveRel A B _) | ||
| pathIsoBijectiveRel .Iso.fun = pathToBijectiveRel | ||
| pathIsoBijectiveRel .Iso.inv = BijectiveRelToPath | ||
| pathIsoBijectiveRel .Iso.rightInv (R , Rbij) = BijectiveRelEq λ a b → ua-ungluePath-Equiv _ ∙ₑ rightEquivPath Rbij a b | ||
| pathIsoBijectiveRel .Iso.leftInv P = cong ua (path→BijectiveRel→Equiv P) ∙ ua-pathToEquiv P | ||
|
|
||
| path≡BijectiveRel : (A ≡ B) ≡ BijectiveRel A B _ | ||
| path≡BijectiveRel = isoToPath pathIsoBijectiveRel | ||
| Original file line number | Diff line number | Diff line change |
|---|---|---|
|
|
@@ -15,8 +15,6 @@ This file proves a variety of basic results about paths: | |
|
|
||
| - Direct definitions of lower h-levels | ||
|
|
||
| - Export natural numbers | ||
|
|
||
| - Export universe lifting | ||
|
|
||
| -} | ||
|
|
@@ -485,18 +483,36 @@ is2Groupoid A = ∀ a b → isGroupoid (Path A a b) | |
| singlP : (A : I → Type ℓ) (a : A i0) → Type _ | ||
| singlP A a = Σ[ x ∈ A i1 ] PathP A a x | ||
|
|
||
| singlP' : (A : I → Type ℓ) (a : A i1) → Type _ | ||
| singlP' A a = Σ[ x ∈ A i0 ] PathP A x a | ||
|
Collaborator
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. Do we really need all of this duplication of singlP?
Contributor
Author
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. I needed
Collaborator
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. It can't be defined as
Contributor
Author
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more.
Collaborator
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. Why would it need reflection? In
Contributor
Author
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. oh yeah I didn't think of that! for some reason I had assumed that we would need the 2LTT strict equality, but I forgot that the inductive equality works just as well.
Collaborator
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. The reflection there is still helpful, it's just to allow for an abbreviation of writing an equivalence that is strict without manually invoking |
||
|
|
||
| singl : (a : A) → Type _ | ||
| singl {A = A} a = singlP (λ _ → A) a | ||
|
|
||
| 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) | ||
| isContrSingl a .fst = _ , refl | ||
| isContrSingl a .snd (x , p) i = _ , λ j → p (i ∧ j) | ||
|
|
||
| isContrSingl' : (a : A) → isContr (singl' a) | ||
| isContrSingl' a .fst = _ , refl | ||
| isContrSingl' a .snd (x , p) i = _ , λ j → p (~ 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 | ||
| isContrSinglP A a .snd (x , p) i = _ , λ k → fill A (λ j → λ where | ||
| (i = i0) → transport-filler (λ i → A i) a j | ||
| (i = i1) → p j | ||
| ) (inS a) k | ||
|
|
||
| isContrSinglP' : (A : I → Type ℓ) (a : A i1) → isContr (singlP' A a) | ||
| isContrSinglP' A a .fst = _ , symP (transport-filler (λ i → A (~ i)) a) | ||
| isContrSinglP' A a .snd (x , p) i = _ , λ k → fill (λ i → A (~ i)) (λ j → λ where | ||
| (i = i0) → transport-filler (λ i → A (~ i)) a j | ||
| (i = i1) → p (~ j) | ||
| ) (inS a) (~ k) | ||
|
|
||
| -- Higher cube types | ||
|
|
||
|
|
@@ -508,6 +524,13 @@ SquareP : | |
| → Type ℓ | ||
| SquareP A a₀₋ a₁₋ a₋₀ a₋₁ = PathP (λ i → PathP (λ j → A i j) (a₋₀ i) (a₋₁ i)) a₀₋ a₁₋ | ||
|
|
||
| -- This is the type of squares: | ||
| -- a₀₀ =====> a₀₁ | ||
| -- || || | ||
| -- || || | ||
| -- \/ \/ | ||
| -- a₁₀ =====> a₁₁ | ||
|
|
||
| Square : | ||
| {a₀₀ a₀₁ : A} (a₀₋ : a₀₀ ≡ a₀₁) | ||
| {a₁₀ a₁₁ : A} (a₁₋ : a₁₀ ≡ a₁₁) | ||
|
|
@@ -587,31 +610,29 @@ isProp→PathP : ∀ {B : I → Type ℓ} → ((i : I) → isProp (B i)) | |
| → 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 | ||
| isContr→isProp (c , h) a b = sym (h a) ∙ h b | ||
|
|
||
| isContr→isSet' : isContr A → isSet' A | ||
| isContr→isSet' (c , h) p q r s i j = hcomp (λ k → λ where | ||
| (i = i0) → h (p j) k | ||
| (i = i1) → h (q j) k | ||
| (j = i0) → h (r i) k | ||
| (j = i1) → h (s i) k | ||
| ) c | ||
|
|
||
| isContr→isSet : isContr A → isSet A | ||
| isContr→isSet c = isSet'→isSet (isContr→isSet' c) | ||
|
|
||
| isPropIsContr : isProp (isContr A) | ||
| isPropIsContr (c0 , h0) (c1 , h1) i .fst = h0 c1 i | ||
| isPropIsContr (c0 , h0) (c1 , h1) i .snd y = isContr→isSet' (c0 , h0) (h0 y) (h1 y) (h0 c1) refl i | ||
|
|
||
| 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 h a = isContr→isSet (a , h a) 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 | ||
| isProp→isSet' h {a} = isContr→isSet' (a , h a) | ||
|
|
||
| isPropIsProp : isProp (isProp A) | ||
| isPropIsProp f g i a b = isProp→isSet f a b (f a b) (g a b) i | ||
|
|
||
Uh oh!
There was an error while loading. Please reload this page.