Skip to content

Commit

Permalink
Merge PR #12627: Fix Canonical with universe polymorphism and primiti…
Browse files Browse the repository at this point in the history
…ve projection

Reviewed-by: ejgallego
Ack-by: gares
  • Loading branch information
gares committed Jul 8, 2020
2 parents 421b221 + 1e92ed4 commit 769823c
Show file tree
Hide file tree
Showing 4 changed files with 14 additions and 1 deletion.
2 changes: 1 addition & 1 deletion pretyping/recordops.ml
Original file line number Diff line number Diff line change
Expand Up @@ -185,7 +185,7 @@ let rec cs_pattern_of_constr env t =
| Rel n -> Default_cs, Some n, []
| Prod (_,a,b) when Vars.noccurn 1 b -> Prod_cs, None, [a; Vars.lift (-1) b]
| Proj (p, c) ->
let { Environ.uj_type = ty } = Typeops.infer env c in
let ty = Retyping.get_type_of_constr env c in
let _, params = Inductive.find_rectype env ty in
Const_cs (GlobRef.ConstRef (Projection.constant p)), None, params @ [c]
| Sort s -> Sort_cs (Sorts.family s), None, []
Expand Down
3 changes: 3 additions & 0 deletions pretyping/retyping.ml
Original file line number Diff line number Diff line change
Expand Up @@ -260,6 +260,9 @@ let get_type_of ?(polyprop=true) ?(lax=false) env sigma c =
(* Makes an unsafe judgment from a constr *)
let get_judgment_of env evc c = { uj_val = c; uj_type = get_type_of env evc c }

let get_type_of_constr ?polyprop ?lax env ?(uctx=UState.from_env env) c =
EConstr.Unsafe.to_constr (get_type_of ?polyprop ?lax env (Evd.from_ctx uctx) (EConstr.of_constr c))

(* Returns sorts of a context *)
let sorts_of_context env evc ctxt =
let rec aux = function
Expand Down
4 changes: 4 additions & 0 deletions pretyping/retyping.mli
Original file line number Diff line number Diff line change
Expand Up @@ -30,6 +30,10 @@ exception RetypeError of retype_error
val get_type_of :
?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> types

(** No-evar version of [get_type_of] *)
val get_type_of_constr : ?polyprop:bool -> ?lax:bool
-> env -> ?uctx:UState.t -> Constr.t -> Constr.types

val get_sort_of :
?polyprop:bool -> env -> evar_map -> types -> Sorts.t

Expand Down
6 changes: 6 additions & 0 deletions test-suite/bugs/closed/bug_12528.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
Set Primitive Projections.
Set Universe Polymorphism.
Record ptd := Ptd { t : Type ; p : t }.
Definition type := Ptd Type (unit:Type).
Definition type' := Ptd Type (p type).
Canonical type'.

0 comments on commit 769823c

Please sign in to comment.