aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--test-suite/bugs/closed/6770.v7
-rw-r--r--test-suite/success/Fixpoint.v30
-rw-r--r--vernac/comFixpoint.ml2
3 files changed, 38 insertions, 1 deletions
diff --git a/test-suite/bugs/closed/6770.v b/test-suite/bugs/closed/6770.v
new file mode 100644
index 000000000..9bcc74083
--- /dev/null
+++ b/test-suite/bugs/closed/6770.v
@@ -0,0 +1,7 @@
+Section visibility.
+
+ Let Fixpoint by_proof (n:nat) : True.
+ Proof. exact I. Defined.
+End visibility.
+
+Fail Check by_proof.
diff --git a/test-suite/success/Fixpoint.v b/test-suite/success/Fixpoint.v
index 5fc703cf0..efb32ef6f 100644
--- a/test-suite/success/Fixpoint.v
+++ b/test-suite/success/Fixpoint.v
@@ -91,3 +91,33 @@ apply Cons2.
exact b.
apply (ex1 (S n) (negb b)).
Defined.
+
+Section visibility.
+
+ Let Fixpoint imm (n:nat) : True := I.
+
+ Let Fixpoint by_proof (n:nat) : True.
+ Proof. exact I. Defined.
+End visibility.
+
+Fail Check imm.
+Fail Check by_proof.
+
+Module Import mod_local.
+ Fixpoint imm_importable (n:nat) : True := I.
+
+ Local Fixpoint imm_local (n:nat) : True := I.
+
+ Fixpoint by_proof_importable (n:nat) : True.
+ Proof. exact I. Defined.
+
+ Local Fixpoint by_proof_local (n:nat) : True.
+ Proof. exact I. Defined.
+End mod_local.
+
+Check imm_importable.
+Fail Check imm_local.
+Check mod_local.imm_local.
+Check by_proof_importable.
+Fail Check by_proof_local.
+Check mod_local.by_proof_local.
diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml
index ea731b34c..b5b8697d2 100644
--- a/vernac/comFixpoint.ml
+++ b/vernac/comFixpoint.ml
@@ -254,7 +254,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC)
fixdefs) in
let evd = Evd.from_ctx ctx in
- Lemmas.start_proof_with_initialization (Global,poly,DefinitionBody Fixpoint)
+ Lemmas.start_proof_with_initialization (local,poly,DefinitionBody Fixpoint)
evd pl (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ()))
else begin
(* We shortcut the proof process *)