aboutsummaryrefslogtreecommitdiffstats
path: root/driver/Complements.v
blob: 3f32cc6314a4542775dfaed3d82ee6091153cd18 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
(* *********************************************************************)
(*                                                                     *)
(*              The Compcert verified compiler                         *)
(*                                                                     *)
(*          Xavier Leroy, INRIA Paris-Rocquencourt                     *)
(*                                                                     *)
(*  Copyright Institut National de Recherche en Informatique et en     *)
(*  Automatique.  All rights reserved.  This file is distributed       *)
(*  under the terms of the INRIA Non-Commercial License Agreement.     *)
(*                                                                     *)
(* *********************************************************************)

(** Corollaries of the main semantic preservation theorem. *)

Require Import Classical.
Require Import Coqlib.
Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Determinism.
Require Import Csyntax.
Require Import Csem.
Require Import Asm.
Require Import Compiler.
Require Import Errors.

(** * Determinism of Asm semantics *)

(** In this section, we show that the semantics for the Asm language
  are deterministic, provided that the program is executed against a
  deterministic world, as formalized in module [Determinism]. *)

Remark extcall_arguments_deterministic:
  forall rs m sg args args',
  extcall_arguments rs m sg args ->
  extcall_arguments rs m sg args' -> args = args'.
Proof.
  assert (
    forall rs m ll args,
    extcall_args rs m ll args ->
    forall args', extcall_args rs m ll args' -> args = args').
  induction 1; intros.
  inv H. auto.
  inv H1. decEq; eauto.
  inv H; inv H4; congruence.
  unfold extcall_arguments; intros; eauto.
Qed.

Lemma step_internal_deterministic:
  forall ge s t1 s1 t2 s2,
  Asm.step ge s t1 s1 -> Asm.step ge s t2 s2 -> matching_traces t1 t2 ->
  s1 = s2 /\ t1 = t2.
Proof.
  intros. inv H; inv H0.
  assert (c0 = c) by congruence.
  assert (i0 = i) by congruence.
  assert (rs'0 = rs') by congruence.
  assert (m'0 = m') by congruence.
  subst. auto.
  replace i with (Pbuiltin ef args res) in H5 by congruence. simpl in H5. inv H5.
  congruence.
  replace i with (Pbuiltin ef args res) in H12 by congruence. simpl in H12. inv H12.
  rewrite H2 in H7; inv H7. rewrite H3 in H8; inv H8. rewrite H4 in H9; inv H9. 
  exploit external_call_determ. eexact H5. eexact H12. auto. 
  intros [A [B C]]. subst. auto. 
  congruence.
  congruence.
  congruence.
  assert (ef0 = ef) by congruence. subst ef0.
  assert (args0 = args). eapply extcall_arguments_deterministic; eauto. subst args0.
  exploit external_call_determ. 
  eexact H4. eexact H9. auto. 
  intros [A [B C]]. subst. 
  intuition congruence.
Qed.

Lemma initial_state_deterministic:
  forall p s1 s2,
  initial_state p s1 -> initial_state p s2 -> s1 = s2.
Proof.
  intros. inv H; inv H0. f_equal. congruence.
Qed.

Lemma final_state_not_step:
  forall ge st r, final_state st r -> nostep step ge st.
Proof.
  unfold nostep. intros. red; intros. inv H. inv H0.
  unfold Vzero in H1. congruence.
  unfold Vzero in H1. congruence.
  unfold Vzero in H1. congruence.
Qed.

Lemma final_state_deterministic:
  forall st r1 r2, final_state st r1 -> final_state st r2 -> r1 = r2.
Proof.
  intros. inv H; inv H0. congruence.
Qed.

Theorem asm_exec_program_deterministic:
  forall p w beh1 beh2,
  Asm.exec_program p beh1 -> Asm.exec_program p beh2 ->
  possible_behavior w beh1 -> possible_behavior w beh2 ->
  beh1 = beh2.
Proof.
  intros. 
  eapply (program_behaves_deterministic _ _ step (initial_state p) final_state); eauto.
  exact step_internal_deterministic.
  exact (initial_state_deterministic p).
  exact final_state_deterministic.
  exact final_state_not_step.
Qed.

(** * Additional semantic preservation property *)

(** Combining the semantic preservation theorem from module [Compiler]
  with the determinism of Asm executions, we easily obtain
  additional, stronger semantic preservation properties.
  The first property states that, when compiling a Clight
  program that has well-defined semantics, all possible executions
  of the resulting Asm code correspond to an execution of
  the source Clight program. *)

Theorem transf_c_program_is_refinement:
  forall p tp w,
  transf_c_program p = OK tp ->
  (exists b, Csem.exec_program p b /\ possible_behavior w b /\ not_wrong b) ->
  (forall b, Asm.exec_program tp b -> possible_behavior w b -> Csem.exec_program p b).
Proof.
  intros. destruct H0 as [b0 [A [B C]]]. 
  assert (Asm.exec_program tp b0).
    eapply transf_c_program_correct; eauto.
  assert (b = b0). eapply asm_exec_program_deterministic; eauto.
  subst b0. auto.
Qed.

Section SPECS_PRESERVED.

(** The second additional results shows that if one execution
  of the source Clight program satisfies a given specification
  (a predicate on the observable behavior of the program),
  then all executions of the produced Asm program satisfy
  this specification as well.  *) 

Variable spec: program_behavior -> Prop.

Hypothesis spec_not_wrong: forall b, spec b -> not_wrong b.

Theorem transf_c_program_preserves_spec:
  forall p tp w,
  transf_c_program p = OK tp ->
  (exists b, Csem.exec_program p b /\ possible_behavior w b /\ spec b) ->
  (forall b, Asm.exec_program tp b -> possible_behavior w b -> spec b).
Proof.
  intros. destruct H0 as [b1 [A [B C]]].
  assert (Asm.exec_program tp b1).
    eapply transf_c_program_correct; eauto.
  assert (b1 = b). eapply asm_exec_program_deterministic; eauto. 
  subst b1. auto.
Qed.

End SPECS_PRESERVED.