X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;f=src%2FGeneral.v;h=96ac1baf758f6721787292c58d588c4e66f8a3c9;hb=32436fdf380f7f2efc7a70896268509e7b3e0d6f;hp=b0266f4ae9258c4534eb5265e1fce33f4f3c6a09;hpb=2f9327bdfddb3908dc279c02db25ad527b37e17a;p=coq-hetmet.git diff --git a/src/General.v b/src/General.v index b0266f4..96ac1ba 100644 --- a/src/General.v +++ b/src/General.v @@ -41,6 +41,13 @@ Notation "[ q ]" := (@T_Leaf (option _) (Some q)) : tree_scope. Notation "[ ]" := (@T_Leaf (option _) None) : tree_scope. Notation "[]" := (@T_Leaf (option _) None) : tree_scope. +Fixpoint InT {A} (a:A) (t:Tree ??A) {struct t} : Prop := + match t with + | T_Leaf None => False + | T_Leaf (Some x) => x = a + | T_Branch b1 b2 => InT a b1 \/ InT a b2 + end. + Open Scope tree_scope. Fixpoint mapTree {a b:Type}(f:a->b)(t:@Tree a) : @Tree b := @@ -52,7 +59,7 @@ Fixpoint mapOptionTree {a b:Type}(f:a->b)(t:@Tree ??a) : @Tree ??b := match t with | T_Leaf None => T_Leaf None | T_Leaf (Some x) => T_Leaf (Some (f x)) - | T_Branch l r => T_Branch (mapOptionTree f l) (mapOptionTree f r) + | T_Branch l r => T_Branch (mapOptionTree f l) (mapOptionTree f r) end. Fixpoint mapTreeAndFlatten {a b:Type}(f:a->@Tree b)(t:@Tree a) : @Tree b := match t with @@ -129,6 +136,14 @@ Lemma mapOptionTree_compose : forall A B C (f:A->B)(g:B->C)(l:Tree ??A), reflexivity. Qed. +Lemma mapOptionTree_extensional {A}{B}(f g:A->B) : (forall a, f a = g a) -> (forall t, mapOptionTree f t = mapOptionTree g t). + intros. + induction t. + destruct a; auto. + simpl; rewrite H; auto. + simpl; rewrite IHt1; rewrite IHt2; auto. + Qed. + Open Scope string_scope. Fixpoint treeToString {T}{TT:ToString T}(t:Tree ??T) : string := match t with @@ -223,6 +238,17 @@ Lemma mapOptionTree_comp a b c (f:a->b) (g:b->c) q : (mapOptionTree g (mapOption reflexivity. Qed. +Lemma leaves_unleaves {T}(t:list T) : leaves (unleaves t) = t. + induction t; auto. + simpl. + rewrite IHt; auto. + Qed. + +Lemma mapleaves' {T:Type}(t:list T){Q}{f:T->Q} : unleaves (map f t) = mapOptionTree f (unleaves t). + induction t; simpl in *; auto. + rewrite IHt; auto. + Qed. + (* handy facts: map preserves the length of a list *) Lemma map_on_nil : forall A B (s1:list A) (f:A->B), nil = map f s1 -> s1=nil. intros. @@ -322,6 +348,11 @@ Inductive distinct {T:Type} : list T -> Prop := | distinct_nil : distinct nil | distinct_cons : forall a ax, (@In _ a ax -> False) -> distinct ax -> distinct (a::ax). +Inductive distinctT {T:Type} : Tree ??T -> Prop := +| distinctT_nil : distinctT [] +| distinctT_leaf : forall t, distinctT [t] +| distinctT_cons : forall t1 t2, (forall v, InT v t1 -> InT v t2 -> False) -> distinctT (t1,,t2). + Lemma in_decidable {VV:Type}{eqdVV:EqDecidable VV} : forall (v:VV)(lv:list VV), sumbool (In v lv) (not (In v lv)). intros. induction lv. @@ -366,6 +397,10 @@ Lemma distinct_decidable {VV:Type}{eqdVV:EqDecidable VV} : forall (lv:list VV), apply n; auto. Defined. +Lemma distinctT_decidable {VV:Type}{eqdVV:EqDecidable VV} : forall (lv:Tree ??VV), sumbool (distinctT lv) (not (distinctT lv)). + admit. + Defined. + Lemma map_preserves_length {A}{B}(f:A->B)(l:list A) : (length l) = (length (map f l)). induction l; auto. simpl. @@ -849,3 +884,33 @@ Coercion FMT : FreshMonad >-> Funclass. Variable Prelude_error : forall {A}, string -> A. Extract Inlined Constant Prelude_error => "Prelude.error". + + + + +Ltac eqd_dec_refl X := + destruct (eqd_dec X X) as [eqd_dec1 | eqd_dec2]; + [ clear eqd_dec1 | set (eqd_dec2 (refl_equal _)) as eqd_dec2'; inversion eqd_dec2' ]. + +Lemma unleaves_injective : forall T (t1 t2:list T), unleaves t1 = unleaves t2 -> t1 = t2. + intros T. + induction t1; intros. + destruct t2. + auto. + inversion H. + destruct t2. + inversion H. + simpl in H. + inversion H. + set (IHt1 _ H2) as q. + rewrite q. + reflexivity. + Qed. + +Lemma fst_zip : forall T Q n (v1:vec T n)(v2:vec Q n), vec_map (@fst _ _) (vec_zip v1 v2) = v1. + admit. + Defined. + +Lemma snd_zip : forall T Q n (v1:vec T n)(v2:vec Q n), vec_map (@snd _ _) (vec_zip v1 v2) = v2. + admit. + Defined.