refactor(library): simplify theorems using improved tactics
This commit is contained in:
parent
f13ca3cd9a
commit
7e875c8d85
9 changed files with 19 additions and 19 deletions
|
@ -418,10 +418,10 @@ definition has_decidable_eq {A : Type} [H : decidable_eq A] : ∀ l₁ l₂ : li
|
||||||
match H a b with
|
match H a b with
|
||||||
| inl Hab :=
|
| inl Hab :=
|
||||||
match has_decidable_eq l₁ l₂ with
|
match has_decidable_eq l₁ l₂ with
|
||||||
| inl He := inl (eq.rec_on Hab (eq.rec_on He rfl))
|
| inl He := inl (by congruence; repeat assumption)
|
||||||
| inr Hn := inr (λ H, list.no_confusion H (λ Hab Ht, absurd Ht Hn))
|
| inr Hn := inr (by intro H; injection H; contradiction)
|
||||||
end
|
end
|
||||||
| inr Hnab := inr (λ H, list.no_confusion H (λ Hab Ht, absurd Hab Hnab))
|
| inr Hnab := inr (by intro H; injection H; contradiction)
|
||||||
end
|
end
|
||||||
|
|
||||||
/- quasiequal a l l' means that l' is exactly l, with a added
|
/- quasiequal a l l' means that l' is exactly l, with a added
|
||||||
|
|
|
@ -64,16 +64,16 @@ namespace pos_num
|
||||||
| (bit0 a) one := inr (by contradiction)
|
| (bit0 a) one := inr (by contradiction)
|
||||||
| (bit0 a) (bit0 b) :=
|
| (bit0 a) (bit0 b) :=
|
||||||
match decidable_eq a b with
|
match decidable_eq a b with
|
||||||
| inl H₁ := inl (eq.rec_on H₁ rfl)
|
| inl H₁ := inl (by rewrite H₁)
|
||||||
| inr H₁ := inr (λ H, pos_num.no_confusion H (λ H₂, absurd H₂ H₁))
|
| inr H₁ := inr (by intro H; injection H; contradiction)
|
||||||
end
|
end
|
||||||
| (bit0 a) (bit1 b) := inr (by contradiction)
|
| (bit0 a) (bit1 b) := inr (by contradiction)
|
||||||
| (bit1 a) one := inr (by contradiction)
|
| (bit1 a) one := inr (by contradiction)
|
||||||
| (bit1 a) (bit0 b) := inr (by contradiction)
|
| (bit1 a) (bit0 b) := inr (by contradiction)
|
||||||
| (bit1 a) (bit1 b) :=
|
| (bit1 a) (bit1 b) :=
|
||||||
match decidable_eq a b with
|
match decidable_eq a b with
|
||||||
| inl H₁ := inl (eq.rec_on H₁ rfl)
|
| inl H₁ := inl (by rewrite H₁)
|
||||||
| inr H₁ := inr (λ H, pos_num.no_confusion H (λ H₂, absurd H₂ H₁))
|
| inr H₁ := inr (by intro H; injection H; contradiction)
|
||||||
end
|
end
|
||||||
|
|
||||||
local notation a < b := (lt a b = tt)
|
local notation a < b := (lt a b = tt)
|
||||||
|
|
|
@ -34,6 +34,6 @@ namespace option
|
||||||
| (some v₁) (some v₂) :=
|
| (some v₁) (some v₂) :=
|
||||||
match H v₁ v₂ with
|
match H v₁ v₂ with
|
||||||
| inl e := by left; congruence; assumption
|
| inl e := by left; congruence; assumption
|
||||||
| inr n := by right; intro h; injection h; refine absurd _ n; assumption
|
| inr n := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
end option
|
end option
|
||||||
|
|
|
@ -24,9 +24,9 @@ namespace prod
|
||||||
| inl e₁ :=
|
| inl e₁ :=
|
||||||
match h₂ b b' with
|
match h₂ b b' with
|
||||||
| inl e₂ := by left; congruence; repeat assumption
|
| inl e₂ := by left; congruence; repeat assumption
|
||||||
| inr n₂ := by right; intro h; injection h; refine absurd _ n₂; assumption
|
| inr n₂ := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
| inr n₁ := by right; intro h; injection h; refine absurd _ n₁; assumption
|
| inr n₁ := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
|
|
||||||
definition swap {A : Type} : A × A → A × A
|
definition swap {A : Type} : A × A → A × A
|
||||||
|
|
|
@ -28,7 +28,7 @@ begin
|
||||||
apply (@by_cases (a₇ = b₇)), intros,
|
apply (@by_cases (a₇ = b₇)), intros,
|
||||||
apply (@by_cases (a₈ = b₈)), intros,
|
apply (@by_cases (a₈ = b₈)), intros,
|
||||||
left, congruence, repeat assumption,
|
left, congruence, repeat assumption,
|
||||||
repeat (intro n; right; intro h; injection h; refine absurd _ n; assumption)
|
repeat (intro n; right; intro h; injection h; contradiction)
|
||||||
end
|
end
|
||||||
|
|
||||||
open string
|
open string
|
||||||
|
@ -42,7 +42,7 @@ definition decidable_eq_string [instance] : ∀ s₁ s₂ : string, decidable (s
|
||||||
| inl e₁ :=
|
| inl e₁ :=
|
||||||
match decidable_eq_string r₁ r₂ with
|
match decidable_eq_string r₁ r₂ with
|
||||||
| inl e₂ := by left; congruence; repeat assumption
|
| inl e₂ := by left; congruence; repeat assumption
|
||||||
| inr n₂ := by right; intro h; injection h; refine absurd _ n₂; assumption
|
| inr n₂ := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
| inr n₁ := by right; intro h; injection h; refine absurd _ n₁; assumption
|
| inr n₁ := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
|
|
|
@ -33,6 +33,6 @@ namespace subtype
|
||||||
begin
|
begin
|
||||||
apply (@by_cases (v₁ = v₂)),
|
apply (@by_cases (v₁ = v₂)),
|
||||||
{intro e, revert p₁, rewrite e, intro p₁, left, congruence},
|
{intro e, revert p₁, rewrite e, intro p₁, left, congruence},
|
||||||
{intro n, right, intro h, injection h, refine absurd _ n, assumption}
|
{intro n, right, intro h, injection h, contradiction}
|
||||||
end
|
end
|
||||||
end subtype
|
end subtype
|
||||||
|
|
|
@ -41,13 +41,13 @@ namespace sum
|
||||||
| has_decidable_eq (inl a₁) (inl a₂) :=
|
| has_decidable_eq (inl a₁) (inl a₂) :=
|
||||||
match h₁ a₁ a₂ with
|
match h₁ a₁ a₂ with
|
||||||
| decidable.inl hp := by left; congruence; assumption
|
| decidable.inl hp := by left; congruence; assumption
|
||||||
| decidable.inr hn := by right; intro h; injection h; refine absurd _ hn; assumption
|
| decidable.inr hn := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
| has_decidable_eq (inl a₁) (inr b₂) := by right; contradiction
|
| has_decidable_eq (inl a₁) (inr b₂) := by right; contradiction
|
||||||
| has_decidable_eq (inr b₁) (inl a₂) := by right; contradiction
|
| has_decidable_eq (inr b₁) (inl a₂) := by right; contradiction
|
||||||
| has_decidable_eq (inr b₁) (inr b₂) :=
|
| has_decidable_eq (inr b₁) (inr b₂) :=
|
||||||
match h₂ b₁ b₂ with
|
match h₂ b₁ b₂ with
|
||||||
| decidable.inl hp := by left; congruence; assumption
|
| decidable.inl hp := by left; congruence; assumption
|
||||||
| decidable.inr hn := by right; intro h; injection h; refine absurd _ hn; assumption
|
| decidable.inr hn := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
end sum
|
end sum
|
||||||
|
|
|
@ -263,8 +263,8 @@ namespace vector
|
||||||
| inl Hab :=
|
| inl Hab :=
|
||||||
match decidable_eq v₁ v₂ with
|
match decidable_eq v₁ v₂ with
|
||||||
| inl He := by left; congruence; repeat assumption
|
| inl He := by left; congruence; repeat assumption
|
||||||
| inr Hn := by right; intro h; injection h; refine absurd _ Hn; assumption
|
| inr Hn := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
| inr Hnab := by right; intro h; injection h; refine absurd _ Hnab; assumption
|
| inr Hnab := by right; intro h; injection h; contradiction
|
||||||
end
|
end
|
||||||
end vector
|
end vector
|
||||||
|
|
|
@ -35,7 +35,7 @@ namespace nat
|
||||||
| has_decidable_eq (succ x) (succ y) :=
|
| has_decidable_eq (succ x) (succ y) :=
|
||||||
match has_decidable_eq x y with
|
match has_decidable_eq x y with
|
||||||
| inl xeqy := inl (by rewrite xeqy)
|
| inl xeqy := inl (by rewrite xeqy)
|
||||||
| inr xney := inr (λ h : succ x = succ y, by injection h with xeqy; exact absurd xeqy xney)
|
| inr xney := inr (by intro h; injection h; contradiction)
|
||||||
end
|
end
|
||||||
|
|
||||||
-- less-than is well-founded
|
-- less-than is well-founded
|
||||||
|
|
Loading…
Reference in a new issue