| author | huffman | 
| Tue, 24 Feb 2009 11:10:05 -0800 | |
| changeset 30081 | 46b9c8ae3897 | 
| parent 29269 | 5c25a2012975 | 
| child 30173 | eabece26b89b | 
| permissions | -rw-r--r-- | 
| 19 | 1  | 
(* Title: Pure/ROOT.ML  | 
| 0 | 2  | 
|
| 16842 | 3  | 
Pure Isabelle.  | 
| 0 | 4  | 
*)  | 
5  | 
||
| 
26109
 
c69c3559355b
more elaborate structure Distribution (filled-in by makedist);
 
wenzelm 
parents: 
25953 
diff
changeset
 | 
6  | 
structure Distribution = (*filled-in by makedist*)  | 
| 
 
c69c3559355b
more elaborate structure Distribution (filled-in by makedist);
 
wenzelm 
parents: 
25953 
diff
changeset
 | 
7  | 
struct  | 
| 
 
c69c3559355b
more elaborate structure Distribution (filled-in by makedist);
 
wenzelm 
parents: 
25953 
diff
changeset
 | 
8  | 
val version = "Isabelle repository version";  | 
| 
27642
 
c0db1220b071
structure Distribution: swapped default for is_official;
 
wenzelm 
parents: 
27546 
diff
changeset
 | 
9  | 
val is_official = false;  | 
| 
 
c0db1220b071
structure Distribution: swapped default for is_official;
 
wenzelm 
parents: 
27546 
diff
changeset
 | 
10  | 
val changelog = "";  | 
| 
26109
 
c69c3559355b
more elaborate structure Distribution (filled-in by makedist);
 
wenzelm 
parents: 
25953 
diff
changeset
 | 
11  | 
end;  | 
| 11835 | 12  | 
|
| 
23825
 
e0372eba47b7
simplified loading of ML files -- no static forward references;
 
wenzelm 
parents: 
23696 
diff
changeset
 | 
13  | 
(*if true then some tools will OMIT some proofs*)  | 
| 
 
e0372eba47b7
simplified loading of ML files -- no static forward references;
 
wenzelm 
parents: 
23696 
diff
changeset
 | 
14  | 
val quick_and_dirty = ref false;  | 
| 
 
e0372eba47b7
simplified loading of ML files -- no static forward references;
 
wenzelm 
parents: 
23696 
diff
changeset
 | 
15  | 
|
| 12248 | 16  | 
print_depth 10;  | 
| 0 | 17  | 
|
| 
5017
 
786a17461ab9
moved table.ML, object.ML, seq.ML, name_space.ML to General;
 
wenzelm 
parents: 
5004 
diff
changeset
 | 
18  | 
(*basic tools*)  | 
| 21396 | 19  | 
use "General/basics.ML";  | 
| 0 | 20  | 
use "library.ML";  | 
| 22592 | 21  | 
|
| 
5017
 
786a17461ab9
moved table.ML, object.ML, seq.ML, name_space.ML to General;
 
wenzelm 
parents: 
5004 
diff
changeset
 | 
22  | 
cd "General"; use "ROOT.ML"; cd "..";  | 
| 28200 | 23  | 
cd "Concurrent"; use "ROOT.ML"; cd "..";  | 
| 28120 | 24  | 
|
| 14781 | 25  | 
(*fundamental structures*)  | 
| 20075 | 26  | 
use "name.ML";  | 
| 0 | 27  | 
use "term.ML";  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29263 
diff
changeset
 | 
28  | 
use "term_ord.ML";  | 
| 20507 | 29  | 
use "term_subst.ML";  | 
| 29263 | 30  | 
use "old_term.ML";  | 
| 24257 | 31  | 
use "logic.ML";  | 
| 
14823
 
ebb8499d0fd2
moved print_mode to General/output.ML; load General/pretty.ML early;
 
wenzelm 
parents: 
14781 
diff
changeset
 | 
