X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fmodels%2Fq_function.ma;h=b1ff639e44b959bd6a2a482c001df38d552316f9;hb=e0b4028cb1f8423b40d5f9ad396f10f42db86f0e;hp=bf2eb82d8f210ffdb8e9c3e362bb5fa7da2e6278;hpb=d309714a7f00acfae311fa24612e57e9be085ff3;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/models/q_function.ma b/helm/software/matita/contribs/dama/dama/models/q_function.ma index bf2eb82d8..b1ff639e4 100644 --- a/helm/software/matita/contribs/dama/dama/models/q_function.ma +++ b/helm/software/matita/contribs/dama/dama/models/q_function.ma @@ -13,53 +13,22 @@ (**************************************************************************) include "nat_ordered_set.ma". -include "models/q_bars.ma". +include "models/q_shift.ma". -lemma initial_shift_same_values: - ∀l1:q_f.∀init.init < start l1 → - same_values l1 - (mk_q_f init (〈\fst (unpos (start l1 - init) ?),OQ〉:: bars l1)). -[apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption] -intros; generalize in ⊢ (? ? (? ? (? ? (? ? ? (? ? ? (? ? %)) ?) ?))); intro; -cases (unpos (start l1-init) H1); intro input; -simplify in ⊢ (? ? ? (? ? ? (? ? ? (? (? ? (? ? (? ? ? % ?) ?)) ?)))); -cases (value (mk_q_f init (〈w,OQ〉::bars l1)) input) (v1 Hv1); -cases Hv1 (HV1 HV1 HV1 HV1); cases HV1 (Hi1 Hv11 Hv12); clear HV1 Hv1; -[1: cut (input < start l1) as K;[2: apply (q_lt_trans ??? Hi1 H)] - rewrite > (value_OQ_l ?? K); simplify; symmetry; assumption; -|2: cut (start l1 + sum_bases (bars l1) (len (bars l1)) ≤ input) as K;[2: - simplify in Hi1; apply (q_le_trans ???? Hi1); rewrite > H2; - rewrite > q_plus_sym in ⊢ (? ? (? ? %)); - rewrite > q_plus_assoc; rewrite > q_elim_minus; - rewrite > q_plus_sym in ⊢ (? ? (? (? ? %) ?)); - rewrite > q_plus_assoc; rewrite < q_elim_minus; - rewrite > q_plus_minus; rewrite > q_plus_sym in ⊢ (? ? (? % ?)); - rewrite > q_plus_OQ; apply q_eq_to_le; reflexivity;] - rewrite > (value_OQ_r ?? K); simplify; symmetry; assumption; -|3: simplify in Hi1; destruct Hi1; -|4: - -STOP - -qed. - - - alias symbol "pi2" = "pair pi2". alias symbol "pi1" = "pair pi1". definition rebase_spec ≝ ∀l1,l2:q_f.∃p:q_f × q_f. And4 - (*len (bars (\fst p)) = len (bars (\snd p))*) (start (\fst p) = start (\snd p)) - (same_bases (\fst p) (\snd p)) + (same_bases (bars (\fst p)) (bars (\snd p))) (same_values l1 (\fst p)) (same_values l2 (\snd p)). definition rebase_spec_simpl ≝ λstart.λl1,l2:list bar.λp:(list bar) × (list bar). And3 - (same_bases (mk_q_f start (\fst p)) (mk_q_f start (\snd p))) + (same_bases (\fst p) (\snd p)) (same_values (mk_q_f start l1) (mk_q_f start (\fst p))) (same_values (mk_q_f start l2) (mk_q_f start (\snd p))). @@ -72,6 +41,8 @@ definition eject ≝ coercion eject. definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h. coercion inject with 0 1 nocomposites. + +axiom devil : False. definition rebase: rebase_spec. intros 2 (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2; @@ -116,7 +87,7 @@ in aux : ∀l1,l2,m.∃z.∀s.spec s l1 l2 m z); unfold spec; [1,2: assumption; |3: intro; apply (H3 input); |4: intro; rewrite > H in H4; - rewrite > (H4 input); reflexivity;] + rewrite > (H4 input) in ⊢ (? ? % ?); reflexivity;] |2: letin l2' ≝ (〈\fst (unpos (s2-s1) ?),OQ〉::l2);[ apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption] @@ -126,23 +97,97 @@ in aux : ∀l1,l2,m.∃z.∀s.spec s l1 l2 m z); unfold spec; [1: reflexivity |2: assumption; |3: assumption; - |4: intro; rewrite < (H4 input); clear H3 H4 H2 w; - cases (value (mk_q_f s1 l2') input); - cases (q_cmp input (start (mk_q_f s1 l2'))) in H1; - whd in ⊢ (% → ?); - [1: intros; cases H2; clear H2; whd in ⊢ (??? %); - cases (value (mk_q_f s2 l2) input); - cases (q_cmp input (start (mk_q_f s2 l2))) in H2; - whd in ⊢ (% → ?); - [1: intros; cases H6; clear H6; change with (w1 = w); - - (* TODO *) ]] + |4: intro; + rewrite > (initial_shift_same_values (mk_q_f s2 l2) s1 H input) in ⊢ (? ? % ?); + rewrite < (H4 input)in ⊢ (? ? ? %); reflexivity;] + |3: letin l1' ≝ (〈\fst (unpos (s1-s2) ?),OQ〉::l1);[ + apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ; + assumption] + cases (aux l1' l2 (S (len l1' + len l2))); + cases (H1 s2 (le_n ?)); clear H1 aux; + exists [apply 〈mk_q_f s2 (\fst w), mk_q_f s2 (\snd w)〉] split; + [1: reflexivity + |2: assumption; + |4: assumption; + |3: intro; simplify in ⊢ (? ? ? (? ? ? (? ? ? (? % ?)))); + rewrite > (initial_shift_same_values (mk_q_f s1 l1) s2 H input) in ⊢ (? ? % ?); + rewrite < (H3 input) in ⊢ (? ? ? %); reflexivity;]] |1,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption; -|3:(* TODO *) -|4:(* TODO *) -|5:(* TODO *) -|6:(* TODO *) -|7:(* TODO *) -|8: intros; cases (?:False); apply (not_le_Sn_O ? H1);] +|8: intros; cases (?:False); apply (not_le_Sn_O ? H1); +|3: intros; generalize in match (unpos ??); intro X; cases X; clear X; + simplify in ⊢ (???? (??? (??? (??? (?? (? (?? (??? % ?) ?) ??)))) ?)); + simplify in ⊢ (???? (???? (??? (??? (?? (? (?? (??? % ?) ?) ??)))))); + clear H4; cases (aux (〈w,\snd b〉::l4) l5 n1); clear aux; + cut (len (〈w,\snd b〉::l4) + len l5 < n1) as K;[2: + simplify in H5; simplify; rewrite > sym_plus in H5; simplify in H5; + rewrite > sym_plus in H5; apply le_S_S_to_le; apply H5;] + split; + [1: simplify in ⊢ (? % ?); simplify in ⊢ (? ? %); + cases (H4 s K); clear K H4; intro input; cases input; [reflexivity] + simplify; apply H7; + |2: simplify in ⊢ (? ? %); cases (H4 s K); clear H4 K H5 spec; + intro; + (* input < s + b1 || input >= s + b1 *) + |3: simplify in ⊢ (? ? %);] +|4: intros; generalize in match (unpos ??); intro X; cases X; clear X; + (* duale del 3 *) +|5: intros; (* triviale, caso in cui non fa nulla *) +|6,7: (* casi base in cui allunga la lista più corta *) +] +elim devil; qed. + +include "Q/q/qtimes.ma". + +let rec area (l:list bar) on l ≝ + match l with + [ nil ⇒ OQ + | cons he tl ⇒ area tl + Qpos (\fst he) * ⅆ[OQ,\snd he]]. + +alias symbol "pi1" = "exT \fst". +alias symbol "minus" = "Q minus". +alias symbol "exists" = "CProp exists". +definition minus_spec_bar ≝ + λf,g,h:list bar. + same_bases f g → len f = len g → + ∀s,i:ℚ. \snd (\fst (value (mk_q_f s h) i)) = + \snd (\fst (value (mk_q_f s f) i)) - \snd (\fst (value (mk_q_f s g) i)). + +definition minus_spec ≝ + λf,g:q_f. + ∃h:q_f. + ∀i:ℚ. \snd (\fst (value h i)) = + \snd (\fst (value f i)) - \snd (\fst (value g i)). + +definition eject_bar : ∀P:list bar → CProp.(∃l:list bar.P l) → list bar ≝ + λP.λp.match p with [ex_introT x _ ⇒ x]. +definition inject_bar ≝ ex_introT (list bar). + +coercion inject_bar with 0 1 nocomposites. +coercion eject_bar with 0 0 nocomposites. + +lemma minus_q_f : ∀f,g. minus_spec f g. +intros; +letin aux ≝ ( + let rec aux (l1, l2 : list bar) on l1 ≝ + match l1 with + [ nil ⇒ [] + | cons he1 tl1 ⇒ + match l2 with + [ nil ⇒ [] + | cons he2 tl2 ⇒ 〈\fst he1, \snd he1 - \snd he2〉 :: aux tl1 tl2]] + in aux : ∀l1,l2 : list bar.∃h.minus_spec_bar l1 l2 h); +[2: intros 4; simplify in H3; destruct H3; +|3: intros 4; simplify in H3; cases l1 in H2; [2: intro X; simplify in X; destruct X] + intros; rewrite > (value_OQ_e (mk_q_f s []) i); [2: reflexivity] + rewrite > q_elim_minus; rewrite > q_plus_OQ; reflexivity; +|1: cases (aux l2 l3); unfold in H2; intros 4; + simplify in ⊢ (? ? (? ? ? (? ? ? (? % ?))) ?); + cases (q_cmp i (s + Qpos (\fst b))); + + + +definition excess ≝ + λf,g.∃i.\snd (\fst (value f i)) < \snd (\fst (value g i)). +