X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fmodels%2Fq_function.ma;h=e2187b51014799aed94ac412bef1ac57f24056d9;hb=d03c932e859d59c0ae381f941b4003d744b6b106;hp=8f0f472a9a48b5658595f26da4a598234afa09b4;hpb=ca41435a6021292ccba239aa173651c0be705b45;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 8f0f472a9..e2187b510 100644 --- a/helm/software/matita/contribs/dama/dama/models/q_function.ma +++ b/helm/software/matita/contribs/dama/dama/models/q_function.ma @@ -12,128 +12,124 @@ (* *) (**************************************************************************) -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. - -record q_f : Type ≝ { - start : ℚ; - bars: list (ℚ × ℚ) (* base, height *) -}. - -axiom qp : ℚ → ℚ → ℚ. - -interpretation "Q plus" 'plus x y = (qp x y). - -axiom qm : ℚ → ℚ → ℚ. - -interpretation "Q minus" 'minus x y = (qm x y). - -axiom qlt : ℚ → ℚ → CProp. - -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" 'le x y = (qle x y). - -notation "'nth'" with precedence 90 for @{'nth}. -notation < "'nth' \nbsp l \nbsp d \nbsp i" with precedence 71 -for @{'nth_appl $l $d $i}. -interpretation "list nth" 'nth = (cic:/matita/list/list/nth.con _). -interpretation "list nth" 'nth_appl l d i = (cic:/matita/list/list/nth.con _ l d i). - -notation < "\rationals \sup 2" non associative with precedence 90 for @{'q2}. -interpretation "Q x Q" 'q2 = (Prod Q Q). - -let rec make_list (A:Type) (def:nat→A) (n:nat) on n ≝ - match n with - [ O ⇒ [] - | S m ⇒ def m :: make_list A def m]. - -notation "'mk_list'" with precedence 90 for @{'mk_list}. -interpretation "'mk_list'" 'mk_list = (make_list _). -notation < "'mk_list' \nbsp f \nbsp n" -with precedence 71 for @{'mk_list_appl $f $n}. -interpretation "'mk_list' appl" 'mk_list_appl f n = (make_list _ f n). - -definition q0 : ℚ × ℚ ≝ 〈OQ,OQ〉. -notation < "0 \sub \rationals" with precedence 90 for @{'q0}. -interpretation "q0" 'q0 = q0. - -notation < "[ \rationals \sup 2]" with precedence 90 for @{'lq2}. -interpretation "lq2" 'lq2 = (list (Prod Q Q)). -notation < "[ \rationals \sup 2] \sup 2" with precedence 90 for @{'lq22}. -interpretation "lq22" 'lq22 = (Prod (list (Prod Q Q)) (list (Prod Q Q))). - - -notation "'len'" with precedence 90 for @{'len}. -interpretation "len" 'len = length. -notation < "'len' \nbsp l" with precedence 70 for @{'len_appl $l}. -interpretation "len appl" 'len_appl l = (length _ l). - -definition eject ≝ - λP.λp:∃x:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).P x.match p with [ex_introT p _ ⇒ p]. -coercion cic:/matita/dama/models/q_function/eject.con. -definition inject ≝ - λP.λp:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).λh:P p. ex_introT ? P p h. -(*coercion inject with 0 1 nocomposites.*) -coercion cic:/matita/dama/models/q_function/inject.con 0 1 nocomposites. - -definition cb0h ≝ (λl.mk_list (λi.〈\fst (nth l q0 i),OQ〉) (length ? l)). +include "nat_ordered_set.ma". +include "models/q_shift.ma". alias symbol "pi2" = "pair pi2". alias symbol "pi1" = "pair pi1". -definition rebase: - q_f → q_f → - ∃p:q_f × q_f.∀i. - \fst (nth (bars (\fst p)) q0 i) = - \fst (nth (bars (\snd p)) q0 i). -intros (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2; -letin spec ≝ (λl1,l2:list (ℚ × ℚ).λm:nat.λz:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).True); +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_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_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. + +definition rebase: rebase_spec. +intros 2 (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2; +letin spec ≝ ( + λ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 ⇒ 〈[],[]〉 -| S m ⇒ +[ O ⇒ 〈 nil ? , nil ? 〉 +| S m ⇒ match l1 with [ nil ⇒ 〈cb0h l2, l2〉 | cons he1 tl1 ⇒ 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 _ ⇒ + [ q_eq _ ⇒ let rc ≝ aux tl1 tl2 m in - 〈he1 :: \fst rc,he2 :: \snd rc〉 - | q_lt _ ⇒ + 〈he1 :: \fst rc,he2 :: \snd rc〉 + | 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); - -cases (q_cmp s1 s2); -[1: apply (mk_q_f s1); -|2: apply (mk_q_f s1); cases l2; - [1: letin l2' ≝ ( -[1: (* offset: the smallest one *) - cases +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 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); 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: assumption; + |4: intro; rewrite > (initial_shift_same_values (mk_q_f s2 l2) s1 H input); + rewrite < (H4 input); 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; rewrite > (initial_shift_same_values (mk_q_f s1 l1) s2 H input); + rewrite < (H3 input); 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 *) +] +qed.