[Agda] Understanding Glue

Thierry Coquand Thierry.Coquand at cse.gu.se
Wed May 20 17:02:54 CEST 2020


 Indeed, it would be clearer to present the Glue, unglue construction at the pre-type
level. If we have A and a partial map  w: T -> A defined on phi then we can find an extension

 G = Glue (A, phi, w)    and    unglue :  G -> A

such that (G,unglue) = (T,w) on the extent phi.
 This can be done for any function w.


 The main Lemma is then that if A, T are fibrant and w is an equivalence, then G is fibrant
and unglue is an equivalence.
(This expresses that the type   Sigma (X:U) Equiv X A   is contractible, which is one way
to state univalence.)

Best, Thierry



On 20 May 2020, at 11:34, Thorsten Altenkirch <Thorsten.Altenkirch at nottingham.ac.uk<mailto:Thorsten.Altenkirch at nottingham.ac.uk>> wrote:

Hi,

I am trying to understand the Glue construction better. I was wondering whether Glue just means that a certain pre-type is fibrant (i.e. admits composition, i.e. transp and hcomp). I was trying in agda

MyGlue : (A : Set) {φ : I}
     → Partial φ (Σ[ T ∈ Set ] T ≃' A) → Set
MyGlue A {φ} Te = Σ (PartialP φ λ o → fst (Te o)) (λ t → A [ φ ↦ (λ p → fst (snd (Te p)) (t p)) ])

I have switched off universe checking, hence Setw = Set. Basically I have used the domain of glue as the definition of Glue. However, this doesn’t seem enough to derive

ua : {A B : Set} → (p : A ≃' B) → A ≡ B

because we need the extra definitional equality

MyGlue A i1 (T , e) = T

Is there a way to avoid this? This is not a provable equality either, I think.

Cheers,
Thorsten

This message and any attachment are intended solely for the addressee
and may contain confidential information. If you have received this
message in error, please contact the sender and delete the email and
attachment.

Any views or opinions expressed by the author of this email do not
necessarily reflect the views of the University of Nottingham. Email
communications with the University of Nottingham may be monitored
where permitted by law.






-------------- next part --------------
An HTML attachment was scrubbed...
URL: <http://lists.chalmers.se/pipermail/agda/attachments/20200520/0fba6143/attachment.html>


More information about the Agda mailing list