]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/Z/dirichlet_product.ma
Reorganization of results.
[helm.git] / matita / library / Z / dirichlet_product.ma
index edc2037618215b7a5f0bb5ddf07d1d142dab9039..6328bc7dca8ca34f3a836d2f7b54af8a4a204ada 100644 (file)
@@ -23,417 +23,6 @@ definition dirichlet_product : (nat \to Z) \to (nat \to Z) \to nat \to Z \def
 sigma_p (S n) 
  (\lambda d.divides_b d n) (\lambda d. (f d)*(g (div n d))).
 
-(* da spostare *)
-
-(* spostati in div_and_mod.ma
-theorem mod_SO: \forall n:nat. mod n (S O) = O.
-intro.
-apply sym_eq.
-apply le_n_O_to_eq.
-apply le_S_S_to_le.
-apply lt_mod_m_m.
-apply le_n.
-qed.
-theorem div_SO: \forall n:nat. div n (S O) = n.
-intro.
-rewrite > (div_mod ? (S O)) in \vdash (? ? ? %)
-  [rewrite > mod_SO.
-   rewrite < plus_n_O.
-   apply times_n_SO
-  |apply le_n
-  ]
-qed.*)
-
-theorem and_true: \forall a,b:bool. 
-andb a b =true \to a =true \land b= true.
-intro.elim a
-  [split
-    [reflexivity|assumption]
-  |apply False_ind.
-   apply not_eq_true_false.
-   apply sym_eq.
-   assumption
-  ]
-qed.
-
-theorem lt_times_plus_times: \forall a,b,n,m:nat. 
-a < n \to b < m \to a*m + b < n*m.
-intros 3.
-apply (nat_case n)
-  [intros.apply False_ind.apply (not_le_Sn_O ? H)
-  |intros.simplify.
-   rewrite < sym_plus.
-   unfold.
-   change with (S b+a*m1 \leq m1+m*m1).
-   apply le_plus
-    [assumption
-    |apply le_times
-      [apply le_S_S_to_le.assumption
-      |apply le_n
-      ]
-    ]
-  ]
-qed.
-
-theorem divides_to_divides_b_true1 : \forall n,m:nat.
-O < m \to n \divides m \to divides_b n m = true.
-intro.
-elim (le_to_or_lt_eq O n (le_O_n n))
-  [apply divides_to_divides_b_true
-    [assumption|assumption]
-  |apply False_ind.
-   rewrite < H in H2.
-   elim H2.
-   simplify in H3.
-   apply (not_le_Sn_O O).
-   rewrite > H3 in H1.
-   assumption
-  ]
-qed.
-(* spostato in primes.ma
-theorem divides_to_div: \forall n,m.divides n m \to m/n*n = m.
-intro.
-elim (le_to_or_lt_eq O n (le_O_n n))
-  [rewrite > plus_n_O.
-   rewrite < (divides_to_mod_O ? ? H H1).
-   apply sym_eq.
-   apply div_mod.
-   assumption
-  |elim H1.
-   generalize in match H2.
-   rewrite < H.
-   simplify.
-   intro.
-   rewrite > H3.
-   reflexivity
-  ]
-qed.*)
-(* spostato in div_and_mod.ma
-theorem le_div: \forall n,m. O < n \to m/n \le m.
-intros.
-rewrite > (div_mod m n) in \vdash (? ? %)
-  [apply (trans_le ? (m/n*n))
-    [rewrite > times_n_SO in \vdash (? % ?).
-     apply le_times
-      [apply le_n|assumption]
-    |apply le_plus_n_r
-    ]
-  |assumption
-  ]
-qed.*)
-  
-theorem sigma_p2_eq: 
-\forall g: nat \to nat \to Z.
-\forall h11,h12,h21,h22: nat \to nat \to nat. 
-\forall n,m.
-\forall p11,p21:nat \to bool.
-\forall p12,p22:nat \to nat \to bool.
-(\forall i,j. i < n \to j < m \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 < n \land h12 i j < m) \to
-(\forall i,j. i < n \to j < m \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) < n \land (h22 i j) < m) \to
-sigma_p n p11 (\lambda x:nat .sigma_p m (p12 x) (\lambda y. g x y)) =
-sigma_p n p21 (\lambda x:nat .sigma_p m (p22 x) (\lambda y. g (h11 x y) (h12 x y))).
-intros.
-rewrite < sigma_p2'.
-rewrite < sigma_p2'.
-apply sym_eq.
-letin h := (\lambda x.(h11 (x/m) (x\mod m))*m + (h12 (x/m) (x\mod m))).
-letin h1 := (\lambda x.(h21 (x/m) (x\mod m))*m + (h22 (x/m) (x\mod m))).
-apply (trans_eq ? ? 
-  (sigma_p (n*m) (\lambda x:nat.p21 (x/m)\land p22 (x/m) (x\mod m))
-  (\lambda x:nat.g ((h x)/m) ((h x)\mod m))))
-  [clear h.clear h1.
-   apply eq_sigma_p1
-    [intros.reflexivity
-    |intros.
-     cut (O < m)
-      [cut (x/m < n)
-        [cut (x \mod m < m)
-          [elim (and_true ? ? H3).
-           elim (H ? ? Hcut1 Hcut2 H4 H5).
-           elim H6.clear H6.
-           elim H8.clear H8.
-           elim H6.clear H6.
-           elim H8.clear H8.
-           apply eq_f2
-            [apply sym_eq.
-             apply div_plus_times.
-             assumption
-            |autobatch
-            ]
-          |apply lt_mod_m_m.
-           assumption
-          ]
-        |apply (lt_times_n_to_lt m)
-          [assumption
-          |apply (le_to_lt_to_lt ? x)
-            [apply (eq_plus_to_le ? ? (x \mod m)).
-             apply div_mod.
-             assumption
-            |assumption
-            ]
-          ]
-        ]
-      |apply not_le_to_lt.unfold.intro.
-       generalize in match H2.
-       apply (le_n_O_elim ? H4).
-       rewrite < times_n_O.
-       apply le_to_not_lt.
-       apply le_O_n  
-      ]      
-    ]
-  |apply (eq_sigma_p_gh ? h h1);intros
-    [cut (O < m)
-      [cut (i/m < n)
-        [cut (i \mod m < m)
-          [elim (and_true ? ? H3).
-           elim (H ? ? Hcut1 Hcut2 H4 H5).
-           elim H6.clear H6.
-           elim H8.clear H8.
-           elim H6.clear H6.
-           elim H8.clear H8.
-           cut ((h11 (i/m) (i\mod m)*m+h12 (i/m) (i\mod m))/m = 
-                 h11 (i/m) (i\mod m))
-            [cut ((h11 (i/m) (i\mod m)*m+h12 (i/m) (i\mod m))\mod m =
-                  h12 (i/m) (i\mod m))
-              [rewrite > Hcut3.
-               rewrite > Hcut4.
-               rewrite > H6.
-               rewrite > H12.
-               reflexivity
-              |apply mod_plus_times. 
-               assumption
-              ]
-            |apply div_plus_times.
-             assumption
-            ]
-          |apply lt_mod_m_m.
-           assumption
-          ]
-        |apply (lt_times_n_to_lt m)
-          [assumption
-          |apply (le_to_lt_to_lt ? i)
-            [apply (eq_plus_to_le ? ? (i \mod m)).
-             apply div_mod.
-             assumption
-            |assumption
-            ]
-          ]
-        ]
-      |apply not_le_to_lt.unfold.intro.
-       generalize in match H2.
-       apply (le_n_O_elim ? H4).
-       rewrite < times_n_O.
-       apply le_to_not_lt.
-       apply le_O_n  
-      ]      
-    |cut (O < m)
-      [cut (i/m < n)
-        [cut (i \mod m < m)
-          [elim (and_true ? ? H3).
-           elim (H ? ? Hcut1 Hcut2 H4 H5).
-           elim H6.clear H6.
-           elim H8.clear H8.
-           elim H6.clear H6.
-           elim H8.clear H8.
-           cut ((h11 (i/m) (i\mod m)*m+h12 (i/m) (i\mod m))/m = 
-                 h11 (i/m) (i\mod m))
-            [cut ((h11 (i/m) (i\mod m)*m+h12 (i/m) (i\mod m))\mod m =
-                  h12 (i/m) (i\mod m))
-              [rewrite > Hcut3.
-               rewrite > Hcut4.
-               rewrite > H10.
-               rewrite > H11.
-               apply sym_eq.
-               apply div_mod.
-               assumption
-              |apply mod_plus_times. 
-               assumption
-              ]
-            |apply div_plus_times.
-             assumption
-            ]
-          |apply lt_mod_m_m.
-           assumption
-          ]
-        |apply (lt_times_n_to_lt m)
-          [assumption
-          |apply (le_to_lt_to_lt ? i)
-            [apply (eq_plus_to_le ? ? (i \mod m)).
-             apply div_mod.
-             assumption
-            |assumption
-            ]
-          ]
-        ]
-      |apply not_le_to_lt.unfold.intro.
-       generalize in match H2.
-       apply (le_n_O_elim ? H4).
-       rewrite < times_n_O.
-       apply le_to_not_lt.
-       apply le_O_n  
-      ]      
-    |cut (O < m)
-      [cut (i/m < n)
-        [cut (i \mod m < m)
-          [elim (and_true ? ? H3).
-           elim (H ? ? Hcut1 Hcut2 H4 H5).
-           elim H6.clear H6.
-           elim H8.clear H8.
-           elim H6.clear H6.
-           elim H8.clear H8.
-           apply lt_times_plus_times
-            [assumption|assumption]
-          |apply lt_mod_m_m.
-           assumption
-          ]
-        |apply (lt_times_n_to_lt m)
-          [assumption
-          |apply (le_to_lt_to_lt ? i)
-            [apply (eq_plus_to_le ? ? (i \mod m)).
-             apply div_mod.
-             assumption
-            |assumption
-            ]
-          ]
-        ]
-      |apply not_le_to_lt.unfold.intro.
-       generalize in match H2.
-       apply (le_n_O_elim ? H4).
-       rewrite < times_n_O.
-       apply le_to_not_lt.
-       apply le_O_n  
-      ]
-    |cut (O < m)
-      [cut (j/m < n)
-        [cut (j \mod m < m)
-          [elim (and_true ? ? H3).
-           elim (H1 ? ? Hcut1 Hcut2 H4 H5).
-           elim H6.clear H6.
-           elim H8.clear H8.
-           elim H6.clear H6.
-           elim H8.clear H8.
-           cut ((h21 (j/m) (j\mod m)*m+h22 (j/m) (j\mod m))/m = 
-                 h21 (j/m) (j\mod m))
-            [cut ((h21 (j/m) (j\mod m)*m+h22 (j/m) (j\mod m))\mod m =
-                  h22 (j/m) (j\mod m))
-              [rewrite > Hcut3.
-               rewrite > Hcut4.
-               rewrite > H6.
-               rewrite > H12.
-               reflexivity
-              |apply mod_plus_times. 
-               assumption
-              ]
-            |apply div_plus_times.
-             assumption
-            ]
-          |apply lt_mod_m_m.
-           assumption
-          ] 
-        |apply (lt_times_n_to_lt m)
-          [assumption
-          |apply (le_to_lt_to_lt ? j)
-            [apply (eq_plus_to_le ? ? (j \mod m)).
-             apply div_mod.
-             assumption
-            |assumption
-            ]
-          ]
-        ]
-      |apply not_le_to_lt.unfold.intro.
-       generalize in match H2.
-       apply (le_n_O_elim ? H4).
-       rewrite < times_n_O.
-       apply le_to_not_lt.
-       apply le_O_n  
-      ] 
-    |cut (O < m)
-      [cut (j/m < n)
-        [cut (j \mod m < m)
-          [elim (and_true ? ? H3).
-           elim (H1 ? ? Hcut1 Hcut2 H4 H5).
-           elim H6.clear H6.
-           elim H8.clear H8.
-           elim H6.clear H6.
-           elim H8.clear H8.
-           cut ((h21 (j/m) (j\mod m)*m+h22 (j/m) (j\mod m))/m = 
-                 h21 (j/m) (j\mod m))
-            [cut ((h21 (j/m) (j\mod m)*m+h22 (j/m) (j\mod m))\mod m =
-                  h22 (j/m) (j\mod m))
-              [rewrite > Hcut3.
-               rewrite > Hcut4.               
-               rewrite > H10.
-               rewrite > H11.
-               apply sym_eq.
-               apply div_mod.
-               assumption
-              |apply mod_plus_times. 
-               assumption
-              ]
-            |apply div_plus_times.
-             assumption
-            ]
-          |apply lt_mod_m_m.
-           assumption
-          ] 
-        |apply (lt_times_n_to_lt m)
-          [assumption
-          |apply (le_to_lt_to_lt ? j)
-            [apply (eq_plus_to_le ? ? (j \mod m)).
-             apply div_mod.
-             assumption
-            |assumption
-            ]
-          ]
-        ]
-      |apply not_le_to_lt.unfold.intro.
-       generalize in match H2.
-       apply (le_n_O_elim ? H4).
-       rewrite < times_n_O.
-       apply le_to_not_lt.
-       apply le_O_n  
-      ] 
-    |cut (O < m)
-      [cut (j/m < n)
-        [cut (j \mod m < m)
-          [elim (and_true ? ? H3).
-           elim (H1 ? ? Hcut1 Hcut2 H4 H5).
-           elim H6.clear H6.
-           elim H8.clear H8.
-           elim H6.clear H6.
-           elim H8.clear H8.
-           apply (lt_times_plus_times ? ? ? m)
-            [assumption|assumption]
-          |apply lt_mod_m_m.
-           assumption
-          ] 
-        |apply (lt_times_n_to_lt m)
-          [assumption
-          |apply (le_to_lt_to_lt ? j)
-            [apply (eq_plus_to_le ? ? (j \mod m)).
-             apply div_mod.
-             assumption
-            |assumption
-            ]
-          ]
-        ]
-      |apply not_le_to_lt.unfold.intro.
-       generalize in match H2.
-       apply (le_n_O_elim ? H4).
-       rewrite < times_n_O.
-       apply le_to_not_lt.
-       apply le_O_n  
-      ]
-    ]
-  ]
-qed.
-     
 (* dirichlet_product is associative only up to extensional equality *)
 theorem associative_dirichlet_product: 
 \forall f,g,h:nat\to Z.\forall n:nat.O < n \to
