(* *)
(**************************************************************************)
+include "russell_support.ma".
include "models/q_bars.ma".
-alias symbol "pi2" = "pair pi2".
-alias symbol "pi1" = "pair pi1".
definition rebase_spec ≝
- ∀l1,l2:q_f.∃p:q_f × q_f.
- And4
- (start (\fst p) = start (\snd p))
+ λl1,l2:q_f.λp:q_f × q_f.
+ And3
(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 (\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))).
+definition same_values_simpl ≝
+ λl1,l2.∀p1,p2,p3,p4,p5,p6.same_values (mk_q_f l1 p1 p2 p3) (mk_q_f l2 p4 p5 p6).
-(* a local letin makes russell fail *)
-definition cb0h : list bar → list bar ≝
- λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l).
+alias symbol "pi2" = "pair pi2".
+alias symbol "pi1" = "pair pi1".
+definition rebase_spec_aux ≝
+ λl1,l2:list bar.λp:(list bar) × (list bar).
+ sorted q2_lt l1 → sorted q2_lt l2 →
+ (l1 ≠ [] → \snd (\nth l1 ▭ (pred (\len l1))) = 〈OQ,OQ〉) →
+ (l2 ≠ [] → \snd (\nth l2 ▭ (pred (\len l2))) = 〈OQ,OQ〉) →
+ And4
+ (nth_base l1 O = nth_base (\fst p) O ∨
+ nth_base l2 O = nth_base (\fst p) O)
+ (sorted q2_lt (\fst p) ∧ sorted q2_lt (\snd p))
+ ((l1 ≠ [] → \snd (\nth (\fst p) ▭ (pred (\len (\fst p)))) = 〈OQ,OQ〉) ∧
+ (l2 ≠ [] → \snd (\nth (\snd p) ▭ (pred (\len (\snd p)))) = 〈OQ,OQ〉))
+ (And3
+ (same_bases (\fst p) (\snd p))
+ (same_values_simpl l1 (\fst p))
+ (same_values_simpl l2 (\snd p))).
definition eject ≝
λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
axiom devil : False.
-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".
+definition copy ≝
+ λl:list bar.make_list ? (λn.〈nth_base l (n - \len l),〈OQ,OQ〉〉) (\len l).
+
+lemma copy_rebases:
+ ∀l1.rebase_spec_aux l1 [] 〈l1, copy l1〉.
+intros; cases l1; intros 4;
+[1: split; [left; reflexivity]; split; try assumption; unfold; intros;
+ unfold same_values; intros; reflexivity;
+|2: rewrite > H2; [2: intro X; destruct X] clear H2 H3;
+ split; [left; reflexivity] split;
+ unfold same_values_simpl; unfold same_values; intros; try reflexivity;
+ try assumption; [4: normalize in p2; destruct p2|2: cases H2; reflexivity;]
+ simplify; clear H1;
+ [1: elim (\len l) in H; simplify; [apply (sorted_one q2_lt);]
+
+
+
+
+definition rebase: ∀l1,l2:q_f.∃p:q_f × q_f.rebase_spec l1 l2 p.
+intros 2 (f1 f2); cases f1 (b1 Hs1 Hb1 He1); cases f2 (b2 Hs2 Hb2 He2); clear f1 f2;
+alias symbol "plus" = "natural plus".
+alias symbol "pi2" = "pair pi2".
+alias symbol "pi1" = "pair pi1".
+alias symbol "minus" = "Q minus".
letin aux ≝ (
-let rec aux (l1,l2:list bar) (n:nat) on n : (list bar) × (list bar) ≝
+let rec aux (l1,l2:list bar) (n : nat) on n : (list bar) × (list bar) ≝
match n with
-[ O ⇒ 〈 nil ? , nil ? 〉
-| S m ⇒
+[ O ⇒ 〈[], []〉
+| S m ⇒
match l1 with
- [ nil ⇒ 〈cb0h l2, l2〉
+ [ nil ⇒ 〈copy l2, l2〉
| cons he1 tl1 ⇒
match l2 with
- [ nil ⇒ 〈l1, cb0h l1〉
+ [ nil ⇒ 〈l1, copy l1〉
| cons he2 tl2 ⇒
- let base1 ≝ Qpos (\fst he1) in
- let base2 ≝ Qpos (\fst he2) in
- let height1 ≝ (\snd he1) in
- let height2 ≝ (\snd he2) in
+ let base1 ≝ \fst he1 in
+ let base2 ≝ \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 Hp ⇒
- let rest ≝ base2 - base1 in
- let rc ≝ aux tl1 (〈\fst (unpos rest ?),height2〉 :: tl2) m in
- 〈〈\fst he1,height1〉 :: \fst rc,〈\fst he1,height2〉 :: \snd rc〉
+ [ q_leq Hp1 ⇒
+ match q_cmp base2 base1 with
+ [ q_leq Hp2 ⇒
+ let rc ≝ aux tl1 tl2 m in
+ 〈he1 :: \fst rc,he2 :: \snd rc〉
+ | q_gt 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 Hp ⇒
let rest ≝ base1 - base2 in
- 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.∀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) 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: 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;
+ let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
+ 〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉]]]]
+in aux : ∀l1,l2,m.∃z.m = \len l1 + \len l2 → rebase_spec_aux l1 l2 z);
+[7: clearbody aux; cases (aux b1 b2 (\len b1 + \len b2)) (w Hw); clear aux;
+ cases (Hw (refl_eq ??) Hs1 Hs2 (λ_.He1) (λ_.He2)); clear Hw; cases H1; cases H2; cases H3; clear H3 H1 H2;
+ exists [constructor 1;constructor 1;[apply (\fst w)|5:apply (\snd w)]] try assumption;
+ [1,3: apply hide; cases H (X X); try rewrite < (H8 O); try rewrite < X; assumption
+ |2,4: apply hide;[apply H6|apply H7]intro X;[rewrite > X in Hb1|rewrite > X in Hb2]
+ normalize in Hb1 Hb2; [destruct Hb1|destruct Hb2]]
+ unfold; unfold same_values; simplify in ⊢ (? (? % %) ? ?);
+ simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉);
+ split; [assumption; |apply H9;|apply H10]
+|6: intro ABS; unfold; intros 4; clear H1 H2;
+ cases l in ABS H3; intros 1; [2: simplify in H1; destruct H1]
+ cases l1 in H4 H1; intros; [2: simplify in H2; destruct H2]
+ split; [left;reflexivity|split; apply (sorted_nil q2_lt);|split; assumption;]
+ split; unfold; intros; unfold same_values; intros; reflexivity;
+|5: unfold rebase_spec_aux; intros; cases l1 in H2 H4 H6; intros; [ simplify in H2; destruct H2;]
+ lapply H6 as H7; [2: intro X; destruct X] clear H6 H5;
+ rewrite > H7; split; [right; simplify;
+
+ split; [left;reflexivity]
+ split;
+
+,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;