refactor(library,hott): use/test new 'contradiction' tactic in the standard and hott libraries
This commit is contained in:
parent
9c8a63caec
commit
9760968b45
11 changed files with 30 additions and 30 deletions
|
@ -37,7 +37,7 @@ namespace nat
|
|||
nat.rec_on m
|
||||
(λ n, nat.cases_on n
|
||||
(inl rfl)
|
||||
(λ m, inr (λ (e : succ m = zero), down (nat.no_confusion e))))
|
||||
(λ m, inr (by contradiction)))
|
||||
(λ (m' : nat) (ih : ∀n, decidable (n = m')) (n : nat), nat.cases_on n
|
||||
(inr (λ h, down (nat.no_confusion h)))
|
||||
(λ (n' : nat),
|
||||
|
@ -79,8 +79,8 @@ namespace nat
|
|||
definition not_lt_zero (a : nat) : ¬ a < zero :=
|
||||
have aux : ∀ {b}, a < b → b = zero → empty, from
|
||||
λ b H, lt.cases_on H
|
||||
(λ heq, down (nat.no_confusion heq))
|
||||
(λ b h₁ heq, down (nat.no_confusion heq)),
|
||||
(by contradiction)
|
||||
(by contradiction),
|
||||
λ H, aux H rfl
|
||||
|
||||
definition zero_lt_succ (a : nat) : zero < succ a :=
|
||||
|
|
|
@ -52,7 +52,7 @@ theorem find_discr_cons_of_eq {f g : A → B} {a : A} (l : list A) : f a = g a
|
|||
assume eq, if_pos eq
|
||||
|
||||
theorem ne_of_find_discr_eq_some {f g : A → B} {a : A} : ∀ {l}, find_discr f g l = some a → f a ≠ g a
|
||||
| [] e := option.no_confusion e
|
||||
| [] e := by contradiction
|
||||
| (x::l) e := by_cases
|
||||
(λ h : f x = g x,
|
||||
have aux : find_discr f g l = some a, by rewrite [find_discr_cons_of_eq l h at e]; exact e,
|
||||
|
|
|
@ -108,10 +108,10 @@ int.cases_on a
|
|||
int.cases_on b
|
||||
(take n,
|
||||
if H : m = n then inl (congr_arg of_nat H) else inr (take H1, H (of_nat.inj H1)))
|
||||
(take n', inr (assume H, int.no_confusion H)))
|
||||
(take n', inr (by contradiction)))
|
||||
(take m',
|
||||
int.cases_on b
|
||||
(take n, inr (assume H, int.no_confusion H))
|
||||
(take n, inr (by contradiction))
|
||||
(take n',
|
||||
(if H : m' = n' then inl (congr_arg neg_succ_of_nat H) else
|
||||
inr (take H1, H (neg_succ_of_nat.inj H1)))))
|
||||
|
|
|
@ -406,8 +406,8 @@ end nth
|
|||
open decidable
|
||||
definition has_decidable_eq {A : Type} [H : decidable_eq A] : ∀ l₁ l₂ : list A, decidable (l₁ = l₂)
|
||||
| [] [] := inl rfl
|
||||
| [] (b::l₂) := inr (λ H, list.no_confusion H)
|
||||
| (a::l₁) [] := inr (λ H, list.no_confusion H)
|
||||
| [] (b::l₂) := inr (by contradiction)
|
||||
| (a::l₁) [] := inr (by contradiction)
|
||||
| (a::l₁) (b::l₂) :=
|
||||
match H a b with
|
||||
| inl Hab :=
|
||||
|
|
|
@ -24,17 +24,17 @@ theorem eq_nil_of_perm_nil {l₁ : list A} (p : [] ~ l₁) : l₁ = [] :=
|
|||
have gen : ∀ (l₂ : list A) (p : l₂ ~ l₁), l₂ = [] → l₁ = [], from
|
||||
take l₂ p, perm.induction_on p
|
||||
(λ h, h)
|
||||
(λ x y l₁ l₂ p₁ r₁, list.no_confusion r₁)
|
||||
(λ x y l e, list.no_confusion e)
|
||||
(λ x y l₁ l₂ p₁ r₁, by contradiction)
|
||||
(λ x y l e, by contradiction)
|
||||
(λ l₁ l₂ l₃ p₁ p₂ r₁ r₂ e, r₂ (r₁ e)),
|
||||
gen [] p rfl
|
||||
|
||||
theorem not_perm_nil_cons (x : A) (l : list A) : ¬ [] ~ (x::l) :=
|
||||
have gen : ∀ (l₁ l₂ : list A) (p : l₁ ~ l₂), l₁ = [] → l₂ = (x::l) → false, from
|
||||
take l₁ l₂ p, perm.induction_on p
|
||||
(λ e₁ e₂, list.no_confusion e₂)
|
||||
(λ x l₁ l₂ p₁ r₁ e₁ e₂, list.no_confusion e₁)
|
||||
(λ x y l e₁ e₂, list.no_confusion e₁)
|
||||
(λ e₁ e₂, by contradiction)
|
||||
(λ x l₁ l₂ p₁ r₁ e₁ e₂, by contradiction)
|
||||
(λ x y l e₁ e₂, by contradiction)
|
||||
(λ l₁ l₂ l₃ p₁ p₂ r₁ r₂ e₁ e₂,
|
||||
begin
|
||||
rewrite [e₂ at *, e₁ at *],
|
||||
|
|
|
@ -47,7 +47,7 @@ nat.induction_on x
|
|||
/- successor and predecessor -/
|
||||
|
||||
theorem succ_ne_zero (n : ℕ) : succ n ≠ 0 :=
|
||||
assume H, nat.no_confusion H
|
||||
by contradiction
|
||||
|
||||
-- add_rewrite succ_ne_zero
|
||||
|
||||
|
|
|
@ -61,17 +61,17 @@ namespace pos_num
|
|||
|
||||
theorem decidable_eq [instance] : ∀ (a b : pos_num), decidable (a = b)
|
||||
| one one := inl rfl
|
||||
| one (bit0 b) := inr (λ H, pos_num.no_confusion H)
|
||||
| one (bit1 b) := inr (λ H, pos_num.no_confusion H)
|
||||
| (bit0 a) one := inr (λ H, pos_num.no_confusion H)
|
||||
| one (bit0 b) := inr (by contradiction)
|
||||
| one (bit1 b) := inr (by contradiction)
|
||||
| (bit0 a) one := inr (by contradiction)
|
||||
| (bit0 a) (bit0 b) :=
|
||||
match decidable_eq a b with
|
||||
| inl H₁ := inl (eq.rec_on H₁ rfl)
|
||||
| inr H₁ := inr (λ H, pos_num.no_confusion H (λ H₂, absurd H₂ H₁))
|
||||
end
|
||||
| (bit0 a) (bit1 b) := inr (λ H, pos_num.no_confusion H)
|
||||
| (bit1 a) one := inr (λ H, pos_num.no_confusion H)
|
||||
| (bit1 a) (bit0 b) := inr (λ H, pos_num.no_confusion H)
|
||||
| (bit0 a) (bit1 b) := inr (by contradiction)
|
||||
| (bit1 a) one := inr (by contradiction)
|
||||
| (bit1 a) (bit0 b) := inr (by contradiction)
|
||||
| (bit1 a) (bit1 b) :=
|
||||
match decidable_eq a b with
|
||||
| inl H₁ := inl (eq.rec_on H₁ rfl)
|
||||
|
|
|
@ -20,7 +20,7 @@ namespace option
|
|||
not_false
|
||||
|
||||
theorem none_ne_some {A : Type} (a : A) : none ≠ some a :=
|
||||
assume H, option.no_confusion H
|
||||
by contradiction
|
||||
|
||||
theorem some.inj {A : Type} {a₁ a₂ : A} (H : some a₁ = some a₂) : a₁ = a₂ :=
|
||||
option.no_confusion H (λe, e)
|
||||
|
|
|
@ -22,10 +22,10 @@ namespace sum
|
|||
variables {A B : Type}
|
||||
|
||||
definition inl_ne_inr (a : A) (b : B) : inl a ≠ inr b :=
|
||||
assume H, sum.no_confusion H
|
||||
by contradiction
|
||||
|
||||
definition inr_ne_inl (b : B) (a : A) : inr b ≠ inl a :=
|
||||
assume H, sum.no_confusion H
|
||||
by contradiction
|
||||
|
||||
definition inl_inj {a₁ a₂ : A} : intro_left B a₁ = intro_left B a₂ → a₁ = a₂ :=
|
||||
assume H, sum.no_confusion H (λe, e)
|
||||
|
@ -45,8 +45,8 @@ namespace sum
|
|||
| decidable.inl hp := decidable.inl (hp ▸ rfl)
|
||||
| decidable.inr hn := decidable.inr (λ he, absurd (inl_inj he) hn)
|
||||
end
|
||||
| has_decidable_eq (inl a₁) (inr b₂) := decidable.inr (λ e, sum.no_confusion e)
|
||||
| has_decidable_eq (inr b₁) (inl a₂) := decidable.inr (λ e, sum.no_confusion e)
|
||||
| has_decidable_eq (inl a₁) (inr b₂) := decidable.inr (by contradiction)
|
||||
| has_decidable_eq (inr b₁) (inl a₂) := decidable.inr (by contradiction)
|
||||
| has_decidable_eq (inr b₁) (inr b₂) :=
|
||||
match h₂ b₁ b₂ with
|
||||
| decidable.inl hp := decidable.inl (hp ▸ rfl)
|
||||
|
|
|
@ -44,8 +44,8 @@ namespace nat
|
|||
(acc.intro zero (λ (y : nat) (hlt : y < zero),
|
||||
have aux : ∀ {n₁}, y < n₁ → zero = n₁ → acc lt y, from
|
||||
λ n₁ hlt, nat.lt.cases_on hlt
|
||||
(λ heq, nat.no_confusion heq)
|
||||
(λ b hlt heq, nat.no_confusion heq),
|
||||
(by contradiction)
|
||||
(by contradiction),
|
||||
aux hlt rfl))
|
||||
(λ (n : nat) (ih : acc lt n),
|
||||
acc.intro (succ n) (λ (m : nat) (hlt : m < succ n),
|
||||
|
@ -68,8 +68,8 @@ namespace nat
|
|||
definition not_lt_zero (a : nat) : ¬ a < zero :=
|
||||
have aux : ∀ {b}, a < b → b = zero → false, from
|
||||
λ b H, lt.cases_on H
|
||||
(λ heq, nat.no_confusion heq)
|
||||
(λ b h₁ heq, nat.no_confusion heq),
|
||||
(by contradiction)
|
||||
(by contradiction),
|
||||
λ H, aux H rfl
|
||||
|
||||
definition zero_lt_succ (a : nat) : zero < succ a :=
|
||||
|
|
|
@ -68,7 +68,7 @@ assert zero_eq_one : 0 = 1, from calc
|
|||
... = β (α m) : aux α (zω_eq_znkω m (M f + 1))
|
||||
... = β (M f + 1) : by rewrite znkω_bound
|
||||
... = 1 : by rewrite znkω_bound,
|
||||
nat.no_confusion zero_eq_one
|
||||
by contradiction
|
||||
end
|
||||
|
||||
/-
|
||||
|
|
Loading…
Reference in a new issue