src/Pure/Isar/locale.ML
changeset 12263 6f2acf10e2a2
parent 12143 dc42d17c5b53
child 12273 7fb9840d358d
equal deleted inserted replaced
12262:11ff5f47df6e 12263:6f2acf10e2a2
     3     Author:     Markus Wenzel, TU Muenchen
     3     Author:     Markus Wenzel, TU Muenchen
     4     License:    GPL (GNU GENERAL PUBLIC LICENSE)
     4     License:    GPL (GNU GENERAL PUBLIC LICENSE)
     5 
     5 
     6 Locales -- Isar proof contexts as meta-level predicates, with local
     6 Locales -- Isar proof contexts as meta-level predicates, with local
     7 syntax and implicit structures.  Draws basic ideas from Florian
     7 syntax and implicit structures.  Draws basic ideas from Florian
     8 Kammueller's original version of locales, but uses the rich
     8 Kammüller's original version of locales, but uses the rich
     9 infrastructure of Isar instead of the raw meta-logic.
     9 infrastructure of Isar instead of the raw meta-logic.
    10 *)
    10 *)
    11 
    11 
    12 signature BASIC_LOCALE =
    12 signature BASIC_LOCALE =
    13 sig
    13 sig
    17 
    17 
    18 signature LOCALE =
    18 signature LOCALE =
    19 sig
    19 sig
    20   include BASIC_LOCALE
    20   include BASIC_LOCALE
    21   type context
    21   type context
    22   type expression
    22   datatype expression =
       
    23     Locale of string |
       
    24     Merge of expression list |
       
    25     Rename of expression * string option list
    23   datatype ('typ, 'term, 'fact, 'att) elem =
    26   datatype ('typ, 'term, 'fact, 'att) elem =
    24     Fixes of (string * 'typ option * mixfix option) list |
    27     Fixes of (string * 'typ option * mixfix option) list |
    25     Assumes of ((string * 'att list) * ('term * ('term list * 'term list)) list) list |
    28     Assumes of ((string * 'att list) * ('term * ('term list * 'term list)) list) list |
    26     Defines of ((string * 'att list) * ('term * 'term list)) list |
    29     Defines of ((string * 'att list) * ('term * 'term list)) list |
    27     Notes of ((string * 'att list) * ('fact * 'att list) list) list |
    30     Notes of ((string * 'att list) * ('fact * 'att list) list) list |
    50 
    53 
    51 (** locale elements and locales **)
    54 (** locale elements and locales **)
    52 
    55 
    53 type context = ProofContext.context;
    56 type context = ProofContext.context;
    54 
    57 
    55 type expression = string;
    58 datatype expression =
       
    59   Locale of string |
       
    60   Merge of expression list |
       
    61   Rename of expression * string option list;
    56 
    62 
    57 datatype ('typ, 'term, 'fact, 'att) elem =
    63 datatype ('typ, 'term, 'fact, 'att) elem =
    58   Fixes of (string * 'typ option * mixfix option) list |
    64   Fixes of (string * 'typ option * mixfix option) list |
    59   Assumes of ((string * 'att list) * ('term * ('term list * 'term list)) list) list |
    65   Assumes of ((string * 'att list) * ('term * ('term list * 'term list)) list) list |
    60   Defines of ((string * 'att list) * ('term * 'term list)) list |
    66   Defines of ((string * 'att list) * ('term * 'term list)) list |
    63 
    69 
    64 type 'att element = (string, string, string, 'att) elem;
    70 type 'att element = (string, string, string, 'att) elem;
    65 type 'att element_i = (typ, term, thm list, 'att) elem;
    71 type 'att element_i = (typ, term, thm list, 'att) elem;
    66 
    72 
    67 type locale =
    73 type locale =
    68   {imports: expression list, elements: context attribute element_i list, closed: bool};
    74  {imports: string list,
    69 
    75   elements: context attribute element_i list,
    70 fun make_locale imports elements closed =
    76   params: (string * typ) list,
    71   {imports = imports, elements = elements, closed = closed}: locale;
    77   text: ((string * typ) list * term) option,
    72 
    78   closed: bool};
    73 fun close_locale {imports, elements, closed = _} = make_locale imports elements true;
    79 
       
    80 fun make_locale imports elements params text closed =
       
    81  {imports = imports, elements = elements, params = params,
       
    82   text = text, closed = closed}: locale;
    74 
    83 
    75 
    84 
    76 
    85 
    77 (** theory data **)
    86 (** theory data **)
    78 
    87 
    83   val name = "Isar/locales";
    92   val name = "Isar/locales";
    84   type T = NameSpace.T * locale Symtab.table;
    93   type T = NameSpace.T * locale Symtab.table;
    85 
    94 
    86   val empty = (NameSpace.empty, Symtab.empty);
    95   val empty = (NameSpace.empty, Symtab.empty);
    87   val copy = I;
    96   val copy = I;
    88   fun finish (space, locales) = (space, Symtab.map close_locale locales);
    97 
       
    98   fun close {imports, elements, params, text, closed = _} =
       
    99     make_locale imports elements params text true;
       
   100   fun finish (space, locales) = (space, Symtab.map close locales);
       
   101 
    89   val prep_ext = I;
   102   val prep_ext = I;
    90   fun merge ((space1, locales1), (space2, locales2)) =
   103   fun merge ((space1, locales1), (space2, locales2)) =
    91       (NameSpace.merge (space1, space2), Symtab.merge (K true) (locales1, locales2));
   104       (NameSpace.merge (space1, space2), Symtab.merge (K true) (locales1, locales2));
    92 
   105 
    93   fun print _ (space, locales) =
   106   fun print _ (space, locales) =
   119 
   132 
   120 
   133 
   121 
   134 
   122 (** internalize elements **)
   135 (** internalize elements **)
   123 
   136 
   124 (* read_elem *)
   137 (* read *)
       
   138 
       
   139 fun read_expression ctxt (Locale xname) = Locale (intern (ProofContext.sign_of ctxt) xname)
       
   140   | read_expression ctxt (Merge exprs) = Merge (map (read_expression ctxt) exprs)
       
   141   | read_expression ctxt (Rename (expr, xs)) = Rename (read_expression ctxt expr, xs);
   125 
   142 
   126 fun read_elem ctxt =
   143 fun read_elem ctxt =
   127  fn Fixes fixes =>
   144  fn Fixes fixes =>
   128       let val vars =
   145       let val vars =
   129         #2 (foldl_map ProofContext.read_vars (ctxt, map (fn (x, T, _) => ([x], T)) fixes))
   146         #2 (foldl_map ProofContext.read_vars (ctxt, map (fn (x, T, _) => ([x], T)) fixes))
   134       let val propps =
   151       let val propps =
   135         #2 (ProofContext.read_propp (ctxt, map (fn (_, (t, ps)) => [(t, (ps, []))]) defs))
   152         #2 (ProofContext.read_propp (ctxt, map (fn (_, (t, ps)) => [(t, (ps, []))]) defs))
   136       in Defines (map #1 defs ~~ map (fn [(t', (ps', []))] => (t', ps')) propps) end
   153       in Defines (map #1 defs ~~ map (fn [(t', (ps', []))] => (t', ps')) propps) end
   137   | Notes facts =>
   154   | Notes facts =>
   138       Notes (map (apsnd (map (apfst (ProofContext.get_thms ctxt)))) facts)
   155       Notes (map (apsnd (map (apfst (ProofContext.get_thms ctxt)))) facts)
   139   | Uses xname => Uses (intern (ProofContext.sign_of ctxt) xname);
   156   | Uses expr => Uses (read_expression ctxt expr);
   140 
   157 
   141 
   158 
   142 (* prepare attributes *)
   159 (* prepare attributes *)
   143 
   160 
   144 local fun int_att attrib (x, srcs) = (x, map attrib srcs) in
   161 local fun int_att attrib (x, srcs) = (x, map attrib srcs) in
   146 fun attribute _ (Fixes fixes) = Fixes fixes
   163 fun attribute _ (Fixes fixes) = Fixes fixes
   147   | attribute attrib (Assumes asms) = Assumes (map (apfst (int_att attrib)) asms)
   164   | attribute attrib (Assumes asms) = Assumes (map (apfst (int_att attrib)) asms)
   148   | attribute attrib (Defines defs) = Defines (map (apfst (int_att attrib)) defs)
   165   | attribute attrib (Defines defs) = Defines (map (apfst (int_att attrib)) defs)
   149   | attribute attrib (Notes facts) =
   166   | attribute attrib (Notes facts) =
   150       Notes (map (apfst (int_att attrib) o apsnd (map (int_att attrib))) facts)
   167       Notes (map (apfst (int_att attrib) o apsnd (map (int_att attrib))) facts)
   151   | attribute _ (Uses name) = Uses name;
   168   | attribute _ (Uses expr) = Uses expr;
   152 
   169 
   153 end;
   170 end;
       
   171 
       
   172 
       
   173 (** renaming **)
       
   174 
       
   175 fun rename ren x = if_none (assoc_string (ren, x)) x;
       
   176 
       
   177 fun rename_term ren (Free (x, T)) = Free (rename ren x, T)
       
   178   | rename_term ren (t $ u) = rename_term ren t $ rename_term ren u
       
   179   | rename_term ren (Abs (x, T, t)) = Abs (x, T, rename_term ren t)
       
   180   | rename_term _ a = a;
       
   181 
       
   182 fun rename_thm ren th =
       
   183   let
       
   184     val {sign, hyps, prop, maxidx, ...} = Thm.rep_thm th;
       
   185     val cert = Thm.cterm_of sign;
       
   186     val (xs, Ts) = Library.split_list (foldl Drule.add_frees ([], prop :: hyps));
       
   187     val xs' = map (rename ren) xs;
       
   188     fun cert_frees names = map (cert o Free) (names ~~ Ts);
       
   189     fun cert_vars names = map (cert o Var o apfst (rpair (maxidx + 1))) (names ~~ Ts);
       
   190   in
       
   191     if xs = xs' then th
       
   192     else
       
   193       th
       
   194       |> Drule.implies_intr_list (map cert hyps)
       
   195       |> Drule.forall_intr_list (cert_frees xs)
       
   196       |> Drule.forall_elim_list (cert_vars xs)
       
   197       |> Thm.instantiate ([], cert_vars xs ~~ cert_frees xs')
       
   198       |> (fn th' => Drule.implies_elim_list th' (map (Thm.assume o cert o rename_term ren) hyps))
       
   199   end;
       
   200 
       
   201 
       
   202 fun rename_elem ren (Fixes fixes) = Fixes (map (fn (x, T, mx) =>
       
   203       (rename ren x, T, if mx = None then mx else Some Syntax.NoSyn)) fixes)
       
   204   | rename_elem ren (Assumes asms) = Assumes (map (apsnd (map (fn (t, (ps, qs)) =>
       
   205       (rename_term ren t, (map (rename_term ren) ps, map (rename_term ren) qs))))) asms)
       
   206   | rename_elem ren (Defines defs) = Defines (map (apsnd (fn (t, ps) =>
       
   207       (rename_term ren t, map (rename_term ren) ps))) defs)
       
   208   | rename_elem ren (Notes facts) = Notes (map (apsnd (map (apfst (map (rename_thm ren))))) facts)
       
   209   | rename_elem ren (Uses expr) = sys_error "FIXME";
   154 
   210 
   155 
   211 
   156 
   212 
   157 (** activate locales **)
   213 (** activate locales **)
   158 
   214 
   165   | activate (ctxt, Defines defs) = #1 (ProofContext.assume_i ProofContext.export_def
   221   | activate (ctxt, Defines defs) = #1 (ProofContext.assume_i ProofContext.export_def
   166       (map (fn ((name, atts), (t, ps)) =>
   222       (map (fn ((name, atts), (t, ps)) =>
   167         let val (c, t') = ProofContext.cert_def ctxt t
   223         let val (c, t') = ProofContext.cert_def ctxt t
   168         in ((if name = "" then Thm.def_name c else name, atts), [(t', (ps, []))]) end) defs) ctxt)
   224         in ((if name = "" then Thm.def_name c else name, atts), [(t', (ps, []))]) end) defs) ctxt)
   169   | activate (ctxt, Notes facts) = #1 (ProofContext.have_thmss facts ctxt)
   225   | activate (ctxt, Notes facts) = #1 (ProofContext.have_thmss facts ctxt)
   170   | activate (ctxt, Uses name) = activate_locale_i name ctxt
   226   | activate (ctxt, Uses expr) = activate_expression [] (ctxt, expr)
   171 
   227 
   172 and activate_elements_i elems ctxt = foldl activate (ctxt, elems)
   228 and activate_elements_i elems ctxt = foldl activate (ctxt, elems)
   173 
   229 
   174 and activate_locale_elements (ctxt, name) =
   230 and activate_expression rs (ctxt, Locale name) = activate_loc rs name ctxt
       
   231   | activate_expression rs (ctxt, Merge exprs) = foldl (activate_expression rs) (ctxt, exprs)
       
   232   | activate_expression rs (ctxt, Rename (expr, xs)) = activate_expression (xs :: rs) (ctxt, expr)
       
   233 
       
   234 and activate_locale_elements rs (ctxt, name) =
   175   let
   235   let
   176     val thy = ProofContext.theory_of ctxt;
   236     val thy = ProofContext.theory_of ctxt;
   177     val {elements, ...} = the_locale thy name;    (*exception ERROR*)
   237     val {elements, params, ...} = the_locale thy name;    (*exception ERROR*)
       
   238     val param_names = map #1 params;
       
   239 
       
   240     fun replace (opt_x :: xs, y :: ys) = if_none opt_x y :: replace (xs, ys)
       
   241       | replace ([], ys) = ys
       
   242       | replace (_ :: _, []) = raise ProofContext.CONTEXT
       
   243           ("Too many parameters in renaming of locale " ^ quote name, ctxt);
       
   244 
       
   245     val elements' =
       
   246       if null rs then elements
       
   247       else map (rename_elem (param_names ~~ foldr replace (rs, param_names))) elements;
   178   in
   248   in
   179     activate_elements_i elements ctxt handle ProofContext.CONTEXT (msg, c) =>
   249     activate_elements_i elements' ctxt handle ProofContext.CONTEXT (msg, c) =>
   180       raise ProofContext.CONTEXT (msg ^ "\nThe error(s) above occurred in locale " ^
   250       raise ProofContext.CONTEXT (msg ^ "\nThe error(s) above occurred in locale " ^
   181         quote (cond_extern (Theory.sign_of thy) name), c)
   251         quote (cond_extern (Theory.sign_of thy) name), c)
   182   end
   252   end
   183 
   253 
   184 and activate_locale_i name ctxt =
   254 and activate_loc rs name ctxt =
   185   activate_locale_elements (foldl activate_locale_elements
   255   activate_locale_elements rs (foldl (activate_locale_elements rs)
   186     (ctxt, #imports (the_locale (ProofContext.theory_of ctxt) name)), name);
   256     (ctxt, #imports (the_locale (ProofContext.theory_of ctxt) name)), name);
   187 
   257 
   188 
   258 
   189 fun activate_elements elems ctxt =
   259 fun activate_elements elems ctxt =
   190   foldl ((fn (ctxt, elem) => activate (ctxt, read_elem ctxt elem))) (ctxt, elems);
   260   foldl ((fn (ctxt, elem) => activate (ctxt, read_elem ctxt elem))) (ctxt, elems);
   191 
   261 
       
   262 val activate_locale_i = activate_loc [];
       
   263 
   192 fun activate_locale xname ctxt =
   264 fun activate_locale xname ctxt =
   193   activate_locale_i (intern (ProofContext.sign_of ctxt) xname) ctxt;
   265   activate_locale_i (intern (ProofContext.sign_of ctxt) xname) ctxt;
   194 
   266 
   195 
   267 
   196 
   268 
   198 
   270 
   199 fun pretty_locale thy xname =
   271 fun pretty_locale thy xname =
   200   let
   272   let
   201     val sg = Theory.sign_of thy;
   273     val sg = Theory.sign_of thy;
   202     val name = intern sg xname;
   274     val name = intern sg xname;
   203     val {imports, elements, closed = _} = the_locale thy name;
   275     val {imports, elements, params = _, text = _, closed = _} = the_locale thy name;
   204     val locale_ctxt = ProofContext.init thy |> activate_locale_i name;
   276     val locale_ctxt = ProofContext.init thy |> activate_locale_i name;
   205 
   277 
   206     val prt_typ = Pretty.quote o ProofContext.pretty_typ locale_ctxt;
   278     val prt_typ = Pretty.quote o ProofContext.pretty_typ locale_ctxt;
   207     val prt_term = Pretty.quote o ProofContext.pretty_term locale_ctxt;
   279     val prt_term = Pretty.quote o ProofContext.pretty_term locale_ctxt;
   208     val prt_thm = Pretty.quote o ProofContext.pretty_thm locale_ctxt;
   280     val prt_thm = Pretty.quote o ProofContext.pretty_thm locale_ctxt;
   225 
   297 
   226     fun prt_fact (("", _), ths) = Pretty.block (Pretty.breaks (map prt_thm (flat (map fst ths))))
   298     fun prt_fact (("", _), ths) = Pretty.block (Pretty.breaks (map prt_thm (flat (map fst ths))))
   227       | prt_fact ((a, _), ths) = Pretty.block
   299       | prt_fact ((a, _), ths) = Pretty.block
   228           (Pretty.breaks (Pretty.str (a ^ ":") :: map prt_thm (flat (map fst ths))));
   300           (Pretty.breaks (Pretty.str (a ^ ":") :: map prt_thm (flat (map fst ths))));
   229 
   301 
       
   302     fun prt_expr (Locale name) = Pretty.str (cond_extern sg name)
       
   303       | prt_expr (Merge exprs) = Pretty.enclose "(" ")"
       
   304           (flat (separate [Pretty.str " +", Pretty.brk 1]
       
   305             (map (single o prt_expr) exprs)))
       
   306       | prt_expr (Rename (expr, xs)) = Pretty.enclose "(" ")"
       
   307           (Pretty.breaks (prt_expr expr :: map (fn x => Pretty.str (if_none x "_")) xs));
       
   308 
   230     fun prt_elem (Fixes fixes) = Pretty.big_list "fixes" (map prt_fix fixes)
   309     fun prt_elem (Fixes fixes) = Pretty.big_list "fixes" (map prt_fix fixes)
   231       | prt_elem (Assumes asms) = Pretty.big_list "assumes" (map prt_asm asms)
   310       | prt_elem (Assumes asms) = Pretty.big_list "assumes" (map prt_asm asms)
   232       | prt_elem (Defines defs) = Pretty.big_list "defines" (map prt_def defs)
   311       | prt_elem (Defines defs) = Pretty.big_list "defines" (map prt_def defs)
   233       | prt_elem (Notes facts) = Pretty.big_list "notes" (map prt_fact facts)
   312       | prt_elem (Notes facts) = Pretty.big_list "notes" (map prt_fact facts)
   234       | prt_elem (Uses name) = Pretty.str ("uses " ^ cond_extern sg name);
   313       | prt_elem (Uses expr) = Pretty.block [Pretty.str "uses", Pretty.brk 1, prt_expr expr];
   235 
   314 
   236     val prt_header = Pretty.block (Pretty.str ("locale " ^ cond_extern sg name ^ " =") ::
   315     val prt_header = Pretty.block (Pretty.str ("locale " ^ cond_extern sg name ^ " =") ::
   237        (if null imports then [] else
   316        (if null imports then [] else
   238        (Pretty.str " " :: flat (separate [Pretty.str " +", Pretty.brk 1]
   317        (Pretty.str " " :: flat (separate [Pretty.str " +", Pretty.brk 1]
   239            (map (single o Pretty.str o cond_extern sg) imports)) @ [Pretty.str " +"])));
   318            (map (single o Pretty.str o cond_extern sg) imports)) @ [Pretty.str " +"])));
   271     val _ =
   350     val _ =
   272       if is_none (get_locale thy name) then () else
   351       if is_none (get_locale thy name) then () else
   273       error ("Duplicate definition of locale " ^ quote name);
   352       error ("Duplicate definition of locale " ^ quote name);
   274 
   353 
   275     val imports = map (prep_locale sign) raw_imports;
   354     val imports = map (prep_locale sign) raw_imports;
   276     val imports_ctxt = foldl activate_locale_elements (ProofContext.init thy, imports);
   355     val imports_ctxt = foldl (activate_locale_elements []) (ProofContext.init thy, imports);
   277     fun prep (ctxt, raw_elem) =
   356     fun prep (ctxt, raw_elem) =
   278       let val elem = closeup ctxt (prep_elem ctxt raw_elem)
   357       let val elem = closeup ctxt (prep_elem ctxt raw_elem)
   279       in (activate (ctxt, elem), elem) end;
   358       in (activate (ctxt, elem), elem) end;
   280     val (locale_ctxt, elems) = foldl_map prep (imports_ctxt, raw_elems);
   359     val (locale_ctxt, elems) = foldl_map prep (imports_ctxt, raw_elems);
       
   360     val params = [];  (* FIXME *)
       
   361     val text = None;  (* FIXME *)
   281   in
   362   in
   282     thy
   363     thy
   283     |> declare_locale name
   364     |> declare_locale name
   284     |> put_locale name (make_locale imports elems false)
   365     |> put_locale name (make_locale imports elems params text false)
   285   end;
   366   end;
   286 
   367 
   287 val add_locale = gen_add_locale intern read_elem;
   368 val add_locale = gen_add_locale intern read_elem;
   288 val add_locale_i = gen_add_locale (K I) (K I);
   369 val add_locale_i = gen_add_locale (K I) (K I);
   289 
   370 
   291 
   372 
   292 (** store results **)
   373 (** store results **)
   293 
   374 
   294 fun add_thmss name args thy =
   375 fun add_thmss name args thy =
   295   let
   376   let
   296     val {imports, elements, closed} = the_locale thy name;
   377     val {imports, elements, params, text, closed} = the_locale thy name;
   297     val _ = conditional closed (fn () =>
   378     val _ = conditional closed (fn () =>
   298       error ("Cannot store results in closed locale: " ^ quote name));
   379       error ("Cannot store results in closed locale: " ^ quote name));
   299     val note = Notes (map (fn ((a, ths), atts) =>
   380     val note = Notes (map (fn ((a, ths), atts) =>
   300       ((a, atts), [(map (curry Thm.name_thm a) ths, [])])) args);
   381       ((a, atts), [(map (curry Thm.name_thm a) ths, [])])) args);
   301   in
   382   in
   302     activate (thy |> ProofContext.init |> activate_locale_i name, note);    (*test attributes!*)
   383     activate (thy |> ProofContext.init |> activate_locale_i name, note);    (*test attributes!*)
   303     thy |> put_locale name (make_locale imports (elements @ [note]) closed)
   384     thy |> put_locale name (make_locale imports (elements @ [note]) params text closed)
   304   end;
   385   end;
   305 
   386 
   306 
   387 
   307 
   388 
   308 (** locale theory setup **)
   389 (** locale theory setup **)