X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fmodels%2Fq_bars.ma;h=65066590f4baef3f754305ba709df1353d149bae;hb=591ffe6f23ec9d2a4d368d2c1e7b213986189e44;hp=2aae66a5bc05555ac2f5e7bee7703a26d739ea57;hpb=d34d7101ba59a19fa030fe9a5c6b3d563efc8f3d;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/models/q_bars.ma b/helm/software/matita/contribs/dama/dama/models/q_bars.ma index 2aae66a5b..65066590f 100644 --- a/helm/software/matita/contribs/dama/dama/models/q_bars.ma +++ b/helm/software/matita/contribs/dama/dama/models/q_bars.ma @@ -14,7 +14,7 @@ include "nat_ordered_set.ma". include "models/q_support.ma". -include "models/list_support.ma". +include "models/list_support.ma". include "cprop_connectives.ma". definition bar ≝ ℚ × ℚ. @@ -29,146 +29,151 @@ interpretation "q0" 'empty_bar = empty_bar. notation < "\ldots\rect\square\EmptySmallSquare\ldots" with precedence 90 for @{'lq2}. interpretation "lq2" 'lq2 = (list bar). -inductive sorted : list bar → Prop ≝ -| sorted_nil : sorted [] -| sorted_one : ∀x. sorted [x] -| sorted_cons : ∀x,y,tl. \fst x < \fst y → sorted (y::tl) → sorted (x::y::tl). +definition q2_lt := mk_rel bar (λx,y:bar.\fst x < \fst y). -definition nth_base ≝ λf,n. \fst (nth f ▭ n). -definition nth_height ≝ λf,n. \snd (nth f ▭ n). +interpretation "bar lt" 'lt x y = (rel_op _ q2_lt x y). + +lemma q2_trans : ∀a,b,c:bar. a < b → b < c → a < c. +intros 3; cases a; cases b; cases c; unfold q2_lt; simplify; intros; +apply (q_lt_trans ??? H H1); +qed. + +definition q2_trel := mk_trans_rel bar q2_lt q2_trans. + +interpretation "bar lt" 'lt x y = (FunClass_2_OF_trans_rel q2_trel x y). + +definition canonical_q_lt : rel bar → trans_rel ≝ λx:rel bar.q2_trel. + +coercion canonical_q_lt with nocomposites. + +interpretation "bar lt" 'lt x y = (FunClass_2_OF_trans_rel (canonical_q_lt _) x y). + +definition nth_base ≝ λf,n. \fst (\nth f ▭ n). +definition nth_height ≝ λf,n. \snd (\nth f ▭ n). record q_f : Type ≝ { bars: list bar; - increasing_bars : sorted bars; + bars_sorted : sorted q2_lt bars; bars_begin_OQ : nth_base bars O = OQ; - bars_tail_OQ : nth_height bars (pred (len bars)) = OQ + bars_end_OQ : nth_height bars (pred (\len bars)) = OQ }. -lemma nth_nil: ∀T,i.∀def:T. nth [] def i = def. -intros; elim i; simplify; [reflexivity;] assumption; qed. - -lemma all_bases_positives : ∀f:q_f.∀i.i < len (bars f) → OQ < nth_base (bars f) i. -intro f; elim (increasing_bars f); -[1: unfold nth_base; rewrite > nth_nil; apply (q_pos_OQ one); -|2: cases i in H; [2: cases (?:False); +lemma len_bases_gt_O: ∀f.O < \len (bars f). +intros; generalize in match (bars_begin_OQ f); cases (bars f); intros; +[2: simplify; apply le_S_S; apply le_O_n; +|1: normalize in H; destruct H;] +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 (len_gt_non_empty ?? (len_bases_gt_O f)); intros; +cases (cmp_nat (\len l) i); +[2: lapply (sorted_tail_bigger q2_lt ?? ▭ H ? H2) as K; + simplify in H1; rewrite < H1; apply K; +|1: simplify; elim l in i H2;[simplify; rewrite > nth_nil; apply (q_pos_OQ one)] + cases n in H3; intros; [simplify in H3; cases (not_le_Sn_O ? H3)] + apply (H2 n1); simplify in H3; apply (le_S_S_to_le ?? H3);] qed. -definition eject_Q ≝ - λP.λp:∃x:ℚ.P x.match p with [ex_introT p _ ⇒ p]. -coercion eject_Q. -definition inject_Q ≝ λP.λp:ℚ.λh:P p. ex_introT ? P p h. -coercion inject_Q with 0 1 nocomposites. - -definition value_spec : q_f → ℚ → ℚ → Prop ≝ - λf,i,q. - ∃j. q = nth_height (bars f) j ∧ - (nth_base (bars f) j < i ∧ - ∀n.j < n → n < len (bars f) → i ≤ nth_base (bars f) n). +(* +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. +*) + +(* +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. +*) + + +inductive value_spec (f : q_f) (i : ℚ) : ℚ → nat → CProp ≝ + value_of : ∀q,j. + nth_height (bars f) j = q → + nth_base (bars f) j < i → + (∀n.j < n → n < \len (bars f) → i ≤ nth_base (bars f) n) → value_spec f i q j. + + +inductive 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. -definition value : ∀f:q_f.∀i:ratio.∃p:ℚ.value_spec f (Qpos i) p. +definition value : ∀f:q_f.∀i:ratio.∃p:ℚ.∃j.value_spec f (Qpos i) p j. intros; -alias symbol "lt" (instance 5) = "Q less than". -alias symbol "leq" = "Q less or equal than". -letin value_spec_aux ≝ ( - λf,i,q.∃j. q = nth_height f j ∧ - (nth_base f j < i ∧ ∀n.j < n → n < len f → i ≤ nth_base f n)); -letin value ≝ ( - let rec value (acc: ℚ) (l : list bar) on l : ℚ ≝ - match l with - [ nil ⇒ acc - | cons x tl ⇒ - match q_cmp (\fst x) (Qpos i) with - [ q_leq _ ⇒ value (\snd x) tl - | q_gt _ ⇒ acc]] - in value : - ∀acc,l.∃p:ℚ. OQ ≤ acc → value_spec_aux l (Qpos i) p); -[4: clearbody value; cases (value OQ (bars f)) (p Hp); exists[apply p]; - cases (Hp (q_le_n ?)) (j Hj); cases Hj (Hjp H); cases H (Hin Hmax); - clear Hp value value_spec_aux Hj H; exists [apply j]; split[2:split;intros;] - try apply Hmax; assumption; -|1: intro Hacc; clear H2; cases (value (\snd b) l1) (j Hj); - cases (q_cmp (\snd b) (Qpos i)) (Hib Hib); - [1: cases (Hj Hib) (w Hw); simplify in ⊢ (? ? ? %); clear Hib Hj; - exists [apply (S w)] cases Hw; cases H3; clear Hw H3; - split; try assumption; split; try assumption; intros; - apply (q_le_trans ??? (H5 (pred n) ??)); [3: apply q_le_n] - - - - -[5: clearbody value; - cases (q_cmp i (start f)); - [2: exists [apply 〈O,OQ〉] simplify; constructor 1; split; try assumption; - try reflexivity; apply q_lt_to_le; assumption; - |1: cases (bars f); [exists [apply 〈O,OQ〉] simplify; constructor 3; split;try assumption;reflexivity;] - cases (value ⅆ[i,start f] (b::l)) (p Hp); - cases (Hp (q_dist_ge_OQ ? ?)); clear Hp value; [cases H1; destruct H2] - cases H1; clear H1; lapply (sum_bases_O (b::l) (\fst p)) as H1; - [2: apply (q_le_trans ??? H2); rewrite > H; apply q_eq_to_le; - rewrite > q_d_x_x; reflexivity; - |1: exists [apply p] simplify; constructor 4; rewrite > H1; split; - try split; try rewrite > q_d_x_x; try autobatch depth=2; - [1: rewrite > H; rewrite > q_plus_sym; apply q_lt_plus; - rewrite > q_plus_minus; apply q_lt_plus_trans; [apply sum_bases_ge_OQ] - apply q_pos_lt_OQ; - |2: rewrite > H; rewrite > q_d_x_x; apply q_eq_to_le; reflexivity; - |3: rewrite > H; rewrite > q_d_x_x; apply q_lt_plus_trans; - try apply sum_bases_ge_OQ; apply q_pos_lt_OQ;]] - |3: cases (q_cmp i (start f+sum_bases (bars f) (len (bars f)))); - [1: exists [apply 〈O,OQ〉] simplify; constructor 2; split; try assumption; - try reflexivity; rewrite > H1; apply q_eq_to_le; reflexivity; - |3: exists [apply 〈O,OQ〉] simplify; constructor 2; split; try assumption; - try reflexivity; apply q_lt_to_le; assumption; - |2: generalize in match (refl_eq ? (bars f): bars f = bars f); - generalize in match (bars f) in ⊢ (??? % → %); intro X; cases X; clear X; - intros; - [1: exists [apply 〈O,OQ〉] simplify; constructor 3; split; reflexivity; - |2: cases (value ⅆ[i,start f] (b::l)) (p Hp); - cases (Hp (q_dist_ge_OQ ? ?)); clear Hp value; [cases H3;destruct H4] - cases H3; clear H3; - exists [apply p]; constructor 4; split; try split; try assumption; - [1: intro X; destruct X; - |2: apply q_lt_to_le; assumption; - |3: rewrite < H2; assumption; - |4: cases (cmp_nat (\fst p) (len (bars f))); - [1:apply lt_to_le;rewrite H3;rewrite < H2;apply le_n] - cases (?:False); cases (\fst p) in H3 H4 H6; clear H5; - [1: intros; apply (not_le_Sn_O ? H5); - |2: rewrite > q_d_sym; rewrite > q_d_noabs; [2: apply q_lt_to_le; assumption] - intros; lapply (q_lt_inj_plus_r ?? (Qopp (start f)) H1); clear H1; - generalize in match Hletin; - rewrite > (q_plus_sym (start f)); rewrite < q_plus_assoc; - do 2 rewrite < q_elim_minus; rewrite > q_plus_minus; - rewrite > q_plus_OQ; intro K; apply (q_lt_corefl (i-start f)); - apply (q_lt_le_trans ???? H3); rewrite < H2; - apply (q_lt_trans ??? K); apply sum_bases_increasing; - assumption;]]]]] -|1,3: intros; right; split; - [1,4: clear H2; cases (value (q-Qpos (\fst b)) l1); - cases (H2 (q_le_to_diff_ge_OQ ?? (? H1))); - [1: intro; apply q_lt_to_le;assumption; - |3: simplify; cases H4; apply q_le_minus; assumption; - |2,5: simplify; cases H4; rewrite > H5; rewrite > H6; - apply q_le_minus; apply sum_bases_empty_nat_of_q_le_q; - |4: intro X; rewrite > X; apply q_eq_to_le; reflexivity; - |*: simplify; apply q_le_minus; cases H4; assumption;] - |2,5: cases (value (q-Qpos (\fst b)) l1); - cases (H4 (q_le_to_diff_ge_OQ ?? (? H1))); - [1,4: intros; [apply q_lt_to_le|apply q_eq_to_le;symmetry] assumption; - |3,6: cases H5; simplify; change with (q < sum_bases l1 (S (\fst w)) + Qpos (\fst b)); - apply q_lt_plus; assumption; - |2,5: simplify; cases H5; rewrite > H6; simplify; rewrite > H7; - apply q_lt_plus; apply sum_bases_empty_nat_of_q_le_q_one;] - |*: cases (value (q-Qpos (\fst b)) l1); simplify; - cases (H4 (q_le_to_diff_ge_OQ ?? (? H1))); - [1,4: intros; [apply q_lt_to_le|apply q_eq_to_le;symmetry] assumption; - |3,6: cases H5; assumption; - |*: cases H5; rewrite > H6; rewrite > H8; - elim (\fst w); [1,3:reflexivity;] simplify; assumption;]] -|2: clear value H2; simplify; intros; right; split; [assumption|3:reflexivity] - rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption; -|4: intros; left; split; reflexivity;] -qed. +letin P ≝ + (λx:bar.match q_cmp (Qpos i) (\fst x) with[ q_leq _ ⇒ true| q_gt _ ⇒ false]); +exists [apply (nth_height (bars f) (pred (find ? P (bars f) ▭)));] +exists [apply (pred (find ? P (bars f) ▭))] apply value_of; +[1: reflexivity +|2: cases (cases_find bar P (bars f) ▭); + [1: cases i1 in H H1 H2 H3; simplify; intros; + [1: generalize in match (bars_begin_OQ f); + cases (len_gt_non_empty ?? (len_bases_gt_O f)); simplify; intros; + rewrite > H4; apply q_pos_OQ; + |2: cases (len_gt_non_empty ?? (len_bases_gt_O f)) in H3; + intros; lapply (H3 n (le_n ?)) as K; unfold P in K; + cases (q_cmp (Qpos i) (\fst (\nth (x::l) ▭ n))) in K; + simplify; intros; [destruct H5] assumption] + |2: destruct H; cases (len_gt_non_empty ?? (len_bases_gt_O f)) in H2; + simplify; intros; lapply (H (\len l) (le_n ?)) as K; clear H; + unfold P in K; cases (q_cmp (Qpos i) (\fst (\nth (x::l) ▭ (\len l)))) in K; + simplify; intros; [destruct H2] assumption;] +|3: intro; cases (cases_find bar P (bars f) ▭); intros; + [1: generalize in match (bars_sorted f); + cases (list_break ??? H) in H1; rewrite > H6; + rewrite < H1; simplify; rewrite > nth_len; unfold P; + cases (q_cmp (Qpos i) (\fst x)); simplify; + intros (X Hs); [2: destruct X] clear X; + cases (sorted_pivot q2_lt ??? ▭ Hs); + cut (\len l1 ≤ n) as Hn; [2: + rewrite > H1; cases i1 in H4; simplify; intro X; [2: assumption] + apply lt_to_le; assumption;] + unfold nth_base; rewrite > (nth_append_ge_len ????? Hn); + cut (n - \len l1 < \len (x::l2)) as K; [2: + simplify; rewrite > H1; rewrite > (?:\len l2 = \len (bars f) - \len (l1 @ [x]));[2: + rewrite > H6; repeat rewrite > len_append; simplify; + repeat rewrite < plus_n_Sm; rewrite < plus_n_O; simplify; + rewrite > sym_plus; rewrite < minus_plus_m_m; reflexivity;] + rewrite > len_append; rewrite > H1; simplify; rewrite < plus_n_SO; + apply le_S_S; clear H1 H6 H7 Hs H8 H9 Hn x l2 l1 H4 H3 H2 H P i; + elim (\len (bars f)) in i1 n H5; [cases (not_le_Sn_O ? H);] + simplify; cases n2; [ repeat rewrite < minus_n_O; apply le_S_S_to_le; assumption] + cases n1 in H1; [intros; rewrite > eq_minus_n_m_O; apply le_O_n] + intros; simplify; apply H; apply le_S_S_to_le; assumption;] + cases (n - \len l1) in K; simplify; intros; [ assumption] + lapply (H9 ? (le_S_S_to_le ?? H10)) as W; apply (q_le_trans ??? H7); + apply q_lt_to_le; apply W; + |2: cases (not_le_Sn_n i1); rewrite > H in ⊢ (??%); + apply (trans_le ??? ? H4); cases i1 in H3; intros; apply le_S_S; + [ apply le_O_n; | assumption]]] +qed. lemma value_OQ_l: ∀l,i.i < start l → \snd (\fst (value l i)) = OQ.