]> matita.cs.unibo.it Git - helm.git/commitdiff
...
authorEnrico Tassi <enrico.tassi@inria.fr>
Thu, 4 Sep 2008 13:36:50 +0000 (13:36 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Thu, 4 Sep 2008 13:36:50 +0000 (13:36 +0000)
helm/software/matita/contribs/dama/dama/models/q_function.ma
helm/software/matita/library/formal_topology/concrete_spaces.ma

index 47dccaf5d486402bc91eab72c6b528d82affbbac..33898007aa36fcb02675461889834e4a984c8327 100644 (file)
@@ -114,8 +114,221 @@ definition rebase_spec ≝
     (same_bases (bars (\fst p)) (bars (\snd p)))
     (same_values l1 (\fst p)) 
     (same_values l2 (\snd p)).
+
+definition last ≝
+ λT:Type.λd.λl:list T. \nth l d (pred (\len l)).
+
+notation > "\last" non associative with precedence 90 for @{'last}.
+notation < "\last d l" non associative with precedence 70 for @{'last2 $d $l}.
+interpretation "list last" 'last = (last _). 
+interpretation "list last applied" 'last2 d l = (last _ d l).
+
+definition head ≝
+ λT:Type.λd.λl:list T.\nth l d O.
+
+notation > "\hd" non associative with precedence 90 for @{'hd}.
+notation < "\hd d l" non associative with precedence 70 for @{'hd2 $d $l}.
+interpretation "list head" 'hd = (head _).
+interpretation "list head applied" 'hd2 d l = (head _ d l).
+
+alias symbol "lt" = "bar lt".
+lemma inversion_sorted:
+  ∀a,l. sorted q2_lt (a::l) → a < \hd ▭ l ∨ l = [].
+intros 2; elim l; [right;reflexivity] left; inversion H1; intros;
+[1,2:destruct H2| destruct H5; assumption]
+qed.
+
+lemma inversion_sorted2:
+  ∀a,b,l. sorted q2_lt (a::b::l) → a < b.
+intros; inversion H; intros; [1,2:destruct H1] destruct H4; assumption; 
+qed.
+
+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).
+
+inductive rebase_cases : list bar → list bar → (list bar) × (list bar) → Prop ≝
+| rb_fst_full  : ∀b,h1,h2,xs,ys,r1,r2. 〈b,h1〉 < \hd ▭ ys → 
+   \snd(\last ▭ (〈b,h1〉::xs)) = \snd(\last ▭ (〈b,h1〉::r1)) →
+   \snd(\last ▭ (〈b,h1〉::xs)) = \snd(\last ▭ (〈b,h2〉::r2)) →
+   rebase_cases (〈b,h1〉::xs) ys 〈〈b,h1〉::r1,〈b,h2〉::r2〉
+| rb_snd_full  : ∀b,h1,h2,xs,ys,r1,r2. 〈b,h1〉 < \hd ▭ xs →
+   \snd(\last ▭ (〈b,h1〉::ys)) = \snd(\last ▭ (〈b,h1〉::r2)) →
+   \snd(\last ▭ (〈b,h1〉::ys)) = \snd(\last ▭ (〈b,h2〉::r1)) →
+   rebase_cases xs (〈b,h1〉::ys) 〈〈b,h2〉::r1,〈b,h1〉::r2〉  
+| rb_all_full  : ∀b,h1,h2,h3,h4,xs,ys,r1,r2.
+   \snd(\last ▭ (〈b,h1〉::xs)) = \snd(\last ▭ (〈b,h3〉::r1)) →
+   \snd(\last ▭ (〈b,h2〉::ys)) = \snd(\last ▭ (〈b,h4〉::r2)) →
+   rebase_cases (〈b,h1〉::xs) (〈b,h2〉::ys) 〈〈b,h3〉::r1,〈b,h4〉::r2〉 
+| rb_all_empty : rebase_cases [] [] 〈[],[]〉.
+
+alias symbol "pi2" = "pair pi2".
+alias symbol "pi1" = "pair pi1".
+alias symbol "leq" = "natural 'less or equal to'".
+inductive rebase_spec_aux_p (l1, l2:list bar) (p:(list bar) × (list bar)) : Prop ≝
+| prove_rebase_spec_aux:
+   rebase_cases l1 l2 p →
+   (sorted q2_lt (\fst p)) →
+   (sorted q2_lt (\snd p)) → True → True → (*
+   ((l1 ≠ [] ∧ \snd (\last ▭ (\fst p)) = 〈OQ,OQ〉) ∨
+    (l1 = [] ∧ ∀i.\snd (\nth (\fst p) ▭ i) = 〈OQ,OQ〉)) → 
+   ((l2 ≠ [] ∧ \snd (\last ▭ (\snd p)) = 〈OQ,OQ〉) ∨
+    (l2 = [] ∧ ∀i.\snd (\nth (\snd p) ▭ i) = 〈OQ,OQ〉)) →*)
+   (same_bases (\fst p) (\snd p)) →
+   (same_values_simpl l1 (\fst p)) → 
+   (same_values_simpl l2 (\snd p)) →  
+   (*\len (\fst p) ≤ \len l1 + \len l2 → *)
+   (*\len (\fst p) = \len (\snd p) → *)
+   rebase_spec_aux_p l1 l2 p.   
+
+lemma sort_q2lt_same_base:
+  ∀b,h1,h2,l. sorted q2_lt (〈b,h1〉::l) → sorted q2_lt (〈b,h2〉::l).
+intros; cases (inversion_sorted ?? H); [2: rewrite > H1; apply (sorted_one q2_lt)]
+lapply (sorted_tail q2_lt ?? H) as K; clear H; cases l in H1 K; simplify; intros;
+[apply (sorted_one q2_lt);|apply (sorted_cons q2_lt);[2: assumption] apply H]
+qed.
+
+lemma aux_preserves_sorting:
+ ∀b,b3,l2,l3,w. rebase_cases l2 l3 w → 
+  sorted q2_lt (b::l2) → sorted q2_lt (b3::l3) → \fst b3 = \fst b →
+  sorted q2_lt (\fst w) → sorted q2_lt (\snd w) → same_bases (\fst w) (\snd w) → 
+  sorted q2_lt (b :: \fst w).
+intros 6; cases H; simplify; intros;
+[ apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H4);   
+| apply (sorted_cons q2_lt); [2:apply (sort_q2lt_same_base ???? H7);]
+  whd; rewrite < H6; apply (inversion_sorted2 ??? H5);
+| apply (sorted_cons q2_lt); [2: assumption] apply (inversion_sorted2 ??? H3);
+| apply (sorted_one q2_lt);]
+qed.   
+
+lemma aux_preserves_sorting2:
+ ∀b,b3,l2,l3,w. rebase_cases l2 l3 w → 
+  sorted q2_lt (b::l2) → sorted q2_lt (b3::l3) → \fst b3 = \fst b →
+  sorted q2_lt (\fst w) → sorted q2_lt (\snd w) → same_bases (\fst w) (\snd w) → 
+  sorted q2_lt (b :: \snd w).
+intros 6; cases H; simplify; intros;
+[ apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H4);   
+| apply (sorted_cons q2_lt); [2:assumption] 
+  whd; rewrite < H6; apply (inversion_sorted2 ??? H5);
+| apply (sorted_cons q2_lt); [2: assumption] apply (inversion_sorted2 ??? H3);
+| apply (sorted_one q2_lt);]
+qed.   
+
+definition rebase_spec_aux ≝ 
+ λl1,l2:list bar.λp:(list bar) × (list bar).
+ sorted q2_lt l1 → (\snd (\last ▭ l1) = 〈OQ,OQ〉) →
+ sorted q2_lt l2 → (\snd (\last ▭ l2) = 〈OQ,OQ〉) → 
+   rebase_spec_aux_p l1 l2 p.
+    
+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: ∀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 "leq" = "natural 'less or equal to'".
+alias symbol "minus" = "Q minus".
+letin aux ≝ ( 
+let rec aux (l1,l2:list bar) (n : nat) on n : (list bar) × (list bar) ≝
+match n with
+[ O ⇒ 〈[], []〉
+| S m ⇒
+  match l1 with
+  [ nil ⇒ 〈copy l2, l2〉
+  | cons he1 tl1 ⇒
+     match l2 with
+     [ nil ⇒ 〈l1, copy l1〉
+     | cons he2 tl2 ⇒  
+         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_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 (〈rest,height1〉 :: tl1) tl2 m in
+             〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉]]]]
+in aux : ∀l1,l2,m.∃z.\len l1 + \len l2 ≤ m → rebase_spec_aux l1 l2 z);
+[7: clearbody aux; cases (aux b1 b2 (\len b1 + \len b2)) (res Hres);
+    exists; [split; constructor 1; [apply (\fst res)|5:apply (\snd res)]]
+    [1,4: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); assumption;
+    |2,5: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); lapply (H5 O) as K;
+          clear H1 H2 H3 H4 H5 H6 H7 Hres aux; unfold nth_base;
+          cases H in K He1 He2 Hb1 Hb2; simplify; intros; assumption;
+    |3,6: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); 
+          cases H in He1 He2; simplify; intros;
+          [4,8: assumption;
+          |1,6,7: rewrite < H9; assumption;
+          |2,3,5: rewrite < H10; [2: symmetry] assumption;]]
+    split; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); unfold same_values; intros; 
+    [1: assumption
+    |2: simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉); apply H6;
+    |3: simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉); apply H7]
+|3: cut (\fst b3 = \fst b) as E; [2: apply q_le_to_le_to_eq; assumption]
+    clear H6 H5 H4 H3 He2 Hb2 Hs2 b2 He1 Hb1 Hs1 b1; cases (aux l2 l3 n1); 
+    clear aux; intro K; simplify in K; rewrite <plus_n_Sm in K; 
+    lapply le_S_S_to_le to K as W; lapply lt_to_le to W as R; 
+    simplify in match (? ≪w,H3≫); intros; cases (H3 R); clear H3 R W K;
+      [2,4: apply (sorted_tail q2_lt);[apply b|3:apply b3]assumption;
+      |3: cases l2 in H5; simplify; intros; try reflexivity; assumption; 
+      |5: cases l3 in H7; simplify; intros; try reflexivity; assumption;]
+    constructor 1; simplify in match (\fst 〈?,?〉); simplify in match (\snd 〈?,?〉);
+    [1: cases b in E H5 H7 H11 H14; cases b3; intros (E H5 H7 H11 H14); simplify in E; 
+        destruct E; constructor 3;
+        [ cases H8 in H5 H7; intros; [1,3:assumption] simplify;
+          [ rewrite > H16; rewrite < H7; symmetry; apply H17; | reflexivity]
+        | cases H8 in H5 H7; simplify; intros; [2,3: assumption] 
+          [ rewrite < H7; rewrite > H16; apply H17; | reflexivity]]
+    |2: apply (aux_preserves_sorting ? b3 ??? H8); assumption;
+    |3: apply (aux_preserves_sorting2 ? b3 ??? H8); try assumption;
+        try reflexivity; cases (inversion_sorted ?? H4);[2:rewrite >H3; apply (sorted_one q2_lt);]
+        cases l2 in H3 H4; intros; [apply (sorted_one q2_lt)]
+        apply (sorted_cons q2_lt);[2:apply (sorted_tail q2_lt ?? H3);] whd;
+        rewrite > E; assumption;
+    |4: apply I 
+    |5: apply I
+    |6: intro; elim i; intros; simplify; solve [symmetry;assumption|apply H13]
+    |7: unfold; intros; apply H14;
+    |8: 
+    
+    clear H15 H14 H11 H12 H7 H5; cases H8; clear H8; cases H3; clear H3;
+        [1: apply (move_head ?? H4 H5 ?? H9); symmetry; assumption;
+        |2: apply (move_head ??? H5 ?? H9); [apply (soted_q2lt_rewrite_hd ??? E  H6)]
+            symmetry; rewrite > (H13 O); assumption;
+        |3: apply (soted_q2lt_rewrite_hd ??? E); apply (move_head ?? H6 H7); [symmetry;] assumption;  
+        |4: rewrite > H8; apply (sorted_one q2_lt);]
+    
     
+     cases l2 in H4 H8 H16; cases l3 in H6;
+        intros; cases H5; clear H5; cases H7; clear H7;
+        [1,2: cases (q_lt_corefl ? H5);
+        |3: rewrite >(?:\fst w = []); [apply (sorted_one q2_lt)]
+            simplify in H6; lapply (le_n_O_to_eq ? H6) as K;
+            cases (\fst w) in K; simplify; intro X [reflexivity|destruct X]
+        |4:
+        
+         rewrite >H8 in H5; cases (\fst w) in H9 H5; intros [apply (sorted_one q2_lt)]
+                
 
+
+
+
+
+
+
+
+    
 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).
 
@@ -163,6 +376,7 @@ intros; elim l1; intros 4;
     |2: intro; symmetry; apply (copy_same_bases (a::l));]]
 qed.
 
