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=591ffe6f23ec9d2a4d368d2c1e7b213986189e44;hp=c458d418d3c46cf91efcd84b13b1ce2e3604c22e;hpb=5d5a1cc683e634f7b86cd1b8f36a52e204bcee32;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 c458d418d..b1ff639e4 100644 --- a/helm/software/matita/contribs/dama/dama/models/q_function.ma +++ b/helm/software/matita/contribs/dama/dama/models/q_function.ma @@ -12,205 +12,46 @@ (* *) (**************************************************************************) -include "Q/q/q.ma". -include "list/list.ma". -include "cprop_connectives.ma". - - -notation "\rationals" non associative with precedence 99 for @{'q}. -interpretation "Q" 'q = Q. - -definition bar ≝ ratio × ℚ. (* base (Qpos) , height *) -record q_f : Type ≝ { start : ℚ; bars: list bar }. - -axiom qp : ℚ → ℚ → ℚ. -axiom qm : ℚ → ℚ → ℚ. -axiom qlt : ℚ → ℚ → CProp. - -interpretation "Q plus" 'plus x y = (qp x y). -interpretation "Q minus" 'minus x y = (qm x y). -interpretation "Q less than" 'lt x y = (qlt x y). - -inductive q_comparison (a,b:ℚ) : CProp ≝ - | q_eq : a = b → q_comparison a b - | q_lt : a < b → q_comparison a b - | q_gt : b < a → q_comparison a b. - -axiom q_cmp:∀a,b:ℚ.q_comparison a b. - -definition qle ≝ λa,b:ℚ.a = b ∨ a < b. - -interpretation "Q less or equal than" 'leq x y = (qle x y). - -axiom q_le_minus: ∀a,b,c:ℚ. a ≤ c - b → a + b ≤ c. -axiom q_lt_plus: ∀a,b,c:ℚ. a - b < c → a < c + b. - -axiom dist : ℚ → ℚ → ℚ. - - -interpretation "list nth" 'nth = (nth _). -interpretation "list nth" 'nth_appl l d i = (nth _ l d i). -notation "'nth'" with precedence 90 for @{'nth}. -notation < "'nth' \nbsp term 90 l \nbsp term 90 d \nbsp term 90 i" -with precedence 69 for @{'nth_appl $l $d $i}. - -notation < "\rationals \sup 2" non associative with precedence 90 for @{'q2}. -interpretation "Q x Q" 'q2 = (Prod Q Q). - -definition make_list ≝ - λA:Type.λdef:nat→A. - let rec make_list (n:nat) on n ≝ - match n with [ O ⇒ nil ? | S m ⇒ def m :: make_list m] - in make_list. - -interpretation "'mk_list' appl" 'mk_list_appl f n = (make_list _ f n). -interpretation "'mk_list'" 'mk_list = (make_list _). -notation "'mk_list'" with precedence 90 for @{'mk_list}. -notation < "'mk_list' \nbsp term 90 f \nbsp term 90 n" -with precedence 69 for @{'mk_list_appl $f $n}. - - -definition empty_bar : bar ≝ 〈one,OQ〉. -notation "\rect" with precedence 90 for @{'empty_bar}. -interpretation "q0" 'empty_bar = empty_bar. - -notation < "\ldots\rect\square\EmptySmallSquare\ldots" with precedence 90 for @{'lq2}. -interpretation "lq2" 'lq2 = (list bar). - -notation "'len'" with precedence 90 for @{'len}. -interpretation "len" 'len = (length _). -notation < "'len' \nbsp term 90 l" with precedence 69 for @{'len_appl $l}. -interpretation "len appl" 'len_appl l = (length _ l). - -(* a local letin makes russell fail *) -definition cb0h ≝ (λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l)). - -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. - -let rec sum_bases (l:list bar) (i:nat)on i ≝ - match i with - [ O ⇒ OQ - | S m ⇒ - match l with - [ nil ⇒ sum_bases l m + Qpos one - | cons x tl ⇒ sum_bases tl m + Qpos (\fst x)]]. - -definition eject1 ≝ - λP.λp:∃x:nat × ℚ.P x.match p with [ex_introT p _ ⇒ p]. -coercion eject1. -definition inject1 ≝ λP.λp:nat × ℚ.λh:P p. ex_introT ? P p h. -coercion inject1 with 0 1 nocomposites. - -definition value : - ∀f:q_f.∀i:ℚ.∃p:nat × ℚ. - match q_cmp i (start f) with - [ q_lt _ ⇒ \snd p = OQ - | _ ⇒ - And3 - (sum_bases (bars f) (\fst p) ≤ i - start f) - (i - start f < sum_bases (bars f) (S (\fst p))) - (\snd p = \snd (nth (bars f) ▭ (\fst p)))]. -intros; -alias symbol "pi2" = "pair pi2". -alias symbol "pi1" = "pair pi1". -letin value ≝ ( - let rec value (p: ℚ) (l : list bar) on l ≝ - match l with - [ nil ⇒ 〈O,OQ〉 - | cons x tl ⇒ - match q_cmp p (Qpos (\fst x)) with - [ q_lt _ ⇒ 〈O, \snd x〉 - | _ ⇒ - let rc ≝ value (p - Qpos (\fst x)) tl in - 〈S (\fst rc),\snd rc〉]] - in value : - ∀acc,l.∃p:nat × ℚ. OQ ≤ acc → - And3 - (sum_bases l (\fst p) ≤ acc) - (acc < sum_bases l (S (\fst p))) - (\snd p = \snd (nth l ▭ (\fst p)))); -[5: clearbody value; - cases (q_cmp i (start f)); - [2: exists [apply 〈O,OQ〉] simplify; reflexivity; - |*: cases (value (i - start f) (bars f)) (p Hp); - cases Hp; clear Hp value; - exists[1,3:apply p]; simplify; split; assumption;] -|1,3: intros; split; - [1,4: clear H2; cases (value (q-Qpos (\fst b)) l1); cases H2; clear H2; - simplify; apply q_le_minus; assumption; - |2,5: cases (value (q-Qpos (\fst b)) l1); cases H3; clear H3 H2 value; - change with (q < sum_bases l1 (S (\fst w)) + Qpos (\fst b)); - apply q_lt_plus; assumption; - |*: cases (value (q-Qpos (\fst b)) l1); simplify; cases H3; clear H3 value H2; - assumption;] -|2: clear value H2; simplify; split; - [1: - - -definition same_shape ≝ - λl1,l2:q_f. - ∀input.∃col. - - And3 - (sum_bases (bars l2) j ≤ offset - start l2) - (offset - start l2 ≤ sum_bases (bars l2) (S j)) - (\snd (nth (bars l2)) q0 j) = \snd (nth (bars l1) q0 i). - -…┐─┌┐… -\ldots\boxdl\boxh\boxdr\boxdl\ldots +include "nat_ordered_set.ma". +include "models/q_shift.ma". 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)) - (∀i.\fst (nth (bars (\fst p)) q0 i) = \fst (nth (bars (\snd p)) q0 i)) - (∀i,offset. - sum_bases (bars l1) i ≤ offset - start l1 → - offset - start l1 ≤ sum_bases (bars l1) (S i) → - ∃j. - And3 - (sum_bases (bars (\fst p)) j ≤ offset - start (\fst p)) - (offset - start (\fst p) ≤ sum_bases (bars (\fst p)) (S j)) - (\snd (nth (bars (\fst p)) q0 j) = \snd (nth (bars l1) q0 i)) ∧ - And3 - (sum_bases (bars (\snd p)) j ≤ offset - start (\snd p)) - (offset - start (\snd p) ≤ sum_bases (bars (\snd p)) (S j)) - (\snd (nth (bars (\snd p)) q0 j) = \snd (nth (bars l2) q0 i))). + (same_bases (bars (\fst p)) (bars (\snd p))) + (same_values l1 (\fst p)) + (same_values l2 (\snd p)). definition rebase_spec_simpl ≝ - λl1,l2:list (ℚ × ℚ).λp:(list (ℚ × ℚ)) × (list (ℚ × ℚ)). - len ( (\fst p)) = len ( (\snd p)) ∧ - (∀i. - \fst (nth ( (\fst p)) q0 i) = - \fst (nth ( (\snd p)) q0 i)) ∧ - ∀i,offset. - sum_bases ( l1) i ≤ offset ∧ - offset ≤ sum_bases ( l1) (S i) - → - ∃j. - sum_bases ( (\fst p)) j ≤ offset ∧ - offset ≤ sum_bases ((\fst p)) (S j) ∧ - \snd (nth ( (\fst p)) q0 j) = - \snd (nth ( l1) q0 i). + λstart.λl1,l2:list bar.λp:(list bar) × (list bar). + And3 + (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))). + +(* a local letin makes russell fail *) +definition cb0h : list bar → list bar ≝ + λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l). definition eject ≝ λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p]. 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; letin spec ≝ ( - λl1,l2:list (ℚ × ℚ).λm:nat.λz:(list (ℚ × ℚ)) × (list (ℚ × ℚ)). - len l1 + len l2 < m → rebase_spec_simpl l1 l2 z); + λs.λl1,l2.λm.λz.len l1 + len l2 < m → rebase_spec_simpl s l1 l2 z); +alias symbol "pi1" (instance 34) = "exT \fst". +alias symbol "pi1" (instance 21) = "exT \fst". letin aux ≝ ( -let rec aux (l1,l2:list (ℚ × ℚ)) (n:nat) on n : (list (ℚ × ℚ)) × (list (ℚ × ℚ)) ≝ +let rec aux (l1,l2:list bar) (n:nat) on n : (list bar) × (list bar) ≝ match n with [ O ⇒ 〈 nil ? , nil ? 〉 | S m ⇒ @@ -220,43 +61,133 @@ match n with match l2 with [ nil ⇒ 〈l1, cb0h l1〉 | cons he2 tl2 ⇒ - let base1 ≝ (\fst he1) in - let base2 ≝ (\fst he2) in + let base1 ≝ Qpos (\fst he1) in + let base2 ≝ Qpos (\fst he2) in let height1 ≝ (\snd he1) in let height2 ≝ (\snd he2) in match q_cmp base1 base2 with [ q_eq _ ⇒ let rc ≝ aux tl1 tl2 m in 〈he1 :: \fst rc,he2 :: \snd rc〉 - | q_lt _ ⇒ + | q_lt Hp ⇒ let rest ≝ base2 - base1 in - let rc ≝ aux tl1 (〈rest,height2〉 :: tl2) m in - 〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉 - | q_gt _ ⇒ + let rc ≝ aux tl1 (〈\fst (unpos rest ?),height2〉 :: tl2) m in + 〈〈\fst he1,height1〉 :: \fst rc,〈\fst he1,height2〉 :: \snd rc〉 + | q_gt Hp ⇒ let rest ≝ base1 - base2 in - let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in - 〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉 + let rc ≝ aux (〈\fst (unpos rest ?),height1〉 :: tl1) tl2 m in + 〈〈\fst he2,height1〉 :: \fst rc,〈\fst he2,height2〉 :: \snd rc〉 ]]]] -in aux : ∀l1,l2,m.∃z.spec l1 l2 m z); unfold spec; -[7: clearbody aux; unfold spec in aux; clear spec; +in aux : ∀l1,l2,m.∃z.∀s.spec s l1 l2 m z); unfold spec; +[9: clearbody aux; unfold spec in aux; clear spec; cases (q_cmp s1 s2); [1: cases (aux l1 l2 (S (len l1 + len l2))); - cases (H1 (le_n ?)); clear H1; - exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s2 (\snd w)〉] repeat split; - [1: cases H2; assumption; + cases (H1 s1 (le_n ?)); clear H1; + exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s2 (\snd w)〉] split; + [1,2: assumption; + |3: intro; apply (H3 input); + |4: intro; rewrite > H in H4; + 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] + cases (aux l1 l2' (S (len l1 + len l2'))); + cases (H1 s1 (le_n ?)); clear H1 aux; + exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s1 (\snd w)〉] split; + [1: reflexivity |2: assumption; - |3: cases H2; assumption; - |4: intros; cases (H3 i (offset - s1)); - [2: - - -|1,2: simplify; generalize in ⊢ (? ? (? (? ? (? ? ? (? ? %)))) (? (? ? (? ? ? (? ? %))))); intro X; - cases X (rc OK); clear X; simplify; apply eq_f; assumption; -|3: cases (aux l4 l5 n1) (rc OK); simplify; apply eq_f; assumption; -|4,5: simplify; unfold cb0h; rewrite > len_mk_list; reflexivity; -|6: reflexivity] -clearbody aux; unfold spec in aux; clear spec; + |3: assumption; + |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; +|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))); + -qed. \ No newline at end of file +definition excess ≝ + λf,g.∃i.\snd (\fst (value f i)) < \snd (\fst (value g i)). +