]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/decidable_kit/eqtype.ma
demodulate takes an extra argument 'all', if present it attempts to solve
[helm.git] / helm / software / matita / library / decidable_kit / eqtype.ma
index 35f404df35aee4414ec589e47801c4cecbeba5a7..4c84e85ca1be5a2b1df1dc5fde5d74c7aa81de7d 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. 
@@ -138,7 +135,7 @@ 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 ≝