src/Pure/Isar/expression.ML
author ballarin
Fri, 21 Nov 2008 18:01:39 +0100
changeset 28872 686963dbf6cd
parent 28862 53f13f763d4f
child 28879 db2816a37a34
permissions -rw-r--r--
add_locale functional.
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
    ID:         $Id$
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     3
    Author:     Clemens Ballarin, TU Muenchen
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     4
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
     5
New locale development --- experimental.
28697
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
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     8
signature EXPRESSION =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
     9
sig
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    10
  type 'term map
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    11
  type 'morph expr
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    12
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    13
  val empty_expr: 'morph expr
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    14
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    15
  type expression = (string * string map) expr * (Name.binding * string option * mixfix) list
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    16
(*  type expression_i = Morphism.morphism expr * (Name.binding * typ option * mixfix) list *)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    17
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    18
  (* Declaring locales *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    19
  val add_locale: string -> bstring -> expression -> Element.context list -> theory ->
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    20
    string * Proof.context
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    21
(*
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    22
  val add_locale_i: string -> bstring -> expression_i -> Element.context_i list -> theory ->
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    23
    string * Proof.context
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    24
*)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    25
  (* Debugging and development *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    26
  val parse_expression: OuterParse.token list -> expression * OuterParse.token list
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    27
end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    28
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    29
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    30
structure Expression: EXPRESSION =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    31
struct
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    32
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    33
datatype ctxt = datatype Element.ctxt;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    34
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    35
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    36
(*** Expressions ***)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    37
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    38
datatype 'term map =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    39
  Positional of 'term option list |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    40
  Named of (string * 'term) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    41
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    42
datatype 'morph expr = Expr of (string * 'morph) list;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    43
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
    44
type expression = (string * string map) expr * (Name.binding * string option * mixfix) list;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    45
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    46
val empty_expr = Expr [];
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    47
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    48
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    49
(** Parsing and printing **)
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    50
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    51
local
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    52
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    53
structure P = OuterParse;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    54
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    55
val loc_keyword = P.$$$ "fixes" || P.$$$ "constrains" || P.$$$ "assumes" ||
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    56
   P.$$$ "defines" || P.$$$ "notes";
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    57
fun plus1_unless test scan =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    58
  scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan));
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
val prefix = P.name --| P.$$$ ":";
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    61
val named = P.name -- (P.$$$ "=" |-- P.term);
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    62
val position = P.maybe P.term;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    63
val instance = P.$$$ "where" |-- P.and_list1 named >> Named ||
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    64
  Scan.repeat1 position >> Positional;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    65
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    66
in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    67
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    68
val parse_expression =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    69
  let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    70
    fun expr2 x = P.xname x;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    71
    fun expr1 x = (Scan.optional prefix "" -- expr2 --
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    72
      Scan.optional instance (Named []) >> (fn ((p, l), i) => (l, (p, i)))) x;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    73
    fun expr0 x = (plus1_unless loc_keyword expr1 >> Expr) x;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    74
  in expr0 -- P.for_fixes end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    75
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    76
end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
    77
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    78
fun pretty_expr thy (Expr expr) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    79
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    80
    fun pretty_pos NONE = Pretty.str "_"
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    81
      | pretty_pos (SOME x) = Pretty.str x;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    82
    fun pretty_named (x, y) = [Pretty.str x, Pretty.brk 1, Pretty.str "=",
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    83
          Pretty.brk 1, Pretty.str y] |> Pretty.block;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    84
    fun pretty_ren (Positional ps) = take_suffix is_none ps |> snd |>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    85
          map pretty_pos |> Pretty.breaks
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    86
      | pretty_ren (Named []) = []
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    87
      | pretty_ren (Named ps) = Pretty.str "where" :: Pretty.brk 1 ::
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    88
          (ps |> map pretty_named |> Pretty.separate "and");
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    89
    fun pretty_rename (loc, ("", ren)) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    90
          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
    91
      | pretty_rename (loc, (prfx, ren)) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    92
          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
    93
            Pretty.brk 1 :: pretty_ren ren);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    94
  in Pretty.separate "+" (map pretty_rename expr) |> Pretty.block end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    95
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    96
fun err_in_expr thy msg (Expr expr) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    97
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    98
    val err_msg =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
    99
      if null expr then msg
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   100
      else msg ^ "\n" ^ Pretty.string_of (Pretty.block
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   101
        [Pretty.str "The above error(s) occurred in expression:", Pretty.brk 1,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   102
          pretty_expr thy (Expr expr)])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   103
  in error err_msg end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   104
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   105
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   106
(** Internalise locale names in expr **)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   107
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   108
fun intern thy (Expr instances) = Expr (map (apfst (NewLocale.intern thy)) instances);
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   109
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   110
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   111
(** Parameters of expression.
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   112
   Sanity check of instantiations.
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   113
   Positional instantiations are extended to match full length of parameter list. **)
