src/HOL/Statespace/state_space.ML
author haftmann
Fri, 19 Jun 2009 17:23:21 +0200
changeset 31723 f5cafe803b55
parent 30515 4120fc59dd85
child 32194 966e166d039d
permissions -rw-r--r--
discontinued ancient tradition to suffix certain ML module names with "_package"
     1 (*  Title:      HOL/Statespace/state_space.ML
     2     Author:     Norbert Schirmer, TU Muenchen
     3 *)
     4 
     5 signature STATE_SPACE =
     6   sig
     7     val KN : string
     8     val distinct_compsN : string
     9     val getN : string
    10     val putN : string
    11     val injectN : string
    12     val namespaceN : string
    13     val projectN : string
    14     val valuetypesN : string
    15 
    16     val namespace_definition :
    17        bstring ->
    18        typ ->
    19        Expression.expression ->
    20        string list -> string list -> theory -> theory
    21 
    22     val define_statespace :
    23        string list ->
    24        string ->
    25        (string list * bstring * (string * string) list) list ->
    26        (string * string) list -> theory -> theory
    27     val define_statespace_i :
    28        string option ->
    29        string list ->
    30        string ->
    31        (typ list * bstring * (string * string) list) list ->
    32        (string * typ) list -> theory -> theory
    33 
    34     val statespace_decl :
    35        OuterParse.token list ->
    36        ((string list * bstring) *
    37          ((string list * xstring * (bstring * bstring) list) list *
    38           (bstring * string) list)) * OuterParse.token list
    39 
    40 
    41     val neq_x_y : Proof.context -> term -> term -> thm option
    42     val distinctNameSolver : Simplifier.solver
    43     val distinctTree_tac :
    44        Proof.context -> term * int -> Tactical.tactic
    45     val distinct_simproc : Simplifier.simproc
    46 
    47 
    48     val get_comp : Context.generic -> string -> (typ * string) Option.option
    49     val get_silent : Context.generic -> bool
    50     val set_silent : bool -> Context.generic -> Context.generic
    51 
    52     val gen_lookup_tr : Proof.context -> term -> string -> term
    53     val lookup_swap_tr : Proof.context -> term list -> term
    54     val lookup_tr : Proof.context -> term list -> term
    55     val lookup_tr' : Proof.context -> term list -> term
    56 
    57     val gen_update_tr :
    58        bool -> Proof.context -> string -> term -> term -> term
    59     val update_tr : Proof.context -> term list -> term
    60     val update_tr' : Proof.context -> term list -> term
    61   end;
    62 
    63 structure StateSpace : STATE_SPACE =
    64 struct
    65 
    66 (* Theorems *)
    67 
    68 (* Names *)
    69 val distinct_compsN = "distinct_names"
    70 val namespaceN = "_namespace"
    71 val valuetypesN = "_valuetypes"
    72 val projectN = "project"
    73 val injectN = "inject"
    74 val getN = "get"
    75 val putN = "put"
    76 val project_injectL = "StateSpaceLocale.project_inject";
    77 val KN = "StateFun.K_statefun"
    78 
    79 
    80 (* Library *)
    81 
    82 fun fold1 f xs = fold f (tl xs) (hd xs)
    83 fun fold1' f [] x = x
    84   | fold1' f xs _ = fold1 f xs
    85 
    86 fun sublist_idx eq xs ys =
    87   let
    88     fun sublist n xs ys =
    89          if is_prefix eq xs ys then SOME n
    90          else (case ys of [] => NONE
    91                | (y::ys') => sublist (n+1) xs ys')
    92   in sublist 0 xs ys end;
    93 
    94 fun is_sublist eq xs ys = is_some (sublist_idx eq xs ys);
    95 
    96 fun sorted_subset eq [] ys = true
    97   | sorted_subset eq (x::xs) [] = false
    98   | sorted_subset eq (x::xs) (y::ys) = if eq (x,y) then sorted_subset eq xs ys
    99                                        else sorted_subset eq (x::xs) ys;
   100 
   101 
   102 
   103 type namespace_info =
   104  {declinfo: (typ*string) Termtab.table, (* type, name of statespace *)
   105   distinctthm: thm Symtab.table,
   106   silent: bool
   107  };
   108 
   109 structure NameSpaceArgs =
   110 struct
   111   type T = namespace_info;
   112   val empty = {declinfo = Termtab.empty, distinctthm = Symtab.empty, silent = false};
   113   val extend = I;
   114   fun merge pp ({declinfo=declinfo1, distinctthm=distinctthm1, silent=silent1},
   115                 {declinfo=declinfo2, distinctthm=distinctthm2, silent=silent2}) =
   116                 {declinfo = Termtab.merge (K true) (declinfo1, declinfo2),
   117                  distinctthm = Symtab.merge (K true) (distinctthm1, distinctthm2),
   118                  silent = silent1 andalso silent2}
   119 end;
   120 
   121 structure NameSpaceData = GenericDataFun(NameSpaceArgs);
   122 
   123 fun make_namespace_data declinfo distinctthm silent =
   124      {declinfo=declinfo,distinctthm=distinctthm,silent=silent};
   125 
   126 
   127 fun delete_declinfo n ctxt =
   128   let val {declinfo,distinctthm,silent} = NameSpaceData.get ctxt;
   129   in NameSpaceData.put
   130        (make_namespace_data (Termtab.delete_safe n declinfo) distinctthm silent) ctxt
   131   end;
   132 
   133 
   134 fun update_declinfo (n,v) ctxt =
   135   let val {declinfo,distinctthm,silent} = NameSpaceData.get ctxt;
   136   in NameSpaceData.put
   137       (make_namespace_data (Termtab.update (n,v) declinfo) distinctthm silent) ctxt
   138   end;
   139 
   140 fun set_silent silent ctxt =
   141   let val {declinfo,distinctthm,...} = NameSpaceData.get ctxt;
   142   in NameSpaceData.put
   143       (make_namespace_data declinfo distinctthm silent) ctxt
   144   end;
   145 
   146 val get_silent = #silent o NameSpaceData.get;
   147 
   148 fun prove_interpretation_in ctxt_tac (name, expr) thy =
   149    thy
   150    |> Expression.sublocale_cmd name expr
   151    |> Proof.global_terminal_proof
   152          (Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt),Position.none), NONE)
   153    |> ProofContext.theory_of
   154 
   155 fun add_locale name expr elems thy =
   156   thy 
   157   |> Expression.add_locale (Binding.name name) (Binding.name name) expr elems
   158   |> snd
   159   |> LocalTheory.exit;
   160 
   161 fun add_locale_cmd name expr elems thy =
   162   thy 
   163   |> Expression.add_locale_cmd (Binding.name name) Binding.empty expr elems
   164   |> snd
   165   |> LocalTheory.exit;
   166 
   167 type statespace_info =
   168  {args: (string * sort) list, (* type arguments *)
   169   parents: (typ list * string * string option list) list,
   170              (* type instantiation, state-space name, component renamings *)
   171   components: (string * typ) list,
   172   types: typ list (* range types of state space *)
   173  };
   174 
   175 structure StateSpaceArgs =
   176 struct
   177   val name = "HOL/StateSpace";
   178   type T = statespace_info Symtab.table;
   179   val empty = Symtab.empty;
   180   val extend = I;
   181 
   182   fun merge pp (nt1,nt2) = Symtab.merge (K true) (nt1, nt2);
   183 end;
   184 
   185 structure StateSpaceData = GenericDataFun(StateSpaceArgs);
   186 
   187 fun add_statespace name args parents components types ctxt =
   188      StateSpaceData.put
   189       (Symtab.update_new (name, {args=args,parents=parents,
   190                                 components=components,types=types}) (StateSpaceData.get ctxt))
   191       ctxt;
   192 
   193 fun get_statespace ctxt name =
   194       Symtab.lookup (StateSpaceData.get ctxt) name;
   195 
   196 
   197 fun lookupI eq xs n =
   198   (case AList.lookup eq xs n of
   199      SOME v => v
   200    | NONE => n);
   201 
   202 fun mk_free ctxt name =
   203   if Variable.is_fixed ctxt name orelse Variable.is_declared ctxt name
   204   then
   205     let val n' = lookupI (op =) (Variable.fixes_of ctxt) name
   206     in SOME (Free (n',ProofContext.infer_type ctxt n')) end
   207   else NONE
   208 
   209 
   210 fun get_dist_thm ctxt name = Symtab.lookup (#distinctthm (NameSpaceData.get ctxt)) name;
   211 fun get_comp ctxt name =
   212      Option.mapPartial
   213        (Termtab.lookup (#declinfo (NameSpaceData.get ctxt)))
   214        (mk_free (Context.proof_of ctxt) name);
   215 
   216 
   217 (*** Tactics ***)
   218 
   219 fun neq_x_y ctxt x y =
   220   (let
   221     val dist_thm = the (get_dist_thm (Context.Proof ctxt) (#1 (dest_Free x)));
   222     val ctree = cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
   223     val tree = term_of ctree;
   224     val x_path = the (DistinctTreeProver.find_tree x tree);
   225     val y_path = the (DistinctTreeProver.find_tree y tree);
   226     val thm = DistinctTreeProver.distinctTreeProver dist_thm x_path y_path;
   227   in SOME thm
   228   end handle Option => NONE)
   229 
   230 fun distinctTree_tac ctxt
   231       (Const ("Trueprop",_) $
   232         (Const ("Not", _) $ (Const ("op =", _) $ (x as Free _)$ (y as Free _))), i) =
   233   (case (neq_x_y ctxt x y) of
   234      SOME neq => rtac neq i
   235    | NONE => no_tac)
   236   | distinctTree_tac _ _ = no_tac;
   237 
   238 val distinctNameSolver = mk_solver' "distinctNameSolver"
   239      (fn ss => case try Simplifier.the_context ss of
   240                  SOME ctxt => SUBGOAL (distinctTree_tac ctxt)
   241                 | NONE => fn i => no_tac)
   242 
   243 val distinct_simproc =
   244   Simplifier.simproc @{theory HOL} "StateSpace.distinct_simproc" ["x = y"]
   245     (fn thy => fn ss => (fn (Const ("op =",_)$(x as Free _)$(y as Free _)) =>
   246         (case try Simplifier.the_context ss of
   247           SOME ctxt => Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq])
   248                        (neq_x_y ctxt x y)
   249         | NONE => NONE)
   250       | _ => NONE))
   251 
   252 local
   253   val ss = HOL_basic_ss
   254 in
   255 fun interprete_parent name dist_thm_name parent_expr thy =
   256   let
   257 
   258     fun solve_tac ctxt (_,i) st =
   259       let
   260         val distinct_thm = ProofContext.get_thm ctxt dist_thm_name;
   261         val goal = List.nth (cprems_of st,i-1);
   262         val rule = DistinctTreeProver.distinct_implProver distinct_thm goal;
   263       in EVERY [rtac rule i] st
   264       end
   265 
   266     fun tac ctxt = EVERY [Locale.intro_locales_tac true ctxt [],
   267                           ALLGOALS (SUBGOAL (solve_tac ctxt))]
   268 
   269   in thy
   270      |> prove_interpretation_in tac (name,parent_expr)
   271   end;
   272 
   273 end;
   274 
   275 fun namespace_definition name nameT parent_expr parent_comps new_comps thy =
   276   let
   277     val all_comps = parent_comps @ new_comps;
   278     val vars = (map (fn n => (Binding.name n, NONE, NoSyn)) all_comps);
   279     val full_name = Sign.full_bname thy name;
   280     val dist_thm_name = distinct_compsN;
   281 
   282     val dist_thm_full_name = dist_thm_name;
   283     fun comps_of_thm thm = prop_of thm
   284              |> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free);
   285 
   286     fun type_attr phi (ctxt,thm) =
   287       (case ctxt of Context.Theory _ => (ctxt,thm)
   288        | _ =>
   289         let
   290           val {declinfo,distinctthm=tt,silent} = (NameSpaceData.get ctxt);
   291 	  val all_names = comps_of_thm thm;
   292           fun upd name tt =
   293                (case (Symtab.lookup tt name) of
   294                  SOME dthm => if sorted_subset (op =) (comps_of_thm dthm) all_names
   295                               then Symtab.update (name,thm) tt else tt
   296                 | NONE => Symtab.update (name,thm) tt)
   297 
   298           val tt' = tt |> fold upd all_names;
   299           val activate_simproc =
   300               Output.no_warnings
   301                (Simplifier.map_ss (fn ss => ss addsimprocs [distinct_simproc]));
   302           val ctxt' =
   303               ctxt
   304               |> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent}
   305               |> activate_simproc
   306         in (ctxt',thm)
   307         end)
   308 
   309     val attr = Attrib.internal type_attr;
   310 
   311     val assumes = Element.Assumes [((Binding.name dist_thm_name,[attr]),
   312                     [(HOLogic.Trueprop $
   313                       (Const ("DistinctTreeProver.all_distinct",
   314                                  Type ("DistinctTreeProver.tree",[nameT]) --> HOLogic.boolT) $
   315                       DistinctTreeProver.mk_tree (fn n => Free (n,nameT)) nameT
   316                                 (sort fast_string_ord all_comps)),
   317                       ([]))])];
   318   in thy
   319      |> add_locale name ([],vars) [assumes]
   320      |> ProofContext.theory_of
   321      |> interprete_parent name dist_thm_full_name parent_expr
   322   end;
   323 
   324 fun encode_dot x = if x= #"." then #"_" else x;
   325 
   326 fun encode_type (TFree (s, _)) = s
   327   | encode_type (TVar ((s,i),_)) = "?" ^ s ^ string_of_int i
   328   | encode_type (Type (n,Ts)) =
   329       let
   330         val Ts' = fold1' (fn x => fn y => x ^ "_" ^ y) (map encode_type Ts) "";
   331         val n' = String.map encode_dot n;
   332       in if Ts'="" then n' else Ts' ^ "_" ^ n' end;
   333 
   334 fun project_name T = projectN ^"_"^encode_type T;
   335 fun inject_name T = injectN ^"_"^encode_type T;
   336 
   337 fun project_free T pT V = Free (project_name T, V --> pT);
   338 fun inject_free T pT V = Free (inject_name T, pT --> V);
   339 
   340 fun get_name n = getN ^ "_" ^ n;
   341 fun put_name n = putN ^ "_" ^ n;
   342 fun get_const n T nT V = Free (get_name n, (nT --> V) --> T);
   343 fun put_const n T nT V = Free (put_name n, T --> (nT --> V) --> (nT --> V));
   344 
   345 fun lookup_const T nT V = Const ("StateFun.lookup",(V --> T) --> nT --> (nT --> V) --> T);
   346 fun update_const T nT V =
   347  Const ("StateFun.update",
   348           (V --> T) --> (T --> V) --> nT --> (T --> T) --> (nT --> V) --> (nT --> V));
   349 
   350 fun K_const T = Const ("StateFun.K_statefun",T --> T --> T);
   351 
   352 val no_syn = #3 (Syntax.no_syn ((),()));
   353 
   354 
   355 fun add_declaration name decl thy =
   356   thy
   357   |> TheoryTarget.init name
   358   |> (fn lthy => LocalTheory.declaration (decl lthy) lthy)
   359   |> LocalTheory.exit_global;
   360 
   361 fun parent_components thy (Ts, pname, renaming) =
   362   let
   363     val ctxt = Context.Theory thy;
   364     fun rename [] xs = xs
   365       | rename (NONE::rs)  (x::xs) = x::rename rs xs
   366       | rename (SOME r::rs) ((x,T)::xs) = (r,T)::rename rs xs;
   367     val {args,parents,components,...} =
   368 	      the (Symtab.lookup (StateSpaceData.get ctxt) pname);
   369     val inst = map fst args ~~ Ts;
   370     val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);
   371     val parent_comps =
   372      List.concat (map (fn (Ts',n,rs) => parent_components thy (map subst Ts',n,rs)) parents);
   373     val all_comps = rename renaming (parent_comps @ map (apsnd subst) components);
   374   in all_comps end;
   375 
   376 fun take_upto i xs = List.take(xs,i) handle Subscript => xs;
   377 
   378 fun statespace_definition state_type args name parents parent_comps components thy =
   379   let
   380     val full_name = Sign.full_bname thy name;
   381     val all_comps = parent_comps @ components;
   382 
   383     val components' = map (fn (n,T) => (n,(T,full_name))) components;
   384     val all_comps' = map (fn (n,T) => (n,(T,full_name))) all_comps;
   385 
   386     fun parent_expr (_,n,rs) = (suffix namespaceN n,((n,false),Expression.Positional rs));
   387         (* FIXME: a more specific renaming-prefix (including parameter names) may be nicer *)
   388     val parents_expr = map parent_expr parents;
   389     fun distinct_types Ts =
   390       let val tab = fold (fn T => fn tab => Typtab.update (T,()) tab) Ts Typtab.empty;
   391       in map fst (Typtab.dest tab) end;
   392 
   393     val Ts = distinct_types (map snd all_comps);
   394     val arg_names = map fst args;
   395     val valueN = Name.variant arg_names "'value";
   396     val nameN = Name.variant (valueN::arg_names) "'name";
   397     val valueT = TFree (valueN, Sign.defaultS thy);
   398     val nameT = TFree (nameN, Sign.defaultS thy);
   399     val stateT = nameT --> valueT;
   400     fun projectT T = valueT --> T;
   401     fun injectT T = T --> valueT;
   402     val locinsts = map (fn T => (project_injectL,
   403                     (("",false),Expression.Positional 
   404                              [SOME (Free (project_name T,projectT T)), 
   405                               SOME (Free ((inject_name T,injectT T)))]))) Ts;
   406     val locs = flat (map (fn T => [(Binding.name (project_name T),NONE,NoSyn),
   407                                      (Binding.name (inject_name T),NONE,NoSyn)]) Ts);
   408     val constrains = List.concat
   409          (map (fn T => [(project_name T,projectT T),(inject_name T,injectT T)]) Ts);
   410 
   411     fun interprete_parent_valuetypes (Ts, pname, _) thy =
   412       let
   413         val {args,types,...} =
   414              the (Symtab.lookup (StateSpaceData.get (Context.Theory thy)) pname);
   415         val inst = map fst args ~~ Ts;
   416         val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);
   417         val pars = List.concat (map ((fn T => [project_name T,inject_name T]) o subst) types);
   418 
   419         val expr = ([(suffix valuetypesN name, 
   420                      (("",false),Expression.Positional (map SOME pars)))],[]);
   421       in
   422         prove_interpretation_in (ALLGOALS o solve_tac o Assumption.all_prems_of)
   423           (suffix valuetypesN name, expr) thy
   424       end;
   425 
   426     fun interprete_parent (_, pname, rs) =
   427       let
   428         val expr = ([(pname, (("",false), Expression.Positional rs))],[])
   429       in prove_interpretation_in
   430            (fn ctxt => Locale.intro_locales_tac false ctxt [])
   431            (full_name, expr) end;
   432 
   433     fun declare_declinfo updates lthy phi ctxt =
   434       let
   435         fun upd_prf ctxt =
   436           let
   437             fun upd (n,v) =
   438               let
   439                 val nT = ProofContext.infer_type (LocalTheory.target_of lthy) n
   440               in Context.proof_map
   441                   (update_declinfo (Morphism.term phi (Free (n,nT)),v))
   442               end;
   443           in ctxt |> fold upd updates end;
   444 
   445       in Context.mapping I upd_prf ctxt end;
   446 
   447    fun string_of_typ T =
   448       setmp show_sorts true
   449        (PrintMode.setmp [] (Syntax.string_of_typ (ProofContext.init thy))) T;
   450    val fixestate = (case state_type of
   451          NONE => []
   452        | SOME s =>
   453           let
   454 	    val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)];
   455             val cs = Element.Constrains
   456                        (map (fn (n,T) =>  (n,string_of_typ T))
   457                          ((map (fn (n,_) => (n,nameT)) all_comps) @
   458                           constrains))
   459           in [fx,cs] end
   460        )
   461 
   462 
   463   in thy
   464      |> namespace_definition
   465            (suffix namespaceN name) nameT (parents_expr,[])
   466            (map fst parent_comps) (map fst components)
   467      |> Context.theory_map (add_statespace full_name args parents components [])
   468      |> add_locale (suffix valuetypesN name) (locinsts,locs) []
   469      |> ProofContext.theory_of 
   470      |> fold interprete_parent_valuetypes parents
   471      |> add_locale_cmd name
   472               ([(suffix namespaceN full_name ,(("",false),Expression.Named [])),
   473                 (suffix valuetypesN full_name,(("",false),Expression.Named  []))],[]) fixestate
   474      |> ProofContext.theory_of 
   475      |> fold interprete_parent parents
   476      |> add_declaration (SOME full_name) (declare_declinfo components')
   477   end;
   478 
   479 
   480 (* prepare arguments *)
   481 
   482 fun read_raw_parent ctxt raw_T =
   483   (case ProofContext.read_typ_abbrev ctxt raw_T of
   484     Type (name, Ts) => (Ts, name)
   485   | T => error ("Bad parent statespace specification: " ^ Syntax.string_of_typ ctxt T));
   486 
   487 fun gen_define_statespace prep_typ state_space args name parents comps thy =
   488   let (* - args distinct
   489          - only args may occur in comps and parent-instantiations
   490          - number of insts must match parent args
   491          - no duplicate renamings
   492          - renaming should occur in namespace
   493       *)
   494     val _ = writeln ("Defining statespace " ^ quote name ^ " ...");
   495 
   496     val ctxt = ProofContext.init thy;
   497 
   498     fun add_parent (Ts,pname,rs) env =
   499       let
   500         val full_pname = Sign.full_bname thy pname;
   501         val {args,components,...} =
   502               (case get_statespace (Context.Theory thy) full_pname of
   503                 SOME r => r
   504                | NONE => error ("Undefined statespace " ^ quote pname));
   505 
   506 
   507         val (Ts',env') = fold_map (prep_typ ctxt) Ts env
   508             handle ERROR msg => cat_error msg
   509                     ("The error(s) above occured in parent statespace specification "
   510                     ^ quote pname);
   511         val err_insts = if length args <> length Ts' then
   512             ["number of type instantiation(s) does not match arguments of parent statespace "
   513               ^ quote pname]
   514             else [];
   515 
   516         val rnames = map fst rs
   517         val err_dup_renamings = (case duplicates (op =) rnames of
   518              [] => []
   519             | dups => ["Duplicate renaming(s) for " ^ commas dups])
   520 
   521         val cnames = map fst components;
   522         val err_rename_unknowns = (case (filter (fn n => not (n mem cnames))) rnames of
   523               [] => []
   524              | rs => ["Unknown components " ^ commas rs]);
   525 
   526 
   527         val rs' = map (AList.lookup (op =) rs o fst) components;
   528         val errs =err_insts @ err_dup_renamings @ err_rename_unknowns
   529       in if null errs then ((Ts',full_pname,rs'),env')
   530          else error (cat_lines (errs @ ["in parent statespace " ^ quote pname]))
   531       end;
   532 
   533     val (parents',env) = fold_map add_parent parents [];
   534 
   535     val err_dup_args =
   536          (case duplicates (op =) args of
   537             [] => []
   538           | dups => ["Duplicate type argument(s) " ^ commas dups]);
   539 
   540 
   541     val err_dup_components =
   542          (case duplicates (op =) (map fst comps) of
   543            [] => []
   544           | dups => ["Duplicate state-space components " ^ commas dups]);
   545 
   546     fun prep_comp (n,T) env =
   547       let val (T', env') = prep_typ ctxt T env handle ERROR msg =>
   548        cat_error msg ("The error(s) above occured in component " ^ quote n)
   549       in ((n,T'), env') end;
   550 
   551     val (comps',env') = fold_map prep_comp comps env;
   552 
   553     val err_extra_frees =
   554       (case subtract (op =) args (map fst env') of
   555         [] => []
   556       | extras => ["Extra free type variable(s) " ^ commas extras]);
   557 
   558     val defaultS = Sign.defaultS thy;
   559     val args' = map (fn x => (x, AList.lookup (op =) env x |> the_default defaultS)) args;
   560 
   561 
   562     fun fst_eq ((x:string,_),(y,_)) = x = y;
   563     fun snd_eq ((_,t:typ),(_,u)) = t = u;
   564 
   565     val raw_parent_comps = (List.concat (map (parent_components thy) parents'));
   566     fun check_type (n,T) =
   567           (case distinct (snd_eq) (filter (curry fst_eq (n,T)) raw_parent_comps) of
   568              []  => []
   569            | [_] => []
   570            | rs  => ["Different types for component " ^ n ^": " ^ commas
   571                        (map (Pretty.string_of o Display.pretty_ctyp o ctyp_of thy o snd) rs)])
   572 
   573     val err_dup_types = List.concat (map check_type (duplicates fst_eq raw_parent_comps))
   574 
   575     val parent_comps = distinct (fst_eq) raw_parent_comps;
   576     val all_comps = parent_comps @ comps';
   577     val err_comp_in_parent = (case duplicates (op =) (map fst all_comps) of
   578                [] => []
   579              | xs => ["Components already defined in parents: " ^ commas xs]);
   580     val errs = err_dup_args @ err_dup_components @ err_extra_frees @
   581                err_dup_types @ err_comp_in_parent;
   582   in if null errs
   583      then thy |> statespace_definition state_space args' name parents' parent_comps comps'
   584      else error (cat_lines errs)
   585   end
   586   handle ERROR msg => cat_error msg ("Failed to define statespace " ^ quote name);
   587 
   588 val define_statespace = gen_define_statespace Record.read_typ NONE;
   589 val define_statespace_i = gen_define_statespace Record.cert_typ;
   590 
   591 
   592 (*** parse/print - translations ***)
   593 
   594 
   595 local
   596 fun map_get_comp f ctxt (Free (name,_)) =
   597       (case (get_comp ctxt name) of
   598         SOME (T,_) => f T T dummyT
   599       | NONE => (Syntax.free "arbitrary"(*; error "context not ready"*)))
   600   | map_get_comp _ _ _ = Syntax.free "arbitrary";
   601 
   602 val get_comp_projection = map_get_comp project_free;
   603 val get_comp_injection  = map_get_comp inject_free;
   604 
   605 fun name_of (Free (n,_)) = n;
   606 in
   607 
   608 fun gen_lookup_tr ctxt s n =
   609       (case get_comp (Context.Proof ctxt) n of
   610         SOME (T,_) =>
   611            Syntax.const "StateFun.lookup"$Syntax.free (project_name T)$Syntax.free n$s
   612        | NONE =>
   613            if get_silent (Context.Proof ctxt)
   614 	   then Syntax.const "StateFun.lookup" $ Syntax.const "undefined" $ Syntax.free n $ s
   615            else raise TERM ("StateSpace.gen_lookup_tr: component " ^ n ^ " not defined",[]));
   616 
   617 fun lookup_tr ctxt [s,Free (n,_)] = gen_lookup_tr ctxt s n;
   618 fun lookup_swap_tr ctxt [Free (n,_),s] = gen_lookup_tr ctxt s n;
   619 
   620 fun lookup_tr' ctxt [_$Free (prj,_),n as (_$Free (name,_)),s] =
   621      ( case get_comp (Context.Proof ctxt) name of
   622          SOME (T,_) =>  if prj=project_name T then
   623                            Syntax.const "_statespace_lookup" $ s $ n
   624                         else raise Match
   625       | NONE => raise Match)
   626   | lookup_tr' _ ts = raise Match;
   627 
   628 fun gen_update_tr id ctxt n v s =
   629   let
   630     fun pname T = if id then "Fun.id" else project_name T
   631     fun iname T = if id then "Fun.id" else inject_name T
   632    in
   633      (case get_comp (Context.Proof ctxt) n of
   634        SOME (T,_) => Syntax.const "StateFun.update"$
   635                    Syntax.free (pname T)$Syntax.free (iname T)$
   636                    Syntax.free n$(Syntax.const KN $ v)$s
   637       | NONE =>
   638          if get_silent (Context.Proof ctxt)
   639          then Syntax.const "StateFun.update"$
   640                    Syntax.const "undefined" $ Syntax.const "undefined" $
   641                    Syntax.free n $ (Syntax.const KN $ v) $ s
   642          else raise TERM ("StateSpace.gen_update_tr: component " ^ n ^ " not defined",[]))
   643    end;
   644 
   645 fun update_tr ctxt [s,Free (n,_),v] = gen_update_tr false ctxt n v s;
   646 
   647 fun update_tr' ctxt [_$Free (prj,_),_$Free (inj,_),n as (_$Free (name,_)),(Const (k,_)$v),s] =
   648      if Long_Name.base_name k = Long_Name.base_name KN then
   649         (case get_comp (Context.Proof ctxt) name of
   650           SOME (T,_) => if inj=inject_name T andalso prj=project_name T then
   651                            Syntax.const "_statespace_update" $ s $ n $ v
   652                         else raise Match
   653          | NONE => raise Match)
   654      else raise Match
   655   | update_tr' _ _ = raise Match;
   656 
   657 end;
   658 
   659 
   660 (*** outer syntax *)
   661 
   662 local structure P = OuterParse and K = OuterKeyword in
   663 
   664 val type_insts =
   665   P.typ >> single ||
   666   P.$$$ "(" |-- P.!!! (P.list1 P.typ --| P.$$$ ")")
   667 
   668 val comp = P.name -- (P.$$$ "::" |-- P.!!! P.typ);
   669 fun plus1_unless test scan =
   670   scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan));
   671 
   672 val mapsto = P.$$$ "=";
   673 val rename = P.name -- (mapsto |-- P.name);
   674 val renames =  Scan.optional (P.$$$ "[" |-- P.!!! (P.list1 rename --| P.$$$ "]")) [];
   675 
   676 
   677 val parent = ((type_insts -- P.xname) || (P.xname >> pair [])) -- renames
   678              >> (fn ((insts,name),renames) => (insts,name,renames))
   679 
   680 
   681 val statespace_decl =
   682    P.type_args -- P.name --
   683     (P.$$$ "=" |--
   684      ((Scan.repeat1 comp >> pair []) ||
   685       (plus1_unless comp parent --
   686         Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 comp)) [])))
   687 
   688 val statespace_command =
   689   OuterSyntax.command "statespace" "define state space" K.thy_decl
   690   (statespace_decl >> (fn ((args,name),(parents,comps)) =>
   691                            Toplevel.theory (define_statespace args name parents comps)))
   692 
   693 end;
   694 
   695 end;