src/Pure/Isar/expression.ML
author ballarin
Mon, 08 Dec 2008 14:18:29 +0100
changeset 29019 8e7d6f959bd7
parent 29018 17538bdef546
child 29020 3e95d28114a1
permissions -rw-r--r--
Explicitly close up defines.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     1
(*  Title:      Pure/Isar/expression.ML
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     2
    Author:     Clemens Ballarin, TU Muenchen
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     3
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
     4
New locale development --- experimental.
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     5
*)
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     6
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     7
signature EXPRESSION =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     8
sig
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
     9
  datatype 'term map = Positional of 'term option list | Named of (string * 'term) list;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    10
  type 'term expr = (string * (string * 'term map)) list;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    11
  type expression = string expr * (Binding.T * string option * mixfix) list;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    12
  type expression_i = term expr * (Binding.T * typ option * mixfix) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    13
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
    14
  (* Processing of context statements *)
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    15
  val read_statement: Element.context list -> (string * string list) list list ->
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    16
    Proof.context ->  (term * term list) list list * Proof.context;
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    17
  val cert_statement: Element.context_i list -> (term * term list) list list ->
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    18
    Proof.context -> (term * term list) list list * Proof.context;
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
    19
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    20
  (* Declaring locales *)
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
    21
  val add_locale_cmd: string -> bstring -> expression -> Element.context list -> theory ->
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    22
    string * Proof.context
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
    23
  val add_locale: string -> bstring -> expression_i -> Element.context_i list -> theory ->
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    24
    string * Proof.context
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    25
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    26
  (* Interpretation *)
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
    27
  val sublocale_cmd: string -> expression -> theory -> Proof.state;
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
    28
  val sublocale: string -> expression_i -> theory -> Proof.state;
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
    29
  val interpretation_cmd: expression -> theory -> Proof.state;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
    30
  val interpretation: expression_i -> theory -> Proof.state;
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
    31
  val interpret_cmd: expression -> bool -> Proof.state -> Proof.state;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
    32
  val interpret: expression_i -> bool -> Proof.state -> Proof.state;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
    33
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    34
  (* Debugging and development *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    35
  val parse_expression: OuterParse.token list -> expression * OuterParse.token list
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
    36
    (* FIXME to spec_parse.ML *)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    37
end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    38
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    39
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    40
structure Expression : EXPRESSION =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    41
struct
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    42
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    43
datatype ctxt = datatype Element.ctxt;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    44
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    45
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    46
(*** Expressions ***)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    47
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    48
datatype 'term map =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    49
  Positional of 'term option list |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    50
  Named of (string * 'term) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    51
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    52
type 'term expr = (string * (string * 'term map)) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    53
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    54
type expression = string expr * (Binding.T * string option * mixfix) list;
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
    55
type expression_i = term expr * (Binding.T * typ option * mixfix) list;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    56
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    57
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    58
(** Parsing and printing **)
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    59
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    60
local
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    61
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    62
structure P = OuterParse;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    63
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    64
val loc_keyword = P.$$$ "fixes" || P.$$$ "constrains" || P.$$$ "assumes" ||
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    65
   P.$$$ "defines" || P.$$$ "notes";
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    66
fun plus1_unless test scan =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    67
  scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan));
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    68
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    69
val prefix = P.name --| P.$$$ ":";
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    70
val named = P.name -- (P.$$$ "=" |-- P.term);
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    71
val position = P.maybe P.term;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    72
val instance = P.$$$ "where" |-- P.and_list1 named >> Named ||
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    73
  Scan.repeat1 position >> Positional;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    74
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    75
in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    76
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    77
val parse_expression =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    78
  let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    79
    fun expr2 x = P.xname x;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    80
    fun expr1 x = (Scan.optional prefix "" -- expr2 --
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    81
      Scan.optional instance (Named []) >> (fn ((p, l), i) => (l, (p, i)))) x;
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    82
    fun expr0 x = (plus1_unless loc_keyword expr1) x;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    83
  in expr0 -- P.for_fixes end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    84
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    85
end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    86
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
    87
