src/Pure/Pure.thy
author wenzelm
Mon, 04 Apr 2016 22:13:08 +0200
changeset 62856 3f97aa4580c6
parent 62849 caaa2fc4040d
child 62859 b2f951051472
permissions -rw-r--r--
tuned -- more explicit sections;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48929
05d4e5f660ae entity markup for theory Pure, to enable hyperlinks etc.;
wenzelm
parents: 48891
diff changeset
     1
(*  Title:      Pure/Pure.thy
05d4e5f660ae entity markup for theory Pure, to enable hyperlinks etc.;
wenzelm
parents: 48891
diff changeset
     2
    Author:     Makarius
05d4e5f660ae entity markup for theory Pure, to enable hyperlinks etc.;
wenzelm
parents: 48891
diff changeset
     3
05d4e5f660ae entity markup for theory Pure, to enable hyperlinks etc.;
wenzelm
parents: 48891
diff changeset
     4
Final stage of bootstrapping Pure, based on implicit background theory.
05d4e5f660ae entity markup for theory Pure, to enable hyperlinks etc.;
wenzelm
parents: 48891
diff changeset
     5
*)
05d4e5f660ae entity markup for theory Pure, to enable hyperlinks etc.;
wenzelm
parents: 48891
diff changeset
     6
48638
22d65e375c01 more standard bootstrapping of Pure.thy;
wenzelm
parents: 29606
diff changeset
     7
theory Pure
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
     8
  keywords
61579
634cd44bb1d3 symbolic syntax "\<comment> text";
wenzelm
parents: 61566
diff changeset
     9
    "!!" "!" "+" "--" ":" ";" "<" "<=" "=" "=>" "?" "[" "\<comment>" "\<equiv>"
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    10
    "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>"
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    11
    "\<subseteq>" "]" "assumes" "attach" "binder" "constrains"
61671
20d4cd2ceab2 prefer "rewrites" and "defines" to note rewrite morphisms
haftmann
parents: 61670
diff changeset
    12
    "defines" "rewrites" "fixes" "for" "identifier" "if" "in" "includes" "infix"
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58918
diff changeset
    13
    "infixl" "infixr" "is" "notes" "obtains" "open" "output"
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61338
diff changeset
    14
    "overloaded" "pervasive" "premises" "private" "qualified" "rewrites"
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61338
diff changeset
    15
    "shows" "structure" "unchecked" "where" "when" "|"
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58928
diff changeset
    16
  and "text" "txt" :: document_body
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58928
diff changeset
    17
  and "text_raw" :: document_raw
57506
f5dbec155914 suppress completion of obscure keyword;
wenzelm
parents: 57442
diff changeset
    18
  and "default_sort" :: thy_decl == ""
f5dbec155914 suppress completion of obscure keyword;
wenzelm
parents: 57442
diff changeset
    19
  and "typedecl" "type_synonym" "nonterminal" "judgment"
62169
a6047f511de7 removed old 'defs' command;
wenzelm
parents: 61890
diff changeset
    20
    "consts" "syntax" "no_syntax" "translations" "no_translations"
55385
169e12bbf9a3 discontinued axiomatic 'classes', 'classrel', 'arities';
wenzelm
parents: 55152
diff changeset
    21
    "definition" "abbreviation" "type_notation" "no_type_notation" "notation"
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 61252
diff changeset
    22
    "no_notation" "axiomatization" "lemmas" "declare"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    23
    "hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
    24
  and "ML_file" "ML_file_debug" "ML_file_no_debug" :: thy_load % "ML"
60957
574254152856 support for ML files with/without debugger information;
wenzelm
parents: 60749
diff changeset
    25
  and "SML_file" "SML_file_debug" "SML_file_no_debug" :: thy_load % "ML"
56618
874bdedb2313 added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
wenzelm
parents: 56275
diff changeset
    26
  and "SML_import" "SML_export" :: thy_decl % "ML"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    27
  and "ML_prf" :: prf_decl % "proof"  (* FIXME % "ML" ?? *)
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    28
  and "ML_val" "ML_command" :: diag % "ML"
55762
27a45aec67a0 suppress completion of obscure keyword, avoid confusion with plain "simp";
wenzelm
parents: 55516
diff changeset
    29
  and "simproc_setup" :: thy_decl % "ML" == ""
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    30
  and "setup" "local_setup" "attribute_setup" "method_setup"
55762
27a45aec67a0 suppress completion of obscure keyword, avoid confusion with plain "simp";
wenzelm
parents: 55516
diff changeset
    31
    "declaration" "syntax_declaration"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    32
    "parse_ast_translation" "parse_translation" "print_translation"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    33
    "typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    34
  and "bundle" :: thy_decl
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    35
  and "include" "including" :: prf_decl
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    36
  and "print_bundles" :: diag
59901
840d03805755 added command 'experiment';
wenzelm
parents: 58999
diff changeset
    37
  and "context" "locale" "experiment" :: thy_decl_block
51224
c3e99efacb67 back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
wenzelm
parents: 50603
diff changeset
    38
  and "interpret" :: prf_goal % "proof"
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61853
diff changeset
    39
  and "interpretation" "global_interpretation" "sublocale" :: thy_goal
58800
bfed1c26caed explicit keyword category for commands that may start a block;
wenzelm
parents: 58660
diff changeset
    40
  and "class" :: thy_decl_block
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    41
  and "subclass" :: thy_goal
58800
bfed1c26caed explicit keyword category for commands that may start a block;
wenzelm
parents: 58660
diff changeset
    42
  and "instantiation" :: thy_decl_block
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    43
  and "instance" :: thy_goal
58800
bfed1c26caed explicit keyword category for commands that may start a block;
wenzelm
parents: 58660
diff changeset
    44
  and "overloading" :: thy_decl_block
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    45
  and "code_datatype" :: thy_decl
61338
de610e8df459 added 'proposition' command;
wenzelm
parents: 61337
diff changeset
    46
  and "theorem" "lemma" "corollary" "proposition" :: thy_goal
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 61252
diff changeset
    47
  and "schematic_goal" :: thy_goal
58800
bfed1c26caed explicit keyword category for commands that may start a block;
wenzelm
parents: 58660
diff changeset
    48
  and "notepad" :: thy_decl_block
50128
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49569
diff changeset
    49
  and "have" :: prf_goal % "proof"
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49569
diff changeset
    50
  and "hence" :: prf_goal % "proof" == "then have"
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49569
diff changeset
    51
  and "show" :: prf_asm_goal % "proof"
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49569
diff changeset
    52
  and "thus" :: prf_asm_goal % "proof" == "then show"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    53
  and "then" "from" "with" :: prf_chain % "proof"
60371
8a5cfdda1b98 added Isar command 'supply';
wenzelm
parents: 60317
diff changeset
    54
  and "note" :: prf_decl % "proof"
8a5cfdda1b98 added Isar command 'supply';
wenzelm
parents: 60317
diff changeset
    55
  and "supply" :: prf_script % "proof"
8a5cfdda1b98 added Isar command 'supply';
wenzelm
parents: 60317
diff changeset
    56
  and "using" "unfolding" :: prf_decl % "proof"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    57
  and "fix" "assume" "presume" "def" :: prf_asm % "proof"
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60371
diff changeset
    58
  and "consider" :: prf_goal % "proof"
53371
47b23c582127 more explicit indication of 'guess' as improper Isar (aka "script") element;
wenzelm
parents: 52549
diff changeset
    59
  and "obtain" :: prf_asm_goal % "proof"
60624
5b6552e12421 clarified keyword categories;
wenzelm
parents: 60623
diff changeset
    60
  and "guess" :: prf_script_asm_goal % "proof"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    61
  and "let" "write" :: prf_decl % "proof"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    62
  and "case" :: prf_asm % "proof"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    63
  and "{" :: prf_open % "proof"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    64
  and "}" :: prf_close % "proof"
60694
b3fa4a8cdb5f clarified text folds: proof ... qed counts as extra block;
wenzelm
parents: 60624
diff changeset
    65
  and "next" :: next_block % "proof"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    66
  and "qed" :: qed_block % "proof"
62312
5e5a881ebc12 command '\<proof>' is an alias for 'sorry', with different typesetting;
wenzelm
parents: 62169
diff changeset
    67
  and "by" ".." "." "sorry" "\<proof>" :: "qed" % "proof"
53571
e58ca0311c0f more explicit indication of 'done' as proof script element;
wenzelm
parents: 53371
diff changeset
    68
  and "done" :: "qed_script" % "proof"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    69
  and "oops" :: qed_global % "proof"
50128
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49569
diff changeset
    70
  and "defer" "prefer" "apply" :: prf_script % "proof"
599c935aac82 alternative completion for outer syntax keywords;
wenzelm
parents: 49569
diff changeset
    71
  and "apply_end" :: prf_script % "proof" == ""
60624
5b6552e12421 clarified keyword categories;
wenzelm
parents: 60623
diff changeset
    72
  and "subgoal" :: prf_script_goal % "proof"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    73
  and "proof" :: prf_block % "proof"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    74
  and "also" "moreover" :: prf_decl % "proof"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    75
  and "finally" "ultimately" :: prf_chain % "proof"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    76
  and "back" :: prf_script % "proof"
61252
c165f0472d57 separate command 'print_definitions';
wenzelm
parents: 60957
diff changeset
    77
  and "help" "print_commands" "print_options" "print_context" "print_theory"
