src/Tools/isac/BaseDefinitions/Know_Store.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60508 ce09935439b3
child 60521 23c40bb1bdbf
permissions -rw-r--r--
polish naming in Rewrite_Order
walther@59887
     1
(*  Title:      src/Tools/isac/Know_Store.thy
wneuper@59586
     2
    Author:     Mathias Lehnfeld
walther@59871
     3
walther@60154
     4
Calc work on Problem employing MethodC; both are collected in a respective Store;
walther@59890
     5
The collections' structure is independent from the dependency graph of Isabelle's theories
walther@59890
     6
in Knowledge.
walther@59917
     7
Store is also used by Thy_Write, required for Isac's Java front end, irrelevant for PIDE.
walther@59890
     8
walther@59887
     9
The files (in "xxxxx-def.sml") contain definitions required for Know_Store;
walther@59871
    10
they also include minimal code required for other "xxxxx-def.sml" files.
walther@59871
    11
These files have companion files "xxxxx.sml" with all further code,
walther@59871
    12
located at appropriate positions in the file structure.
walther@59871
    13
walther@59871
    14
The separation of "xxxxx-def.sml" from "xxxxx.sml" should be overcome by
walther@59871
    15
appropriate use of polymorphic high order functions.
wneuper@59586
    16
*)
wneuper@59586
    17
wenzelm@60286
    18
theory Know_Store
wenzelm@60286
    19
  imports Complex_Main
wenzelm@60313
    20
  keywords "rule_set_knowledge" "calculation" :: thy_decl
wenzelm@60223
    21
begin
wneuper@59586
    22
wenzelm@60223
    23
setup \<open>
wenzelm@60223
    24
  ML_Antiquotation.conditional \<^binding>\<open>isac_test\<close>
wenzelm@60223
    25
    (fn _ => Options.default_bool \<^system_option>\<open>isac_test\<close>)
wenzelm@60223
    26
\<close>
wenzelm@60223
    27
wneuper@59586
    28
ML_file libraryC.sml
walther@59880
    29
ML_file theoryC.sml
walther@59868
    30
ML_file unparseC.sml
walther@59874
    31
ML_file "rule-def.sml"
walther@59875
    32
ML_file "thmC-def.sml"
walther@59919
    33
ML_file "eval-def.sml"       (*rename identifiers by use of struct.id*)
walther@59867
    34
ML_file "rewrite-order.sml"  (*rename identifiers by use of struct.id*)
wneuper@59586
    35
ML_file rule.sml
walther@59909
    36
ML_file "error-pattern-def.sml"
walther@59850
    37
ML_file "rule-set.sml"
walther@59882
    38
walther@59891
    39
ML_file "store.sml"
walther@59892
    40
ML_file "check-unique.sml"
walther@59985
    41
ML_file "references-def.sml"
walther@59945
    42
ML_file "model-pattern.sml"
walther@59893
    43
ML_file "problem-def.sml"
walther@59893
    44
ML_file "method-def.sml"
walther@59893
    45
ML_file "cas-def.sml"
Walther@60505
    46
ML_file "formalise.sml"
Walther@60505
    47
ML_file "example.sml"
walther@59917
    48
ML_file "thy-write.sml"
walther@59882
    49
wneuper@59586
    50
ML \<open>
wneuper@59586
    51
\<close> ML \<open>
wneuper@59587
    52
\<close> ML \<open>
wneuper@59586
    53
\<close>
wneuper@59586
    54
section \<open>Knowledge elements for problems and methods\<close>
wneuper@59586
    55
ML \<open>
walther@59887
    56