fun pretty_expr thy expr =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    88
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    89
    fun pretty_pos NONE = Pretty.str "_"
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    90
      | pretty_pos (SOME x) = Pretty.str x;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    91
    fun pretty_named (x, y) = [Pretty.str x, Pretty.brk 1, Pretty.str "=",
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    92
          Pretty.brk 1, Pretty.str y] |> Pretty.block;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    93
    fun pretty_ren (Positional ps) = take_suffix is_none ps |> snd |>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    94
          map pretty_pos |> Pretty.breaks
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    95
      | pretty_ren (Named []) = []
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    96
      | pretty_ren (Named ps) = Pretty.str "where" :: Pretty.brk 1 ::
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    97
          (ps |> map pretty_named |> Pretty.separate "and");
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    98
    fun pretty_rename (loc, ("", ren)) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    99
          Pretty.block (Pretty.str (NewLocale.extern thy loc) :: Pretty.brk 1 :: pretty_ren ren) 
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   100
      | pretty_rename (loc, (prfx, ren)) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   101
          Pretty.block (Pretty.str prfx :: Pretty.brk 1 :: Pretty.str (NewLocale.extern thy loc) ::
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   102
            Pretty.brk 1 :: pretty_ren ren);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   103
  in Pretty.separate "+" (map pretty_rename expr) |> Pretty.block end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   104
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   105
fun err_in_expr thy msg expr =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   106
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   107
    val err_msg =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   108
      if null expr then msg
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   109
      else msg ^ "\n" ^ Pretty.string_of (Pretty.block
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   110
        [Pretty.str "The above error(s) occurred in expression:", Pretty.brk 1,
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   111
          pretty_expr thy expr])
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   112
  in error err_msg end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   113
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   114
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   115
(** Internalise locale names in expr **)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   116
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   117
fun intern thy instances =  map (apfst (NewLocale.intern thy)) instances;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   118
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   119
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   120
(** Parameters of expression.
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   121
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   122
   Sanity check of instantiations and extraction of implicit parameters.
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   123
   The latter only occurs iff strict = false.
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   124
   Positional instantiations are extended to match full length of parameter list
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   125
   of instantiated locale. **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   126
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   127
fun parameters_of thy strict (expr, fixed) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   128
  let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   129
    fun reject_dups message xs =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   130
      let val dups = duplicates (op =) xs
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   131
      in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   132
        if null dups then () else error (message ^ commas dups)
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   133
      end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   134
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   135
    fun match_bind (n, b) = (n = Name.name_of b);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   136
    fun bind_eq (b1, b2) = (Name.name_of b1 = Name.name_of b2);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   137
      (* FIXME: cannot compare bindings for equality. *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   138
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   139
    fun params_loc loc =
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   140
          (NewLocale.params_of thy loc |> map (fn (p, _, mx) => (p, mx)), loc);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   141
    fun params_inst (expr as (loc, (prfx, Positional insts))) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   142
          let
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   143
            val (ps, loc') = params_loc loc;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   144
	    val d = length ps - length insts;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   145
	    val insts' =
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
   146
	      if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
   147
                quote (NewLocale.extern thy loc))
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   148
	      else insts @ replicate d NONE;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   149
            val ps' = (ps ~~ insts') |>
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   150
              map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE);
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   151
          in (ps', (loc', (prfx, Positional insts'))) end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   152
      | params_inst (expr as (loc, (prfx, Named insts))) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   153
          let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   154
            val _ = reject_dups "Duplicate instantiation of the following parameter(s): "
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   155
              (map fst insts);
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   156
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   157
            val (ps, loc') = params_loc loc;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   158
            val ps' = fold (fn (p, _) => fn ps =>
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   159
              if AList.defined match_bind ps p then AList.delete match_bind p ps
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   160
              else error (quote p ^" not a parameter of instantiated expression.")) insts ps;
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   161
          in (ps', (loc', (prfx, Named insts))) end;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   162
    fun params_expr is =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   163
          let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   164
            val (is', ps') = fold_map (fn i => fn ps =>
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   165
              let
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   166
                val (ps', i') = params_inst i;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   167
                val ps'' = AList.join bind_eq (fn p => fn (mx1, mx2) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   168
                  (* FIXME: should check for bindings being the same.
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   169
                     Instead we check for equal name and syntax. *)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   170
                  if mx1 = mx2 then mx1
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   171
                  else error ("Conflicting syntax for parameter" ^ quote (Binding.display p) ^
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   172
                    " in expression.")) (ps, ps')
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   173
              in (i', ps'') end) is []
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   174
          in (ps', is') end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   175
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   176
    val (implicit, expr') = params_expr expr;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   177
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   178
    val implicit' = map (#1 #> Name.name_of) implicit;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   179
    val fixed' = map (#1 #> Name.name_of) fixed;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   180
    val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   181
    val implicit'' = if strict then []
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   182
      else let val _ = reject_dups
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   183
          "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed')
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   184
        in map (fn (b, mx) => (b, NONE, mx)) implicit end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   185
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   186
  in (expr', implicit'' @ fixed) end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   187
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   188
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   189
(** Read instantiation **)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   190
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   191
(* Parse positional or named instantiation *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   192
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   193
local
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   194
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   195
fun prep_inst parse_term parms (Positional insts) ctxt =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   196
      (insts ~~ parms) |> map (fn
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   197
        (NONE, p) => Syntax.parse_term ctxt p |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   198
        (SOME t, _) => parse_term ctxt t)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   199
  | prep_inst parse_term parms (Named insts) ctxt =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   200
      parms |> map (fn p => case AList.lookup (op =) insts p of
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   201
        SOME t => parse_term ctxt t |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   202
        NONE => Syntax.parse_term ctxt p);
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   203
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   204
in
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   205
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   206
fun parse_inst x = prep_inst Syntax.parse_term x;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   207
fun make_inst x = prep_inst (K I) x;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   208
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   209
end;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   210
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   211
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   212
(* Instantiation morphism *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   213
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   214
fun inst_morph (parm_names, parm_types) (prfx, insts') ctxt =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   215
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   216
    (* parameters *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   217
    val type_parm_names = fold Term.add_tfreesT parm_types [] |> map fst;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   218
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   219
    (* type inference and contexts *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   220
    val parm_types' = map (TypeInfer.paramify_vars o Logic.varifyT) parm_types;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   221
    val type_parms = fold Term.add_tvarsT parm_types' [] |> map (Logic.mk_type o TVar);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   222
    val arg = type_parms @ map2 TypeInfer.constrain parm_types' insts';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   223
    val res = Syntax.check_terms ctxt arg;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   224
    val ctxt' = ctxt |> fold Variable.auto_fixes res;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   225
    
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   226
    (* instantiation *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   227
    val (type_parms'', res') = chop (length type_parms) res;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   228
    val insts'' = (parm_names ~~ res') |> map_filter
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   229
      (fn (inst as (x, Free (y, _))) => if x = y then NONE else SOME inst |
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   230
        inst => SOME inst);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   231
    val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   232
    val inst = Symtab.make insts'';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   233
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   234
    (Element.inst_morphism (ProofContext.theory_of ctxt) (instT, inst) $>
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   235
      Morphism.binding_morphism (Binding.qualify prfx), ctxt')
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   236
  end;
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   237
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   238
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   239
(*** Locale processing ***)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   240
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   241
(** Parsing **)
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   242
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   243
fun parse_elem prep_typ prep_term ctxt elem =
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   244
  Element.map_ctxt {binding = I, var = I, typ = prep_typ ctxt,
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   245
    term = prep_term ctxt, fact = I, attrib = I} elem;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   246
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   247
fun parse_concl prep_term ctxt concl =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   248
  (map o map) (fn (t, ps) =>
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   249
    (prep_term ctxt, map (prep_term ctxt) ps)) concl;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   250
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   251
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   252
(** Simultaneous type inference: instantiations + elements + conclusion **)
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   253
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   254
local
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   255
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   256
fun mk_type T = (Logic.mk_type T, []);
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   257
fun mk_term t = (t, []);
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   258
fun mk_propp (p, pats) = (Syntax.type_constraint propT p, pats);
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   259
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   260
fun dest_type (T, []) = Logic.dest_type T;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   261
fun dest_term (t, []) = t;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   262
fun dest_propp (p, pats) = (p, pats);
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   263
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   264
fun extract_inst (_, (_, ts)) = map mk_term ts;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   265
fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs));
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   266
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   267
fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   268
  | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   269
  | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   270
  | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   271
  | extract_elem (Notes _) = [];
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   272
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   273
fun restore_elem (Fixes fixes, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   274
      (fixes ~~ css) |> map (fn ((x, _, mx), cs) =>
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   275
        (x, cs |> map dest_type |> try hd, mx)) |> Fixes
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   276
  | restore_elem (Constrains csts, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   277
      (csts ~~ css) |> map (fn ((x, _), cs) =>
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   278
        (x, cs |> map dest_type |> hd)) |> Constrains
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   279
  | restore_elem (Assumes asms, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   280
      (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   281
  | restore_elem (Defines defs, css) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   282
      (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   283
  | restore_elem (Notes notes, _) = Notes notes;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   284
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   285
fun check cs context =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   286
  let
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   287
    fun prep (_, pats) (ctxt, t :: ts) =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   288
      let val ctxt' = Variable.auto_fixes t ctxt
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   289
      in
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   290
        ((t, Syntax.check_props (ProofContext.set_mode ProofContext.mode_pattern ctxt') pats),
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   291
          (ctxt', ts))
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   292
      end
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   293
    val (cs', (context', _)) = fold_map prep cs
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   294
      (context, Syntax.check_terms
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   295
        (ProofContext.set_mode ProofContext.mode_schematic context) (map fst cs));
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   296
  in (cs', context') end;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   297
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   298
in
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   299
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   300
fun check_autofix insts elems concl ctxt =
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   301
  let
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   302
    val inst_cs = map extract_inst insts;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   303
    val elem_css = map extract_elem elems;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   304
    val concl_cs = (map o map) mk_propp concl;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   305
    (* Type inference *)
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   306
    val (inst_cs' :: css', ctxt') =
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   307
      (fold_burrow o fold_burrow) check (inst_cs :: elem_css @ [concl_cs]) ctxt;
28936
f1647bf418f5 No resolution of patterns within context statements.
ballarin
parents: 28903
diff changeset
   308
    val (elem_css', [concl_cs']) = chop (length elem_css) css';
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   309
  in
28936
f1647bf418f5 No resolution of patterns within context statements.
ballarin
parents: 28903
diff changeset
   310
    (map restore_inst (insts ~~ inst_cs'), map restore_elem (elems ~~ elem_css'),
f1647bf418f5 No resolution of patterns within context statements.
ballarin
parents: 28903
diff changeset
   311
      concl_cs', ctxt')
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   312
  end;
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   313
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   314
end;
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   315
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   316
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   317
(** Prepare locale elements **)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   318
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   319
fun declare_elem prep_vars (Fixes fixes) ctxt =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   320
      let val (vars, _) = prep_vars fixes ctxt
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   321
      in ctxt |> ProofContext.add_fixes_i vars |> snd end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   322
  | declare_elem prep_vars (Constrains csts) ctxt =
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   323
      ctxt |> prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) |> snd
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   324
  | declare_elem _ (Assumes _) ctxt = ctxt
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   325
  | declare_elem _ (Defines _) ctxt = ctxt
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   326
  | declare_elem _ (Notes _) ctxt = ctxt;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   327
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   328
(** Finish locale elements, extract specification text **)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   329
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   330
val norm_term = Envir.beta_norm oo Term.subst_atomic;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   331
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   332
fun abstract_thm thy eq =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   333
  Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   334
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   335
fun bind_def ctxt eq (xs, env, ths) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   336
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   337
    val ((y, T), b) = LocalDefs.abs_def eq;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   338
    val b' = norm_term env b;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   339
    val th = abstract_thm (ProofContext.theory_of ctxt) eq;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   340
    fun err msg = error (msg ^ ": " ^ quote y);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   341
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   342
    exists (fn (x, _) => x = y) xs andalso
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   343
      err "Attempt to define previously specified variable";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   344
    exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   345
      err "Attempt to redefine variable";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   346
    (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   347
  end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   348
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   349
fun eval_text _ _ (Fixes _) text = text
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   350
  | eval_text _ _ (Constrains _) text = text
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   351
  | eval_text _ is_ext (Assumes asms)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   352
        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   353
      let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   354
        val ts = maps (map #1 o #2) asms;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   355
        val ts' = map (norm_term env) ts;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   356
        val spec' =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   357
          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   358
          else ((exts, exts'), (ints @ ts, ints' @ ts'));
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   359
      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   360
  | eval_text ctxt _ (Defines defs) (spec, binds) =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   361
      (spec, fold (bind_def ctxt o #1 o #2) defs binds)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   362
  | eval_text _ _ (Notes _) text = text;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   363
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   364
fun closeup _ _ false elem = elem
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   365
  | closeup ctxt parms true elem =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   366
      let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   367
        fun close_frees t =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   368
          let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   369
            val rev_frees =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   370
              Term.fold_aterms (fn Free (x, T) =>
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   371
                if AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t [];
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   372
          in Term.list_all_free (rev rev_frees, t) end;
29019
8e7d6f959bd7 Explicitly close up defines.
ballarin
parents: 29018
diff changeset
   373
  (* FIXME consider closing in syntactic phase *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   374
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   375
        fun no_binds [] = []
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   376
          | no_binds _ = error "Illegal term bindings in context element";
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   377
      in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   378
        (case elem of
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   379
          Assumes asms => Assumes (asms |> map (fn (a, propps) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   380
            (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   381
        | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
29019
8e7d6f959bd7 Explicitly close up defines.
ballarin
parents: 29018
diff changeset
   382
            (a, (close_frees (#2 (LocalDefs.cert_def ctxt (close_frees t))), no_binds ps))))
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   383
        | e => e)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   384
      end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   385
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   386
fun finish_primitive parms _ (Fixes fixes) = Fixes (map (fn (binding, _, mx) =>
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   387
      let val x = Name.name_of binding
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   388
      in (binding, AList.lookup (op =) parms x, mx) end) fixes)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   389
  | finish_primitive _ _ (Constrains _) = Constrains []
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   390
  | finish_primitive _ close (Assumes asms) = close (Assumes asms)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   391
  | finish_primitive _ close (Defines defs) = close (Defines defs)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   392
  | finish_primitive _ _ (Notes facts) = Notes facts;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   393
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   394
fun finish_inst ctxt parms do_close (loc, (prfx, inst)) text =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   395
  let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   396
    val thy = ProofContext.theory_of ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   397
    val (parm_names, parm_types) = NewLocale.params_of thy loc |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   398
      map (fn (b, SOME T, _) => (Name.name_of b, T)) |> split_list;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   399
    val (asm, defs) = NewLocale.specification_of thy loc;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   400
    val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   401
    val asm' = Option.map (Morphism.term morph) asm;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   402
    val defs' = map (Morphism.term morph) defs;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   403
    val text' = text |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   404
      (if is_some asm
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   405
        then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   406
        else I) |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   407
      (if not (null defs)
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   408
        then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   409
        else I)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   410
(* FIXME clone from new_locale.ML *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   411
  in ((loc, morph), text') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   412
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   413
fun finish_elem ctxt parms do_close elem text =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   414
  let
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   415
    val elem' = finish_primitive parms (closeup ctxt parms do_close) elem;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   416
    val text' = eval_text ctxt true elem' text;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   417
  in (elem', text') end
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   418
  
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   419
fun finish ctxt parms do_close insts elems text =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   420
  let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   421
    val (deps, text') = fold_map (finish_inst ctxt parms do_close) insts text;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   422
    val (elems', text'') = fold_map (finish_elem ctxt parms do_close) elems text';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   423
  in (deps, elems', text'') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   424
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   425
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   426
(** Process full context statement: instantiations + elements + conclusion **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   427
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   428
(* Interleave incremental parsing and type inference over entire parsed stretch. *)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   429
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   430
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   431
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   432
fun prep_full_context_statement parse_typ parse_prop parse_inst prep_vars prep_expr
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   433
    strict do_close context raw_import raw_elems raw_concl =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   434
  let
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   435
    val thy = ProofContext.theory_of context;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   436
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   437
    val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import);
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   438
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   439
    fun prep_inst (loc, (prfx, inst)) (i, insts, ctxt) =
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   440
      let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   441
        val (parm_names, parm_types) = NewLocale.params_of thy loc |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   442
          map (fn (b, SOME T, _) => (Name.name_of b, T)) |> split_list;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   443
        val inst' = parse_inst parm_names inst ctxt;
28885
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   444
        val parm_types' = map (TypeInfer.paramify_vars o
6f6bf52e75bb Expression types cleaned up, proper treatment of term patterns.
ballarin
parents: 28879
diff changeset
   445
          Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT) parm_types;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   446
        val inst'' = map2 TypeInfer.constrain parm_types' inst';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   447
        val insts' = insts @ [(loc, (prfx, inst''))];
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   448
        val (insts'', _, _, ctxt' (* FIXME not used *) ) = check_autofix insts' [] [] ctxt;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   449
        val inst''' = insts'' |> List.last |> snd |> snd;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   450
        val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   451
        val ctxt'' = NewLocale.activate_declarations thy (loc, morph) ctxt;
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   452
      in (i+1, insts', ctxt'') end;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   453
  
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   454
    fun prep_elem raw_elem (insts, elems, ctxt) =
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   455
      let
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   456
        val ctxt' = declare_elem prep_vars raw_elem ctxt;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   457
        val elems' = elems @ [parse_elem parse_typ parse_prop ctxt' raw_elem];
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   458
        (* FIXME term bindings *)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   459
        val (_, _, _, ctxt'') = check_autofix insts elems' [] ctxt';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   460
      in (insts, elems', ctxt') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   461
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   462
    fun prep_concl raw_concl (insts, elems, ctxt) =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   463
      let
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   464
        val concl = (map o map) (fn (t, ps) =>
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   465
          (parse_prop ctxt t, map (parse_prop ctxt) ps)) raw_concl;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   466
      in check_autofix insts elems concl ctxt end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   467
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   468
    val fors = prep_vars fixed context |> fst;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   469
    val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   470
    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], NewLocale.clear_local_idents ctxt);
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   471
    val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   472
    val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   473
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   474
    (* Retrieve parameter types *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   475
    val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Name.name_of o #1) fixes) |
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   476
      _ => fn ps => ps) (Fixes fors :: elems) [];
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   477
    val (Ts, ctxt') = fold_map ProofContext.inferred_param xs ctxt; 
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   478
    val parms = xs ~~ Ts;  (* params from expression and elements *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   479
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   480
    val Fixes fors' = finish_primitive parms I (Fixes fors);
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   481
    val (deps, elems', text) = finish ctxt' parms do_close insts elems ((([], []), ([], [])), ([], [], []));
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   482
    (* text has the following structure:
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   483
           (((exts, exts'), (ints, ints')), (xs, env, defs))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   484
       where
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   485
         exts: external assumptions (terms in assumes elements)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   486
         exts': dito, normalised wrt. env
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   487
         ints: internal assumptions (terms in assumptions from insts)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   488
         ints': dito, normalised wrt. env
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   489
         xs: the free variables in exts' and ints' and rhss of definitions,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   490
           this includes parameters except defined parameters
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   491
         env: list of term pairs encoding substitutions, where the first term
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   492
           is a free variable; substitutions represent defines elements and
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   493
           the rhs is normalised wrt. the previous env
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   494
         defs: theorems representing the substitutions from defines elements
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   495
           (thms are normalised wrt. env).
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   496
       elems is an updated version of raw_elems:
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   497
         - type info added to Fixes and modified in Constrains
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   498
         - axiom and definition statement replaced by corresponding one
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   499
           from proppss in Assumes and Defines
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   500
         - Facts unchanged
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   501
       *)
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   502
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   503
  in ((fors', deps, elems', concl), (parms, text)) end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   504
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   505
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   506
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   507
fun read_full_context_statement x =
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   508
  prep_full_context_statement Syntax.parse_typ Syntax.parse_prop parse_inst
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   509
  ProofContext.read_vars intern x;
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   510
fun cert_full_context_statement x =
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   511
  prep_full_context_statement (K I) (K I) make_inst ProofContext.cert_vars (K I) x;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   512
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   513
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   514
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   515
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   516
(* Context statement: elements + conclusion *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   517
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   518
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   519
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   520
fun prep_statement prep activate raw_elems raw_concl context =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   521
  let
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   522
     val ((_, _, elems, concl), _) = prep true false context ([], []) raw_elems raw_concl;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   523
     val (_, context') = activate elems (ProofContext.set_stmt true context);
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   524
  in (concl, context') end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   525
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   526
in
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   527
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   528
fun read_statement x = prep_statement read_full_context_statement Element.activate x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   529
fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   530
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   531
end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   532
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   533
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   534
(* Locale declaration: import + elements *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   535
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   536
local
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   537
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   538
fun prep_declaration prep activate raw_import raw_elems context =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   539
  let
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   540
    val ((fixed, deps, elems, _), (parms, (spec, (_, _, defs)))) =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   541
      prep false true context raw_import raw_elems [];
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   542
    (* Declare parameters and imported facts *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   543
    val context' = context |>
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   544
      ProofContext.add_fixes_i fixed |> snd |>
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   545
      NewLocale.clear_local_idents |> fold NewLocale.activate_local_facts deps;
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   546
    val ((elems', _), _) = activate elems (ProofContext.set_stmt true context');
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   547
  in ((fixed, deps, elems'), (parms, spec, defs)) end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   548
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   549
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   550
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   551
fun read_declaration x = prep_declaration read_full_context_statement Element.activate x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   552
fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   553
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   554
end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   555
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   556
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   557
(* Locale expression to set up a goal *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   558
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   559
local
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   560
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   561
fun props_of thy (name, morph) =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   562
  let
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   563
    val (asm, defs) = NewLocale.specification_of thy name;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   564
  in
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   565
    (case asm of NONE => defs | SOME asm => asm :: defs) |> map (Morphism.term morph)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   566
  end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   567
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   568
fun prep_goal_expression prep expression context =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   569
  let
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   570
    val thy = ProofContext.theory_of context;
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
   571
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   572
    val ((fixed, deps, _, _), _) = prep true true context expression [] [];
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   573
    (* proof obligations *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   574
    val propss = map (props_of thy) deps;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   575
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   576
    val goal_ctxt = context |>
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   577
      ProofContext.add_fixes_i fixed |> snd |>
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   578
      (fold o fold) Variable.auto_fixes propss;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   579
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   580
    val export = Variable.export_morphism goal_ctxt context;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   581
    val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   582
(*    val exp_term = TermSubst.zero_var_indexes o Morphism.term export; *)
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   583
    val exp_term = Drule.term_rule thy (singleton exp_fact);
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   584
    val exp_typ = Logic.type_map exp_term;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   585
    val export' =
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   586
      Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   587
  in ((propss, deps, export'), goal_ctxt) end;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   588
    
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   589
in
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   590
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   591
fun read_goal_expression x = prep_goal_expression read_full_context_statement x;
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   592
fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x;
28879
db2816a37a34 Read/cert_statement for theorem statements.
ballarin
parents: 28872
diff changeset
   593
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   594
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   595
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   596
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   597
(*** Locale declarations ***)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   598
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   599
(* axiomsN: name of theorem set with destruct rules for locale predicates,
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   600
     also name suffix of delta predicates and assumptions. *)
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   601
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   602
val axiomsN = "axioms";
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   603
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   604
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   605
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   606
(* introN: name of theorems for introduction rules of locale and
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   607
     delta predicates *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   608
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   609
val introN = "intro";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   610
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   611
fun atomize_spec thy ts =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   612
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   613
    val t = Logic.mk_conjunction_balanced ts;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   614
    val body = ObjectLogic.atomize_term thy t;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   615
    val bodyT = Term.fastype_of body;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   616
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   617
    if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   618
    else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   619
  end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   620
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   621
(* achieve plain syntax for locale predicates (without "PROP") *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   622
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   623
fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   624
  if length args = n then
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   625
    Syntax.const "_aprop" $
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   626
      Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   627
  else raise Match);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   628
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   629
(* define one predicate including its intro rule and axioms
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   630
   - bname: predicate name
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   631
   - parms: locale parameters
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   632
   - defs: thms representing substitutions from defines elements
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   633
   - ts: terms representing locale assumptions (not normalised wrt. defs)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   634
   - norm_ts: terms representing locale assumptions (normalised wrt. defs)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   635
   - thy: the theory
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   636
*)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   637
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   638
fun def_pred bname parms defs ts norm_ts thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   639
  let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   640
    val name = Sign.full_bname thy bname;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   641
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   642
    val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   643
    val env = Term.add_term_free_names (body, []);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   644
    val xs = filter (member (op =) env o #1) parms;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   645
    val Ts = map #2 xs;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   646
    val extraTs = (Term.term_tfrees body \\ fold Term.add_tfreesT Ts [])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   647
      |> Library.sort_wrt #1 |> map TFree;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   648
    val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   649
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   650
    val args = map Logic.mk_type extraTs @ map Free xs;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   651
    val head = Term.list_comb (Const (name, predT), args);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   652
    val statement = ObjectLogic.ensure_propT thy head;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   653
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   654
    val ([pred_def], defs_thy) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   655
      thy
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   656
      |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   657
      |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   658
      |> PureThy.add_defs false
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   659
        [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   660
    val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   661
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   662
    val cert = Thm.cterm_of defs_thy;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   663
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   664
    val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   665
      MetaSimplifier.rewrite_goals_tac [pred_def] THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   666
      Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   667
      Tactic.compose_tac (false,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   668
        Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   669
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   670
    val conjuncts =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   671
      (Drule.equal_elim_rule2 OF [body_eq,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   672
        MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   673
      |> Conjunction.elim_balanced (length ts);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   674
    val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   675
      Element.prove_witness defs_ctxt t
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   676
       (MetaSimplifier.rewrite_goals_tac defs THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   677
        Tactic.compose_tac (false, ax, 0) 1));
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   678
  in ((statement, intro, axioms), defs_thy) end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   679
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   680
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   681
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   682
(* CB: main predicate definition function *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   683
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   684
fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   685
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   686
    val (a_pred, a_intro, a_axioms, thy'') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   687
      if null exts then (NONE, NONE, [], thy)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   688
      else
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   689
        let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   690
          val aname = if null ints then pname else pname ^ "_" ^ axiomsN;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   691
          val ((statement, intro, axioms), thy') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   692
            thy
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   693
            |> def_pred aname parms defs exts exts';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   694
          val (_, thy'') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   695
            thy'
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   696
            |> Sign.add_path aname
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   697
            |> Sign.no_base_names
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   698
            |> PureThy.note_thmss Thm.internalK
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   699
              [((Binding.name introN, []), [([intro], [NewLocale.unfold_attrib])])]
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   700
            ||> Sign.restore_naming thy';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   701
          in (SOME statement, SOME intro, axioms, thy'') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   702
    val (b_pred, b_intro, b_axioms, thy'''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   703
      if null ints then (NONE, NONE, [], thy'')
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   704
      else
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   705
        let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   706
          val ((statement, intro, axioms), thy''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   707
            thy''
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   708
            |> def_pred pname parms defs (ints @ the_list a_pred) (ints' @ the_list a_pred);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   709
          val (_, thy'''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   710
            thy'''
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   711
            |> Sign.add_path pname
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   712
            |> Sign.no_base_names
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   713
            |> PureThy.note_thmss Thm.internalK
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   714
                 [((Binding.name introN, []), [([intro], [NewLocale.intro_attrib])]),
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   715
                  ((Binding.name axiomsN, []),
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   716
                    [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   717
            ||> Sign.restore_naming thy''';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   718
        in (SOME statement, SOME intro, axioms, thy'''') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   719
  in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   720
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   721
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   722
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   723
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   724
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   725
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   726
fun assumes_to_notes (Assumes asms) axms =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   727
      fold_map (fn (a, spec) => fn axs =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   728
          let val (ps, qs) = chop (length spec) axs
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   729
          in ((a, [(ps, [])]), qs) end) asms axms
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   730
      |> apfst (curry Notes Thm.assumptionK)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   731
  | assumes_to_notes e axms = (e, axms);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   732
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   733
fun defines_to_notes thy (Defines defs) defns =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   734
    let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   735
      val defs' = map (fn (_, (def, _)) => def) defs
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   736
      val notes = map (fn (a, (def, _)) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   737
        (a, [([assume (cterm_of thy def)], [])])) defs
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   738
    in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   739
      (Notes (Thm.definitionK, notes), defns @ defs')
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   740
    end
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   741
  | defines_to_notes _ e defns = (e, defns);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   742
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   743
fun gen_add_locale prep_decl
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   744
    bname predicate_name raw_imprt raw_body thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   745
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   746
    val thy_ctxt = ProofContext.init thy;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   747
    val name = Sign.full_bname thy bname;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   748
    val _ = NewLocale.test_locale thy name andalso
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   749
      error ("Duplicate definition of locale " ^ quote name);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   750
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   751
    val ((fixed, deps, body_elems), text as (parms, ((_, exts'), _), defs)) =
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   752
      prep_decl raw_imprt raw_body thy_ctxt;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   753
    val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   754
      define_preds predicate_name text thy;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   755
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   756
    val extraTs = fold Term.add_tfrees exts' [] \\ fold Term.add_tfreesT (map snd parms) [];
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   757
    val _ = if null extraTs then ()
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   758
      else warning ("Additional type variable(s) in locale specification " ^ quote bname);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   759
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   760
    val satisfy = Element.satisfy_morphism b_axioms;
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   761
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   762
    val params = fixed @
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   763
      (body_elems |> map_filter (fn Fixes fixes => SOME fixes | _ => NONE) |> flat);
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   764
    val asm = if is_some b_statement then b_statement else a_statement;
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   765
    val (body_elems', defns) = fold_map (defines_to_notes thy') body_elems [];
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   766
    val notes = body_elems' |>
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   767
      (fn elems => fold_map assumes_to_notes elems (map Element.conclude_witness a_axioms)) |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   768
      fst |> map (Element.morph_ctxt satisfy) |>
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   769
      map_filter (fn Notes notes => SOME notes | _ => NONE) |>
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   770
      (if is_some asm
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28951
diff changeset
   771
        then cons (Thm.internalK, [((Binding.name (bname ^ "_" ^ axiomsN), []),
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   772
          [([assume (cterm_of thy' (the asm))], [(Attrib.internal o K) NewLocale.witness_attrib])])])
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   773
        else I);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   774
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   775
    val deps' = map (fn (l, morph) => (l, morph $> satisfy)) deps;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   776
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   777
    val loc_ctxt = thy' |>
28991
694227dd3e8c dropped NameSpace.declare_base
haftmann
parents: 28965
diff changeset
   778
      NewLocale.register_locale bname (extraTs, params)
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   779
        (asm, map prop_of defs) ([], [])
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   780
        (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev) |>
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   781
      NewLocale.init name
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   782
  in (name, loc_ctxt) end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   783
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   784
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   785
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   786
val add_locale_cmd = gen_add_locale read_declaration;
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   787
val add_locale = gen_add_locale cert_declaration;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   788
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   789
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   790
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   791
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   792
(*** Interpretation ***)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   793
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   794
(** Witnesses and goals **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   795
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   796
fun prep_propp propss = propss |> map (map (rpair [] o Element.mark_witness));
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   797
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   798
fun prep_result propps thmss =
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   799
  ListPair.map (fn (props, thms) => map2 Element.make_witness props thms) (propps, thmss);
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   800
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   801
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   802
(** Interpretation between locales: declaring sublocale relationships **)
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   803
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   804
local
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   805
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   806
fun gen_sublocale prep_expr intern
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   807
    raw_target expression thy =
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   808
  let
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   809
    val target = intern thy raw_target;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   810
    val target_ctxt = NewLocale.init target thy;
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   811
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   812
    val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
28898
530c7d28a962 Proper treatment of expressions with free arguments.
ballarin
parents: 28895
diff changeset
   813
    
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   814
    fun store_dep ((name, morph), thms) =
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   815
      NewLocale.add_dependency target (name, morph $> Element.satisfy_morphism thms $> export);
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   816
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   817
    fun after_qed results =
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   818
      ProofContext.theory (
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   819
        (* store dependencies *)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   820
        fold store_dep (deps ~~ prep_result propss results) #>
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   821
        (* propagate registrations *)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   822
        (fn thy => fold_rev (fn reg => NewLocale.activate_global_facts reg)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   823
          (NewLocale.get_global_registrations thy) thy));
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   824
  in
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   825
    goal_ctxt |>
28951
e89dde5f365c Sublocale: removed public after_qed; identifiers private to NewLocale.
ballarin
parents: 28936
diff changeset
   826
      Proof.theorem_i NONE after_qed (prep_propp propss) |>
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   827
      Element.refine_witness |> Seq.hd
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   828
  end;
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   829
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   830
in
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   831
28902
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   832
fun sublocale_cmd x = gen_sublocale read_goal_expression NewLocale.intern x;
2019bcc9d8bf Ahere to modern naming conventions; proper treatment of internal vs external names.
ballarin
parents: 28898
diff changeset
   833
fun sublocale x = gen_sublocale cert_goal_expression (K I) x;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   834
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   835
end;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   836
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   837
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   838
(** Interpretation in theories **)
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   839
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   840
local
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   841
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   842
fun gen_interpretation prep_expr
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   843
    expression thy =
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   844
  let
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   845
    val ctxt = ProofContext.init thy;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   846
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   847
    val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   848
    
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   849
    fun store_reg ((name, morph), thms) =
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   850
      let
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   851
        val morph' = morph $> Element.satisfy_morphism thms $> export;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   852
      in
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   853
        NewLocale.add_global_registration (name, morph') #>
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   854
        NewLocale.activate_global_facts (name, morph')
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   855
      end;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   856
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   857
    fun after_qed results =
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   858
      ProofContext.theory (fold store_reg (regs ~~ prep_result propss results));
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   859
  in
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   860
    goal_ctxt |>
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   861
      Proof.theorem_i NONE after_qed (prep_propp propss) |>
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   862
      Element.refine_witness |> Seq.hd
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   863
  end;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   864
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   865
in
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   866
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   867
fun interpretation_cmd x = gen_interpretation read_goal_expression x;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   868
fun interpretation x = gen_interpretation cert_goal_expression x;
28895
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   869
4e2914c2f8c5 Sublocale command.
ballarin
parents: 28885
diff changeset
   870
end;
28903
b3fc3a62247a Intro_locales_tac to simplify goals involving locale predicates.
ballarin
parents: 28902
diff changeset
   871
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   872
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   873
(** Interpretation in proof contexts **)
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   874
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   875
local
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   876
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   877
fun gen_interpret prep_expr
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   878
    expression int state =
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   879
  let
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   880
    val _ = Proof.assert_forward_or_chain state;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   881
    val ctxt = Proof.context_of state;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   882
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   883
    val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   884
    
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   885
    fun store_reg ((name, morph), thms) =
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   886
      let
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   887
        val morph' = morph $> Element.satisfy_morphism thms $> export;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   888
      in
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   889
        NewLocale.activate_local_facts (name, morph')
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   890
      end;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   891
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   892
    fun after_qed results =
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   893
      Proof.map_context (fold store_reg (regs ~~ prep_result propss results)) #> Seq.single;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   894
  in
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   895
    state |> Proof.map_context (K goal_ctxt) |>
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   896
      Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   897
      "interpret" NONE after_qed (map (pair (Binding.empty, [])) (prep_propp propss)) |>
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   898
      Element.refine_witness |> Seq.hd
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   899
  end;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   900
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   901
in
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   902
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   903
fun interpret_cmd x = gen_interpret read_goal_expression x;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   904
fun interpret x = gen_interpret cert_goal_expression x;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   905
28993
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   906
end;
829e684b02ef Interpretation in theories including interaction with subclass relation.
ballarin
parents: 28951
diff changeset
   907
29018
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   908
end;
17538bdef546 Interpretation in proof contexts.
ballarin
parents: 28994
diff changeset
   909