28697
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
fun parameters_of thy (expr, fixed) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   116
  let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   117
    fun reject_dups message xs =
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   118
      let val dups = duplicates (op =) xs
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   119
      in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   120
        if null dups then () else error (message ^ commas dups)
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   121
      end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   122
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   123
    fun match_bind (n, b) = (n = Name.name_of b);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   124
    fun bind_eq (b1, b2) = (Name.name_of b1 = Name.name_of b2);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   125
      (* FIXME: cannot compare bindings for equality. *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   126
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   127
    fun params_loc loc =
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   128
          (NewLocale.params_of thy loc |> map (fn (p, _, mx) => (p, mx)), loc);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   129
    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
   130
          let
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   131
            val (ps, loc') = params_loc loc;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   132
	    val d = length ps - length insts;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   133
	    val insts' =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   134
	      if d < 0 then err_in_expr thy "More arguments than parameters in instantiation."
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   135
                (Expr [expr])
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   136
	      else insts @ replicate d NONE;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   137
            val ps' = (ps ~~ insts') |>
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   138
              map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE);
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   139
          in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   140
            (ps', (loc', (prfx, Positional insts')))
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   141
          end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   142
      | params_inst (expr as (loc, (prfx, Named insts))) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   143
          let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   144
            val _ = reject_dups "Duplicate instantiation of the following parameter(s): "
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   145
              (map fst insts);
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   146
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   147
            val (ps, loc') = params_loc loc;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   148
            val ps' = fold (fn (p, _) => fn ps =>
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   149
              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
   150
              else error (quote p ^" not a parameter of instantiated expression.")) insts ps;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   151
          in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   152
            (ps', (loc', (prfx, Named insts)))
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   153
          end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   154
    fun params_expr (Expr is) =
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   155
          let
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   156
            val (is', ps') = fold_map (fn i => fn ps =>
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   157
              let
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   158
                val (ps', i') = params_inst i;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   159
                val ps'' = AList.join bind_eq (fn p => fn (mx1, mx2) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   160
                  (* FIXME: should check for bindings being the same.
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   161
                     Instead we check for equal name and syntax. *)
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   162
                  if mx1 = mx2 then mx1
28862
53f13f763d4f tuned name bindings
haftmann
parents: 28859
diff changeset
   163
                  else error ("Conflicting syntax for parameter" ^ quote (Name.display p) ^
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   164
                    " in expression.")) (ps, ps')
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   165
              in (i', ps'') end) is []
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   166
          in
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   167
            (ps', Expr is')
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   168
          end;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   169
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   170
    val (parms, expr') = params_expr expr;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   171
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   172
    val parms' = map (#1 #> Name.name_of) parms;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   173
    val fixed' = map (#1 #> Name.name_of) fixed;
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   174
    val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   175
    val _ = reject_dups "Parameter(s) declared simultaneously in expression and for clause: " (parms' @ fixed');
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   176
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   177
  in (expr', (map (fn (b, mx) => (b, NONE, mx)) parms @ fixed)) end;
28697
140bfb63f893 New-style locale expressions with instantiation (new file expression.ML).
ballarin
parents:
diff changeset
   178
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   179
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   180
(** Read instantiation **)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   181
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   182
(* Parse positional or named instantiation *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   183
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   184
local
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   185
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   186
fun prep_inst parse_term parms (Positional insts) ctxt =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   187
      (insts ~~ parms) |> map (fn
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   188
        (NONE, p) => Syntax.parse_term ctxt p |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   189
        (SOME t, _) => parse_term ctxt t)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   190
  | prep_inst parse_term parms (Named insts) ctxt =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   191
      parms |> map (fn p => case AList.lookup (op =) insts p of
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   192
        SOME t => parse_term ctxt t |
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   193
        NONE => Syntax.parse_term ctxt p);
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   194
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   195
in
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   196
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   197
fun parse_inst x = prep_inst Syntax.parse_term x;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   198
fun make_inst x = prep_inst (K I) x;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   199
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   200
end;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   201
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   202
(* Prepare type inference problem for Syntax.check_terms *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   203
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   204
fun varify_indexT i ty = ty |> Term.map_atyps
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   205
  (fn TFree (a, S) => TVar ((a, i), S)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   206
    | TVar (ai, _) => raise TYPE ("Illegal schematic variable: " ^
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   207
        quote (Term.string_of_vname ai), [ty], []));
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   208
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   209
(* Instantiation morphism *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   210
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   211
fun inst_morph (parm_names, parm_types) (prfx, insts') ctxt =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   212
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   213
    (* parameters *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   214
    val type_parm_names = fold Term.add_tfreesT parm_types [] |> map fst;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   215
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   216
    (* type inference and contexts *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   217
    val parm_types' = map (TypeInfer.paramify_vars o Logic.varifyT) parm_types;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   218
    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
   219
    val arg = type_parms @ map2 TypeInfer.constrain parm_types' insts';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   220
    val res = Syntax.check_terms ctxt arg;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   221
    val ctxt' = ctxt |> fold Variable.auto_fixes res;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   222
    
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   223
    (* instantiation *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   224
    val (type_parms'', res') = chop (length type_parms) res;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   225
    val insts'' = (parm_names ~~ res') |> map_filter
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   226
      (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
   227
        inst => SOME inst);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   228
    val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   229
    val inst = Symtab.make insts'';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   230
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   231
    (Element.inst_morphism (ProofContext.theory_of ctxt) (instT, inst) $>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   232
      Morphism.name_morphism (Name.qualified prfx), ctxt')
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   233
  end;
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   234
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   235
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   236
(*** Locale processing ***)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   237
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   238
(** Parsing **)
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   239
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   240
fun parse_elem prep_typ prep_term ctxt elem =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   241
  Element.map_ctxt {name = I, var = I, typ = prep_typ ctxt,
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   242
    term = prep_term ctxt, fact = I, attrib = I} elem;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   243
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   244
fun parse_concl prep_term ctxt concl =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   245
  (map o map) (fn (t, ps) =>
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   246
    (prep_term ctxt, map (prep_term ctxt) ps)) concl;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   247
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   248
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   249
(** Type checking **)
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
fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map (Logic.mk_type #> rpair [])) fixes
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   252
  | extract_elem (Constrains csts) = map (#2 #> single #> map (Logic.mk_type #> rpair [])) csts
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   253
  | extract_elem (Assumes asms) = map #2 asms
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   254
  | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [(t, ps)]) defs
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   255
  | extract_elem (Notes _) = [];
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   256
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   257
fun restore_elem (Fixes fixes, propps) =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   258
      (fixes ~~ propps) |> map (fn ((x, _, mx), propp) =>
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   259
        (x, propp |> map (fst #> Logic.dest_type) |> try hd, mx)) |> Fixes
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   260
  | restore_elem (Constrains csts, propps) =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   261
      (csts ~~ propps) |> map (fn ((x, _), propp) =>
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   262
        (x, propp |> map (fst #> Logic.dest_type) |> hd)) |> Constrains
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   263
  | restore_elem (Assumes asms, propps) =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   264
      (asms ~~ propps) |> map (fn ((b, _), propp) => (b, propp)) |> Assumes
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   265
  | restore_elem (Defines defs, propps) =
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   266
      (defs ~~ propps) |> map (fn ((b, _), [propp]) => (b, propp)) |> Defines
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   267
  | restore_elem (Notes notes, _) = Notes notes;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   268
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   269
fun check_autofix insts elems concl ctxt =
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   270
  let
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   271
    val instss = map (snd o snd) insts |> (map o map) (fn t => (t, []));
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   272
    val elemss = elems |> map extract_elem;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   273
    val all_terms' = (burrow o burrow_fst) (Syntax.check_terms ctxt) (concl @ instss @ flat elemss); 
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   274
(*    val (ctxt', all_props') = ProofContext.check_propp_schematic (ctxt, concl @ flat propss); *)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   275
    val ctxt'' = (fold o fold) (fn (t, _) => Variable.auto_fixes t) all_terms' ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   276
    val (concl', mores') = chop (length concl) all_terms';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   277
    val (insts', elems') = chop (length instss) mores';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   278
  in (insts' |> (map o map) fst |> curry (op ~~) insts |> map (fn ((l, (p, _)), is) => (l, (p, is))),
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   279
    elems' |> unflat elemss |> curry (op ~~) elems |> map restore_elem, concl', ctxt'') end;
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   280
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   281
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   282
(** Prepare locale elements **)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   283
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   284
fun declare_elem prep_vars (Fixes fixes) ctxt =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   285
      let val (vars, _) = prep_vars fixes ctxt
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   286
      in ctxt |> ProofContext.add_fixes_i vars |> snd end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   287
  | declare_elem prep_vars (Constrains csts) ctxt =
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   288
      ctxt |> prep_vars (map (fn (x, T) => (Name.binding x, SOME T, NoSyn)) csts) |> snd
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   289
  | declare_elem _ (Assumes _) ctxt = ctxt
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   290
  | declare_elem _ (Defines _) ctxt = ctxt
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   291
  | declare_elem _ (Notes _) ctxt = ctxt;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   292
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   293
(** Finish locale elements, extract specification text **)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   294
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   295
val norm_term = Envir.beta_norm oo Term.subst_atomic;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   296
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   297
fun abstract_thm thy eq =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   298
  Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   299
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   300
fun bind_def ctxt eq (xs, env, ths) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   301
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   302
    val ((y, T), b) = LocalDefs.abs_def eq;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   303
    val b' = norm_term env b;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   304
    val th = abstract_thm (ProofContext.theory_of ctxt) eq;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   305
    fun err msg = error (msg ^ ": " ^ quote y);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   306
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   307
    exists (fn (x, _) => x = y) xs andalso
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   308
      err "Attempt to define previously specified variable";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   309
    exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   310
      err "Attempt to redefine variable";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   311
    (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   312
  end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   313
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   314
fun eval_text _ _ (Fixes _) text = text
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   315
  | eval_text _ _ (Constrains _) text = text
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   316
  | eval_text _ is_ext (Assumes asms)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   317
        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   318
      let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   319
        val ts = maps (map #1 o #2) asms;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   320
        val ts' = map (norm_term env) ts;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   321
        val spec' =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   322
          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   323
          else ((exts, exts'), (ints @ ts, ints' @ ts'));
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   324
      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   325
  | eval_text ctxt _ (Defines defs) (spec, binds) =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   326
      (spec, fold (bind_def ctxt o #1 o #2) defs binds)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   327
  | eval_text _ _ (Notes _) text = text;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   328
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   329
fun closeup _ _ false elem = elem
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   330
  | closeup ctxt parms true elem =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   331
      let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   332
        fun close_frees t =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   333
          let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   334
            val rev_frees =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   335
              Term.fold_aterms (fn Free (x, T) =>
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   336
                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
   337
          in Term.list_all_free (rev rev_frees, t) end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   338
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   339
        fun no_binds [] = []
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   340
          | no_binds _ = error "Illegal term bindings in context element";
28795
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
        (case elem of
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   343
          Assumes asms => Assumes (asms |> map (fn (a, propps) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   344
            (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   345
        | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   346
            (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   347
        | e => e)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   348
      end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   349
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   350
fun finish_primitive parms _ (Fixes fixes) = Fixes (map (fn (binding, _, mx) =>
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   351
      let val x = Name.name_of binding
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   352
      in (binding, AList.lookup (op =) parms x, mx) end) fixes)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   353
  | finish_primitive _ _ (Constrains _) = Constrains []
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   354
  | finish_primitive _ close (Assumes asms) = close (Assumes asms)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   355
  | finish_primitive _ close (Defines defs) = close (Defines defs)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   356
  | finish_primitive _ _ (Notes facts) = Notes facts;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   357
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   358
fun finish_inst ctxt parms do_close (loc, (prfx, inst)) text =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   359
  let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   360
    val thy = ProofContext.theory_of ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   361
    val (parm_names, parm_types) = NewLocale.params_of thy loc |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   362
      map (fn (b, SOME T, _) => (Name.name_of b, T)) |> split_list;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   363
    val (asm, defs) = NewLocale.specification_of thy loc;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   364
    val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   365
    val asm' = Option.map (Morphism.term morph) asm;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   366
    val defs' = map (Morphism.term morph) defs;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   367
    val text' = text |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   368
      (if is_some asm
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   369
        then eval_text ctxt false (Assumes [(Attrib.no_binding, [(the asm', [])])])
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   370
        else I) |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   371
      (if not (null defs)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   372
        then eval_text ctxt false (Defines (map (fn def => (Attrib.no_binding, (def, []))) defs'))
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   373
        else I)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   374
(* FIXME clone from new_locale.ML *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   375
  in ((loc, morph), text') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   376
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   377
fun finish_elem ctxt parms do_close elem text =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   378
  let
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   379
    val elem' = finish_primitive parms (closeup ctxt parms do_close) elem;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   380
    val text' = eval_text ctxt true elem' text;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   381
  in (elem', text') end
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   382
  
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   383
fun finish ctxt parms do_close insts elems text =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   384
  let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   385
    val (deps, text') = fold_map (finish_inst ctxt parms do_close) insts text;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   386
    val (elems', text'') = fold_map (finish_elem ctxt parms do_close) elems text';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   387
  in (deps, elems', text'') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   388
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   389
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   390
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   391
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   392
(* nice, but for now not needed
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   393
fun fold_suffixes f [] y = f [] y
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   394
  | fold_suffixes f (x :: xs) y = f (x :: xs) (f xs y);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   395
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   396
fun fold_prefixes f xs y = fold_suffixes (f o rev) (rev xs) y;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   397
*)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   398
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   399
fun prep_elems parse_typ parse_prop parse_inst prep_vars
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   400
    do_close context fixed raw_insts raw_elems raw_concl =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   401
  let
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   402
    val thy = ProofContext.theory_of context;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   403
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   404
    fun prep_inst (loc, (prfx, inst)) (i, ids, insts, ctxt) =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   405
      let
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   406
        val (parm_names, parm_types) = NewLocale.params_of thy loc |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   407
          map (fn (b, SOME T, _) => (Name.name_of b, T)) |> split_list;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   408
        val inst' = parse_inst parm_names inst ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   409
        val parm_types' = map (TypeInfer.paramify_vars o varify_indexT i) parm_types;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   410
        val inst'' = map2 TypeInfer.constrain parm_types' inst';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   411
        val insts' = insts @ [(loc, (prfx, inst''))];
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   412
        val (insts'', _, _, ctxt') = check_autofix insts' [] [] ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   413
        val inst''' = insts'' |> List.last |> snd |> snd;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   414
        val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   415
        val (ids', ctxt'') = NewLocale.activate_declarations loc morph thy ids ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   416
      in (i+1, ids', insts', ctxt'') end;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   417
  
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   418
    fun prep_elem raw_elem (insts, elems, ctxt) =
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   419
      let
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   420
        val ctxt' = declare_elem prep_vars raw_elem ctxt;
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   421
        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
   422
        (* FIXME term bindings *)
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   423
        val (_, _, _, ctxt'') = check_autofix insts elems' [] ctxt';
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   424
      in (insts, elems', ctxt') end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   425
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   426
    fun prep_concl raw_concl (insts, elems, ctxt) =
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   427
      let
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   428
        val concl = (map o map) (fn (t, ps) =>
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   429
          (parse_prop ctxt t, map (parse_prop ctxt) ps)) raw_concl;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   430
      in check_autofix insts elems concl ctxt end;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   431
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   432
    val fors = prep_vars fixed context |> fst;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   433
    val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   434
    val (_, _, insts', ctxt') = fold prep_inst raw_insts (0, NewLocale.empty, [], ctxt);
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   435
    val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   436
    val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   437
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   438
    (* Retrieve parameter types *)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   439
    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
   440
      _ => fn ps => ps) (Fixes fors :: elems) [];
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   441
    val (Ts, ctxt') = fold_map ProofContext.inferred_param xs ctxt; 
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   442
    val parms = xs ~~ Ts;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   443
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   444
    val Fixes fors' = finish_primitive parms I (Fixes fors);
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   445
    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
   446
    (* text has the following structure:
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   447
           (((exts, exts'), (ints, ints')), (xs, env, defs))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   448
       where
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   449
         exts: external assumptions (terms in assumes elements)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   450
         exts': dito, normalised wrt. env
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   451
         ints: internal assumptions (terms in assumptions from insts)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   452
         ints': dito, normalised wrt. env
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   453
         xs: the free variables in exts' and ints' and rhss of definitions,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   454
           this includes parameters except defined parameters
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   455
         env: list of term pairs encoding substitutions, where the first term
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   456
           is a free variable; substitutions represent defines elements and
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   457
           the rhs is normalised wrt. the previous env
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   458
         defs: theorems representing the substitutions from defines elements
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   459
           (thms are normalised wrt. env).
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   460
       elems is an updated version of raw_elems:
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   461
         - type info added to Fixes and modified in Constrains
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   462
         - axiom and definition statement replaced by corresponding one
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   463
           from proppss in Assumes and Defines
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   464
         - Facts unchanged
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   465
       *)
28852
5ddea758679b Type inference for elements through syntax module.
ballarin
parents: 28832
diff changeset
   466
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   467
  in ((parms, fors', deps, elems', concl), text) end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   468
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   469
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   470
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   471
fun read_elems x = prep_elems Syntax.parse_typ Syntax.parse_prop parse_inst
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   472
  ProofContext.read_vars x;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   473
fun cert_elems x = prep_elems (K I) (K I) make_inst
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   474
  ProofContext.cert_vars x;
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   475
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   476
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   477
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   478
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   479
(* full context statements: import + elements + conclusion *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   480
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   481
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   482
28832
cf7237498e7a Activate elements moved to element.ML.
ballarin
parents: 28795
diff changeset
   483
fun prep_context_statement prep_expr prep_elems
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   484
    do_close imprt elements raw_concl context =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   485
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   486
    val thy = ProofContext.theory_of context;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   487
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   488
    val (Expr expr, fixed) = parameters_of thy (apfst (prep_expr thy) imprt);
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   489
    val ctxt = context |> ProofContext.add_fixes fixed |> snd;
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   490
    (* FIXME push inside prep_elems *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   491
    val ((parms, fors, deps, elems, concl), (spec, (_, _, defs))) =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   492
      prep_elems do_close context fixed expr elements raw_concl;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   493
    (* FIXME activate deps *)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   494
    val ((elems', _), ctxt') =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   495
      Element.activate elems (ProofContext.set_stmt true ctxt);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   496
  in
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   497
    (((fors, deps), (ctxt', elems'), (parms, spec, defs)), concl)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   498
  end
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   499
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   500
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   501
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   502
fun read_context imprt body ctxt =
28859
d50b523c55db Deleted debug message (PolyML).
ballarin
parents: 28852
diff changeset
   503
  #1 (prep_context_statement intern read_elems true imprt body [] ctxt);
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   504
(*
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   505
fun cert_context imprt body ctxt =
28832
cf7237498e7a Activate elements moved to element.ML.
ballarin
parents: 28795
diff changeset
   506
  #1 (prep_context_statement (K I) cert_elems true imprt body [] ctxt);
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   507
*)
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   508
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   509
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   510
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   511
(** Dependencies **)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   512
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   513
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   514
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   515
(*** Locale declarations ***)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   516
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   517
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   518
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   519
(* introN: name of theorems for introduction rules of locale and
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   520
     delta predicates;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   521
   axiomsN: name of theorem set with destruct rules for locale predicates,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   522
     also name suffix of delta predicates. *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   523
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   524
val introN = "intro";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   525
val axiomsN = "axioms";
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   526
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   527
fun atomize_spec thy ts =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   528
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   529
    val t = Logic.mk_conjunction_balanced ts;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   530
    val body = ObjectLogic.atomize_term thy t;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   531
    val bodyT = Term.fastype_of body;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   532
  in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   533
    if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   534
    else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   535
  end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   536
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   537
(* achieve plain syntax for locale predicates (without "PROP") *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   538
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   539
fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   540
  if length args = n then
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   541
    Syntax.const "_aprop" $
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   542
      Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   543
  else raise Match);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   544
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   545
(* CB: define one predicate including its intro rule and axioms
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   546
   - bname: predicate name
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   547
   - parms: locale parameters
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   548
   - defs: thms representing substitutions from defines elements
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   549
   - ts: terms representing locale assumptions (not normalised wrt. defs)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   550
   - norm_ts: terms representing locale assumptions (normalised wrt. defs)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   551
   - thy: the theory
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   552
*)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   553
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   554
fun def_pred bname parms defs ts norm_ts thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   555
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   556
    val name = Sign.full_name thy bname;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   557
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   558
    val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   559
    val env = Term.add_term_free_names (body, []);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   560
    val xs = filter (member (op =) env o #1) parms;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   561
    val Ts = map #2 xs;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   562
    val extraTs = (Term.term_tfrees body \\ fold Term.add_tfreesT Ts [])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   563
      |> Library.sort_wrt #1 |> map TFree;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   564
    val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   565
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   566
    val args = map Logic.mk_type extraTs @ map Free xs;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   567
    val head = Term.list_comb (Const (name, predT), args);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   568
    val statement = ObjectLogic.ensure_propT thy head;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   569
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   570
    val ([pred_def], defs_thy) =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   571
      thy
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   572
      |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   573
      |> Sign.declare_const [] ((Name.binding bname, predT), NoSyn) |> snd
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   574
      |> PureThy.add_defs false
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   575
        [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   576
    val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   577
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   578
    val cert = Thm.cterm_of defs_thy;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   579
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   580
    val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   581
      MetaSimplifier.rewrite_goals_tac [pred_def] THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   582
      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
   583
      Tactic.compose_tac (false,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   584
        Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   585
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   586
    val conjuncts =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   587
      (Drule.equal_elim_rule2 OF [body_eq,
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   588
        MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   589
      |> Conjunction.elim_balanced (length ts);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   590
    val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   591
      Element.prove_witness defs_ctxt t
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   592
       (MetaSimplifier.rewrite_goals_tac defs THEN
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   593
        Tactic.compose_tac (false, ax, 0) 1));
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   594
  in ((statement, intro, axioms), defs_thy) 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
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   597
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   598
(* CB: main predicate definition function *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   599
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   600
fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   601
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   602
    val (a_pred, a_intro, a_axioms, thy'') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   603
      if null exts then (NONE, NONE, [], thy)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   604
      else
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   605
        let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   606
          val aname = if null ints then pname else pname ^ "_" ^ axiomsN;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   607
          val ((statement, intro, axioms), thy') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   608
            thy
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   609
            |> def_pred aname parms defs exts exts';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   610
          val (_, thy'') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   611
            thy'
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   612
            |> Sign.add_path aname
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   613
            |> Sign.no_base_names
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   614
            |> PureThy.note_thmss Thm.internalK [((Name.binding introN, []), [([intro], [])])]
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   615
            ||> Sign.restore_naming thy';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   616
          in (SOME statement, SOME intro, axioms, thy'') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   617
    val (b_pred, b_intro, b_axioms, thy'''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   618
      if null ints then (NONE, NONE, [], thy'')
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   619
      else
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   620
        let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   621
          val ((statement, intro, axioms), thy''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   622
            thy''
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   623
            |> 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
   624
          val (_, thy'''') =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   625
            thy'''
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   626
            |> Sign.add_path pname
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   627
            |> Sign.no_base_names
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   628
            |> PureThy.note_thmss Thm.internalK
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   629
                 [((Name.binding introN, []), [([intro], [])]),
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   630
                  ((Name.binding axiomsN, []),
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   631
                    [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   632
            ||> Sign.restore_naming thy''';
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   633
        in (SOME statement, SOME intro, axioms, thy'''') end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   634
  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
   635
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   636
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   637
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   638
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   639
local
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   640
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   641
fun assumes_to_notes (Assumes asms) axms =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   642
      fold_map (fn (a, spec) => fn axs =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   643
          let val (ps, qs) = chop (length spec) axs
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   644
          in ((a, [(ps, [])]), qs) end) asms axms
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   645
      |> apfst (curry Notes Thm.assumptionK)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   646
  | assumes_to_notes e axms = (e, axms);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   647
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   648
fun defines_to_notes thy (Defines defs) defns =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   649
    let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   650
      val defs' = map (fn (_, (def, _)) => def) defs
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   651
      val notes = map (fn (a, (def, _)) =>
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   652
        (a, [([assume (cterm_of thy def)], [])])) defs
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   653
    in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   654
      (Notes (Thm.definitionK, notes), defns @ defs')
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   655
    end
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   656
  | defines_to_notes _ e defns = (e, defns);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   657
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   658
fun gen_add_locale prep_context
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   659
    bname predicate_name raw_imprt raw_body thy =
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   660
  let
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   661
    val thy_ctxt = ProofContext.init thy;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   662
    val name = Sign.full_name thy bname;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   663
    val _ = NewLocale.test_locale thy name andalso
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   664
      error ("Duplicate definition of locale " ^ quote name);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   665
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   666
    val ((fors, deps), (_, body_elems), text as (parms, ((_, exts'), _), defs)) =
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   667
      prep_context raw_imprt raw_body thy_ctxt;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   668
    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
   669
      define_preds predicate_name text thy;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   670
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   671
    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
   672
    val _ = if null extraTs then ()
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   673
      else warning ("Additional type variable(s) in locale specification " ^ quote bname);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   674
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   675
    val satisfy = Element.satisfy_morphism b_axioms;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   676
    val params = fors @
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   677
      (body_elems |> map_filter (fn Fixes fixes => SOME fixes | _ => NONE) |> flat);
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   678
    val (body_elems', defns) = fold_map (defines_to_notes thy) body_elems [];
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   679
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   680
    val notes = body_elems' |>
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   681
      (fn elems => fold_map assumes_to_notes elems (map Element.conclude_witness a_axioms)) |>
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   682
      fst |> map (Element.morph_ctxt satisfy) |>
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   683
      map_filter (fn Notes notes => SOME notes | _ => NONE);
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   684
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   685
    val deps' = map (fn (l, morph) => (l, morph $> satisfy)) deps;
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   686
28795
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   687
    val loc_ctxt = thy' |>
28872
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   688
      NewLocale.register_locale name (extraTs, params)
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   689
        (if is_some b_statement then b_statement else a_statement, map prop_of defs) ([], [])
686963dbf6cd add_locale functional.
ballarin
parents: 28862
diff changeset
   690
        (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
   691
      NewLocale.init name
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   692
  in (name, loc_ctxt) end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   693
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   694
in
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   695
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   696
val add_locale = gen_add_locale read_context;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   697
(* val add_locale_i = gen_add_locale cert_context; *)
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   698
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   699
end;
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   700
6891e273c33b Initial part of locale reimplementation.
ballarin
parents: 28701
diff changeset
   701
end;