]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/decidable_kit/eqtype.ma
nasty change in the lexer/parser:
[helm.git] / helm / software / matita / library / decidable_kit / eqtype.ma
index 88721632f04270377d0de4dfbdb4d60ca585f2b1..6aa3ff59fd2bf57600a558665606f7aa20e8b626 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/decidable_kit/eqtype/".
-
 include "decidable_kit/decidable.ma".
 include "datatypes/constructors.ma".
 
@@ -52,7 +50,7 @@ definition nat_eqType : eqType ≝ mk_eqType ? ? eqbP.
 (* XXX le coercion nel cast non vengono inserite *)
 definition nat_canonical_eqType : nat → nat_eqType := 
   λn : nat.(n : sort nat_eqType).
-coercion cic:/matita/decidable_kit/eqtype/nat_canonical_eqType.con.
+coercion nat_canonical_eqType.
 
 definition bcmp ≝ λx,y:bool. match x with [ true ⇒ y | false => notb y ].
 
@@ -70,7 +68,7 @@ qed.
 definition bool_eqType : eqType ≝  mk_eqType ? ? bcmpP.
 definition bool_canonical_eqType : bool → bool_eqType := 
   λb : bool.(b : sort bool_eqType).
-coercion cic:/matita/decidable_kit/eqtype/bool_canonical_eqType.con.
+coercion bool_canonical_eqType.
 
 (* ### subtype of an eqType ### *)
 
@@ -83,8 +81,7 @@ notation "{x, h}"
   right associative with precedence 80
 for @{'sig $x $h}.
 
-interpretation "sub_eqType" 'sig x h = 
-  (cic:/matita/decidable_kit/eqtype/sigma.ind#xpointer(1/1/1) _ _ x h).
+interpretation "sub_eqType" 'sig x h =  (mk_sigma ? ? x h).
 
 (* restricting an eqType gives an eqType *)
 lemma sigma_eq_dec : ∀d:eqType.∀p.∀x,y:sigma d p. 
@@ -95,7 +92,7 @@ cases x (s ps); cases y (t pt); simplify; intros (Est);
 [1: constructor 1; generalize in match ps; rewrite > Est; intros (pt');
     rewrite < (pirrel ? ? ? pt pt' (eqType_decidable bool_eqType)); reflexivity;
 |2: constructor 2; unfold Not; intros (H); destruct H;
-    cases (Est); assumption;]  
+    cases (Est); reflexivity]  
 qed. 
 
 definition sub_eqType ≝ λd : eqType.λp. mk_eqType ? ? (sigma_eq_dec d p).
@@ -132,14 +129,13 @@ cases a; cases b; simplify; intros (H);
 |2,3,5: destruct H;
 |4: rewrite > (b2pT ? ? (eqP d ? ?) H); reflexivity; 
 |6,7: unfold Not; intros (H1); destruct H1
-|8: unfold Not; intros (H1); destruct H1; 
-    rewrite > Hcut in H;  rewrite > cmp_refl in H; destruct H;]
+|8: unfold Not; intros (H1); destruct H1;rewrite > cmp_refl in H; destruct H;]
 qed.
 
 definition option_eqType : eqType → eqType ≝ λd:eqType.mk_eqType ? ? (ocmpP d).
 definition option_canonical_eqType : ∀d:eqType.d → option_eqType d ≝
   λd:eqType.λx:d.(Some ? x : sort (option_eqType d)).
-coercion cic:/matita/decidable_kit/eqtype/option_canonical_eqType.con.
+coercion option_canonical_eqType.
 
 (* belle le coercions! *)
 definition test_canonical_option_eqType ≝