(* Knowledge (and Exercises) are held by "Know_Store" in Isac's Java front-end.
wneuper@59586
    57
  In the front-end Knowledge comprises theories, problems and methods.
wneuper@59586
    58
  Elements of problems and methods are defined in theories alongside
wneuper@59586
    59
  the development of respective language elements. 
wneuper@59586
    60
  However, the structure of methods and problems is independent from theories' 
wneuper@59586
    61
  deductive structure. Thus respective structures are built in Build_Thydata.thy.
wneuper@59586
    62
wneuper@59586
    63
  Most elements of problems and methods are implemented in "Knowledge/", but some
walther@59887
    64
  of them are implemented in "ProgLang/" already; thus "Know_Store.thy" got this
wneuper@59586
    65
  location in the directory structure.
wneuper@59586
    66
wneuper@59586
    67
  get_* retrieves all * of the respective theory PLUS of all ancestor theories.
wneuper@59586
    68
*)
wneuper@59586
    69
signature KESTORE_ELEMS =
wneuper@59586
    70
sig
Walther@60509
    71
  val get_rew_ord: theory -> Rewrite_Ord.T list
Walther@60509
    72
  val add_rew_ord: Rewrite_Ord.T list -> theory -> theory
walther@59879
    73
  val get_rlss: theory -> (Rule_Set.id * (ThyC.id * Rule_Set.T)) list
walther@59879
    74
  val add_rlss: (Rule_Set.id * (ThyC.id * Rule_Set.T)) list -> theory -> theory
walther@59919
    75
  val get_calcs: theory -> (Eval_Def.prog_calcID * (Eval_Def.calID * Eval_Def.eval_fn)) list
walther@59919
    76
  val add_calcs: (Eval_Def.prog_calcID * (Eval_Def.calID * Eval_Def.eval_fn)) list -> theory -> theory
walther@59896
    77
  val get_cas: theory -> CAS_Def.T list
walther@59896
    78
  val add_cas: CAS_Def.T list -> theory -> theory
Walther@60495
    79
  val get_pbls: theory -> Probl_Def.store
Walther@60502
    80
  val add_pbls: Proof.context -> (Probl_Def.T * References_Def.id) list -> theory -> theory
walther@59895
    81
  val get_mets: theory -> Meth_Def.store
Walther@60502
    82
  val add_mets: Proof.context -> (Meth_Def.T * References_Def.id) list -> theory -> theory
Walther@60505
    83
  val get_expls: theory -> Example.store
Walther@60505
    84
  val add_expls: (Example.T * Store.key) list -> theory -> theory
walther@59917
    85
  val get_thes: theory -> (Thy_Write.thydata Store.node) list
walther@59917
    86
  val add_thes: (Thy_Write.thydata * Thy_Write.theID) list -> theory -> theory (* thydata dropped at existing elems *)
walther@59917
    87
  val insert_fillpats: (Thy_Write.theID * Error_Pattern_Def.fill_in list) list -> theory -> theory 
wneuper@59586
    88
  val get_ref_thy: unit -> theory
wneuper@59586
    89
  val set_ref_thy: theory -> unit
Walther@60508
    90
end;
wneuper@59586
    91
Walther@60506
    92
structure KEStore_Elems(**): KESTORE_ELEMS(**) =
wneuper@59586
    93
struct
wneuper@59586
    94
  fun union_overwrite eq l1 l2 = fold (insert eq) l2 (*..swapped..*) l1;
wneuper@59586
    95
wneuper@59586
    96
  structure Data = Theory_Data (
Walther@60506
    97
    type T = (string * (LibraryC.subst -> term * term -> bool)) list;
Walther@60506
    98
    val empty = [];
Walther@60506
    99
    val extend = I;
Walther@60506
   100
    val merge = merge Rewrite_Ord.equal;
Walther@60506
   101
    );  
Walther@60506
   102
  fun get_rew_ord thy = Data.get thy
Walther@60506
   103
  fun add_rew_ord rlss = Data.map (union_overwrite Rewrite_Ord.equal rlss)
Walther@60506
   104
Walther@60506
   105
  structure Data = Theory_Data (
walther@59879
   106
    type T = (Rule_Set.id * (ThyC.id * Rule_Set.T)) list;
wneuper@59586
   107
    val empty = [];
wneuper@59586
   108
    val extend = I;
walther@59852
   109
    val merge = Rule_Set.to_kestore;
wneuper@59586
   110
    );  
