src/Tools/isac/CalcElements/KEStore.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Fri, 23 Aug 2019 17:23:25 +0200
changeset 59587 f59488210ffa
parent 59586 5dad05602c23
child 59592 99c8d2ff63eb
permissions -rw-r--r--
remove Delete.thy
wneuper@59586
     1
(*  Title:      src/Tools/isac/KEStore.thy
wneuper@59586
     2
    Author:     Mathias Lehnfeld
wneuper@59586
     3
*)
wneuper@59586
     4
wneuper@59586
     5
theory KEStore imports Complex_Main
wneuper@59586
     6
wneuper@59586
     7
begin
wneuper@59586
     8
ML_file libraryC.sml
wneuper@59586
     9
ML_file rule.sml
wneuper@59586
    10
ML_file calcelems.sml
wneuper@59586
    11
wneuper@59586
    12
ML \<open>
wneuper@59586
    13
\<close> ML \<open>
wneuper@59587
    14
\<close> ML \<open>
wneuper@59586
    15
\<close>
wneuper@59586
    16
section \<open>Knowledge elements for problems and methods\<close>
wneuper@59586
    17
ML \<open>
wneuper@59586
    18
(* Knowledge (and Exercises) are held by "KEStore" in Isac's Java front-end.
wneuper@59586
    19
  In the front-end Knowledge comprises theories, problems and methods.
wneuper@59586
    20
  Elements of problems and methods are defined in theories alongside
wneuper@59586
    21
  the development of respective language elements. 
wneuper@59586
    22
  However, the structure of methods and problems is independent from theories' 
wneuper@59586
    23
  deductive structure. Thus respective structures are built in Build_Thydata.thy.
wneuper@59586
    24
wneuper@59586
    25
  Most elements of problems and methods are implemented in "Knowledge/", but some
wneuper@59586
    26
  of them are implemented in "ProgLang/" already; thus "KEStore.thy" got this
wneuper@59586
    27
  location in the directory structure.
wneuper@59586
    28
wneuper@59586
    29
  get_* retrieves all * of the respective theory PLUS of all ancestor theories.
wneuper@59586
    30
*)
wneuper@59586
    31
signature KESTORE_ELEMS =
wneuper@59586
    32
sig
wneuper@59586
    33
  val get_rlss: theory -> (Rule.rls' * (Rule.theory' * Rule.rls)) list
wneuper@59586
    34
  val add_rlss: (Rule.rls' * (Rule.theory' * Rule.rls)) list -> theory -> theory
wneuper@59586
    35
  val get_calcs: theory -> (Rule.prog_calcID * (Rule.calID * Rule.eval_fn)) list
wneuper@59586
    36
  val add_calcs: (Rule.prog_calcID * (Rule.calID * Rule.eval_fn)) list -> theory -> theory
wneuper@59586
    37
  val get_cas: theory -> Celem.cas_elem list
wneuper@59586
    38
  val add_cas: Celem.cas_elem list -> theory -> theory
wneuper@59586
    39
  val get_ptyps: theory -> Celem.ptyps
wneuper@59586
    40
  val add_pbts: (Celem.pbt * Celem.pblID) list -> theory -> theory
wneuper@59586
    41
  val get_mets: theory -> Celem.mets
wneuper@59586
    42
  val add_mets: (Celem.met * Celem.metID) list -> theory -> theory
wneuper@59586
    43
  val get_thes: theory -> (Celem.thydata Celem.ptyp) list
wneuper@59586
    44
  val add_thes: (Celem.thydata * Celem.theID) list -> theory -> theory (* thydata dropped at existing elems *)
wneuper@59586
    45
  val insert_fillpats: (Celem.theID * Celem.fillpat list) list -> theory -> theory 
wneuper@59586
    46
  val get_ref_thy: unit -> theory
wneuper@59586
    47
  val set_ref_thy: theory -> unit
wneuper@59586
    48
end;                               
wneuper@59586
    49
wneuper@59586
    50
structure KEStore_Elems: KESTORE_ELEMS =
wneuper@59586
    51
struct
wneuper@59586
    52
  fun union_overwrite eq l1 l2 = fold (insert eq) l2 (*..swapped..*) l1;
wneuper@59586
    53
wneuper@59586
    54
  structure Data = Theory_Data (
wneuper@59586
    55
    type T = (Rule.rls' * (Rule.theory' * Rule.rls)) list;
wneuper@59586
    56
    val empty = [];
wneuper@59586
    57
    val extend = I;
wneuper@59586
    58
    val merge = Celem.merge_rlss;
wneuper@59586
    59
    );  
wneuper@59586
    60
  fun get_rlss thy = Data.get thy
wneuper@59586
    61
  fun add_rlss rlss = Data.map (union_overwrite Celem.rls_eq rlss)
wneuper@59586
    62
