]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/basics/deqsets.ma
dc0044edc08efed553791e47ab5862e989a9350f
[helm.git] / matita / matita / lib / basics / deqsets.ma
1 (*
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.                     
5     ||I||                                                                 
6     ||T||  
7     ||A||  
8     \   /  This file is distributed under the terms of the       
9      \ /   GNU General Public License Version 2   
10       V_______________________________________________________________ *)
11
12 include "basics/types.ma".
13 include "basics/bool.ma".
14
15 (****** DeqSet: a set with a decidbale equality ******)
16
17 record DeqSet : Type[1] ≝ { carr :> Type[0];
18    eqb: carr → carr → bool;
19    eqb_true: ∀x,y. (eqb x y = true) ↔ (x = y)
20 }.
21
22 notation "a == b" non associative with precedence 45 for @{ 'eqb $a $b }.
23 interpretation "eqb" 'eqb a b = (eqb ? a b).
24
25 notation "\P H" non associative with precedence 90 
26   for @{(proj1 … (eqb_true ???) $H)}. 
27
28 notation "\b H" non associative with precedence 90 
29   for @{(proj2 … (eqb_true ???) $H)}. 
30   
31 notation < "𝐃" non associative with precedence 90 
32  for @{'bigD}.
33 interpretation "DeqSet" 'bigD = (mk_DeqSet ???).
34   
35 lemma eqb_false: ∀S:DeqSet.∀a,b:S. 
36   (eqb ? a b) = false ↔ a ≠ b.
37 #S #a #b % #H 
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)
40   ]
41 qed.
42  
43 notation "\Pf H" non associative with precedence 90 
44   for @{(proj1 … (eqb_false ???) $H)}. 
45
46 notation "\bf H" non associative with precedence 90 
47   for @{(proj2 … (eqb_false ???) $H)}. 
48   
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)]
52 qed.
53
54 definition beqb ≝ λb1,b2.
55   match b1 with [ true ⇒ b2 | false ⇒ notb b2].
56
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/
60 qed. 
61
62 definition DeqBool ≝ mk_DeqSet bool beqb beqb_true.
63
64 alias symbol "hint_decl" (instance 1) = "hint_decl_Type1".
65 unification hint  0 ≔ ; 
66     X ≟ mk_DeqSet bool beqb beqb_true
67 (* ---------------------------------------- *) ⊢ 
68     bool ≡ carr X.
69     
70 unification hint  0 ≔ b1,b2:bool; 
71     X ≟ mk_DeqSet bool beqb beqb_true
72 (* ---------------------------------------- *) ⊢ 
73     beqb b1 b2 ≡ eqb X b1 b2.
74     
75 example exhint: ∀b:bool. (b == false) = true → b = false. 
76 #b #H @(\P H).
77 qed.
78
79 (* option *)
80
81 definition eq_option ≝
82   λA:DeqSet.λa1,a2:option A.
83     match a1 with 
84     [ None ⇒ match a2 with [None ⇒ true | _ ⇒ false]
85     | Some a1' ⇒ match a2 with [None ⇒ false | Some a2' ⇒ a1'==a2']].
86
87 lemma eq_option_true: ∀A:DeqSet.∀a1,a2:option A.
88   eq_option A a1 a2 = true ↔ a1 = a2.
89 #A *
90   [* 
91     [% //
92     |#a1 % normalize #H destruct 
93     ]
94   |#a1 *  
95     [normalize % #H destruct
96     |#a2 normalize %
97       [#Heq >(\P Heq) //
98       |#H destruct @(\b ?) //
99       ]
100   ]
101 qed.
102
103 definition DeqOption ≝ λA:DeqSet.
104   mk_DeqSet (option A) (eq_option A) (eq_option_true A).
105   
106 unification hint  0 ≔ C; 
107     T ≟ carr C,
108     X ≟ DeqOption C
109 (* ---------------------------------------- *) ⊢ 
110     option T ≡ carr X.
111
112 unification hint  0 ≔ T,a1,a2; 
113     X ≟ DeqOption T
114 (* ---------------------------------------- *) ⊢ 
115     eq_option T a1 a2 ≡ eqb X a1 a2.
116
117 (* pairs *)
118 definition eq_pairs ≝
119   λA,B:DeqSet.λp1,p2:A×B.(\fst p1 == \fst p2) ∧ (\snd p1 == \snd p2).
120
121 lemma eq_pairs_true: ∀A,B:DeqSet.∀p1,p2:A×B.
122   eq_pairs A B p1 p2 = true ↔ p1 = p2.
123 #A #B * #a1 #b1 * #a2 #b2 %
124   [#H cases (andb_true …H) #eqa #eqb >(\P eqa) >(\P eqb) //
125   |#H destruct normalize >(\b (refl … a2)) >(\b (refl … b2)) //
126   ]
127 qed.
128
129 definition DeqProd ≝ λA,B:DeqSet.
130   mk_DeqSet (A×B) (eq_pairs A B) (eq_pairs_true A B).
131   
132 unification hint  0 ≔ C1,C2; 
133     T1 ≟ carr C1,
134     T2 ≟ carr C2,
135     X ≟ DeqProd C1 C2
136 (* ---------------------------------------- *) ⊢ 
137     T1×T2 ≡ carr X.
138
139 unification hint  0 ≔ T1,T2,p1,p2; 
140     X ≟ DeqProd T1 T2
141 (* ---------------------------------------- *) ⊢ 
142     eq_pairs T1 T2 p1 p2 ≡ eqb X p1 p2.
143
144 example hint2: ∀b1,b2. 
145   〈b1,true〉==〈false,b2〉=true → 〈b1,true〉=〈false,b2〉.
146 #b1 #b2 #H @(\P H)
147 qed.
148
149 (* sum *)
150 definition eq_sum ≝
151   λA,B:DeqSet.λp1,p2:A+B.
152   match p1 with
153   [ inl a1 ⇒ match p2 with
154     [ inl a2 ⇒ a1 == a2 | inr b2 ⇒ false ]
155   | inr b1 ⇒ match p2 with 
156     [ inl a2 ⇒ false | inr b2 ⇒ b1 == b2 ]
157   ].
158
159 lemma eq_sum_true: ∀A,B:DeqSet.∀p1,p2:A+B.
160   eq_sum A B p1 p2 = true ↔ p1 = p2.
161 #A #B * 
162   [#a1 * 
163     [#a2 normalize % 
164       [#eqa >(\P eqa) // | #H destruct @(\b ?) //]
165     |#b2 normalize % #H destruct 
166     ]
167   |#b1 *
168     [#a2 normalize % #H destruct
169     |#b2 normalize %
170       [#eqb >(\P eqb) // | #H destruct @(\b ?) //]
171     ]
172   ]
173 qed.
174
175 definition DeqSum ≝ λA,B:DeqSet.
176   mk_DeqSet (A+B) (eq_sum A B) (eq_sum_true A B).
177   
178 unification hint  0 ≔ C1,C2; 
179     T1 ≟ carr C1,
180     T2 ≟ carr C2,
181     X ≟ DeqSum C1 C2
182 (* ---------------------------------------- *) ⊢ 
183     T1+T2 ≡ carr X.
184
185 unification hint  0 ≔ T1,T2,p1,p2; 
186     X ≟ DeqSum T1 T2
187 (* ---------------------------------------- *) ⊢ 
188     eq_sum T1 T2 p1 p2 ≡ eqb X p1 p2.
189
190 (* sigma *)
191 definition eq_sigma ≝ 
192   λA:DeqSet.λP:A→Prop.λp1,p2:Σx:A.P x.
193   match p1 with 
194   [mk_Sig a1 h1 ⇒ 
195     match p2 with 
196     [mk_Sig a2 h2 ⇒ a1==a2]].
197  
198 (* uses proof irrelevance *)
199 lemma eq_sigma_true: ∀A:DeqSet.∀P.∀p1,p2:Σx.P x.
200   eq_sigma A P p1 p2 = true ↔ p1 = p2.
201 #A #P * #a1 #Ha1 * #a2 #Ha2 %
202   [normalize #eqa generalize in match Ha1; >(\P eqa) // 
203   |#H >H @(\b ?) //
204   ]
205 qed.
206
207 definition DeqSig ≝ λA:DeqSet.λP:A→Prop.
208   mk_DeqSet (Σx:A.P x) (eq_sigma A P) (eq_sigma_true A P).
209   
210 unification hint  0 ≔ C,P; 
211     T ≟ carr C,
212     X ≟ DeqSig C P
213 (* ---------------------------------------- *) ⊢ 
214     Σx:T.P x ≡ carr X.
215
216 unification hint  0 ≔ T,P,p1,p2; 
217     X ≟ DeqSig T P
218 (* ---------------------------------------- *) ⊢ 
219     eq_sigma T P p1 p2 ≡ eqb X p1 p2.
220