aboutsummaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorMatthieu Sozeau2017-07-28 15:31:28 +0200
committerMatthieu Sozeau2019-02-08 10:55:51 +0100
commitc039d78bd098a499a34038e64bd1e5fbe280d7f3 (patch)
tree8bcc7cd765321a9df77c3888bd3d3a28a35438f2 /pretyping
parent46b671c7473385ec7747a796e85b3cf704d000c6 (diff)
locus: Add an AtLeastOneOccurrence constructor.
The semantics is obviously that it is an error if not at least one occurrence is found (natural semantics for rewriting for example).
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/evarconv.ml2
-rw-r--r--pretyping/find_subterm.ml2
-rw-r--r--pretyping/locus.ml1
-rw-r--r--pretyping/locusops.ml17
-rw-r--r--pretyping/locusops.mli2
-rw-r--r--pretyping/tacred.ml1
-rw-r--r--pretyping/unification.ml2
7 files changed, 19 insertions, 8 deletions
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index aa30ed8d2e..f8e5e0759b 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -1148,7 +1148,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
| (id,_,c,cty,evsref,filter,occs)::subst ->
let set_var k =
match occs with
- | Some Locus.AllOccurrences -> mkVar id
+ | Some (Locus.AtLeastOneOccurrence | Locus.AllOccurrences) -> mkVar id
| Some _ -> user_err Pp.(str "Selection of specific occurrences not supported")
| None ->
let evty = set_holes evdref cty subst in
diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml
index b16087031b..d150f8e1cb 100644
--- a/pretyping/find_subterm.ml
+++ b/pretyping/find_subterm.ml
@@ -125,7 +125,7 @@ let replace_term_occ_gen_modulo sigma occs like_first test bywhat cl occ t =
end;
add_subst t subst; incr pos;
(* Check nested matching subterms *)
- if occs != Locus.AllOccurrences && occs != Locus.NoOccurrences then
+ if not (Locusops.is_all_occurrences occs) && occs != Locus.NoOccurrences then
begin nested := true; ignore (subst_below k t); nested := false end;
(* Do the effective substitution *)
Vars.lift k (bywhat ()))
diff --git a/pretyping/locus.ml b/pretyping/locus.ml
index 37dd120c1a..087a6b9174 100644
--- a/pretyping/locus.ml
+++ b/pretyping/locus.ml
@@ -20,6 +20,7 @@ type 'a or_var =
type 'a occurrences_gen =
| AllOccurrences
+ | AtLeastOneOccurrence
| AllOccurrencesBut of 'a list (** non-empty *)
| NoOccurrences
| OnlyOccurrences of 'a list (** non-empty *)
diff --git a/pretyping/locusops.ml b/pretyping/locusops.ml
index 6b6a3f8a9f..aaa4ce684d 100644
--- a/pretyping/locusops.ml
+++ b/pretyping/locusops.ml
@@ -19,15 +19,17 @@ let occurrences_map f = function
| AllOccurrencesBut l ->
let l' = f l in
if l' = [] then AllOccurrences else AllOccurrencesBut l'
- | (NoOccurrences|AllOccurrences) as o -> o
+ | (NoOccurrences|AllOccurrences|AtLeastOneOccurrence) as o -> o
let convert_occs = function
+ | AtLeastOneOccurrence -> (false,[])
| AllOccurrences -> (false,[])
| AllOccurrencesBut l -> (false,l)
| NoOccurrences -> (true,[])
| OnlyOccurrences l -> (true,l)
let is_selected occ = function
+ | AtLeastOneOccurrence -> true
| AllOccurrences -> true
| AllOccurrencesBut l -> not (Int.List.mem occ l)
| OnlyOccurrences l -> Int.List.mem occ l
@@ -46,6 +48,11 @@ let is_nowhere = function
| { onhyps=Some[]; concl_occs=NoOccurrences } -> true
| _ -> false
+let is_all_occurrences = function
+ | AtLeastOneOccurrence
+ | AllOccurrences -> true
+ | _ -> false
+
(** Clause conversion functions, parametrized by a hyp enumeration function *)
(** From [clause] to [simple_clause] *)
@@ -61,12 +68,12 @@ let simple_clause_of enum_hyps cl =
List.map Option.make (enum_hyps ())
| Some l ->
List.map (fun ((occs,id),w) ->
- if occs <> AllOccurrences then error_occurrences ();
+ if not (is_all_occurrences occs) then error_occurrences ();
if w = InHypValueOnly then error_body_selection ();
Some id) l in
if cl.concl_occs = NoOccurrences then hyps
else
- if cl.concl_occs <> AllOccurrences then error_occurrences ()
+ if not (is_all_occurrences cl.concl_occs) then error_occurrences ()
else None :: hyps
(** From [clause] to [concrete_clause] *)
@@ -111,7 +118,7 @@ let clause_with_generic_occurrences cls =
List.for_all
(function ((AllOccurrences,_),_) -> true | _ -> false) hyps in
let concl = match cls.concl_occs with
- | AllOccurrences | NoOccurrences -> true
+ | AtLeastOneOccurrence | AllOccurrences | NoOccurrences -> true
| _ -> false in
hyps && concl
@@ -122,6 +129,6 @@ let clause_with_generic_context_selection cls =
List.for_all
(function ((AllOccurrences,_),InHyp) -> true | _ -> false) hyps in
let concl = match cls.concl_occs with
- | AllOccurrences | NoOccurrences -> true
+ | AtLeastOneOccurrence | AllOccurrences | NoOccurrences -> true
| _ -> false in
hyps && concl
diff --git a/pretyping/locusops.mli b/pretyping/locusops.mli
index a07c018c32..ac15fe1018 100644
--- a/pretyping/locusops.mli
+++ b/pretyping/locusops.mli
@@ -21,6 +21,8 @@ val convert_occs : occurrences -> bool * int list
val is_selected : int -> occurrences -> bool
+val is_all_occurrences : 'a occurrences_gen -> bool
+
(** Usual clauses *)
val allHypsAndConcl : 'a clause_expr
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index 2308a541fb..5db571433a 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -1105,6 +1105,7 @@ let unfoldoccs env sigma (occs,name) c =
| AllOccurrences -> unfold env sigma name c
| OnlyOccurrences l -> unfo true l
| AllOccurrencesBut l -> unfo false l
+ | AtLeastOneOccurrence -> unfo false []
(* Unfold reduction tactic: *)
let unfoldn loccname env sigma c =
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index ac0b58b92b..ce97912b84 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -1649,7 +1649,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
match occurrences_of_hyp hyp occs with
| NoOccurrences, InHyp ->
(push_named_context_val d sign,depdecls)
- | AllOccurrences, InHyp as occ ->
+ | (AllOccurrences | AtLeastOneOccurrence), InHyp as occ ->
let occ = if likefirst then LikeFirst else AtOccs occ in
let newdecl = replace_term_occ_decl_modulo sigma occ test mkvarid d in
if Context.Named.Declaration.equal (EConstr.eq_constr sigma) d newdecl