+
 definition eject ≝
   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
 coercion eject.
index 568f3c0528488ec22c45a10378ebb0e5861f8435..0fc4d365b6398e30ef77a266f8e85acab96513fe 100644 (file)
@@ -66,9 +66,21 @@ qed.
 
 interpretation "fintersectsS" 'fintersects U V = (fintersectsS _ U V).
 
-(*
+
 definition relS: ∀o: basic_pair. concr o → Ω \sup (form o) → CProp.
- apply (λo:basic_pair.λx:concr o.λS: Ω \sup (form o).∃y: form o.y ∈ S ∧ x ⊩ y);
+
+ apply (λo:basic_pair.λx:concr o.λS: Ω \sup (form o).∃y: form o.y ∈ S ∧ 
+   (* OK: FunClass_2_OF_binary_relation (concr o) (form o) (rel o) x y *)
+   ?);
+   change in x with (carr1 (setoid1_of_setoid (concr o)));
+   apply (FunClass_2_OF_binary_relation ?? (rel ?) x y); 
+x ⊩ y);
+ rel (concr o) o -> binary_relation ...
+ rel ? = seotid1_OF_setoid ?
+ carr rel ? = Type_OF_objs1 (concr o) -> 
+         carr (setoid1_of_REL (concr o))
 
 interpretation "basic pair relation for subsets" 'Vdash2 x y = (relS _ x y).
 interpretation "basic pair relation for subsets (non applied)" 'Vdash = (relS _).