refactor(hott): remove most 'context' commands from the HoTT library
This commit is contained in:
parent
670eac9d50
commit
76bf8de91a
7 changed files with 18 additions and 18 deletions
|
@ -41,12 +41,12 @@ namespace binary
|
|||
definition right_distributive := ∀a b c, (a + b) * c = a * c + b * c
|
||||
end
|
||||
|
||||
context
|
||||
section
|
||||
variable {A : Type}
|
||||
variable {f : A → A → A}
|
||||
variable H_comm : commutative f
|
||||
variable H_assoc : associative f
|
||||
infixl `*` := f
|
||||
local infixl `*` := f
|
||||
theorem left_comm : ∀a b c, a*(b*c) = b*(a*c) :=
|
||||
take a b c, calc
|
||||
a*(b*c) = (a*b)*c : H_assoc
|
||||
|
@ -60,11 +60,11 @@ namespace binary
|
|||
... = (a*c)*b : H_assoc
|
||||
end
|
||||
|
||||
context
|
||||
section
|
||||
variable {A : Type}
|
||||
variable {f : A → A → A}
|
||||
variable H_assoc : associative f
|
||||
infixl `*` := f
|
||||
local infixl `*` := f
|
||||
theorem assoc4helper (a b c d) : (a*b)*(c*d) = a*((b*c)*d) :=
|
||||
calc
|
||||
(a*b)*(c*d) = a*(b*(c*d)) : H_assoc
|
||||
|
|
|
@ -14,7 +14,7 @@ import .funext_varieties .ua
|
|||
|
||||
open eq function prod is_trunc sigma equiv is_equiv unit
|
||||
|
||||
context
|
||||
section
|
||||
universe variables l
|
||||
|
||||
private theorem ua_isequiv_postcompose {A B : Type.{l}} {C : Type}
|
||||
|
|
|
@ -95,7 +95,7 @@ namespace is_equiv
|
|||
is_equiv.mk (inv f) sect' retr' adj'
|
||||
end
|
||||
|
||||
context
|
||||
section
|
||||
parameters {A B : Type} (f : A → B) (g : B → A)
|
||||
(ret : f ∘ g ∼ id) (sec : g ∘ f ∼ id)
|
||||
|
||||
|
|
|
@ -16,7 +16,7 @@ private definition const {A B : Type} (f : A → B) := ∀ x y, f x = f y
|
|||
private definition coll (A : Type) := Σ f : A → A, const f
|
||||
private definition path_coll (A : Type) := ∀ x y : A, coll (x = y)
|
||||
|
||||
context
|
||||
section
|
||||
parameter {A : Type}
|
||||
hypothesis [h : decidable_eq A]
|
||||
variables {x y : A}
|
||||
|
|
|
@ -56,10 +56,10 @@ namespace prod
|
|||
intro : ∀{a₁ b₁ a₂ b₂}, Ra a₁ a₂ → Rb b₁ b₂ → rprod (a₁, b₁) (a₂, b₂)
|
||||
end
|
||||
|
||||
context
|
||||
section
|
||||
parameters {A B : Type}
|
||||
parameters {Ra : A → A → Type} {Rb : B → B → Type}
|
||||
infix `≺`:50 := lex Ra Rb
|
||||
local infix `≺`:50 := lex Ra Rb
|
||||
|
||||
definition lex.accessible {a} (aca : acc Ra a) (acb : ∀b, acc Rb b): ∀b, acc (lex Ra Rb) (a, b) :=
|
||||
acc.rec_on aca
|
||||
|
|
|
@ -26,9 +26,9 @@ namespace well_founded
|
|||
definition apply [coercion] {A : Type} {R : A → A → Type} (wf : well_founded R) : ∀a, acc R a :=
|
||||
take a, well_founded.rec_on wf (λp, p) a
|
||||
|
||||
context
|
||||
section
|
||||
parameters {A : Type} {R : A → A → Type}
|
||||
infix `≺`:50 := R
|
||||
local infix `≺`:50 := R
|
||||
|
||||
hypothesis [Hwf : well_founded R]
|
||||
|
||||
|
@ -98,7 +98,7 @@ well_founded.intro (λ (a : A),
|
|||
|
||||
-- Subrelation of a well-founded relation is well-founded
|
||||
namespace subrelation
|
||||
context
|
||||
section
|
||||
parameters {A : Type} {R Q : A → A → Type}
|
||||
parameters (H₁ : subrelation Q R)
|
||||
parameters (H₂ : well_founded R)
|
||||
|
@ -115,7 +115,7 @@ end subrelation
|
|||
|
||||
-- The inverse image of a well-founded relation is well-founded
|
||||
namespace inv_image
|
||||
context
|
||||
section
|
||||
parameters {A B : Type} {R : B → B → Type}
|
||||
parameters (f : A → B)
|
||||
parameters (H : well_founded R)
|
||||
|
@ -136,9 +136,9 @@ end inv_image
|
|||
|
||||
-- The transitive closure of a well-founded relation is well-founded
|
||||
namespace tc
|
||||
context
|
||||
section
|
||||
parameters {A : Type} {R : A → A → Type}
|
||||
notation `R⁺` := tc R
|
||||
local notation `R⁺` := tc R
|
||||
|
||||
definition accessible {z} (ac: acc R z) : acc R⁺ z :=
|
||||
acc.rec_on ac
|
||||
|
|
|
@ -228,9 +228,9 @@ namespace sigma
|
|||
: (Σa, B a) ≃ (Σa', B' a') :=
|
||||
equiv.mk (sigma_functor f g) !is_equiv_sigma_functor
|
||||
|
||||
context
|
||||
attribute inv [irreducible]
|
||||
attribute function.compose [irreducible] --this is needed for the following class inference problem
|
||||
section
|
||||
local attribute inv [irreducible]
|
||||
local attribute function.compose [irreducible] --this is needed for the following class inference problem
|
||||
definition sigma_equiv_sigma (Hf : A ≃ A') (Hg : Π a, B a ≃ B' (to_fun Hf a)) :
|
||||
(Σa, B a) ≃ (Σa', B' a') :=
|
||||
sigma_equiv_sigma_of_is_equiv (to_fun Hf) (λ a, to_fun (Hg a))
|
||||
|
|
Loading…
Reference in a new issue