aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel
diff options
context:
space:
mode:
authorGravatar barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7>2006-05-03 12:02:44 +0000
committerGravatar barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7>2006-05-03 12:02:44 +0000
commita9df5940e0b362f627abf500e056e197e403c8e1 (patch)
tree81e6aad40d9343085f3d5acdeb719a4fd11e1858 /kernel
parentc599e13efd9e7b00c39d5a3812d8360124257f4b (diff)
bug #1096: whd_stack on one arg of conversion had side-effect on the other arg
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@8782 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'kernel')
-rw-r--r--kernel/closure.ml9
-rw-r--r--kernel/reduction.ml16
2 files changed, 16 insertions, 9 deletions
diff --git a/kernel/closure.ml b/kernel/closure.ml
index 625d3b0d0..abbc66077 100644
--- a/kernel/closure.ml
+++ b/kernel/closure.ml
@@ -554,7 +554,7 @@ let rec lft_fconstr n ft =
| FFix(fx,e) -> {norm=Cstr; term=FFix(fx,subs_shft(n,e))}
| FCoFix(cfx,e) -> {norm=Cstr; term=FCoFix(cfx,subs_shft(n,e))}
| FLIFT(k,m) -> lft_fconstr (n+k) m
- | FLOCKED -> anomaly "lft_constr found locked term"
+ | FLOCKED -> assert false
| _ -> {norm=ft.norm; term=FLIFT(n,ft)}
let lift_fconstr k f =
if k=0 then f else lft_fconstr k f
@@ -772,8 +772,7 @@ let rec to_constr constr_fun lfts v =
let fr = mk_clos2 env t in
let unfv = update v (fr.norm,fr.term) in
to_constr constr_fun lfts unfv
- | FLOCKED -> (*anomaly "Closure.to_constr: found locked term"*)
-mkVar(id_of_string"_LOCK_")
+ | FLOCKED -> assert false (*mkVar(id_of_string"_LOCK_")*)
(* This function defines the correspondance between constr and
fconstr. When we find a closure whose substitution is the identity,
@@ -949,7 +948,7 @@ let contract_fix_vect fix =
(bds.(i),
(fun j -> { norm = Cstr; term = FCoFix ((j,rdcl),env) }),
env, Array.length bds)
- | _ -> anomaly "Closure.contract_fix_vect: not a (co)fixpoint"
+ | _ -> assert false
in
let rec subst_bodies_from_i i env =
if i = nfix then
@@ -969,7 +968,7 @@ let rec knh m stk =
match m.term with
| FLIFT(k,a) -> knh a (zshift k stk)
| FCLOS(t,e) -> knht e t (zupdate m stk)
- | FLOCKED -> anomaly "Closure.knh: found lock"
+ | FLOCKED -> assert false
| FApp(a,b) -> knh a (append_stack b (zupdate m stk))
| FCases(ci,p,t,br) -> knh t (Zcase(ci,p,br)::zupdate m stk)
| FFix(((ri,n),(_,_,_)),_) ->
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 2467e941f..4aa7c5753 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -291,10 +291,18 @@ and eqappr cv_pb infos appr1 appr2 cuniv =
convert_stacks infos lft1 lft2 v1 v2 u2
else raise NotConvertible
- | ( (FLetIn _, _) | (_, FLetIn _) | (FCases _,_) | (_,FCases _)
- | (FApp _,_) | (_,FApp _) | (FCLOS _, _) | (_,FCLOS _)
- | (FLIFT _, _) | (_,FLIFT _) | (FLOCKED,_) | (_,FLOCKED)) ->
- anomaly "Unexpected term returned by fhnf"
+ (* Can happen because whd_stack on one arg may have side-effects
+ on the other arg and coulb be no more in hnf... *)
+ | ( (FLetIn _, _) | (FCases _,_) | (FApp _,_)
+ | (FCLOS _, _) | (FLIFT _, _)) ->
+ eqappr cv_pb infos (lft1, whd_stack infos hd1 v1) appr2 cuniv
+
+ | ( (_, FLetIn _) | (_,FCases _) | (_,FApp _)
+ | (_,FCLOS _) | (_,FLIFT _)) ->
+ eqappr cv_pb infos (lft1, whd_stack infos hd1 v1) appr2 cuniv
+
+ (* Should not happen because whd_stack unlocks references *)
+ | ((FLOCKED,_) | (_,FLOCKED)) -> assert false
| _ -> raise NotConvertible