chore(library/hott) change naming to leo's naming proposal

This commit is contained in:
Jakob von Raumer 2014-11-26 16:45:12 -05:00 committed by Leonardo de Moura
parent 7374149758
commit 4587e46c67
2 changed files with 11 additions and 11 deletions

View file

@ -77,7 +77,7 @@ context
(λ xy, prod.rec_on xy
(λ b c p, path.rec_on p idp))))
theorem ua_implies_funext_nondep {A : Type} {B : Type.{l+1}}
theorem nondep_funext_from_ua {A : Type} {B : Type.{l+1}}
: Π {f g : A → B}, f g → f ≈ g :=
(λ (f g : A → B) (p : f g),
let d := λ (x : A), dpair (f x , f x) idp in
@ -104,7 +104,7 @@ end
-- Now we use this to prove weak funext, which as we know
-- implies (with dependent eta) also the strong dependent funext.
universe variables l k
theorem ua_implies_weak_funext [ua3 : ua_type.{k+1}] [ua4 : ua_type.{k+2}] : weak_funext.{l+1 k+1} :=
theorem weak_funext_from_ua [ua3 : ua_type.{k+1}] [ua4 : ua_type.{k+2}] : weak_funext.{l+1 k+1} :=
(λ (A : Type) (P : A → Type) allcontr,
let U := (λ (x : A), unit) in
have pequiv : Π (x : A), P x ≃ U x,
@ -113,10 +113,10 @@ theorem ua_implies_weak_funext [ua3 : ua_type.{k+1}] [ua4 : ua_type.{k+2}] : wea
from (λ x, @is_equiv.inv _ _
equiv_path ua_type.inst (pequiv x)),
have p : P ≈ U,
from @ua_implies_funext_nondep _ A Type P U psim,
from @nondep_funext_from_ua _ A Type P U psim,
have tU' : is_contr (A → unit),
from is_contr.mk (λ x, ⋆)
(λ f, @ua_implies_funext_nondep _ A unit (λ x, ⋆) f
(λ f, @nondep_funext_from_ua _ A unit (λ x, ⋆) f
(λ x, unit.rec_on (f x) idp)),
have tU : is_contr (Π x, U x),
from tU',
@ -126,5 +126,5 @@ theorem ua_implies_weak_funext [ua3 : ua_type.{k+1}] [ua4 : ua_type.{k+2}] : wea
)
-- In the following we will proof function extensionality using the univalence axiom
definition ua_implies_funext [instance] [ua ua2 : ua_type] : funext :=
weak_funext_implies_funext (@ua_implies_weak_funext ua ua2)
definition funext_from_ua [instance] [ua ua2 : ua_type] : funext :=
funext_from_weak_funext (@weak_funext_from_ua ua ua2)

View file

@ -25,7 +25,7 @@ definition weak_funext.{l k} :=
-- The obvious implications are Funext -> NaiveFunext -> WeakFunext
-- TODO: Get class inference to work locally
definition funext_implies_naive_funext [F : funext] : naive_funext :=
definition naive_funext_from_funext [F : funext] : naive_funext :=
(λ A P f g h,
have Fefg: is_equiv (@apD10 A P f g),
from (@funext.ap F A P f g),
@ -33,7 +33,7 @@ definition funext_implies_naive_funext [F : funext] : naive_funext :=
eq1
)
definition naive_funext_implies_weak_funext : naive_funext → weak_funext :=
definition weak_funext_from_naive_funext : naive_funext → weak_funext :=
(λ nf A P (Pc : Πx, is_contr (P x)),
let c := λx, center (P x) in
is_contr.mk c (λ f,
@ -92,7 +92,7 @@ end
-- Now the proof is fairly easy; we can just use the same induction principle on both sides.
universe variables l k
theorem weak_funext_implies_funext (wf : weak_funext.{l+1 k+1}) : funext.{l+1 k+1} :=
theorem funext_from_weak_funext (wf : weak_funext.{l+1 k+1}) : funext.{l+1 k+1} :=
funext.mk (λ A B f g,
let eq_to_f := (λ g' x, f ≈ g') in
let sim2path := htpy_ind _ f eq_to_f idp in
@ -106,5 +106,5 @@ theorem weak_funext_implies_funext (wf : weak_funext.{l+1 k+1}) : funext.{l+1 k+
from (λ h, path.rec_on h (htpy_ind_beta _ f _ idp)),
is_equiv.adjointify apD10 (sim2path g) sect retr)
definition naive_funext_implies_funext : naive_funext -> funext :=
compose weak_funext_implies_funext naive_funext_implies_weak_funext
definition funext_from_naive_funext : naive_funext -> funext :=
compose funext_from_weak_funext weak_funext_from_naive_funext