added Isar/net_rules.ML;
authorwenzelm
Sun, 27 Feb 2000 15:07:53 +0100
changeset 8298 9b089bc07f69
parent 8297 f5fdb69ad4d2
child 8299 52d9f64841d6
added Isar/net_rules.ML;
src/Pure/IsaMakefile
src/Pure/Isar/ROOT.ML
src/Pure/Isar/net_rules.ML
--- a/src/Pure/IsaMakefile	Thu Feb 24 21:33:32 2000 +0100
+++ b/src/Pure/IsaMakefile	Sun Feb 27 15:07:53 2000 +0100
@@ -23,30 +23,31 @@
 
 Pure: $(OUT)/Pure
 
-$(OUT)/Pure: General/ROOT.ML General/buffer.ML General/file.ML \
-  General/graph.ML General/history.ML General/name_space.ML \
-  General/object.ML General/path.ML General/position.ML \
-  General/pretty.ML General/scan.ML General/seq.ML General/source.ML \
-  General/symbol.ML General/table.ML General/url.ML Interface/ROOT.ML \
-  Interface/isamode.ML Interface/proof_general.ML Isar/ROOT.ML \
-  Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML \
-  Isar/comment.ML Isar/isar.ML Isar/isar_cmd.ML Isar/isar_syn.ML \
-  Isar/isar_thy.ML Isar/local_defs.ML Isar/method.ML Isar/obtain.ML \
-  Isar/outer_lex.ML Isar/outer_parse.ML Isar/outer_syntax.ML \
-  Isar/proof.ML Isar/proof_context.ML Isar/proof_data.ML \
-  Isar/proof_history.ML Isar/session.ML Isar/skip_proof.ML \
-  Isar/toplevel.ML ML-Systems/mlworks.ML ML-Systems/polyml.ML \
-  ML-Systems/smlnj-0.93.ML ML-Systems/smlnj.ML ROOT.ML Syntax/ROOT.ML \
-  Syntax/ast.ML Syntax/lexicon.ML Syntax/mixfix.ML Syntax/parser.ML \
-  Syntax/printer.ML Syntax/syn_ext.ML Syntax/syn_trans.ML \
-  Syntax/syntax.ML Syntax/token_trans.ML Syntax/type_ext.ML \
-  Thy/ROOT.ML Thy/html.ML Thy/latex.ML Thy/present.ML Thy/thm_deps.ML \
-  Thy/thm_database.ML Thy/thy_info.ML Thy/thy_load.ML Thy/thy_parse.ML \
-  Thy/thy_scan.ML Thy/thy_syn.ML axclass.ML basis.ML context.ML \
-  deriv.ML display.ML drule.ML envir.ML goals.ML install_pp.ML \
-  library.ML locale.ML logic.ML net.ML pattern.ML pure.ML pure_thy.ML \
-  search.ML sign.ML sorts.ML tactic.ML tctical.ML term.ML theory.ML \
-  theory_data.ML thm.ML type.ML type_infer.ML unify.ML
+$(OUT)/Pure: General/ROOT.ML General/buffer.ML General/file.ML		\
+  General/graph.ML General/history.ML General/name_space.ML		\
+  General/object.ML General/path.ML General/position.ML			\
+  General/pretty.ML General/scan.ML General/seq.ML General/source.ML	\
+  General/symbol.ML General/table.ML General/url.ML Interface/ROOT.ML	\
+  Interface/isamode.ML Interface/proof_general.ML Isar/ROOT.ML		\
+  Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML	\
+  Isar/comment.ML Isar/isar.ML Isar/isar_cmd.ML Isar/isar_syn.ML	\
+  Isar/isar_thy.ML Isar/local_defs.ML Isar/method.ML Isar/net_rules.ML	\
+  Isar/obtain.ML Isar/outer_lex.ML Isar/outer_parse.ML			\
+  Isar/outer_syntax.ML Isar/proof.ML Isar/proof_context.ML		\
+  Isar/proof_data.ML Isar/proof_history.ML Isar/session.ML		\
+  Isar/skip_proof.ML Isar/toplevel.ML ML-Systems/mlworks.ML		\
+  ML-Systems/polyml.ML ML-Systems/smlnj-0.93.ML ML-Systems/smlnj.ML	\
+  ROOT.ML Syntax/ROOT.ML Syntax/ast.ML Syntax/lexicon.ML		\
+  Syntax/mixfix.ML Syntax/parser.ML Syntax/printer.ML			\
+  Syntax/syn_ext.ML Syntax/syn_trans.ML Syntax/syntax.ML		\
+  Syntax/token_trans.ML Syntax/type_ext.ML Thy/ROOT.ML Thy/html.ML	\
+  Thy/latex.ML Thy/present.ML Thy/thm_deps.ML Thy/thm_database.ML	\
+  Thy/thy_info.ML Thy/thy_load.ML Thy/thy_parse.ML Thy/thy_scan.ML	\
+  Thy/thy_syn.ML axclass.ML basis.ML context.ML deriv.ML display.ML	\
+  drule.ML envir.ML goals.ML install_pp.ML library.ML locale.ML		\
+  logic.ML net.ML pattern.ML pure.ML pure_thy.ML search.ML sign.ML	\
+  sorts.ML tactic.ML tctical.ML term.ML theory.ML theory_data.ML	\
+  thm.ML type.ML type_infer.ML unify.ML
 	@./mk
 
 
