ROOT.ML
author wenzelm
Wed, 21 Sep 1994 15:40:41 +0200
changeset 145 a9f7ff3a464c
parent 129 0bba840aa07c
child 148 13b15899c528
permissions -rw-r--r--
minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/ROOT
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     3
    Author: 	Tobias Nipkow
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     5
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     6
Adds Classical Higher-order Logic to a database containing pure Isabelle. 
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
128
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    13
init_thy_reader();
10
7972e16d2dd3 changes for new Readthy
clasohm
parents: 0
diff changeset
    14
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    15
print_depth 1;  
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    16
eta_contract := true;
17
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    17
use_thy "HOL";
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    18
use "../Provers/hypsubst.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    19
use "../Provers/classical.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    20
use "../Provers/simplifier.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    21
use "../Provers/splitter.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    22
use "../Provers/ind.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    23
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    24
(** Applying HypsubstFun to generate hyp_subst_tac **)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    25
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    26
structure Hypsubst_Data =
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    27
  struct
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    28
  (*Take apart an equality judgement; otherwise raise Match!*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    29
  fun dest_eq (Const("Trueprop",_) $ (Const("op =",_)  $ t $ u)) = (t,u);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    30
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    31
  val imp_intr = impI
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    32
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    33
  (*etac rev_cut_eq moves an equality to be the last premise. *)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    34
  val rev_cut_eq = prove_goal HOL.thy "[| a=b;  a=b ==> R |] ==> R"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    35
   (fn prems => [ REPEAT(resolve_tac prems 1) ]);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    36
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    37
  val rev_mp = rev_mp
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    38
  val subst = subst 
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    39
  val sym = sym
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    40
  end;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    41
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    42
structure Hypsubst = HypsubstFun(Hypsubst_Data);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    43
open Hypsubst;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    44
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    45
(** Applying ClassicalFun to create a classical prover **)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    46
structure Classical_Data = 
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    47
  struct
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    48
  val sizef = size_of_thm
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    49
  val mp = mp
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    50
  val not_elim = notE
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    51
  val swap = swap
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    52
  val hyp_subst_tacs=[hyp_subst_tac]
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    53
  end;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    54
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    55
structure Classical = ClassicalFun(Classical_Data);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    56
open Classical;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    57
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    58
(*Propositional rules*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    59
val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI]
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    60
                       addSEs [conjE,disjE,impCE,FalseE,iffE];
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    61
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    62
(*Quantifier rules*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    63
val HOL_cs = prop_cs addSIs [allI] addIs [exI,ex1I] 
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    64
                     addSEs [exE,ex1E] addEs [allE];
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    65
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    66
val HOL_dup_cs = prop_cs addSIs [allI] addIs [exCI,ex1I] 
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    67
                         addSEs [exE,ex1E] addEs [all_dupE];
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    68
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    69
structure HOL_Induction = InductionFun(struct val spec=spec end);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    70
open HOL_Induction;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    71
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    72
use     "simpdata.ML";
17
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    73
use_thy "Ord";
119
93dc86ccee28 HOL/Makefile, ROOT: updated for new .thy files
lcp
parents: 101
diff changeset
    74
use_thy "subset";
93dc86ccee28 HOL/Makefile, ROOT: updated for new .thy files
lcp
parents: 101
diff changeset
    75
use_thy "equalities";
17
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    76
use_thy "Prod";
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
    77
use_thy "Sum";
128
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    78
use_thy "Gfp";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    79
use_thy "Nat";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    80
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    81
(* Add user sections *)
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    82
use "../Pure/section_utils.ML";
129
0bba840aa07c New version of datatype.ML with primrec (Norbert).
nipkow
parents: 128
diff changeset
    83
use "datatype.ML";
128
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    84
use "ind_syntax.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    85
use "add_ind_def.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    86
use "intr_elim.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    87
use "indrule.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    88
use "Inductive.ML";
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    89
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    90
structure ThySyn = ThySynFun
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    91
 (val user_keywords = ["|", "datatype", "primrec",
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    92
		       "inductive", "coinductive", "intrs", 
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    93
		       "monos", "con_defs"]
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    94
  and user_sections = [("inductive",  inductive_decl ""),
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    95
		       ("coinductive",  inductive_decl "Co"),
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    96
		       ("datatype",  datatype_decls),
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    97
		       ("primrec", primrec_decl)]);
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    98
init_thy_reader ();
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
    99
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 119
diff changeset
   100
use_thy "Finite";
17
8c0c3a67d9e1 changed use_thy's parameter to exact theory name
clasohm
parents: 10
diff changeset
   101
use_thy "LList";
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   102
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   103
use "../Pure/install_pp.ML";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   104
print_depth 8;  
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   105
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   106
val HOL_build_completed = ();	(*indicate successful build*)