wneuper@59586
   111
  fun get_rlss thy = Data.get thy
walther@59852
   112
  fun add_rlss rlss = Data.map (union_overwrite Rule_Set.equal rlss)
wneuper@59586
   113
wneuper@59586
   114
  structure Data = Theory_Data (
walther@59919
   115
    type T = (Eval_Def.prog_calcID * (Eval_Def.calID * Eval_Def.eval_fn)) list;
wneuper@59586
   116
    val empty = [];
wneuper@59586
   117
    val extend = I;
walther@59919
   118
    val merge = merge Eval_Def.calc_eq;
wneuper@59586
   119
    );                                                              
wneuper@59586
   120
  fun get_calcs thy = Data.get thy
walther@59919
   121
  fun add_calcs calcs = Data.map (union_overwrite Eval_Def.calc_eq calcs)
wneuper@59586
   122
wneuper@59586
   123
  structure Data = Theory_Data (
walther@59985
   124
    type T = (term * (References_Def.T * (term list -> (term * term list) list))) list;
wneuper@59586
   125
    val empty = [];
wneuper@59586
   126
    val extend = I;
walther@59896
   127
    val merge = merge CAS_Def.equal;
wneuper@59586
   128
    );                                                              
wneuper@59586
   129
  fun get_cas thy = Data.get thy
walther@59896
   130
  fun add_cas cas = Data.map (union_overwrite CAS_Def.equal cas)
wneuper@59586
   131
wneuper@59586
   132
  structure Data = Theory_Data (
walther@59894
   133
    type T = Probl_Def.store;
walther@59894
   134
    val empty = [Probl_Def.empty_store];
wneuper@59586
   135
    val extend = I;
walther@59897
   136
    val merge = Store.merge;
wneuper@59586
   137
    );
Walther@60495
   138
  fun get_pbls thy = Data.get thy;
Walther@60502
   139
  fun add_pbls ctxt pbts thy =
Walther@60502
   140
    let
Walther@60502
   141
      fun add_pbt (pbt as {guh,...}, pblID) =
Walther@60502
   142
        (* the pblID has the leaf-element as first; better readability achieved *)
Walther@60502
   143
        (if (Config.get ctxt check_unique) then Probl_Def.check_unique guh (Data.get thy) else ();
Walther@60502
   144
          rev pblID |> Store.insert pblID pbt);
Walther@60502
   145
    in Data.map (fold add_pbt pbts) thy end;
wneuper@59586
   146
wneuper@59586
   147
  structure Data = Theory_Data (
walther@59895
   148
    type T = Meth_Def.store;
walther@59895
   149
    val empty = [Meth_Def.empty_store];
wneuper@59586
   150
    val extend = I;
walther@59897
   151
    val merge = Store.merge;
wneuper@59586
   152
    );
wneuper@59586
   153
  val get_mets = Data.get;
Walther@60502
   154
  fun add_mets ctxt mets thy =
Walther@60502
   155
    let
Walther@60502
   156
      fun add_met (met as {guh,...}, metID) =
Walther@60502
   157
        (if (Config.get ctxt check_unique) then Meth_Def.check_unique guh (Data.get thy) else ();
Walther@60502
   158
          Store.insert metID met metID);
Walther@60502
   159
    in Data.map (fold add_met mets) thy end;
wneuper@59586
   160
wneuper@59586
   161
  structure Data = Theory_Data (
Walther@60505
   162
    type T = Example.store;
Walther@60505
   163
    val empty = [Example.empty_store];
Walther@60505
   164
    val extend = I;
Walther@60505
   165
    val merge = Store.merge;
Walther@60505
   166
    );
Walther@60505
   167
  val get_expls = Data.get;
Walther@60505
   168
  fun add_expls expls thy =
Walther@60505
   169
    let
Walther@60505
   170
      fun add_expl (expl, expl_id) = Store.insert expl_id expl expl_id;
Walther@60505
   171
    in Data.map (fold add_expl expls) thy end;
