src/ZF/simpdata.ML
author wenzelm
Tue, 20 May 1997 19:29:50 +0200
changeset 3257 4e3724e0659f
parent 2876 02c12d4c8b97
child 3425 fc4ca570d185
permissions -rw-r--r--
README generation;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     1
(*  Title:      ZF/simpdata
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
     6
Rewriting for ZF set theory: specialized extraction of rewrites from theorems
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
     9
(** Rewriting **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    11
(*For proving rewrite rules*)
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    12
fun prove_fun s = 
2876
02c12d4c8b97 Uses ZF.thy again, to make that theory usable
paulson
parents: 2496
diff changeset
    13
    (writeln s;  prove_goal ZF.thy s
2496
40efb87985b5 Removal of needless "addIs [equality]", etc.
paulson
parents: 2482
diff changeset
    14
       (fn prems => [ (cut_facts_tac prems 1), (Fast_tac 1) ]));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    16
(*Are all these really suitable?*)
2482
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    17
val ball_simps = map prove_fun
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    18
    ["(ALL x:0.P(x)) <-> True",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    19
     "(ALL x:succ(i).P(x)) <-> P(i) & (ALL x:i.P(x))",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    20
     "(ALL x:cons(a,B).P(x)) <-> P(a) & (ALL x:B.P(x))",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    21
     "(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: 2469
diff changeset
    22
     "(ALL x:Union(A).P(x)) <-> (ALL y:A. ALL x:y. P(x))",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    23
     "(ALL x:Collect(A,Q).P(x)) <-> (ALL x:A. Q(x) --> P(x))"];
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    24
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    25
val bex_simps = map prove_fun
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    26
    ["(EX x:0.P(x)) <-> False",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    27
     "(EX x:succ(i).P(x)) <-> P(i) | (EX x:i.P(x))",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    28
     "(EX x:cons(a,B).P(x)) <-> P(a) | (EX x:B.P(x))",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    29
     "(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: 2469
diff changeset
    30
     "(EX x:Union(A).P(x)) <-> (EX y:A. EX x:y.  P(x))",
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    31
     "(EX x:Collect(A,Q).P(x)) <-> (EX x:A. Q(x) & P(x))"];
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    32
87383dd9f4b5 Default rewrite rules for quantification over Collect(A,P)
paulson
parents: 2469
diff changeset
    33
Addsimps (ball_simps @ bex_simps);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    35
Addsimps (map prove_fun
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    36
	  ["{x:0. P(x)} = 0",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    37
	   "{x:A. False} = 0",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    38
	   "{x:A. True} = A",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    39
	   "RepFun(0,f) = 0",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    40
	   "RepFun(succ(i),f) = cons(f(i), RepFun(i,f))",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    41
	   "RepFun(cons(a,B),f) = cons(f(a), RepFun(B,f))"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    43
Addsimps (map prove_fun
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    44
	  ["0 Un A = A",  "A Un 0 = A",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    45
	   "0 Int A = 0", "A Int 0 = 0",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    46
	   "0-A = 0",     "A-0 = A",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    47
	   "Union(0) = 0",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    48
	   "Union(cons(b,A)) = b Un Union(A)",
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    49
	   "Inter({b}) = b"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
(** New version of mk_rew_rules **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
(*Should False yield False<->True, or should it solve goals some other way?*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
1036
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    55
(*Analyse a theorem to atomic rewrite rules*)
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    56
fun atomize (conn_pairs, mem_pairs) th = 
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    57
  let fun tryrules pairs t =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    58
          case head_of t of
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    59
              Const(a,_) => 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    60
                (case assoc(pairs,a) of
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    61
                     Some rls => flat (map (atomize (conn_pairs, mem_pairs))
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    62
                                       ([th] RL rls))
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    63
                   | None     => [th])
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    64
            | _ => [th]
1036
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    65
  in case concl_of th of 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    66
         Const("Trueprop",_) $ P => 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    67
            (case P of
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    68
                 Const("op :",_) $ a $ b => tryrules mem_pairs b
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    69
               | Const("True",_)         => []
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    70
               | Const("False",_)        => []
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    71
               | A => tryrules conn_pairs A)
1036
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    72
       | _                       => [th]
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    73
  end;
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    74
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
(*Analyse a rigid formula*)
1036
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    76
val ZF_conn_pairs =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    77
  [("Ball",     [bspec]), 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    78
   ("All",      [spec]),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    79
   ("op -->",   [mp]),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    80
   ("op &",     [conjunct1,conjunct2])];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
(*Analyse a:b, where b is rigid*)
1036
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    83
val ZF_mem_pairs = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    84
  [("Collect",  [CollectD1,CollectD2]),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    85
   ("op -",     [DiffD1,DiffD2]),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1036
diff changeset
    86
   ("op Int",   [IntD1,IntD2])];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
1036
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    88
val ZF_atomize = atomize (ZF_conn_pairs, ZF_mem_pairs);
0d28f4bc8a44 Recoded function atomize so that new ways of creating
lcp
parents: 855
diff changeset
    89
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    90
simpset := !simpset setmksimps (map mk_meta_eq o ZF_atomize o gen_all);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    91
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    92
val ZF_ss = !simpset;