src/Pure/Isar/context_rules.ML
author wenzelm
Sat, 09 Aug 2008 22:43:46 +0200
changeset 27809 a1e409db516b
parent 26463 9283b4185fdf
child 29582 0950f4f0d0cd
permissions -rw-r--r--
unified Args.T with OuterLex.token, renamed some operations;
     1 (*  Title:      Pure/Isar/context_rules.ML
     2     ID:         $Id$
     3     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
     4 
     5 Declarations of intro/elim/dest rules in Pure (see also
     6 Provers/classical.ML for a more specialized version of the same idea).
     7 *)
     8 
     9 signature CONTEXT_RULES =
    10 sig
    11   type netpair
    12   type T
    13   val netpair_bang: Proof.context -> netpair
    14   val netpair: Proof.context -> netpair
    15   val orderlist: ((int * int) * 'a) list -> 'a list
    16   val find_rules_netpair: bool -> thm list -> term -> netpair -> thm list
    17   val find_rules: bool -> thm list -> term -> Proof.context -> thm list list
    18   val print_rules: Proof.context -> unit
    19   val addSWrapper: ((int -> tactic) -> int -> tactic) -> theory -> theory
    20   val addWrapper: ((int -> tactic) -> int -> tactic) -> theory -> theory
    21   val Swrap: Proof.context -> (int -> tactic) -> int -> tactic
    22   val wrap: Proof.context -> (int -> tactic) -> int -> tactic
    23   val intro_bang: int option -> attribute
    24   val elim_bang: int option -> attribute
    25   val dest_bang: int option -> attribute
    26   val intro: int option -> attribute
    27   val elim: int option -> attribute
    28   val dest: int option -> attribute
    29   val intro_query: int option -> attribute
    30   val elim_query: int option -> attribute
    31   val dest_query: int option -> attribute
    32   val rule_del: attribute
    33   val add_args:
    34     (int option -> attribute) -> (int option -> attribute) -> (int option -> attribute) ->
    35     Attrib.src -> attribute
    36 end;
    37 
    38 structure ContextRules: CONTEXT_RULES =
    39 struct
    40 
    41 
    42 (** rule declaration contexts **)
    43 
    44 (* rule kinds *)
    45 
    46 val intro_bangK = (0, false);
    47 val elim_bangK = (0, true);
    48 val introK = (1, false);
    49 val elimK = (1, true);
    50 val intro_queryK = (2, false);
    51 val elim_queryK = (2, true);
    52 
    53 val kind_names =
    54  [(intro_bangK, "safe introduction rules (intro!)"),
    55   (elim_bangK, "safe elimination rules (elim!)"),
    56   (introK, "introduction rules (intro)"),
    57   (elimK, "elimination rules (elim)"),
    58   (intro_queryK, "extra introduction rules (intro?)"),
    59   (elim_queryK, "extra elimination rules (elim?)")];
    60 
    61 val rule_kinds = map #1 kind_names;
    62 val rule_indexes = distinct (op =) (map #1 rule_kinds);
    63 
    64 
    65 (* context data *)
    66 
    67 type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net;
    68 val empty_netpairs: netpair list = replicate (length rule_indexes) (Net.empty, Net.empty);
    69 
    70 datatype T = Rules of
    71  {next: int,
    72   rules: (int * ((int * bool) * thm)) list,
    73   netpairs: netpair list,
    74   wrappers: (((int -> tactic) -> int -> tactic) * stamp) list *
    75     (((int -> tactic) -> int -> tactic) * stamp) list};
    76 
    77 fun make_rules next rules netpairs wrappers =
    78   Rules {next = next, rules = rules, netpairs = netpairs, wrappers = wrappers};
    79 
    80 fun add_rule (i, b) opt_w th (Rules {next, rules, netpairs, wrappers}) =
    81   let val w = (case opt_w of SOME w => w | NONE => Tactic.subgoals_of_brl (b, th)) in
    82     make_rules (next - 1) ((w, ((i, b), th)) :: rules)
    83       (nth_map i (Tactic.insert_tagged_brl ((w, next), (b, th))) netpairs) wrappers
    84   end;
    85 
    86 fun del_rule th (rs as Rules {next, rules, netpairs, wrappers}) =
    87   let
    88     fun eq_th (_, (_, th')) = Thm.eq_thm_prop (th, th');
    89     fun del b netpair = Tactic.delete_tagged_brl (b, th) netpair handle Net.DELETE => netpair;
    90   in
    91     if not (exists eq_th rules) then rs
    92     else make_rules next (filter_out eq_th rules) (map (del false o del true) netpairs) wrappers
    93   end;
    94 
    95 structure Rules = GenericDataFun
    96 (
    97   type T = T;
    98   val empty = make_rules ~1 [] empty_netpairs ([], []);
    99   val extend = I;
   100 
   101   fun merge _ (Rules {rules = rules1, wrappers = (ws1, ws1'), ...},
   102       Rules {rules = rules2, wrappers = (ws2, ws2'), ...}) =
   103     let
   104       val wrappers =
   105         (Library.merge (eq_snd (op =)) (ws1, ws2), Library.merge (eq_snd (op =)) (ws1', ws2'));
   106       val rules = Library.merge (fn ((_, (k1, th1)), (_, (k2, th2))) =>
   107           k1 = k2 andalso Thm.eq_thm_prop (th1, th2)) (rules1, rules2);
   108       val next = ~ (length rules);
   109       val netpairs = fold (fn (n, (w, ((i, b), th))) =>
   110           nth_map i (Tactic.insert_tagged_brl ((w, n), (b, th))))
   111         (next upto ~1 ~~ rules) empty_netpairs;
   112     in make_rules (next - 1) rules netpairs wrappers end;
   113 );
   114 
   115 fun print_rules ctxt =
   116   let
   117     val Rules {rules, ...} = Rules.get (Context.Proof ctxt);
   118     fun prt_kind (i, b) =
   119       Pretty.big_list ((the o AList.lookup (op =) kind_names) (i, b) ^ ":")
   120         (map_filter (fn (_, (k, th)) =>
   121             if k = (i, b) then SOME (ProofContext.pretty_thm ctxt th) else NONE)
   122           (sort (int_ord o pairself fst) rules));
   123   in Pretty.writeln (Pretty.chunks (map prt_kind rule_kinds)) end;
   124 
   125 
   126 (* access data *)
   127 
   128 fun netpairs ctxt = let val Rules {netpairs, ...} = Rules.get (Context.Proof ctxt) in netpairs end;
   129 val netpair_bang = hd o netpairs;
   130 val netpair = hd o tl o netpairs;
   131 
   132 
   133 (* retrieving rules *)
   134 
   135 fun untaglist [] = []
   136   | untaglist [(k : int * int, x)] = [x]
   137   | untaglist ((k, x) :: (rest as (k', x') :: _)) =
   138       if k = k' then untaglist rest
   139       else x :: untaglist rest;
   140 
   141 fun orderlist brls =
   142   untaglist (sort (prod_ord int_ord int_ord o pairself fst) brls);
   143 
   144 fun orderlist_no_weight brls =
   145   untaglist (sort (int_ord o pairself (snd o fst)) brls);
   146 
   147 fun may_unify weighted t net =
   148   map snd ((if weighted then orderlist else orderlist_no_weight) (Net.unify_term net t));
   149 
   150 fun find_erules _ [] = K []
   151   | find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact));
   152 
   153 fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal);
   154 
   155 fun find_rules_netpair weighted facts goal (inet, enet) =
   156   find_erules weighted facts enet @ find_irules weighted goal inet;
   157 
   158 fun find_rules weighted facts goals =
   159   map (find_rules_netpair weighted facts goals) o netpairs;
   160 
   161 
   162 (* wrappers *)
   163 
   164 fun gen_add_wrapper upd w =
   165   Context.theory_map (Rules.map (fn (rs as Rules {next, rules, netpairs, wrappers}) =>
   166     make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers)));
   167 
   168 val addSWrapper = gen_add_wrapper Library.apfst;
   169 val addWrapper = gen_add_wrapper Library.apsnd;
   170 
   171 
   172 fun gen_wrap which ctxt =
   173   let val Rules {wrappers, ...} = Rules.get (Context.Proof ctxt)
   174   in fold_rev fst (which wrappers) end;
   175 
   176 val Swrap = gen_wrap #1;
   177 val wrap = gen_wrap #2;
   178 
   179 
   180 
   181 (** attributes **)
   182 
   183 (* add and del rules *)
   184 
   185 fun rule_del (x, th) =
   186   (Rules.map (del_rule th o del_rule (Tactic.make_elim th)) x, th);
   187 
   188 fun rule_add k view opt_w =
   189   (fn (x, th) => (Rules.map (add_rule k opt_w (view th)) x, th)) o rule_del;
   190 
   191 val intro_bang  = rule_add intro_bangK I;
   192 val elim_bang   = rule_add elim_bangK I;
   193 val dest_bang   = rule_add elim_bangK Tactic.make_elim;
   194 val intro       = rule_add introK I;
   195 val elim        = rule_add elimK I;
   196 val dest        = rule_add elimK Tactic.make_elim;
   197 val intro_query = rule_add intro_queryK I;
   198 val elim_query  = rule_add elim_queryK I;
   199 val dest_query  = rule_add elim_queryK Tactic.make_elim;
   200 
   201 val _ = Context.>> (Context.map_theory
   202   (snd o PureThy.add_thms [(("", Drule.equal_intr_rule), [intro_query NONE])]));
   203 
   204 
   205 (* concrete syntax *)
   206 
   207 fun add_args a b c = Attrib.syntax
   208   (Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) --
   209     Scan.option OuterParse.nat) >> (fn (f, n) => f n));
   210 
   211 val rule_atts =
   212  [("intro", add_args intro_bang intro intro_query, "declaration of introduction rule"),
   213   ("elim", add_args elim_bang elim elim_query, "declaration of elimination rule"),
   214   ("dest", add_args dest_bang dest dest_query, "declaration of destruction rule"),
   215   ("rule", Attrib.syntax (Scan.lift Args.del >> K rule_del),
   216     "remove declaration of intro/elim/dest rule")];
   217 
   218 val _ = Context.>> (Context.map_theory (Attrib.add_attributes rule_atts));
   219 
   220 end;