# HG changeset patch # User nipkow # Date 1236868286 -3600 # Node ID 0a41b0662264f5f8990385348f3a3039b4a2f75c # Parent f1cb00030d4f4fbac4e0667c94396b10bee7e9d1 added div lemmas diff -r f1cb00030d4f -r 0a41b0662264 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Thu Mar 12 00:02:30 2009 +0100 +++ b/src/HOL/Divides.thy Thu Mar 12 15:31:26 2009 +0100 @@ -295,6 +295,27 @@ end +lemma div_mult_div_if_dvd: "(y::'a::{semiring_div,no_zero_divisors}) dvd x \ + z dvd w \ (x div y) * (w div z) = (x * w) div (y * z)" +unfolding dvd_def + apply clarify + apply (case_tac "y = 0") + apply simp + apply (case_tac "z = 0") + apply simp + apply (simp add: algebra_simps) + apply (subst mult_assoc [symmetric]) + apply (simp add: no_zero_divisors) +done + + +lemma div_power: "(y::'a::{semiring_div,no_zero_divisors,recpower}) dvd x \ + (x div y)^n = x^n div y^n" +apply (induct n) + apply simp +apply(simp add: div_mult_div_if_dvd dvd_power_same) +done + class ring_div = semiring_div + comm_ring_1 begin