refactor(library): replace decidable_eq with abbreviation

This commit is contained in:
Leonardo de Moura 2014-09-09 16:07:07 -07:00
parent c92a9b8808
commit 9b9adf8831
13 changed files with 22 additions and 25 deletions

View file

@ -139,8 +139,8 @@ namespace bool
inhabited.mk ff
theorem has_decidable_eq [protected] [instance] : decidable_eq bool :=
decidable_eq.intro (λ (a b : bool),
take a b : bool,
rec_on a
(rec_on b (inl rfl) (inr ff_ne_tt))
(rec_on b (inr (ne.symm ff_ne_tt)) (inl rfl)))
(rec_on b (inr (ne.symm ff_ne_tt)) (inl rfl))
end bool

View file

@ -206,7 +206,7 @@ exists_intro (pr1 (rep a))
definition of_nat (n : ) : := psub (pair n 0)
theorem has_decidable_eq [instance] [protected] : decidable_eq :=
decidable_eq.intro (λ (a b : ), _)
take a b : , _
opaque_hint (hiding int)
coercion of_nat

View file

@ -106,7 +106,7 @@ induction_on n
(take k IH H, IH (succ_inj H))
theorem has_decidable_eq [instance] [protected] : decidable_eq :=
decidable_eq.intro (λ (n m : ),
take n m : ,
have general : ∀n, decidable (n = m), from
rec_on m
(take n,
@ -123,7 +123,7 @@ have general : ∀n, decidable (n = m), from
have H1 : succ n' ≠ succ m', from
assume Heq, absurd (succ_inj Heq) Hne,
inr H1))),
general n)
general n
theorem two_step_induction_on {P : → Prop} (a : ) (H1 : P 0) (H2 : P 1)
(H3 : ∀ (n : ) (IH1 : P n) (IH2 : P (succ n)), P (succ (succ n))) : P a :=

View file

@ -39,12 +39,12 @@ namespace option
inhabited.mk none
theorem has_decidable_eq [protected] [instance] {A : Type} (H : decidable_eq A) : decidable_eq (option A) :=
decidable_eq.intro (λ (o₁ o₂ : option A),
take o₁ o₂ : option A,
rec_on o₁
(rec_on o₂ (inl rfl) (take a₂, (inr (none_ne_some a₂))))
(take a₁ : A, rec_on o₂
(inr (ne.symm (none_ne_some a₁)))
(take a₂ : A, decidable.rec_on (H a₁ a₂)
(assume Heq : a₁ = a₂, inl (Heq ▸ rfl))
(assume Hne : a₁ ≠ a₂, inr (assume Hn : some a₁ = some a₂, absurd (equal Hn) Hne)))))
(assume Hne : a₁ ≠ a₂, inr (assume Hn : some a₁ = some a₂, absurd (equal Hn) Hne))))
end option

View file

@ -49,11 +49,11 @@ section
inhabited.destruct H1 (λa, inhabited.destruct H2 (λb, inhabited.mk (pair a b)))
theorem has_decidable_eq [protected] [instance] (H1 : decidable_eq A) (H2 : decidable_eq B) : decidable_eq (A × B) :=
decidable_eq.intro (λ (u v : A × B),
take u v : A × B,
have H3 : u = v ↔ (pr1 u = pr1 v) ∧ (pr2 u = pr2 v), from
iff.intro
(assume H, H ▸ and.intro rfl rfl)
(assume H, and.elim H (assume H4 H5, equal H4 H5)),
decidable_iff_equiv _ (iff.symm H3))
decidable_iff_equiv _ (iff.symm H3)
end
end prod

View file

@ -42,9 +42,9 @@ section
inhabited.mk (tag a H)
theorem has_decidable_eq [protected] [instance] (H : decidable_eq A) : decidable_eq {x, P x} :=
decidable_eq.intro (λ (a1 a2 : {x, P x}),
take a1 a2 : {x, P x},
have H1 : (a1 = a2) ↔ (elt_of a1 = elt_of a2), from
iff.intro (assume H, eq.subst H rfl) (assume H, equal H),
decidable_iff_equiv _ (iff.symm H1))
decidable_iff_equiv _ (iff.symm H1)
end
end subtype

View file

@ -57,7 +57,7 @@ namespace sum
theorem has_eq_decidable [protected] [instance] {A B : Type} (H1 : decidable_eq A) (H2 : decidable_eq B) :
decidable_eq (A ⊎ B) :=
decidable_eq.intro (λ (s1 s2 : A ⊎ B),
take s1 s2 : A ⊎ B,
rec_on s1
(take a1, show decidable (inl B a1 = s2), from
rec_on s2
@ -78,5 +78,5 @@ namespace sum
(take b2, show decidable (inr A b1 = inr A b2), from
decidable.rec_on (H2 b1 b2)
(assume Heq : b1 = b2, decidable.inl (Heq ▸ rfl))
(assume Hne : b1 ≠ b2, decidable.inr (mt inr_inj Hne)))))
(assume Hne : b1 ≠ b2, decidable.inr (mt inr_inj Hne))))
end sum

View file

@ -19,5 +19,5 @@ namespace unit
inhabited.mk ⋆
theorem has_decidable_eq [protected] [instance] : decidable_eq unit :=
decidable_eq.intro (λ (a b : unit), inl (equal a b))
take (a b : unit), inl (equal a b)
end unit

View file

@ -96,8 +96,4 @@ decidable_iff_equiv Ha (eq_to_iff H)
end decidable
inductive decidable_eq [class] (A : Type) : Type :=
intro : (Π x y : A, decidable (x = y)) → decidable_eq A
theorem of_decidable_eq [instance] [coercion] {A : Type} (H : decidable_eq A) (x y : A) : decidable (x = y) :=
decidable_eq.rec (λ H, H) H x y
abbreviation decidable_eq (A : Type) := Π (a b : A), decidable (a = b)

View file

@ -1,3 +1,3 @@
VISIT coe.lean
WAIT
INFO 5 33
INFO 5 16

View file

@ -1,13 +1,13 @@
-- BEGINWAIT
-- ENDWAIT
-- BEGININFO
-- TYPE|5|33
-- TYPE|5|16
decidable (eq a b)
-- ACK
-- SYNTH|5|33
-- SYNTH|5|16
int.has_decidable_eq a b
-- ACK
-- SYMBOL|5|33
-- SYMBOL|5|16
_
-- ACK
-- ENDINFO

View file

@ -2,7 +2,7 @@ import data.int
open int
theorem has_decidable_eq [instance] [protected] : decidable_eq :=
decidable_eq.intro (λ (a b : ), _)
take (a b : ), _
variable n : nat
variable i : int

View file

@ -9,5 +9,6 @@ section
theorem tst1 (H : Πx y, decidable (R x y)) (a b c : A) : decidable (R a b ∧ R b a)
theorem tst2 (H : decidable_bin_rel R) (a b c : A) : decidable (R a b ∧ R b a)
theorem tst2 (H : decidable_bin_rel R) (a b c : A) : decidable (R a b ∧ R b a R b b) :=
_
end