Skip to content

Commit

Permalink
Ltac2: support "head" reduction flag
Browse files Browse the repository at this point in the history
Close coq#18209
  • Loading branch information
SkySkimmer committed Nov 7, 2023
1 parent f98da71 commit fcf4ed7
Show file tree
Hide file tree
Showing 7 changed files with 28 additions and 4 deletions.
1 change: 1 addition & 0 deletions plugins/ltac2/g_ltac2.mlg
Original file line number Diff line number Diff line change
Expand Up @@ -773,6 +773,7 @@ GRAMMAR EXTEND Gram
| IDENT "cofix" -> { CAst.make ~loc @@ QCofix }
| IDENT "zeta" -> { CAst.make ~loc @@ QZeta }
| IDENT "delta"; d = delta_flag -> { d }
| IDENT "head" -> { CAst.make ~loc @@ QHead }
] ]
;
refglobal:
Expand Down
1 change: 1 addition & 0 deletions plugins/ltac2/tac2qexpr.mli
Original file line number Diff line number Diff line change
Expand Up @@ -128,6 +128,7 @@ type red_flag_r =
| QZeta
| QConst of reference or_anti list CAst.t
| QDeltaBut of reference or_anti list CAst.t
| QHead

type red_flag = red_flag_r CAst.t

Expand Down
9 changes: 9 additions & 0 deletions plugins/ltac2/tac2quote.ml
Original file line number Diff line number Diff line change
Expand Up @@ -357,6 +357,7 @@ let make_red_flag l =
| [] -> red
| {v=flag} :: lf ->
let red = match flag with
| QHead -> { red with rStrength = Head }
| QBeta -> { red with rBeta = true }
| QMatch -> { red with rMatch = true }
| QFix -> { red with rFix = true }
Expand Down Expand Up @@ -388,10 +389,18 @@ let of_reference r =
in
of_anti of_ref r

let of_strength ?loc s =
let s = let open Genredexpr in match s with
| Norm -> std_core "Norm"
| Head -> std_core "Head"
in
constructor ?loc s []

let of_strategy_flag {loc;v=flag} =
let open Genredexpr in
let flag = make_red_flag flag in
CAst.make ?loc @@ CTacRec (None, [
std_proj "rStrength", of_strength ?loc flag.rStrength;
std_proj "rBeta", of_bool ?loc flag.rBeta;
std_proj "rMatch", of_bool ?loc flag.rMatch;
std_proj "rFix", of_bool ?loc flag.rFix;
Expand Down
9 changes: 7 additions & 2 deletions plugins/ltac2/tac2stdlib.ml
Original file line number Diff line number Diff line change
Expand Up @@ -60,10 +60,15 @@ let to_clause v = match Value.to_tuple v with

let clause = make_to_repr to_clause

let to_red_strength = function
| ValInt 0 -> Norm
| ValInt 1 -> Head
| _ -> assert false

let to_red_flag v = match Value.to_tuple v with
| [| beta; iota; fix; cofix; zeta; delta; const |] ->
| [| strength; beta; iota; fix; cofix; zeta; delta; const |] ->
{
rStrength = Norm; (* specifying rStrength is not yet supported *)
rStrength = to_red_strength strength;
rBeta = Value.to_bool beta;
rMatch = Value.to_bool iota;
rFix = Value.to_bool fix;
Expand Down
5 changes: 5 additions & 0 deletions test-suite/output/reduction.v
Original file line number Diff line number Diff line change
Expand Up @@ -25,3 +25,8 @@ Axiom ignore : forall {T}, T -> unit.
Eval simpl head in ignore (fun x => 1 + x).
Eval cbn head in ignore (fun x => 1 + x).
Eval lazy head in ignore (fun x => 1 + x).

Require Import Ltac2.Ltac2.

Ltac2 Eval eval lazy in (2 + 2).
Ltac2 Eval eval lazy head in (2 + 2).
4 changes: 2 additions & 2 deletions user-contrib/Ltac2/RedFlags.v
Original file line number Diff line number Diff line change
Expand Up @@ -19,12 +19,12 @@ End Notations.

Ltac2 none := {
rBeta := false; rMatch := false; rFix := false; rCofix := false;
rZeta := false; rDelta := false; rConst := [];
rZeta := false; rDelta := false; rConst := []; rStrength := Norm;
}.

Ltac2 all : t := {
rBeta := true; rMatch := true; rFix := true; rCofix := true;
rZeta := true; rDelta := true; rConst := [];
rZeta := true; rDelta := true; rConst := []; rStrength := Norm;
}.

Ltac2 beta : t := red_flags:(beta).
Expand Down
3 changes: 3 additions & 0 deletions user-contrib/Ltac2/Std.v
Original file line number Diff line number Diff line change
Expand Up @@ -43,7 +43,10 @@ Ltac2 Type reference := [
| ConstructRef (constructor)
].

Ltac2 Type strength := [ Norm | Head ].

Ltac2 Type red_flags := {
rStrength : strength;
rBeta : bool;
rMatch : bool;
rFix : bool;
Expand Down

0 comments on commit fcf4ed7

Please sign in to comment.