Proof beautification
authornipkow
Mon, 15 Aug 1994 16:12:35 +0200
changeset 518 4530c45370b4
parent 517 a9f93400f307
child 519 98b88551e102
Proof beautification
src/ZF/IMP/Com.thy
src/ZF/IMP/Equiv.ML
--- a/src/ZF/IMP/Com.thy	Fri Aug 12 18:45:33 1994 +0200
+++ b/src/ZF/IMP/Com.thy	Mon Aug 15 16:12:35 1994 +0200
@@ -109,10 +109,12 @@
     semi    "[| <c0,sigma> -c-> sigma2; <c1,sigma2> -c-> sigma1 |] ==> \
 \            <c0 ; c1, sigma> -c-> sigma1"
 
-    ifc1     "[| c1: com; <b,sigma> -b-> 1; <c0,sigma> -c-> sigma1 |] ==> \
+    ifc1     "[| b:bexp; c1:com; sigma:loc->nat;   \
+\		 <b,sigma> -b-> 1; <c0,sigma> -c-> sigma1 |] ==> \
 \             <ifc b then c0 else c1, sigma> -c-> sigma1"
 
-    ifc0     "[| c0 : com; <b,sigma> -b-> 0; <c1,sigma> -c-> sigma1 |] ==> \
+    ifc0     "[| b:bexp; c0:com; sigma:loc->nat;   \
+\		 <b,sigma> -b-> 0; <c1,sigma> -c-> sigma1 |] ==> \
 \             <ifc b then c0 else c1, sigma> -c-> sigma1"
 
     while0   "[| c: com; <b, sigma> -b-> 0 |] ==> \
--- a/src/ZF/IMP/Equiv.ML	Fri Aug 12 18:45:33 1994 +0200
+++ b/src/ZF/IMP/Equiv.ML	Mon Aug 15 16:12:35 1994 +0200
@@ -5,23 +5,20 @@
 *)
 
 val prems = goal Equiv.thy "[| a: aexp; sigma: loc -> nat |] ==> \
-\ <a,sigma> -a-> n <-> n = A(a,sigma) ";
+\ <a,sigma> -a-> n <-> A(a,sigma) = n";
 by (res_inst_tac [("x","n")] spec 1);                       (* quantify n *)
 by (res_inst_tac [("x","a")] aexp.induct 1);                (* struct. ind. *)
 by (resolve_tac prems 1);                                   (* type prem. *)
-by (safe_tac ZF_cs);                        		    (* allI,-->,<-- *)
 by (rewrite_goals_tac A_rewrite_rules);			    (* rewr. Den.   *)
-by (TRYALL (fast_tac (ZF_cs addSIs (evala.intrs@prems)) )); (* <== *)
-by (TRYALL (fast_tac (ZF_cs addSEs aexp_elim_cases)));      (* ==> *)
+by (TRYALL (fast_tac (ZF_cs addSIs (evala.intrs@prems)
+                            addSEs aexp_elim_cases)));
 val aexp_iff = result();
 
 
-val aexp_rew_rules_cs = ZF_cs addIs  op_type_intrs@[aexp_iff RS iffD1 RS sym];
-
-val aexp1 = prove_goal Equiv.thy		    (* elim the prems *)
-        "<a,sigma> -a-> n ==> A(a,sigma) = n"	    (* destruction rule *)
-     (fn prems => [(fast_tac (aexp_rew_rules_cs addSIs prems) 1)]);
-
+val aexp1 = prove_goal Equiv.thy
+        "[| <a,sigma> -a-> n; a: aexp; sigma: loc -> nat |] \
+        \ ==> A(a,sigma) = n"	    (* destruction rule *)
+     (fn prems => [(fast_tac (ZF_cs addSIs ((aexp_iff RS iffD1)::prems)) 1)]);
 val aexp2 = aexp_iff RS iffD2;
 
 
