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=8dad0c436b9b224bcb73e7527e86bba5fde54e1b;hpb=8f4162a9db17a597d4fba49eb957009fc0268378;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 8dad0c436..ea701bb9d 100644 --- a/helm/software/matita/contribs/dama/dama/models/list_support.ma +++ b/helm/software/matita/contribs/dama/dama/models/list_support.ma @@ -68,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 : @@ -97,21 +97,6 @@ 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)); alias id "le_plus_n" = "cic:/matita/nat/le_arith/le_plus_n.con". @@ -162,7 +147,7 @@ intros 2 (r l); elim l; |2: apply H3; assumption]] qed. -inductive cases_bool (p:bool) : bool → CProp ≝ +coinductive cases_bool (p:bool) : bool → CProp ≝ | case_true : p = true → cases_bool p true | cases_false : p = false → cases_bool p false. @@ -170,7 +155,23 @@ lemma case_b : ∀A:Type.∀f:A → bool. ∀x.cases_bool (f x) (f x). intros; cases (f x);[left;|right] reflexivity; qed. -include "cprop_connectives.ma". +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]. @@ -178,7 +179,7 @@ coercion eject_N. definition inject_N ≝ λP.λp:nat.λh:P p. ex_introT ? P p h. coercion inject_N with 0 1 nocomposites. -inductive find_spec (T:Type) (P:T→bool) (l:list T) (d:T) (res : nat) : nat → CProp ≝ +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