refactor(library): avoid 'context' command in the standard library
This commit is contained in:
parent
ce153d01f9
commit
670eac9d50
11 changed files with 28 additions and 28 deletions
|
@ -44,12 +44,12 @@ namespace binary
|
|||
definition left_commutative [reducible] {B : Type} (f : A → B → B) := ∀ a₁ a₂ b, f a₁ (f a₂ b) = f a₂ (f a₁ b)
|
||||
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 : left_commutative f :=
|
||||
take a b c, calc
|
||||
a*(b*c) = (a*b)*c : H_assoc
|
||||
|
@ -63,11 +63,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
|
||||
|
|
|
@ -79,7 +79,7 @@ namespace category
|
|||
(λ a b f, @subsingleton.elim (set_hom a b) _ _ _)
|
||||
local attribute discrete_category [reducible]
|
||||
definition Discrete_category (A : Type) [H : decidable_eq A] := Mk (discrete_category A)
|
||||
context
|
||||
section
|
||||
local attribute discrete_category [instance]
|
||||
include H
|
||||
theorem discrete_category.endomorphism {a b : A} (f : a ⟶ b) : a = b :=
|
||||
|
|
|
@ -489,7 +489,7 @@ section
|
|||
abs.by_cases H1 H2
|
||||
|
||||
-- the triangle inequality
|
||||
context
|
||||
section
|
||||
private lemma aux1 {a b : A} (H1 : a + b ≥ 0) (H2 : a ≥ 0) : abs (a + b) ≤ abs a + abs b :=
|
||||
decidable.by_cases
|
||||
(assume H3 : b ≥ 0,
|
||||
|
|
|
@ -128,8 +128,8 @@ calc
|
|||
sub_nat_nat m n = nat.cases_on 0 (of_nat (m - n)) _ : H1 ▸ rfl
|
||||
... = of_nat (m - n) : rfl
|
||||
|
||||
context
|
||||
attribute sub_nat_nat [reducible]
|
||||
section
|
||||
local attribute sub_nat_nat [reducible]
|
||||
theorem sub_nat_nat_of_lt {m n : ℕ} (H : m < n) :
|
||||
sub_nat_nat m n = neg_succ_of_nat (pred (n - m)) :=
|
||||
have H1 : n - m = succ (pred (n - m)), from (succ_pred_of_pos (sub_pos_of_lt H))⁻¹,
|
||||
|
@ -276,9 +276,9 @@ calc
|
|||
... = abstr (repr b) : abstr_eq H
|
||||
... = b : abstr_repr
|
||||
|
||||
context
|
||||
attribute abstr [reducible]
|
||||
attribute dist [reducible]
|
||||
section
|
||||
local attribute abstr [reducible]
|
||||
local attribute dist [reducible]
|
||||
theorem nat_abs_abstr (p : ℕ × ℕ) : nat_abs (abstr p) = dist (pr1 p) (pr2 p) :=
|
||||
let m := pr1 p, n := pr2 p in
|
||||
or.elim (@le_or_gt n m)
|
||||
|
@ -458,8 +458,8 @@ have H3 : pabs (padd (repr a) (repr b)) ≤ pabs (repr a) + pabs (repr b),
|
|||
from !dist_add_add_le_add_dist_dist,
|
||||
H⁻¹ ▸ H1⁻¹ ▸ H2⁻¹ ▸ H3
|
||||
|
||||
context
|
||||
attribute nat_abs [reducible]
|
||||
section
|
||||
local attribute nat_abs [reducible]
|
||||
theorem mul_nat_abs (a b : ℤ) : nat_abs (a * b) = #nat (nat_abs a) * (nat_abs b) :=
|
||||
int.cases_on a
|
||||
(take m,
|
||||
|
|
|
@ -16,7 +16,7 @@ import data.subtype data.nat.basic data.nat.order
|
|||
open nat subtype decidable well_founded
|
||||
|
||||
namespace nat
|
||||
context find_x
|
||||
section find_x
|
||||
parameter {p : nat → Prop}
|
||||
|
||||
private definition lbp (x : nat) : Prop := ∀ y, y < x → ¬ p y
|
||||
|
|
|
@ -32,7 +32,7 @@ namespace function
|
|||
mk_equivalence (@equiv A B) (@equiv.refl A B) (@equiv.symm A B) (@equiv.trans A B)
|
||||
end function
|
||||
|
||||
context
|
||||
section
|
||||
open quot
|
||||
variables {A : Type} {B : A → Type}
|
||||
|
||||
|
|
|
@ -273,8 +273,8 @@ namespace nat
|
|||
|
||||
notation a * b := mul a b
|
||||
|
||||
context
|
||||
attribute sub [reducible]
|
||||
section
|
||||
local attribute sub [reducible]
|
||||
definition succ_sub_succ_eq_sub (a b : nat) : succ a - succ b = a - b :=
|
||||
nat.induction_on b
|
||||
rfl
|
||||
|
|
|
@ -46,10 +46,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 → Prop} {Rb : B → B → Prop}
|
||||
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
|
||||
|
|
|
@ -42,10 +42,10 @@ namespace sigma
|
|||
| right : ∀a {b₁ b₂}, Rb a b₁ b₂ → lex ⟨a, b₁⟩ ⟨a, b₂⟩
|
||||
end
|
||||
|
||||
context
|
||||
section
|
||||
parameters {A : Type} {B : A → Type}
|
||||
parameters {Ra : A → A → Prop} {Rb : Π a : A, B a → B a → Prop}
|
||||
infix `≺`:50 := lex Ra Rb
|
||||
local infix `≺`:50 := lex Ra Rb
|
||||
|
||||
definition lex.accessible {a} (aca : acc Ra a) (acb : ∀a, well_founded (Rb a)) : ∀ (b : B a), acc (lex Ra Rb) ⟨a, b⟩ :=
|
||||
acc.rec_on aca
|
||||
|
|
|
@ -25,9 +25,9 @@ namespace well_founded
|
|||
definition apply [coercion] {A : Type} {R : A → A → Prop} (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 → Prop}
|
||||
infix `≺`:50 := R
|
||||
local infix `≺`:50 := R
|
||||
|
||||
hypothesis [Hwf : well_founded R]
|
||||
|
||||
|
@ -81,7 +81,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 → Prop}
|
||||
parameters (H₁ : subrelation Q R)
|
||||
parameters (H₂ : well_founded R)
|
||||
|
@ -98,7 +98,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 → Prop}
|
||||
parameters (f : A → B)
|
||||
parameters (H : well_founded R)
|
||||
|
@ -119,9 +119,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 → Prop}
|
||||
notation `R⁺` := tc R
|
||||
local notation `R⁺` := tc R
|
||||
|
||||
definition accessible {z} (ac: acc R z) : acc R⁺ z :=
|
||||
acc.rec_on ac
|
||||
|
|
|
@ -8,7 +8,7 @@ open eq.ops nonempty inhabited
|
|||
-- Diaconescu’s theorem
|
||||
-- Show that Excluded middle follows from
|
||||
-- Hilbert's choice operator, function extensionality and Prop extensionality
|
||||
context
|
||||
section
|
||||
parameter p : Prop
|
||||
|
||||
private definition u [reducible] := epsilon (λx, x = true ∨ p)
|
||||
|
|
Loading…
Reference in a new issue