--- a/src/Pure/Isar/ROOT.ML	Thu Feb 24 21:33:32 2000 +0100
+++ b/src/Pure/Isar/ROOT.ML	Sun Feb 27 15:07:53 2000 +0100
@@ -13,6 +13,7 @@
 use "proof_history.ML";
 use "args.ML";
 use "attrib.ML";
+use "net_rules.ML";
 use "method.ML";
 
 (*derived proof elements*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/net_rules.ML	Sun Feb 27 15:07:53 2000 +0100
@@ -0,0 +1,74 @@
+(*  Title:      Pure/Isar/net_rules.ML
+    ID:         $Id$
+    Author:     Markus Wenzel, TU Muenchen
+
+Efficient storage of rules: preserves order, prefers later entries.
+*)
+
+signature NET_RULES =
+sig
+  type 'a T
+  val rules: 'a T -> 'a list
+  val may_unify: 'a T -> term -> 'a list
+  val init: ('a * 'a -> bool) -> ('a -> term) -> 'a T
+  val merge: 'a T * 'a T -> 'a T
+  val delete: 'a -> 'a T -> 'a T
+  val insert: 'a -> 'a T -> 'a T
+  val deletes: 'a list -> 'a T -> 'a T
+  val inserts: 'a list -> 'a T -> 'a T
+  val init_intro: thm T
+  val init_elim: thm T
+end;
+
+structure NetRules: NET_RULES =
+struct
+
+(* datatype rules *)
+
+datatype 'a T =
+  Rules of {
+    eq: 'a * 'a -> bool,
+    index: 'a -> term,
+    rules: 'a list,
+    next: int,
+    net: (int * 'a) Net.net};
+
+fun mk_rules eq index rules next net =
+  Rules {eq = eq, index = index, rules = rules, next = next, net = net};
+
+fun rules (Rules {rules = rs, ...}) = rs;
+fun may_unify (Rules {net, ...}) tm = Tactic.orderlist (Net.unify_term net tm);
+
+
+(* build rules *)
+
+fun init eq index = mk_rules eq index [] ~1 Net.empty;
+
+fun add rule (Rules {eq, index, rules, next, net}) =
+  mk_rules eq index (rule :: rules) (next - 1)
+    (Net.insert_term ((index rule, (next, rule)), net, K false));
+
+fun make eq index rules = foldr (uncurry add) (rules, init eq index);
+
+
+fun merge (Rules {eq, index, rules = rules1, ...}, Rules {rules = rules2, ...}) =
+  make eq index (Library.generic_merge eq I I rules1 rules2);
+
+fun delete rule (rs as Rules {eq, index, rules, next, net}) =
+  if not (Library.gen_mem eq (rule, rules)) then rs
+  else mk_rules eq index (Library.gen_rem eq (rules, rule)) next
+    (Net.delete_term ((index rule, (0, rule)), net, eq o pairself #2));
+
+fun insert rule rs = add rule (delete rule rs);    (*ensure new rule gets precedence*)
+
+fun deletes rules rs = foldr (uncurry delete) (rules, rs);
+fun inserts rules rs = foldr (uncurry insert) (rules, rs);
+
+
+(* intro/elim rules *)
+
+val init_intro = init Thm.eq_thm Thm.concl_of;
+val init_elim = init Thm.eq_thm Thm.major_prem_of;
+
+
+end;