diff options
| author | Gaëtan Gilbert | 2020-06-06 21:21:50 +0200 |
|---|---|---|
| committer | Gaëtan Gilbert | 2020-06-06 21:22:08 +0200 |
| commit | bc842a94b18f3a09ea1127b0bd3e08599dd97305 (patch) | |
| tree | 61976d3a6e22cba7a551b4ac75d52c371e965300 | |
| parent | 9c26e583668827bff5159e7671c406ace8b2f3ae (diff) | |
Fix uncaught NotArity in inductive type
Fixes #12390
| -rw-r--r-- | test-suite/bugs/closed/bug_12390.v | 3 | ||||
| -rw-r--r-- | vernac/comInductive.ml | 2 |
2 files changed, 4 insertions, 1 deletions
diff --git a/test-suite/bugs/closed/bug_12390.v b/test-suite/bugs/closed/bug_12390.v new file mode 100644 index 0000000000..a8fda29a6f --- /dev/null +++ b/test-suite/bugs/closed/bug_12390.v @@ -0,0 +1,3 @@ +Fail Inductive foo : forall P, P := . +Fail Inductive bar : nat := . +Fail Inductive baz : _ := . diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml index 4242f06844..95489c9132 100644 --- a/vernac/comInductive.ml +++ b/vernac/comInductive.ml @@ -117,7 +117,7 @@ let intern_ind_arity env sigma ind = let pretype_ind_arity env sigma (loc, c, impls, pseudo_poly) = let sigma,t = understand_tcc env sigma ~expected_type:IsType c in match Reductionops.sort_of_arity env sigma t with - | exception Invalid_argument _ -> + | exception Reduction.NotArity -> user_err ?loc (str "Not an arity") | s -> let concl = match pseudo_poly with |