@@ -488,6 +77,8 @@ apply (trans_eq ? ?
         (\lambda d,d1:nat.d/d1)
         (S n)
         (S n)
+        (S n)
+        (S n)
         ?
         ?
         (\lambda d,d1:nat.divides_b d1 d)
@@ -773,44 +364,6 @@ intro.apply (nat_case n)
   ]
 qed.
 
-(* da spostare in times *)
-definition Zone \def pos O.
-
-theorem Ztimes_Zone_l: \forall z:Z. Ztimes Zone z = z.
-intro.unfold Zone.simplify.
-elim z;simplify
-  [reflexivity
-  |rewrite < plus_n_O.reflexivity
-  |rewrite < plus_n_O.reflexivity
-  ]
-qed.
-
-theorem Ztimes_Zone_r: \forall z:Z. Ztimes z Zone = z.
-intro.
-rewrite < sym_Ztimes.
-apply Ztimes_Zone_l.
-qed.
-
-theorem injective_Zplus_l: \forall x:Z.injective Z Z (\lambda y.y+x).
-intro.simplify.intros (z y).
-rewrite < Zplus_z_OZ.
-rewrite < (Zplus_z_OZ y).
-rewrite < (Zplus_Zopp x).
-rewrite < (Zplus_Zopp x).
-rewrite < assoc_Zplus.
-rewrite < assoc_Zplus.
-apply eq_f2
-  [assumption|reflexivity]
-qed.
-
-theorem injective_Zplus_r: \forall x:Z.injective Z Z (\lambda y.x+y).
-intro.simplify.intros (z y).
-apply (injective_Zplus_l x).
-rewrite < sym_Zplus.
-rewrite > H.
-apply sym_Zplus.
-qed.
-
 theorem sigma_p_OZ:
 \forall p: nat \to bool.\forall n.sigma_p n p (\lambda m.OZ) = OZ.
 intros.elim n
