aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--test-suite/bugs/closed/5300.v39
-rw-r--r--toplevel/assumptions.ml21
2 files changed, 58 insertions, 2 deletions
diff --git a/test-suite/bugs/closed/5300.v b/test-suite/bugs/closed/5300.v
new file mode 100644
index 000000000..18202df50
--- /dev/null
+++ b/test-suite/bugs/closed/5300.v
@@ -0,0 +1,39 @@
+Module Test1.
+
+ Module Type Foo.
+ Parameter t : unit.
+ End Foo.
+
+ Module Bar : Foo.
+ Module Type Rnd. Definition t' : unit := tt. End Rnd.
+ Module Rnd_inst : Rnd. Definition t' : unit := tt. End Rnd_inst.
+ Definition t : unit.
+ exact Rnd_inst.t'.
+ Qed.
+ End Bar.
+
+ Print Assumptions Bar.t.
+End Test1.
+
+Module Test2.
+ Module Type Foo.
+ Parameter t1 : unit.
+ Parameter t2 : unit.
+ End Foo.
+
+ Module Bar : Foo.
+ Inductive ind := .
+ Definition t' : unit := tt.
+ Definition t1 : unit.
+ Proof.
+ exact ((fun (_ : ind -> False) => tt) (fun H => match H with end)).
+ Qed.
+ Definition t2 : unit.
+ Proof.
+ exact t'.
+ Qed.
+ End Bar.
+
+ Print Assumptions Bar.t1.
+ Print Assumptions Bar.t1.
+End Test2.
diff --git a/toplevel/assumptions.ml b/toplevel/assumptions.ml
index 45c539e22..a865ee987 100644
--- a/toplevel/assumptions.ml
+++ b/toplevel/assumptions.ml
@@ -42,6 +42,11 @@ let rec search_cst_label lab = function
| (l, SFBconst cb) :: _ when Label.equal l lab -> cb
| _ :: fields -> search_cst_label lab fields
+let rec search_mind_label lab = function
+ | [] -> raise Not_found
+ | (l, SFBmind mind) :: _ when Label.equal l lab -> mind
+ | _ :: fields -> search_mind_label lab fields
+
(* TODO: using [empty_delta_resolver] below is probably slightly incorrect. But:
a) I don't see currently what should be used instead
b) this shouldn't be critical for Print Assumption. At worse some
@@ -133,6 +138,18 @@ let lookup_constant cst =
else lookup_constant_in_impl cst (Some cb)
with Not_found -> lookup_constant_in_impl cst None
+let lookup_mind_in_impl mind =
+ try
+ let mp,dp,lab = repr_kn (canonical_mind mind) in
+ let fields = memoize_fields_of_mp mp in
+ search_mind_label lab fields
+ with Not_found ->
+ anomaly (str "Print Assumption: unknown inductive " ++ MutInd.print mind)
+
+let lookup_mind mind =
+ try Global.lookup_mind mind
+ with Not_found -> lookup_mind_in_impl mind
+
(** Graph traversal of an object, collecting on the way the dependencies of
traversed objects *)
@@ -204,7 +221,7 @@ and traverse_inductive (curr, data, ax2ty) mind obj =
where I_0, I_1, ... are in the same mutual definition and c_ij
are all their constructors. *)
if Refmap_env.mem firstind_ref data then data, ax2ty else
- let mib = Global.lookup_mind mind in
+ let mib = lookup_mind mind in
(* Collects references of parameters *)
let param_ctx = mib.mind_params_ctxt in
let nparam = List.length param_ctx in
@@ -308,7 +325,7 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
else
accu
| IndRef (m,_) | ConstructRef ((m,_),_) ->
- let mind = Global.lookup_mind m in
+ let mind = lookup_mind m in
if mind.mind_typing_flags.check_guarded then
accu
else