| author | wenzelm | 
| Mon, 06 Nov 2000 22:50:50 +0100 | |
| changeset 10404 | 93efbb62667c | 
| parent 9907 | 473a6604da94 | 
| child 11233 | 34c81a796ee3 | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: ZF/simpdata | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1991 University of Cambridge | |
| 5 | ||
| 2469 | 6 | Rewriting for ZF set theory: specialized extraction of rewrites from theorems | 
| 0 | 7 | *) | 
| 8 | ||
| 2469 | 9 | (** Rewriting **) | 
| 0 | 10 | |
| 3425 | 11 | local | 
| 12 | (*For proving rewrite rules*) | |
| 9570 
e16e168984e1
installation of cancellation simprocs for the integers
 paulson parents: 
7570diff
changeset | 13 | fun prover s = (prove_goal (the_context ()) s (fn _ => [Blast_tac 1])); | 
| 3425 | 14 | |
| 15 | in | |
| 0 | 16 | |
| 3425 | 17 | val ball_simps = map prover | 
| 18 | ["(ALL x:A. P(x) | Q) <-> ((ALL x:A. P(x)) | Q)", | |
| 19 | "(ALL x:A. P | Q(x)) <-> (P | (ALL x:A. Q(x)))", | |
| 20 | "(ALL x:A. P --> Q(x)) <-> (P --> (ALL x:A. Q(x)))", | |
| 21 | "(ALL x:A. P(x) --> Q) <-> ((EX x:A. P(x)) --> Q)", | |
| 22 | "(ALL x:0.P(x)) <-> True", | |
| 3840 | 23 | "(ALL x:succ(i).P(x)) <-> P(i) & (ALL x:i. P(x))", | 
| 24 | "(ALL x:cons(a,B).P(x)) <-> P(a) & (ALL x:B. P(x))", | |
| 2482 
87383dd9f4b5
Default rewrite rules for quantification over Collect(A,P)
 paulson parents: 
2469diff
changeset | 25 | "(ALL x:RepFun(A,f). P(x)) <-> (ALL y:A. P(f(y)))", | 
| 
87383dd9f4b5
Default rewrite rules for quantification over Collect(A,P)
 paulson parents: 
2469diff
changeset | 26 | "(ALL x:Union(A).P(x)) <-> (ALL y:A. ALL x:y. P(x))", | 
| 3859 | 27 | "(ALL x:Collect(A,Q).P(x)) <-> (ALL x:A. Q(x) --> P(x))", | 
| 28 | "(~(ALL x:A. P(x))) <-> (EX x:A. ~P(x))"]; | |
| 2482 
87383dd9f4b5
Default rewrite rules for quantification over Collect(A,P)
 paulson parents: 
2469diff
changeset | 29 | |
| 3425 | 30 | val ball_conj_distrib = | 
| 31 | prover "(ALL x:A. P(x) & Q(x)) <-> ((ALL x:A. P(x)) & (ALL x:A. Q(x)))"; | |
| 32 | ||
| 33 | val bex_simps = map prover | |
| 34 | ["(EX x:A. P(x) & Q) <-> ((EX x:A. P(x)) & Q)", | |
| 35 | "(EX x:A. P & Q(x)) <-> (P & (EX x:A. Q(x)))", | |
| 36 | "(EX x:0.P(x)) <-> False", | |
| 3840 | 37 | "(EX x:succ(i).P(x)) <-> P(i) | (EX x:i. P(x))", | 
| 38 | "(EX x:cons(a,B).P(x)) <-> P(a) | (EX x:B. P(x))", | |
| 2482 
87383dd9f4b5
Default rewrite rules for quantification over Collect(A,P)
 paulson parents: 
2469diff
changeset | 39 | "(EX x:RepFun(A,f). P(x)) <-> (EX y:A. P(f(y)))", | 
| 
87383dd9f4b5
Default rewrite rules for quantification over Collect(A,P)
 paulson parents: 
2469diff
changeset | 40 | "(EX x:Union(A).P(x)) <-> (EX y:A. EX x:y. P(x))", | 
| 3859 | 41 | "(EX x:Collect(A,Q).P(x)) <-> (EX x:A. Q(x) & P(x))", | 
| 42 | "(~(EX x:A. P(x))) <-> (ALL x:A. ~P(x))"]; | |
| 2482 
87383dd9f4b5
Default rewrite rules for quantification over Collect(A,P)
 paulson parents: 
