src/Tools/isac/BaseDefinitions/Know_Store.thy
author Walther Neuper <walther.neuper@jku.at>
Thu, 14 May 2020 09:30:40 +0200
changeset 59976 950922a768ca
parent 59962 6a59d252345d
child 59985 9aaeab7d38b6
permissions -rw-r--r--
start renaming Specification -> References;

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