]> matita.cs.unibo.it Git - helm.git/blob - matita/library/decidable_kit/eqtype.ma
tagged 0.5.0-rc1
[helm.git] / matita / library / decidable_kit / eqtype.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/decidable_kit/eqtype/".
16
17 include "decidable_kit/decidable.ma".
18 include "datatypes/constructors.ma".
19
20 (* ### types with a decidable AND rewrite (leibniz) compatible ### *)
21
22 definition eq_compatible ≝ λT:Type.λa,b:T. reflect (a = b).
23     
24 record eqType : Type ≝ {
25   sort :> Type;
26   cmp : sort → sort → bool ;
27   eqP : ∀x,y:sort. eq_compatible sort x y (cmp x y)
28 }.
29
30 lemma cmp_refl : ∀d:eqType.∀x. cmp d x x = true.
31 intros (d x); generalize in match (eqP d x x); intros (H);
32 unfold in H; (* inversion non fa whd!!! *)
33 inversion H; [ intros; reflexivity | intros (ABS); cases (ABS (refl_eq ? ?)) ]
34 qed. 
35
36 (* to use streicher *)
37 lemma eqType_decidable : ∀d:eqType. decT d.
38 intros (d); unfold decT; intros (x y); unfold decidable;
39 generalize in match (eqP d x y); intros (H);
40 cases H; [left | right] assumption;
41 qed.
42
43 lemma eqbP : ∀x,y:nat. eq_compatible nat x y (eqb x y).
44 intros (x y);
45 generalize in match (refl_eq ? (eqb x y));
46 generalize in match (eqb x y) in ⊢ (? ? ? % → %); intros 1 (b);
47 cases b; intros (H); [ apply reflect_true | apply reflect_false ];
48 [ apply (eqb_true_to_eq ? ? H) | apply (eqb_false_to_not_eq ? ? H) ]
49 qed.
50
51 definition nat_eqType : eqType ≝ mk_eqType ? ? eqbP.
52 (* XXX le coercion nel cast non vengono inserite *)
53 definition nat_canonical_eqType : nat → nat_eqType := 
54   λn : nat.(n : sort nat_eqType).
55 coercion cic:/matita/decidable_kit/eqtype/nat_canonical_eqType.con.
56
57 definition bcmp ≝ λx,y:bool. match x with [ true ⇒ y | false => notb y ].
58
59 lemma bcmpP : ∀x,y:bool. eq_compatible bool x y (bcmp x y).
60 intros (x y);
61 generalize in match (refl_eq ? (bcmp x y));
62 generalize in match (bcmp x y) in ⊢ (? ? ? % → %); intros 1 (b);
63 cases b; intros (H); [ apply reflect_true | apply reflect_false ];
64 generalize in match H; clear H;
65 cases x; cases y; simplify; intros (H); [1,4: reflexivity];
66 (* non va: try destruct H; *)
67 [1,2,3,6: destruct H | *: unfold Not; intros (H1); destruct H1]
68 qed.
69
70 definition bool_eqType : eqType ≝  mk_eqType ? ? bcmpP.
71 definition bool_canonical_eqType : bool → bool_eqType := 
72   λb : bool.(b : sort bool_eqType).
73 coercion cic:/matita/decidable_kit/eqtype/bool_canonical_eqType.con.
74
75 (* ### subtype of an eqType ### *)
76
77 record sigma (d : eqType) (p : d -> bool) : Type ≝ {
78   sval : d;
79   sprop : p sval = true
80 }.
81
82 notation "{x, h}"
83   right associative with precedence 80
84 for @{'sig $x $h}.
85
86 interpretation "sub_eqType" 'sig x h = 
87   (cic:/matita/decidable_kit/eqtype/sigma.ind#xpointer(1/1/1) _ _ x h).
88
89 (* restricting an eqType gives an eqType *)
90 lemma sigma_eq_dec : ∀d:eqType.∀p.∀x,y:sigma d p. 
91   eq_compatible ? x y (cmp ? (sval ? ? x) (sval ? ?y)).
92 intros (d p x y);
93 cases (eqP d (sval ? ? x) (sval ? ? y)); generalize in match H; clear H; 
94 cases x (s ps); cases y (t pt); simplify; intros (Est); 
95 [1: constructor 1; generalize in match ps; rewrite > Est; intros (pt');
96     rewrite < (pirrel ? ? ? pt pt' (eqType_decidable bool_eqType)); reflexivity;
97 |2: constructor 2; unfold Not; intros (H); destruct H;
98     cases (Est); reflexivity]  
99 qed. 
100
101 definition sub_eqType ≝ λd : eqType.λp. mk_eqType ? ? (sigma_eq_dec d p).
102
103 inductive in_sub (d : eqType) (p : d → bool) (x : d) : option (sigma d p) → Type ≝
104 | in_sig : ∀s : sigma d p. sval ? ? s = x → in_sub d p x (Some ? s)
105 | out_sig : p x = false → in_sub d p x (None ?).
106
107 definition if_p ≝ λd:eqType.λp:d→bool.λx:d. 
108   match (eqP bool_eqType (p x) true) with 
109   [ (reflect_true H) ⇒ Some ? {?,H} 
110   | (reflect_false _) ⇒ None ?].
111
112 lemma in_sub_eq : ∀d,p,x. in_sub d p x (if_p d p x).
113 intros (d p x); unfold if_p;
114 cases (eqP bool_eqType (p x) true); simplify;
115 [ apply in_sig; reflexivity; | apply out_sig; apply (p2bF ? ? (idP ?) H)]
116 qed.  
117
118 definition ocmp : ∀d:eqType.∀a,b:option d.bool ≝
119   λd:eqType.λa,b:option d.
120   match a with 
121   [ None ⇒ match b with [ None ⇒ true | (Some _) ⇒ false] 
122   | (Some x) ⇒ match b with [ None ⇒ false | (Some y) ⇒ cmp d x y]].  
123   
124 lemma ocmpP : ∀d:eqType.∀a,b:option d.eq_compatible ? a b (ocmp d a b).
125 intros (d a b);
126 generalize in match (refl_eq ? (ocmp ? a b));
127 generalize in match (ocmp ? a b) in ⊢ (? ? ? % → %); intros 1 (c);
128 cases c; intros (H); [ apply reflect_true | apply reflect_false ];
129 generalize in match H; clear H;
130 cases a; cases b; simplify; intros (H);
131 [1: reflexivity;
132 |2,3,5: destruct H;
133 |4: rewrite > (b2pT ? ? (eqP d ? ?) H); reflexivity; 
134 |6,7: unfold Not; intros (H1); destruct H1
135 |8: unfold Not; intros (H1); destruct H1;rewrite > cmp_refl in H; destruct H;]
136 qed.
137
138 definition option_eqType : eqType → eqType ≝ λd:eqType.mk_eqType ? ? (ocmpP d).
139 definition option_canonical_eqType : ∀d:eqType.d → option_eqType d ≝
140   λd:eqType.λx:d.(Some ? x : sort (option_eqType d)).
141 coercion cic:/matita/decidable_kit/eqtype/option_canonical_eqType.con.
142
143 (* belle le coercions! *)
144 definition test_canonical_option_eqType ≝ 
145   (eq (option_eqType nat_eqType) O (S O)).
146
147 lemma cmpP : ∀d:eqType.∀x,y:d.∀P:bool → Prop. 
148   (x=y → P true) → (cmp d x y = false → P false) → P (cmp d x y).
149 intros; cases (eqP ? x y); [2:apply H1; apply (p2bF ? ? (eqP d ? ?))] autobatch;
150 qed.
151
152 lemma cmpC : ∀d:eqType.∀x,y:d. cmp d x y = cmp d y x.
153 intros; apply (cmpP ? x y); apply (cmpP ? y x); [1,4: intros; reflexivity]
154 [intros (H2 H1) | intros (H1 H2)] rewrite > H1 in H2; rewrite > cmp_refl in H2; 
155 destruct H2.      
156 qed.