wneuper@59586
    63
  structure Data = Theory_Data (
wneuper@59586
    64
    type T = (Rule.prog_calcID * (Rule.calID * Rule.eval_fn)) list;
wneuper@59586
    65
    val empty = [];
wneuper@59586
    66
    val extend = I;
wneuper@59586
    67
    val merge = merge Rule.calc_eq;
wneuper@59586
    68
    );                                                              
wneuper@59586
    69
  fun get_calcs thy = Data.get thy
wneuper@59586
    70
  fun add_calcs calcs = Data.map (union_overwrite Rule.calc_eq calcs)
wneuper@59586
    71
wneuper@59586
    72
  structure Data = Theory_Data (
wneuper@59586
    73
    type T = (term * (Celem.spec * (term list -> (term * term list) list))) list;
wneuper@59586
    74
    val empty = [];
wneuper@59586
    75
    val extend = I;
wneuper@59586
    76
    val merge = merge Celem.cas_eq;
wneuper@59586
    77
    );                                                              
wneuper@59586
    78
  fun get_cas thy = Data.get thy
wneuper@59586
    79
  fun add_cas cas = Data.map (union_overwrite Celem.cas_eq cas)
wneuper@59586
    80
wneuper@59586
    81
  structure Data = Theory_Data (
wneuper@59586
    82
    type T = Celem.ptyps;
wneuper@59586
    83
    val empty = [Celem.e_Ptyp];
wneuper@59586
    84
    val extend = I;
wneuper@59586
    85
    val merge = Celem.merge_ptyps;
wneuper@59586
    86
    );
wneuper@59586
    87
  fun get_ptyps thy = Data.get thy;
wneuper@59586
    88
  fun add_pbts pbts thy = let
wneuper@59586
    89
          fun add_pbt (pbt as {guh,...}, pblID) =
wneuper@59586
    90
                (* the pblID has the leaf-element as first; better readability achieved *)
wneuper@59586
    91
                (if (!Celem.check_guhs_unique) then Celem.check_pblguh_unique guh (Data.get thy) else ();
wneuper@59586
    92
                  rev pblID |> Celem.insrt pblID pbt);
wneuper@59586
    93
        in Data.map (fold add_pbt pbts) thy end;
wneuper@59586
    94
wneuper@59586
    95
  structure Data = Theory_Data (
wneuper@59586
    96
    type T = Celem.mets;
wneuper@59586
    97
    val empty = [Celem.e_Mets];
wneuper@59586
    98
    val extend = I;
wneuper@59586
    99
    val merge = Celem.merge_ptyps;
wneuper@59586
   100
    );
wneuper@59586
   101
  val get_mets = Data.get;
wneuper@59586
   102
  fun add_mets mets thy = let
wneuper@59586
   103
          fun add_met (met as {guh,...}, metID) =
wneuper@59586
   104
                (if (!Celem.check_guhs_unique) then Celem.check_metguh_unique guh (Data.get thy) else ();
wneuper@59586
   105
                  Celem.insrt metID met metID);
wneuper@59586
   106
        in Data.map (fold add_met mets) thy end;
wneuper@59586
   107
wneuper@59586
   108
  structure Data = Theory_Data (
wneuper@59586
   109
    type T = (Celem.thydata Celem.ptyp) list;
wneuper@59586
   110
    val empty = [];
wneuper@59586
   111
    val extend = I;
wneuper@59586
   112
    val merge = Celem.merge_ptyps; (* relevant for store_thm, store_rls *)
wneuper@59586
   113
    );                                                              
wneuper@59586
   114
  fun get_thes thy = Data.get thy
wneuper@59586
   115
  fun add_thes thes thy = let
wneuper@59586
   116
    fun add_the (thydata, theID) = Celem.add_thydata ([], theID) thydata
wneuper@59586
   117
  in Data.map (fold add_the thes) thy end;
wneuper@59586
   118
  fun insert_fillpats fis thy =
wneuper@59586
   119
    let
wneuper@59586
   120
      fun update_elem (theID, fillpats) =
wneuper@59586
   121
        let
wneuper@59586
   122
          val hthm = Celem.get_py (Data.get thy) theID theID
wneuper@59586
   123
          val hthm' = Celem.update_hthm hthm fillpats
wneuper@59586
   124
            handle ERROR _ =>
wneuper@59586
   125
              error ("insert_fillpats: " ^ strs2str theID ^ "must address a theorem")
wneuper@59586
   126
        in Celem.update_ptyps theID theID hthm' end
wneuper@59586
   127
    in Data.map (fold update_elem fis) thy end
wneuper@59586
   128
wneuper@59586
   129
  val cur_thy = Synchronized.var "finally_knowledge_complete" @{theory};
wneuper@59586
   130
  fun set_ref_thy thy = Synchronized.change cur_thy (fn _ => thy); (* never RE-set ! *)
