ROOT.ML
author lcp
Thu, 06 Apr 1995 11:27:54 +0200
changeset 241 b67c8e01ae04
parent 240 2209eb5bb56b
child 251 f04b33ce250f
permissions -rw-r--r--
Removed the "exit 1" calls, since now the Makefile does them.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
     1
(*  Title:      HOL/ROOT.ML
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
     3
    Author:     Tobias Nipkow
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     5
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
     6
Adds Classical Higher-order Logic to a database containing Pure Isabelle.
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     7
Should be executed in the subdirectory HOL.
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     8
*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     9
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    10
val banner = "Higher-Order Logic";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    11
writeln banner;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    12
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    13
print_depth 1;
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    14
set eta_contract;
10
7972e16d2dd3 changes for new Readthy
clasohm
parents: 0
diff changeset
    15
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    16
(* Add user sections *)
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    17
use "../Pure/section_utils.ML";
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    18
use "thy_syntax.ML";
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    19
17
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    20
use_thy "HOL";
240
2209eb5bb56b Set up for new hyp_subst_tac.
lcp
parents: 196
diff changeset
    21
use "../Provers/simplifier.ML";
2209eb5bb56b Set up for new hyp_subst_tac.
lcp
parents: 196
diff changeset
    22
use "../Provers/splitter.ML";
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    23
use "../Provers/hypsubst.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    24
use "../Provers/classical.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    25
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    26
(** Applying HypsubstFun to generate hyp_subst_tac **)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    27
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    28
structure Hypsubst_Data =
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    29
  struct
240
2209eb5bb56b Set up for new hyp_subst_tac.
lcp
parents: 196
diff changeset
    30
  structure Simplifier = Simplifier
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    31
  (*Take apart an equality judgement; otherwise raise Match!*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    32
  fun dest_eq (Const("Trueprop",_) $ (Const("op =",_)  $ t $ u)) = (t,u);
240
2209eb5bb56b Set up for new hyp_subst_tac.
lcp
parents: 196
diff changeset
    33
  val eq_reflection = eq_reflection
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    34
  val imp_intr = impI
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    35
  val rev_mp = rev_mp
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    36
  val subst = subst
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    37
  val sym = sym
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    38
  end;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    39
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    40
structure Hypsubst = HypsubstFun(Hypsubst_Data);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    41
open Hypsubst;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    42
166
c59c471126ab HOL/ROOT/HOL_dup_cs: removed as obsolete
lcp
parents: 163
diff changeset
    43
(*** Applying ClassicalFun to create a classical prover ***)
c59c471126ab HOL/ROOT/HOL_dup_cs: removed as obsolete
lcp
parents: 163
diff changeset
    44
structure Classical_Data = 
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    45
  struct
166
c59c471126ab HOL/ROOT/HOL_dup_cs: removed as obsolete
lcp
parents: 163
diff changeset
    46
  val sizef	= size_of_thm
c59c471126ab HOL/ROOT/HOL_dup_cs: removed as obsolete
lcp
parents: 163
diff changeset
    47
  val mp	= mp
c59c471126ab HOL/ROOT/HOL_dup_cs: removed as obsolete
lcp
parents: 163
diff changeset
    48
  val not_elim	= notE
c59c471126ab HOL/ROOT/HOL_dup_cs: removed as obsolete
lcp
parents: 163
diff changeset
    49
  val classical	= classical
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    50
  val hyp_subst_tacs=[hyp_subst_tac]
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    51
  end;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    52
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    53
structure Classical = ClassicalFun(Classical_Data);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    54
open Classical;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    55
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    56
(*Propositional rules*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    57
val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI]
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    58
                       addSEs [conjE,disjE,impCE,FalseE,iffE];
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    59
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    60
(*Quantifier rules*)
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    61
val HOL_cs = prop_cs addSIs [allI] addIs [exI,ex1I]
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    62
                     addSEs [exE,ex1E] addEs [allE];
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    63
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    64
use     "simpdata.ML";
17
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    65
use_thy "Ord";
119
93dc86ccee28 HOL/Makefile, ROOT: updated for new .thy files
lcp
parents: 101
diff changeset
    66
use_thy "subset";
93dc86ccee28 HOL/Makefile, ROOT: updated for new .thy files
lcp
parents: 101
diff changeset
    67
use_thy "equalities";
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    68
use     "hologic.ML";
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    69
use     "subtype.ML";
17
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    70
use_thy "Prod";
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    71
use_thy "Sum";
128
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    72
use_thy "Gfp";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    73
use_thy "Nat";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    74
129
0bba840aa07c New version of datatype.ML with primrec (Norbert).
nipkow
parents: 128
diff changeset
    75
use "datatype.ML";
128
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    76
use "ind_syntax.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    77
use "add_ind_def.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    78
use "intr_elim.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    79
use "indrule.ML";
183
74ce9774c923 Removed obsolete induction package
nipkow
parents: 173
diff changeset
    80
use_thy "Inductive";
128
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    81
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    82
use_thy "Finite";
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 183
diff changeset
    83
use_thy "Sexp";
148
13b15899c528 moved LList* to ex
nipkow
parents: 129
diff changeset
    84
use_thy "List";
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    85
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    86
init_pps ();
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    87
print_depth 8;
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    88
163
edadccb76178 rearranged theory section stuff;
wenzelm
parents: 148
diff changeset
    89
val HOL_build_completed = ();   (*indicate successful build*)