diff options
| author | Gaëtan Gilbert | 2020-01-17 17:28:50 +0100 |
|---|---|---|
| committer | Gaëtan Gilbert | 2020-01-27 13:54:43 +0100 |
| commit | f011a88bd4be4797617741d6829d2530bc29ebdf (patch) | |
| tree | 42baae9f5d4d4e9f27233649e5d4920ada0b0617 /engine/evd.ml | |
| parent | 506b35913103c17e4d27663aa0f977452d5815b0 (diff) | |
schemes: use rigid universes
so for instance
~~~coq
Set Printing All. Set Printing Universes.
Polymorphic Inductive foo@{u v|u<=v} : Type@{u}:= .
Lemma bla@{u v|u < v} : foo@{u v} -> False.
Proof. induction 1. Qed.
~~~
works.
Diffstat (limited to 'engine/evd.ml')
| -rw-r--r-- | engine/evd.ml | 12 |
1 files changed, 6 insertions, 6 deletions
diff --git a/engine/evd.ml b/engine/evd.ml index 8e7d942c37..70f58163fd 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -902,14 +902,14 @@ let make_nonalgebraic_variable evd u = let fresh_sort_in_family ?loc ?(rigid=univ_flexible) evd s = with_context_set ?loc rigid evd (UnivGen.fresh_sort_in_family s) -let fresh_constant_instance ?loc env evd c = - with_context_set ?loc univ_flexible evd (UnivGen.fresh_constant_instance env c) +let fresh_constant_instance ?loc ?(rigid=univ_flexible) env evd c = + with_context_set ?loc rigid evd (UnivGen.fresh_constant_instance env c) -let fresh_inductive_instance ?loc env evd i = - with_context_set ?loc univ_flexible evd (UnivGen.fresh_inductive_instance env i) +let fresh_inductive_instance ?loc ?(rigid=univ_flexible) env evd i = + with_context_set ?loc rigid evd (UnivGen.fresh_inductive_instance env i) -let fresh_constructor_instance ?loc env evd c = - with_context_set ?loc univ_flexible evd (UnivGen.fresh_constructor_instance env c) +let fresh_constructor_instance ?loc ?(rigid=univ_flexible) env evd c = + with_context_set ?loc rigid evd (UnivGen.fresh_constructor_instance env c) let fresh_global ?loc ?(rigid=univ_flexible) ?names env evd gr = with_context_set ?loc rigid evd (UnivGen.fresh_global_instance ?loc ?names env gr) |