32  | 
use "General/pretty.ML";  | 
| 28404 | 33  | 
use "context.ML";  | 
34  | 
use "context_position.ML";  | 
|
| 24235 | 35  | 
use "Syntax/lexicon.ML";  | 
36  | 
use "Syntax/simple_syntax.ML";  | 
|
| 24272 | 37  | 
use "sorts.ML";  | 
38  | 
use "type.ML";  | 
|
| 24113 | 39  | 
use "config.ML";  | 
| 19 | 40  | 
|
| 4949 | 41  | 
(*inner syntax module*)  | 
| 22679 | 42  | 
use "Syntax/ast.ML";  | 
43  | 
use "Syntax/syn_ext.ML";  | 
|
44  | 
use "Syntax/parser.ML";  | 
|
45  | 
use "Syntax/type_ext.ML";  | 
|
46  | 
use "Syntax/syn_trans.ML";  | 
|
47  | 
use "Syntax/mixfix.ML";  | 
|
48  | 
use "Syntax/printer.ML";  | 
|
49  | 
use "Syntax/syntax.ML";  | 
|
50  | 
||
| 27262 | 51  | 
use "type_infer.ML";  | 
| 24574 | 52  | 
use "ML/ml_syntax.ML";  | 
| 0 | 53  | 
|
| 
15825
 
1576f9d3ffae
added content of Pure/Thy/ROOT.ML, Pure/Proof/ROOT.ML;
 
wenzelm 
parents: 
15801 
diff
changeset
 | 
54  | 
(*core of tactical proof system*)  | 
| 18934 | 55  | 
use "envir.ML";  | 
| 18059 | 56  | 
use "consts.ML";  | 
| 24257 | 57  | 
use "primitive_defs.ML";  | 
| 27546 | 58  | 
use "defs.ML";  | 
59  | 
use "net.ML";  | 
|
| 0 | 60  | 
use "sign.ML";  | 
61  | 
use "pattern.ML";  | 
|
62  | 
use "unify.ML";  | 
|
| 1528 | 63  | 
use "theory.ML";  | 
| 24664 | 64  | 
use "interpretation.ML";  | 
| 11511 | 65  | 
use "proofterm.ML";  | 
| 0 | 66  | 
use "thm.ML";  | 
| 22361 | 67  | 
use "more_thm.ML";  | 
| 26279 | 68  | 
use "facts.ML";  | 
| 3986 | 69  | 
use "pure_thy.ML";  | 
| 19589 | 70  | 
use "display.ML";  | 
| 0 | 71  | 
use "drule.ML";  | 
| 22233 | 72  | 
use "morphism.ML";  | 
| 19898 | 73  | 
use "variable.ML";  | 
| 24833 | 74  | 
use "conv.ML";  | 
| 0 | 75  | 
use "tctical.ML";  | 
| 1582 | 76  | 
use "search.ML";  | 
| 21708 | 77  | 
use "tactic.ML";  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
14823 
diff
changeset
 | 
78  | 
use "meta_simplifier.ML";  | 
| 19417 | 79  | 
use "conjunction.ML";  | 
| 20225 | 80  | 
use "assumption.ML";  | 
| 17963 | 81  | 
use "goal.ML";  | 
| 24963 | 82  | 
use "axclass.ML";  | 
| 0 | 83  | 
|
| 22108 | 84  | 
(*the main Isar system*)  | 
| 6178 | 85  | 
cd "Isar"; use "ROOT.ML"; cd "..";  | 
| 20207 | 86  | 
use "subgoal.ML";  | 
| 5834 | 87  | 
|
| 13402 | 88  | 
use "Proof/extraction.ML";  | 
| 11511 | 89  | 
|
| 17467 | 90  | 
cd "Tools"; use "ROOT.ML"; cd "..";  | 
91  | 
||
| 24455 | 92  | 
use "codegen.ML";  | 
93  | 
||
| 
12778
 
3120e338ffae
Interface/proof_general.ML move to proof_general.ML;
 
wenzelm 
parents: 
12248 
diff
changeset
 | 
94  | 
(*configuration for Proof General*)  | 
| 21941 | 95  | 
cd "ProofGeneral"; use "ROOT.ML"; cd "..";  | 
| 16781 | 96  | 
|
| 
23825
 
e0372eba47b7
simplified loading of ML files -- no static forward references;
 
wenzelm 
parents: 
23696 
diff
changeset
 | 
97  | 
use "pure_setup.ML";  |