Skip to content

Commit

Permalink
Tactics: introduce postprocess_type
Browse files Browse the repository at this point in the history
The attributes postprocess_with and postprocess_for_extraction_with
allow a user to use a custom tactic to transform a definition after
typechecking (to something equivalent); usually the used tactic does
some controlled normalization and unfolding. However, the tactic is only
applied on the body of a letbinding and not the type.

This patch adds an attribute `postprocess_type` to indicate that the
tactic should be applied to the type too.
  • Loading branch information
mtzguido committed Feb 7, 2025
1 parent 33c02b7 commit 02ede3d
Show file tree
Hide file tree
Showing 8 changed files with 751 additions and 10 deletions.
14 changes: 12 additions & 2 deletions examples/tactics/Postprocess.fst
Original file line number Diff line number Diff line change
Expand Up @@ -19,11 +19,21 @@ open FStar.Tactics.V2

assume val foo : int -> int
assume val lem : unit -> Lemma (foo 1 == foo 2)
let tau () = apply_lemma (`lem)
let tau () =
grewrite (`(foo 1)) (`(foo 2));
trefl ();
apply_lemma (`lem);
()

[@@(postprocess_with tau)]
[@@postprocess_with tau]
let x : int = foo 1

[@@postprocess_with tau]
let x' : (z:int{z == foo 1}) = foo 1

[@@postprocess_with tau; postprocess_type]
let x'' : (z:int{z == foo 1}) = foo 1

[@@(postprocess_for_extraction_with tau)]
let y : int = foo 1

