aboutsummaryrefslogtreecommitdiffhomepage
path: root/checker
diff options
context:
space:
mode:
authorGravatar Gaëtan Gilbert <gaetan.gilbert@skyskimmer.net>2017-11-09 17:27:58 +0100
committerGravatar Gaëtan Gilbert <gaetan.gilbert@skyskimmer.net>2018-02-10 01:34:19 +0100
commit35fba70509d9fb219b2a88f8e7bd246b2671b39b (patch)
tree21cdb9aa905ef8f6375aa6cd119254794233d0ac /checker
parent10d3d803e6b57024dd15df7d61670ce42260948a (diff)
Simplification: cumulativity information is variance information.
Since cumulativity of an inductive type is the universe constraints which make a term convertible with its universe-renamed copy, the only constraints we can get are between a universe and its copy. As such we do not need to be able to represent arbitrary constraints between universes and copied universes in a double-sized ucontext, instead we can just keep around an array describing whether a bound universe is covariant, invariant or irrelevant (CIC has no contravariant conversion rule). Printing is fairly obtuse and should be improved: when we print the CumulativityInfo we add marks to the universes of the instance: = for invariant, + for covariant and * for irrelevant. ie Record Foo@{i j k} := { foo : Type@{i} -> Type@{j} }. Print Foo. gives Cumulative Record Foo : Type@{max(i+1, j+1)} := Build_Foo { foo : Type@{i} -> Type@{j} } (* =i +j *k |= *)
Diffstat (limited to 'checker')
-rw-r--r--checker/indtypes.ml17
-rw-r--r--checker/reduction.ml17
-rw-r--r--checker/univ.ml34
-rw-r--r--checker/univ.mli15
-rw-r--r--checker/values.ml4
5 files changed, 67 insertions, 20 deletions
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
index 4de597766..1807ae0ec 100644
--- a/checker/indtypes.ml
+++ b/checker/indtypes.ml
@@ -564,14 +564,23 @@ let check_subtyping cumi paramsctxt env inds =
We must produce the substitution σ : [ Var(i) -> Var (i + n) | 0 <= i < n]
and push the constraints [ Var(n) ... Var(2n - 1) |- cst{σ} ], together
with the cumulativity constraints [ cumul_cst ]. *)
- let len = AUContext.size (ACumulativityInfo.univ_context cumi) in
- let inst = Instance.of_array (Array.init len (fun i -> Level.var (i + len))) in
+ let uctx = ACumulativityInfo.univ_context cumi in
+ let len = AUContext.size uctx in
+ let inst = Instance.of_array @@ Array.init len (fun i -> Level.var (i + len)) in
+
let other_context = ACumulativityInfo.univ_context cumi in
let uctx_other = UContext.make (inst, AUContext.instantiate inst other_context) in
- let cumul_context = AUContext.repr (ACumulativityInfo.subtyp_context cumi) in
- let cumul_cst = UContext.constraints cumul_context in
+ let cumul_cst =
+ Array.fold_left_i (fun i csts var ->
+ match var with
+ | Variance.Irrelevant -> csts
+ | Variance.Covariant -> Constraint.add (Level.var i,Le,Level.var (i+len)) csts
+ | Variance.Invariant -> Constraint.add (Level.var i,Eq,Level.var (i+len)) csts)
+ Constraint.empty (ACumulativityInfo.variance cumi)
+ in
let env = Environ.push_context uctx_other env in
let env = Environ.add_constraints cumul_cst env in
+
(* process individual inductive types: *)
Array.iter (fun { mind_user_lc = lc; mind_arity = arity } ->
match arity with
diff --git a/checker/reduction.ml b/checker/reduction.ml
index 29bb8901e..2297c90b3 100644
--- a/checker/reduction.ml
+++ b/checker/reduction.ml
@@ -158,24 +158,17 @@ let compare_stacks f fmind lft1 stk1 lft2 stk2 =
let convert_inductive_instances cv_pb cumi u u' univs =
let len_instance =
Univ.AUContext.size (Univ.ACumulativityInfo.univ_context cumi) in
- let ind_subtypctx = Univ.ACumulativityInfo.subtyp_context cumi in
if not ((len_instance = Univ.Instance.length u) &&
(len_instance = Univ.Instance.length u')) then
anomaly (Pp.str "Invalid inductive subtyping encountered!")
else
- let comp_cst =
- let comp_subst = (Univ.Instance.append u u') in
- Univ.AUContext.instantiate comp_subst ind_subtypctx
- in
+ let variance = Univ.ACumulativityInfo.variance cumi in
let comp_cst =
match cv_pb with
- CONV ->
- let comp_cst' =
- let comp_subst = (Univ.Instance.append u' u) in
- Univ.AUContext.instantiate comp_subst ind_subtypctx
- in
- Univ.Constraint.union comp_cst comp_cst'
- | CUMUL -> comp_cst
+ | CONV ->
+ Univ.Variance.eq_constraints variance u u' Univ.Constraint.empty
+ | CUMUL ->
+ Univ.Variance.leq_constraints variance u u' Univ.Constraint.empty
in
if (Univ.check_constraints comp_cst univs) then () else raise NotConvertible
diff --git a/checker/univ.ml b/checker/univ.ml
index 7d01657df..46b3ce680 100644
--- a/checker/univ.ml
+++ b/checker/univ.ml
@@ -1003,12 +1003,42 @@ end
type abstract_universe_context = AUContext.t
+module Variance =
+struct
+ (** A universe position in the instance given to a cumulative
+ inductive can be the following. Note there is no Contravariant
+ case because [forall x : A, B <= forall x : A', B'] requires [A =
+ A'] as opposed to [A' <= A]. *)
+ type t = Irrelevant | Covariant | Invariant
+
+ let leq_constraint csts variance u u' =
+ match variance with
+ | Irrelevant -> csts
+ | Covariant -> Constraint.add (u, Le, u') csts
+ | Invariant -> Constraint.add (u, Eq, u') csts
+
+ let eq_constraint csts variance u u' =
+ match variance with
+ | Irrelevant -> csts
+ | Covariant | Invariant -> Constraint.add (u, Eq, u') csts
+
+ let leq_constraints variance u u' csts =
+ let len = Array.length u in
+ assert (len = Array.length u' && len = Array.length variance);
+ Array.fold_left3 leq_constraint csts variance u u'
+
+ let eq_constraints variance u u' csts =
+ let len = Array.length u in
+ assert (len = Array.length u' && len = Array.length variance);
+ Array.fold_left3 eq_constraint csts variance u u'
+end
+
module CumulativityInfo =
struct
- type t = universe_context * universe_context
+ type t = universe_context * Variance.t array
let univ_context (univcst, subtypcst) = univcst
- let subtyp_context (univcst, subtypcst) = subtypcst
+ let variance (univs, variance) = variance
end
diff --git a/checker/univ.mli b/checker/univ.mli
index 21c94d952..8c0685e0b 100644
--- a/checker/univ.mli
+++ b/checker/univ.mli
@@ -218,12 +218,25 @@ end
type abstract_universe_context = AUContext.t
+module Variance :
+sig
+ (** A universe position in the instance given to a cumulative
+ inductive can be the following. Note there is no Contravariant
+ case because [forall x : A, B <= forall x : A', B'] requires [A =
+ A'] as opposed to [A' <= A]. *)
+ type t = Irrelevant | Covariant | Invariant
+
+ val leq_constraints : t array -> Instance.t constraint_function
+ val eq_constraints : t array -> Instance.t constraint_function
+end
+
+
module ACumulativityInfo :
sig
type t
val univ_context : t -> abstract_universe_context
- val subtyp_context : t -> abstract_universe_context
+ val variance : t -> Variance.t array
end
diff --git a/checker/values.ml b/checker/values.ml
index 313067cb6..55e1cdb6f 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -110,10 +110,12 @@ let v_cstrs =
(v_tuple "univ_constraint"
[|v_level;v_enum "order_request" 3;v_level|]))
+let v_variance = v_enum "variance" 3
+
let v_instance = Annot ("instance", Array v_level)
let v_context = v_tuple "universe_context" [|v_instance;v_cstrs|]
let v_abs_context = v_context (* only for clarity *)
-let v_abs_cum_info = v_tuple "cumulativity_info" [|v_abs_context; v_context|]
+let v_abs_cum_info = v_tuple "cumulativity_info" [|v_abs_context; Array v_variance|]
let v_context_set = v_tuple "universe_context_set" [|v_hset v_level;v_cstrs|]
(** kernel/term *)