Skip to content

Commit

Permalink
Suppress erasable warning for props.
Browse files Browse the repository at this point in the history
  • Loading branch information
gebner committed Jan 10, 2025
1 parent f2ab935 commit 292723c
Show file tree
Hide file tree
Showing 5 changed files with 42 additions and 7 deletions.
7 changes: 7 additions & 0 deletions src/typechecker/FStarC.TypeChecker.Env.fst
Original file line number Diff line number Diff line change
Expand Up @@ -1056,6 +1056,13 @@ let rec non_informative env t =
| Tm_meta {tm} -> non_informative env tm
| _ -> None

let rec non_informative_sort t =
match (U.unrefine t).n with
| Tm_fvar fv when fv_eq_lid fv Const.prop_lid -> true
| Tm_arrow {comp=c} -> non_informative_sort (comp_result c)
| Tm_meta {tm} -> non_informative_sort tm
| _ -> false

let num_effect_indices env name r =
let sig_t = name |> lookup_effect_lid env |> SS.compress in
match sig_t.n with
Expand Down
5 changes: 5 additions & 0 deletions src/typechecker/FStarC.TypeChecker.Env.fsti
Original file line number Diff line number Diff line change
Expand Up @@ -346,6 +346,11 @@ val fv_has_erasable_attr : env -> fv -> bool
and any `x: t ...` can be erased to `i`. *)
val non_informative : env -> typ -> option term

(* `non_informative_sort t` is `true` if the type family `t: ... -> Type` only ranges over noninformative types,
i.e., any `x: s ...` such that `s ... : t ...` can be erased.
(practically, this means that `t` is of the form `... -> prop`) *)
val non_informative_sort : typ -> bool

val try_lookup_effect_lid : env -> lident -> option term
val lookup_effect_lid : env -> lident -> term
val lookup_effect_abbrev : env -> universes -> lident -> option (binders & comp)
Expand Down
15 changes: 15 additions & 0 deletions src/typechecker/FStarC.TypeChecker.Normalize.fst
Original file line number Diff line number Diff line change
Expand Up @@ -2821,6 +2821,21 @@ let non_info_norm env t =
in
non_informative env (normalize steps env t)

let non_info_sort_norm env t =
let steps = [UnfoldUntil (Env.delta_depth_of_fv env (S.fvconst PC.prop_lid)); // do not unfold prop
AllowUnboundUniverses;
EraseUniverses;
Primops;
Beta; Iota;
HNF;
(* We could use Weak too were it not that we need
* to descend in the codomain of arrows. *)
Unascribe; //remove ascriptions
ForExtraction //and refinement types
]
in
non_informative_sort (normalize steps env t)

(*
* Ghost T to Pure T promotion
*
Expand Down
1 change: 1 addition & 0 deletions src/typechecker/FStarC.TypeChecker.Normalize.fsti
Original file line number Diff line number Diff line change
Expand Up @@ -35,6 +35,7 @@ val unfold_whnf': steps -> Env.env -> term -> term
val unfold_whnf: Env.env -> term -> term
val reduce_uvar_solutions:Env.env -> term -> term
val non_info_norm: Env.env -> term -> option term
val non_info_sort_norm: Env.env -> term -> bool

(*
* The maybe versions of ghost_to_pure only promote
Expand Down
21 changes: 14 additions & 7 deletions src/typechecker/FStarC.TypeChecker.Quals.fst
Original file line number Diff line number Diff line change
Expand Up @@ -210,13 +210,20 @@ let check_erasable env quals (r:Range.range) se =
let has_iface_val = match DsEnv.iface_decls (Env.dsenv env) (Env.current_module env) with
| Some iface_decls -> iface_decls |> BU.for_some (Parser.AST.decl_is_val (ident_of_lid lbname.fv_name.v))
| None -> false in
let _, body, _ = U.abs_formals lb.lbdef in
if has_iface_val && Some? (non_info_norm_weak env body) then log_issue lbname Error_MustEraseMissing [
text (BU.format1 "Values of type `%s` will be erased during extraction, \
but its interface hides this fact." (show lbname));
text (BU.format1 "Add the `erasable` \
attribute to the `val %s` declaration for this symbol in the interface" (show lbname));
]
let val_decl = Env.try_lookup_val_decl env lbname.fv_name.v in
if has_iface_val && Some? val_decl then
let _, body, _ = U.abs_formals lb.lbdef in
let Some ((us, t), _) = val_decl in
let known_to_be_erasable =
let env = Env.push_univ_vars env us in
N.non_info_sort_norm env t in
if not known_to_be_erasable && Some? (non_info_norm_weak env body) then
log_issue lbname Error_MustEraseMissing [
text (BU.format1 "Values of type `%s` will be erased during extraction, \
but its interface hides this fact." (show lbname));
text (BU.format1 "Add the `erasable` \
attribute to the `val %s` declaration for this symbol in the interface" (show lbname));
]
| _ -> ()
end;
if se_has_erasable_attr
Expand Down

0 comments on commit 292723c

Please sign in to comment.