refactor(library): use '[protected]' modifier
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
e14814d4bf
commit
f891485a26
8 changed files with 32 additions and 32 deletions
|
@ -33,13 +33,13 @@ section
|
|||
|
||||
variable {T : Type}
|
||||
|
||||
theorem list_induction_on {P : list T → Prop} (l : list T) (Hnil : P nil)
|
||||
theorem induction_on [protected] {P : list T → Prop} (l : list T) (Hnil : P nil)
|
||||
(Hind : forall x : T, forall l : list T, forall H : P l, P (cons x l)) : P l :=
|
||||
list_rec Hnil Hind l
|
||||
|
||||
theorem list_cases_on {P : list T → Prop} (l : list T) (Hnil : P nil)
|
||||
theorem cases_on [protected] {P : list T → Prop} (l : list T) (Hnil : P nil)
|
||||
(Hcons : forall x : T, forall l : list T, P (cons x l)) : P l :=
|
||||
list_induction_on l Hnil (take x l IH, Hcons x l)
|
||||
induction_on l Hnil (take x l IH, Hcons x l)
|
||||
|
||||
notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l
|
||||
|
||||
|
@ -57,12 +57,12 @@ theorem nil_concat {t : list T} : nil ++ t = t
|
|||
theorem cons_concat {x : T} {s t : list T} : (x :: s) ++ t = x :: (s ++ t)
|
||||
|
||||
theorem concat_nil {t : list T} : t ++ nil = t :=
|
||||
list_induction_on t rfl
|
||||
induction_on t rfl
|
||||
(take (x : T) (l : list T) (H : concat l nil = l),
|
||||
show concat (cons x l) nil = cons x l, from H ▸ rfl)
|
||||
|
||||
theorem concat_assoc {s t u : list T} : s ++ t ++ u = s ++ (t ++ u) :=
|
||||
list_induction_on s rfl
|
||||
induction_on s rfl
|
||||
(take x l,
|
||||
assume H : concat (concat l t) u = concat l (concat t u),
|
||||
calc
|
||||
|
@ -80,7 +80,7 @@ theorem length_nil : length (@nil T) = 0 := rfl
|
|||
theorem length_cons {x : T} {t : list T} : length (x :: t) = succ (length t)
|
||||
|
||||
theorem length_concat {s t : list T} : length (s ++ t) = length s + length t :=
|
||||
list_induction_on s
|
||||
induction_on s
|
||||
(calc
|
||||
length (concat nil t) = length t : rfl
|
||||
... = zero + length t : {add_zero_left⁻¹}
|
||||
|
@ -120,7 +120,7 @@ theorem reverse_cons {x : T} {l : list T} : reverse (x :: l) = append x (reverse
|
|||
theorem reverse_singleton {x : T} : reverse [x] = [x]
|
||||
|
||||
theorem reverse_concat {s t : list T} : reverse (s ++ t) = (reverse t) ++ (reverse s) :=
|
||||
list_induction_on s (concat_nil⁻¹)
|
||||
induction_on s (concat_nil⁻¹)
|
||||
(take x s,
|
||||
assume IH : reverse (s ++ t) = concat (reverse t) (reverse s),
|
||||
calc
|
||||
|
@ -130,7 +130,7 @@ list_induction_on s (concat_nil⁻¹)
|
|||
... = reverse t ++ (reverse (x :: s)) : rfl)
|
||||
|
||||
theorem reverse_reverse {l : list T} : reverse (reverse l) = l :=
|
||||
list_induction_on l rfl
|
||||
induction_on l rfl
|
||||
(take x l',
|
||||
assume H: reverse (reverse l') = l',
|
||||
show reverse (reverse (x :: l')) = x :: l', from
|
||||
|
@ -141,7 +141,7 @@ list_induction_on l rfl
|
|||
... = x :: l' : rfl)
|
||||
|
||||
theorem append_eq_reverse_cons {x : T} {l : list T} : append x l = reverse (x :: reverse l) :=
|
||||
list_induction_on l rfl
|
||||
induction_on l rfl
|
||||
(take y l',
|
||||
assume H : append x l' = reverse (x :: reverse l'),
|
||||
calc
|
||||
|
@ -160,7 +160,7 @@ theorem head_nil {x : T} : head x (@nil T) = x
|
|||
theorem head_cons {x x' : T} {t : list T} : head x' (x :: t) = x
|
||||
|
||||
theorem head_concat {s t : list T} {x : T} : s ≠ nil → (head x (s ++ t) = head x s) :=
|
||||
list_cases_on s
|
||||
cases_on s
|
||||
(take H : nil ≠ nil, absurd rfl H)
|
||||
(take x s, take H : cons x s ≠ nil,
|
||||
calc
|
||||
|
@ -175,7 +175,7 @@ theorem tail_nil : tail (@nil T) = nil
|
|||
theorem tail_cons {x : T} {l : list T} : tail (cons x l) = l
|
||||
|
||||
theorem cons_head_tail {x : T} {l : list T} : l ≠ nil → (head x l) :: (tail l) = l :=
|
||||
list_cases_on l
|
||||
cases_on l
|
||||
(assume H : nil ≠ nil, absurd rfl H)
|
||||
(take x l, assume H : cons x l ≠ nil, rfl)
|
||||
|
||||
|
@ -192,7 +192,7 @@ theorem mem_nil {x : T} : x ∈ nil ↔ false := iff_rfl
|
|||
theorem mem_cons {x y : T} {l : list T} : mem x (cons y l) ↔ (x = y ∨ mem x l) := iff_rfl
|
||||
|
||||
theorem mem_concat_imp_or {x : T} {s t : list T} : x ∈ s ++ t → x ∈ s ∨ x ∈ t :=
|
||||
list_induction_on s or_inr
|
||||
induction_on s or_inr
|
||||
(take y s,
|
||||
assume IH : x ∈ s ++ t → x ∈ s ∨ x ∈ t,
|
||||
assume H1 : x ∈ (y :: s) ++ t,
|
||||
|
@ -201,7 +201,7 @@ list_induction_on s or_inr
|
|||
iff_elim_right or_assoc H3)
|
||||
|
||||
theorem mem_or_imp_concat {x : T} {s t : list T} : x ∈ s ∨ x ∈ t → x ∈ s ++ t :=
|
||||
list_induction_on s
|
||||
induction_on s
|
||||
(take H, or_elim H false_elim (assume H, H))
|
||||
(take y s,
|
||||
assume IH : x ∈ s ∨ x ∈ t → x ∈ s ++ t,
|
||||
|
@ -217,7 +217,7 @@ theorem mem_concat {x : T} {s t : list T} : x ∈ s ++ t ↔ x ∈ s ∨ x ∈ t
|
|||
:= iff_intro mem_concat_imp_or mem_or_imp_concat
|
||||
|
||||
theorem mem_split {x : T} {l : list T} : x ∈ l → ∃s t : list T, l = s ++ (x :: t) :=
|
||||
list_induction_on l
|
||||
induction_on l
|
||||
(take H : x ∈ nil, false_elim (iff_elim_left mem_nil H))
|
||||
(take y l,
|
||||
assume IH : x ∈ l → ∃s t : list T, l = s ++ (x :: t),
|
||||
|
@ -249,8 +249,8 @@ list_induction_on l
|
|||
|
||||
-- theorem not_mem_find (l : list T) (x : T) : ¬ mem x l → find x l = length l
|
||||
-- :=
|
||||
-- @list_induction_on T (λl, ¬ mem x l → find x l = length l) l
|
||||
-- -- list_induction_on l
|
||||
-- @induction_on T (λl, ¬ mem x l → find x l = length l) l
|
||||
-- -- induction_on l
|
||||
-- (assume P1 : ¬ mem x nil,
|
||||
-- show find x nil = length nil, from
|
||||
-- calc
|
||||
|
|
|
@ -29,11 +29,11 @@ theorem nat_rec_zero {P : ℕ → Type} (x : P zero) (f : ∀m, P m → P (succ
|
|||
theorem nat_rec_succ {P : ℕ → Type} (x : P zero) (f : ∀m, P m → P (succ m)) (n : ℕ) :
|
||||
nat_rec x f (succ n) = f n (nat_rec x f n)
|
||||
|
||||
theorem induction_on {P : ℕ → Prop} (a : ℕ) (H1 : P zero) (H2 : ∀ (n : ℕ) (IH : P n), P (succ n)) :
|
||||
theorem induction_on [protected] {P : ℕ → Prop} (a : ℕ) (H1 : P zero) (H2 : ∀ (n : ℕ) (IH : P n), P (succ n)) :
|
||||
P a :=
|
||||
nat_rec H1 H2 a
|
||||
|
||||
definition rec_on {P : ℕ → Type} (n : ℕ) (H1 : P zero) (H2 : ∀m, P m → P (succ m)) : P n :=
|
||||
definition rec_on [protected] {P : ℕ → Type} (n : ℕ) (H1 : P zero) (H2 : ∀m, P m → P (succ m)) : P n :=
|
||||
nat_rec H1 H2 n
|
||||
|
||||
|
||||
|
|
|
@ -432,7 +432,7 @@ theorem ge_decidable [instance] (n m : ℕ) : decidable (n ≥ m)
|
|||
|
||||
-- ### misc
|
||||
|
||||
theorem strong_induction_on {P : nat → Prop} (n : ℕ) (H : ∀n, (∀m, m < n → P m) → P n) : P n :=
|
||||
theorem strong_induction_on [protected] {P : nat → Prop} (n : ℕ) (H : ∀n, (∀m, m < n → P m) → P n) : P n :=
|
||||
have H1 : ∀ {n m : nat}, m < n → P m, from
|
||||
take n,
|
||||
induction_on n
|
||||
|
@ -448,7 +448,7 @@ have H1 : ∀ {n m : nat}, m < n → P m, from
|
|||
(assume H4: m = n', H4⁻¹ ▸ H2)),
|
||||
H1 self_lt_succ
|
||||
|
||||
theorem case_strong_induction_on {P : nat → Prop} (a : nat) (H0 : P 0)
|
||||
theorem case_strong_induction_on [protected] {P : nat → Prop} (a : nat) (H0 : P 0)
|
||||
(Hind : ∀(n : nat), (∀m, m ≤ n → P m) → P (succ n)) : P a :=
|
||||
strong_induction_on a (
|
||||
take n,
|
||||
|
|
|
@ -11,11 +11,11 @@ inductive option (A : Type) : Type :=
|
|||
none {} : option A,
|
||||
some : A → option A
|
||||
|
||||
theorem induction_on {A : Type} {p : option A → Prop} (o : option A)
|
||||
theorem induction_on [protected] {A : Type} {p : option A → Prop} (o : option A)
|
||||
(H1 : p none) (H2 : ∀a, p (some a)) : p o :=
|
||||
option_rec H1 H2 o
|
||||
|
||||
definition rec_on {A : Type} {C : option A → Type} (o : option A)
|
||||
definition rec_on [protected] {A : Type} {C : option A → Type} (o : option A)
|
||||
(H1 : C none) (H2 : ∀a, C (some a)) : C o :=
|
||||
option_rec H1 H2 o
|
||||
|
||||
|
|
|
@ -24,11 +24,11 @@ namespace sum_plus_notation
|
|||
infixr `+`:25 := sum -- conflicts with notation for addition
|
||||
end sum_plus_notation
|
||||
|
||||
abbreviation rec_on {A B : Type} {C : (A ⊎ B) → Type} (s : A ⊎ B)
|
||||
abbreviation rec_on [protected] {A B : Type} {C : (A ⊎ B) → Type} (s : A ⊎ B)
|
||||
(H1 : ∀a : A, C (inl B a)) (H2 : ∀b : B, C (inr A b)) : C s :=
|
||||
sum_rec H1 H2 s
|
||||
|
||||
abbreviation cases_on {A B : Type} {P : (A ⊎ B) → Prop} (s : A ⊎ B)
|
||||
abbreviation cases_on [protected] {A B : Type} {P : (A ⊎ B) → Prop} (s : A ⊎ B)
|
||||
(H1 : ∀a : A, P (inl B a)) (H2 : ∀b : B, P (inr A b)) : P s :=
|
||||
sum_rec H1 H2 s
|
||||
|
||||
|
|
|
@ -23,14 +23,14 @@ notation x `≈` y:50 `:>`:0 A:0 := @path A x y -- TODO: is this right?
|
|||
notation `idp`:max := idpath _ -- TODO: can we / should we use `1`?
|
||||
|
||||
namespace path
|
||||
abbreviation induction_on {A : Type} {a b : A} (p : a ≈ b)
|
||||
abbreviation induction_on [protected] {A : Type} {a b : A} (p : a ≈ b)
|
||||
{C : Π (b : A) (p : a ≈ b), Type} (H : C a (idpath a)) : C b p :=
|
||||
path_rec H p
|
||||
end path
|
||||
|
||||
|
||||
-- TODO: should all this be in namespace path?
|
||||
using path
|
||||
using path (induction_on)
|
||||
|
||||
-- Concatenation and inverse
|
||||
-- -------------------------
|
||||
|
|
|
@ -16,10 +16,10 @@ inl trivial
|
|||
theorem false_decidable [instance] : decidable false :=
|
||||
inr not_false_trivial
|
||||
|
||||
theorem induction_on {p : Prop} {C : Prop} (H : decidable p) (H1 : p → C) (H2 : ¬p → C) : C :=
|
||||
theorem induction_on [protected] {p : Prop} {C : Prop} (H : decidable p) (H1 : p → C) (H2 : ¬p → C) : C :=
|
||||
decidable_rec H1 H2 H
|
||||
|
||||
definition rec_on [inline] {p : Prop} {C : Type} (H : decidable p) (H1 : p → C) (H2 : ¬p → C) :
|
||||
definition rec_on [protected] [inline] {p : Prop} {C : Type} (H : decidable p) (H1 : p → C) (H2 : ¬p → C) :
|
||||
C :=
|
||||
decidable_rec H1 H2 H
|
||||
|
||||
|
|
|
@ -8,7 +8,7 @@ import logic.classes.decidable tools.tactic
|
|||
using decidable tactic eq_ops
|
||||
|
||||
definition ite (c : Prop) {H : decidable c} {A : Type} (t e : A) : A :=
|
||||
rec_on H (assume Hc, t) (assume Hnc, e)
|
||||
decidable.rec_on H (assume Hc, t) (assume Hnc, e)
|
||||
|
||||
notation `if` c `then` t `else` e:45 := ite c t e
|
||||
|
||||
|
@ -38,11 +38,11 @@ if_neg not_false_trivial
|
|||
|
||||
theorem if_cond_congr {c₁ c₂ : Prop} {H₁ : decidable c₁} {H₂ : decidable c₂} (Heq : c₁ ↔ c₂) {A : Type} (t e : A)
|
||||
: (if c₁ then t else e) = (if c₂ then t else e) :=
|
||||
rec_on H₁
|
||||
(assume Hc₁ : c₁, rec_on H₂
|
||||
decidable.rec_on H₁
|
||||
(assume Hc₁ : c₁, decidable.rec_on H₂
|
||||
(assume Hc₂ : c₂, if_pos Hc₁ ⬝ (if_pos Hc₂)⁻¹)
|
||||
(assume Hnc₂ : ¬c₂, absurd (iff_elim_left Heq Hc₁) Hnc₂))
|
||||
(assume Hnc₁ : ¬c₁, rec_on H₂
|
||||
(assume Hnc₁ : ¬c₁, decidable.rec_on H₂
|
||||
(assume Hc₂ : c₂, absurd (iff_elim_right Heq Hc₂) Hnc₁)
|
||||
(assume Hnc₂ : ¬c₂, if_neg Hnc₁ ⬝ (if_neg Hnc₂)⁻¹))
|
||||
|
||||
|
|
Loading…
Reference in a new issue