]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/iteration2.ma
Restructuring.
[helm.git] / helm / software / matita / library / nat / iteration2.ma
index e00bb442089752014f6ed9c630bf7ad445b867d9..c5236ff4a6f8678907c1143bf47ffd18d5c1dcf5 100644 (file)
@@ -191,7 +191,165 @@ apply (eq_iter_p_gen_gh nat O plus ? ? ? g h h1 n n1 p1 p2)
 ]
 qed.
 
+theorem eq_sigma_p_pred: 
+\forall n,p,g. p O = true \to
+sigma_p (S n) (\lambda i.p (pred i)) (\lambda i.g(pred i)) = 
+plus (sigma_p n p g) (g O).
+intros.
+unfold sigma_p.
+apply eq_iter_p_gen_pred
+  [assumption
+  |apply symmetricIntPlus
+  |apply associative_plus
+  |intros.apply sym_eq.apply plus_n_O
+  ]
+qed.
+
+(* monotonicity *)
+theorem le_sigma_p: 
+\forall n:nat. \forall p:nat \to bool. \forall g1,g2:nat \to nat.
+(\forall i. i < n \to p i = true \to g1 i \le g2 i ) \to 
+sigma_p n p g1 \le sigma_p n p g2.
+intros.
+generalize in match H.
+elim n
+  [apply le_n.
+  |apply (bool_elim ? (p n1));intros
+    [rewrite > true_to_sigma_p_Sn
+      [rewrite > true_to_sigma_p_Sn in ⊢ (? ? %)
+        [apply le_plus
+          [apply H2[apply le_n|assumption]
+          |apply H1.
+           intros.
+           apply H2[apply le_S.assumption|assumption]
+          ]
+        |assumption
+        ]
+      |assumption
+      ]
+    |rewrite > false_to_sigma_p_Sn
+      [rewrite > false_to_sigma_p_Sn in ⊢ (? ? %)
+        [apply H1.
+         intros.
+         apply H2[apply le_S.assumption|assumption]
+        |assumption
+        ]
+      |assumption
+      ]
+    ]
+  ]
+qed.
 
+theorem lt_sigma_p: 
+\forall n:nat. \forall p:nat \to bool. \forall g1,g2:nat \to nat.
+(\forall i. i < n \to p i = true \to g1 i \le g2 i ) \to
+(\exists i. i < n \and (p i = true) \and (g1 i < g2 i)) \to
+sigma_p n p g1 < sigma_p n p g2.
+intros 4.
+elim n
+  [elim H1.clear H1.
+   elim H2.clear H2.
+   elim H1.clear H1.
+   apply False_ind.
+   apply (lt_to_not_le ? ? H2).
+   apply le_O_n
+  |apply (bool_elim ? (p n1));intros
+    [apply (bool_elim ? (leb (S (g1 n1)) (g2 n1)));intros
+      [rewrite > true_to_sigma_p_Sn
+        [rewrite > true_to_sigma_p_Sn in ⊢ (? ? %)
+          [change with 
+           (S (g1 n1)+sigma_p n1 p g1 \le g2 n1+sigma_p n1 p g2).
+           apply le_plus
+            [apply leb_true_to_le.assumption
+            |apply le_sigma_p.intros.
+             apply H1
+              [apply lt_to_le.apply le_S_S.assumption
+              |assumption
+              ]
+            ]
+          |assumption
+          ]
+        |assumption
+        ]
+      |rewrite > true_to_sigma_p_Sn
+        [rewrite > true_to_sigma_p_Sn in ⊢ (? ? %)
+          [unfold lt.
+           rewrite > plus_n_Sm.
+           apply le_plus
+            [apply H1
+              [apply le_n
+              |assumption
+              ]
+            |apply H
+              [intros.apply H1
+                [apply lt_to_le.apply le_S_S.assumption
+                |assumption
+                ]
+              |elim H2.clear H2.
+               elim H5.clear H5.
+               elim H2.clear H2.
+               apply (ex_intro ? ? a).
+               split
+                [split
+                  [elim (le_to_or_lt_eq a n1)
+                    [assumption
+                    |absurd (g1 a < g2 a)
+                      [assumption
+                      |apply leb_false_to_not_le.
+                       rewrite > H2.
+                       assumption
+                      ]
+                    |apply le_S_S_to_le.
+                     assumption
+                    ]
+                  |assumption
+                  ]
+                |assumption
+                ]
+              ]
+            ]
+          |assumption
+          ]
+        |assumption
+        ]
+      ]
+    |rewrite > false_to_sigma_p_Sn
+      [rewrite > false_to_sigma_p_Sn in ⊢ (? ? %)
+        [apply H
+          [intros.apply H1
+            [apply lt_to_le.apply le_S_S.assumption
+            |assumption
+            ]
+          |elim H2.clear H2.
+           elim H4.clear H4.
+           elim H2.clear H2.
+           apply (ex_intro ? ? a).
+           split
+            [split
+              [elim (le_to_or_lt_eq a n1)
+                [assumption
+                |apply False_ind.
+                 apply not_eq_true_false.
+                 rewrite < H6.
+                 rewrite < H3.
+                 rewrite < H2. 
+                 reflexivity
+                |apply le_S_S_to_le.
+                 assumption
+                ]
+              |assumption
+              ]
+            |assumption
+            ]
+          ]
+        |assumption
+        ]
+      |assumption
+      ]
+    ]
+  ]
+qed.
+          
 theorem sigma_p_divides: 
 \forall n,m,p:nat.O < n \to prime p \to Not (divides p n) \to 
 \forall g: nat \to nat.
@@ -611,4 +769,36 @@ apply iter_p_gen_knm
   |assumption
   ]
 qed.
-  
\ No newline at end of file
+  
+  
+theorem sigma_p2_eq: 
+\forall g: nat \to nat \to nat.
+\forall h11,h12,h21,h22: nat \to nat \to nat. 
+\forall n1,m1,n2,m2.
+\forall p11,p21:nat \to bool.
+\forall p12,p22:nat \to nat \to bool.
+(\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to 
+p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
+\land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
+\land h11 i j < n1 \land h12 i j < m1) \to
+(\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to 
+p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
+\land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
+\land (h21 i j) < n2 \land (h22 i j) < m2) \to
+sigma_p n1 p11 (\lambda x:nat .sigma_p m1 (p12 x) (\lambda y. g x y)) =
+sigma_p n2 p21 (\lambda x:nat .sigma_p m2 (p22 x) (\lambda y. g (h11 x y) (h12 x y))).
+intros.
+unfold sigma_p.
+unfold sigma_p in \vdash (? ? (? ? ? ? (\lambda x:?.%) ? ?) ?).
+unfold sigma_p in \vdash (? ? ? (? ? ? ? (\lambda x:?.%) ? ?)).
+
+apply(iter_p_gen_2_eq nat O plus ? ? ? g h11 h12 h21 h22 n1 m1 n2 m2 p11 p21 p12 p22)
+[ apply symmetricIntPlus
+| apply associative_plus
+| intro.
+  rewrite < (plus_n_O).
+  reflexivity
+| assumption
+| assumption
+]
+qed.
\ No newline at end of file