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