@@ -887,64 +440,6 @@ elim n
   ]
 qed.           
 
-(* da spostare *)
-theorem notb_notb: \forall b:bool. notb (notb b) = b.
-intros.
-elim b;reflexivity.
-qed.
-
-theorem injective_notb: injective bool bool notb.
-unfold injective.
-intros.
-rewrite < notb_notb.
-rewrite < (notb_notb y).
-apply eq_f.
-assumption.
-qed.
-
-theorem divides_div: \forall d,n. divides d n \to divides (n/d) n.
-intros.
-apply (witness ? ? d).
-apply sym_eq.
-apply divides_to_div.
-assumption.
-qed.
-
-theorem divides_b_div_true: 
-\forall d,n. O < n \to 
-  divides_b d n = true \to divides_b (n/d) n = true.
-intros.
-apply divides_to_divides_b_true1
-  [assumption
-  |apply divides_div.
-   apply divides_b_true_to_divides.
-   assumption
-  ]
-qed.
-
-(* spostato in primes.ma (non in div_and_mod.ma perche' serve il predicato divides)
-theorem div_div: \forall n,d:nat. O < n \to divides d n \to 
-n/(n/d) = d.
-intros.
-apply (inj_times_l1 (n/d))
-  [apply (lt_times_n_to_lt d)
-    [apply (divides_to_lt_O ? ? H H1).
-    |rewrite > divides_to_div;assumption
-    ]
-  |rewrite > divides_to_div
-    [rewrite > sym_times.
-     rewrite > divides_to_div
-      [reflexivity
-      |assumption
-      ]
-    |apply (witness ? ? d).
-     apply sym_eq.
-     apply divides_to_div.
-     assumption
-    ]
-  ]
-qed.
-*)
 theorem commutative_dirichlet_product: \forall f,g:nat \to Z.\forall n. O < n \to 
 dirichlet_product f g n = dirichlet_product g f n.
 intros.