Skip to content

Commit a94edc5

Browse files
committed
Add Ctypes.link_match_program_gen
A more general version of the link_match_program linking theorem. It supports match_fundef relations parameterized by the source compilation unit.
1 parent e9f40aa commit a94edc5

File tree

1 file changed

+52
-0
lines changed

1 file changed

+52
-0
lines changed

cfrontend/Ctypes.v

Lines changed: 52 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1516,6 +1516,57 @@ Global Opaque Linker_program.
15161516

15171517
(** ** Commutation between linking and program transformations *)
15181518

1519+
Section LINK_MATCH_PROGRAM_GEN.
1520+
1521+
Context {F G: Type}.
1522+
Variable match_fundef: program F -> fundef F -> fundef G -> Prop.
1523+
1524+
Hypothesis link_match_fundef:
1525+
forall ctx1 ctx2 f1 tf1 f2 tf2 f,
1526+
link f1 f2 = Some f ->
1527+
match_fundef ctx1 f1 tf1 -> match_fundef ctx2 f2 tf2 ->
1528+
exists tf, link tf1 tf2 = Some tf /\ (match_fundef ctx1 f tf \/ match_fundef ctx2 f tf).
1529+
1530+
Let match_program (p: program F) (tp: program G) : Prop :=
1531+
Linking.match_program_gen match_fundef eq p p tp
1532+
/\ prog_types tp = prog_types p.
1533+
1534+
Theorem link_match_program_gen:
1535+
forall p1 p2 tp1 tp2 p,
1536+
link p1 p2 = Some p -> match_program p1 tp1 -> match_program p2 tp2 ->
1537+
exists tp, link tp1 tp2 = Some tp /\ match_program p tp.
1538+
Proof.
1539+
intros until p; intros L [M1 T1] [M2 T2].
1540+
exploit link_linkorder; eauto. intros [LO1 LO2].
1541+
Local Transparent Linker_program.
1542+
simpl in L; unfold link_program in L.
1543+
destruct (link (program_of_program p1) (program_of_program p2)) as [pp|] eqn:LP; try discriminate.
1544+
assert (A: exists tpp,
1545+
link (program_of_program tp1) (program_of_program tp2) = Some tpp
1546+
/\ Linking.match_program_gen match_fundef eq p pp tpp).
1547+
{ eapply Linking.link_match_program; eauto.
1548+
- intros.
1549+
Local Transparent Linker_types.
1550+
simpl in *. destruct (type_eq v1 v2); inv H. exists v; rewrite dec_eq_true; auto.
1551+
}
1552+
destruct A as (tpp & TLP & MP).
1553+
simpl; unfold link_program. rewrite TLP.
1554+
destruct (lift_option (link (prog_types p1) (prog_types p2))) as [[typs EQ]|EQ]; try discriminate.
1555+
destruct (link_build_composite_env (prog_types p1) (prog_types p2) typs
1556+
(prog_comp_env p1) (prog_comp_env p2) (prog_comp_env_eq p1)
1557+
(prog_comp_env_eq p2) EQ) as (env & P & Q).
1558+
rewrite <- T1, <- T2 in EQ.
1559+
destruct (lift_option (link (prog_types tp1) (prog_types tp2))) as [[ttyps EQ']|EQ']; try congruence.
1560+
assert (ttyps = typs) by congruence. subst ttyps.
1561+
destruct (link_build_composite_env (prog_types tp1) (prog_types tp2) typs
1562+
(prog_comp_env tp1) (prog_comp_env tp2) (prog_comp_env_eq tp1)
1563+
(prog_comp_env_eq tp2) EQ') as (tenv & R & S).
1564+
assert (tenv = env) by congruence. subst tenv.
1565+
econstructor; split; eauto. inv L. split; auto.
1566+
Qed.
1567+
1568+
End LINK_MATCH_PROGRAM_GEN.
1569+
15191570
Section LINK_MATCH_PROGRAM.
15201571

15211572
Context {F G: Type}.
@@ -1571,3 +1622,4 @@ Local Transparent Linker_program.
15711622
Qed.
15721623

15731624
End LINK_MATCH_PROGRAM.
1625+

0 commit comments

Comments
 (0)