wneuper@59586
   131
  fun get_ref_thy () = Synchronized.value cur_thy;
wneuper@59586
   132
end;
wneuper@59586
   133
\<close>
wneuper@59586
   134
wneuper@59586
   135
section \<open>Re-use existing access functions for knowledge elements\<close>
wneuper@59586
   136
text \<open>
wneuper@59586
   137
  The independence of problems' and methods' structure enforces the accesse
wneuper@59586
   138
  functions to use "Isac", the final theory which comprises all knowledge defined.
wneuper@59586
   139
\<close>
wneuper@59586
   140
ML \<open>
wneuper@59586
   141
val get_ref_thy = KEStore_Elems.get_ref_thy;
wneuper@59586
   142
wneuper@59586
   143
fun assoc_rls (rls' : Rule.rls') =
wneuper@59586
   144
  case AList.lookup (op =) (KEStore_Elems.get_rlss (Rule.Thy_Info_get_theory "Isac")) rls' of
wneuper@59586
   145
    SOME (_, rls) => rls
wneuper@59586
   146
  | NONE => raise ERROR ("rls \""^ rls' ^ "\" missing in KEStore.\n" ^
wneuper@59586
   147
    "TODO exception hierarchy needs to be established.")
wneuper@59586
   148
wneuper@59586
   149
fun assoc_rls' thy (rls' : Rule.rls') =
wneuper@59586
   150
  case AList.lookup (op =) (KEStore_Elems.get_rlss thy) rls' of
wneuper@59586
   151
    SOME (_, rls) => rls
wneuper@59586
   152
  | NONE => raise ERROR ("rls \""^ rls' ^ "\" missing in KEStore.\n" ^
wneuper@59586
   153
    "TODO exception hierarchy needs to be established.")
wneuper@59586
   154
wneuper@59586
   155
fun assoc_calc thy calID = let
wneuper@59586
   156
    fun ass ([], key) =
wneuper@59586
   157
          error ("assoc_calc: '" ^ key ^ "' not found in theory " ^ (Context.theory_name thy))
wneuper@59586
   158
      | ass ((calc, (keyi, _)) :: pairs, key) =
wneuper@59586
   159
          if key = keyi then calc else ass (pairs, key);
wneuper@59586
   160
  in ass (thy |> KEStore_Elems.get_calcs, calID) end;
wneuper@59586
   161
wneuper@59586
   162
