| changeset 57512 | cc97b347b301 |
| parent 55736 | f1ed1e9cd080 |
| child 58881 | b9556a055632 |
--- a/src/HOL/Library/Code_Target_Nat.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Code_Target_Nat.thy Fri Jul 04 20:18:47 2014 +0200 @@ -112,7 +112,7 @@ from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp show ?thesis by (simp add: Let_def divmod_nat_div_mod of_nat_add [symmetric]) - (simp add: * mult_commute of_nat_mult add_commute) + (simp add: * mult.commute of_nat_mult add.commute) qed declare of_nat_code_if [code]