clasohm@1465: (* Title: HOL/ex/cla clasohm@969: ID: $Id$ clasohm@1465: Author: Lawrence C Paulson, Cambridge University Computer Laboratory clasohm@969: Copyright 1994 University of Cambridge clasohm@969: clasohm@969: Higher-Order Logic: predicate calculus problems clasohm@969: clasohm@969: Taken from FOL/cla.ML; beware of precedence of = vs <-> clasohm@969: *) clasohm@969: clasohm@969: writeln"File HOL/ex/cla."; clasohm@969: wenzelm@4089: context HOL.thy; (*Boosts efficiency by omitting redundant rules*) paulson@1912: clasohm@969: goal HOL.thy "(P --> Q | R) --> (P-->Q) | (P-->R)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*If and only if*) clasohm@969: paulson@2997: goal HOL.thy "(P=Q) = (Q = (P::bool))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: goal HOL.thy "~ (P = (~P))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: clasohm@969: (*Sample problems from clasohm@969: F. J. Pelletier, clasohm@969: Seventy-Five Problems for Testing Automatic Theorem Provers, clasohm@969: J. Automated Reasoning 2 (1986), 191-216. clasohm@969: Errata, JAR 4 (1988), 236-236. clasohm@969: clasohm@969: The hardest problems -- judging by experience with several theorem provers, clasohm@969: including matrix ones -- are 34 and 43. clasohm@969: *) clasohm@969: clasohm@969: writeln"Pelletier's examples"; clasohm@969: (*1*) clasohm@969: goal HOL.thy "(P-->Q) = (~Q --> ~P)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*2*) clasohm@969: goal HOL.thy "(~ ~ P) = P"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*3*) clasohm@969: goal HOL.thy "~(P-->Q) --> (Q-->P)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*4*) clasohm@969: goal HOL.thy "(~P-->Q) = (~Q --> P)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*5*) clasohm@969: goal HOL.thy "((P|Q)-->(P|R)) --> (P|(Q-->R))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*6*) clasohm@969: goal HOL.thy "P | ~ P"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*7*) clasohm@969: goal HOL.thy "P | ~ ~ ~ P"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*8. Peirce's law*) clasohm@969: goal HOL.thy "((P-->Q) --> P) --> P"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*9*) clasohm@969: goal HOL.thy "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*10*) clasohm@969: goal HOL.thy "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P=Q)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*11. Proved in each direction (incorrectly, says Pelletier!!) *) paulson@4061: goal HOL.thy "P=(P::bool)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*12. "Dijkstra's law"*) clasohm@969: goal HOL.thy "((P = Q) = R) = (P = (Q = R))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*13. Distributive law*) clasohm@969: goal HOL.thy "(P | (Q & R)) = ((P | Q) & (P | R))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*14*) clasohm@969: goal HOL.thy "(P = Q) = ((Q | ~P) & (~Q|P))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*15*) clasohm@969: goal HOL.thy "(P --> Q) = (~P | Q)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*16*) clasohm@969: goal HOL.thy "(P-->Q) | (Q-->P)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*17*) clasohm@969: goal HOL.thy "((P & (Q-->R))-->S) = ((~P | Q | S) & (~P | ~R | S))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Classical Logic: examples with quantifiers"; clasohm@969: clasohm@969: goal HOL.thy "(! x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: wenzelm@3842: goal HOL.thy "(? x. P-->Q(x)) = (P --> (? x. Q(x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: wenzelm@3842: goal HOL.thy "(? x. P(x)-->Q) = ((! x. P(x)) --> Q)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: wenzelm@3842: goal HOL.thy "((! x. P(x)) | Q) = (! x. P(x) | Q)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*From Wishnu Prasetya*) clasohm@969: goal HOL.thy clasohm@969: "(!s. q(s) --> r(s)) & ~r(s) & (!s. ~r(s) & ~q(s) --> p(t) | q(t)) \ clasohm@969: \ --> p(t) | r(t)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: clasohm@969: writeln"Problems requiring quantifier duplication"; clasohm@969: clasohm@969: (*Needs multiple instantiation of the quantifier.*) clasohm@969: goal HOL.thy "(! x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: (*Needs double instantiation of the quantifier*) clasohm@969: goal HOL.thy "? x. P(x) --> P(a) & P(b)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: goal HOL.thy "? z. P(z) --> (! x. P(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: goal HOL.thy "? x. (? y. P(y)) --> P(x)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Hard examples with quantifiers"; clasohm@969: clasohm@969: writeln"Problem 18"; clasohm@969: goal HOL.thy "? y. ! x. P(y)-->P(x)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 19"; clasohm@969: goal HOL.thy "? x. ! y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 20"; clasohm@969: goal HOL.thy "(! x y. ? z. ! w. (P(x)&Q(y)-->R(z)&S(w))) \ clasohm@969: \ --> (? x y. P(x) & Q(y)) --> (? z. R(z))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 21"; clasohm@969: goal HOL.thy "(? x. P-->Q(x)) & (? x. Q(x)-->P) --> (? x. P=Q(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 22"; clasohm@969: goal HOL.thy "(! x. P = Q(x)) --> (P = (! x. Q(x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 23"; clasohm@969: goal HOL.thy "(! x. P | Q(x)) = (P | (! x. Q(x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 24"; clasohm@969: goal HOL.thy "~(? x. S(x)&Q(x)) & (! x. P(x) --> Q(x)|R(x)) & \ wenzelm@3842: \ (~(? x. P(x)) --> (? x. Q(x))) & (! x. Q(x)|R(x) --> S(x)) \ clasohm@969: \ --> (? x. P(x)&R(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 25"; clasohm@969: goal HOL.thy "(? x. P(x)) & \ clasohm@969: \ (! x. L(x) --> ~ (M(x) & R(x))) & \ clasohm@969: \ (! x. P(x) --> (M(x) & L(x))) & \ clasohm@969: \ ((! x. P(x)-->Q(x)) | (? x. P(x)&R(x))) \ clasohm@969: \ --> (? x. Q(x)&P(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 26"; clasohm@1465: goal HOL.thy "((? x. p(x)) = (? x. q(x))) & \ clasohm@1465: \ (! x. ! y. p(x) & q(y) --> (r(x) = s(y))) \ clasohm@969: \ --> ((! x. p(x)-->r(x)) = (! x. q(x)-->s(x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 27"; clasohm@969: goal HOL.thy "(? x. P(x) & ~Q(x)) & \ clasohm@969: \ (! x. P(x) --> R(x)) & \ clasohm@969: \ (! x. M(x) & L(x) --> P(x)) & \ clasohm@969: \ ((? x. R(x) & ~ Q(x)) --> (! x. L(x) --> ~ R(x))) \ clasohm@969: \ --> (! x. M(x) --> ~L(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 28. AMENDED"; clasohm@969: goal HOL.thy "(! x. P(x) --> (! x. Q(x))) & \ clasohm@969: \ ((! x. Q(x)|R(x)) --> (? x. Q(x)&S(x))) & \ wenzelm@3842: \ ((? x. S(x)) --> (! x. L(x) --> M(x))) \ clasohm@969: \ --> (! x. P(x) & L(x) --> M(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; clasohm@969: goal HOL.thy "(? x. F(x)) & (? y. G(y)) \ clasohm@969: \ --> ( ((! x. F(x)-->H(x)) & (! y. G(y)-->J(y))) = \ clasohm@969: \ (! x y. F(x) & G(y) --> H(x) & J(y)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 30"; clasohm@969: goal HOL.thy "(! x. P(x) | Q(x) --> ~ R(x)) & \ clasohm@969: \ (! x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ clasohm@969: \ --> (! x. S(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 31"; wenzelm@3842: goal HOL.thy "~(? x. P(x) & (Q(x) | R(x))) & \ clasohm@969: \ (? x. L(x) & P(x)) & \ clasohm@969: \ (! x. ~ R(x) --> M(x)) \ clasohm@969: \ --> (? x. L(x) & M(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 32"; clasohm@969: goal HOL.thy "(! x. P(x) & (Q(x)|R(x))-->S(x)) & \ clasohm@969: \ (! x. S(x) & R(x) --> L(x)) & \ clasohm@969: \ (! x. M(x) --> R(x)) \ clasohm@969: \ --> (! x. P(x) & M(x) --> L(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 33"; clasohm@969: goal HOL.thy "(! x. P(a) & (P(x)-->P(b))-->P(c)) = \ clasohm@969: \ (! x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: paulson@1768: writeln"Problem 34 AMENDED (TWICE!!)"; clasohm@969: (*Andrews's challenge*) clasohm@1465: goal HOL.thy "((? x. ! y. p(x) = p(y)) = \ paulson@3019: \ ((? x. q(x)) = (! y. p(y)))) = \ paulson@3019: \ ((? x. ! y. q(x) = q(y)) = \ paulson@3019: \ ((? x. p(x)) = (! y. q(y))))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 35"; clasohm@969: goal HOL.thy "? x y. P x y --> (! u v. P u v)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 36"; clasohm@969: goal HOL.thy "(! x. ? y. J x y) & \ clasohm@969: \ (! x. ? y. G x y) & \ clasohm@1465: \ (! x y. J x y | G x y --> \ clasohm@969: \ (! z. J y z | G y z --> H x z)) \ clasohm@969: \ --> (! x. ? y. H x y)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 37"; clasohm@969: goal HOL.thy "(! z. ? w. ! x. ? y. \ wenzelm@3842: \ (P x z -->P y w) & P y z & (P y w --> (? u. Q u w))) & \ clasohm@969: \ (! x z. ~(P x z) --> (? y. Q y z)) & \ clasohm@969: \ ((? x y. Q x y) --> (! x. R x x)) \ clasohm@969: \ --> (! x. ? y. R x y)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 38"; clasohm@969: goal HOL.thy clasohm@1465: "(! x. p(a) & (p(x) --> (? y. p(y) & r x y)) --> \ clasohm@1465: \ (? z. ? w. p(z) & r x w & r w z)) = \ clasohm@1465: \ (! x. (~p(a) | p(x) | (? z. ? w. p(z) & r x w & r w z)) & \ paulson@1716: \ (~p(a) | ~(? y. p(y) & r x y) | \ clasohm@969: \ (? z. ? w. p(z) & r x w & r w z)))"; paulson@2891: by (Blast_tac 1); (*beats fast_tac!*) paulson@1716: result(); clasohm@969: clasohm@969: writeln"Problem 39"; clasohm@969: goal HOL.thy "~ (? x. ! y. F y x = (~ F y y))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 40. AMENDED"; clasohm@969: goal HOL.thy "(? y. ! x. F x y = F x x) \ clasohm@969: \ --> ~ (! x. ? y. ! z. F z y = (~ F z x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 41"; clasohm@1465: goal HOL.thy "(! z. ? y. ! x. f x y = (f x z & ~ f x x)) \ clasohm@969: \ --> ~ (? z. ! x. f x z)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 42"; clasohm@969: goal HOL.thy "~ (? y. ! x. p x y = (~ (? z. p x z & p z x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: paulson@2891: writeln"Problem 43!!"; clasohm@969: goal HOL.thy clasohm@1465: "(! x::'a. ! y::'a. q x y = (! z. p z x = (p z y::bool))) \ clasohm@969: \ --> (! x. (! y. q x y = (q y x::bool)))"; paulson@2891: by (Blast_tac 1); paulson@3347: result(); clasohm@969: clasohm@969: writeln"Problem 44"; clasohm@1465: goal HOL.thy "(! x. f(x) --> \ clasohm@969: \ (? y. g(y) & h x y & (? y. g(y) & ~ h x y))) & \ clasohm@1465: \ (? x. j(x) & (! y. g(y) --> h x y)) \ clasohm@969: \ --> (? x. j(x) & ~f(x))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 45"; clasohm@969: goal HOL.thy clasohm@1465: "(! x. f(x) & (! y. g(y) & h x y --> j x y) \ clasohm@1465: \ --> (! y. g(y) & h x y --> k(y))) & \ clasohm@1465: \ ~ (? y. l(y) & k(y)) & \ clasohm@1465: \ (? x. f(x) & (! y. h x y --> l(y)) \ clasohm@1465: \ & (! y. g(y) & h x y --> j x y)) \ clasohm@969: \ --> (? x. f(x) & ~ (? y. g(y) & h x y))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: clasohm@969: writeln"Problems (mainly) involving equality or functions"; clasohm@969: clasohm@969: writeln"Problem 48"; clasohm@969: goal HOL.thy "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 49 NOT PROVED AUTOMATICALLY"; clasohm@969: (*Hard because it involves substitution for Vars; clasohm@969: the type constraint ensures that x,y,z have the same type as a,b,u. *) clasohm@969: goal HOL.thy "(? x y::'a. ! z. z=x | z=y) & P(a) & P(b) & (~a=b) \ wenzelm@3842: \ --> (! u::'a. P(u))"; paulson@4153: by (Classical.Safe_tac); clasohm@969: by (res_inst_tac [("x","a")] allE 1); clasohm@969: by (assume_tac 1); clasohm@969: by (res_inst_tac [("x","b")] allE 1); clasohm@969: by (assume_tac 1); paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 50"; clasohm@969: (*What has this to do with equality?*) wenzelm@3842: goal HOL.thy "(! x. P a x | (! y. P x y)) --> (? x. ! y. P x y)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 51"; clasohm@969: goal HOL.thy clasohm@969: "(? z w. ! x y. P x y = (x=z & y=w)) --> \ clasohm@969: \ (? z. ! x. ? w. (! y. P x y = (y=w)) = (x=z))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 52"; clasohm@969: (*Almost the same as 51. *) clasohm@969: goal HOL.thy clasohm@969: "(? z w. ! x y. P x y = (x=z & y=w)) --> \ clasohm@969: \ (? w. ! y. ? z. (! x. P x y = (x=z)) = (y=w))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 55"; clasohm@969: clasohm@969: (*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988). clasohm@969: fast_tac DISCOVERS who killed Agatha. *) clasohm@969: goal HOL.thy "lives(agatha) & lives(butler) & lives(charles) & \ clasohm@969: \ (killed agatha agatha | killed butler agatha | killed charles agatha) & \ clasohm@969: \ (!x y. killed x y --> hates x y & ~richer x y) & \ clasohm@969: \ (!x. hates agatha x --> ~hates charles x) & \ clasohm@969: \ (hates agatha agatha & hates agatha charles) & \ clasohm@969: \ (!x. lives(x) & ~richer x agatha --> hates butler x) & \ clasohm@969: \ (!x. hates agatha x --> hates butler x) & \ clasohm@969: \ (!x. ~hates x agatha | ~hates x butler | ~hates x charles) --> \ clasohm@969: \ killed ?who agatha"; paulson@2922: by (Fast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 56"; clasohm@969: goal HOL.thy clasohm@969: "(! x. (? y. P(y) & x=f(y)) --> P(x)) = (! x. P(x) --> P(f(x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 57"; clasohm@969: goal HOL.thy clasohm@969: "P (f a b) (f b c) & P (f b c) (f a c) & \ clasohm@969: \ (! x y z. P x y & P y z --> P x z) --> P (f a b) (f a c)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 58 NOT PROVED AUTOMATICALLY"; clasohm@969: goal HOL.thy "(! x y. f(x)=g(y)) --> (! x y. f(f(x))=f(g(y)))"; clasohm@969: val f_cong = read_instantiate [("f","f")] arg_cong; wenzelm@4089: by (fast_tac (claset() addIs [f_cong]) 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 59"; clasohm@969: goal HOL.thy "(! x. P(x) = (~P(f(x)))) --> (? x. P(x) & ~P(f(x)))"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: clasohm@969: writeln"Problem 60"; clasohm@969: goal HOL.thy clasohm@969: "! x. P x (f x) = (? y. (! z. P z y --> P z (f x)) & P x y)"; paulson@2891: by (Blast_tac 1); clasohm@969: result(); clasohm@969: paulson@2715: writeln"Problem 62 as corrected in JAR 18 (1997), page 135"; paulson@1404: goal HOL.thy clasohm@1465: "(ALL x. p a & (p x --> p(f x)) --> p(f(f x))) = \ clasohm@1465: \ (ALL x. (~ p a | p x | p(f(f x))) & \ paulson@1404: \ (~ p a | ~ p(f x) | p(f(f x))))"; paulson@2891: by (Blast_tac 1); paulson@1404: result(); paulson@1404: paulson@1712: (*From Rudnicki, Obvious Inferences, JAR 3 (1987), 383-393. paulson@1712: It does seem obvious!*) paulson@1712: goal Prod.thy paulson@1712: "(ALL x. F(x) & ~G(x) --> (EX y. H(x,y) & J(y))) & \ paulson@1712: \ (EX x. K(x) & F(x) & (ALL y. H(x,y) --> K(y))) & \ paulson@1712: \ (ALL x. K(x) --> ~G(x)) --> (EX x. K(x) --> ~G(x))"; paulson@2891: by (Blast_tac 1); paulson@1712: result(); paulson@1712: paulson@1712: goal Prod.thy paulson@1712: "(ALL x y. R(x,y) | R(y,x)) & \ paulson@1712: \ (ALL x y. S(x,y) & S(y,x) --> x=y) & \ paulson@1712: \ (ALL x y. R(x,y) --> S(x,y)) --> (ALL x y. S(x,y) --> R(x,y))"; paulson@2891: by (Blast_tac 1); paulson@1712: result(); paulson@1712: paulson@1712: paulson@1712: clasohm@969: writeln"Reached end of file.";