2469diff
changeset | 43 | |
| 3425 | 44 | val bex_disj_distrib = | 
| 45 | prover "(EX x:A. P(x) | Q(x)) <-> ((EX x:A. P(x)) | (EX x:A. Q(x)))"; | |
| 46 | ||
| 47 | val Rep_simps = map prover | |
| 5202 | 48 |     ["{x. y:0, R(x,y)} = 0",	(*Replace*)
 | 
| 49 |      "{x:0. P(x)} = 0",		(*Collect*)
 | |
| 3425 | 50 |      "{x:A. False} = 0",
 | 
| 51 |      "{x:A. True} = A",
 | |
| 5202 | 52 | "RepFun(0,f) = 0", (*RepFun*) | 
| 3425 | 53 | "RepFun(succ(i),f) = cons(f(i), RepFun(i,f))", | 
| 54 | "RepFun(cons(a,B),f) = cons(f(a), RepFun(B,f))"] | |
| 0 | 55 | |
| 3425 | 56 | val misc_simps = map prover | 
| 57 | ["0 Un A = A", "A Un 0 = A", | |
| 58 | "0 Int A = 0", "A Int 0 = 0", | |
| 59 | "0-A = 0", "A-0 = A", | |
| 60 | "Union(0) = 0", | |
| 61 | "Union(cons(b,A)) = b Un Union(A)", | |
| 62 |      "Inter({b}) = b"]
 | |
| 0 | 63 | |
| 3425 | 64 | end; | 
| 65 | ||
| 9907 | 66 | bind_thms ("ball_simps", ball_simps);
 | 
| 67 | bind_thm ("ball_conj_distrib", ball_conj_distrib);
 | |
| 68 | bind_thms ("bex_simps", bex_simps);
 | |
| 69 | bind_thm ("bex_disj_distrib", bex_disj_distrib);
 | |
| 70 | bind_thms ("Rep_simps", Rep_simps);
 | |
| 71 | bind_thms ("misc_simps", misc_simps);
 | |
| 72 | ||
| 3425 | 73 | Addsimps (ball_simps @ bex_simps @ Rep_simps @ misc_simps); | 
| 74 | ||
| 0 | 75 | |
| 76 | (** New version of mk_rew_rules **) | |
| 77 | ||
| 78 | (*Should False yield False<->True, or should it solve goals some other way?*) | |
| 79 | ||
| 1036 | 80 | (*Analyse a theorem to atomic rewrite rules*) | 
| 81 | fun atomize (conn_pairs, mem_pairs) th = | |
| 82 | let fun tryrules pairs t = | |
| 1461 | 83 | case head_of t of | 
| 84 | Const(a,_) => | |
| 85 | (case assoc(pairs,a) of | |
| 86 | Some rls => flat (map (atomize (conn_pairs, mem_pairs)) | |
| 87 | ([th] RL rls)) | |
| 88 | | None => [th]) | |
| 89 | | _ => [th] | |
| 1036 | 90 | in case concl_of th of | 
| 1461 | 91 |          Const("Trueprop",_) $ P => 
 | 
| 92 | (case P of | |
| 93 |                  Const("op :",_) $ a $ b => tryrules mem_pairs b
 | |
| 94 |                | Const("True",_)         => []
 | |
| 95 |                | Const("False",_)        => []
 | |
| 96 | | A => tryrules conn_pairs A) | |
| 1036 | 97 | | _ => [th] | 
| 98 | end; | |
| 99 | ||
| 0 | 100 | (*Analyse a rigid formula*) | 
| 1036 | 101 | val ZF_conn_pairs = | 
| 1461 | 102 |   [("Ball",     [bspec]), 
 | 
| 103 |    ("All",      [spec]),
 | |
| 104 |    ("op -->",   [mp]),
 | |
| 105 |    ("op &",     [conjunct1,conjunct2])];
 | |
| 0 | 106 | |
| 107 | (*Analyse a:b, where b is rigid*) | |
| 1036 | 108 | val ZF_mem_pairs = | 
| 1461 | 109 |   [("Collect",  [CollectD1,CollectD2]),
 | 
| 110 |    ("op -",     [DiffD1,DiffD2]),
 | |
| 111 |    ("op Int",   [IntD1,IntD2])];
 | |
| 0 | 112 | |
| 1036 | 113 | val ZF_atomize = atomize (ZF_conn_pairs, ZF_mem_pairs); | 
| 114 | ||
| 5553 | 115 | simpset_ref() := simpset() setmksimps (map mk_eq o ZF_atomize o gen_all) | 
| 9570 
e16e168984e1
installation of cancellation simprocs for the integers
 paulson parents: 
7570diff
changeset | 116 | addcongs [if_weak_cong] | 
| 
e16e168984e1
installation of cancellation simprocs for the integers
 paulson parents: 
7570diff
changeset | 117 | addsplits [split_if] | 
| 7570 | 118 | setSolver (mk_solver "types" Type_solver_tac); | 
| 2469 | 119 | |
| 4091 | 120 | val ZF_ss = simpset(); |