Skip to content

Commit 3b94c02

Browse files
authored
Merge pull request #593 from Tragicus/cspr
Update cs hook
2 parents 7c9d321 + b4660ad commit 3b94c02

File tree

5 files changed

+151
-76
lines changed

5 files changed

+151
-76
lines changed

Changelog.md

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -23,6 +23,7 @@ Requires Elpi 1.18.2 and Coq 8.19.
2323

2424
### Apps
2525
- TC: avoid declaring options twice (could make vscoq2 fail)
26+
- CS: `cs` now takes a context, a term that is the projection of some structure applied to the parameters of the structure, a term to put a structure on and the solution to return
2627

2728
## [2.0.1] - 29/12/2023
2829

apps/cs/src/coq_elpi_cs_hook.mlg

Lines changed: 47 additions & 25 deletions
Original file line numberDiff line numberDiff line change
@@ -10,48 +10,72 @@ module Evarconv = Evarconv
1010
module Evarconv_hacked = Evarconv_hacked
1111

1212

13-
let elpi_cs_hook program env sigma lhs rhs =
14-
let (lhead, largs) = EConstr.decompose_app sigma lhs in
15-
let (rhead, rargs) = EConstr.decompose_app sigma rhs in
16-
if (EConstr.isConst sigma lhead && Structures.Structure.is_projection (fst (EConstr.destConst sigma lhead))) ||
17-
(EConstr.isConst sigma rhead && Structures.Structure.is_projection (fst (EConstr.destConst sigma rhead)))
18-
then begin
13+
let elpi_cs_hook program env sigma (t1, sk1) (t2, sk2) =
1914
let loc = API.Ast.Loc.initial "(unknown)" in
2015
let atts = [] in
21-
(*let sigma, ty = Typing.type_of env sigma lhs in*)
22-
let sigma, (ty, _) = Evarutil.new_type_evar env sigma Evd.univ_flexible in
23-
let { Coqlib.eq } = Coqlib.build_coq_eq_data () in
24-
let sigma, eq = EConstr.fresh_global env sigma eq in
25-
let t = EConstr.mkApp (eq,[|ty;lhs;rhs|]) in
26-
let sigma, goal = Evarutil.new_evar env sigma t in
16+
let () = Feedback.msg_info (Pp.str "cs hook start") in
17+
let (proji, u), arg =
18+
match Termops.global_app_of_constr sigma t1 with
19+
| (Names.GlobRef.ConstRef proji, u), arg -> (proji, u), arg
20+
| (proji, _), _ -> let () = Feedback.msg_info Pp.(str "proj is not const" ++ Names.GlobRef.print proji) in raise Not_found in
21+
let () = Feedback.msg_info (Pp.str "cs hook got proj") in
22+
let structure = Structures.Structure.find_from_projection proji in
23+
let () = Feedback.msg_info (Pp.str "cs hook got structure") in
24+
let params1, c1, extra_args1 =
25+
match arg with
26+
| Some c -> (* A primitive projection applied to c *)
27+
let ty = Retyping.get_type_of ~lax:true env sigma c in
28+
let (i,u), ind_args =
29+
(* Are we sure that ty is not an evar? *)
30+
Inductiveops.find_mrectype env sigma ty
31+
in ind_args, c, sk1
32+
| None ->
33+
match Reductionops.Stack.strip_n_app structure.nparams sk1 with
34+
| Some (params1, c1, extra_args1) -> (Option.get @@ Reductionops.Stack.list_of_app_stack params1), c1, extra_args1
35+
| _ -> raise Not_found in
36+
let () = Feedback.msg_info Pp.(str "cs hook got params & arg " ++ int (List.length sk2) ++ str " " ++ int (List.length extra_args1)) in
37+
let sk2, extra_args2 =
38+
if Reductionops.Stack.args_size sk2 = Reductionops.Stack.args_size extra_args1 then [], sk2
39+
else match Reductionops.Stack.strip_n_app (Reductionops.Stack.args_size sk2 - Reductionops.Stack.args_size extra_args1 - 1) sk2 with
40+
| None -> raise Not_found
41+
| Some (l',el,s') -> ((Option.get @@ Reductionops.Stack.list_of_app_stack l') @ [el],s') in
42+
let rhs = Reductionops.Stack.zip sigma (t2, Reductionops.Stack.append_app_list sk2 Reductionops.Stack.empty) in
43+
let sigma, goal = Evarutil.new_evar env sigma (Retyping.get_type_of env sigma c1) in
2744
let goal_evar, _ = EConstr.destEvar sigma goal in
2845
let query ~depth state =
2946
let state, (loc, q), gls =
30-
Coq_elpi_HOAS.goals2query sigma [goal_evar] loc ~main:(Coq_elpi_HOAS.Solve [])
47+
Coq_elpi_HOAS.goals2query sigma [goal_evar] loc ~main:(Coq_elpi_HOAS.Solve [EConstr.mkApp (EConstr.mkConstU (proji, EConstr.EInstance.empty), Array.of_list params1); rhs])
3148
~in_elpi_tac_arg:Coq_elpi_arg_HOAS.in_elpi_tac_econstr ~depth state in
3249
let state, qatts = atts2impl loc Summary.Stage.Interp ~depth state atts q in
3350
let state = API.State.set Coq_elpi_builtins.tactic_mode state true in
3451
state, (loc, qatts), gls
35-
in
52+
in let () = Feedback.msg_info Pp.(str "compile solver") in
3653
match Interp.get_and_compile program with
3754
| None -> None
3855
| Some (cprogram, _) ->
39-
match Interp.run ~static_check:false cprogram (`Fun query) with
40-
| API.Execute.Success solution ->
56+
let () = Feedback.msg_info Pp.(str "run solver\n") in
57+
begin try match Interp.run ~static_check:false cprogram (`Fun query) with
58+
| API.Execute.Success solution -> let () = Feedback.msg_info Pp.(str "found solution\n") in
4159
let gls = Evar.Set.singleton goal_evar in
4260
let sigma, _, _ = Coq_elpi_HOAS.solution2evd sigma solution gls in
43-
let ty_evar, _ = EConstr.destEvar sigma ty in
44-
Some (Evd.remove (Evd.remove sigma ty_evar) goal_evar)
61+
if Evd.is_defined sigma goal_evar then
62+
let lhs = Reductionops.Stack.zip sigma (EConstr.mkConstU (proji, EConstr.EInstance.empty), Reductionops.Stack.append_app_list (params1 @ [goal]) Reductionops.Stack.empty) in
63+
let lhs = Reductionops.whd_const proji env sigma lhs in
64+
let lhs = Reductionops.whd_betaiotazeta env sigma lhs in
65+
let hh, sk1 = EConstr.decompose_app sigma lhs in
66+
let () = Feedback.msg_info Pp.(str "aha" ++ Printer.pr_econstr_env env sigma lhs) in
67+
let h2, sk2 = EConstr.decompose_app sigma rhs in
68+
let _, params = EConstr.decompose_app sigma (Retyping.get_type_of env sigma goal) in
69+
Some (sigma, (hh, h2), goal, [], (Array.to_list params, params1), (Array.to_list sk1, Array.to_list sk2), (extra_args1, extra_args2), c1, (None, rhs))
70+
else None
4571
| API.Execute.NoMoreSteps
46-
| API.Execute.Failure -> None
47-
| exception (Coq_elpi_utils.LtacFail (level, msg)) -> None
48-
end
49-
else None
72+
| API.Execute.Failure -> let () = Feedback.msg_info Pp.(str "solver failed\n") in None
73+
with e -> let () = Feedback.msg_info Pp.(str "solver crashed\n") in raise e end
74+
| exception e -> let () = Feedback.msg_info Pp.(str "compiler crashed\n") in raise e
5075

5176
let add_cs_hook =
5277
let cs_hook_program = Summary.ref ~name:"elpi-cs" None in
5378
let cs_hook env sigma proj pat =
54-
Feedback.msg_info (Pp.str "run");
5579
match !cs_hook_program with
5680
| None -> None
5781
| Some h -> elpi_cs_hook h env sigma proj pat in
@@ -60,8 +84,6 @@ let add_cs_hook =
6084
let inCs =
6185
let cache program =
6286
cs_hook_program := Some program;
63-
Feedback.msg_info (Pp.str "activate");
64-
6587
Evarconv_hacked.activate_hook ~name in
6688
let open Libobject in
6789
declare_object

apps/cs/src/evarconv_hacked.ml

Lines changed: 82 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -488,7 +488,10 @@ let rec ise_stack2 no_app env evd f sk1 sk2 =
488488
|_, _ -> fail (UnifFailure (i,(* Maybe improve: *) NotSameHead))
489489
in ise_rev_stack2 false evd (List.rev sk1) (List.rev sk2)
490490

491-
type hook = Environ.env -> Evd.evar_map -> EConstr.t -> EConstr.t -> Evd.evar_map option
491+
type hook = Environ.env -> Evd.evar_map -> (EConstr.t * Reductionops.Stack.t) -> (EConstr.t * Reductionops.Stack.t) -> (Evd.evar_map * (EConstr.t * EConstr.t) * EConstr.t * EConstr.t list *
492+
(EConstr.t list * EConstr.t list) * (EConstr.t list * EConstr.t list) *
493+
(Reductionops.Stack.t * Reductionops.Stack.t) * EConstr.constr *
494+
(int option * EConstr.constr)) option
492495

493496
let all_hooks = ref (CString.Map.empty : hook CString.Map.t)
494497

@@ -591,6 +594,38 @@ let infer_conv_noticing_evars ~pb ~ts env sigma t1 t2 =
591594
if !has_evar then None
592595
else Some (UnifFailure (sigma, UnifUnivInconsistency e))
593596

597+
let pr_econstr = ref (fun _ _ _ -> Pp.str "unable to print econstr")
598+
599+
(* TODO: move to a proper place *)
600+
let rec split_at n acc l =
601+
if n = 0 then (List.rev acc, l)
602+
else match l with
603+
| [] -> (List.rev acc, l)
604+
| h :: t -> split_at (n-1) (h :: acc) t
605+
let split_at n l = split_at n [] l
606+
607+
let try_simplify_proj_construct flags env evd v k sk =
608+
match k with (* try unfolding an applied projection on the rhs *)
609+
| Proj (p, _, c) -> begin
610+
let c = whd_all env evd c in (* reduce argument *)
611+
try let (hd, args) = destApp evd c in
612+
if isConstruct evd hd then Some (whd_nored_state env evd (args.(Projection.npars p + Projection.arg p), sk))
613+
else None
614+
with _ -> None
615+
end
616+
| Const (cn, _) when Structures.Structure.is_projection cn -> begin
617+
match split_at (Structures.Structure.projection_nparams cn) (Option.default [] (Stack.list_of_app_stack sk)) with
618+
| (_, []) -> None
619+
| (_, c :: _) -> begin
620+
let c = whd_all env evd c in
621+
try let (hd, _) = destApp evd c in
622+
if isConstruct evd hd then
623+
Some (whd_betaiota_deltazeta_for_iota_state flags.open_ts env evd (v,sk))
624+
else None
625+
with _ -> None
626+
end end
627+
| _ -> None
628+
594629
let rec evar_conv_x flags env evd pbty term1 term2 =
595630
let term1 = whd_head_evar evd term1 in
596631
let term2 = whd_head_evar evd term2 in
@@ -923,7 +958,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
923958
flex_maybeflex false ev2 appr2 appr1 v1
924959

925960
| MaybeFlexible v1, MaybeFlexible v2 -> begin
926-
match EConstr.kind evd term1, EConstr.kind evd term2 with
961+
let k1 = EConstr.kind evd term1 in
962+
let k2 = EConstr.kind evd term2 in
963+
begin
964+
match k1, k2 with
927965
| LetIn (na1,b1,t1,c'1), LetIn (na2,b2,t2,c'2) ->
928966
let f1 i = (* FO *)
929967
ise_and i
@@ -996,16 +1034,22 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
9961034
| None ->
9971035
UnifFailure (i,NotSameHead)
9981036
and f2 i =
999-
(try
1037+
(match try_simplify_proj_construct flags env evd v1 k1 sk1, try_simplify_proj_construct flags env evd v2 k2 sk2 with
1038+
| Some x1, Some x2 -> UnifFailure (i,NoCanonicalStructure)
1039+
| Some x1, None -> UnifFailure (i,NoCanonicalStructure)
1040+
| None, Some x2 -> UnifFailure (i,NoCanonicalStructure)
1041+
| _, _ -> (try
10001042
if not flags.with_cs then raise Not_found
10011043
else conv_record flags env
10021044
(try check_conv_record env i appr1 appr2
1003-
with Not_found -> check_conv_record env i appr2 appr1)
1004-
with Not_found ->
1005-
let sigma = i in
1006-
match apply_hooks env sigma (Stack.zip sigma appr1) (Stack.zip sigma appr2) with
1007-
| Some sigma -> Success sigma
1008-
| None -> UnifFailure (i,NoCanonicalStructure))
1045+
with Not_found -> begin match (apply_hooks env i appr1 appr2) with
1046+
| Some r -> r
1047+
| None -> begin try check_conv_record env i appr2 appr1
1048+
with Not_found -> begin match (apply_hooks env i appr2 appr1) with
1049+
| Some r -> r
1050+
| None -> raise Not_found
1051+
end end end)
1052+
with Not_found -> UnifFailure (i,NoCanonicalStructure)))
10091053
and f3 i =
10101054
(* heuristic: unfold second argument first, exception made
10111055
if the first argument is a beta-redex (expand a constant
@@ -1044,7 +1088,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
10441088
flags.open_ts env i (v2,sk2))
10451089
in
10461090
ise_try evd [f1; f2; f3]
1047-
end
1091+
end end
10481092

10491093
| Rigid, Rigid when EConstr.isLambda evd term1 && EConstr.isLambda evd term2 ->
10501094
let (na1,c1,c'1) = EConstr.destLambda evd term1 in
@@ -1062,14 +1106,21 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
10621106
| Rigid, Flexible ev2 -> flex_rigid false ev2 appr2 appr1
10631107

10641108
| MaybeFlexible v1, Rigid ->
1109+
let k1 = EConstr.kind evd term1 in begin
1110+
let () = debug_unification (fun () -> Pp.(v 0 (str "v1 maybeflexible against rigid" ++ !pr_econstr env evd v1 ++ cut ()))) in
1111+
match try_simplify_proj_construct flags env evd v1 k1 sk1 with
1112+
| Some x1 -> evar_eqappr_x flags env evd pbty x1 appr2
1113+
| None ->
10651114
let f3 i =
10661115
(try
10671116
if not flags.with_cs then raise Not_found
1068-
else conv_record flags env (check_conv_record env i appr1 appr2)
1069-
with Not_found -> let sigma = i in
1070-
match apply_hooks env sigma (Stack.zip sigma appr1) (Stack.zip sigma appr2) with
1071-
| Some sigma -> Success sigma
1072-
| None -> UnifFailure (i,NoCanonicalStructure))
1117+
else conv_record flags env (
1118+
try check_conv_record env i appr1 appr2 with
1119+
| Not_found -> begin match apply_hooks env i appr1 appr2 with
1120+
| Some r -> r
1121+
| None -> raise Not_found
1122+
end)
1123+
with Not_found -> UnifFailure (i,NoCanonicalStructure))
10731124

10741125
and f4 i =
10751126
evar_eqappr_x flags env i pbty
@@ -1078,22 +1129,32 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
10781129
appr2
10791130
in
10801131
ise_try evd [f3; f4]
1132+
end
10811133

10821134
| Rigid, MaybeFlexible v2 ->
1135+
let k2 = EConstr.kind evd term2 in begin
1136+
let () = debug_unification (fun () -> Pp.(v 0 (str "rigid against v2 maybeflexible" ++ !pr_econstr env evd v2 ++ cut ()))) in
1137+
match try_simplify_proj_construct flags env evd v2 k2 sk2 with
1138+
| Some x2 -> let () = debug_unification (fun () -> Pp.(v 0 (str "reduced to" ++ !pr_econstr env evd (let (x, _) = x2 in x)))) in evar_eqappr_x flags env evd pbty appr1 x2
1139+
| None ->
10831140
let f3 i =
10841141
(try
10851142
if not flags.with_cs then raise Not_found
1086-
else conv_record flags env (check_conv_record env i appr2 appr1)
1087-
with Not_found -> let sigma = i in
1088-
match apply_hooks env sigma (Stack.zip sigma appr1) (Stack.zip sigma appr2) with
1089-
| Some sigma -> Success sigma
1090-
| None -> UnifFailure (i,NoCanonicalStructure))
1143+
else conv_record flags env (
1144+
try check_conv_record env i appr2 appr1 with
1145+
| Not_found -> begin let () = debug_unification (fun () -> Pp.(v 0 (str "ask cs hook"))) in match apply_hooks env i appr2 appr1 with
1146+
| Some r -> r
1147+
| None -> raise Not_found
1148+
end
1149+
| e -> let () = Feedback.msg_info Pp.(str "cs hook crashed") in failwith "argh")
1150+
with Not_found -> UnifFailure (i,NoCanonicalStructure))
10911151
and f4 i =
10921152
evar_eqappr_x flags env i pbty appr1
10931153
(whd_betaiota_deltazeta_for_iota_state
10941154
flags.open_ts env i (v2,sk2))
10951155
in
10961156
ise_try evd [f3; f4]
1157+
end
10971158

