| author | wenzelm | 
| Sat, 25 Jun 2022 10:27:42 +0200 | |
| changeset 75619 | 9639c3867b86 | 
| parent 69597 | ff784d5a5bfb | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 43158 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 68776 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 nipkow parents: 
67406diff
changeset | 3 | subsection "Security Typing of Commands" | 
| 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 nipkow parents: 
67406diff
changeset | 4 | |
| 43158 | 5 | theory Sec_Typing imports Sec_Type_Expr | 
| 6 | begin | |
| 7 | ||
| 68776 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 nipkow parents: 
67406diff
changeset | 8 | subsubsection "Syntax Directed Typing" | 
| 43158 | 9 | |
| 10 | inductive sec_type :: "nat \<Rightarrow> com \<Rightarrow> bool" ("(_/ \<turnstile> _)" [0,0] 50) where
 | |
| 11 | Skip: | |
| 12 | "l \<turnstile> SKIP" | | |
| 13 | Assign: | |
| 50342 | 14 | "\<lbrakk> sec x \<ge> sec a; sec x \<ge> l \<rbrakk> \<Longrightarrow> l \<turnstile> x ::= a" | | 
| 47818 | 15 | Seq: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 16 | "\<lbrakk> l \<turnstile> c\<^sub>1; l \<turnstile> c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile> c\<^sub>1;;c\<^sub>2" | | 
| 43158 | 17 | If: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 18 | "\<lbrakk> max (sec b) l \<turnstile> c\<^sub>1; max (sec b) l \<turnstile> c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2" | | 
| 43158 | 19 | While: | 
| 50342 | 20 | "max (sec b) l \<turnstile> c \<Longrightarrow> l \<turnstile> WHILE b DO c" | 
| 43158 | 21 | |
| 22 | code_pred (expected_modes: i => i => bool) sec_type . | |
| 23 | ||
| 24 | value "0 \<turnstile> IF Less (V ''x1'') (V ''x'') THEN ''x1'' ::= N 0 ELSE SKIP" | |
| 25 | value "1 \<turnstile> IF Less (V ''x1'') (V ''x'') THEN ''x'' ::= N 0 ELSE SKIP" | |
| 26 | value "2 \<turnstile> IF Less (V ''x1'') (V ''x'') THEN ''x1'' ::= N 0 ELSE SKIP" | |
| 27 | ||
| 28 | inductive_cases [elim!]: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 29 | "l \<turnstile> x ::= a" "l \<turnstile> c\<^sub>1;;c\<^sub>2" "l \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2" "l \<turnstile> WHILE b DO c" | 
| 43158 | 30 | |
| 31 | ||
| 67406 | 32 | text\<open>An important property: anti-monotonicity.\<close> | 
| 43158 | 33 | |
| 34 | lemma anti_mono: "\<lbrakk> l \<turnstile> c; l' \<le> l \<rbrakk> \<Longrightarrow> l' \<turnstile> c" | |
| 45015 | 35 | apply(induction arbitrary: l' rule: sec_type.induct) | 
| 43158 | 36 | apply (metis sec_type.intros(1)) | 
| 37 | apply (metis le_trans sec_type.intros(2)) | |
| 38 | apply (metis sec_type.intros(3)) | |
| 39 | apply (metis If le_refl sup_mono sup_nat_def) | |
| 40 | apply (metis While le_refl sup_mono sup_nat_def) | |
| 41 | done | |
| 42 | ||
| 43 | lemma confinement: "\<lbrakk> (c,s) \<Rightarrow> t; l \<turnstile> c \<rbrakk> \<Longrightarrow> s = t (< l)" | |
| 45015 | 44 | proof(induction rule: big_step_induct) | 
| 43158 | 45 | case Skip thus ?case by simp | 
| 46 | next | |
| 47 | case Assign thus ?case by auto | |
| 48 | next | |
| 47818 | 49 | case Seq thus ?case by auto | 
| 43158 | 50 | next | 
| 51 | case (IfTrue b s c1) | |
| 50342 | 52 | hence "max (sec b) l \<turnstile> c1" by auto | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 53 | hence "l \<turnstile> c1" by (metis max.cobounded2 anti_mono) | 
| 45015 | 54 | thus ?case using IfTrue.IH by metis | 
| 43158 | 55 | next | 
| 56 | case (IfFalse b s c2) | |
| 50342 | 57 | hence "max (sec b) l \<turnstile> c2" by auto | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 58 | hence "l \<turnstile> c2" by (metis max.cobounded2 anti_mono) | 
| 45015 | 59 | thus ?case using IfFalse.IH by metis | 
| 43158 | 60 | next | 
| 61 | case WhileFalse thus ?case by auto | |
| 62 | next | |
| 63 | case (WhileTrue b s1 c) | |
| 50342 | 64 | hence "max (sec b) l \<turnstile> c" by auto | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 65 | hence "l \<turnstile> c" by (metis max.cobounded2 anti_mono) | 
| 43158 | 66 | thus ?case using WhileTrue by metis | 
| 67 | qed | |
| 68 | ||
| 69 | ||
| 70 | theorem noninterference: | |
| 71 | "\<lbrakk> (c,s) \<Rightarrow> s'; (c,t) \<Rightarrow> t'; 0 \<turnstile> c; s = t (\<le> l) \<rbrakk> | |
| 72 | \<Longrightarrow> s' = t' (\<le> l)" | |
| 45015 | 73 | proof(induction arbitrary: t t' rule: big_step_induct) | 
| 43158 | 74 | case Skip thus ?case by auto | 
| 75 | next | |
| 76 | case (Assign x a s) | |
| 77 | have [simp]: "t' = t(x := aval a t)" using Assign by auto | |
| 67406 | 78 | have "sec x >= sec a" using \<open>0 \<turnstile> x ::= a\<close> by auto | 
| 43158 | 79 | show ?case | 
| 80 | proof auto | |
| 81 | assume "sec x \<le> l" | |
| 67406 | 82 | with \<open>sec x >= sec a\<close> have "sec a \<le> l" by arith | 
| 43158 | 83 | thus "aval a s = aval a t" | 
| 67406 | 84 | by (rule aval_eq_if_eq_le[OF \<open>s = t (\<le> l)\<close>]) | 
| 43158 | 85 | next | 
| 86 | fix y assume "y \<noteq> x" "sec y \<le> l" | |
| 67406 | 87 | thus "s y = t y" using \<open>s = t (\<le> l)\<close> by simp | 
| 43158 | 88 | qed | 
| 89 | next | |
| 47818 | 90 | case Seq thus ?case by blast | 
| 43158 | 91 | next | 
| 92 | case (IfTrue b s c1 s' c2) | |
| 67406 | 93 | have "sec b \<turnstile> c1" "sec b \<turnstile> c2" using \<open>0 \<turnstile> IF b THEN c1 ELSE c2\<close> by auto | 
| 43158 | 94 | show ?case | 
| 95 | proof cases | |
| 50342 | 96 | assume "sec b \<le> l" | 
| 67406 | 97 | hence "s = t (\<le> sec b)" using \<open>s = t (\<le> l)\<close> by auto | 
| 98 | hence "bval b t" using \<open>bval b s\<close> by(simp add: bval_eq_if_eq_le) | |
| 99 | with IfTrue.IH IfTrue.prems(1,3) \<open>sec b \<turnstile> c1\<close> anti_mono | |
| 43158 | 100 | show ?thesis by auto | 
| 101 | next | |
| 50342 | 102 | assume "\<not> sec b \<le> l" | 
| 103 | have 1: "sec b \<turnstile> IF b THEN c1 ELSE c2" | |
| 67406 | 104 | by(rule sec_type.intros)(simp_all add: \<open>sec b \<turnstile> c1\<close> \<open>sec b \<turnstile> c2\<close>) | 
| 105 | from confinement[OF \<open>(c1, s) \<Rightarrow> s'\<close> \<open>sec b \<turnstile> c1\<close>] \<open>\<not> sec b \<le> l\<close> | |
| 43158 | 106 | have "s = s' (\<le> l)" by auto | 
| 107 | moreover | |
| 67406 | 108 | from confinement[OF \<open>(IF b THEN c1 ELSE c2, t) \<Rightarrow> t'\<close> 1] \<open>\<not> sec b \<le> l\<close> | 
| 43158 | 109 | have "t = t' (\<le> l)" by auto | 
| 67406 | 110 | ultimately show "s' = t' (\<le> l)" using \<open>s = t (\<le> l)\<close> by auto | 
| 43158 | 111 | qed | 
| 112 | next | |
| 113 | case (IfFalse b s c2 s' c1) | |
| 67406 | 114 | have "sec b \<turnstile> c1" "sec b \<turnstile> c2" using \<open>0 \<turnstile> IF b THEN c1 ELSE c2\<close> by auto | 
| 43158 | 115 | show ?case | 
| 116 | proof cases | |
| 50342 | 117 | assume "sec b \<le> l" | 
| 67406 | 118 | hence "s = t (\<le> sec b)" using \<open>s = t (\<le> l)\<close> by auto | 
| 119 | hence "\<not> bval b t" using \<open>\<not> bval b s\<close> by(simp add: bval_eq_if_eq_le) | |
| 120 | with IfFalse.IH IfFalse.prems(1,3) \<open>sec b \<turnstile> c2\<close> anti_mono | |
| 43158 | 121 | show ?thesis by auto | 
| 122 | next | |
| 50342 | 123 | assume "\<not> sec b \<le> l" | 
| 124 | have 1: "sec b \<turnstile> IF b THEN c1 ELSE c2" | |
| 67406 | 125 | by(rule sec_type.intros)(simp_all add: \<open>sec b \<turnstile> c1\<close> \<open>sec b \<turnstile> c2\<close>) | 
| 126 | from confinement[OF big_step.IfFalse[OF IfFalse(1,2)] 1] \<open>\<not> sec b \<le> l\<close> | |
| 43158 | 127 | have "s = s' (\<le> l)" by auto | 
| 128 | moreover | |
| 67406 | 129 | from confinement[OF \<open>(IF b THEN c1 ELSE c2, t) \<Rightarrow> t'\<close> 1] \<open>\<not> sec b \<le> l\<close> | 
| 43158 | 130 | have "t = t' (\<le> l)" by auto | 
| 67406 | 131 | ultimately show "s' = t' (\<le> l)" using \<open>s = t (\<le> l)\<close> by auto | 
| 43158 | 132 | qed | 
| 133 | next | |
| 134 | case (WhileFalse b s c) | |
| 50342 | 135 | have "sec b \<turnstile> c" using WhileFalse.prems(2) by auto | 
| 43158 | 136 | show ?case | 
| 137 | proof cases | |
| 50342 | 138 | assume "sec b \<le> l" | 
| 67406 | 139 | hence "s = t (\<le> sec b)" using \<open>s = t (\<le> l)\<close> by auto | 
| 140 | hence "\<not> bval b t" using \<open>\<not> bval b s\<close> by(simp add: bval_eq_if_eq_le) | |
| 43158 | 141 | with WhileFalse.prems(1,3) show ?thesis by auto | 
| 142 | next | |
| 50342 | 143 | assume "\<not> sec b \<le> l" | 
| 144 | have 1: "sec b \<turnstile> WHILE b DO c" | |
| 67406 | 145 | by(rule sec_type.intros)(simp_all add: \<open>sec b \<turnstile> c\<close>) | 
| 146 | from confinement[OF \<open>(WHILE b DO c, t) \<Rightarrow> t'\<close> 1] \<open>\<not> sec b \<le> l\<close> | |
| 43158 | 147 | have "t = t' (\<le> l)" by auto | 
| 67406 | 148 | thus "s = t' (\<le> l)" using \<open>s = t (\<le> l)\<close> by auto | 
| 43158 | 149 | qed | 
| 150 | next | |
| 151 | case (WhileTrue b s1 c s2 s3 t1 t3) | |
| 152 | let ?w = "WHILE b DO c" | |
| 67406 | 153 | have "sec b \<turnstile> c" using \<open>0 \<turnstile> WHILE b DO c\<close> by auto | 
| 43158 | 154 | show ?case | 
| 155 | proof cases | |
| 50342 | 156 | assume "sec b \<le> l" | 
| 67406 | 157 | hence "s1 = t1 (\<le> sec b)" using \<open>s1 = t1 (\<le> l)\<close> by auto | 
| 43158 | 158 | hence "bval b t1" | 
| 67406 | 159 | using \<open>bval b s1\<close> by(simp add: bval_eq_if_eq_le) | 
| 43158 | 160 | then obtain t2 where "(c,t1) \<Rightarrow> t2" "(?w,t2) \<Rightarrow> t3" | 
| 67406 | 161 | using \<open>(?w,t1) \<Rightarrow> t3\<close> by auto | 
| 162 | from WhileTrue.IH(2)[OF \<open>(?w,t2) \<Rightarrow> t3\<close> \<open>0 \<turnstile> ?w\<close> | |
| 163 | WhileTrue.IH(1)[OF \<open>(c,t1) \<Rightarrow> t2\<close> anti_mono[OF \<open>sec b \<turnstile> c\<close>] | |
| 164 | \<open>s1 = t1 (\<le> l)\<close>]] | |
| 43158 | 165 | show ?thesis by simp | 
| 166 | next | |
| 50342 | 167 | assume "\<not> sec b \<le> l" | 
| 67406 | 168 | have 1: "sec b \<turnstile> ?w" by(rule sec_type.intros)(simp_all add: \<open>sec b \<turnstile> c\<close>) | 
| 169 | from confinement[OF big_step.WhileTrue[OF WhileTrue.hyps] 1] \<open>\<not> sec b \<le> l\<close> | |
| 43158 | 170 | have "s1 = s3 (\<le> l)" by auto | 
| 171 | moreover | |
| 67406 | 172 | from confinement[OF \<open>(WHILE b DO c, t1) \<Rightarrow> t3\<close> 1] \<open>\<not> sec b \<le> l\<close> | 
| 43158 | 173 | have "t1 = t3 (\<le> l)" by auto | 
| 67406 | 174 | ultimately show "s3 = t3 (\<le> l)" using \<open>s1 = t1 (\<le> l)\<close> by auto | 
| 43158 | 175 | qed | 
| 176 | qed | |
| 177 | ||
| 178 | ||
| 68776 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 nipkow parents: 
67406diff
changeset | 179 | subsubsection "The Standard Typing System" | 
| 43158 | 180 | |
| 69597 | 181 | text\<open>The predicate \<^prop>\<open>l \<turnstile> c\<close> is nicely intuitive and executable. The | 
| 43158 | 182 | standard formulation, however, is slightly different, replacing the maximum | 
| 183 | computation by an antimonotonicity rule. We introduce the standard system now | |
| 67406 | 184 | and show the equivalence with our formulation.\<close> | 
| 43158 | 185 | |
| 186 | inductive sec_type' :: "nat \<Rightarrow> com \<Rightarrow> bool" ("(_/ \<turnstile>'' _)" [0,0] 50) where
 | |
| 187 | Skip': | |
| 188 | "l \<turnstile>' SKIP" | | |
| 189 | Assign': | |
| 50342 | 190 | "\<lbrakk> sec x \<ge> sec a; sec x \<ge> l \<rbrakk> \<Longrightarrow> l \<turnstile>' x ::= a" | | 
| 47818 | 191 | Seq': | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 192 | "\<lbrakk> l \<turnstile>' c\<^sub>1; l \<turnstile>' c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' c\<^sub>1;;c\<^sub>2" | | 
| 43158 | 193 | If': | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 194 | "\<lbrakk> sec b \<le> l; l \<turnstile>' c\<^sub>1; l \<turnstile>' c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' IF b THEN c\<^sub>1 ELSE c\<^sub>2" | | 
| 43158 | 195 | While': | 
| 50342 | 196 | "\<lbrakk> sec b \<le> l; l \<turnstile>' c \<rbrakk> \<Longrightarrow> l \<turnstile>' WHILE b DO c" | | 
| 43158 | 197 | anti_mono': | 
| 198 | "\<lbrakk> l \<turnstile>' c; l' \<le> l \<rbrakk> \<Longrightarrow> l' \<turnstile>' c" | |
| 199 | ||
| 200 | lemma sec_type_sec_type': "l \<turnstile> c \<Longrightarrow> l \<turnstile>' c" | |
| 45015 | 201 | apply(induction rule: sec_type.induct) | 
| 43158 | 202 | apply (metis Skip') | 
| 203 | apply (metis Assign') | |
| 47818 | 204 | apply (metis Seq') | 
| 54864 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 205 | apply (metis max.commute max.absorb_iff2 nat_le_linear If' anti_mono') | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 206 | by (metis less_or_eq_imp_le max.absorb1 max.absorb2 nat_le_linear While' anti_mono') | 
| 43158 | 207 | |
| 208 | ||
| 209 | lemma sec_type'_sec_type: "l \<turnstile>' c \<Longrightarrow> l \<turnstile> c" | |
| 45015 | 210 | apply(induction rule: sec_type'.induct) | 
| 43158 | 211 | apply (metis Skip) | 
| 212 | apply (metis Assign) | |
| 47818 | 213 | apply (metis Seq) | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 214 | apply (metis max.absorb2 If) | 
| 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 215 | apply (metis max.absorb2 While) | 
| 43158 | 216 | by (metis anti_mono) | 
| 217 | ||
| 68776 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 nipkow parents: 
67406diff
changeset | 218 | subsubsection "A Bottom-Up Typing System" | 
| 43158 | 219 | |
| 220 | inductive sec_type2 :: "com \<Rightarrow> level \<Rightarrow> bool" ("(\<turnstile> _ : _)" [0,0] 50) where
 | |
| 221 | Skip2: | |
| 222 | "\<turnstile> SKIP : l" | | |
| 223 | Assign2: | |
| 50342 | 224 | "sec x \<ge> sec a \<Longrightarrow> \<turnstile> x ::= a : sec x" | | 
| 47818 | 225 | Seq2: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 226 | "\<lbrakk> \<turnstile> c\<^sub>1 : l\<^sub>1; \<turnstile> c\<^sub>2 : l\<^sub>2 \<rbrakk> \<Longrightarrow> \<turnstile> c\<^sub>1;;c\<^sub>2 : min l\<^sub>1 l\<^sub>2 " | | 
| 43158 | 227 | If2: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 228 | "\<lbrakk> sec b \<le> min l\<^sub>1 l\<^sub>2; \<turnstile> c\<^sub>1 : l\<^sub>1; \<turnstile> c\<^sub>2 : l\<^sub>2 \<rbrakk> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52382diff
changeset | 229 | \<Longrightarrow> \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2 : min l\<^sub>1 l\<^sub>2" | | 
| 43158 | 230 | While2: | 
| 50342 | 231 | "\<lbrakk> sec b \<le> l; \<turnstile> c : l \<rbrakk> \<Longrightarrow> \<turnstile> WHILE b DO c : l" | 
| 43158 | 232 | |
| 233 | ||
| 234 | lemma sec_type2_sec_type': "\<turnstile> c : l \<Longrightarrow> l \<turnstile>' c" | |
| 45015 | 235 | apply(induction rule: sec_type2.induct) | 
| 43158 | 236 | apply (metis Skip') | 
| 237 | apply (metis Assign' eq_imp_le) | |
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 238 | apply (metis Seq' anti_mono' min.cobounded1 min.cobounded2) | 
| 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 239 | apply (metis If' anti_mono' min.absorb2 min.absorb_iff1 nat_le_linear) | 
| 43158 | 240 | by (metis While') | 
| 241 | ||
| 242 | lemma sec_type'_sec_type2: "l \<turnstile>' c \<Longrightarrow> \<exists> l' \<ge> l. \<turnstile> c : l'" | |
| 45015 | 243 | apply(induction rule: sec_type'.induct) | 
| 43158 | 244 | apply (metis Skip2 le_refl) | 
| 245 | apply (metis Assign2) | |
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
53015diff
changeset | 246 | apply (metis Seq2 min.boundedI) | 
| 43158 | 247 | apply (metis If2 inf_greatest inf_nat_def le_trans) | 
| 248 | apply (metis While2 le_trans) | |
| 249 | by (metis le_trans) | |
| 250 | ||
| 251 | end |