aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics
diff options
context:
space:
mode:
authorGravatar soubiran <soubiran@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-01-12 16:12:44 +0000
committerGravatar soubiran <soubiran@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-01-12 16:12:44 +0000
commitb264811fbed75caff5deb2b6eb78a327dc134f68 (patch)
treee36bbb300f21531d621ffcbc49aaebc6cc266c94 /tactics
parent1a197792a437ca0a797c5dcec099d64b550028f8 (diff)
Added module sharing support for typeclasses and hints (pri_auto_tactic).
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12655 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'tactics')
-rw-r--r--tactics/auto.ml20
1 files changed, 18 insertions, 2 deletions
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 8447dfdef..45e384a41 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -99,10 +99,26 @@ type search_entry = stored_data list * stored_data list * Bounded_net.t
let empty_se = ([],[],Bounded_net.create ())
+let eq_pri_auto_tactic x y =
+ if x.pri = y.pri && x.pat = y.pat then
+ match x.code,y.code with
+ | Res_pf(cstr,_),Res_pf(cstr1,_) ->
+ eq_constr cstr cstr1
+ | ERes_pf(cstr,_),ERes_pf(cstr1,_) ->
+ eq_constr cstr cstr1
+ | Give_exact cstr,Give_exact cstr1 ->
+ eq_constr cstr cstr1
+ | Res_pf_THEN_trivial_fail(cstr,_)
+ ,Res_pf_THEN_trivial_fail(cstr1,_) ->
+ eq_constr cstr cstr1
+ | _,_ -> false
+ else
+ false
+
let add_tac pat t st (l,l',dn) =
match pat with
- | None -> if not (List.mem t l) then (insert t l, l', dn) else (l, l', dn)
- | Some pat -> if not (List.mem t l') then (l, insert t l', Bounded_net.add st dn (pat,t)) else (l, l', dn)
+ | None -> if not (List.exists (eq_pri_auto_tactic t) l) then (insert t l, l', dn) else (l, l', dn)
+ | Some pat -> if not (List.exists (eq_pri_auto_tactic t) l') then (l, insert t l', Bounded_net.add st dn (pat,t)) else (l, l', dn)
let rebuild_dn st (l,l',dn) =
(l, l', List.fold_left (fun dn t -> Bounded_net.add (Some st) dn (Option.get t.pat, t))