# HG changeset patch
# User haftmann
# Date 1527012869 0
# Node ID 8b284d24f43464b7d6e47f643d077e23884c0bb3
# Parent  54a1278737352ede356ce42ff63f22cdb386940f
automatic classical rule to derive a dvd b from b mod a = 0

diff -r 54a127873735 -r 8b284d24f434 src/HOL/Rings.thy
--- a/src/HOL/Rings.thy	Tue May 22 18:14:29 2018 +0000
+++ b/src/HOL/Rings.thy	Tue May 22 18:14:29 2018 +0000
@@ -1712,7 +1712,7 @@
   "b mod a = 0" if "a dvd b"
   using that minus_div_mult_eq_mod [of b a] by simp
 
-lemma mod_0_imp_dvd: 
+lemma mod_0_imp_dvd [dest!]: 
   "b dvd a" if "a mod b = 0"
 proof -
   have "b dvd (a div b) * b" by simp