-
Notifications
You must be signed in to change notification settings - Fork 1
/
Copy pathTyping.agda
148 lines (106 loc) · 5.49 KB
/
Typing.agda
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
{- Typing rules of CC2 -}
open import Common.Types
module CC2.Typing where
open import Data.Nat
open import Data.Product renaming (_,_ to ⟨_,_⟩)
open import Data.Maybe
open import Data.List
open import Relation.Binary.PropositionalEquality using (_≡_; refl)
open import Syntax
open import Common.Utils
open import Common.Types
open import Common.Coercions
open import CoercionExpr.SecurityLevel renaming (∥_∥ to ∥_∥ₗ)
open import Memory.HeapContext
open import CC2.Syntax
infix 4 _;_;_;_⊢_⇐_
data _;_;_;_⊢_⇐_ : Context → HeapContext → Label → StaticLabel → Term → Type → Set where
⊢const : ∀ {Γ Σ gc ℓv ι} {k : rep ι} {ℓ}
-------------------------------------------- Const
→ Γ ; Σ ; gc ; ℓv ⊢ $ k ⇐ ` ι of l ℓ
⊢addr : ∀ {Γ Σ gc ℓv n T ℓ ℓ̂}
→ lookup-Σ Σ (a⟦ ℓ̂ ⟧ n) ≡ just T
------------------------------------------------------------------- Addr
→ Γ ; Σ ; gc ; ℓv ⊢ addr n ⇐ Ref (T of l ℓ̂) of l ℓ
⊢var : ∀ {Γ Σ gc ℓv A x}
→ Γ ∋ x ⦂ A
----------------------------- Var
→ Γ ; Σ ; gc ; ℓv ⊢ ` x ⇐ A
⊢lam : ∀ {Γ Σ gc ℓv g A B N ℓ}
→ (∀ {ℓv} → A ∷ Γ ; Σ ; g ; ℓv ⊢ N ⇐ B)
------------------------------------------------------------------- Lambda
→ Γ ; Σ ; gc ; ℓv ⊢ ƛ N ⇐ ⟦ g ⟧ A ⇒ B of l ℓ
⊢app : ∀ {Γ Σ ℓc ℓv A B C L M ℓ}
→ Γ ; Σ ; l ℓc ; ℓv ⊢ L ⇐ ⟦ l (ℓc ⋎ ℓ) ⟧ A ⇒ B of l ℓ
→ Γ ; Σ ; l ℓc ; ℓv ⊢ M ⇐ A
→ C ≡ stamp B (l ℓ)
------------------------------------------------------ AppStatic
→ Γ ; Σ ; l ℓc ; ℓv ⊢ app L M A B ℓ ⇐ C
⊢app⋆ : ∀ {Γ Σ gc ℓv A T L M}
→ Γ ; Σ ; gc ; ℓv ⊢ L ⇐ ⟦ ⋆ ⟧ A ⇒ (T of ⋆) of ⋆
→ Γ ; Σ ; gc ; ℓv ⊢ M ⇐ A
------------------------------------------------------ App!
→ Γ ; Σ ; gc ; ℓv ⊢ app⋆ L M A T ⇐ T of ⋆
⊢if : ∀ {Γ Σ ℓc ℓv A B L M N ℓ}
→ Γ ; Σ ; l ℓc ; ℓv ⊢ L ⇐ ` Bool of l ℓ
→ (∀ {ℓv} → Γ ; Σ ; l (ℓc ⋎ ℓ) ; ℓv ⊢ M ⇐ A)
→ (∀ {ℓv} → Γ ; Σ ; l (ℓc ⋎ ℓ) ; ℓv ⊢ N ⇐ A)
→ B ≡ stamp A (l ℓ)
--------------------------------------------------------- If
→ Γ ; Σ ; l ℓc ; ℓv ⊢ if L A ℓ M N ⇐ B
⊢if⋆ : ∀ {Γ Σ gc ℓv T L M N}
→ Γ ; Σ ; gc ; ℓv ⊢ L ⇐ ` Bool of ⋆
→ (∀ {ℓv} → Γ ; Σ ; ⋆ ; ℓv ⊢ M ⇐ T of ⋆)
→ (∀ {ℓv} → Γ ; Σ ; ⋆ ; ℓv ⊢ N ⇐ T of ⋆)
--------------------------------------------------------- If!
→ Γ ; Σ ; gc ; ℓv ⊢ if⋆ L T M N ⇐ T of ⋆
⊢let : ∀ {Γ Σ gc ℓv M N A B}
→ Γ ; Σ ; gc ; ℓv ⊢ M ⇐ A
→ (∀ {ℓv} → A ∷ Γ ; Σ ; gc ; ℓv ⊢ N ⇐ B)
--------------------------------------------- Let
→ Γ ; Σ ; gc ; ℓv ⊢ `let M A N ⇐ B
⊢ref : ∀ {Γ Σ ℓc ℓv M T ℓ}
→ Γ ; Σ ; l ℓc ; ℓv ⊢ M ⇐ T of l ℓ
→ ℓc ≼ ℓ
---------------------------------------------------------------- RefStatic
→ Γ ; Σ ; l ℓc ; ℓv ⊢ ref⟦ ℓ ⟧ M ⇐ Ref (T of l ℓ) of l low
⊢ref? : ∀ {Γ Σ ℓv M T ℓ p}
→ Γ ; Σ ; ⋆ ; ℓv ⊢ M ⇐ T of l ℓ
---------------------------------------------------------------- Ref?
→ Γ ; Σ ; ⋆ ; ℓv ⊢ ref?⟦ ℓ ⟧ M p ⇐ Ref (T of l ℓ) of l low
⊢deref : ∀ {Γ Σ gc ℓv M A B ℓ}
→ Γ ; Σ ; gc ; ℓv ⊢ M ⇐ Ref A of l ℓ
→ B ≡ stamp A (l ℓ)
------------------------------------- Deref
→ Γ ; Σ ; gc ; ℓv ⊢ ! M A ℓ ⇐ B
⊢deref⋆ : ∀ {Γ Σ gc ℓv M T}
→ Γ ; Σ ; gc ; ℓv ⊢ M ⇐ Ref (T of ⋆) of ⋆
------------------------------------- Deref!
→ Γ ; Σ ; gc ; ℓv ⊢ !⋆ M T ⇐ T of ⋆
⊢assign : ∀ {Γ Σ ℓc ℓv L M T ℓ ℓ̂}
→ Γ ; Σ ; l ℓc ; ℓv ⊢ L ⇐ Ref (T of l ℓ̂) of l ℓ
→ Γ ; Σ ; l ℓc ; ℓv ⊢ M ⇐ T of l ℓ̂
→ ℓc ≼ ℓ̂ → ℓ ≼ ℓ̂
------------------------------------------------------------- AssignStatic
→ Γ ; Σ ; l ℓc ; ℓv ⊢ assign L M T ℓ̂ ℓ ⇐ ` Unit of l low
⊢assign? : ∀ {Γ Σ gc ℓv L M T ĝ p}
→ Γ ; Σ ; gc ; ℓv ⊢ L ⇐ Ref (T of ĝ) of ⋆
→ Γ ; Σ ; gc ; ℓv ⊢ M ⇐ T of ĝ
------------------------------------------------------------- Assign?
→ Γ ; Σ ; gc ; ℓv ⊢ assign? L M T ĝ p ⇐ ` Unit of l low
{- Note: L has ⋆ so that we can get its security using ∥_∥ at runtime -}
⊢prot : ∀ {Γ Σ gc gc′ ℓv A B M ℓ} {PC} {vc : LVal PC}
→ let ℓv′ = ∥ PC ∥ vc in
Γ ; Σ ; gc′ ; ℓv′ ⊢ M ⇐ A
→ ⊢ PC ⇐ gc′
→ ℓv ⋎ ℓ ≼ ℓv′
→ B ≡ stamp A (l ℓ)
---------------------------------------------------- Prot
→ Γ ; Σ ; gc ; ℓv ⊢ prot PC vc ℓ M A ⇐ B
⊢cast : ∀ {Γ Σ gc ℓv A B M} {c : Cast A ⇒ B}
→ Γ ; Σ ; gc ; ℓv ⊢ M ⇐ A
----------------------------------------- Cast
→ Γ ; Σ ; gc ; ℓv ⊢ M ⟨ c ⟩ ⇐ B
⊢blame : ∀ {Γ Σ gc ℓv A p}
------------------------------------ Blame
→ Γ ; Σ ; gc ; ℓv ⊢ blame p ⇐ A