src/FOL/ex/If.ML
author paulson
Thu, 21 Mar 1996 13:02:26 +0100
changeset 1601 0ef6ea27ab15
parent 1459 d12da312eff4
child 2469 b50b8c0eec01
permissions -rw-r--r--
Changes required by removal of the theory argument of Theorem

(*  Title:      FOL/ex/if
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1991  University of Cambridge

For ex/if.thy.  First-Order Logic: the 'if' example
*)

open If;
open Cla;    (*in case structure Int is open!*)

val prems = goalw If.thy [if_def]
    "[| P ==> Q; ~P ==> R |] ==> if(P,Q,R)";
by (fast_tac (FOL_cs addIs prems) 1);
qed "ifI";

val major::prems = goalw If.thy [if_def]
   "[| if(P,Q,R);  [| P; Q |] ==> S; [| ~P; R |] ==> S |] ==> S";
by (cut_facts_tac [major] 1);
by (fast_tac (FOL_cs addIs prems) 1);
qed "ifE";


goal If.thy
    "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))";
by (rtac iffI 1);
by (etac ifE 1);
by (etac ifE 1);
by (rtac ifI 1);
by (rtac ifI 1);

choplev 0;
val if_cs = FOL_cs addSIs [ifI] addSEs[ifE];
by (fast_tac if_cs 1);
qed "if_commute";


goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))";
by (fast_tac if_cs 1);
qed "nested_ifs";

choplev 0;
by (rewtac if_def);
by (fast_tac FOL_cs 1);
result();


(*An invalid formula.  High-level rules permit a simpler diagnosis*)
goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))";
by (fast_tac if_cs 1) handle ERROR => writeln"Failed, as expected";
(*Check that subgoals remain: proof failed.*)
getgoal 1; 
by (REPEAT (step_tac if_cs 1));

choplev 0;
by (rewtac if_def);
by (fast_tac FOL_cs 1) handle ERROR => writeln"Failed, as expected";
(*Check that subgoals remain: proof failed.*)
getgoal 1; 
by (REPEAT (step_tac FOL_cs 1));



writeln"Reached end of file.";