@@ -20,6 +20,13 @@ Arguments Φ _ {_}.
2020
2121Global Opaque φ.
2222
23+ Module WrappedLink.
24+ Record t : Type := {
25+ A : Set ;
26+ H : Link A;
27+ }.
28+ End WrappedLink.
29+
2330Smpl Create of_ty.
2431
2532Module OfTy.
@@ -877,6 +884,17 @@ Module Run.
877884 {{ k (SuccessOrPanic.to_value value_inter) 🔽 R, Output }}
878885 ) ->
879886 {{ LowM.CallClosure ty closure args k 🔽 R, Output }}
887+ | CallClosureWrapped
888+ (ty : Ty.t) (f : list Value.t -> M) (args : list Value.t) (k : Value.t + Exception.t -> M)
889+ (wrapped_link : WrappedLink.t) :
890+ let Output' : Set := wrapped_link.(WrappedLink.A) in
891+ let _ := wrapped_link.(WrappedLink.H) in
892+ let closure := Value.Closure (existS (_, _) f) in
893+ {{ f args 🔽 Output', Output' }} ->
894+ (forall (value_inter : SuccessOrPanic.t Output'),
895+ {{ k (SuccessOrPanic.to_value value_inter) 🔽 R, Output }}
896+ ) ->
897+ {{ LowM.CallClosure ty closure args k 🔽 R, Output }}
880898 | CallLogicalOp
881899 (op : LogicalOp.t) (lhs : bool) (rhs : M) (k : Value.t + Exception.t -> M) :
882900 {{ rhs 🔽 R, bool }} ->
@@ -893,6 +911,16 @@ Module Run.
893911 {{ k (Output.to_value value_inter) 🔽 R, Output }}
894912 ) ->
895913 {{ LowM.Let ty e k 🔽 R, Output }}
914+ | LetWrapped
915+ (ty : Ty.t) (e : M) (k : Value.t + Exception.t -> M)
916+ (wrapped_link : WrappedLink.t) :
917+ let Output' : Set := wrapped_link.(WrappedLink.A) in
918+ let _ := wrapped_link.(WrappedLink.H) in
919+ {{ e 🔽 R, Output' }} ->
920+ (forall (value_inter : Output.t R Output'),
921+ {{ k (Output.to_value value_inter) 🔽 R, Output }}
922+ ) ->
923+ {{ LowM.Let ty e k 🔽 R, Output }}
896924 | Loop
897925 (ty : Ty.t) (body : M) (k : Value.t + Exception.t -> M)
898926 (of_ty : OfTy.t ty) :
@@ -1068,6 +1096,15 @@ Proof.
10681096 | H : forall _ : SuccessOrPanic.t Output', _ |- _ => apply (H output')
10691097 end.
10701098 }
1099+ { (* CallClosureWrapped *)
1100+ eapply LowM.Call. {
1101+ exact (evaluate _ _ _ _ _ run).
1102+ }
1103+ intros output'; eapply evaluate.
1104+ match goal with
1105+ | H : forall _ : SuccessOrPanic.t Output', _ |- _ => apply (H output')
1106+ end.
1107+ }
10711108 { (* CallLogicalOp *)
10721109 match goal with
10731110 | H : forall _ : Output.t _ bool, _ |- _ => rename H into H_k
@@ -1111,6 +1148,15 @@ Proof.
11111148 | H : forall _ : Output.t _ Output', _ |- _ => apply (H output')
11121149 end.
11131150 }
1151+ { (* LetWrapped *)
1152+ eapply LowM.Let. {
1153+ exact (evaluate _ _ _ _ _ run).
1154+ }
1155+ intros output'; eapply evaluate.
1156+ match goal with
1157+ | H : forall _ : Output.t _ Output', _ |- _ => apply (H output')
1158+ end.
1159+ }
11141160 { (* Loop *)
11151161 eapply LowM.Loop. {
11161162 exact (evaluate _ _ _ _ _ run).
@@ -1244,8 +1290,8 @@ Ltac run_symbolic_closure :=
12441290 ].
12451291
12461292Ltac run_symbolic_closure_auto :=
1247- unshelve eapply Run.CallClosure ; [
1248- now repeat smpl of_ty |
1293+ unshelve eapply Run.CallClosureWrapped ; [
1294+ econstructor; shelve |
12491295 try prepare_call;
12501296 (
12511297 (
@@ -1354,16 +1400,15 @@ Ltac rewrite_cast_integer :=
13541400 |]
13551401 end .
13561402
1357- Ltac run_symbolic_let :=
1358- unshelve eapply Run.Let ; [repeat smpl of_ty | | cbn; intros []].
1359-
13601403Ltac run_symbolic_loop :=
13611404 unshelve eapply Run.Loop; [
13621405 smpl of_ty |
13631406 |
13641407 cbn; intros []
13651408 ].
13661409
1410+ Smpl Create run_symbolic.
1411+
13671412Ltac run_symbolic_one_step_immediate :=
13681413 match goal with
13691414 | |- {{ _ 🔽 _, _ }} =>
@@ -1380,31 +1425,41 @@ Ltac run_symbolic_one_step_immediate :=
13801425 run_symbolic_get_trait_method ||
13811426 run_symbolic_closure_auto ||
13821427 run_symbolic_logical_op ||
1383- run_symbolic_let ||
13841428 run_sub_pointer ||
13851429 run_symbolic_loop ||
1386- fold @LowM.let_
1430+ fold @LowM.let_ ||
1431+ smpl run_symbolic
13871432 end .
13881433
1389- Smpl Create run_symbolic.
1434+ Ltac run_symbolic_inner :=
1435+ (run_symbolic_one_step_immediate; run_symbolic_inner) ||
1436+ (unshelve eapply Run.LetWrapped; [
1437+ econstructor; shelve |
1438+ run_symbolic_inner |
1439+ cbn; intros [];
1440+ run_symbolic_inner
1441+ ]) ||
1442+ match goal with
1443+ | |- context[match Output.Exception.to_exception ?exception with _ => _ end ] =>
1444+ destruct exception; run_symbolic_inner
1445+ end ||
1446+ match goal with
1447+ | |- {{ ?expression 🔽 _, _ }} =>
1448+ match expression with
1449+ | context [match ?expression with _ => _ end ] =>
1450+ destruct expression; run_symbolic_inner
1451+ end
1452+ end ||
1453+ (* We do our best to make progress, but can stop at any point to debug. *)
1454+ idtac.
13901455
1391- (** We should use this tactic instead of the ones above, as this one calls all the others. *)
13921456Ltac run_symbolic :=
1393- progress (repeat (
1394- run_symbolic_one_step_immediate ||
1395- smpl run_symbolic ||
1396- match goal with
1397- | |- context[match Output.Exception.to_exception ?exception with _ => _ end ] =>
1398- destruct exception; run_symbolic
1399- end ||
1400- match goal with
1401- | |- {{ ?expression 🔽 _, _ }} =>
1402- match expression with
1403- | context [match ?expression with _ => _ end ] =>
1404- destruct expression; run_symbolic
1405- end
1406- end
1407- )).
1457+ (unshelve (progress run_symbolic_inner));
1458+ (* There might be remaining types for branches that are never reached. *)
1459+ try (
1460+ exact Empty_set ||
1461+ typeclasses eauto
1462+ ).
14081463
14091464Axiom is_discriminant_tuple_eq :
14101465 forall
0 commit comments