Expand Down
1 change: 1 addition & 0 deletions src/parser/FStarC.Parser.Const.fst
Original file line number Diff line number Diff line change
Expand Up @@ -448,6 +448,7 @@ let bv_lid = lid_of_path (["FStar"; "Stubs"; "Reflection"; "Types"; "bv"]) FStar
let fv_lid = lid_of_path (["FStar"; "Stubs"; "Reflection"; "Types"; "fv"]) FStarC.Range.dummyRange
let norm_step_lid = psconst "norm_step"
let postprocess_with = p2l ["FStar"; "Tactics"; "Effect"; "postprocess_with"]
let postprocess_type = p2l ["FStar"; "Tactics"; "Effect"; "postprocess_type"]
let preprocess_with = p2l ["FStar"; "Tactics"; "Effect"; "preprocess_with"]
let postprocess_extr_with = p2l ["FStar"; "Tactics"; "Effect"; "postprocess_for_extraction_with"]
let term_lid = p2l ["FStar"; "Stubs"; "Reflection"; "Types"; "term"]
Expand Down
37 changes: 29 additions & 8 deletions src/typechecker/FStarC.TypeChecker.Tc.fst
Original file line number Diff line number Diff line change
Expand Up @@ -255,16 +255,27 @@ let proc_check_with (attrs:list attribute) (kont : unit -> 'a) : 'a =
| _ -> failwith "ill-formed `check_with`"

let handle_postprocess_with_attr (env:Env.env) (ats:list attribute)
: (list attribute & option term)
: (list attribute & option (term & bool))
= (* Extract the postprocess_with *)
match U.extract_attr' PC.postprocess_with ats with
| None -> ats, None
| Some (ats, [tau, None]) ->
ats, Some tau
let ats, on_type_too =
match U.extract_attr' PC.postprocess_type ats with
| None -> ats, false
| Some (ats, []) -> ats, true
| _ ->
Errors.log_issue env Errors.Warning_UnrecognizedAttribute [
text <| BU.format1 "Ill-formed application of `%s`" (show PC.postprocess_type);
];
ats, false
in
ats, Some (tau, on_type_too)
| Some (ats, args) ->
Errors.log_issue env Errors.Warning_UnrecognizedAttribute
(BU.format1 "Ill-formed application of `%s`" (show PC.postprocess_with));
ats, None
Errors.log_issue env Errors.Warning_UnrecognizedAttribute [
text <| BU.format1 "Ill-formed application of `%s`" (show PC.postprocess_with);
];
ats, None

let store_sigopts (se:sigelt) : sigelt =
{ se with sigopts = Some (Options.get_vconfig ()) }
Expand Down Expand Up @@ -459,14 +470,24 @@ let tc_sig_let env r se lbs lids : list sigelt & list sigelt & Env.env =
(* remove the postprocess_with, if any *)
let se = { se with sigattrs = attrs } in

let postprocess_lb (tau:term) (lb:letbinding) : letbinding =
let postprocess_lb (tau:term) (on_type_too : bool) (lb:letbinding) : letbinding =
let s, univnames = SS.univ_var_opening lb.lbunivs in
let lbdef = SS.subst s lb.lbdef in
let lbtyp = SS.subst s lb.lbtyp in
let env = Env.push_univ_vars env univnames in

let lbdef = Env.postprocess env tau lbtyp lbdef in
let lbdef = SS.close_univ_vars univnames lbdef in
{ lb with lbdef = lbdef }

let lbtyp =
if on_type_too then
let u = env.universe_of env lbtyp in
let lbtyp = Env.postprocess env tau lbtyp lbtyp in
SS.close_univ_vars univnames lbtyp
else lbtyp
in

{ lb with lbdef = lbdef; lbtyp = lbtyp }
in
let env' =
match (SS.compress e).n with
Expand Down Expand Up @@ -494,7 +515,7 @@ let tc_sig_let env r se lbs lids : list sigelt & list sigelt & Env.env =
// Postprocess the letbindings with the tactic, if any
let lbs = (fst lbs,
(match post_tau with
| Some tau -> List.map (postprocess_lb tau) (snd lbs)
| Some (tau, on_type_too) -> List.map (postprocess_lb tau on_type_too) (snd lbs)
| None -> (snd lbs)))
in

Expand Down
2 changes: 2 additions & 0 deletions tests/tactics/Makefile
Original file line number Diff line number Diff line change
Expand Up @@ -12,3 +12,5 @@ Hacl_Meta.ml:
Hacl.Meta.Use.fst.test: Hacl.Meta.Use.fst Hacl_Meta.ml
$(FSTAR) $< --load Hacl.Meta
touch $@

$(OUTPUT_DIR)/Postprocess.fst.output: FSTAR_ARGS+=--dump_module Postprocess --ugly
149 changes: 149 additions & 0 deletions tests/tactics/Postprocess.fst
Original file line number Diff line number Diff line change
@@ -0,0 +1,149 @@
(*
Copyright 2008-2018 Microsoft Research
Licensed under the Apache License, Version 2.0 (the "License");
you may not use this file except in compliance with the License.
You may obtain a copy of the License at
http://www.apache.org/licenses/LICENSE-2.0
Unless required by applicable law or agreed to in writing, software
distributed under the License is distributed on an "AS IS" BASIS,
WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied.
See the License for the specific language governing permissions and
limitations under the License.
*)
module Postprocess

open FStar.Tactics.V2

assume val foo : int -> int
assume val lem : unit -> Lemma (foo 1 == foo 2)
let tau () =
grewrite (`(foo 1)) (`(foo 2));
trefl ();
apply_lemma (`lem);
()

[@@postprocess_with tau]
let x : int = foo 1

[@@postprocess_with tau]
let x' : (z:int{z == foo 1}) = foo 1

[@@postprocess_with tau; postprocess_type]
let x'' : (z:int{z == foo 1}) = foo 1

[@@(postprocess_for_extraction_with tau)]
let y : int = foo 1

let _ = assert (x == foo 2)
let _ = assert (y == foo 1) // but `foo 2` in extracted code

(* More hardcore transformations *)

noeq
type t1 =
| A1 : t1
| B1 : int -> t1
| C1 : (int -> t1) -> t1

noeq
type t2 =
| A2 : t2
| B2 : int -> t2
| C2 : (int -> t2) -> t2

let rec lift : t1 -> t2 =
function
| A1 -> A2
| B1 i -> B2 i
| C1 f -> C2 (fun x -> lift (f x))

let lemA () : Lemma (lift A1 == A2) = ()
let lemB x : Lemma (lift (B1 x) == (B2 x)) = ()
let lemC ($f: int -> t1) : Lemma (lift (C1 f) == C2 (fun x -> lift (f x))) by (compute ()) = ()

(* These could really be polymorphic *)
let congB #i #j (_ : squash (i == j)) : Lemma (B2 i == B2 j) = ()
let congC #f #g (_ : squash (f == g)) : Lemma (C2 f == C2 g) = ()

let xx = C1 (function
| 0 -> A1
| 5 -> B1 42
| x -> B1 24)

open FStar.FunctionalExtensionality

let q_as_lem (#a:Type) (#b:a -> Type) (p:squash (forall x. b x)) (x:a)
: Lemma (b x)
= ()

let congruence_fun #a (#b:a -> Type) (f g:(x:a -> b x)) (x:squash (forall x. f x == g x)) :
Lemma (ensures (fun (x:a) -> f x) == (fun (x:a) -> g x)) =
assert ((fun (x:a) -> f x) == (fun (x:a) -> g x))
by (l_to_r [quote (q_as_lem x)];
trefl())

let apply_feq_lem #a #b ($f $g : a -> b) : Lemma (requires (forall x. f x == g x))
(ensures ((fun x -> f x) == (fun x -> g x))) = congruence_fun f g ()

let fext () = apply_lemma (`apply_feq_lem); dismiss (); ignore (forall_intros ())

let _onL a b c (_ : squash (a == b)) (_ : squash (b == c)) : Lemma (a == c) = ()
let onL () = apply_lemma (`_onL)

// invariant: takes goals of shape squash (E == ?u) and solves them
let rec push_lifts' (u:unit) : Tac unit =
match term_as_formula (cur_goal ()) with
| Comp (Eq _) lhs rhs ->
begin
match inspect lhs with
| Tv_App h t ->
begin match inspect h with
| Tv_FVar fv ->
if fv_to_string fv = `%lift
then case_analyze (fst t)
else fail "not a lift (1)"
| _ -> fail "not a lift (2)"
end

| Tv_Abs _ _ ->
fext ();
push_lifts' ()

| _ -> fail "not a lift (3)"
end
| _ ->
fail "not an equality"
and case_analyze (lhs:term) : Tac unit =
let ap l =
onL (); apply_lemma l
in
let lhs = norm_term [weak;hnf;primops;delta] lhs in
let head, args = collect_app lhs in
begin match inspect head with
| Tv_FVar fv ->
if fv_to_string fv = `%A1 then (apply_lemma (`lemA))
else if fv_to_string fv = `%B1 then (ap (`lemB); apply_lemma (`congB); push_lifts' ())
else if fv_to_string fv = `%C1 then (ap (`lemC); apply_lemma (`congC); push_lifts' ())
else (tlabel "unknown fv"; trefl ())
| _ ->
tlabel "head unk";
trefl ()
end

let push_lifts () : Tac unit =
push_lifts' ();
(* dump "after"; *)
()

//#push-options "--tactic_trace_d 2"
[@@(postprocess_with push_lifts)]
let yy = lift xx

[@@(postprocess_with push_lifts)]
let zz1 = lift (C1 (fun y -> (C1 (fun x -> A1))))

[@@(postprocess_for_extraction_with push_lifts)]
let zz2 = lift (C1 (fun y -> (C1 (fun x -> A1))))
Loading

0 comments on commit 02ede3d

Please sign in to comment.