]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/arithmetics/chinese_reminder.ma
we added a definition and a couple of lemmas
[helm.git] / matita / matita / lib / arithmetics / chinese_reminder.ma
index 0837e429e2eb4e97d0d558934eea16e5db98dd30..ec20f30dac3c73aa53bc0c8771cbd4239d85195b 100644 (file)
@@ -30,7 +30,7 @@ cut (∃c,d.c*n - d*m = 1 ∨ d*m - c*n = 1) [<pnm @eq_minus_gcd]
      <commutative_plus <plus_minus
        [@(eq_times_plus_to_congruent … posm) //
        |applyS monotonic_le_times_r @(transitive_le ? ((a+b*m)*d)) // 
-        applyS monotonic_le_times_r apply (transitive_le … (b*m)) /2/
+        applyS monotonic_le_times_r @(transitive_le … (b*m)) /2/
        ]
     |(* congruent to b *)
      -pnm (cut (d*m = c*n-1))
@@ -118,9 +118,9 @@ theorem mod_cr_pair : ∀m,n,a,b. a < m → b < n → gcd n m = 1 →
 #m #n #a #b #ltam #ltbn #pnm 
 cut (andb (eqb ((cr_pair m n a b) \mod m) a) 
          (eqb ((cr_pair m n a b) \mod n) b) = true)
-  [@f_min_true cases(congruent_ab_lt m n a b ?? pnm)
+  [whd in match (cr_pair m n a b); @f_min_true cases(congruent_ab_lt m n a b ?? pnm)
     [#x * * #cxa #cxb #ltx @(ex_intro ?? x) % /2/
-     >eq_to_eqb_true 
+     >eq_to_eqb_true
       [>eq_to_eqb_true // <(lt_to_eq_mod …ltbn) //
       |<(lt_to_eq_mod …ltam) //
       ]