# HG changeset patch # User paulson # Date 1119974164 -7200 # Node ID 0115764233e4ec2b79195106e18a857033962dae # Parent e7df213a1918b933437c073733a743e03e60810d stylistic improvements diff -r e7df213a1918 -r 0115764233e4 src/HOL/ex/Classical.thy --- a/src/HOL/ex/Classical.thy Tue Jun 28 16:12:03 2005 +0200 +++ b/src/HOL/ex/Classical.thy Tue Jun 28 17:56:04 2005 +0200 @@ -706,26 +706,27 @@ --> (\x. f x --> g x)" by blast -text{*Problem 47. Schubert's Steamroller*} - text{*26 clauses; 63 Horn clauses - 87094 inferences so far. Searching to depth 36*} -lemma "(\x. P1 x --> P0 x) & (\x. P1 x) & - (\x. P2 x --> P0 x) & (\x. P2 x) & - (\x. P3 x --> P0 x) & (\x. P3 x) & - (\x. P4 x --> P0 x) & (\x. P4 x) & - (\x. P5 x --> P0 x) & (\x. P5 x) & - (\x. Q1 x --> Q0 x) & (\x. Q1 x) & - (\x. P0 x --> ((\y. Q0 y-->R x y) | - (\y. P0 y & S y x & - (\z. Q0 z&R y z) --> R x y))) & - (\x y. P3 y & (P5 x|P4 x) --> S x y) & - (\x y. P3 x & P2 y --> S x y) & - (\x y. P2 x & P1 y --> S x y) & - (\x y. P1 x & (P2 y|Q1 y) --> ~R x y) & - (\x y. P3 x & P4 y --> R x y) & - (\x y. P3 x & P5 y --> ~R x y) & - (\x. (P4 x|P5 x) --> (\y. Q0 y & R x y)) - --> (\x y. P0 x & P0 y & (\z. Q1 z & R y z & R x y))" +text{*Problem 47. Schubert's Steamroller. + 26 clauses; 63 Horn clauses. + 87094 inferences so far. Searching to depth 36*} +lemma "(\x. wolf x \ animal x) & (\x. wolf x) & + (\x. fox x \ animal x) & (\x. fox x) & + (\x. bird x \ animal x) & (\x. bird x) & + (\x. caterpillar x \ animal x) & (\x. caterpillar x) & + (\x. snail x \ animal x) & (\x. snail x) & + (\x. grain x \ plant x) & (\x. grain x) & + (\x. animal x \ + ((\y. plant y \ eats x y) \ + (\y. animal y & smaller_than y x & + (\z. plant z & eats y z) \ eats x y))) & + (\x y. bird y & (snail x \ caterpillar x) \ smaller_than x y) & + (\x y. bird x & fox y \ smaller_than x y) & + (\x y. fox x & wolf y \ smaller_than x y) & + (\x y. wolf x & (fox y \ grain y) \ ~eats x y) & + (\x y. bird x & caterpillar y \ eats x y) & + (\x y. bird x & snail y \ ~eats x y) & + (\x. (caterpillar x \ snail x) \ (\y. plant y & eats x y)) + \ (\x y. animal x & animal y & (\z. grain z & eats y z & eats x y))" by (tactic{*safe_best_meson_tac 1*}) --{*Nearly twice as fast as @{text meson}, which performs iterative deepening rather than best-first search*} diff -r e7df213a1918 -r 0115764233e4 src/HOL/ex/NatSum.thy --- a/src/HOL/ex/NatSum.thy Tue Jun 28 16:12:03 2005 +0200 +++ b/src/HOL/ex/NatSum.thy Tue Jun 28 17:56:04 2005 +0200 @@ -18,7 +18,6 @@ *} lemmas [simp] = - lessThan_Suc atMost_Suc setsum_op_ivl_Suc setsum_cl_ivl_Suc left_distrib right_distrib left_diff_distrib right_diff_distrib --{*for true subtraction*} diff_mult_distrib diff_mult_distrib2 --{*for type nat*} @@ -28,10 +27,8 @@ squared. *} -lemma sum_of_odds: "(\i \ {0..i=0..i=0..i=0..i=0..n. i) = n * Suc n" - apply (induct n) - apply auto - done + by (induct n, auto) lemma sum_of_squares: "6 * (\i=0..n. i * i) = n * Suc n * Suc (2 * n)" - apply (induct n) - apply auto - done + by (induct n, auto) lemma sum_of_cubes: "4 * (\i=0..n. i * i * i) = n * n * Suc n * Suc n" - apply (induct n) - apply auto - done + by (induct n, auto) text {* @@ -95,24 +80,20 @@ done text {* - Tow alternative proofs, with a change of variables and much more + Two alternative proofs, with a change of variables and much more subtraction, performed using the integers. *} lemma int_sum_of_fourth_powers: "30 * int (\i=0..i=0..i=0.. (k - 1) * (\i=0..