c165f0472d57 separate command 'print_definitions';
wenzelm
parents: 60957
diff changeset
    78
    "print_definitions" "print_syntax" "print_abbrevs" "print_defn_rules"
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    79
    "print_theorems" "print_locales" "print_classes" "print_locale"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    80
    "print_interps" "print_dependencies" "print_attributes"
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    81
    "print_simpset" "print_rules" "print_trans_rules" "print_methods"
56069
451d5b73f8cf simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents: 55762
diff changeset
    82
    "print_antiquotations" "print_ML_antiquotations" "thy_deps"
58845
8451eddc4d67 removed obsolete Proof General commands;
wenzelm
parents: 58842
diff changeset
    83
    "locale_deps" "class_deps" "thm_deps" "print_term_bindings"
57415
e721124f1b1e command 'print_term_bindings' supersedes 'print_binds';
wenzelm
parents: 56864
diff changeset
    84
    "print_facts" "print_cases" "print_statement" "thm" "prf" "full_prf"
e721124f1b1e command 'print_term_bindings' supersedes 'print_binds';
wenzelm
parents: 56864
diff changeset
    85
    "prop" "term" "typ" "print_codesetup" "unused_thms" :: diag
58845
8451eddc4d67 removed obsolete Proof General commands;
wenzelm
parents: 58842
diff changeset
    86
  and "display_drafts" "print_state" :: diag
48646
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 48641
diff changeset
    87
  and "welcome" :: diag
48641
92b48b8abfe4 more standard bootstrapping of Pure outer syntax;
wenzelm
parents: 48638
diff changeset
    88
  and "end" :: thy_end % "theory"
56797
32963b43a538 suppress slightly odd completions of "real";
wenzelm
parents: 56618
diff changeset
    89
  and "realizers" :: thy_decl == ""
32963b43a538 suppress slightly odd completions of "real";
wenzelm
parents: 56618
diff changeset
    90
  and "realizability" :: thy_decl == ""
32963b43a538 suppress slightly odd completions of "real";
wenzelm
parents: 56618
diff changeset
    91
  and "extract_type" "extract" :: thy_decl
48646
91281e9472d8 more official command specifications, including source position;
wenzelm
parents: 48641
diff changeset
    92
  and "find_theorems" "find_consts" :: diag
57886
7cae177c9084 support for named collections of theorems in canonical order;
wenzelm
parents: 57506
diff changeset
    93
  and "named_theorems" :: thy_decl
48638
22d65e375c01 more standard bootstrapping of Pure.thy;
wenzelm
parents: 29606
diff changeset
    94
begin
15803
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
    95
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
    96
section \<open>Isar commands\<close>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
    97
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
    98
