src/Pure/Isar/target_context.ML
author paulson <lp15@cam.ac.uk>
Sun, 15 Nov 2020 13:06:24 +0000
changeset 72608 ad45ae49be85
parent 72536 589645894305
child 72952 09479be1fe2a
permissions -rw-r--r--
more de-applying
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Isar/target_context.ML
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     3
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     4
Handling of named and nested targets at the Isar toplevel via context begin / end blocks.
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     5
*)
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     6
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     7
signature TARGET_CONTEXT =
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
     8
sig
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
     9
  val context_begin_named_cmd: (xstring * Position.T) list -> xstring * Position.T ->
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    10
    theory -> local_theory
72505
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    11
  val end_named_cmd: local_theory -> Context.generic
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    12
  val switch_named_cmd: (xstring * Position.T) option -> Context.generic ->
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    13
    (local_theory -> Context.generic) * local_theory
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    14
  val context_begin_nested_cmd: (xstring * Position.T) list -> Element.context list ->
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    15
    Context.generic -> local_theory
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    16
  val end_nested_cmd: local_theory -> Context.generic
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    17
end;
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    18
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    19
structure Target_Context: TARGET_CONTEXT =
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    20
struct
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    21
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    22
fun prep_includes thy =
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    23
  map (Bundle.check (Proof_Context.init_global thy));
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    24
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    25
fun context_begin_named_cmd raw_includes ("-", _) thy =
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    26
      Named_Target.init (prep_includes thy raw_includes) "" thy
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    27
  | context_begin_named_cmd raw_includes name thy =
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    28
      Named_Target.init (prep_includes thy raw_includes) (Locale.check thy name) thy;
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    29
72505
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    30
val end_named_cmd = Context.Theory o Local_Theory.exit_global;
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    31
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    32
fun switch_named_cmd NONE (Context.Theory thy) =
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    33
      (end_named_cmd, Named_Target.theory_init thy)
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    34
  | switch_named_cmd (SOME name) (Context.Theory thy) =
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    35
      (end_named_cmd, context_begin_named_cmd [] name thy)
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    36
  | switch_named_cmd NONE (Context.Proof lthy) =
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    37
      (Context.Proof o Local_Theory.reset, lthy)
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    38
  | switch_named_cmd (SOME name) (Context.Proof lthy) =
72505
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    39
      (Context.Proof o Named_Target.reinit lthy o Local_Theory.exit_global,
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    40
        (context_begin_named_cmd [] name o Local_Theory.exit_global_nonbrittle) lthy);
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    41
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    42
fun includes raw_includes lthy =
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    43
  Bundle.includes (map (Bundle.check lthy) raw_includes) lthy;
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    44
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    45
fun context_begin_nested_cmd raw_includes raw_elems gthy =
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    46
  gthy
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    47
  |> Context.cases Named_Target.theory_init Local_Theory.assert
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
    48
  |> includes raw_includes
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    49
  |> Expression.read_declaration ([], []) I raw_elems
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    50
  |> (#4 o fst)
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    51
  |> Local_Theory.begin_nested
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    52
  |> snd;
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    53
72505
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    54
fun end_nested_cmd lthy =
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    55
  let
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    56
    val lthy' = Local_Theory.end_nested lthy
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    57
  in
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    58
    if Named_Target.is_theory lthy'
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    59
    then Context.Theory (Local_Theory.exit_global lthy')
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    60
    else Context.Proof lthy'
974071d873ba tuned interfaces
haftmann
parents: 72502
diff changeset
    61
  end;
72453
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    62
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    63
end;
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    64
e4dde7beab39 dedicated module for toplevel target handling
haftmann
parents:
diff changeset
    65
structure Local_Theory : LOCAL_THEORY = struct open Local_Theory; end;