@@ -37,37 +34,22 @@
 
 
 val prems = goal Equiv.thy "[| b: bexp; sigma: loc -> nat |] ==> \
-\ <b,sigma> -b-> w <-> w = B(b,sigma) ";
-
+\ <b,sigma> -b-> w <-> B(b,sigma) = w";
 by (res_inst_tac [("x","w")] spec 1);			(* quantify w *)
 by (res_inst_tac [("x","b")] bexp.induct 1);		(* struct. ind. *)
 by (resolve_tac prems 1);				(* type prem. *)
-by (safe_tac ZF_cs);                                  	(* allI,-->,<-- *)
 by (rewrite_goals_tac B_rewrite_rules);			(* rewr. Den.   *)
-by (TRYALL (fast_tac 					(* <== *)
-            (ZF_cs addSIs (evalb.intrs@prems@[aexp2])) ));
-by (TRYALL (fast_tac ((ZF_cs addSDs [aexp1]) addSEs bexp_elim_cases)));
-								(* ==> *)
+by (TRYALL (fast_tac (ZF_cs addSIs (evalb.intrs@prems@[aexp2])
+                            addSDs [aexp1] addSEs bexp_elim_cases)));
 val bexp_iff = result();
 
-
-val bexp_rew_rules_cs = ZF_cs addIs  op_type_intrs@[bexp_iff RS iffD1 RS sym];
-
 val bexp1 = prove_goal Equiv.thy
-        "<b,sigma> -b-> w ==> B(b,sigma) = w"
-     (fn prems => [(fast_tac (bexp_rew_rules_cs addSIs prems) 1)]);
+        "[| <b,sigma> -b-> w; b: bexp; sigma: loc -> nat |]\
+        \ ==> B(b,sigma) = w"
+     (fn prems => [(fast_tac (ZF_cs addSIs ((bexp_iff RS iffD1)::prems)) 1)]);
+val bexp2 = bexp_iff RS iffD2;
 
-val bexp2 = prove_goal Equiv.thy 
-    "[| B(b,sigma) = w; b : bexp; sigma : loc -> nat |] ==> <b,sigma> -b-> w"
-    (fn prems => 
-    [(cut_facts_tac prems 1), 
-     (fast_tac (ZF_cs addIs ([bexp_iff RS iffD2])) 1)]);
-
-
-
-val prems = goal Equiv.thy
-	"<c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
-by (cut_facts_tac prems 1);
+goal Equiv.thy "!!c. <c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
 
 (* start with rule induction *)
 be (evalc.mutual_induct RS spec RS spec RS spec RSN (2,rev_mp)) 1;
@@ -104,7 +86,7 @@
                    addEs [C_type,C_type_fst];
 
 val [prem] = goal Equiv.thy
-    "c : com ==> ALL x. x:C(c) --> <c,fst(x)> -c-> snd(x)";
+    "c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";
 br (prem RS com.induct) 1;
 by (rewrite_tac C_rewrite_rules);
 by (safe_tac com_cs);
@@ -117,15 +99,15 @@
 by (fast_tac com_cs 1);
 
 (* comp *)
-by (REPEAT (EVERY [(etac allE 1),(etac impE 1),(atac 1)]));
+by (REPEAT (EVERY [(dtac bspec 1),(atac 1)]));
 by (asm_full_simp_tac ZF_ss 1);
 by (fast_tac com_cs 1);
 
 (* while *)
-by (EVERY [(forward_tac [Gamma_bnd_mono] 1),(etac induct 1),(atac 1)]);
+by (EVERY1 [forward_tac [Gamma_bnd_mono], etac induct, atac]);
 by (rewrite_goals_tac [Gamma_def]);  
 by (safe_tac com_cs);
-by (EVERY [(etac allE 1),(etac impE 1),(atac 1)]);
+by (EVERY1 [dtac bspec, atac]);
 by (ALLGOALS (asm_full_simp_tac ZF_ss));
 
 (* while, if *)
@@ -136,7 +118,7 @@
 (**** Proof of Equivalence ****)
 
 val com_iff_cs = ZF_cs addIs [C_subset RS subsetD]
-                       addEs [com2 RS spec RS impE]
+                       addEs [com2 RS bspec]
                        addDs [com1];
 
 goal Equiv.thy