]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/basics/deqsets.ma
Added a turing/universal directory for the universal turing machine (and
[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 unification hint  0 ≔ ; 
65     X ≟ mk_DeqSet bool beqb beqb_true
66 (* ---------------------------------------- *) ⊢ 
67     bool ≡ carr X.
68     
69 unification hint  0 ≔ b1,b2:bool; 
70     X ≟ mk_DeqSet bool beqb beqb_true
71 (* ---------------------------------------- *) ⊢ 
72     beqb b1 b2 ≡ eqb X b1 b2.
73     
74 example exhint: ∀b:bool. (b == false) = true → b = false. 
75 #b #H @(\P H).
76 qed.
77
78 (* pairs *)
79 definition eq_pairs ≝
80   λA,B:DeqSet.λp1,p2:A×B.(\fst p1 == \fst p2) ∧ (\snd p1 == \snd p2).
81
82 lemma eq_pairs_true: ∀A,B:DeqSet.∀p1,p2:A×B.
83   eq_pairs A B p1 p2 = true ↔ p1 = p2.
84 #A #B * #a1 #b1 * #a2 #b2 %
85   [#H cases (andb_true …H) #eqa #eqb >(\P eqa) >(\P eqb) //
86   |#H destruct normalize >(\b (refl … a2)) >(\b (refl … b2)) //
87   ]
88 qed.
89
90 definition DeqProd ≝ λA,B:DeqSet.
91   mk_DeqSet (A×B) (eq_pairs A B) (eq_pairs_true A B).
92   
93 unification hint  0 ≔ C1,C2; 
94     T1 ≟ carr C1,
95     T2 ≟ carr C2,
96     X ≟ DeqProd C1 C2
97 (* ---------------------------------------- *) ⊢ 
98     T1×T2 ≡ carr X.
99
100 unification hint  0 ≔ T1,T2,p1,p2; 
101     X ≟ DeqProd T1 T2
102 (* ---------------------------------------- *) ⊢ 
103     eq_pairs T1 T2 p1 p2 ≡ eqb X p1 p2.
104
105 example hint2: ∀b1,b2. 
106   〈b1,true〉==〈false,b2〉=true → 〈b1,true〉=〈false,b2〉.
107 #b1 #b2 #H @(\P H)
108 qed.
109
110 (* sum *)
111 definition eq_sum ≝
112   λA,B:DeqSet.λp1,p2:A+B.
113   match p1 with
114   [ inl a1 ⇒ match p2 with
115     [ inl a2 ⇒ a1 == a2 | inr b2 ⇒ false ]
116   | inr b1 ⇒ match p2 with 
117     [ inl a2 ⇒ false | inr b2 ⇒ b1 == b2 ]
118   ].
119
120 lemma eq_sum_true: ∀A,B:DeqSet.∀p1,p2:A+B.
121   eq_sum A B p1 p2 = true ↔ p1 = p2.
122 #A #B * 
123   [#a1 * 
124     [#a2 normalize % 
125       [#eqa >(\P eqa) // | #H destruct @(\b ?) //]
126     |#b2 normalize % #H destruct 
127     ]
128   |#b1 *
129     [#a2 normalize % #H destruct
130     |#b2 normalize %
131       [#eqb >(\P eqb) // | #H destruct @(\b ?) //]
132     ]
133   ]
134 qed.
135
136 definition DeqSum ≝ λA,B:DeqSet.
137   mk_DeqSet (A+B) (eq_sum A B) (eq_sum_true A B).
138   
139 unification hint  0 ≔ C1,C2; 
140     T1 ≟ carr C1,
141     T2 ≟ carr C2,
142     X ≟ DeqSum C1 C2
143 (* ---------------------------------------- *) ⊢ 
144     T1+T2 ≡ carr X.
145
146 unification hint  0 ≔ T1,T2,p1,p2; 
147     X ≟ DeqSum T1 T2
148 (* ---------------------------------------- *) ⊢ 
149     eq_sum T1 T2 p1 p2 ≡ eqb X p1 p2.
150