theory Sec_TypingT imports Sec_Type_Expr begin subsection "A Termination-Sensitive Syntax Directed System" inductive sec_type :: "nat ⇒ com ⇒ bool" ("(_/ ⊢ _)" [0,0] 50) where Skip: "l ⊢ SKIP" | Assign: "⟦ sec x ≥ sec a; sec x ≥ l ⟧ ⟹ l ⊢ x ::= a" | Seq: "l ⊢ c⇩1 ⟹ l ⊢ c⇩2 ⟹ l ⊢ c⇩1;;c⇩2" | If: "⟦ max (sec b) l ⊢ c⇩1; max (sec b) l ⊢ c⇩2 ⟧ ⟹ l ⊢ IF b THEN c⇩1 ELSE c⇩2" | While: "sec b = 0 ⟹ 0 ⊢ c ⟹ 0 ⊢ WHILE b DO c" code_pred (expected_modes: i => i => bool) sec_type . inductive_cases [elim!]: "l ⊢ x ::= a" "l ⊢ c⇩1;;c⇩2" "l ⊢ IF b THEN c⇩1 ELSE c⇩2" "l ⊢ WHILE b DO c" lemma anti_mono: "l ⊢ c ⟹ l' ≤ l ⟹ l' ⊢ c" apply(induction arbitrary: l' rule: sec_type.induct) apply (metis sec_type.intros(1)) apply (metis le_trans sec_type.intros(2)) apply (metis sec_type.intros(3)) apply (metis If le_refl sup_mono sup_nat_def) by (metis While le_0_eq) lemma confinement: "(c,s) ⇒ t ⟹ l ⊢ c ⟹ s = t (< l)" proof(induction rule: big_step_induct) case Skip thus ?case by simp next case Assign thus ?case by auto next case Seq thus ?case by auto next case (IfTrue b s c1) hence "max (sec b) l ⊢ c1" by auto hence "l ⊢ c1" by (metis max.cobounded2 anti_mono) thus ?case using IfTrue.IH by metis next case (IfFalse b s c2) hence "max (sec b) l ⊢ c2" by auto hence "l ⊢ c2" by (metis max.cobounded2 anti_mono) thus ?case using IfFalse.IH by metis next case WhileFalse thus ?case by auto next case (WhileTrue b s1 c) hence "l ⊢ c" by auto thus ?case using WhileTrue by metis qed lemma termi_if_non0: "l ⊢ c ⟹ l ≠ 0 ⟹ ∃ t. (c,s) ⇒ t" apply(induction arbitrary: s rule: sec_type.induct) apply (metis big_step.Skip) apply (metis big_step.Assign) apply (metis big_step.Seq) apply (metis IfFalse IfTrue le0 le_antisym max.cobounded2) apply simp done theorem noninterference: "(c,s) ⇒ s' ⟹ 0 ⊢ c ⟹ s = t (≤ l) ⟹ ∃ t'. (c,t) ⇒ t' ∧ s' = t' (≤ l)" proof(induction arbitrary: t rule: big_step_induct) case Skip thus ?case by auto next case (Assign x a s) have "sec x >= sec a" using `0 ⊢ x ::= a` by auto have "(x ::= a,t) ⇒ t(x := aval a t)" by auto moreover have "s(x := aval a s) = t(x := aval a t) (≤ l)" proof auto assume "sec x ≤ l" with `sec x ≥ sec a` have "sec a ≤ l" by arith thus "aval a s = aval a t" by (rule aval_eq_if_eq_le[OF `s = t (≤ l)`]) next fix y assume "y ≠ x" "sec y ≤ l" thus "s y = t y" using `s = t (≤ l)` by simp qed ultimately show ?case by blast next case Seq thus ?case by blast next case (IfTrue b s c1 s' c2) have "sec b ⊢ c1" "sec b ⊢ c2" using `0 ⊢ IF b THEN c1 ELSE c2` by auto obtain t' where t': "(c1, t) ⇒ t'" "s' = t' (≤ l)" using IfTrue.IH[OF anti_mono[OF `sec b ⊢ c1`] `s = t (≤ l)`] by blast show ?case proof cases assume "sec b ≤ l" hence "s = t (≤ sec b)" using `s = t (≤ l)` by auto hence "bval b t" using `bval b s` by(simp add: bval_eq_if_eq_le) thus ?thesis by (metis t' big_step.IfTrue) next assume "¬ sec b ≤ l" hence 0: "sec b ≠ 0" by arith have 1: "sec b ⊢ IF b THEN c1 ELSE c2" by(rule sec_type.intros)(simp_all add: `sec b ⊢ c1` `sec b ⊢ c2`) from confinement[OF big_step.IfTrue[OF IfTrue(1,2)] 1] `¬ sec b ≤ l` have "s = s' (≤ l)" by auto moreover from termi_if_non0[OF 1 0, of t] obtain t' where t': "(IF b THEN c1 ELSE c2,t) ⇒ t'" .. moreover from confinement[OF t' 1] `¬ sec b ≤ l` have "t = t' (≤ l)" by auto ultimately show ?case using `s = t (≤ l)` by auto qed next case (IfFalse b s c2 s' c1) have "sec b ⊢ c1" "sec b ⊢ c2" using `0 ⊢ IF b THEN c1 ELSE c2` by auto obtain t' where t': "(c2, t) ⇒ t'" "s' = t' (≤ l)" using IfFalse.IH[OF anti_mono[OF `sec b ⊢ c2`] `s = t (≤ l)`] by blast show ?case proof cases assume "sec b ≤ l" hence "s = t (≤ sec b)" using `s = t (≤ l)` by auto hence "¬ bval b t" using `¬ bval b s` by(simp add: bval_eq_if_eq_le) thus ?thesis by (metis t' big_step.IfFalse) next assume "¬ sec b ≤ l" hence 0: "sec b ≠ 0" by arith have 1: "sec b ⊢ IF b THEN c1 ELSE c2" by(rule sec_type.intros)(simp_all add: `sec b ⊢ c1` `sec b ⊢ c2`) from confinement[OF big_step.IfFalse[OF IfFalse(1,2)] 1] `¬ sec b ≤ l` have "s = s' (≤ l)" by auto moreover from termi_if_non0[OF 1 0, of t] obtain t' where t': "(IF b THEN c1 ELSE c2,t) ⇒ t'" .. moreover from confinement[OF t' 1] `¬ sec b ≤ l` have "t = t' (≤ l)" by auto ultimately show ?case using `s = t (≤ l)` by auto qed next case (WhileFalse b s c) hence [simp]: "sec b = 0" by auto have "s = t (≤ sec b)" using `s = t (≤ l)` by auto hence "¬ bval b t" using `¬ bval b s` by (metis bval_eq_if_eq_le le_refl) with WhileFalse.prems(2) show ?case by auto next case (WhileTrue b s c s'' s') let ?w = "WHILE b DO c" from `0 ⊢ ?w` have [simp]: "sec b = 0" by auto have "0 ⊢ c" using `0 ⊢ WHILE b DO c` by auto from WhileTrue.IH(1)[OF this `s = t (≤ l)`] obtain t'' where "(c,t) ⇒ t''" and "s'' = t'' (≤l)" by blast from WhileTrue.IH(2)[OF `0 ⊢ ?w` this(2)] obtain t' where "(?w,t'') ⇒ t'" and "s' = t' (≤l)" by blast from `bval b s` have "bval b t" using bval_eq_if_eq_le[OF `s = t (≤l)`] by auto show ?case using big_step.WhileTrue[OF `bval b t` `(c,t) ⇒ t''` `(?w,t'') ⇒ t'`] by (metis `s' = t' (≤ l)`) qed subsection "The Standard Termination-Sensitive System" text{* The predicate @{prop"l ⊢ c"} is nicely intuitive and executable. The standard formulation, however, is slightly different, replacing the maximum computation by an antimonotonicity rule. We introduce the standard system now and show the equivalence with our formulation. *} inductive sec_type' :: "nat ⇒ com ⇒ bool" ("(_/ ⊢'' _)" [0,0] 50) where Skip': "l ⊢' SKIP" | Assign': "⟦ sec x ≥ sec a; sec x ≥ l ⟧ ⟹ l ⊢' x ::= a" | Seq': "l ⊢' c⇩1 ⟹ l ⊢' c⇩2 ⟹ l ⊢' c⇩1;;c⇩2" | If': "⟦ sec b ≤ l; l ⊢' c⇩1; l ⊢' c⇩2 ⟧ ⟹ l ⊢' IF b THEN c⇩1 ELSE c⇩2" | While': "⟦ sec b = 0; 0 ⊢' c ⟧ ⟹ 0 ⊢' WHILE b DO c" | anti_mono': "⟦ l ⊢' c; l' ≤ l ⟧ ⟹ l' ⊢' c" lemma sec_type_sec_type': "l ⊢ c ⟹ l ⊢' c" apply(induction rule: sec_type.induct) apply (metis Skip') apply (metis Assign') apply (metis Seq') apply (metis max.commute max.absorb_iff2 nat_le_linear If' anti_mono') by (metis While') lemma sec_type'_sec_type: "l ⊢' c ⟹ l ⊢ c" apply(induction rule: sec_type'.induct) apply (metis Skip) apply (metis Assign) apply (metis Seq) apply (metis max.absorb2 If) apply (metis While) by (metis anti_mono) corollary sec_type_eq: "l ⊢ c ⟷ l ⊢' c" by (metis sec_type'_sec_type sec_type_sec_type') end