]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/binomial.ma
updating the structures for sorts
[helm.git] / helm / software / matita / library / nat / binomial.ma
index 9505761488e3dd9568d51036ff877471c21215de..83ef0acbb99d54f43ec260f16dc60447966be230 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/nat/binomial".
-
 include "nat/iteration2.ma".
 include "nat/factorial2.ma".
 
@@ -141,6 +139,28 @@ rewrite > (lt_to_lt_to_eq_div_div_times_times ? ? (S k)) in ⊢ (? ? ? (? % ?))
  apply lt_times;apply lt_O_fact]
 qed.
 
+theorem lt_O_bc: \forall n,m. m \le n \to O < bc n m.
+intro.elim n
+  [apply (le_n_O_elim ? H).
+   simplify.apply le_n
+  |elim (le_to_or_lt_eq ? ? H1)
+    [generalize in match H2.cases m;intro
+      [rewrite > bc_n_O.apply le_n
+      |rewrite > bc1
+        [apply (trans_le ? (bc n1 n2))
+          [apply H.apply le_S_S_to_le.apply lt_to_le.assumption
+          |apply le_plus_n_r
+          ]
+        |apply le_S_S_to_le.assumption
+        ]
+      ]
+    |rewrite > H2.
+     rewrite > bc_n_n.
+     apply le_n
+    ]
+  ]
+qed.
+
 theorem exp_plus_sigma_p:\forall a,b,n.
 exp (a+b) n = sigma_p (S n) (\lambda k.true) 
   (\lambda k.(bc n k)*(exp a (n-k))*(exp b k)).
@@ -226,4 +246,15 @@ apply eq_sigma_p;intros
    rewrite < times_n_SO.
    reflexivity
   ]
-qed.
\ No newline at end of file
+qed.
+
+theorem exp_Sn_SSO: \forall n. exp (S n) 2 = S((exp n 2) + 2*n).
+intros.simplify.
+rewrite < times_n_SO.
+rewrite < plus_n_O.
+rewrite < sym_times.simplify.
+rewrite < assoc_plus.
+rewrite < sym_plus.
+reflexivity.
+qed.
+