src/Tools/isac/BaseDefinitions/Know_Store.thy
author Walther Neuper <walther.neuper@jku.at>
Sun, 19 Apr 2020 16:43:53 +0200
changeset 59891 68396aa5821f
parent 59890 ba0757da0dc8
child 59892 b8cfae027755
permissions -rw-r--r--
proper names for Celem1, Celem3
walther@59887
     1
(*  Title:      src/Tools/isac/Know_Store.thy
wneuper@59586
     2
    Author:     Mathias Lehnfeld
walther@59871
     3
walther@59890
     4
Calc work on Problem employing Method; 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@59890
     7
Store is also used by Thy_Html, 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
walther@59887
    18
theory Know_Store imports Complex_Main
wneuper@59586
    19
wneuper@59586
    20
begin
wneuper@59586
    21
ML_file libraryC.sml
walther@59880
    22
ML_file theoryC.sml
walther@59868
    23
ML_file unparseC.sml
walther@59874
    24
ML_file "rule-def.sml"
walther@59875
    25
ML_file "thmC-def.sml"
walther@59867
    26
ML_file "exec-def.sml"       (*rename identifiers by use of struct.id*)
walther@59867
    27
ML_file "rewrite-order.sml"  (*rename identifiers by use of struct.id*)
wneuper@59586
    28
ML_file rule.sml
walther@59867
    29
ML_file "error-fill-def.sml" (*rename identifiers by use of struct.id*)
walther@59850
    30
ML_file "rule-set.sml"
walther@59882
    31
walther@59884
    32
ML_file "celem-0.sml"  (*survey Celem.*)
walther@59891
    33
ML_file "store.sml"
walther@59891
    34
ML_file "specification.sml"
walther@59884
    35
ML_file "celem-4.sml"  (*pat*)
walther@59884
    36
ML_file "celem-5.sml"  (*ptyps*)
walther@59884
    37
ML_file "celem-6.sml"  (*mets*)
walther@59884
    38
ML_file "celem-7.sml"  (*cas_elem*)
walther@59884
    39
ML_file "celem-8.sml"  (*thydata*)
walther@59884
    40
ML_file "celem-91.sml" (*check_guhs_unique*)
walther@59882
    41
walther@59884
    42
ML_file tracing.sml
wneuper@59586
    43
ML \<open>
wneuper@59586
    44
\<close> ML \<open>
wneuper@59587
    45
\<close> ML \<open>
wneuper@59586
    46
\<close>
wneuper@59586
    47
section \<open>Knowledge elements for problems and methods\<close>
wneuper@59586
    48
ML \<open>
walther@59887
    49
(* Knowledge (and Exercises) are held by "Know_Store" in Isac's Java front-end.
wneuper@59586
    50
  In the front-end Knowledge comprises theories, problems and methods.
wneuper@59586
    51
  Elements of problems and methods are defined in theories alongside
wneuper@59586
    52
  the development of respective language elements. 
wneuper@59586
    53
  However, the structure of methods and problems is independent from theories' 
wneuper@59586
    54
  deductive structure. Thus respective structures are built in Build_Thydata.thy.
wneuper@59586
    55
wneuper@59586
    56
  Most elements of problems and methods are implemented in "Knowledge/", but some
walther@59887
    57
  of them are implemented in "ProgLang/" already; thus "Know_Store.thy" got this
wneuper@59586
    58
  location in the directory structure.
wneuper@59586
    59
wneuper@59586
    60
  get_* retrieves all * of the respective theory PLUS of all ancestor theories.
wneuper@59586
    61
*)
wneuper@59586
    62
signature KESTORE_ELEMS =
wneuper@59586
    63
sig
walther@59879
    64
  val get_rlss: theory -> (Rule_Set.id * (ThyC.id * Rule_Set.T)) list
walther@59879
    65
  val add_rlss: (Rule_Set.id * (ThyC.id * Rule_Set.T)) list -> theory -> theory
walther@59854
    66
  val get_calcs: theory -> (Exec_Def.prog_calcID * (Exec_Def.calID * Exec_Def.eval_fn)) list
walther@59854
    67
  val add_calcs: (Exec_Def.prog_calcID * (Exec_Def.calID * Exec_Def.eval_fn)) list -> theory -> theory
walther@59889
    68
  val get_cas: theory -> Celem7.cas_elem list
walther@59889
    69
  val add_cas: Celem7.cas_elem list -> theory -> theory
walther@59889
    70
  val get_ptyps: theory -> Celem5.ptyps
walther@59891
    71
  val add_pbts: (Celem5.pbt * Spec.pblID) list -> theory -> theory
walther@59889
    72
  val get_mets: theory -> Celem6.mets
walther@59891
    73
  val add_mets: (Celem6.met * Spec.metID) list -> theory -> theory
walther@59890
    74
  val get_thes: theory -> (Celem8.thydata Store.ptyp) list
walther@59889
    75
  val add_thes: (Celem8.thydata * Celem8.theID) list -> theory -> theory (* thydata dropped at existing elems *)
walther@59889
    76
  val insert_fillpats: (Celem8.theID * Error_Fill_Def.fillpat list) list -> theory -> theory 
wneuper@59586
    77
  val get_ref_thy: unit -> theory
wneuper@59586
    78
  val set_ref_thy: theory -> unit
wneuper@59586
    79
end;                               
wneuper@59586
    80
wneuper@59586
    81
structure KEStore_Elems: KESTORE_ELEMS =
wneuper@59586
    82
struct
wneuper@59586
    83
  fun union_overwrite eq l1 l2 = fold (insert eq) l2 (*..swapped..*) l1;
wneuper@59586
    84
wneuper@59586
    85
  structure Data = Theory_Data (
walther@59879
    86
    type T = (Rule_Set.id * (ThyC.id * Rule_Set.T)) list;
wneuper@59586
    87
    val empty = [];
wneuper@59586
    88
    val extend = I;
walther@59852
    89
    val merge = Rule_Set.to_kestore;
wneuper@59586
    90
    );  
wneuper@59586
    91
  fun get_rlss thy = Data.get thy
walther@59852
    92
  fun add_rlss rlss = Data.map (union_overwrite Rule_Set.equal rlss)
wneuper@59586
    93
wneuper@59586
    94
  structure Data = Theory_Data (
walther@59854
    95
    type T = (Exec_Def.prog_calcID * (Exec_Def.calID * Exec_Def.eval_fn)) list;
wneuper@59586
    96
    val empty = [];
wneuper@59586
    97
    val extend = I;
walther@59853
    98
    val merge = merge Exec_Def.calc_eq;
wneuper@59586
    99
    );                                                              
wneuper@59586
   100
  fun get_calcs thy = Data.get thy
walther@59853
   101
  fun add_calcs calcs = Data.map (union_overwrite Exec_Def.calc_eq calcs)
wneuper@59586
   102
wneuper@59586
   103
  structure Data = Theory_Data (
walther@59891
   104
    type T = (term * (Spec.spec * (term list -> (term * term list) list))) list;
wneuper@59586
   105
    val empty = [];
wneuper@59586
   106
    val extend = I;
walther@59889
   107
    val merge = merge Celem7.cas_eq;
wneuper@59586
   108
    );                                                              
wneuper@59586
   109
  fun get_cas thy = Data.get thy
walther@59889
   110
  fun add_cas cas = Data.map (union_overwrite Celem7.cas_eq cas)
wneuper@59586
   111
wneuper@59586
   112
  structure Data = Theory_Data (
walther@59889
   113
    type T = Celem5.ptyps;
walther@59889
   114
    val empty = [Celem5.e_Ptyp];
wneuper@59586
   115
    val extend = I;
walther@59890
   116
    val merge = Store.merge_ptyps;
wneuper@59586
   117
    );
wneuper@59586
   118
  fun get_ptyps thy = Data.get thy;
wneuper@59586
   119
  fun add_pbts pbts thy = let
wneuper@59586
   120
          fun add_pbt (pbt as {guh,...}, pblID) =
wneuper@59586
   121
                (* the pblID has the leaf-element as first; better readability achieved *)
walther@59889
   122
                (if (!Celem91.check_guhs_unique) then Celem91.check_pblguh_unique guh (Data.get thy) else ();
walther@59890
   123
                  rev pblID |> Store.insrt pblID pbt);
wneuper@59586
   124
        in Data.map (fold add_pbt pbts) thy end;
wneuper@59586
   125
wneuper@59586
   126
  structure Data = Theory_Data (
walther@59889
   127
    type T = Celem6.mets;
walther@59889
   128
    val empty = [Celem6.e_Mets];
wneuper@59586
   129
    val extend = I;
walther@59890
   130
    val merge = Store.merge_ptyps;
wneuper@59586
   131
    );
wneuper@59586
   132
  val get_mets = Data.get;
wneuper@59586
   133
  fun add_mets mets thy = let
wneuper@59586
   134
          fun add_met (met as {guh,...}, metID) =
walther@59889
   135
                (if (!Celem91.check_guhs_unique) then Celem91.check_metguh_unique guh (Data.get thy) else ();
walther@59890
   136
                  Store.insrt metID met metID);
wneuper@59586
   137
        in Data.map (fold add_met mets) thy end;
wneuper@59586
   138
wneuper@59586
   139
  structure Data = Theory_Data (
walther@59890
   140
    type T = (Celem8.thydata Store.ptyp) list;
wneuper@59586
   141
    val empty = [];
wneuper@59586
   142
    val extend = I;
walther@59890
   143
    val merge = Store.merge_ptyps; (* relevant for store_thm, store_rls *)
wneuper@59586
   144
    );                                                              
wneuper@59586
   145
  fun get_thes thy = Data.get thy
wneuper@59586
   146
  fun add_thes thes thy = let
walther@59889
   147
    fun add_the (thydata, theID) = Celem8.add_thydata ([], theID) thydata
wneuper@59586
   148
  in Data.map (fold add_the thes) thy end;
wneuper@59586
   149
  fun insert_fillpats fis thy =
wneuper@59586
   150
    let
wneuper@59586
   151
      fun update_elem (theID, fillpats) =
wneuper@59586
   152
        let
walther@59890
   153
          val hthm = Store.get_py (Data.get thy) theID theID
walther@59889
   154
          val hthm' = Celem8.update_hthm hthm fillpats
wneuper@59586
   155
            handle ERROR _ =>
wneuper@59586
   156
              error ("insert_fillpats: " ^ strs2str theID ^ "must address a theorem")
walther@59890
   157
        in Store.update_ptyps theID theID hthm' end
wneuper@59586
   158
    in Data.map (fold update_elem fis) thy end
wneuper@59586
   159
wneuper@59586
   160
  val cur_thy = Synchronized.var "finally_knowledge_complete" @{theory};
wneuper@59586
   161
  fun set_ref_thy thy = Synchronized.change cur_thy (fn _ => thy); (* never RE-set ! *)
wneuper@59586
   162
  fun get_ref_thy () = Synchronized.value cur_thy;
wneuper@59586
   163
end;
wneuper@59586
   164
\<close>
wneuper@59586
   165
wneuper@59586
   166
section \<open>Re-use existing access functions for knowledge elements\<close>
wneuper@59586
   167
text \<open>
wneuper@59586
   168
  The independence of problems' and methods' structure enforces the accesse
wneuper@59592
   169
  functions to use "Isac_Knowledge", the final theory which comprises all knowledge defined.
wneuper@59586
   170
\<close>
wneuper@59586
   171
ML \<open>
wneuper@59586
   172
val get_ref_thy = KEStore_Elems.get_ref_thy;
wneuper@59586
   173
walther@59867
   174
fun assoc_rls (rls' : Rule_Set.id) =
walther@59879
   175
  case AList.lookup (op =) (KEStore_Elems.get_rlss (ThyC.get_theory "Isac_Knowledge")) rls' of
wneuper@59586
   176
    SOME (_, rls) => rls
walther@59887
   177
  | NONE => raise ERROR ("rls \""^ rls' ^ "\" missing in Know_Store.\n" ^
wneuper@59586
   178
    "TODO exception hierarchy needs to be established.")
wneuper@59586
   179
walther@59867
   180
fun assoc_rls' thy (rls' : Rule_Set.id) =
wneuper@59586
   181
  case AList.lookup (op =) (KEStore_Elems.get_rlss thy) rls' of
wneuper@59586
   182
    SOME (_, rls) => rls
walther@59887
   183
  | NONE => raise ERROR ("rls \""^ rls' ^ "\" missing in Know_Store.\n" ^
wneuper@59586
   184
    "TODO exception hierarchy needs to be established.")
wneuper@59586
   185
wneuper@59586
   186
fun assoc_calc thy calID = let
wneuper@59586
   187
    fun ass ([], key) =
wneuper@59586
   188
          error ("assoc_calc: '" ^ key ^ "' not found in theory " ^ (Context.theory_name thy))
wneuper@59586
   189
      | ass ((calc, (keyi, _)) :: pairs, key) =
wneuper@59586
   190
          if key = keyi then calc else ass (pairs, key);
wneuper@59586
   191
  in ass (thy |> KEStore_Elems.get_calcs, calID) end;
wneuper@59586
   192
wneuper@59586
   193
fun assoc_calc' thy key = let
wneuper@59586
   194
    fun ass ([], key') =
wneuper@59586
   195
          error ("assoc_calc': '" ^ key' ^ "' not found in theory " ^ (Context.theory_name thy))
wneuper@59586
   196
      | ass ((all as (keyi, _)) :: pairs, key') =
wneuper@59586
   197
          if key' = keyi then all else ass (pairs, key');
wneuper@59586
   198
  in ass (KEStore_Elems.get_calcs thy, key) end;
wneuper@59586
   199
wneuper@59586
   200
fun assoc_cas thy key = assoc (KEStore_Elems.get_cas thy, key);
wneuper@59586
   201
wneuper@59586
   202
fun get_ptyps () = get_ref_thy () |> KEStore_Elems.get_ptyps;
wneuper@59586
   203
fun get_mets () = get_ref_thy () |> KEStore_Elems.get_mets;
wneuper@59586
   204
fun get_thes () = get_ref_thy () |> KEStore_Elems.get_thes;
wneuper@59586
   205
\<close>
wneuper@59586
   206
setup \<open>KEStore_Elems.add_rlss 
walther@59852
   207
  [("empty", (Context.theory_name @{theory}, Rule_Set.empty)), 
walther@59850
   208
  ("e_rrls", (Context.theory_name @{theory}, Rule_Set.e_rrls))]\<close>
wneuper@59586
   209
wneuper@59586
   210
section \<open>determine sequence of main parts in thehier\<close>
wneuper@59586
   211
setup \<open>
wneuper@59586
   212
KEStore_Elems.add_thes
walther@59889
   213
  [(Celem8.Html {guh = Celem8.part2guh ["IsacKnowledge"], html = "",
wneuper@59586
   214
    mathauthors = ["Isac team"], coursedesign = []}, ["IsacKnowledge"]),
walther@59889
   215
  (Celem8.Html {guh = Celem8.part2guh ["Isabelle"], html = "",
wneuper@59586
   216
    mathauthors = ["Isabelle team, TU Munich"], coursedesign = []}, ["Isabelle"]),
walther@59889
   217
  (Celem8.Html {guh = Celem8.part2guh ["IsacScripts"], html = "",
wneuper@59586
   218
    mathauthors = ["Isac team"], coursedesign = []}, ["IsacScripts"])]
wneuper@59586
   219
\<close>
wneuper@59586
   220
wneuper@59586
   221
section \<open>Functions for checking KEStore_Elems\<close>
wneuper@59586
   222
ML \<open>
walther@59851
   223
fun short_string_of_rls Rule_Set.Empty = "Erls"
walther@59851
   224
  | short_string_of_rls (Rule_Def.Repeat {calc, rules, ...}) =
wneuper@59586
   225
    "Rls {#calc = " ^ string_of_int (length calc) ^
wneuper@59586
   226
    ", #rules = " ^ string_of_int (length rules) ^ ", ..."
walther@59878
   227
  | short_string_of_rls (Rule_Set.Sequence {calc, rules, ...}) =
wneuper@59586
   228
    "Seq {#calc = " ^ string_of_int (length calc) ^
wneuper@59586
   229
    ", #rules = " ^ string_of_int (length rules) ^ ", ..."
walther@59850
   230
  | short_string_of_rls (Rule_Set.Rrls _) = "Rrls {...}";
wneuper@59586
   231
fun check_kestore_rls (rls', (thyID, rls)) =
wneuper@59586
   232
  "(" ^ rls' ^ ", (" ^ thyID ^ ", " ^ short_string_of_rls rls ^ "))";
wneuper@59586
   233
walther@59852
   234
fun check_kestore_calc ((id, (c, _)) : Rule_Def.calc)  = "(" ^ id ^ ", (" ^ c ^ ", fn))";
wneuper@59586
   235
walther@59846
   236
(* we avoid term_to_string''' defined later *)
walther@59889
   237
fun check_kestore_cas ((t, (s, _)) : Celem7.cas_elem) =
walther@59846
   238
  "(" ^ (Print_Mode.setmp [] (Syntax.string_of_term (Config.put show_markup false
walther@59891
   239
  (Proof_Context.init_global @{theory})))) t ^ ", " ^ Spec.spec2str s ^ ")";
wneuper@59586
   240
wneuper@59586
   241
fun count_kestore_ptyps [] = 0
walther@59890
   242
  | count_kestore_ptyps ((Store.Ptyp (_, _, ps)) :: ps') =
wneuper@59586
   243
      1 + count_kestore_ptyps ps  + count_kestore_ptyps ps';
walther@59890
   244
fun check_kestore_ptyp' strfun (Store.Ptyp (key, pbts, pts)) = "Ptyp (" ^ (quote key) ^ ", " ^
walther@59888
   245
      (strfun pbts) ^ ", " ^ (map (check_kestore_ptyp' strfun) pts |> list2str) ^ ")" |> linefeed;
walther@59889
   246
val check_kestore_ptyp = check_kestore_ptyp' Celem5.pbts2str;
walther@59890
   247
fun ptyp_ord ((Store.Ptyp (s1, _, _)), (Store.Ptyp (s2, _, _))) = string_ord (s1, s2);
walther@59889
   248
fun pbt_ord ({guh = guh'1, ...} : Celem5.pbt, {guh = guh'2, ...} : Celem5.pbt) = string_ord (guh'1, guh'2);
wneuper@59586
   249
fun sort_kestore_ptyp' _ [] = []
walther@59890
   250
  | sort_kestore_ptyp' ordfun ((Store.Ptyp (key, pbts, ps)) :: ps') =
walther@59890
   251
     ((Store.Ptyp (key, sort ordfun pbts, sort_kestore_ptyp' ordfun ps |> sort ptyp_ord))
wneuper@59586
   252
       :: sort_kestore_ptyp' ordfun ps');
wneuper@59586
   253
val sort_kestore_ptyp = sort_kestore_ptyp' pbt_ord;
wneuper@59586
   254
walther@59889
   255
fun metguh2str ({guh,...} : Celem6.met) = guh : string;
walther@59890
   256
fun check_kestore_met (mp: Celem6.met Store.ptyp) =
wneuper@59586
   257
      check_kestore_ptyp' (fn xs => map metguh2str xs |> strs2str) mp;
walther@59889
   258
fun met_ord ({guh = guh'1, ...} : Celem6.met, {guh = guh'2, ...} : Celem6.met) = string_ord (guh'1, guh'2);
wneuper@59586
   259
val sort_kestore_met = sort_kestore_ptyp' met_ord;
wneuper@59586
   260
walther@59889
   261
fun check_kestore_thes thes = ((map writeln) o (map (check_kestore_ptyp' Celem8.thes2str))) thes
wneuper@59586
   262
fun write_thes thydata_list =
wneuper@59586
   263
  thydata_list 
walther@59888
   264
    |> map (fn (id, the) => (Celem8.theID2str id, Celem8.the2str the))
wneuper@59586
   265
    |> map pair2str
wneuper@59586
   266
    |> map writeln
wneuper@59586
   267
\<close>
wneuper@59586
   268
ML \<open>
wneuper@59586
   269
\<close> ML \<open>
wneuper@59586
   270
\<close> ML \<open>
wneuper@59586
   271
\<close>
wneuper@59586
   272
end