src/ZF/ROOT.ML
author paulson
Mon, 07 Aug 2000 10:29:54 +0200
changeset 9548 15bee2731e43
parent 9211 6236c5285bd8
child 9570 e16e168984e1
permissions -rw-r--r--
instantiated Cancel_Numerals for "nat" in ZF
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1361
diff changeset
     1
(*  Title:      ZF/ROOT
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 5
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1361
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Adds Zermelo-Fraenkel Set Theory to a database containing First-Order Logic.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
This theory is the work of Martin Coen, Philippe Noel and Lawrence Paulson.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
val banner = "ZF Set Theory (in FOL)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
writeln banner;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1512
diff changeset
    14
eta_contract:=false;
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1512
diff changeset
    15
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
print_depth 1;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
    18
(*Add user sections for inductive/datatype definitions*)
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    19
use     "thy_syntax";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
    20
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9211
diff changeset
    21
use "~~/src/Provers/Arith/cancel_numerals.ML";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9211
diff changeset
    22
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    23
use_thy "ZF";
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    24
use     "Tools/typechk";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    25
use_thy "mono";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    26
use     "ind_syntax";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    27
use     "Tools/cartprod";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    28
use_thy "Fixedpt";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    29
use     "Tools/inductive_package";
6112
5e4871c5136b datatype package improvements
paulson
parents: 6071
diff changeset
    30
use     "Tools/induct_tacs";
5e4871c5136b datatype package improvements
paulson
parents: 6071
diff changeset
    31
use     "Tools/primrec_package";
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5529
diff changeset
    32
use_thy "QUniv";
6112
5e4871c5136b datatype package improvements
paulson
parents: 6071
diff changeset
    33
use     "Tools/datatype_package";
5529
4a54acae6a15 tidying
paulson
parents: 5511
diff changeset
    34
4a54acae6a15 tidying
paulson
parents: 5511
diff changeset
    35
(*the all-in-one theory*)
9176
8f975d9c1046 simplified slightly by using dependencies better in theories
paulson
parents: 8812
diff changeset
    36
with_path "Integ" use_thy "Main";
5529
4a54acae6a15 tidying
paulson
parents: 5511
diff changeset
    37
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6349
diff changeset
    38
simpset_ref() := simpset() setmksimps (map mk_eq o Ord_atomize o gen_all);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6349
diff changeset
    39
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
print_depth 8;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
5511
7f52fb755581 leaves subgoal package empty
paulson
parents: 4271
diff changeset
    42
Goal "True";  (*leave subgoal package empty*)