Walther@60505
   172
Walther@60505
   173
  structure Data = Theory_Data (
walther@59917
   174
    type T = (Thy_Write.thydata Store.node) list;
wneuper@59586
   175
    val empty = [];
wneuper@59586
   176
    val extend = I;
walther@59897
   177
    val merge = Store.merge; (* relevant for store_thm, store_rls *)
wneuper@59586
   178
    );                                                              
wneuper@59586
   179
  fun get_thes thy = Data.get thy
wneuper@59586
   180
  fun add_thes thes thy = let
walther@59917
   181
    fun add_the (thydata, theID) = Thy_Write.add_thydata ([], theID) thydata
wneuper@59586
   182
  in Data.map (fold add_the thes) thy end;
wneuper@59586
   183
  fun insert_fillpats fis thy =
wneuper@59586
   184
    let
wneuper@59586
   185
      fun update_elem (theID, fillpats) =
wneuper@59586
   186
        let
walther@59897
   187
          val hthm = Store.get (Data.get thy) theID theID
walther@59917
   188
          val hthm' = Thy_Write.update_hthm hthm fillpats
wneuper@59586
   189
            handle ERROR _ =>
walther@59962
   190
              raise ERROR ("insert_fillpats: " ^ strs2str theID ^ "must address a theorem")
walther@59897
   191
        in Store.update theID theID hthm' end
wneuper@59586
   192
    in Data.map (fold update_elem fis) thy end
wneuper@59586
   193
wneuper@59586
   194
  val cur_thy = Synchronized.var "finally_knowledge_complete" @{theory};
wneuper@59586
   195
  fun set_ref_thy thy = Synchronized.change cur_thy (fn _ => thy); (* never RE-set ! *)
wneuper@59586
   196
  fun get_ref_thy () = Synchronized.value cur_thy;
Walther@60506
   197
(**)end(**);
wneuper@59586
   198
\<close>
wneuper@59586
   199
wenzelm@60286
   200
wenzelm@60286
   201
subsection \<open>Isar command syntax\<close>
wenzelm@60286
   202
wenzelm@60286
   203
ML \<open>
wenzelm@60286
   204
local
wenzelm@60286
   205
wenzelm@60286
   206
val parse_rule = Parse.name -- Parse.!!! (\<^keyword>\<open>=\<close> |-- Parse.ML_source);
wenzelm@60286
   207
wenzelm@60286
   208
val ml = ML_Lex.read;
wenzelm@60286
   209
wenzelm@60286
   210
fun ml_rule thy (name, source) =
wenzelm@60286
   211
  ml "(" @ ml (ML_Syntax.print_string name) @ ml ", " @
wenzelm@60286
   212
  ml "(" @ ml (ML_Syntax.print_string (Context.theory_name thy)) @ ml ", " @
wenzelm@60286
   213
  ML_Lex.read_source source @ ml "))";
wenzelm@60286
   214
wenzelm@60286
   215
fun ml_rules thy args =
wenzelm@60286
   216
  ml "Theory.setup (KEStore_Elems.add_rlss [" @
wenzelm@60286
   217
    flat (separate (ml ",") (map (ml_rule thy) args)) @ ml "])";
wenzelm@60286
   218
wenzelm@60286
   219
val _ =
wenzelm@60289
   220
  Outer_Syntax.command \<^command_keyword>\<open>rule_set_knowledge\<close> "register ISAC rule set to Knowledge Store"
wenzelm@60286
   221
    (Parse.and_list1 parse_rule >> (fn args => Toplevel.theory (fn thy =>
wenzelm@60286
   222
      thy |> Context.theory_map
wenzelm@60286
   223
        (ML_Context.expression (Position.thread_data ()) (ml_rules thy args)))));
wenzelm@60286
   224
wenzelm@60313
   225
val calc_name =
wenzelm@60313
   226
  Parse.name -- (\<^keyword>\<open>(\<close> |-- Parse.const --| \<^keyword>\<open>)\<close>) ||