fun assoc_calc' thy key = let
wneuper@59586
   163
    fun ass ([], key') =
wneuper@59586
   164
          error ("assoc_calc': '" ^ key' ^ "' not found in theory " ^ (Context.theory_name thy))
wneuper@59586
   165
      | ass ((all as (keyi, _)) :: pairs, key') =
wneuper@59586
   166
          if key' = keyi then all else ass (pairs, key');
wneuper@59586
   167
  in ass (KEStore_Elems.get_calcs thy, key) end;
wneuper@59586
   168
wneuper@59586
   169
fun assoc_cas thy key = assoc (KEStore_Elems.get_cas thy, key);
wneuper@59586
   170
wneuper@59586
   171
fun get_ptyps () = get_ref_thy () |> KEStore_Elems.get_ptyps;
wneuper@59586
   172
fun get_mets () = get_ref_thy () |> KEStore_Elems.get_mets;
wneuper@59586
   173
fun get_thes () = get_ref_thy () |> KEStore_Elems.get_thes;
wneuper@59586
   174
\<close>
wneuper@59586
   175
setup \<open>KEStore_Elems.add_rlss 
wneuper@59586
   176
  [("e_rls", (Context.theory_name @{theory}, Rule.e_rls)), 
wneuper@59586
   177
  ("e_rrls", (Context.theory_name @{theory}, Rule.e_rrls))]\<close>
wneuper@59586
   178
wneuper@59586
   179
section \<open>determine sequence of main parts in thehier\<close>
wneuper@59586
   180
setup \<open>
wneuper@59586
   181
KEStore_Elems.add_thes
wneuper@59586
   182
  [(Celem.Html {guh = Celem.part2guh ["IsacKnowledge"], html = "",
wneuper@59586
   183
    mathauthors = ["Isac team"], coursedesign = []}, ["IsacKnowledge"]),
wneuper@59586
   184
  (Celem.Html {guh = Celem.part2guh ["Isabelle"], html = "",
wneuper@59586
   185
    mathauthors = ["Isabelle team, TU Munich"], coursedesign = []}, ["Isabelle"]),
wneuper@59586
   186
  (Celem.Html {guh = Celem.part2guh ["IsacScripts"], html = "",
wneuper@59586
   187
    mathauthors = ["Isac team"], coursedesign = []}, ["IsacScripts"])]
wneuper@59586
   188
\<close>
wneuper@59586
   189
wneuper@59586
   190
section \<open>Functions for checking KEStore_Elems\<close>
wneuper@59586
   191
ML \<open>
wneuper@59586
   192
fun short_string_of_rls Rule.Erls = "Erls"
wneuper@59586
   193
  | short_string_of_rls (Rule.Rls {calc, rules, ...}) =
wneuper@59586
   194
    "Rls {#calc = " ^ string_of_int (length calc) ^
wneuper@59586
   195
    ", #rules = " ^ string_of_int (length rules) ^ ", ..."
wneuper@59586
   196
  | short_string_of_rls (Rule.Seq {calc, rules, ...}) =
wneuper@59586
   197
    "Seq {#calc = " ^ string_of_int (length calc) ^
wneuper@59586
   198
    ", #rules = " ^ string_of_int (length rules) ^ ", ..."
wneuper@59586
   199
  | short_string_of_rls (Rule.Rrls _) = "Rrls {...}";
wneuper@59586
   200
fun check_kestore_rls (rls', (thyID, rls)) =
wneuper@59586
   201
  "(" ^ rls' ^ ", (" ^ thyID ^ ", " ^ short_string_of_rls rls ^ "))";
wneuper@59586
   202
wneuper@59586
   203
fun check_kestore_calc ((id, (c, _)) : Rule.calc)  = "(" ^ id ^ ", (" ^ c ^ ", fn))";
wneuper@59586
   204
wneuper@59586
   205
fun check_kestore_cas ((t, (s, _)) : Celem.cas_elem) =
wneuper@59586
   206
  "(" ^ (Rule.term_to_string''' @{theory} t) ^ ", " ^ (Celem.spec2str s) ^ ")";
wneuper@59586
   207
wneuper@59586
   208
fun count_kestore_ptyps [] = 0
wneuper@59586
   209
  | count_kestore_ptyps ((Celem.Ptyp (_, _, ps)) :: ps') =
wneuper@59586
   210
      1 + count_kestore_ptyps ps  + count_kestore_ptyps ps';
wneuper@59586
   211
fun check_kestore_ptyp' strfun (Celem.Ptyp (key, pbts, pts)) = "Ptyp (" ^ (quote key) ^ ", " ^
wneuper@59586
   212
      (strfun pbts) ^ ", " ^ (map (check_kestore_ptyp' strfun) pts |> list2str) ^ ")" |> Celem.linefeed;
wneuper@59586
   213
val check_kestore_ptyp = check_kestore_ptyp' Celem.pbts2str;
wneuper@59586
   214
fun ptyp_ord ((Celem.Ptyp (s1, _, _)), (Celem.Ptyp (s2, _, _))) = string_ord (s1, s2);
wneuper@59586
   215
fun pbt_ord ({guh = guh'1, ...} : Celem.pbt, {guh = guh'2, ...} : Celem.pbt) = string_ord (guh'1, guh'2);
wneuper@59586
   216
fun sort_kestore_ptyp' _ [] = []
wneuper@59586
   217
  | sort_kestore_ptyp' ordfun ((Celem.Ptyp (key, pbts, ps)) :: ps') =
wneuper@59586
   218
     ((Celem.Ptyp (key, sort ordfun pbts, sort_kestore_ptyp' ordfun ps |> sort ptyp_ord))
wneuper@59586
   219
       :: sort_kestore_ptyp' ordfun ps');
wneuper@59586
   220
val sort_kestore_ptyp = sort_kestore_ptyp' pbt_ord;
wneuper@59586
   221
wneuper@59586
   222
fun metguh2str ({guh,...} : Celem.met) = guh : string;
wneuper@59586
   223
fun check_kestore_met (mp: Celem.met Celem.ptyp) =
wneuper@59586
   224
      check_kestore_ptyp' (fn xs => map metguh2str xs |> strs2str) mp;
wneuper@59586
   225
fun met_ord ({guh = guh'1, ...} : Celem.met, {guh = guh'2, ...} : Celem.met) = string_ord (guh'1, guh'2);
wneuper@59586
   226
val sort_kestore_met = sort_kestore_ptyp' met_ord;
wneuper@59586
   227
wneuper@59586
   228
fun check_kestore_thes thes = ((map writeln) o (map (check_kestore_ptyp' Celem.thes2str))) thes
wneuper@59586
   229
fun write_thes thydata_list =
wneuper@59586
   230
  thydata_list 
wneuper@59586
   231
    |> map (fn (id, the) => (Celem.theID2str id, Celem.the2str the))
wneuper@59586
   232
    |> map pair2str
wneuper@59586
   233
    |> map writeln
wneuper@59586
   234
\<close>
wneuper@59586
   235
ML \<open>
wneuper@59586
   236
\<close> ML \<open>
wneuper@59586
   237
\<close> ML \<open>
wneuper@59586
   238
\<close>
wneuper@59586
   239
end