aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--tactics/refine.ml14
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2350.v6
2 files changed, 19 insertions, 1 deletions
diff --git a/tactics/refine.ml b/tactics/refine.ml
index 484954bcc..db89ec6e2 100644
--- a/tactics/refine.ml
+++ b/tactics/refine.ml
@@ -261,6 +261,16 @@ let rec compute_metamap env sigma c = match kind_of_term c with
* Réalise le 3. ci-dessus
*)
+let ensure_products n =
+ let p = ref 0 in
+ let rec aux n =
+ if n = 0 then tclFAIL 0 (mt())
+ else tclORELSE intro (fun gl -> incr p; tclTHEN introf (aux (n-1)) gl) in
+ tclORELSE
+ (aux n)
+ (* Now we know how many red are needed *)
+ (fun gl -> tclDO !p red_in_concl gl)
+
let rec tcc_aux subst (TH (c,mm,sgp) as _th) gl =
let c = substl subst c in
match (kind_of_term c,sgp) with
@@ -339,7 +349,9 @@ let rec tcc_aux subst (TH (c,mm,sgp) as _th) gl =
let fixes = array_map3 (fun f n c -> (out_name f,succ n,c)) fi ni ai in
let firsts,lasts = list_chop j (Array.to_list fixes) in
tclTHENS
- (mutual_fix (out_name fi.(j)) (succ ni.(j)) (firsts@List.tl lasts) j)
+ (tclTHEN
+ (ensure_products (succ ni.(j)))
+ (mutual_fix (out_name fi.(j)) (succ ni.(j)) (firsts@List.tl lasts) j))
(List.map (function
| None -> tclIDTAC
| Some th -> tcc_aux subst th) sgp)
diff --git a/test-suite/bugs/closed/shouldsucceed/2350.v b/test-suite/bugs/closed/shouldsucceed/2350.v
new file mode 100644
index 000000000..e91f22e26
--- /dev/null
+++ b/test-suite/bugs/closed/shouldsucceed/2350.v
@@ -0,0 +1,6 @@
+(* Check that the fix tactic, when called from refine, reduces enough
+ to see the products *)
+
+Definition foo := forall n:nat, n=n.
+Definition bar : foo.
+refine (fix aux (n:nat) := _).