aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorcoqbot-app[bot]2020-11-02 12:59:35 +0000
committerGitHub2020-11-02 12:59:35 +0000
commit35354fcb1d86fc0e8a9372b17e43a2b4a409a00e (patch)
tree7586614d9a161cd943cb26cd170306cdfb34958a
parenta89ac22aec0305b9dbf18dc147b8fe9db3b540c6 (diff)
parent42c1887379f3063e593910c569e3db8f227ac1b9 (diff)
Merge PR #13274: Fix two bugs in conversion of primitive values
Reviewed-by: SkySkimmer
-rw-r--r--kernel/cClosure.ml14
-rw-r--r--test-suite/bugs/closed/bug_13276.v9
2 files changed, 18 insertions, 5 deletions
diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml
index a23ef8fdca..952237ab99 100644
--- a/kernel/cClosure.ml
+++ b/kernel/cClosure.ml
@@ -1165,7 +1165,7 @@ module FNativeEntries =
let mkFloat env f =
check_float env;
- { mark = mark Norm KnownR; term = FFloat f }
+ { mark = mark Cstr KnownR; term = FFloat f }
let mkBool env b =
check_bool env;
@@ -1328,10 +1328,14 @@ let rec knr info tab m stk =
| FFlex(ConstKey (kn,_u as c)) when red_set info.i_flags (fCONST kn) ->
(match ref_value_cache info tab (ConstKey c) with
| Def v -> kni info tab v stk
- | Primitive op when check_native_args op stk ->
- let rargs, a, nargs, stk = get_native_args1 op c stk in
- kni info tab a (Zprimitive(op,c,rargs,nargs)::stk)
- | Undef _ | OpaqueDef _ | Primitive _ -> (set_norm m; (m,stk)))
+ | Primitive op ->
+ if check_native_args op stk then
+ let rargs, a, nargs, stk = get_native_args1 op c stk in
+ kni info tab a (Zprimitive(op,c,rargs,nargs)::stk)
+ else
+ (* Similarly to fix, partially applied primitives are not Norm! *)
+ (m, stk)
+ | Undef _ | OpaqueDef _ -> (set_norm m; (m,stk)))
| FFlex(VarKey id) when red_set info.i_flags (fVAR id) ->
(match ref_value_cache info tab (VarKey id) with
| Def v -> kni info tab v stk
diff --git a/test-suite/bugs/closed/bug_13276.v b/test-suite/bugs/closed/bug_13276.v
new file mode 100644
index 0000000000..15ac7e7b36
--- /dev/null
+++ b/test-suite/bugs/closed/bug_13276.v
@@ -0,0 +1,9 @@
+From Coq Require Import Floats.
+Open Scope float_scope.
+
+Lemma foo :
+ let n := opp 0 in sub n 0 = n.
+Proof.
+cbv.
+apply eq_refl.
+Qed.