wenzelm@60313
   227
  Scan.ahead Parse.name -- Parse.const;
wenzelm@60313
   228
wenzelm@60313
   229
val _ =
wenzelm@60313
   230
  Outer_Syntax.command \<^command_keyword>\<open>calculation\<close>
wenzelm@60313
   231
    "prepare ISAC calculation and register it to Knowledge Store"
wenzelm@60313
   232
    (calc_name -- (\<^keyword>\<open>=\<close> |-- Parse.!!! Parse.ML_source) >> (fn ((calcID, const), source) =>
wenzelm@60313
   233
      Toplevel.theory (fn thy =>
wenzelm@60313
   234
        let
wenzelm@60313
   235
          val ctxt = Proof_Context.init_global thy;
wenzelm@60313
   236
          val Const (c, _) = Proof_Context.read_const {proper = true, strict = true} ctxt const;
wenzelm@60314
   237
          val set_data =
wenzelm@60313
   238
            ML_Context.expression (Input.pos_of source)
wenzelm@60313
   239
              (ml "Theory.setup (Eval_Def.set_data (" @ ML_Lex.read_source source @ ml "))")
wenzelm@60313
   240
            |> Context.theory_map;
wenzelm@60314
   241
          val eval = Eval_Def.the_data (set_data thy);
wenzelm@60313
   242
        in KEStore_Elems.add_calcs [(calcID, (c, eval))] thy end)))
wenzelm@60313
   243
wenzelm@60286
   244
in end;
wenzelm@60286
   245
\<close>
wenzelm@60286
   246
wenzelm@60286
   247
wneuper@59586
   248
section \<open>Re-use existing access functions for knowledge elements\<close>
wneuper@59586
   249
text \<open>
Walther@60508
   250
  The independence of problems' and methods' structure from theory dependency structure
Walther@60508
   251
  enforces the access functions to use "Isac_Knowledge",
Walther@60508
   252
  the final theory which comprises all knowledge defined.
wneuper@59586
   253
\<close>
wneuper@59586
   254
ML \<open>
wneuper@59586
   255
val get_ref_thy = KEStore_Elems.get_ref_thy;
wneuper@59586
   256
Walther@60506
   257
fun assoc' ([], key) = raise ERROR ("ME_Isa: rew_ord \"" ^ key ^ "\" not in system 2")
Walther@60506
   258
  | assoc' ((keyi, xi) :: pairs, key) =
Walther@60506
   259
    if key = keyi then SOME xi else assoc' (pairs, key);
Walther@60506
   260
