2 ||M|| This file is part of HELM, an Hypertextual, Electronic
3 ||A|| Library of Mathematics, developed at the Computer Science
4 ||T|| Department of the University of Bologna, Italy.
8 \ / This file is distributed under the terms of the
9 \ / GNU General Public License Version 2
10 V_______________________________________________________________ *)
12 include "basics/types.ma".
13 include "basics/bool.ma".
15 (****** DeqSet: a set with a decidbale equality ******)
17 record DeqSet : Type[1] ≝ { carr :> Type[0];
18 eqb: carr → carr → bool;
19 eqb_true: ∀x,y. (eqb x y = true) ↔ (x = y)
22 notation "a == b" non associative with precedence 45 for @{ 'eqb $a $b }.
23 interpretation "eqb" 'eqb a b = (eqb ? a b).
25 notation "\P H" non associative with precedence 90
26 for @{(proj1 … (eqb_true ???) $H)}.
28 notation "\b H" non associative with precedence 90
29 for @{(proj2 … (eqb_true ???) $H)}.
31 notation < "𝐃" non associative with precedence 90
33 interpretation "DeqSet" 'bigD = (mk_DeqSet ???).
35 lemma eqb_false: ∀S:DeqSet.∀a,b:S.
36 (eqb ? a b) = false ↔ a ≠ b.
38 [@(not_to_not … not_eq_true_false) #H1 <H @sym_eq @(\b H1)
39 |cases (true_or_false (eqb ? a b)) // #H1 @False_ind @(absurd … (\P H1) H)
43 notation "\Pf H" non associative with precedence 90
44 for @{(proj1 … (eqb_false ???) $H)}.
46 notation "\bf H" non associative with precedence 90
47 for @{(proj2 … (eqb_false ???) $H)}.
49 lemma dec_eq: ∀S:DeqSet.∀a,b:S. a = b ∨ a ≠ b.
50 #S #a #b cases (true_or_false (eqb ? a b)) #H
51 [%1 @(\P H) | %2 @(\Pf H)]
54 definition beqb ≝ λb1,b2.
55 match b1 with [ true ⇒ b2 | false ⇒ notb b2].
57 notation < "a == b" non associative with precedence 45 for @{beqb $a $b }.
58 lemma beqb_true: ∀b1,b2. iff (beqb b1 b2 = true) (b1 = b2).
59 #b1 #b2 cases b1 cases b2 normalize /2/
62 definition DeqBool ≝ mk_DeqSet bool beqb beqb_true.
64 alias symbol "hint_decl" (instance 1) = "hint_decl_Type1".
65 unification hint 0 ≔ ;
66 X ≟ mk_DeqSet bool beqb beqb_true
67 (* ---------------------------------------- *) ⊢
70 unification hint 0 ≔ b1,b2:bool;
71 X ≟ mk_DeqSet bool beqb beqb_true
72 (* ---------------------------------------- *) ⊢
73 beqb b1 b2 ≡ eqb X b1 b2.
75 example exhint: ∀b:bool. (b == false) = true → b = false.
81 definition eq_option ≝
82 λA:DeqSet.λa1,a2:option A.
84 [ None ⇒ match a2 with [None ⇒ true | _ ⇒ false]
85 | Some a1' ⇒ match a2 with [None ⇒ false | Some a2' ⇒ a1'==a2']].
87 lemma eq_option_true: ∀A:DeqSet.∀a1,a2:option A.
88 eq_option A a1 a2 = true ↔ a1 = a2.
92 |#a1 % normalize #H destruct
95 [normalize % #H destruct
98 |#H destruct @(\b ?) //
103 definition DeqOption ≝ λA:DeqSet.
104 mk_DeqSet (option A) (eq_option A) (eq_option_true A).
106 unification hint 0 ≔ C;
109 (* ---------------------------------------- *) ⊢
112 unification hint 0 ≔ T,a1,a2;
114 (* ---------------------------------------- *) ⊢
115 eq_option T a1 a2 ≡ eqb X a1 a2.
119 definition eq_pairs ≝
120 λA,B:DeqSet.λp1,p2:A×B.(\fst p1 == \fst p2) ∧ (\snd p1 == \snd p2).
122 lemma eq_pairs_true: ∀A,B:DeqSet.∀p1,p2:A×B.
123 eq_pairs A B p1 p2 = true ↔ p1 = p2.
124 #A #B * #a1 #b1 * #a2 #b2 %
125 [#H cases (andb_true …H) #eqa #eqb >(\P eqa) >(\P eqb) //
126 |#H destruct normalize >(\b (refl … a2)) >(\b (refl … b2)) //
130 definition DeqProd ≝ λA,B:DeqSet.
131 mk_DeqSet (A×B) (eq_pairs A B) (eq_pairs_true A B).
133 unification hint 0 ≔ C1,C2;
137 (* ---------------------------------------- *) ⊢
140 unification hint 0 ≔ T1,T2,p1,p2;
142 (* ---------------------------------------- *) ⊢
143 eq_pairs T1 T2 p1 p2 ≡ eqb X p1 p2.
145 example hint2: ∀b1,b2.
146 〈b1,true〉==〈false,b2〉=true → 〈b1,true〉=〈false,b2〉.
152 λA,B:DeqSet.λp1,p2:A+B.
154 [ inl a1 ⇒ match p2 with
155 [ inl a2 ⇒ a1 == a2 | inr b2 ⇒ false ]
156 | inr b1 ⇒ match p2 with
157 [ inl a2 ⇒ false | inr b2 ⇒ b1 == b2 ]
160 lemma eq_sum_true: ∀A,B:DeqSet.∀p1,p2:A+B.
161 eq_sum A B p1 p2 = true ↔ p1 = p2.
165 [#eqa >(\P eqa) // | #H destruct @(\b ?) //]
166 |#b2 normalize % #H destruct
169 [#a2 normalize % #H destruct
171 [#eqb >(\P eqb) // | #H destruct @(\b ?) //]
176 definition DeqSum ≝ λA,B:DeqSet.
177 mk_DeqSet (A+B) (eq_sum A B) (eq_sum_true A B).
179 unification hint 0 ≔ C1,C2;
183 (* ---------------------------------------- *) ⊢
186 unification hint 0 ≔ T1,T2,p1,p2;
188 (* ---------------------------------------- *) ⊢
189 eq_sum T1 T2 p1 p2 ≡ eqb X p1 p2.
192 definition eq_sigma ≝
193 λA:DeqSet.λP:A→Prop.λp1,p2:Σx:A.P x.
197 [mk_Sig a2 h2 ⇒ a1==a2]].
199 (* uses proof irrelevance *)
200 lemma eq_sigma_true: ∀A:DeqSet.∀P.∀p1,p2:Σx.P x.
201 eq_sigma A P p1 p2 = true ↔ p1 = p2.
202 #A #P * #a1 #Ha1 * #a2 #Ha2 %
203 [normalize #eqa generalize in match Ha1; >(\P eqa) //
208 definition DeqSig ≝ λA:DeqSet.λP:A→Prop.
209 mk_DeqSet (Σx:A.P x) (eq_sigma A P) (eq_sigma_true A P).
212 unification hint 0 ≔ C,P;
215 (* ---------------------------------------- *) ⊢
218 unification hint 0 ≔ T,P,p1,p2;
220 (* ---------------------------------------- *) ⊢
221 eq_sigma T P p1 p2 ≡ eqb X p1 p2.