X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FGeneral.v;h=db77f344df6bf4e9361bf583700b048e2035fc56;hp=cf9d28dacbbb393974eec1b8e437a2428120d51d;hb=327fd294a7a4d3c02b5491708aadcd9acef9e376;hpb=164cdbf41ca206079b0dcfc18cd13625b286c38c diff --git a/src/General.v b/src/General.v index cf9d28d..db77f34 100644 --- a/src/General.v +++ b/src/General.v @@ -20,6 +20,24 @@ Class EqDecidable (T:Type) := }. Coercion eqd_type : EqDecidable >-> Sortclass. +Instance EqDecidableOption (T:Type)(EQDT:EqDecidable T) : EqDecidable ??T. + apply Build_EqDecidable. + intros. + destruct v1; + destruct v2. + destruct (eqd_dec t t0). + subst. + left; auto. + right. + unfold not; intros. + inversion H. + subst. + apply n. + auto. + right; unfold not; intro; inversion H. + right; unfold not; intro; inversion H. + left; auto. + Defined. Class ToString (T:Type) := { toString : T -> string }. Instance StringToString : ToString string := { toString := fun x => x }. @@ -500,6 +518,47 @@ Definition liftBoolFunc {T}(b:bool)(f:T -> bool) : ??T -> bool := | Some x => f x end. +(* decidable quality on a tree of elements which have decidable equality *) +Definition tree_eq_dec : forall {T:Type}(l1 l2:Tree T)(dec:forall t1 t2:T, sumbool (eq t1 t2) (not (eq t1 t2))), + sumbool (eq l1 l2) (not (eq l1 l2)). + intro T. + intro l1. + induction l1; intros. + destruct l2. + destruct (dec a t). + subst. + left; auto. + right; unfold not; intro; apply n; inversion H; auto. + right. + unfold not; intro. + inversion H. + + destruct l2. + right; unfold not; intro; inversion H. + destruct (IHl1_1 l2_1 dec); + destruct (IHl1_2 l2_2 dec); subst. + left; auto. + right. + unfold not; intro. + inversion H; subst. + apply n; auto. + right. + unfold not; intro. + inversion H; subst. + apply n; auto. + right. + unfold not; intro. + inversion H; subst. + apply n; auto. + Defined. + +Instance EqDecidableTree {T:Type}(eqd:EqDecidable T) : EqDecidable (Tree T). + apply Build_EqDecidable. + intros. + apply tree_eq_dec. + apply eqd_dec. + Defined. + (*******************************************************************************) (* Length-Indexed Lists *)