fun assoc_rew_ord thy ro = ((the o assoc') (KEStore_Elems.get_rew_ord thy, ro))
Walther@60506
   261
  handle ERROR _ => raise ERROR ("ME_Isa: rew_ord \"" ^ ro ^ "\" not in system 1");
Walther@60506
   262
walther@59867
   263
fun assoc_rls (rls' : Rule_Set.id) =
walther@59879
   264
  case AList.lookup (op =) (KEStore_Elems.get_rlss (ThyC.get_theory "Isac_Knowledge")) rls' of
wneuper@59586
   265
    SOME (_, rls) => rls
walther@59887
   266
  | NONE => raise ERROR ("rls \""^ rls' ^ "\" missing in Know_Store.\n" ^
wneuper@59586
   267
    "TODO exception hierarchy needs to be established.")
wneuper@59586
   268
walther@59867
   269
fun assoc_rls' thy (rls' : Rule_Set.id) =
wneuper@59586
   270
  case AList.lookup (op =) (KEStore_Elems.get_rlss thy) rls' of
wneuper@59586
   271
    SOME (_, rls) => rls
walther@59887
   272
  | NONE => raise ERROR ("rls \""^ rls' ^ "\" missing in Know_Store.\n" ^
wneuper@59586
   273
    "TODO exception hierarchy needs to be established.")
wneuper@59586
   274
wneuper@59586
   275
fun assoc_calc thy calID = let
wneuper@59586
   276
    fun ass ([], key) =
walther@59962
   277
          raise ERROR ("assoc_calc: '" ^ key ^ "' not found in theory " ^ (Context.theory_name thy))
wneuper@59586
   278
      | ass ((calc, (keyi, _)) :: pairs, key) =
wneuper@59586
   279
          if key = keyi then calc else ass (pairs, key);
wneuper@59586
   280
  in ass (thy |> KEStore_Elems.get_calcs, calID) end;
wneuper@59586
   281
wneuper@59586
   282
fun assoc_calc' thy key = let
wneuper@59586
   283
    fun ass ([], key') =
walther@59962
   284
          raise ERROR ("assoc_calc': '" ^ key' ^ "' not found in theory " ^ (Context.theory_name thy))
wneuper@59586
   285
      | ass ((all as (keyi, _)) :: pairs, key') =
wneuper@59586
   286
          if key' = keyi then all else ass (pairs, key');
wneuper@59586
   287
  in ass (KEStore_Elems.get_calcs thy, key) end;
wneuper@59586
   288
wneuper@59586
   289
fun assoc_cas thy key = assoc (KEStore_Elems.get_cas thy, key);
wneuper@59586
   290
Walther@60495
   291
fun get_pbls () = get_ref_thy () |> KEStore_Elems.get_pbls;
wneuper@59586
   292
fun get_mets () = get_ref_thy () |> KEStore_Elems.get_mets;
wneuper@59586
   293
fun get_thes () = get_ref_thy () |> KEStore_Elems.get_thes;
Walther@60505
   294
fun get_expls () = get_ref_thy () |> KEStore_Elems.get_expls;
wneuper@59586
   295
\<close>
wenzelm@60286
   296
wenzelm@60289
   297
rule_set_knowledge
wenzelm@60286
   298
  empty = \<open>Rule_Set.empty\<close> and
wenzelm@60286
   299
  e_rrls = \<open>Rule_Set.e_rrls\<close>
wneuper@59586
   300
wneuper@59586
   301
section \<open>determine sequence of main parts in thehier\<close>
wneuper@59586
   302
setup \<open>
wneuper@59586
   303
KEStore_Elems.add_thes
walther@59917
   304
  [(Thy_Write.Html {guh = Thy_Write.part2guh ["IsacKnowledge"], html = "",
wneuper@59586
   305
    mathauthors = ["Isac team"], coursedesign = []}, ["IsacKnowledge"]),
walther@59917
   306
  (Thy_Write.Html {guh = Thy_Write.part2guh ["Isabelle"], html = "",
wneuper@59586
   307
    mathauthors = ["Isabelle team, TU Munich"], coursedesign = []}, ["Isabelle"]),
walther@59917
   308
  (Thy_Write.Html {guh = Thy_Write.part2guh ["IsacScripts"], html = "",
wneuper@59586
   309
    mathauthors = ["Isac team"], coursedesign = []}, ["IsacScripts"])]
wneuper@59586
   310
\<close>
wneuper@59586
   311
wneuper@59586
   312
section \<open>Functions for checking KEStore_Elems\<close>
wneuper@59586
   313
ML \<open>
walther@59851
   314
fun short_string_of_rls Rule_Set.Empty = "Erls"
walther@59851
   315
  | short_string_of_rls (Rule_Def.Repeat {calc, rules, ...}) =
wneuper@59586
   316
    "Rls {#calc = " ^ string_of_int (length calc) ^
wneuper@59586
   317
    ", #rules = " ^ string_of_int (length rules) ^ ", ..."
walther@59878
   318
  | short_string_of_rls (Rule_Set.Sequence {calc, rules, ...}) =
wneuper@59586
   319
    "Seq {#calc = " ^ string_of_int (length calc) ^
wneuper@59586
   320
    ", #rules = " ^ string_of_int (length rules) ^ ", ..."
walther@59850
   321
  | short_string_of_rls (Rule_Set.Rrls _) = "Rrls {...}";
wneuper@59586
   322
fun check_kestore_rls (rls', (thyID, rls)) =
wneuper@59586
   323
  "(" ^ rls' ^ ", (" ^ thyID ^ ", " ^ short_string_of_rls rls ^ "))";
wneuper@59586
   324
walther@59852
   325
fun check_kestore_calc ((id, (c, _)) : Rule_Def.calc)  = "(" ^ id ^ ", (" ^ c ^ ", fn))";
wneuper@59586
   326
walther@59846
   327
(* we avoid term_to_string''' defined later *)
walther@59896
   328
fun check_kestore_cas ((t, (s, _)) : CAS_Def.T) =
walther@59846
   329
  "(" ^ (Print_Mode.setmp [] (Syntax.string_of_term (Config.put show_markup false
walther@59985
   330
  (Proof_Context.init_global @{theory})))) t ^ ", " ^ References_Def.to_string s ^ ")";
wneuper@59586
   331
wneuper@59586
   332
fun count_kestore_ptyps [] = 0
walther@59897
   333
  | count_kestore_ptyps ((Store.Node (_, _, ps)) :: ps') =
wneuper@59586
   334
      1 + count_kestore_ptyps ps  + count_kestore_ptyps ps';
walther@59897
   335
fun check_kestore_ptyp' strfun (Store.Node (key, pbts, pts)) = "Ptyp (" ^ (quote key) ^ ", " ^
walther@59888
   336
      (strfun pbts) ^ ", " ^ (map (check_kestore_ptyp' strfun) pts |> list2str) ^ ")" |> linefeed;
walther@59894
   337
val check_kestore_ptyp = check_kestore_ptyp' Probl_Def.s_to_string;
walther@59897
   338
fun ptyp_ord ((Store.Node (s1, _, _)), (Store.Node (s2, _, _))) = string_ord (s1, s2);
walther@59894
   339
fun pbt_ord ({guh = guh'1, ...} : Probl_Def.T, {guh = guh'2, ...} : Probl_Def.T) = string_ord (guh'1, guh'2);
wneuper@59586
   340
fun sort_kestore_ptyp' _ [] = []
walther@59897
   341
  | sort_kestore_ptyp' ordfun ((Store.Node (key, pbts, ps)) :: ps') =
walther@59897
   342
     ((Store.Node (key, sort ordfun pbts, sort_kestore_ptyp' ordfun ps |> sort ptyp_ord))
wneuper@59586
   343
       :: sort_kestore_ptyp' ordfun ps');
wneuper@59586
   344
val sort_kestore_ptyp = sort_kestore_ptyp' pbt_ord;
wneuper@59586
   345
walther@59895
   346
fun metguh2str ({guh,...} : Meth_Def.T) = guh : string;
walther@59897
   347
fun check_kestore_met (mp: Meth_Def.T Store.node) =
wneuper@59586
   348
      check_kestore_ptyp' (fn xs => map metguh2str xs |> strs2str) mp;
walther@59895
   349
fun met_ord ({guh = guh'1, ...} : Meth_Def.T, {guh = guh'2, ...} : Meth_Def.T) = string_ord (guh'1, guh'2);
wneuper@59586
   350
val sort_kestore_met = sort_kestore_ptyp' met_ord;
wneuper@59586
   351
walther@59917
   352
fun check_kestore_thes thes = ((map writeln) o (map (check_kestore_ptyp' Thy_Write.thes2str))) thes
wneuper@59586
   353
fun write_thes thydata_list =
wneuper@59586
   354
  thydata_list 
walther@59917
   355
    |> map (fn (id, the) => (Thy_Write.theID2str id, Thy_Write.the2str the))
wneuper@59586
   356
    |> map pair2str
wneuper@59586
   357
    |> map writeln
wneuper@59586
   358
\<close>
wneuper@59586
   359
ML \<open>
wneuper@59586
   360
\<close> ML \<open>
wneuper@59586
   361
\<close> ML \<open>
wneuper@59586
   362
\<close>
wneuper@59586
   363
end