X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fmodels%2Flist_support.ma;h=ea701bb9db69cba6b1efa532efc6b01123fef0d8;hb=b5564e329d48efa6c2ca01da18203def26a70294;hp=1125e2a91f119eeaaa0d549336693db925dedace;hpb=910c252965fe17d6b5af92e4658e7d02bac82d58;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/models/list_support.ma b/helm/software/matita/contribs/dama/dama/models/list_support.ma index 1125e2a91..ea701bb9d 100644 --- a/helm/software/matita/contribs/dama/dama/models/list_support.ma +++ b/helm/software/matita/contribs/dama/dama/models/list_support.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "nat/minus.ma". include "list/list.ma". interpretation "list nth" 'nth = (nth _). @@ -41,21 +42,21 @@ lemma len_mk_list : ∀T:Type.∀f:nat→T.∀n.\len (\mk_list f n) = n. intros; elim n; [reflexivity] simplify; rewrite > H; reflexivity; qed. -record rel : Type ≝ { - rel_T :> Type; +record rel (rel_T : Type) : Type ≝ { rel_op :2> rel_T → rel_T → Prop }. record trans_rel : Type ≝ { - o_rel :> rel; - o_tra : ∀x,y,z: o_rel.o_rel x y → o_rel y z → o_rel x z + o_T :> Type; + o_rel :> rel o_T; + o_tra : ∀x,y,z: o_T.o_rel x y → o_rel y z → o_rel x z }. lemma trans: ∀r:trans_rel.∀x,y,z:r.r x y → r y z → r x z. apply o_tra; qed. -inductive sorted (lt : trans_rel): list (rel_T lt) → Prop ≝ +inductive sorted (lt : trans_rel): list (o_T lt) → Prop ≝ | sorted_nil : sorted lt [] | sorted_one : ∀x. sorted lt [x] | sorted_cons : ∀x,y,tl. lt x y → sorted lt (y::tl) → sorted lt (x::y::tl). @@ -67,7 +68,7 @@ lemma len_append: ∀T:Type.∀l1,l2:list T. \len (l1@l2) = \len l1 + \len l2. intros; elim l1; [reflexivity] simplify; rewrite < H; reflexivity; qed. -inductive non_empty_list (A:Type) : list A → Type := +coinductive non_empty_list (A:Type) : list A → Type := | show_head: ∀x,l. non_empty_list A (x::l). lemma len_gt_non_empty : @@ -84,8 +85,8 @@ lemma sorted_skip: ∀r,x,y,l. sorted r (x::y::l) → sorted r (x::l). intros (r x y l H1); inversion H1; intros; [1,2: destruct H] destruct H4; inversion H2; intros; [destruct H4] [1: destruct H4; constructor 2; -|2: destruct H7; constructor 3; [ apply (o_tra ? ??? H H4);] - apply (sorted_tail ??? H2);] +|2: destruct H7; constructor 3; + [ apply (o_tra ? ??? H H4); | apply (sorted_tail ??? H2);]] qed. lemma sorted_tail_bigger : ∀r,x,l,d.sorted r (x::l) → ∀i. i < \len l → r x (\nth l d i). @@ -96,31 +97,20 @@ cases i in H2; destruct H6; simplify; assumption;] qed. -lemma all_bases_positive : ∀f:q_f.∀i. OQ < nth_base (bars f) (S i). -intro f; generalize in match (bars_begin_OQ f); generalize in match (bars_sorted f); -cases (bars_not_nil f); intros; -cases (cmp_nat i (len l)); -[1: lapply (sorted_tail_bigger ?? H ? H2) as K; simplify in H1; - rewrite > H1 in K; apply K; -|2: rewrite > H2; simplify; elim l; simplify; [apply (q_pos_OQ one)] - assumption; -|3: simplify; elim l in i H2;[simplify; rewrite > nth_nil; apply (q_pos_OQ one)] - cases n in H3; intros; [cases (not_le_Sn_O ? H3)] apply (H2 n1); - apply (le_S_S_to_le ?? H3);] -qed. - +(* move in nat/ *) lemma lt_n_plus_n_Sm : ∀n,m:nat.n < n + S m. -intros; rewrite > sym_plus; apply (le_S_S n (m+n)); apply (le_plus_n m n); qed. +intros; rewrite > sym_plus; apply (le_S_S n (m+n)); alias id "le_plus_n" = "cic:/matita/nat/le_arith/le_plus_n.con". +apply (le_plus_n m n); qed. -lemma nth_concat_lt_len: - ∀T:Type.∀l1,l2:list T.∀def.∀i.i < len l1 → nth (l1@l2) def i = nth l1 def i. +lemma nth_append_lt_len: + ∀T:Type.∀l1,l2:list T.∀def.∀i.i < \len l1 → \nth (l1@l2) def i = \nth l1 def i. intros 4; elim l1; [cases (not_le_Sn_O ? H)] cases i in H H1; simplify; intros; [reflexivity| rewrite < H;[reflexivity] apply le_S_S_to_le; apply H1] qed. -lemma nth_concat_ge_len: +lemma nth_append_ge_len: ∀T:Type.∀l1,l2:list T.∀def.∀i. - len l1 ≤ i → nth (l1@l2) def i = nth l2 def (i - len l1). + \len l1 ≤ i → \nth (l1@l2) def i = \nth l2 def (i - \len l1). intros 4; elim l1; [ rewrite < minus_n_O; reflexivity] cases i in H1; simplify; intros; [cases (not_le_Sn_O ? H1)] apply H; apply le_S_S_to_le; apply H1; @@ -128,31 +118,127 @@ qed. lemma nth_len: ∀T:Type.∀l1,l2:list T.∀def,x. - nth (l1@x::l2) def (len l1) = x. + \nth (l1@x::l2) def (\len l1) = x. intros 2; elim l1;[reflexivity] simplify; apply H; qed. -lemma all_bigger_can_concat_bigger: - ∀l1,l2,start,b,x,n. - (∀i.i< len l1 → nth_base l1 i < \fst b) → - (∀i.i< len l2 → \fst b ≤ nth_base l2 i) → - (∀i.i< len l1 → start ≤ i → x ≤ nth_base l1 i) → - start ≤ n → n < len (l1@b::l2) → x ≤ \fst b → x ≤ nth_base (l1@b::l2) n. -intros; cases (cmp_nat n (len l1)); -[1: unfold nth_base; rewrite > (nth_concat_lt_len ????? H6); - apply (H2 n); assumption; -|2: rewrite > H6; unfold nth_base; rewrite > nth_len; assumption; -|3: unfold nth_base; rewrite > nth_concat_ge_len; [2: apply lt_to_le; assumption] - rewrite > len_concat in H4; simplify in H4; rewrite < plus_n_Sm in H4; - lapply linear le_S_S_to_le to H4 as K; rewrite > sym_plus in K; - lapply linear le_plus_to_minus to K as X; - generalize in match X; generalize in match (n - len l1); intro W; cases W; clear W X; - [intros; assumption] intros; - apply (q_le_trans ??? H5); apply (H1 n1); assumption;] -qed. - lemma sorted_head_smaller: - ∀l,p. sorted (p::l) → ∀i.i < len l → \fst p < nth_base l i. -intro l; elim l; intros; [cases (not_le_Sn_O ? H1)] cases i in H2; simplify; intros; + ∀r,l,p,d. sorted r (p::l) → ∀i.i < \len l → r p (\nth l d i). +intros 2 (r l); elim l; [cases (not_le_Sn_O ? H1)] cases i in H2; simplify; intros; [1: inversion H1; [1,2: simplify; intros; destruct H3] intros; destruct H6; assumption; -|2: apply (H p ? n ?); [apply (sorted_skip ??? H1)] apply le_S_S_to_le; apply H2] -qed. +|2: apply (H p ?? n ?); [apply (sorted_skip ???? H1)] apply le_S_S_to_le; apply H2] +qed. + +alias symbol "lt" = "natural 'less than'". +lemma sorted_pivot: + ∀r,l1,l2,p,d. sorted r (l1@p::l2) → + (∀i. i < \len l1 → r (\nth l1 d i) p) ∧ + (∀i. i < \len l2 → r p (\nth l2 d i)). +intros 2 (r l); elim l; +[1: split; [intros; cases (not_le_Sn_O ? H1);] intros; + apply sorted_head_smaller; assumption; +|2: simplify in H1; cases (H ?? d (sorted_tail ??? H1)); + lapply depth = 0 (sorted_head_smaller ??? d H1) as Hs; + split; simplify; intros; + [1: cases i in H4; simplify; intros; + [1: lapply depth = 0 (Hs (\len l1)) as HS; + rewrite > nth_len in HS; apply HS; + rewrite > len_append; simplify; apply lt_n_plus_n_Sm; + |2: apply (H2 n); apply le_S_S_to_le; apply H4] + |2: apply H3; assumption]] +qed. + +coinductive cases_bool (p:bool) : bool → CProp ≝ +| case_true : p = true → cases_bool p true +| cases_false : p = false → cases_bool p false. + +lemma case_b : ∀A:Type.∀f:A → bool. ∀x.cases_bool (f x) (f x). +intros; cases (f x);[left;|right] reflexivity; +qed. + +coinductive break_spec (T : Type) (n : nat) (l : list T) : list T → CProp ≝ +| break_to: ∀l1,x,l2. \len l1 = n → l = l1 @ [x] @ l2 → break_spec T n l l. + +lemma list_break: ∀T,n,l. n < \len l → break_spec T n l l. +intros 2; elim n; +[1: elim l in H; [cases (not_le_Sn_O ? H)] + apply (break_to ?? ? [] a l1); reflexivity; +|2: cases (H l); [2: apply lt_S_to_lt; assumption;] cases l2 in H3; intros; + [1: rewrite < H2 in H1; rewrite > H3 in H1; rewrite > append_nil in H1; + rewrite > len_append in H1; rewrite > plus_n_SO in H1; + cases (not_le_Sn_n ? H1); + |2: apply (break_to ?? ? (l1@[x]) t l3); + [2: simplify; rewrite > associative_append; assumption; + |1: rewrite < H2; rewrite > len_append; rewrite > plus_n_SO; reflexivity]]] +qed. + +include "logic/cprop_connectives.ma". + +definition eject_N ≝ + λP.λp:∃x:nat.P x.match p with [ex_introT p _ ⇒ p]. +coercion eject_N. +definition inject_N ≝ λP.λp:nat.λh:P p. ex_introT ? P p h. +coercion inject_N with 0 1 nocomposites. + +coinductive find_spec (T:Type) (P:T→bool) (l:list T) (d:T) (res : nat) : nat → CProp ≝ +| found: + ∀i. i < \len l → P (\nth l d i) = true → res = i → + (∀j. j < i → P (\nth l d j) = false) → find_spec T P l d res i +| not_found: ∀i. i = \len l → res = i → + (∀j.j < \len l → P (\nth l d j) = false) → find_spec T P l d res i. + +lemma find_lemma : ∀T.∀P:T→bool.∀l:list T.∀d.∃i.find_spec ? P l d i i. +intros; +letin find ≝ ( + let rec aux (acc: nat) (l : list T) on l : nat ≝ + match l with + [ nil ⇒ acc + | cons x tl ⇒ + match P x with + [ false ⇒ aux (S acc) tl + | true ⇒ acc]] + in aux : + ∀acc,l1.∃p:nat. + ∀story. story @ l1 = l → acc = \len story → + find_spec ? P story d acc acc → + find_spec ? P (story @ l1) d p p); +[4: clearbody find; cases (find 0 l); + lapply (H [] (refl_eq ??) (refl_eq ??)) as K; + [2: apply (not_found ?? [] d); intros; try reflexivity; cases (not_le_Sn_O ? H1); + |1: cases K; clear K; + [2: exists[apply (\len l)] + apply not_found; try reflexivity; apply H3; + |1: exists[apply i] apply found; try reflexivity; assumption;]] +|1: intros; cases (aux (S n) l2); simplify; clear aux; + lapply depth = 0 (H5 (story@[t])) as K; clear H5; + change with (find_spec ? P (story @ ([t] @ l2)) d w w); + rewrite < associative_append; apply K; clear K; + [1: rewrite > associative_append; apply H2; + |2: rewrite > H3; rewrite > len_append; rewrite > sym_plus; reflexivity; + |3: cases H4; clear H4; destruct H7; + [2: rewrite > H5; rewrite > (?:S (\len story) = \len (story @ [t])); [2: + rewrite > len_append; rewrite > sym_plus; reflexivity;] + apply not_found; try reflexivity; intros; cases (cmp_nat (S j) (\len story)); + [1: rewrite > (nth_append_lt_len ????? H8); apply H7; apply H8; + |2: rewrite > (nth_append_ge_len ????? (le_S_S_to_le ?? H8)); + rewrite > (?: j - \len story = 0);[assumption] + rewrite > (?:j = \len story);[rewrite > minus_n_n; reflexivity] + apply le_to_le_to_eq; [2: apply le_S_S_to_le; assumption;] + rewrite > len_append in H4;rewrite > sym_plus in H4; + apply le_S_S_to_le; apply H4;] + |1: rewrite < H3 in H5; cases (not_le_Sn_n ? H5);]] +|2: intros; cases H4; clear H4; + [1: destruct H7; rewrite > H3 in H5; cases (not_le_Sn_n ? H5); + |2: apply found; try reflexivity; + [1: rewrite > len_append; simplify; rewrite > H5; apply lt_n_plus_n_Sm; + |2: rewrite > H5; rewrite > nth_append_ge_len; [2: apply le_n] + rewrite < minus_n_n; assumption; + |3: intros; rewrite > H5 in H4; rewrite > nth_append_lt_len; [2:assumption] + apply H7; assumption]] +|3: intros; rewrite > append_nil; assumption;] +qed. + +lemma find : ∀T:Type.∀P:T→bool.∀l:list T.∀d:T.nat. +intros; cases (find_lemma ? f l t); apply w; qed. + +lemma cases_find: ∀T,P,l,d. find_spec T P l d (find T P l d) (find T P l d). +intros; unfold find; cases (find_lemma T P l d); simplify; assumption; qed.