10981159
(* Eta-expansion *)
10991160
| Rigid, _ when isLambda evd term1 && (* if ever ill-typed: *) List.is_empty sk1 ->
@@ -1277,8 +1338,7 @@ and conv_record flags env (evd,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2),c
12771338
(fun i -> evar_conv_x flags env i CONV c1 app);
12781339
(fun i -> exact_ise_stack2 env i (evar_conv_x flags) sk1 sk2);
12791340
test;
1280-
(fun i -> evar_conv_x flags env i CONV h2
1281-
(fst (decompose_app i (substl ks h))))]
1341+
(fun i -> evar_conv_x flags env i CONV h2 (fst (decompose_app i (substl ks h))))]
12821342
else UnifFailure(evd,(*dummy*)NotSameHead)
12831343

12841344
and eta_constructor flags env evd ((ind, i), u) sk1 (term2,sk2) =

apps/cs/tests/test_cs.v

Lines changed: 11 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -3,36 +3,27 @@ From Coq Require Import Bool.
33

44
Elpi Override CS All.
55

6-
Structure S : Type :=
7-
{ sort :> Type }.
6+
Structure S (T : Type) : Type :=
7+
{ sort :> T -> T }.
88

9-
Elpi Accumulate cs.db lp:{{
9+
Elpi Accumulate canonical_solution lp:{{
1010

11-
cs _ {{ sort lp:Sol }} {{ nat }} :-
12-
Sol = {{ Build_S nat }}.
11+
cs _ {{ sort lp:T }} {{ @id lp:T }} {{ Build_S lp:T (@id lp:T) }}.
1312

1413
}}.
1514
Elpi Typecheck canonical_solution.
1615

1716
Check 1.
18-
Check eq_refl _ : (sort _) = nat.
19-
Definition nat1 := nat.
17+
Check eq_refl _ : (sort nat _) = @id nat.
18+
Check 11.
19+
Check eq_refl _ : (sort nat _) 1 = @id nat 1.
20+
Definition id1 := id.
2021
Check 2.
21-
Check eq_refl _ : (sort _) = nat1.
22+
Check eq_refl _ : (sort nat _) = @id1 nat.
2223
Definition sort1 := sort.
2324
Check 3.
24-
Check eq_refl _ : (sort1 _) = nat.
25+
Check eq_refl _ : (sort1 nat _) = @id nat.
2526
Check 4.
26-
Check eq_refl _ : (sort1 _) = nat1.
27-
28-
29-
Elpi Accumulate cs.db lp:{{
30-
31-
cs _ {{ sort lp:Sol }} {{ bool }} :-
32-
% typing is wired in, do we want this?
33-
std.spy(Sol = {{ Prop }}).
34-
35-
}}.
36-
Elpi Typecheck canonical_solution.
27+
Check eq_refl _ : (sort1 nat _) = @id1 nat.
3728

3829

apps/cs/theories/cs.v

Lines changed: 10 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -3,25 +3,26 @@ From elpi Require Import elpi.
33

44
Elpi Db cs.db lp:{{
55

6-
% predicate [canonical-solution Ctx Lhs Rhs] used to unify Lhs with Rhs, with
6+
% predicate [cs Ctx Proj Rhs Sol] used to find Sol such that Proj Sol = Rhs, where
77
% - [Ctx] is the context
8-
% - [Lhs] and [Rhs] are the terms to unify
8+
% - [Proj] is the projector of some structure, applied to the structure's parameters if any
9+
% - [Rhs] the term to find a structure on.
910
:index (0 6 6)
10-
pred cs i:goal-ctx, o:term, o:term.
11+
pred cs i:goal-ctx, i:term, i:term, o:term.
1112

1213
}}.
1314

1415

1516

1617
Elpi Tactic canonical_solution.
17-
Elpi Accumulate lp:{{
18+
Elpi Accumulate Db cs.db.
19+
Elpi Accumulate canonical_solution lp:{{
1820

19-
solve (goal Ctx _ {{ @eq lp:T lp:Lhs lp:Rhs }} _P []) _ :-
20-
cs Ctx Lhs Rhs,
21-
% std.assert! (P = {{ eq_refl lp:Lhs }}) "cs: wrong solution".
21+
solve (goal Ctx _ _Ty Sol [trm Proj, trm Rhs]) _ :-
22+
cs Ctx Proj Rhs Sol,
23+
% std.assert! (P = {{ eq_refl lp:Lhs }}) "cs: wrong solution".
2224
true.
2325

2426
}}.
25-
Elpi Accumulate Db cs.db.
26-
Elpi Typecheck.
27+
Elpi Typecheck canonical_solution.
2728
Elpi CSFallbackTactic canonical_solution.

0 commit comments

Comments
 (0)