diff -r 2e48182cc82c -r cf58b5b794b2 src/HOL/ex/Classical.thy --- a/src/HOL/ex/Classical.thy Sat Dec 26 15:44:14 2015 +0100 +++ b/src/HOL/ex/Classical.thy Sat Dec 26 15:59:27 2015 +0100 @@ -11,9 +11,8 @@ text\The machine "griffon" mentioned below is a 2.5GHz Power Mac G5.\ -text\Taken from @{text "FOL/Classical.thy"}. When porting examples from -first-order logic, beware of the precedence of @{text "="} versus @{text -"\"}.\ +text\Taken from \FOL/Classical.thy\. When porting examples from +first-order logic, beware of the precedence of \=\ versus \\\.\ lemma "(P --> Q | R) --> (P-->Q) | (P-->R)" by blast @@ -430,7 +429,7 @@ (\x. \y. R(x,y)) --> ~ (\x. P x = (\y. R(x,y) --> ~ P y))" by (tactic\Meson.safe_best_meson_tac @{context} 1\) - --\In contrast, @{text meson} is SLOW: 7.6s on griffon\ + \\In contrast, \meson\ is SLOW: 7.6s on griffon\ subsubsection\Pelletier's examples\ @@ -645,7 +644,7 @@ (\x z. ~P x z --> (\y. Q y z)) & ((\x y. Q x y) --> (\x. R x x)) --> (\x. \y. R x y)" -by blast --\causes unification tracing messages\ +by blast \\causes unification tracing messages\ text\Problem 38\ text\Quite hard: 422 Horn clauses!!\ @@ -724,7 +723,7 @@ (\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\Meson.safe_best_meson_tac @{context} 1\) - --\Nearly twice as fast as @{text meson}, + \\Nearly twice as fast as \meson\, which performs iterative deepening rather than best-first search\ text\The Los problem. Circulated by John Harrison\ @@ -756,7 +755,7 @@ text\Problem 55\ text\Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988). - @{text meson} cannot report who killed Agatha.\ + \meson\ cannot report who killed Agatha.\ lemma "lives agatha & lives butler & lives charles & (killed agatha agatha | killed butler agatha | killed charles agatha) & (\x y. killed x y --> hates x y & ~richer x y) & @@ -802,13 +801,13 @@ shows True proof - from a b d have "\x. T(i x x)" by blast - from a b c d have "\x. T(i x (n(n x)))" --\Problem 66\ + from a b c d have "\x. T(i x (n(n x)))" \\Problem 66\ by metis - from a b c d have "\x. T(i (n(n x)) x)" --\Problem 67\ + from a b c d have "\x. T(i (n(n x)) x)" \\Problem 67\ by meson - --\4.9s on griffon. 51061 inferences, depth 21\ + \\4.9s on griffon. 51061 inferences, depth 21\ from a b c' d have "\x. T(i x (n(n x)))" - --\Problem 68: not proved. Listed as satisfiable in TPTP (LCL078-1)\ + \\Problem 68: not proved. Listed as satisfiable in TPTP (LCL078-1)\ oops text\Problem 71, as found in TPTP (SYN007+1.005)\