src/Tools/isac/BaseDefinitions/Know_Store.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 15 Sep 2022 10:07:12 +0200
changeset 60554 210594dd8a5a
parent 60551 3b5be6fae2f0
child 60555 466bcb20f2d7
permissions -rw-r--r--
follow up meeting Makarius 1: new concept of parsing within current ?ML_structure Context?
     1 (*  Title:      src/Tools/isac/Know_Store.thy
     2     Author:     Mathias Lehnfeld
     3 
     4 Store of Theory_Data of all knowledge required by the Lucas-Interpreter.
     5 
     6 Types of elements are defined in "xxxxx-def.sml". These files have companion files "xxxxx.sml" 
     7 with all further code, located at appropriate positions in the file structure.
     8 (The separation of "xxxxx-def.sml" from "xxxxx.sml" should be overcome by
     9 appropriate use of polymorphic high order functions.)
    10 
    11 Notable are Problem.T and MethodC.T; these are trees with a structure different from
    12 Isabelle's theories dependencies.
    13 *)
    14 
    15 theory Know_Store
    16   imports Complex_Main
    17   keywords "rule_set_knowledge" "calculation" :: thy_decl
    18 begin
    19 
    20 setup \<open>
    21   ML_Antiquotation.conditional \<^binding>\<open>isac_test\<close>
    22     (fn _ => Options.default_bool \<^system_option>\<open>isac_test\<close>)
    23 \<close>
    24 
    25 ML_file libraryC.sml
    26 ML_file theoryC.sml
    27 ML_file unparseC.sml
    28 ML_file "rule-def.sml"
    29 ML_file "thmC-def.sml"
    30 ML_file "eval-def.sml"       (*rename identifiers by use of struct.id*)
    31 ML_file "rewrite-order.sml"
    32 ML_file rule.sml
    33 ML_file "error-pattern-def.sml"
    34 ML_file "rule-set.sml"
    35 
    36 ML_file "store.sml"
    37 ML_file "check-unique.sml"
    38 ML_file "references-def.sml"
    39 ML_file "model-pattern.sml"
    40 ML_file "problem-def.sml"
    41 ML_file "method-def.sml"
    42 ML_file "cas-def.sml"
    43 ML_file "formalise.sml"
    44 ML_file "example.sml"
    45 ML_file "thy-write.sml"
    46 
    47 ML \<open>
    48 \<close> ML \<open>
    49 \<close> ML \<open>
    50 \<close>
    51 section \<open>Knowledge elements for problems and methods\<close>
    52 text \<open>
    53   \<open>ML_structure Problem\<close>, \<open>ML_structure MethodC\<close> and \<open>Example\<close>s are held by "Know_Store".
    54 
    55   The structure of \<open>ML_structure Problem\<close> and \<open>ML_structure MethodC\<close> is independent from 
    56   theories' dependency graph. Thus the respective elements are stored as \<open>TermC.as_string\<close>
    57   and parsed on the fly within the current @{ML_structure Context},
    58   while being loaded into \<open>ML_structure Calc\<close>.
    59 
    60   Note: From the children of eb89f586b0b2 onwards the old functions (\<open>term typ_a2real\<close> etc)
    61   are used for "parsing on the fly" until further clarification.
    62 \<open>\<close>
    63   Most elements of \<open>ML_structure Problem\<close> and \<open>ML_structure MethodC\<close> are implemented in 
    64   \<open>directory Knowledge/\<close> but some of them are implemented in \<open>directory ProgLang/\<close>already; 
    65   thus \<open>theory Know_Store\<close> got this location in the directory structure.
    66 
    67   \<open>term Know_Store.get_*\<close> retrieves all * of the respective theory PLUS of all ancestor theories.
    68 \<close> ML \<open>
    69 signature KESTORE_ELEMS(*TODO rename to KNOW_STORE*) =
    70 sig
    71   val get_rew_ords: theory -> Rewrite_Ord.T list
    72   val add_rew_ords: Rewrite_Ord.T list -> theory -> theory
    73   val get_rlss: theory -> (Rule_Set.id * (ThyC.id * Rule_Set.T)) list
    74   val add_rlss: (Rule_Set.id * (ThyC.id * Rule_Set.T)) list -> theory -> theory
    75   val get_calcs: theory -> (Eval_Def.prog_id * (Eval_Def.const_id * Eval_Def.ml_fun)) list
    76   val add_calcs: (Eval_Def.prog_id * (Eval_Def.const_id * Eval_Def.ml_fun)) list -> theory -> theory
    77   val get_cass: theory -> CAS_Def.T list
    78   val add_cass: CAS_Def.T list -> theory -> theory
    79   val get_pbls: theory -> Probl_Def.store
    80   val add_pbls: Proof.context -> (Probl_Def.T * References_Def.id) list -> theory -> theory
    81   val get_mets: theory -> Meth_Def.store
    82   val add_mets: Proof.context -> (Meth_Def.T * References_Def.id) list -> theory -> theory
    83   val get_expls: theory -> Example.store
    84   val add_expls: (Example.T * Store.key) list -> theory -> theory
    85   val get_thes: theory -> (Thy_Write.thydata Store.node) list
    86   val add_thes: (Thy_Write.thydata * Thy_Write.theID) list -> theory -> theory (* thydata dropped at existing elems *)
    87   val insert_fillpats: (Thy_Write.theID * Error_Pattern_Def.fill_in list) list -> theory -> theory 
    88   val get_ref_thy: unit -> theory
    89   val set_ref_thy: theory -> unit
    90 end;
    91 
    92 structure KEStore_Elems(*(*TODO rename to Know_Store*)*): KESTORE_ELEMS(**) =
    93 struct
    94   structure Data = Theory_Data (
    95     type T = Rewrite_Ord.T list;
    96     val empty = [];
    97     val extend = I;
    98     val merge = merge Rewrite_Ord.equal;
    99     );  
   100   fun get_rew_ords thy = Data.get thy
   101   fun add_rew_ords rlss = Data.map (curry (Library.merge Rewrite_Ord.equal) rlss)
   102 
   103   structure Data = Theory_Data (
   104     type T = (Rule_Set.id * (ThyC.id * Rule_Set.T)) list;
   105                             (* ^^^^^ would allow same Rls_Set.id for different thys, NOT impl. *)
   106     val empty = [];
   107     val extend = I;
   108     val merge = Rule_Set.to_kestore;
   109     );  
   110   fun get_rlss thy = Data.get thy
   111   fun add_rlss rlss = Data.map (curry (Library.merge Rule_Set.equal) rlss)
   112 
   113   structure Data = Theory_Data (
   114     type T = (Eval_Def.prog_id * (Eval_Def.const_id * Eval_Def.ml_fun)) list;
   115     val empty = [];
   116     val extend = I;
   117     val merge = merge Eval_Def.equal;
   118     );                                                              
   119   fun get_calcs thy = Data.get thy
   120   fun add_calcs calcs = Data.map (curry (Library.merge Eval_Def.equal) calcs)
   121 
   122   structure Data = Theory_Data (
   123     type T = (term * (References_Def.T * CAS_Def.generate_fn)) list;
   124     val empty = [];
   125     val extend = I;
   126     val merge = merge CAS_Def.equal;
   127     );                                                              
   128   fun get_cass thy = Data.get thy
   129   fun add_cass cas = Data.map (curry (Library.merge CAS_Def.equal) cas)
   130 
   131   structure Data = Theory_Data (
   132     type T = Probl_Def.store;
   133     val empty = [Probl_Def.empty_store];
   134     val extend = I;
   135     val merge = Store.merge;
   136     );
   137   fun get_pbls thy = Data.get thy;
   138   fun add_pbls ctxt pbts thy =
   139     let
   140       fun add_pbt (pbt as {guh,...}, pblID) =
   141         (* the pblID has the leaf-element as first; better readability achieved *)
   142         (if (Config.get ctxt check_unique) then Probl_Def.check_unique guh (Data.get thy) else ();
   143           rev pblID |> Store.insert pblID pbt);
   144     in Data.map (fold add_pbt pbts) thy end;
   145 
   146   structure Data = Theory_Data (
   147     type T = Meth_Def.store;
   148     val empty = [Meth_Def.empty_store];
   149     val extend = I;
   150     val merge = Store.merge;
   151     );
   152   val get_mets = Data.get;
   153   fun add_mets ctxt mets thy =
   154     let
   155       fun add_met (met as {guh,...}, metID) =
   156         (if (Config.get ctxt check_unique) then Meth_Def.check_unique guh (Data.get thy) else ();
   157           Store.insert metID met metID);
   158     in Data.map (fold add_met mets) thy end;
   159 
   160   structure Data = Theory_Data (
   161     type T = Example.store;
   162     val empty = [Example.empty_store];
   163     val extend = I;
   164     val merge = Store.merge;
   165     );
   166   val get_expls = Data.get;
   167   fun add_expls expls thy =
   168     let
   169       fun add_expl (expl, expl_id) = Store.insert expl_id expl expl_id;
   170     in Data.map (fold add_expl expls) thy end;
   171 
   172   structure Data = Theory_Data (
   173     type T = (Thy_Write.thydata Store.node) list;
   174     val empty = [];
   175     val extend = I;
   176     val merge = Store.merge; (* relevant for store_thm, store_rls *)
   177     );                                                              
   178   fun get_thes thy = Data.get thy
   179   fun add_thes thes thy = let
   180     fun add_the (thydata, theID) = Thy_Write.add_thydata ([], theID) thydata
   181   in Data.map (fold add_the thes) thy end;
   182   fun insert_fillpats fis thy =
   183     let
   184       fun update_elem (theID, fillpats) =
   185         let
   186           val hthm = Store.get (Data.get thy) theID theID
   187           val hthm' = Thy_Write.update_hthm hthm fillpats
   188             handle ERROR _ =>
   189               raise ERROR ("insert_fillpats: " ^ strs2str theID ^ "must address a theorem")
   190         in Store.update theID theID hthm' end
   191     in Data.map (fold update_elem fis) thy end
   192 
   193   val cur_thy = Synchronized.var "finally_knowledge_complete" @{theory};
   194   fun set_ref_thy thy = Synchronized.change cur_thy (fn _ => thy); (* never RE-set ! *)
   195   fun get_ref_thy () = Synchronized.value cur_thy;
   196 
   197 (**)end(*struct*);
   198 \<close>
   199 
   200 
   201 subsection \<open>Isar command syntax\<close>
   202 
   203 ML \<open>
   204 local
   205 
   206 val parse_rule = Parse.name -- Parse.!!! (\<^keyword>\<open>=\<close> |-- Parse.ML_source);
   207 
   208 val ml = ML_Lex.read;
   209 
   210 fun ml_rule thy (name, source) =
   211   ml "(" @ ml (ML_Syntax.print_string name) @ ml ", " @
   212   ml "(" @ ml (ML_Syntax.print_string (Context.theory_name thy)) @ ml ", " @
   213   ML_Lex.read_source source @ ml "))";
   214 
   215 fun ml_rules thy args =
   216   ml "Theory.setup (KEStore_Elems.add_rlss [" @
   217     flat (separate (ml ",") (map (ml_rule thy) args)) @ ml "])";
   218 
   219 val _ =
   220   Outer_Syntax.command \<^command_keyword>\<open>rule_set_knowledge\<close> "register ISAC rule set to Knowledge Store"
   221     (Parse.and_list1 parse_rule >> (fn args => Toplevel.theory (fn thy =>
   222       thy |> Context.theory_map
   223         (ML_Context.expression (Position.thread_data ()) (ml_rules thy args)))));
   224 
   225 val calc_name =
   226   Parse.name -- (\<^keyword>\<open>(\<close> |-- Parse.const --| \<^keyword>\<open>)\<close>) ||
   227   Scan.ahead Parse.name -- Parse.const;
   228 
   229 val _ =
   230   Outer_Syntax.command \<^command_keyword>\<open>calculation\<close>
   231     "prepare ISAC calculation and register it to Knowledge Store"
   232     (calc_name -- (\<^keyword>\<open>=\<close> |-- Parse.!!! Parse.ML_source) >> (fn ((calcID, const), source) =>
   233       Toplevel.theory (fn thy =>
   234         let
   235           val ctxt = Proof_Context.init_global thy;
   236           val Const (c, _) = Proof_Context.read_const {proper = true, strict = true} ctxt const;
   237           val set_data =
   238             ML_Context.expression (Input.pos_of source)
   239               (ml "Theory.setup (Eval_Def.ml_fun_to_store (" @ ML_Lex.read_source source @ ml "))")
   240             |> Context.theory_map;
   241           val eval = Eval_Def.ml_fun_from_store (set_data thy);
   242         in KEStore_Elems.add_calcs [(calcID, (c, eval))] thy end)))
   243 
   244 in end;
   245 \<close>
   246 
   247 
   248 section \<open>Re-use existing access functions for knowledge elements\<close>
   249 text \<open>
   250   The independence of problems' and methods' structure from theory dependency structure
   251   enforces the access functions to use "Isac_Knowledge",
   252   the final theory which comprises all knowledge defined.
   253 \<close>
   254 ML \<open>
   255 val get_ref_thy = KEStore_Elems.get_ref_thy;
   256 
   257 (*val get_rew_ord: Proof.context -> string -> Rewrite_Ord.function*)
   258 fun get_rew_ord ctxt (id: Rewrite_Ord.id) = 
   259   case AList.lookup (op =) (KEStore_Elems.get_rew_ords (Proof_Context.theory_of ctxt)) id of
   260     SOME function => function: Rewrite_Ord.function
   261   | NONE => raise ERROR ("rewrite-order \"" ^ id ^ "\" missing in theory \"" ^ 
   262     (ctxt |> Proof_Context.theory_of |> Context.theory_name) ^ "\" (and ancestors)" ^
   263     "\nTODO exception hierarchy needs to be established.")
   264 
   265 (*val get_rls: Proof.context -> Rule_Set.id -> Rule_Def.rule_set*)
   266 fun get_rls ctxt (id : Rule_Set.id) =
   267   case AList.lookup (op =) (KEStore_Elems.get_rlss (Proof_Context.theory_of ctxt)) id of
   268     SOME (_, rls) => rls
   269   | NONE => raise ERROR ("rls \"" ^ id ^ "\" missing in theory \"" ^ 
   270     (ctxt |> Proof_Context.theory_of |> Context.theory_name) ^ "\" (and ancestors)" ^
   271     "\nTODO exception hierarchy needs to be established.")
   272 
   273 (*val get_calc: Proof.context -> Eval_Def.prog_id -> 
   274   Eval_Def.prog_id * (Eval_Def.const_id * Eval_Def.ml_fun)*)
   275 fun get_calc ctxt (id: Eval_Def.prog_id) = 
   276   case AList.lookup (op =) (KEStore_Elems.get_calcs (Proof_Context.theory_of ctxt)) id of
   277     SOME const_id__ml_fun => (id, const_id__ml_fun)
   278   | NONE => raise ERROR ("ml-calculation \"" ^ id ^ "\" missing in theory \"" ^ 
   279     (ctxt |> Proof_Context.theory_of |> Context.theory_name) ^ "\" (and ancestors)" ^
   280     "\nTODO exception hierarchy needs to be established.")
   281 
   282 (*val get_calc_prog_id: Proof.context -> Eval_Def.const_id -> Eval_Def.prog_id*)
   283 fun get_calc_prog_id ctxt (const_id: Eval_Def.const_id) =
   284   let
   285     fun assoc ([], prog_id) =
   286         raise ERROR ("ml-calculation \"" ^ prog_id ^ "\" missing in theory \"" ^ 
   287           (ctxt |> Proof_Context.theory_of |> Context.theory_name) ^ "\" (and ancestors)." ^
   288           "\nThus " ^ quote const_id  ^ " cannot be retrieved." ^
   289           "\nTODO exception hierarchy needs to be established.")
   290       | assoc ((prog_id, (const_id, _)) :: pairs, key) =
   291           if key = const_id then prog_id else assoc (pairs, key);
   292   in assoc (ctxt |> Proof_Context.theory_of |> KEStore_Elems.get_calcs, const_id) end;
   293 
   294 (*val get_cas: Proof.context -> term -> References_Def.T * CAS_Def.generate_fn*)
   295 fun get_cas ctxt tm = 
   296   case AList.lookup (op =) (KEStore_Elems.get_cass (Proof_Context.theory_of ctxt)) tm of
   297     SOME refs__gen_fun => refs__gen_fun: References_Def.T * CAS_Def.generate_fn
   298   | NONE => raise ERROR ("CAS_Cmd \"" ^ 
   299     UnparseC.term_in_thy (get_ref_thy ()) tm ^ "\" missing in theory \"" ^ 
   300     (ctxt |> Proof_Context.theory_of |> Context.theory_name) ^ "\" (and ancestors)" ^
   301     "\nTODO exception hierarchy needs to be established.")
   302 
   303 (*for starting an Exmaple by CAS_Cmd*)
   304 (*val get_cas_global: term -> References_Def.T * CAS_Def.generate_fn*)
   305 fun get_cas_global tm =
   306   let
   307     val thy = get_ref_thy ()
   308   in
   309     case AList.lookup (op =) (KEStore_Elems.get_cass thy) tm of
   310       NONE => (writeln ("CAS_Cmd \"" ^ 
   311           UnparseC.term_in_thy thy tm ^ "\" missing in theory \"" ^ 
   312           (thy |> Context.theory_name) ^ "\" (and ancestors).");
   313         NONE)
   314     | SOME refs__gen_fun => SOME refs__gen_fun
   315   end
   316 
   317 
   318 fun get_pbls () = get_ref_thy () |> KEStore_Elems.get_pbls;
   319 fun get_mets () = get_ref_thy () |> KEStore_Elems.get_mets;
   320 fun get_thes () = get_ref_thy () |> KEStore_Elems.get_thes;
   321 fun get_expls () = get_ref_thy () |> KEStore_Elems.get_expls;
   322 \<close>
   323 
   324 rule_set_knowledge
   325   empty = \<open>Rule_Set.empty\<close> and
   326   e_rrls = \<open>Rule_Set.e_rrls\<close>
   327 
   328 section \<open>determine sequence of main parts in thehier\<close>
   329 setup \<open>
   330 KEStore_Elems.add_thes
   331   [(Thy_Write.Html {guh = Thy_Write.part2guh ["IsacKnowledge"], html = "",
   332     mathauthors = ["Isac team"], coursedesign = []}, ["IsacKnowledge"]),
   333   (Thy_Write.Html {guh = Thy_Write.part2guh ["Isabelle"], html = "",
   334     mathauthors = ["Isabelle team, TU Munich"], coursedesign = []}, ["Isabelle"]),
   335   (Thy_Write.Html {guh = Thy_Write.part2guh ["IsacScripts"], html = "",
   336     mathauthors = ["Isac team"], coursedesign = []}, ["IsacScripts"])]
   337 \<close>
   338 
   339 section \<open>Functions for checking KEStore_Elems\<close>
   340 ML \<open>
   341 fun short_string_of_rls Rule_Set.Empty = "Erls"
   342   | short_string_of_rls (Rule_Def.Repeat {calc, rules, ...}) =
   343     "Rls {#calc = " ^ string_of_int (length calc) ^
   344     ", #rules = " ^ string_of_int (length rules) ^ ", ..."
   345   | short_string_of_rls (Rule_Set.Sequence {calc, rules, ...}) =
   346     "Seq {#calc = " ^ string_of_int (length calc) ^
   347     ", #rules = " ^ string_of_int (length rules) ^ ", ..."
   348   | short_string_of_rls (Rule_Set.Rrls _) = "Rrls {...}";
   349 fun check_kestore_rls (rls', (thyID, rls)) =
   350   "(" ^ rls' ^ ", (" ^ thyID ^ ", " ^ short_string_of_rls rls ^ "))";
   351 
   352 fun check_kestore_calc ((id, (c, _)) : Rule_Def.eval_ml_from_prog)  = "(" ^ id ^ ", (" ^ c ^ ", fn))";
   353 
   354 (* we avoid term_to_string''' defined later *)
   355 fun check_kestore_cas ((t, (s, _)) : CAS_Def.T) =
   356   "(" ^ (Print_Mode.setmp [] (Syntax.string_of_term (Config.put show_markup false
   357   (Proof_Context.init_global @{theory})))) t ^ ", " ^ References_Def.to_string s ^ ")";
   358 
   359 fun count_kestore_ptyps [] = 0
   360   | count_kestore_ptyps ((Store.Node (_, _, ps)) :: ps') =
   361       1 + count_kestore_ptyps ps  + count_kestore_ptyps ps';
   362 fun check_kestore_ptyp' strfun (Store.Node (key, pbts, pts)) = "Ptyp (" ^ (quote key) ^ ", " ^
   363       (strfun pbts) ^ ", " ^ (map (check_kestore_ptyp' strfun) pts |> list2str) ^ ")" |> linefeed;
   364 val check_kestore_ptyp = check_kestore_ptyp' Probl_Def.s_to_string;
   365 fun ptyp_ord ((Store.Node (s1, _, _)), (Store.Node (s2, _, _))) = string_ord (s1, s2);
   366 fun pbt_ord ({guh = guh'1, ...} : Probl_Def.T, {guh = guh'2, ...} : Probl_Def.T) = string_ord (guh'1, guh'2);
   367 fun sort_kestore_ptyp' _ [] = []
   368   | sort_kestore_ptyp' ordfun ((Store.Node (key, pbts, ps)) :: ps') =
   369      ((Store.Node (key, sort ordfun pbts, sort_kestore_ptyp' ordfun ps |> sort ptyp_ord))
   370        :: sort_kestore_ptyp' ordfun ps');
   371 val sort_kestore_ptyp = sort_kestore_ptyp' pbt_ord;
   372 
   373 fun metguh2str ({guh,...} : Meth_Def.T) = guh : string;
   374 fun check_kestore_met (mp: Meth_Def.T Store.node) =
   375       check_kestore_ptyp' (fn xs => map metguh2str xs |> strs2str) mp;
   376 fun met_ord ({guh = guh'1, ...} : Meth_Def.T, {guh = guh'2, ...} : Meth_Def.T) = string_ord (guh'1, guh'2);
   377 val sort_kestore_met = sort_kestore_ptyp' met_ord;
   378 
   379 fun check_kestore_thes thes = ((map writeln) o (map (check_kestore_ptyp' Thy_Write.thes2str))) thes
   380 fun write_thes thydata_list =
   381   thydata_list 
   382     |> map (fn (id, the) => (Thy_Write.theID2str id, Thy_Write.the2str the))
   383     |> map pair2str
   384     |> map writeln
   385 \<close>
   386 ML \<open>
   387 \<close> ML \<open>
   388 \<close> ML \<open>
   389 \<close>
   390 end