| author | wenzelm | 
| Sun, 20 Mar 2011 22:48:08 +0100 | |
| changeset 42017 | 0d4bedb25fc9 | 
| parent 41310 | 65631ca437c9 | 
| child 42455 | 6702c984bf5a | 
| permissions | -rw-r--r-- | 
| 35762 | 1  | 
(* Title: ZF/simpdata.ML  | 
| 0 | 2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
3  | 
Copyright 1991 University of Cambridge  | 
|
4  | 
||
| 35762 | 5  | 
Rewriting for ZF set theory: specialized extraction of rewrites from theorems.  | 
| 0 | 6  | 
*)  | 
7  | 
||
| 12199 | 8  | 
(*** New version of mk_rew_rules ***)  | 
| 0 | 9  | 
|
10  | 
(*Should False yield False<->True, or should it solve goals some other way?*)  | 
|
11  | 
||
| 1036 | 12  | 
(*Analyse a theorem to atomic rewrite rules*)  | 
| 13462 | 13  | 
fun atomize (conn_pairs, mem_pairs) th =  | 
| 1036 | 14  | 
let fun tryrules pairs t =  | 
| 1461 | 15  | 
case head_of t of  | 
| 13462 | 16  | 
Const(a,_) =>  | 
| 17325 | 17  | 
(case AList.lookup (op =) pairs a of  | 
| 32952 | 18  | 
SOME rls => maps (atomize (conn_pairs, mem_pairs)) ([th] RL rls)  | 
19  | 
| NONE => [th])  | 
|
| 1461 | 20  | 
| _ => [th]  | 
| 13462 | 21  | 
in case concl_of th of  | 
| 38513 | 22  | 
         Const(@{const_name Trueprop},_) $ P =>
 | 
| 1461 | 23  | 
(case P of  | 
| 24826 | 24  | 
                 Const(@{const_name mem},_) $ a $ b => tryrules mem_pairs b
 | 
| 38513 | 25  | 
               | Const(@{const_name True},_)         => []
 | 
26  | 
               | Const(@{const_name False},_)        => []
 | 
|
| 1461 | 27  | 
| A => tryrules conn_pairs A)  | 
| 1036 | 28  | 
| _ => [th]  | 
29  | 
end;  | 
|
30  | 
||
| 0 | 31  | 
(*Analyse a rigid formula*)  | 
| 1036 | 32  | 
val ZF_conn_pairs =  | 
| 38513 | 33  | 
  [(@{const_name Ball}, [@{thm bspec}]),
 | 
34  | 
   (@{const_name All}, [@{thm spec}]),
 | 
|
| 41310 | 35  | 
   (@{const_name imp}, [@{thm mp}]),
 | 
36  | 
   (@{const_name conj}, [@{thm conjunct1}, @{thm conjunct2}])];
 | 
|
| 0 | 37  | 
|
38  | 
(*Analyse a:b, where b is rigid*)  | 
|
| 13462 | 39  | 
val ZF_mem_pairs =  | 
| 38513 | 40  | 
  [(@{const_name Collect}, [@{thm CollectD1}, @{thm CollectD2}]),
 | 
41  | 
   (@{const_name Diff}, [@{thm DiffD1}, @{thm DiffD2}]),
 | 
|
42  | 
   (@{const_name Int}, [@{thm IntD1}, @{thm IntD2}])];
 | 
|
| 0 | 43  | 
|
| 1036 | 44  | 
val ZF_atomize = atomize (ZF_conn_pairs, ZF_mem_pairs);  | 
45  | 
||
| 17876 | 46  | 
change_simpset (fn ss =>  | 
| 
36543
 
0e7fc5bf38de
proper context for mksimps etc. -- via simpset of the running Simplifier;
 
wenzelm 
parents: 
35762 
diff
changeset
 | 
47  | 
ss setmksimps (K (map mk_eq o ZF_atomize o gen_all))  | 
| 24893 | 48  | 
  addcongs [@{thm if_weak_cong}]);
 | 
| 
12209
 
09bc6f8456b9
type_solver_tac: use TCSET' to refer to context of goal state (does
 
wenzelm 
parents: 
12199 
diff
changeset
 | 
49  | 
|
| 13462 | 50  | 
local  | 
| 11233 | 51  | 
|
| 24893 | 52  | 
val unfold_bex_tac = unfold_tac [@{thm Bex_def}];
 | 
| 18324 | 53  | 
fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac;  | 
| 11233 | 54  | 
val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac;  | 
55  | 
||
| 24893 | 56  | 
val unfold_ball_tac = unfold_tac [@{thm Ball_def}];
 | 
| 18324 | 57  | 
fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac;  | 
| 11233 | 58  | 
val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac;  | 
59  | 
||
60  | 
in  | 
|
61  | 
||
| 
40241
 
56fad09655a5
discontinued obsolete ML antiquotation @{theory_ref};
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
62  | 
val defBEX_regroup = Simplifier.simproc_global @{theory}
 | 
| 13462 | 63  | 
"defined BEX" ["EX x:A. P(x) & Q(x)"] rearrange_bex;  | 
64  | 
||
| 
40241
 
56fad09655a5
discontinued obsolete ML antiquotation @{theory_ref};
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
65  | 
val defBALL_regroup = Simplifier.simproc_global @{theory}
 | 
| 13462 | 66  | 
"defined BALL" ["ALL x:A. P(x) --> Q(x)"] rearrange_ball;  | 
| 11233 | 67  | 
|
68  | 
end;  | 
|
69  | 
||
| 13462 | 70  | 
Addsimprocs [defBALL_regroup, defBEX_regroup];  | 
71  | 
||
| 12199 | 72  | 
|
| 24893 | 73  | 
val ZF_ss = @{simpset};
 | 
74  |