1 | (* Memory model used in the dynamic semantics of the back-end intermediate |
---|
2 | languages. Inspired by common/Mem.ma, adapted from Compcert *) |
---|
3 | |
---|
4 | (* * This file develops the memory model that is used in the dynamic |
---|
5 | semantics of all the languages used in the compiler. |
---|
6 | It defines a type [mem] of memory states, the following 4 basic |
---|
7 | operations over memory states, and their properties: |
---|
8 | - [load]: read a memory chunk at a given address; |
---|
9 | - [store]: store a memory chunk at a given address; |
---|
10 | - [alloc]: allocate a fresh memory block; |
---|
11 | - [free]: invalidate a memory block. |
---|
12 | *) |
---|
13 | |
---|
14 | include "common/ByteValues.ma". |
---|
15 | include "common/GenMem.ma". |
---|
16 | |
---|
17 | definition bemem ≝ mem. |
---|
18 | |
---|
19 | definition is_addressable : region → bool ≝ λr.match r with |
---|
20 | [ XData ⇒ true | Code ⇒ true | _ ⇒ false ]. |
---|
21 | |
---|
22 | (* CSC: only pointers to XRAM or code memory ATM *) |
---|
23 | definition address ≝ beval × beval. |
---|
24 | |
---|
25 | definition pointer_of_address: address → res pointer ≝ |
---|
26 | λp. let 〈v1,v2〉 ≝ p in pointer_of_bevals [v1;v2]. |
---|
27 | |
---|
28 | definition is_address : (beval × beval) → Prop ≝ λa. |
---|
29 | ∃ptr.pointer_of_address a = OK … ptr. |
---|
30 | |
---|
31 | lemma bevals_of_pointer_is_address : |
---|
32 | ∀p.is_address (list_to_pair … (bevals_of_pointer p) (refl …)). |
---|
33 | * #bl #o |
---|
34 | whd in ⊢ (?%); whd % |
---|
35 | [2: whd in ⊢ (??%?); |
---|
36 | >eq_block_b_b |
---|
37 | change with (eq_bv ???) in match (eq_bv_suffix ?????); |
---|
38 | >eq_bv_refl normalize %] |
---|
39 | qed. |
---|
40 | |
---|
41 | (* For full 8051 memory spaces |
---|
42 | definition pointer_compat' ≝ λb,r. |
---|
43 | match b with |
---|
44 | [ mk_block r' z ⇒ |
---|
45 | if eq_region r' r then True |
---|
46 | else |
---|
47 | match r with |
---|
48 | [ Any ⇒ True |
---|
49 | | XData ⇒ match r' with |
---|
50 | [ PData ⇒ True |
---|
51 | | _ ⇒ False |
---|
52 | ] |
---|
53 | | _ ⇒ False |
---|
54 | ] |
---|
55 | ]. |
---|
56 | |
---|
57 | lemma pointer_compat_to_ind : ∀b,r.pointer_compat' b r → pointer_compat b r. |
---|
58 | ** #z ** // |
---|
59 | qed. |
---|
60 | |
---|
61 | lemma pointer_compat_from_ind : ∀b,r.pointer_compat b r → pointer_compat' b r. |
---|
62 | #b #r #H inversion H |
---|
63 | [ #s #id #EQ1 #EQ2 #EQ3 whd >reflexive_eq_region % |
---|
64 | | #id #EQ1 #EQ2 #EQ3 % |
---|
65 | | #r #id #EQ1 #EQ2 #EQ3 whd @eq_region_elim #EQ4 destruct(EQ4) % |
---|
66 | ] |
---|
67 | qed. |
---|
68 | *) |
---|
69 | (*lemma pointer_of_address_is_safe : ∀a : safe_address.pointer_of_address a = OK … (pointer_of_address_safe a). |
---|
70 | ** #a0 #a1 ***** #z #o |
---|
71 | #Hr |
---|
72 | (* |
---|
73 | ** #a0 #a1 ***** #r #z #Hr #o * lapply (pointer_compat_from_ind ?? Hr) |
---|
74 | cases r in Hr ⊢ %; #Hr * |
---|
75 | *) |
---|
76 | ** #part0 #H0 ** #part1 #H1 *** #EQa0 #EQpart0 #EQa1 #EQpart1 |
---|
77 | destruct normalize |
---|
78 | @eqZb_elim [2,4,6: * #ABS elim (ABS (refl …))] |
---|
79 | change with (eq_bv ???) in match (eq_v ?????); |
---|
80 | @eq_bv_elim [2,4,6: * #ABS elim (ABS (refl …))] |
---|
81 | #_ #_ normalize % |
---|
82 | qed.*) |
---|
83 | |
---|
84 | (*definition safe_address_of_pointer: pointer → res safe_address ≝ λp. |
---|
85 | do a as EQ ← address_of_pointer p ; return «a ,address_of_pointer_OK_to_safe ?? EQ». |
---|
86 | |
---|
87 | lemma address_of_pointer_is_safe : |
---|
88 | ∀p.address_of_pointer p = ! a ← safe_address_of_pointer p ; return (a : address). |
---|
89 | #p normalize % qed. |
---|
90 | (* |
---|
91 | ****#z #H |
---|
92 | lapply (pointer_compat_from_ind ?? H) * #o |
---|
93 | normalize % |
---|
94 | qed.*) |
---|
95 | |
---|
96 | definition code_pointer_of_address: address → res (Σp:pointer. ptype p = Code) ≝ |
---|
97 | code_pointer_of_beval_pair. |
---|
98 | |
---|
99 | definition address_of_code_pointer: (Σp:pointer. ptype p = Code) → address ≝ beval_pair_of_code_pointer. |
---|
100 | |
---|
101 | definition safe_address_of_code_pointer: (Σp:pointer. ptype p = Code) → safe_address ≝ address_of_code_pointer. |
---|
102 | cases H -H * #b #o normalize cases b * #z #EQ destruct |
---|
103 | %{«mk_pointer (mk_block Code z) o,I»} |
---|
104 | % [2: % [2: % [ % [ % ]]]] % |
---|
105 | qed. |
---|
106 | (* |
---|
107 | cases H -H * #r #b #H #o #EQ destruct(EQ) normalize lapply H |
---|
108 | lapply (pointer_compat_from_ind … H) -H cases b * #z * #H |
---|
109 | %{«mk_pointer ? (mk_block Code z) H o,I»} |
---|
110 | % [2: % [2: % [ % [ % ]] % |]|] |
---|
111 | qed.*) |
---|
112 | *) |
---|
113 | |
---|
114 | (* Paolo: why only code pointers and not XRAM too? *) |
---|
115 | (*definition addr_add: address → nat → res address ≝ |
---|
116 | λaddr,n. |
---|
117 | do ptr ← code_pointer_of_address addr ; |
---|
118 | OK … (address_of_code_pointer (shift_pointer 16 ptr (bitvector_of_nat … n))). |
---|
119 | normalize cases ptr #p #E @E |
---|
120 | qed. |
---|
121 | |
---|
122 | definition safe_addr_add: safe_address → nat → res safe_address ≝ |
---|
123 | λaddr,n. |
---|
124 | do ptr ← code_pointer_of_address addr ; |
---|
125 | OK … (safe_address_of_code_pointer (shift_pointer 16 ptr (bitvector_of_nat … n))). |
---|
126 | normalize cases ptr #p #E @E |
---|
127 | qed. |
---|
128 | |
---|
129 | definition addr_sub: address → nat → res address ≝ |
---|
130 | λaddr,n. |
---|
131 | do ptr ← code_pointer_of_address addr ; |
---|
132 | OK … (address_of_code_pointer (neg_shift_pointer 16 ptr (bitvector_of_nat … n))). |
---|
133 | normalize cases ptr #p #E @E |
---|
134 | qed. |
---|
135 | |
---|
136 | definition safe_addr_sub: safe_address → nat → res safe_address ≝ |
---|
137 | λaddr,n. |
---|
138 | do ptr ← code_pointer_of_address addr ; |
---|
139 | OK … (safe_address_of_code_pointer (neg_shift_pointer 16 ptr (bitvector_of_nat … n))). |
---|
140 | normalize cases ptr #p #E @E |
---|
141 | qed. |
---|
142 | *) |
---|