[ assumption
| rewrite > (sym_times c (a/c)).
rewrite > (NdivM_times_M_to_N a c) in \vdash (? % ?)
- [ rewrite < (sym_times a c).
- apply (O_lt_const_to_le_times_const).
+ [ apply (le_times_n c a ?).
assumption
| assumption
| assumption
]
qed.
-(*
-theorem div_times_to_eqSO: \forall a,d:nat.
-O \lt d \to a*d = d \to a = (S O).
-intros.
-apply (inj_times_r1 d)
-[ assumption
-| rewrite > sym_times.
- rewrite < (times_n_SO d).
- assumption
-]
-qed.*)
-
theorem div_mod_minus:
\forall a,b:nat. O \lt b \to
]
qed.
-(*
-theorem sum_div_eq_div: \forall a,b,c:nat.
-O \lt c \to b \lt c \to c \divides a \to (a+b) /c = a/c.
-intros.
-elim H2.
-rewrite > H3.
-rewrite > (sym_times c n2).
-rewrite > (div_plus_times c n2 b)
-[ rewrite > (div_times_ltO n2 c)
- [ reflexivity
- | assumption
- ]
-| assumption
-]
-qed.
-*)
(* A corollary to the division theorem (between natural numbers).
*
]
| assumption
]
-| apply (nat_case1 n)
- [ intros.
- rewrite > H7 in H6.
- rewrite > sym_times in H6.
- simplify in H6.
- rewrite > H6 in H.
- apply False_ind.
- change in H with ((S O) \le O).
- apply (not_le_Sn_O O H)
- | intros.
- apply (lt_O_S m)
+| apply (divides_to_lt_O n c)
+ [ assumption
+ | apply (witness n c b).
+ rewrite > sym_times.
+ assumption
]
]
qed.
| apply (sum_divisor_totient1_aux2);
assumption
]
- | apply (eqb_true_to_eq (gcd (i \mod n) (i/n)) (S O)).
+ | apply (eqb_true_to_eq (gcd (i \mod n) (i/n)) (S O)).
apply (andb_true_true
(eqb (gcd (i\mod n) (i/n)) (S O))
(leb (S (i\mod n)) (i/n))