subsection \<open>Embedded ML text\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
    99
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   100
ML \<open>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   101
local
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   102
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   103
fun ML_file_cmd debug files = Toplevel.generic_theory (fn gthy =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   104
  let
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   105
    val [{src_path, lines, digest, pos}: Token.file] = files (Context.theory_of gthy);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   106
    val provide = Resources.provide (src_path, digest);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   107
    val source = Input.source true (cat_lines lines) (pos, pos);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   108
    val flags: ML_Compiler.flags =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   109
      {SML = false, exchange = false, redirect = true, verbose = true,
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   110
        debug = debug, writeln = writeln, warning = warning};
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   111
  in
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   112
    gthy
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   113
    |> ML_Context.exec (fn () => ML_Context.eval_source flags source)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   114
    |> Local_Theory.propagate_ml_env
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   115
    |> Context.mapping provide (Local_Theory.background_theory provide)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   116
  end);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   117
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   118
fun SML_file_cmd debug files = Toplevel.theory (fn thy =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   119
  let
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   120
    val ([{lines, pos, ...}: Token.file], thy') = files thy;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   121
    val source = Input.source true (cat_lines lines) (pos, pos);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   122
    val flags: ML_Compiler.flags =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   123
      {SML = true, exchange = false, redirect = true, verbose = true,
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   124
        debug = debug, writeln = writeln, warning = warning};
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   125
  in
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   126
    thy' |> Context.theory_map
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   127
      (ML_Context.exec (fn () => ML_Context.eval_source flags source))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   128
  end);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   129
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   130
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   131
  Outer_Syntax.command ("ML_file", @{here}) "read and evaluate Isabelle/ML file"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   132
    (Resources.parse_files "ML_file" >> ML_file_cmd NONE);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   133
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   134
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   135
  Outer_Syntax.command ("ML_file_debug", @{here})
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   136
    "read and evaluate Isabelle/ML file (with debugger information)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   137
    (Resources.parse_files "ML_file_debug" >> ML_file_cmd (SOME true));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   138
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   139
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   140
  Outer_Syntax.command ("ML_file_no_debug", @{here})
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   141
    "read and evaluate Isabelle/ML file (no debugger information)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   142
    (Resources.parse_files "ML_file_no_debug" >> ML_file_cmd (SOME false));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   143
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   144
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   145
  Outer_Syntax.command @{command_keyword SML_file} "read and evaluate Standard ML file"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   146
    (Resources.provide_parse_files "SML_file" >> SML_file_cmd NONE);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   147
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   148
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   149
  Outer_Syntax.command @{command_keyword SML_file_debug}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   150
    "read and evaluate Standard ML file (with debugger information)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   151
    (Resources.provide_parse_files "SML_file_debug" >> SML_file_cmd (SOME true));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   152
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   153
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   154
  Outer_Syntax.command @{command_keyword SML_file_no_debug}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   155
    "read and evaluate Standard ML file (no debugger information)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   156
    (Resources.provide_parse_files "SML_file_no_debug" >> SML_file_cmd (SOME false));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   157
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   158
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   159
  Outer_Syntax.command @{command_keyword SML_export} "evaluate SML within Isabelle/ML environment"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   160
    (Parse.ML_source >> (fn source =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   161
      let
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   162
        val flags: ML_Compiler.flags =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   163
          {SML = true, exchange = true, redirect = false, verbose = true,
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   164
            debug = NONE, writeln = writeln, warning = warning};
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   165
      in
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   166
        Toplevel.theory
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   167
          (Context.theory_map (ML_Context.exec (fn () => ML_Context.eval_source flags source)))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   168
      end));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   169
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   170
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   171
  Outer_Syntax.command @{command_keyword SML_import} "evaluate Isabelle/ML within SML environment"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   172
    (Parse.ML_source >> (fn source =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   173
      let
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   174
        val flags: ML_Compiler.flags =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   175
          {SML = false, exchange = true, redirect = false, verbose = true,
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   176
            debug = NONE, writeln = writeln, warning = warning};
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   177
      in
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   178
        Toplevel.generic_theory
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   179
          (ML_Context.exec (fn () => ML_Context.eval_source flags source) #>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   180
            Local_Theory.propagate_ml_env)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   181
      end));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   182
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   183
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   184
  Outer_Syntax.command @{command_keyword ML_prf} "ML text within proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   185
    (Parse.ML_source >> (fn source =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   186
      Toplevel.proof (Proof.map_context (Context.proof_map
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   187
        (ML_Context.exec (fn () =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   188
            ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source))) #>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   189
          Proof.propagate_ml_env)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   190
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   191
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   192
  Outer_Syntax.command @{command_keyword ML_val} "diagnostic ML text"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   193
    (Parse.ML_source >> Isar_Cmd.ml_diag true);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   194
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   195
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   196
  Outer_Syntax.command @{command_keyword ML_command} "diagnostic ML text (silent)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   197
    (Parse.ML_source >> Isar_Cmd.ml_diag false);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   198
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   199
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   200
  Outer_Syntax.command @{command_keyword setup} "ML setup for global theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   201
    (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.setup));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   202
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   203
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   204
  Outer_Syntax.local_theory @{command_keyword local_setup} "ML setup for local theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   205
    (Parse.ML_source >> Isar_Cmd.local_setup);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   206
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   207
val _ =
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   208
  Outer_Syntax.command @{command_keyword oracle} "declare oracle"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   209
    (Parse.range Parse.name -- (@{keyword "="} |-- Parse.ML_source) >>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   210
      (fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y)));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   211
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   212
val _ =
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   213
  Outer_Syntax.local_theory @{command_keyword attribute_setup} "define attribute in ML"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   214
    (Parse.position Parse.name --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   215
        Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   216
      >> (fn (name, (txt, cmt)) => Attrib.attribute_setup name txt cmt));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   217
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   218
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   219
  Outer_Syntax.local_theory @{command_keyword method_setup} "define proof method in ML"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   220
    (Parse.position Parse.name --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   221
        Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   222
      >> (fn (name, (txt, cmt)) => Method.method_setup name txt cmt));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   223
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   224
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   225
  Outer_Syntax.local_theory @{command_keyword declaration} "generic ML declaration"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   226
    (Parse.opt_keyword "pervasive" -- Parse.ML_source
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   227
      >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = false, pervasive = pervasive} txt));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   228
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   229
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   230
  Outer_Syntax.local_theory @{command_keyword syntax_declaration} "generic ML syntax declaration"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   231
    (Parse.opt_keyword "pervasive" -- Parse.ML_source
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   232
      >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = true, pervasive = pervasive} txt));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   233
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   234
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   235
  Outer_Syntax.local_theory @{command_keyword simproc_setup} "define simproc in ML"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   236
    (Parse.position Parse.name --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   237
      (@{keyword "("} |-- Parse.enum1 "|" Parse.term --| @{keyword ")"} --| @{keyword "="}) --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   238
      Parse.ML_source -- Scan.optional (@{keyword "identifier"} |-- Scan.repeat1 Parse.xname) []
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   239
    >> (fn (((a, b), c), d) => Isar_Cmd.simproc_setup a b c d));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   240
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   241
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   242
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   243
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   244
subsection \<open>Theory commands\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   245
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   246
subsubsection \<open>Sorts and types\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   247
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   248
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   249
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   250
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   251
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   252
  Outer_Syntax.local_theory @{command_keyword default_sort}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   253
    "declare default sort for explicit type variables"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   254
    (Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   255
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   256
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   257
  Outer_Syntax.local_theory @{command_keyword typedecl} "type declaration"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   258
    (Parse.type_args -- Parse.binding -- Parse.opt_mixfix
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   259
      >> (fn ((args, a), mx) =>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   260
          Typedecl.typedecl {final = true} (a, map (rpair dummyS) args, mx) #> snd));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   261
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   262
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   263
  Outer_Syntax.local_theory @{command_keyword type_synonym} "declare type abbreviation"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   264
    (Parse.type_args -- Parse.binding --
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   265
      (@{keyword "="} |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix'))
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   266
      >> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   267
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   268
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   269
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   270
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   271
subsubsection \<open>Consts\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   272
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   273
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   274
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   275
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   276
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   277
  Outer_Syntax.command @{command_keyword judgment} "declare object-logic judgment"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   278
    (Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   279
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   280
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   281
  Outer_Syntax.command @{command_keyword consts} "declare constants"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   282
    (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts_cmd));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   283
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   284
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   285
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   286
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   287
subsubsection \<open>Syntax and translations\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   288
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   289
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   290
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   291
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   292
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   293
  Outer_Syntax.command @{command_keyword nonterminal}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   294
    "declare syntactic type constructors (grammar nonterminal symbols)"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   295
    (Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   296
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   297
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   298
  Outer_Syntax.command @{command_keyword syntax} "add raw syntax clauses"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   299
    (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   300
      >> (Toplevel.theory o uncurry Sign.add_syntax_cmd));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   301
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   302
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   303
  Outer_Syntax.command @{command_keyword no_syntax} "delete raw syntax clauses"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   304
    (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   305
      >> (Toplevel.theory o uncurry Sign.del_syntax_cmd));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   306
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   307
val trans_pat =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   308
  Scan.optional
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   309
    (@{keyword "("} |-- Parse.!!! (Parse.inner_syntax Parse.xname --| @{keyword ")"})) "logic"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   310
    -- Parse.inner_syntax Parse.string;
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   311
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   312
fun trans_arrow toks =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   313
  ((@{keyword "\<rightharpoonup>"} || @{keyword "=>"}) >> K Syntax.Parse_Rule ||
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   314
    (@{keyword "\<leftharpoondown>"} || @{keyword "<="}) >> K Syntax.Print_Rule ||
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   315
    (@{keyword "\<rightleftharpoons>"} || @{keyword "=="}) >> K Syntax.Parse_Print_Rule) toks;
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   316
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   317
val trans_line =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   318
  trans_pat -- Parse.!!! (trans_arrow -- trans_pat)
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   319
    >> (fn (left, (arr, right)) => arr (left, right));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   320
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   321
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   322
  Outer_Syntax.command @{command_keyword translations} "add syntax translation rules"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   323
    (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.translations));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   324
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   325
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   326
  Outer_Syntax.command @{command_keyword no_translations} "delete syntax translation rules"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   327
    (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.no_translations));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   328
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   329
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   330
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   331
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   332
subsubsection \<open>Translation functions\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   333
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   334
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   335
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   336
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   337
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   338
  Outer_Syntax.command @{command_keyword parse_ast_translation}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   339
    "install parse ast translation functions"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   340
    (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_ast_translation));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   341
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   342
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   343
  Outer_Syntax.command @{command_keyword parse_translation}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   344
    "install parse translation functions"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   345
    (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_translation));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   346
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   347
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   348
  Outer_Syntax.command @{command_keyword print_translation}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   349
    "install print translation functions"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   350
    (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_translation));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   351
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   352
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   353
  Outer_Syntax.command @{command_keyword typed_print_translation}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   354
    "install typed print translation functions"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   355
    (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.typed_print_translation));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   356
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   357
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   358
  Outer_Syntax.command @{command_keyword print_ast_translation}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   359
    "install print ast translation functions"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   360
    (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_ast_translation));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   361
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   362
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   363
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   364
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   365
subsubsection \<open>Specifications\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   366
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   367
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   368
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   369
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   370
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   371
  Outer_Syntax.local_theory' @{command_keyword definition} "constant definition"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   372
    (Parse_Spec.constdef >> (fn args => #2 oo Specification.definition_cmd args));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   373
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   374
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   375
  Outer_Syntax.local_theory' @{command_keyword abbreviation} "constant abbreviation"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   376
    (Parse.syntax_mode -- (Scan.option Parse_Spec.constdecl -- Parse.prop)
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   377
      >> (fn (mode, args) => Specification.abbreviation_cmd mode args));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   378
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   379
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   380
  Outer_Syntax.command @{command_keyword axiomatization} "axiomatic constant specification"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   381
    (Scan.optional Parse.fixes [] --
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   382
      Scan.optional (Parse.where_ |-- Parse.!!! (Parse.and_list1 Parse_Spec.specs)) []
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   383
      >> (fn (x, y) => Toplevel.theory (#2 o Specification.axiomatization_cmd x y)));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   384
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   385
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   386
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   387
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   388
subsubsection \<open>Notation\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   389
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   390
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   391
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   392
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   393
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   394
  Outer_Syntax.local_theory @{command_keyword type_notation}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   395
    "add concrete syntax for type constructors"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   396
    (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix)
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   397
      >> (fn (mode, args) => Specification.type_notation_cmd true mode args));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   398
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   399
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   400
  Outer_Syntax.local_theory @{command_keyword no_type_notation}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   401
    "delete concrete syntax for type constructors"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   402
    (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix)
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   403
      >> (fn (mode, args) => Specification.type_notation_cmd false mode args));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   404
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   405
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   406
  Outer_Syntax.local_theory @{command_keyword notation}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   407
    "add concrete syntax for constants / fixed variables"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   408
    (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix)
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   409
      >> (fn (mode, args) => Specification.notation_cmd true mode args));
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   410
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   411
val _ =
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   412
  Outer_Syntax.local_theory @{command_keyword no_notation}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   413
    "delete concrete syntax for constants / fixed variables"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   414
    (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix)
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   415
      >> (fn (mode, args) => Specification.notation_cmd false mode args));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   416
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   417
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   418
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   419
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   420
subsubsection \<open>Theorems\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   421
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   422
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   423
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   424
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   425
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   426
  Outer_Syntax.local_theory' @{command_keyword lemmas} "define theorems"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   427
    (Parse_Spec.name_facts -- Parse.for_fixes >>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   428
      (fn (facts, fixes) => #2 oo Specification.theorems_cmd Thm.theoremK facts fixes));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   429
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   430
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   431
  Outer_Syntax.local_theory' @{command_keyword declare} "declare theorems"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   432
    (Parse.and_list1 Parse.xthms1 -- Parse.for_fixes
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   433
      >> (fn (facts, fixes) =>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   434
          #2 oo Specification.theorems_cmd "" [(Attrib.empty_binding, flat facts)] fixes));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   435
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   436
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   437
  Outer_Syntax.local_theory @{command_keyword named_theorems}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   438
    "declare named collection of theorems"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   439
    (Parse.and_list1 (Parse.binding -- Scan.optional Parse.text "") >>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   440
      fold (fn (b, descr) => snd o Named_Theorems.declare b descr));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   441
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   442
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   443
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   444
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   445
subsubsection \<open>Hide names\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   446
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   447
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   448
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   449
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   450
fun hide_names command_keyword what hide parse prep =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   451
  Outer_Syntax.command command_keyword ("hide " ^ what ^ " from name space")
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   452
    ((Parse.opt_keyword "open" >> not) -- Scan.repeat1 parse >> (fn (fully, args) =>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   453
      (Toplevel.theory (fn thy =>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   454
        let val ctxt = Proof_Context.init_global thy
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   455
        in fold (hide fully o prep ctxt) args thy end))));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   456
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   457
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   458
  hide_names @{command_keyword hide_class} "classes" Sign.hide_class Parse.class
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   459
    Proof_Context.read_class;
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   460
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   461
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   462
  hide_names @{command_keyword hide_type} "types" Sign.hide_type Parse.type_const
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   463
    ((#1 o dest_Type) oo Proof_Context.read_type_name {proper = true, strict = false});
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   464
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   465
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   466
  hide_names @{command_keyword hide_const} "consts" Sign.hide_const Parse.const
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   467
    ((#1 o dest_Const) oo Proof_Context.read_const {proper = true, strict = false});
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   468
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   469
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   470
  hide_names @{command_keyword hide_fact} "facts" Global_Theory.hide_fact
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   471
    (Parse.position Parse.xname) (Global_Theory.check_fact o Proof_Context.theory_of);
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   472
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   473
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   474
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   475
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   476
subsection \<open>Bundled declarations\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   477
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   478
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   479
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   480
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   481
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   482
  Outer_Syntax.local_theory @{command_keyword bundle} "define bundle of declarations"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   483
    ((Parse.binding --| @{keyword "="}) -- Parse.xthms1 -- Parse.for_fixes
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   484
      >> (uncurry Bundle.bundle_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   485
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   486
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   487
  Outer_Syntax.command @{command_keyword include}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   488
    "include declarations from bundle in proof body"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   489
    (Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.proof o Bundle.include_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   490
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   491
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   492
  Outer_Syntax.command @{command_keyword including}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   493
    "include declarations from bundle in goal refinement"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   494
    (Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.proof o Bundle.including_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   495
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   496
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   497
  Outer_Syntax.command @{command_keyword print_bundles}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   498
    "print bundles of declarations"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   499
    (Parse.opt_bang >> (fn b => Toplevel.keep (Bundle.print_bundles b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   500
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   501
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   502
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   503
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   504
subsection \<open>Local theory specifications\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   505
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   506
subsubsection \<open>Specification context\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   507
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   508
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   509
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   510
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   511
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   512
  Outer_Syntax.command @{command_keyword context} "begin local theory context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   513
    ((Parse.position Parse.xname >> (fn name =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   514
        Toplevel.begin_local_theory true (Named_Target.begin name)) ||
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   515
      Scan.optional Parse_Spec.includes [] -- Scan.repeat Parse_Spec.context_element
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   516
        >> (fn (incls, elems) => Toplevel.open_target (#2 o Bundle.context_cmd incls elems)))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   517
      --| Parse.begin);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   518
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   519
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   520
  Outer_Syntax.command @{command_keyword end} "end context"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   521
    (Scan.succeed
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   522
      (Toplevel.exit o Toplevel.end_local_theory o Toplevel.close_target o
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   523
        Toplevel.end_proof (K Proof.end_notepad)));
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   524
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   525
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   526
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   527
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   528
subsubsection \<open>Locales and interpretation\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   529
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   530
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   531
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   532
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   533
val locale_val =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   534
  Parse_Spec.locale_expression --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   535
    Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   536
  Scan.repeat1 Parse_Spec.context_element >> pair ([], []);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   537
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   538
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   539
  Outer_Syntax.command @{command_keyword locale} "define named specification context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   540
    (Parse.binding --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   541
      Scan.optional (@{keyword "="} |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   542
      >> (fn ((name, (expr, elems)), begin) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   543
          Toplevel.begin_local_theory begin
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   544
            (Expression.add_locale_cmd name Binding.empty expr elems #> snd)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   545
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   546
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   547
  Outer_Syntax.command @{command_keyword experiment} "open private specification context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   548
    (Scan.repeat Parse_Spec.context_element --| Parse.begin
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   549
      >> (fn elems =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   550
          Toplevel.begin_local_theory true (Experiment.experiment_cmd elems #> snd)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   551
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   552
val interpretation_args =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   553
  Parse.!!! Parse_Spec.locale_expression --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   554
    Scan.optional
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   555
      (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) [];
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   556
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   557
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   558
  Outer_Syntax.command @{command_keyword interpret}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   559
    "prove interpretation of locale expression in proof context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   560
    (interpretation_args >> (fn (expr, equations) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   561
      Toplevel.proof (Interpretation.interpret_cmd expr equations)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   562
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   563
val interpretation_args_with_defs =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   564
  Parse.!!! Parse_Spec.locale_expression --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   565
    (Scan.optional (@{keyword "defines"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   566
      -- ((Parse.binding -- Parse.opt_mixfix') --| @{keyword "="} -- Parse.term))) [] --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   567
    Scan.optional
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   568
      (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   569
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   570
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   571
  Outer_Syntax.local_theory_to_proof @{command_keyword global_interpretation}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   572
    "prove interpretation of locale expression into global theory"
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   573
    (interpretation_args_with_defs >> (fn (expr, (defs, equations)) =>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   574
      Interpretation.global_interpretation_cmd expr defs equations));
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   575
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   576
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   577
  Outer_Syntax.command @{command_keyword sublocale}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   578
    "prove sublocale relation between a locale and a locale expression"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   579
    ((Parse.position Parse.xname --| (@{keyword "\<subseteq>"} || @{keyword "<"}) --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   580
      interpretation_args_with_defs >> (fn (loc, (expr, (defs, equations))) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   581
        Toplevel.theory_to_proof (Interpretation.global_sublocale_cmd loc expr defs equations)))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   582
    || interpretation_args_with_defs >> (fn (expr, (defs, equations)) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   583
        Toplevel.local_theory_to_proof NONE NONE (Interpretation.sublocale_cmd expr defs equations)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   584
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   585
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   586
  Outer_Syntax.command @{command_keyword interpretation}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   587
    "prove interpretation of locale expression in local theory or into global theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   588
    (interpretation_args >> (fn (expr, equations) =>
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   589
      Toplevel.local_theory_to_proof NONE NONE
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   590
        (Interpretation.isar_interpretation_cmd expr equations)));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   591
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   592
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   593
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   594
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   595
subsubsection \<open>Type classes\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   596
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   597
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   598
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   599
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   600
val class_val =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   601
  Parse_Spec.class_expression --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   602
    Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   603
  Scan.repeat1 Parse_Spec.context_element >> pair [];
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   604
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   605
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   606
  Outer_Syntax.command @{command_keyword class} "define type class"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   607
   (Parse.binding -- Scan.optional (@{keyword "="} |-- class_val) ([], []) -- Parse.opt_begin
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   608
    >> (fn ((name, (supclasses, elems)), begin) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   609
        Toplevel.begin_local_theory begin
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   610
          (Class_Declaration.class_cmd name supclasses elems #> snd)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   611
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   612
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   613
  Outer_Syntax.local_theory_to_proof @{command_keyword subclass} "prove a subclass relation"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   614
    (Parse.class >> Class_Declaration.subclass_cmd);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   615
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   616
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   617
  Outer_Syntax.command @{command_keyword instantiation} "instantiate and prove type arity"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   618
   (Parse.multi_arity --| Parse.begin
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   619
     >> (fn arities => Toplevel.begin_local_theory true (Class.instantiation_cmd arities)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   620
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   621
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   622
  Outer_Syntax.command @{command_keyword instance} "prove type arity or subclass relation"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   623
  ((Parse.class --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   624
    ((@{keyword "\<subseteq>"} || @{keyword "<"}) |-- Parse.!!! Parse.class) >> Class.classrel_cmd ||
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   625
    Parse.multi_arity >> Class.instance_arity_cmd) >> Toplevel.theory_to_proof ||
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   626
    Scan.succeed (Toplevel.local_theory_to_proof NONE NONE (Class.instantiation_instance I)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   627
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   628
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   629
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   630
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   631
subsubsection \<open>Arbitrary overloading\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   632
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   633
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   634
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   635
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   636
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   637
  Outer_Syntax.command @{command_keyword overloading} "overloaded definitions"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   638
   (Scan.repeat1 (Parse.name --| (@{keyword "=="} || @{keyword "\<equiv>"}) -- Parse.term --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   639
      Scan.optional (@{keyword "("} |-- (@{keyword "unchecked"} >> K false) --| @{keyword ")"}) true
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   640
      >> Scan.triple1) --| Parse.begin
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   641
   >> (fn operations => Toplevel.begin_local_theory true (Overloading.overloading_cmd operations)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   642
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   643
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   644
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   645
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   646
subsection \<open>Proof commands\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   647
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   648
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   649
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   650
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   651
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   652
  Outer_Syntax.local_theory_to_proof @{command_keyword notepad} "begin proof context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   653
    (Parse.begin >> K Proof.begin_notepad);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   654
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   655
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   656
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   657
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   658
subsubsection \<open>Statements\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   659
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   660
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   661
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   662
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   663
val structured_statement =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   664
  Parse_Spec.statement -- Parse_Spec.cond_statement -- Parse.for_fixes
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   665
    >> (fn ((shows, (strict, assumes)), fixes) => (strict, fixes, assumes, shows));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   666
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   667
fun theorem spec schematic descr =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   668
  Outer_Syntax.local_theory_to_proof' spec
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   669
    ("state " ^ descr)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   670
    (Scan.optional (Parse_Spec.opt_thm_name ":" --|
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   671
      Scan.ahead (Parse_Spec.includes >> K "" ||
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   672
        Parse_Spec.locale_keyword || Parse_Spec.statement_keyword)) Attrib.empty_binding --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   673
      Scan.optional Parse_Spec.includes [] --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   674
      Parse_Spec.general_statement >> (fn ((a, includes), (elems, concl)) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   675
        ((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   676
          Thm.theoremK NONE (K I) a includes elems concl)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   677
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   678
val _ = theorem @{command_keyword theorem} false "theorem";
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   679
val _ = theorem @{command_keyword lemma} false "lemma";
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   680
val _ = theorem @{command_keyword corollary} false "corollary";
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   681
val _ = theorem @{command_keyword proposition} false "proposition";
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   682
val _ = theorem @{command_keyword schematic_goal} true "schematic goal";
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   683
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   684
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   685
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   686
  Outer_Syntax.command @{command_keyword have} "state local goal"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   687
    (structured_statement >> (fn (a, b, c, d) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   688
      Toplevel.proof' (fn int => Proof.have_cmd a NONE (K I) b c d int #> #2)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   689
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   690
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   691
  Outer_Syntax.command @{command_keyword show} "state local goal, to refine pending subgoals"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   692
    (structured_statement >> (fn (a, b, c, d) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   693
      Toplevel.proof' (fn int => Proof.show_cmd a NONE (K I) b c d int #> #2)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   694
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   695
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   696
  Outer_Syntax.command @{command_keyword hence} "old-style alias of \"then have\""
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   697
    (structured_statement >> (fn (a, b, c, d) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   698
      Toplevel.proof' (fn int => Proof.chain #> Proof.have_cmd a NONE (K I) b c d int #> #2)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   699
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   700
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   701
  Outer_Syntax.command @{command_keyword thus} "old-style alias of  \"then show\""
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   702
    (structured_statement >> (fn (a, b, c, d) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   703
      Toplevel.proof' (fn int => Proof.chain #> Proof.show_cmd a NONE (K I) b c d int #> #2)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   704
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   705
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   706
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   707
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   708
subsubsection \<open>Local facts\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   709
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   710
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   711
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   712
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   713
val facts = Parse.and_list1 Parse.xthms1;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   714
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   715
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   716
  Outer_Syntax.command @{command_keyword then} "forward chaining"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   717
    (Scan.succeed (Toplevel.proof Proof.chain));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   718
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   719
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   720
  Outer_Syntax.command @{command_keyword from} "forward chaining from given facts"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   721
    (facts >> (Toplevel.proof o Proof.from_thmss_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   722
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   723
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   724
  Outer_Syntax.command @{command_keyword with} "forward chaining from given and current facts"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   725
    (facts >> (Toplevel.proof o Proof.with_thmss_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   726
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   727
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   728
  Outer_Syntax.command @{command_keyword note} "define facts"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   729
    (Parse_Spec.name_facts >> (Toplevel.proof o Proof.note_thmss_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   730
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   731
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   732
  Outer_Syntax.command @{command_keyword supply} "define facts during goal refinement (unstructured)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   733
    (Parse_Spec.name_facts >> (Toplevel.proof o Proof.supply_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   734
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   735
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   736
  Outer_Syntax.command @{command_keyword using} "augment goal facts"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   737
    (facts >> (Toplevel.proof o Proof.using_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   738
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   739
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   740
  Outer_Syntax.command @{command_keyword unfolding} "unfold definitions in goal and facts"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   741
    (facts >> (Toplevel.proof o Proof.unfolding_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   742
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   743
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   744
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   745
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   746
subsubsection \<open>Proof context\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   747
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   748
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   749
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   750
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   751
val structured_statement =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   752
  Parse_Spec.statement -- Parse_Spec.if_statement' -- Parse.for_fixes
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   753
    >> (fn ((shows, assumes), fixes) => (fixes, assumes, shows));
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   754
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   755
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   756
  Outer_Syntax.command @{command_keyword fix} "fix local variables (Skolem constants)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   757
    (Parse.fixes >> (Toplevel.proof o Proof.fix_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   758
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   759
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   760
  Outer_Syntax.command @{command_keyword assume} "assume propositions"
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   761
    (structured_statement >> (fn (a, b, c) => Toplevel.proof (Proof.assume_cmd a b c)));
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   762
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   763
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   764
  Outer_Syntax.command @{command_keyword presume} "assume propositions, to be established later"
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   765
    (structured_statement >> (fn (a, b, c) => Toplevel.proof (Proof.presume_cmd a b c)));
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   766
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   767
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   768
  Outer_Syntax.command @{command_keyword def} "local definition (non-polymorphic)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   769
    (Parse.and_list1
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   770
      (Parse_Spec.opt_thm_name ":" --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   771
        ((Parse.binding -- Parse.opt_mixfix) --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   772
          ((@{keyword "\<equiv>"} || @{keyword "=="}) |-- Parse.!!! Parse.termp)))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   773
    >> (Toplevel.proof o Proof.def_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   774
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   775
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   776
  Outer_Syntax.command @{command_keyword consider} "state cases rule"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   777
    (Parse_Spec.obtains >> (Toplevel.proof' o Obtain.consider_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   778
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   779
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   780
  Outer_Syntax.command @{command_keyword obtain} "generalized elimination"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   781
    (Parse.parbinding -- Scan.optional (Parse.fixes --| Parse.where_) [] -- Parse_Spec.statement
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   782
      >> (fn ((x, y), z) => Toplevel.proof' (Obtain.obtain_cmd x y z)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   783
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   784
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   785
  Outer_Syntax.command @{command_keyword guess} "wild guessing (unstructured)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   786
    (Scan.optional Parse.fixes [] >> (Toplevel.proof' o Obtain.guess_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   787
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   788
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   789
  Outer_Syntax.command @{command_keyword let} "bind text variables"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   790
    (Parse.and_list1 (Parse.and_list1 Parse.term -- (@{keyword "="} |-- Parse.term))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   791
      >> (Toplevel.proof o Proof.let_bind_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   792
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   793
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   794
  Outer_Syntax.command @{command_keyword write} "add concrete syntax for constants / fixed variables"
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   795
    (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix)
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   796
    >> (fn (mode, args) => Toplevel.proof (Proof.write_cmd mode args)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   797
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   798
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   799
  Outer_Syntax.command @{command_keyword case} "invoke local context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   800
    (Parse_Spec.opt_thm_name ":" --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   801
      (@{keyword "("} |--
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   802
        Parse.!!! (Parse.position Parse.xname -- Scan.repeat (Parse.maybe Parse.binding)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   803
          --| @{keyword ")"}) ||
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   804
        Parse.position Parse.xname >> rpair []) >> (Toplevel.proof o Proof.case_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   805
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   806
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   807
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   808
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   809
subsubsection \<open>Proof structure\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   810
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   811
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   812
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   813
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   814
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   815
  Outer_Syntax.command @{command_keyword "{"} "begin explicit proof block"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   816
    (Scan.succeed (Toplevel.proof Proof.begin_block));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   817
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   818
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   819
  Outer_Syntax.command @{command_keyword "}"} "end explicit proof block"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   820
    (Scan.succeed (Toplevel.proof Proof.end_block));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   821
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   822
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   823
  Outer_Syntax.command @{command_keyword next} "enter next proof block"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   824
    (Scan.succeed (Toplevel.proof Proof.next_block));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   825
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   826
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   827
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   828
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   829
subsubsection \<open>End proof\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   830
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   831
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   832
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   833
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   834
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   835
  Outer_Syntax.command @{command_keyword qed} "conclude proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   836
    (Scan.option Method.parse >> (fn m =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   837
     (Option.map Method.report m;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   838
      Isar_Cmd.qed m)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   839
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   840
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   841
  Outer_Syntax.command @{command_keyword by} "terminal backward proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   842
    (Method.parse -- Scan.option Method.parse >> (fn (m1, m2) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   843
     (Method.report m1;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   844
      Option.map Method.report m2;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   845
      Isar_Cmd.terminal_proof (m1, m2))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   846
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   847
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   848
  Outer_Syntax.command @{command_keyword ".."} "default proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   849
    (Scan.succeed Isar_Cmd.default_proof);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   850
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   851
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   852
  Outer_Syntax.command @{command_keyword "."} "immediate proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   853
    (Scan.succeed Isar_Cmd.immediate_proof);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   854
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   855
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   856
  Outer_Syntax.command @{command_keyword done} "done proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   857
    (Scan.succeed Isar_Cmd.done_proof);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   858
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   859
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   860
  Outer_Syntax.command @{command_keyword sorry} "skip proof (quick-and-dirty mode only!)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   861
    (Scan.succeed Isar_Cmd.skip_proof);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   862
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   863
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   864
  Outer_Syntax.command @{command_keyword "\<proof>"} "dummy proof (quick-and-dirty mode only!)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   865
    (Scan.succeed Isar_Cmd.skip_proof);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   866
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   867
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   868
  Outer_Syntax.command @{command_keyword oops} "forget proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   869
    (Scan.succeed (Toplevel.forget_proof true));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   870
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   871
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   872
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   873
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   874
subsubsection \<open>Proof steps\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   875
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   876
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   877
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   878
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   879
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   880
  Outer_Syntax.command @{command_keyword defer} "shuffle internal proof state"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   881
    (Scan.optional Parse.nat 1 >> (Toplevel.proof o Proof.defer));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   882
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   883
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   884
  Outer_Syntax.command @{command_keyword prefer} "shuffle internal proof state"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   885
    (Parse.nat >> (Toplevel.proof o Proof.prefer));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   886
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   887
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   888
  Outer_Syntax.command @{command_keyword apply} "initial goal refinement step (unstructured)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   889
    (Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply m))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   890
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   891
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   892
  Outer_Syntax.command @{command_keyword apply_end} "terminal goal refinement step (unstructured)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   893
    (Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply_end m))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   894
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   895
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   896
  Outer_Syntax.command @{command_keyword proof} "backward proof step"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   897
    (Scan.option Method.parse >> (fn m =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   898
      (Option.map Method.report m; Toplevel.proofs (Proof.proof m))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   899
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   900
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   901
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   902
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   903
subsubsection \<open>Subgoal focus\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   904
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   905
ML \<open>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   906
local
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   907
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   908
val opt_fact_binding =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   909
  Scan.optional (Parse.binding -- Parse.opt_attribs || Parse.attribs >> pair Binding.empty)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   910
    Attrib.empty_binding;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   911
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   912
val for_params =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   913
  Scan.optional
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   914
    (@{keyword "for"} |--
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   915
      Parse.!!! ((Scan.option Parse.dots >> is_some) --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   916
        (Scan.repeat1 (Parse.position (Parse.maybe Parse.name)))))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   917
    (false, []);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   918
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   919
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   920
  Outer_Syntax.command @{command_keyword subgoal}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   921
    "focus on first subgoal within backward refinement"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   922
    (opt_fact_binding -- (Scan.option (@{keyword "premises"} |-- Parse.!!! opt_fact_binding)) --
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   923
      for_params >> (fn ((a, b), c) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   924
        Toplevel.proofs (Seq.make_results o Seq.single o #2 o Subgoal.subgoal_cmd a b c)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   925
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   926
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   927
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   928
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   929
subsubsection \<open>Calculation\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   930
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   931
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   932
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   933
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   934
val calculation_args =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   935
  Scan.option (@{keyword "("} |-- Parse.!!! ((Parse.xthms1 --| @{keyword ")"})));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   936
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   937
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   938
  Outer_Syntax.command @{command_keyword also} "combine calculation and current facts"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   939
    (calculation_args >> (Toplevel.proofs' o Calculation.also_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   940
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   941
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   942
  Outer_Syntax.command @{command_keyword finally}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   943
    "combine calculation and current facts, exhibit result"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   944
    (calculation_args >> (Toplevel.proofs' o Calculation.finally_cmd));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   945
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   946
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   947
  Outer_Syntax.command @{command_keyword moreover} "augment calculation by current facts"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   948
    (Scan.succeed (Toplevel.proof' Calculation.moreover));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   949
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   950
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   951
  Outer_Syntax.command @{command_keyword ultimately}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   952
    "augment calculation by current facts, exhibit result"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   953
    (Scan.succeed (Toplevel.proof' Calculation.ultimately));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   954
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   955
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   956
  Outer_Syntax.command @{command_keyword print_trans_rules} "print transitivity rules"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   957
    (Scan.succeed (Toplevel.keep (Calculation.print_rules o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   958
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   959
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   960
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   961
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   962
subsubsection \<open>Proof navigation\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   963
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   964
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   965
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   966
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   967
fun report_back () =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   968
  Output.report [Markup.markup Markup.bad "Explicit backtracking"];
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   969
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   970
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   971
  Outer_Syntax.command @{command_keyword back} "explicit backtracking of proof command"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   972
    (Scan.succeed
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   973
     (Toplevel.actual_proof (fn prf => (report_back (); Proof_Node.back prf)) o
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   974
      Toplevel.skip_proof report_back));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   975
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   976
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   977
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   978
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   979
subsection \<open>Diagnostic commands (for interactive mode only)\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   980
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   981
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
   982
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   983
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   984
val opt_modes =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   985
  Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) [];
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   986
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   987
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   988
  Outer_Syntax.command @{command_keyword help}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   989
    "retrieve outer syntax commands according to name patterns"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   990
    (Scan.repeat Parse.name >>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   991
      (fn pats => Toplevel.keep (fn st => Outer_Syntax.help (Toplevel.theory_of st) pats)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   992
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   993
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   994
  Outer_Syntax.command @{command_keyword print_commands} "print outer syntax commands"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   995
    (Scan.succeed (Toplevel.keep (Outer_Syntax.print_commands o Toplevel.theory_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   996
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   997
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   998
  Outer_Syntax.command @{command_keyword print_options} "print configuration options"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
   999
    (Parse.opt_bang >> (fn b => Toplevel.keep (Attrib.print_options b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1000
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1001
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1002
  Outer_Syntax.command @{command_keyword print_context}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1003
    "print context of local theory target"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1004
    (Scan.succeed (Toplevel.keep (Pretty.writeln_chunks o Toplevel.pretty_context)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1005
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1006
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1007
  Outer_Syntax.command @{command_keyword print_theory}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1008
    "print logical theory contents"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1009
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1010
      Toplevel.keep (Pretty.writeln o Proof_Display.pretty_theory b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1011
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1012
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1013
  Outer_Syntax.command @{command_keyword print_definitions}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1014
    "print dependencies of definitional theory content"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1015
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1016
      Toplevel.keep (Pretty.writeln o Proof_Display.pretty_definitions b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1017
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1018
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1019
  Outer_Syntax.command @{command_keyword print_syntax}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1020
    "print inner syntax of context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1021
    (Scan.succeed (Toplevel.keep (Proof_Context.print_syntax o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1022
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1023
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1024
  Outer_Syntax.command @{command_keyword print_defn_rules}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1025
    "print definitional rewrite rules of context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1026
    (Scan.succeed (Toplevel.keep (Local_Defs.print_rules o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1027
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1028
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1029
  Outer_Syntax.command @{command_keyword print_abbrevs}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1030
    "print constant abbreviations of context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1031
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1032
      Toplevel.keep (Proof_Context.print_abbrevs b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1033
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1034
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1035
  Outer_Syntax.command @{command_keyword print_theorems}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1036
    "print theorems of local theory or proof context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1037
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1038
      Toplevel.keep (Pretty.writeln o Pretty.chunks o Isar_Cmd.pretty_theorems b)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1039
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1040
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1041
  Outer_Syntax.command @{command_keyword print_locales}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1042
    "print locales of this theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1043
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1044
      Toplevel.keep (Locale.print_locales b o Toplevel.theory_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1045
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1046
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1047
  Outer_Syntax.command @{command_keyword print_classes}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1048
    "print classes of this theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1049
    (Scan.succeed (Toplevel.keep (Class.print_classes o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1050
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1051
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1052
  Outer_Syntax.command @{command_keyword print_locale}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1053
    "print locale of this theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1054
    (Parse.opt_bang -- Parse.position Parse.xname >> (fn (b, name) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1055
      Toplevel.keep (fn state => Locale.print_locale (Toplevel.theory_of state) b name)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1056
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1057
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1058
  Outer_Syntax.command @{command_keyword print_interps}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1059
    "print interpretations of locale for this theory or proof context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1060
    (Parse.position Parse.xname >> (fn name =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1061
      Toplevel.keep (fn state => Locale.print_registrations (Toplevel.context_of state) name)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1062
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1063
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1064
  Outer_Syntax.command @{command_keyword print_dependencies}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1065
    "print dependencies of locale expression"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1066
    (Parse.opt_bang -- Parse_Spec.locale_expression >> (fn (b, expr) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1067
      Toplevel.keep (fn state => Expression.print_dependencies (Toplevel.context_of state) b expr)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1068
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1069
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1070
  Outer_Syntax.command @{command_keyword print_attributes}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1071
    "print attributes of this theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1072
    (Parse.opt_bang >> (fn b => Toplevel.keep (Attrib.print_attributes b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1073
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1074
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1075
  Outer_Syntax.command @{command_keyword print_simpset}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1076
    "print context of Simplifier"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1077
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1078
      Toplevel.keep (Pretty.writeln o Simplifier.pretty_simpset b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1079
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1080
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1081
  Outer_Syntax.command @{command_keyword print_rules} "print intro/elim rules"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1082
    (Scan.succeed (Toplevel.keep (Context_Rules.print_rules o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1083
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1084
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1085
  Outer_Syntax.command @{command_keyword print_methods} "print methods of this theory"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1086
    (Parse.opt_bang >> (fn b => Toplevel.keep (Method.print_methods b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1087
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1088
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1089
  Outer_Syntax.command @{command_keyword print_antiquotations}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1090
    "print document antiquotations"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1091
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1092
      Toplevel.keep (Thy_Output.print_antiquotations b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1093
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1094
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1095
  Outer_Syntax.command @{command_keyword print_ML_antiquotations}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1096
    "print ML antiquotations"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1097
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1098
      Toplevel.keep (ML_Context.print_antiquotations b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1099
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1100
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1101
  Outer_Syntax.command @{command_keyword locale_deps} "visualize locale dependencies"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1102
    (Scan.succeed
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1103
      (Toplevel.keep (Toplevel.theory_of #> (fn thy =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1104
        Locale.pretty_locale_deps thy
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1105
        |> map (fn {name, parents, body} =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1106
          ((name, Graph_Display.content_node (Locale.extern thy name) [body]), parents))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1107
        |> Graph_Display.display_graph_old))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1108
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1109
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1110
  Outer_Syntax.command @{command_keyword print_term_bindings}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1111
    "print term bindings of proof context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1112
    (Scan.succeed
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1113
      (Toplevel.keep
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1114
        (Pretty.writeln_chunks o Proof_Context.pretty_term_bindings o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1115
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1116
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1117
  Outer_Syntax.command @{command_keyword print_facts} "print facts of proof context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1118
    (Parse.opt_bang >> (fn b =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1119
      Toplevel.keep (Proof_Context.print_local_facts b o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1120
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1121
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1122
  Outer_Syntax.command @{command_keyword print_cases} "print cases of proof context"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1123
    (Scan.succeed
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1124
      (Toplevel.keep (Pretty.writeln_chunks o Proof_Context.pretty_cases o Toplevel.context_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1125
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1126
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1127
  Outer_Syntax.command @{command_keyword print_statement}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1128
    "print theorems as long statements"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1129
    (opt_modes -- Parse.xthms1 >> Isar_Cmd.print_stmts);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1130
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1131
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1132
  Outer_Syntax.command @{command_keyword thm} "print theorems"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1133
    (opt_modes -- Parse.xthms1 >> Isar_Cmd.print_thms);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1134
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1135
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1136
  Outer_Syntax.command @{command_keyword prf} "print proof terms of theorems"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1137
    (opt_modes -- Scan.option Parse.xthms1 >> Isar_Cmd.print_prfs false);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1138
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1139
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1140
  Outer_Syntax.command @{command_keyword full_prf} "print full proof terms of theorems"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1141
    (opt_modes -- Scan.option Parse.xthms1 >> Isar_Cmd.print_prfs true);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1142
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1143
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1144
  Outer_Syntax.command @{command_keyword prop} "read and print proposition"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1145
    (opt_modes -- Parse.term >> Isar_Cmd.print_prop);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1146
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1147
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1148
  Outer_Syntax.command @{command_keyword term} "read and print term"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1149
    (opt_modes -- Parse.term >> Isar_Cmd.print_term);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1150
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1151
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1152
  Outer_Syntax.command @{command_keyword typ} "read and print type"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1153
    (opt_modes -- (Parse.typ -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1154
      >> Isar_Cmd.print_type);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1155
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1156
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1157
  Outer_Syntax.command @{command_keyword print_codesetup} "print code generator setup"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1158
    (Scan.succeed (Toplevel.keep (Code.print_codesetup o Toplevel.theory_of)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1159
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1160
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1161
  Outer_Syntax.command @{command_keyword print_state}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1162
    "print current proof state (if present)"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1163
    (opt_modes >> (fn modes =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1164
      Toplevel.keep (Print_Mode.with_modes modes (Output.state o Toplevel.string_of_state))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1165
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1166
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1167
  Outer_Syntax.command @{command_keyword welcome} "print welcome message"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1168
    (Scan.succeed (Toplevel.keep (fn _ => writeln (Session.welcome ()))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1169
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1170
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1171
  Outer_Syntax.command @{command_keyword display_drafts}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1172
    "display raw source files with symbols"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1173
    (Scan.repeat1 Parse.path >> (fn names =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1174
      Toplevel.keep (fn _ => ignore (Present.display_drafts (map Path.explode names)))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1175
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1176
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1177
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1178
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1179
subsection \<open>Dependencies\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1180
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1181
ML \<open>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1182
local
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1183
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1184
val theory_bounds =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1185
  Parse.position Parse.theory_xname >> single ||
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1186
  (@{keyword "("} |-- Parse.enum "|" (Parse.position Parse.theory_xname) --| @{keyword ")"});
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1187
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1188
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1189
  Outer_Syntax.command @{command_keyword thy_deps} "visualize theory dependencies"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1190
    (Scan.option theory_bounds -- Scan.option theory_bounds >>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1191
      (fn args => Toplevel.keep (fn st => Thy_Deps.thy_deps_cmd (Toplevel.context_of st) args)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1192
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1193
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1194
val class_bounds =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1195
  Parse.sort >> single ||
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1196
  (@{keyword "("} |-- Parse.enum "|" Parse.sort --| @{keyword ")"});
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1197
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1198
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1199
  Outer_Syntax.command @{command_keyword class_deps} "visualize class dependencies"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1200
    (Scan.option class_bounds -- Scan.option class_bounds >> (fn args =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1201
      Toplevel.keep (fn st => Class_Deps.class_deps_cmd (Toplevel.context_of st) args)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1202
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1203
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1204
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1205
  Outer_Syntax.command @{command_keyword thm_deps} "visualize theorem dependencies"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1206
    (Parse.xthms1 >> (fn args =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1207
      Toplevel.keep (fn st =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1208
        Thm_Deps.thm_deps (Toplevel.theory_of st)
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1209
          (Attrib.eval_thms (Toplevel.context_of st) args))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1210
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1211
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1212
val thy_names =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1213
  Scan.repeat1 (Scan.unless Parse.minus (Parse.position Parse.theory_xname));
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1214
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1215
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1216
  Outer_Syntax.command @{command_keyword unused_thms} "find unused theorems"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1217
    (Scan.option ((thy_names --| Parse.minus) -- Scan.option thy_names) >> (fn opt_range =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1218
        Toplevel.keep (fn st =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1219
          let
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1220
            val thy = Toplevel.theory_of st;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1221
            val ctxt = Toplevel.context_of st;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1222
            fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]);
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1223
            val check = Theory.check ctxt;
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1224
          in
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1225
            Thm_Deps.unused_thms
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1226
              (case opt_range of
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1227
                NONE => (Theory.parents_of thy, [thy])
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1228
              | SOME (xs, NONE) => (map check xs, [thy])
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1229
              | SOME (xs, SOME ys) => (map check xs, map check ys))
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1230
            |> map pretty_thm |> Pretty.writeln_chunks
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1231
          end)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1232
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1233
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1234
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1235
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1236
subsubsection \<open>Find consts and theorems\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1237
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1238
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1239
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1240
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1241
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1242
  Outer_Syntax.command @{command_keyword find_consts}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1243
    "find constants by name / type patterns"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1244
    (Find_Consts.query_parser >> (fn spec =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1245
      Toplevel.keep (fn st =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1246
        Pretty.writeln (Find_Consts.pretty_consts (Toplevel.context_of st) spec))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1247
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1248
val options =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1249
  Scan.optional
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1250
    (Parse.$$$ "(" |--
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1251
      Parse.!!! (Scan.option Parse.nat --
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1252
        Scan.optional (Parse.reserved "with_dups" >> K false) true --| Parse.$$$ ")"))
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1253
    (NONE, true);
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1254
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1255
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1256
  Outer_Syntax.command @{command_keyword find_theorems}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1257
    "find theorems meeting specified criteria"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1258
    (options -- Find_Theorems.query_parser >> (fn ((opt_lim, rem_dups), spec) =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1259
      Toplevel.keep (fn st =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1260
        Pretty.writeln
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1261
          (Find_Theorems.pretty_theorems (Find_Theorems.proof_state st) opt_lim rem_dups spec))));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1262
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1263
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1264
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1265
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1266
subsection \<open>Code generation\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1267
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1268
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1269
local
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1270
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1271
val _ =
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1272
  Outer_Syntax.command @{command_keyword code_datatype}
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1273
    "define set of code datatype constructors"
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1274
    (Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd));
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1275
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1276
in end\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1277
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1278
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1279
subsection \<open>Extraction of programs from proofs\<close>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1280
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1281
ML \<open>
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1282
local
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1283
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1284
val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") [];
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1285
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1286
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1287
  Outer_Syntax.command @{command_keyword realizers}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1288
    "specify realizers for primitive axioms / theorems, together with correctness proof"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1289
    (Scan.repeat1 (Parse.xname -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1290
     (fn xs => Toplevel.theory (fn thy => Extraction.add_realizers
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1291
       (map (fn (((a, vs), s1), s2) => (Global_Theory.get_thm thy a, (vs, s1, s2))) xs) thy)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1292
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1293
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1294
  Outer_Syntax.command @{command_keyword realizability}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1295
    "add equations characterizing realizability"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1296
    (Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_realizes_eqns));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1297
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1298
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1299
  Outer_Syntax.command @{command_keyword extract_type}
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1300
    "add equations characterizing type of extracted program"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1301
    (Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_typeof_eqns));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1302
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1303
val _ =
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1304
  Outer_Syntax.command @{command_keyword extract} "extract terms from proofs"
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1305
    (Scan.repeat1 (Parse.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy =>
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1306
      Extraction.extract (map (apfst (Global_Theory.get_thm thy)) xs) thy)));
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1307
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1308
in end\<close>
62849
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1309
caaa2fc4040d clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
wenzelm
parents: 62848
diff changeset
  1310
62856
3f97aa4580c6 tuned -- more explicit sections;
wenzelm
parents: 62849
diff changeset
  1311
section \<open>Isar attributes\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1312
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1313
attribute_setup tagged =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1314
  \<open>Scan.lift (Args.name -- Args.name) >> Thm.tag\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1315
  "tagged theorem"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1316
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1317
attribute_setup untagged =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1318
  \<open>Scan.lift Args.name >> Thm.untag\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1319
  "untagged theorem"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1320
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1321
attribute_setup kind =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1322
  \<open>Scan.lift Args.name >> Thm.kind\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1323
  "theorem kind"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1324
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1325
attribute_setup THEN =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1326
  \<open>Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- Attrib.thm
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1327
    >> (fn (i, B) => Thm.rule_attribute [B] (fn _ => fn A => A RSN (i, B)))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1328
  "resolution with rule"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1329
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1330
attribute_setup OF =
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1331
  \<open>Attrib.thms >> (fn Bs => Thm.rule_attribute Bs (fn _ => fn A => A OF Bs))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1332
  "rule resolved with facts"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1333
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1334
attribute_setup rename_abs =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1335
  \<open>Scan.lift (Scan.repeat (Args.maybe Args.name)) >> (fn vs =>
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1336
    Thm.rule_attribute [] (K (Drule.rename_bvars' vs)))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1337
  "rename bound variables in abstractions"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1338
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1339
attribute_setup unfolded =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1340
  \<open>Attrib.thms >> (fn ths =>
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1341
    Thm.rule_attribute ths (fn context => Local_Defs.unfold (Context.proof_of context) ths))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1342
  "unfolded definitions"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1343
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1344
attribute_setup folded =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1345
  \<open>Attrib.thms >> (fn ths =>
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1346
    Thm.rule_attribute ths (fn context => Local_Defs.fold (Context.proof_of context) ths))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1347
  "folded definitions"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1348
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1349
attribute_setup consumes =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1350
  \<open>Scan.lift (Scan.optional Parse.int 1) >> Rule_Cases.consumes\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1351
  "number of consumed facts"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1352
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1353
attribute_setup constraints =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1354
  \<open>Scan.lift Parse.nat >> Rule_Cases.constraints\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1355
  "number of equality constraints"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1356
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1357
attribute_setup case_names =
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1358
  \<open>Scan.lift (Scan.repeat1 (Args.name --
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1359
    Scan.optional (@{keyword "["} |-- Scan.repeat1 (Args.maybe Args.name) --| @{keyword "]"}) []))
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1360
    >> (fn cs =>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1361
      Rule_Cases.cases_hyp_names
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1362
        (map #1 cs)
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1363
        (map (map (the_default Rule_Cases.case_hypsN) o #2) cs))\<close>
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1364
  "named rule cases"
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1365
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1366
attribute_setup case_conclusion =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1367
  \<open>Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1368
  "named conclusion of rule cases"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1369
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1370
attribute_setup params =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1371
  \<open>Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1372
  "named rule parameters"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1373
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1374
attribute_setup rule_format =
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1375
  \<open>Scan.lift (Args.mode "no_asm")
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1376
    >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format)\<close>
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1377
  "result put into canonical rule format"
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1378
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1379
attribute_setup elim_format =
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1380
  \<open>Scan.succeed (Thm.rule_attribute [] (K Tactic.make_elim))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1381
  "destruct rule turned into elimination rule format"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1382
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1383
attribute_setup no_vars =
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1384
  \<open>Scan.succeed (Thm.rule_attribute [] (fn context => fn th =>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1385
    let
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1386
      val ctxt = Variable.set_body false (Context.proof_of context);
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1387
      val ((_, [th']), _) = Variable.import true [th] ctxt;
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1388
    in th' end))\<close>
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1389
  "imported schematic variables"
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1390
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1391
attribute_setup atomize =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1392
  \<open>Scan.succeed Object_Logic.declare_atomize\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1393
  "declaration of atomize rule"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1394
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1395
attribute_setup rulify =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1396
  \<open>Scan.succeed Object_Logic.declare_rulify\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1397
  "declaration of rulify rule"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1398
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1399
attribute_setup rotated =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1400
  \<open>Scan.lift (Scan.optional Parse.int 1
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1401
    >> (fn n => Thm.rule_attribute [] (fn _ => rotate_prems n)))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1402
  "rotated theorem premises"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1403
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1404
attribute_setup defn =
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1405
  \<open>Attrib.add_del Local_Defs.defn_add Local_Defs.defn_del\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1406
  "declaration of definitional transformations"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1407
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1408
attribute_setup abs_def =
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61671
diff changeset
  1409
  \<open>Scan.succeed (Thm.rule_attribute [] (fn context =>
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1410
    Local_Defs.meta_rewrite_rule (Context.proof_of context) #> Drule.abs_def))\<close>
55140
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1411
  "abstract over free variables of definitional theorem"
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1412
7eb0c04e4c40 define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
wenzelm
parents: 55030
diff changeset
  1413
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1414
section \<open>Further content for the Pure theory\<close>
20627
30da2841553e revert to previous version;
wenzelm
parents: 20596
diff changeset
  1415
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1416
subsection \<open>Meta-level connectives in assumptions\<close>
15803
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1417
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1418
lemma meta_mp:
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1419
  assumes "PROP P \<Longrightarrow> PROP Q" and "PROP P"
15803
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1420
  shows "PROP Q"
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1421
    by (rule \<open>PROP P \<Longrightarrow> PROP Q\<close> [OF \<open>PROP P\<close>])
15803
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1422
23432
cec811764a38 added meta_impE
nipkow
parents: 22933
diff changeset
  1423
lemmas meta_impE = meta_mp [elim_format]
cec811764a38 added meta_impE
nipkow
parents: 22933
diff changeset
  1424
15803
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1425
lemma meta_spec:
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1426
  assumes "\<And>x. PROP P x"
26958
ed3a58a9eae1 converted to regular application syntax;
wenzelm
parents: 26572
diff changeset
  1427
  shows "PROP P x"
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1428
    by (rule \<open>\<And>x. PROP P x\<close>)
15803
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1429
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1430
lemmas meta_allE = meta_spec [elim_format]
42c75e0c9140 The Pure theory.
wenzelm
parents:
diff changeset
  1431
26570
dbc458262f4c added swap_params;
wenzelm
parents: 26435
diff changeset
  1432
lemma swap_params:
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1433
  "(\<And>x y. PROP P x y) \<equiv> (\<And>y x. PROP P x y)" ..
26570
dbc458262f4c added swap_params;
wenzelm
parents: 26435
diff changeset
  1434
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1435
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1436
subsection \<open>Meta-level conjunction\<close>
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1437
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1438
lemma all_conjunction:
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1439
  "(\<And>x. PROP A x &&& PROP B x) \<equiv> ((\<And>x. PROP A x) &&& (\<And>x. PROP B x))"
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1440
proof
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1441
  assume conj: "\<And>x. PROP A x &&& PROP B x"
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1442
  show "(\<And>x. PROP A x) &&& (\<And>x. PROP B x)"
19121
d7fd5415a781 simplified Pure conjunction;
wenzelm
parents: 19048
diff changeset
  1443
  proof -
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1444
    fix x
26958
ed3a58a9eae1 converted to regular application syntax;
wenzelm
parents: 26572
diff changeset
  1445
    from conj show "PROP A x" by (rule conjunctionD1)
ed3a58a9eae1 converted to regular application syntax;
wenzelm
parents: 26572
diff changeset
  1446
    from conj show "PROP B x" by (rule conjunctionD2)
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1447
  qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1448
next
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1449
  assume conj: "(\<And>x. PROP A x) &&& (\<And>x. PROP B x)"
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1450
  fix x
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28699
diff changeset
  1451
  show "PROP A x &&& PROP B x"
19121
d7fd5415a781 simplified Pure conjunction;
wenzelm
parents: 19048
diff changeset
  1452
  proof -
26958
ed3a58a9eae1 converted to regular application syntax;
wenzelm
parents: 26572
diff changeset
  1453
    show "PROP A x" by (rule conj [THEN conjunctionD1, rule_format])
ed3a58a9eae1 converted to regular application syntax;
wenzelm
parents: 26572
diff changeset
  1454
    show "PROP B x" by (rule conj [THEN conjunctionD2, rule_format])
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1455
  qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1456
qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1457
19121
d7fd5415a781 simplified Pure conjunction;
wenzelm
parents: 19048
diff changeset
  1458
lemma imp_conjunction:
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1459
  "(PROP A \<Longrightarrow> PROP B &&& PROP C) \<equiv> ((PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C))"
18836
3a1e4ee72075 tuned proofs;
wenzelm
parents: 18710
diff changeset
  1460
proof
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1461
  assume conj: "PROP A \<Longrightarrow> PROP B &&& PROP C"
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1462
  show "(PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C)"
19121
d7fd5415a781 simplified Pure conjunction;
wenzelm
parents: 19048
diff changeset
  1463
  proof -
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1464
    assume "PROP A"
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1465
    from conj [OF \<open>PROP A\<close>] show "PROP B" by (rule conjunctionD1)
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1466
    from conj [OF \<open>PROP A\<close>] show "PROP C" by (rule conjunctionD2)
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1467
  qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1468
next
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1469
  assume conj: "(PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C)"
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1470
  assume "PROP A"
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28699
diff changeset
  1471
  show "PROP B &&& PROP C"
19121
d7fd5415a781 simplified Pure conjunction;
wenzelm
parents: 19048
diff changeset
  1472
  proof -
58611
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1473
    from \<open>PROP A\<close> show "PROP B" by (rule conj [THEN conjunctionD1])
d49f3181030e more cartouches;
wenzelm
parents: 58544
diff changeset
  1474
    from \<open>PROP A\<close> show "PROP C" by (rule conj [THEN conjunctionD2])
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1475
  qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1476
qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1477
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1478
lemma conjunction_imp:
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1479
  "(PROP A &&& PROP B \<Longrightarrow> PROP C) \<equiv> (PROP A \<Longrightarrow> PROP B \<Longrightarrow> PROP C)"
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1480
proof
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1481
  assume r: "PROP A &&& PROP B \<Longrightarrow> PROP C"
22933
338c7890c96f tuned proofs;
wenzelm
parents: 21627
diff changeset
  1482
  assume ab: "PROP A" "PROP B"
338c7890c96f tuned proofs;
wenzelm
parents: 21627
diff changeset
  1483
  show "PROP C"
338c7890c96f tuned proofs;
wenzelm
parents: 21627
diff changeset
  1484
  proof (rule r)
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28699
diff changeset
  1485
    from ab show "PROP A &&& PROP B" .
22933
338c7890c96f tuned proofs;
wenzelm
parents: 21627
diff changeset
  1486
  qed
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1487
next
58612
dbe216a75a4b more symbols;
wenzelm
parents: 58611
diff changeset
  1488
  assume r: "PROP A \<Longrightarrow> PROP B \<Longrightarrow> PROP C"
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28699
diff changeset
  1489
  assume conj: "PROP A &&& PROP B"
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1490
  show "PROP C"
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1491
  proof (rule r)
19121
d7fd5415a781 simplified Pure conjunction;
wenzelm
parents: 19048
diff changeset
  1492
    from conj show "PROP A" by (rule conjunctionD1)
d7fd5415a781 simplified Pure conjunction;
wenzelm
parents: 19048
diff changeset
  1493
    from conj show "PROP B" by (rule conjunctionD2)
18466
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1494
  qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1495
qed
389a6f9c31f4 added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents: 18019
diff changeset
  1496
48638
22d65e375c01 more standard bootstrapping of Pure.thy;
wenzelm
parents: 29606
diff changeset
  1497
end