src/Pure/ML_Root.thy
author wenzelm
Tue, 05 Apr 2016 21:51:14 +0200
changeset 62880 76e7d9169b54
parent 62873 src/Pure/ML/ML_Root.thy@2f9c8a18f832
child 62881 20b0cb2f0b87
permissions -rw-r--r--
clarified files;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62880
76e7d9169b54 clarified files;
wenzelm
parents: 62873
diff changeset
     1
(*  Title:      Pure/ML_Root.thy
62868
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
     3
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
     4
Support for ML project ROOT file, with imitation of ML "use" commands.
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
     5
*)
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
     6
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
     7
theory ML_Root
62880
76e7d9169b54 clarified files;
wenzelm
parents: 62873
diff changeset
     8
imports Pure
62868
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
     9
keywords "use" "use_debug" "use_no_debug" :: thy_load
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    10
  and "use_thy" :: thy_load
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    11
begin
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    12
62873
2f9c8a18f832 support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents: 62871
diff changeset
    13
setup \<open>Context.theory_map ML_Env.init_bootstrap\<close>
2f9c8a18f832 support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
wenzelm
parents: 62871
diff changeset
    14
62868
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    15
ML \<open>
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    16
local
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    17
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    18
val _ =
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    19
  Outer_Syntax.command @{command_keyword use}
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    20
    "read and evaluate Isabelle/ML or SML file"
62871
4a6cbe1239fe proper syntax;
wenzelm
parents: 62868
diff changeset
    21
    (Resources.parse_files "use" --| @{keyword ";"} >> ML_File.use NONE);
62868
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    22
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    23
val _ =
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    24
  Outer_Syntax.command @{command_keyword use_debug}
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    25
    "read and evaluate Isabelle/ML or SML file (with debugger information)"
62871
4a6cbe1239fe proper syntax;
wenzelm
parents: 62868
diff changeset
    26
    (Resources.parse_files "use_debug" --| @{keyword ";"} >> ML_File.use (SOME true));
62868
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    27
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    28
val _ =
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    29
  Outer_Syntax.command @{command_keyword use_no_debug}
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    30
    "read and evaluate Isabelle/ML or SML file (no debugger information)"
62871
4a6cbe1239fe proper syntax;
wenzelm
parents: 62868
diff changeset
    31
    (Resources.parse_files "use_no_debug" --| @{keyword ";"} >> ML_File.use (SOME false));
62868
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    32
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    33
val _ =
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    34
  Outer_Syntax.command @{command_keyword use_thy} "undefined dummy command"
62871
4a6cbe1239fe proper syntax;
wenzelm
parents: 62868
diff changeset
    35
    (Parse.path -- @{keyword ";"} >> (fn _ =>
4a6cbe1239fe proper syntax;
wenzelm
parents: 62868
diff changeset
    36
      Toplevel.keep (fn _ => error "Undefined command 'use_thy'")));
62868
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    37
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    38
in end
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    39
\<close>
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    40
61a691db1c4d support for ML project ROOT file, with imitation of ML "use" commands;
wenzelm
parents:
diff changeset
    41
end