Reformatting.

This commit is contained in:
Achim D. Brucker 2016-12-28 20:42:31 +00:00
parent 1ab1f96579
commit ef12bd7397
6 changed files with 2373 additions and 2838 deletions

View File

@ -36,8 +36,10 @@
*****************************************************************************)
subsection {* Elementary Firewall Policy Transformation Rules *}
theory ElementaryRules
imports FWNormalisationCore
theory
ElementaryRules
imports
FWNormalisationCore
begin

View File

@ -35,8 +35,11 @@
*****************************************************************************)
chapter {* Firewall Policy Normalisation *}
theory FWNormalisation
imports NormalisationIPPProofs
theory
FWNormalisation
imports
NormalisationIPPProofs
ElementaryRules
begin
end

View File

@ -92,14 +92,16 @@ datatype ('\<alpha>,'\<beta>) Combinators =
text{*
And define the semantic interpretation of it. For technical reasons, we fix here the type to
policies over IntegerPort addresses. However, we could easily provide definitions for other
address types as well, using a generic consts for the type definition and a primitive recursive
definition for each desired address model. *}
address types as well, using a generic constants for the type definition and a primitive
recursive definition for each desired address model.
*}
subsubsection{* Auxiliary definitions and functions. *}
text{*
This subsubsection defines several functions which are useful later for the combinators, invariants,
This section defines several functions which are useful later for the combinators, invariants,
and proofs.
*}
fun srcNet where
"srcNet (DenyAllFromTo x y) = x"
|"srcNet (AllowPortFromTo x y p) = x"
@ -275,8 +277,8 @@ fun NetsCollected2 where
subsubsection{* Transformations *}
text {*
The following two functions transform a policy into a list of single rules and vice-versa - by
staying on the combinator level.
The following two functions transform a policy into a list of single rules and vice-versa (by
staying on the combinator level).
*}
fun policy2list::"('\<alpha>, '\<beta>) Combinators \<Rightarrow>
@ -309,7 +311,9 @@ definition removeShadowRules1_alternative where
"removeShadowRules1_alternative p =
rev (removeShadowRules1_alternative_rev (rev p))"
text{* Remove all the rules which allow a port, but are shadowed by a deny between these subnets *}
text{*
Remove all the rules which allow a port, but are shadowed by a deny between these subnets.
*}
fun removeShadowRules2:: "(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow>
(('\<alpha>, '\<beta>) Combinators) list"
@ -321,8 +325,10 @@ where
| "removeShadowRules2 (x#y) = x#(removeShadowRules2 y)"
| "removeShadowRules2 [] = []"
text{* Sorting a pocliy. We first need to define an ordering on
rules. This ordering depends on the $Nets\_List$ of a policy. *}
text{*
Sorting a policies: We first need to define an ordering on rules. This ordering depends
on the $Nets\_List$ of a policy.
*}
fun smaller :: "('\<alpha>, '\<beta>) Combinators \<Rightarrow>
('\<alpha>, '\<beta>) Combinators \<Rightarrow>
@ -346,15 +352,10 @@ fun qsort where
lemma qsort_permutes:
"set (qsort xs l) = set xs"
apply (induct xs l rule: qsort.induct)
apply (simp_all)
apply auto
done
by (auto)
lemma set_qsort [simp]: "set (qsort xs l) = set xs"
apply (induct xs l rule: qsort.induct)
apply (simp_all)
apply auto
done
by (simp add: qsort_permutes)
fun insort where
"insort a [] l = [a]"
@ -364,12 +365,10 @@ fun sort where
"sort [] l = []"
| "sort (x#xs) l = insort x (sort xs l) l"
fun sorted where
"sorted [] l \<longleftrightarrow> True" |
"sorted [x] l \<longleftrightarrow> True" |
"sorted (x#y#zs) l \<longleftrightarrow> smaller x y l \<and> sorted (y#zs) l"
"sorted [] l = True"
| "sorted [x] l = True"
| "sorted (x#y#zs) l = (smaller x y l \<and> sorted (y#zs) l)"
fun separate where
"separate (DenyAll#x) = DenyAll#(separate x)"
@ -379,7 +378,8 @@ fun separate where
|"separate x = x"
text {*
Insert the DenyAllFromTo rules, such that traffic between two networks can be tested individually
Insert the DenyAllFromTo rules, such that traffic between two networks can be tested
individually.
*}
fun insertDenies where
@ -389,11 +389,11 @@ fun insertDenies where
(insertDenies xs))"
| "insertDenies [] = []"
text{* Remove duplicate rules. This is especially necessary as
insertDenies might have inserted duplicate rules.
The second function is supposed to work on a list of policies. Only
rules which are duplicated within the same policy are removed. *}
text{*
Remove duplicate rules. This is especially necessary as insertDenies might have inserted
duplicate rules. The second function is supposed to work on a list of policies. Only
rules which are duplicated within the same policy are removed.
*}
fun removeDuplicates where
@ -412,10 +412,8 @@ fun insertDeny where
definition "sort' p l = sort l p"
definition "qsort' p l = qsort l p"
declare dom_eq_empty_conv [simp del]
fun list2policyR::"(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow>
@ -425,7 +423,9 @@ fun list2policyR::"(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow>
|"list2policyR [] = undefined "
text{* We provide the definitions for two address representations. *}
text{*
We provide the definitions for two address representations.
*}
subsubsection{* IntPort *}
@ -437,9 +437,6 @@ where
|"C (AllowPortFromTo x y p) = allow_from_to_port p x y"
|"C (x \<oplus> y) = C x ++ C y"
fun CRotate :: "(adr\<^sub>i\<^sub>p net, port) Combinators \<Rightarrow> (adr\<^sub>i\<^sub>p,DummyContent) packet \<mapsto> unit"
where
" CRotate DenyAll = C DenyAll"
@ -447,24 +444,20 @@ where
|"CRotate (AllowPortFromTo x y p) = C (AllowPortFromTo x y p)"
|"CRotate (x \<oplus> y) = ((CRotate y) ++ ((CRotate x)))"
fun rotatePolicy where
"rotatePolicy DenyAll = DenyAll"
| "rotatePolicy (DenyAllFromTo a b) = DenyAllFromTo a b"
| "rotatePolicy (AllowPortFromTo a b p) = AllowPortFromTo a b p"
| "rotatePolicy (a\<oplus>b) = (rotatePolicy b) \<oplus> (rotatePolicy a)"
lemma check: "rev (policy2list (rotatePolicy p)) = policy2list p"
apply (induct p)
apply simp
apply simp_all
done
by (simp_all)
text{*
All rules appearing at the left of a DenyAllFromTo, have disjunct domains from it (except DenyAll)
All rules appearing at the left of a DenyAllFromTo, have disjunct domains from it
(except DenyAll).
*}
fun (sequential) wellformed_policy2 where
"wellformed_policy2 [] = True"
@ -472,8 +465,10 @@ fun (sequential) wellformed_policy2 where
| "wellformed_policy2 (x#xs) = ((\<forall> c a b. c = DenyAllFromTo a b \<and> c \<in> set xs \<longrightarrow>
Map.dom (C x) \<inter> Map.dom (C c) = {}) \<and> wellformed_policy2 xs)"
text{* An allow rule is disjunct with all rules appearing at the right of it. This invariant is not
necessary as it is a consequence from others, but facilitates some proofs. *}
text{*
An allow rule is disjunct with all rules appearing at the right of it. This invariant is not
necessary as it is a consequence from others, but facilitates some proofs.
*}
fun (sequential) wellformed_policy3::"((adr\<^sub>i\<^sub>p net,port) Combinators) list \<Rightarrow> bool" where
"wellformed_policy3 [] = True"
@ -494,7 +489,6 @@ definition
(rm_MT_rules C) o insertDeny o removeShadowRules1 o
policy2list) p"
definition normalize ::
"(adr\<^sub>i\<^sub>p net, port) Combinators \<Rightarrow>
(adr\<^sub>i\<^sub>p net, port) Combinators list"
@ -508,9 +502,6 @@ definition
(sort (removeShadowRules2 (remdups ((rm_MT_rules C) (insertDeny
(removeShadowRules1 (policy2list p)))))) ((l)))))"
definition normalizeQ ::
"(adr\<^sub>i\<^sub>p net, port) Combinators \<Rightarrow>
(adr\<^sub>i\<^sub>p net, port) Combinators list"
@ -524,12 +515,12 @@ definition
(qsort (removeShadowRules2 (remdups ((rm_MT_rules C) (insertDeny
(removeShadowRules1 (policy2list p)))))) ((l)))))"
text{* Of course, normalize is equal to normalize', the latter looks nicer though. *}
text{*
Of course, normalize is equal to normalize', the latter looks nicer though.
*}
lemma "normalize = normalize'"
by (rule ext, simp add: normalize_def normalize'_def sort'_def)
declare C.simps [simp del]
@ -552,19 +543,20 @@ where
|"Dp (AllowPortFromTo x y p) = Cp (AllowPortFromTo x y p)"
|"Dp (x \<oplus> y) = Cp (y \<oplus> x)"
text{* All rules appearing at the left of a DenyAllFromTo, have disjunct domains from it
(except DenyAll) *}
text{*
All rules appearing at the left of a DenyAllFromTo, have disjunct domains from it
(except DenyAll).
*}
fun (sequential) wellformed_policy2Pr where
"wellformed_policy2Pr [] = True"
| "wellformed_policy2Pr (DenyAll#xs) = wellformed_policy2Pr xs"
| "wellformed_policy2Pr (x#xs) = ((\<forall> c a b. c = DenyAllFromTo a b \<and> c \<in> set xs \<longrightarrow>
Map.dom (Cp x) \<inter> Map.dom (Cp c) = {}) \<and> wellformed_policy2Pr xs)"
text{* An allow rule is disjunct with all rules appearing at the right of it. This invariant is not
necessary as it is a consequence from others, but facilitates some proofs. *}
text{*
An allow rule is disjunct with all rules appearing at the right of it. This invariant is not
necessary as it is a consequence from others, but facilitates some proofs.
*}
fun (sequential) wellformed_policy3Pr::"((adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators) list \<Rightarrow> bool" where
"wellformed_policy3Pr [] = True"
@ -572,9 +564,6 @@ fun (sequential) wellformed_policy3Pr::"((adr\<^sub>i\<^sub>p\<^sub>p net, proto
dom (Cp r) \<inter> dom (Cp (AllowPortFromTo a b p)) = {}) \<and> wellformed_policy3Pr xs)"
| "wellformed_policy3Pr (x#xs) = wellformed_policy3Pr xs"
definition
normalizePr' :: "(adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators
\<Rightarrow> (adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators list" where
@ -583,7 +572,6 @@ definition
(rm_MT_rules Cp) o insertDeny o removeShadowRules1 o
policy2list) p"
definition normalizePr ::
"(adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators
\<Rightarrow> (adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators list" where
@ -605,7 +593,6 @@ definition
(rm_MT_rules Cp) o insertDeny o removeShadowRules1 o
policy2list) p"
definition normalizePrQ ::
"(adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators
\<Rightarrow> (adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators list" where
@ -618,37 +605,33 @@ definition
(qsort (removeShadowRules2 (remdups ((rm_MT_rules Cp) (insertDeny
(removeShadowRules1 (policy2list p)))))) ((l)))))"
text{* Of course, normalize is equal to normalize', the latter looks nicer though. *}
text{*
Of course, normalize is equal to normalize', the latter looks nicer though.
*}
lemma "normalizePr = normalizePr'"
by (rule ext, simp add: normalizePr_def normalizePr'_def sort'_def)
text{* The following definition helps in creating the test specification for the individual parts
of a normalized policy. *}
text{*
The following definition helps in creating the test specification for the individual parts
of a normalized policy.
*}
definition makeFUTPr where
"makeFUTPr FUT p x n =
(packet_Nets x (fst (normBothNets (bothNets p)!n))
(snd(normBothNets (bothNets p)!n)) \<longrightarrow>
FUT x = Cp ((normalizePr p)!Suc n) x)"
declare Cp.simps [simp del]
lemmas PLemmas = C.simps Cp.simps dom_def PolicyCombinators.PolicyCombinators
PortCombinators.PortCombinatorsCore aux
ProtocolPortCombinators.ProtocolCombinatorsCore src_def dest_def in_subnet_def
adr\<^sub>i\<^sub>p\<^sub>pLemmas adr\<^sub>i\<^sub>p\<^sub>pLemmas
lemma aux: "\<lbrakk>x \<noteq> a; y\<noteq>b; (x \<noteq> y \<and> x \<noteq> b) \<or> (a \<noteq> b \<and> a \<noteq> y)\<rbrakk> \<Longrightarrow> {x,a} \<noteq> {y,b}"
by (auto)
lemma aux2: "{a,b} = {b,a}"
by auto
end

View File

@ -36,31 +36,28 @@
*****************************************************************************)
subsection{* Normalisation Proofs (Generic) *}
theory NormalisationGenericProofs
theory
NormalisationGenericProofs
imports
FWNormalisationCore
begin
text {* This theory contains the generic proofs of the normalisation
procedure, i.e. those which are independent from the concrete
semantical interpretation function. *}
text {*
This theory contains the generic proofs of the normalisation procedure, i.e. those which
are independent from the concrete semantical interpretation function.
*}
lemma domNMT: "dom X \<noteq> {} \<Longrightarrow> X \<noteq> \<emptyset>"
apply auto
done
by auto
lemma denyNMT: "deny_all \<noteq> \<emptyset>"
apply (rule domNMT)
apply (simp add: deny_all_def dom_def)
done
by (simp add: deny_all_def dom_def)
lemma wellformed_policy1_charn[rule_format] : "wellformed_policy1 p \<longrightarrow>
DenyAll \<in> set p \<longrightarrow> (\<exists> p'. p = DenyAll # p' \<and> DenyAll \<notin> set p')"
lemma wellformed_policy1_charn[rule_format]:
"wellformed_policy1 p \<longrightarrow> DenyAll \<in> set p \<longrightarrow> (\<exists> p'. p = DenyAll # p' \<and> DenyAll \<notin> set p')"
by(induct p,simp_all)
lemma singleCombinatorsConc: "singleCombinators (x#xs) \<Longrightarrow> singleCombinators xs"
by (case_tac x,simp_all)
@ -69,7 +66,6 @@ lemma aux0_0: "singleCombinators x \<Longrightarrow> \<not> (\<exists> a b. (a\<
apply (rule allI)+
by (case_tac a,simp_all)
lemma aux0_4: "(a \<in> set x \<or> a \<in> set y) = (a \<in> set (x@y))"
by auto
@ -77,27 +73,22 @@ lemma aux0_1: "\<lbrakk>singleCombinators xs; singleCombinators [x]\<rbrakk> \<L
singleCombinators (x#xs)"
by (case_tac x,simp_all)
lemma aux0_6: "\<lbrakk>singleCombinators xs; \<not> (\<exists> a b. x = a \<oplus> b)\<rbrakk> \<Longrightarrow>
singleCombinators(x#xs)"
apply (rule aux0_1,simp_all)
apply (case_tac x,simp_all)
apply auto
done
by auto
lemma aux0_5: " \<not> (\<exists> a b. (a\<oplus>b) \<in> set x) \<Longrightarrow> singleCombinators x"
apply (induct x)
apply simp_all
by (metis aux0_6)
lemma ANDConc[rule_format]: "allNetsDistinct (a#p) \<longrightarrow> allNetsDistinct (p)"
apply (simp add: allNetsDistinct_def)
apply (case_tac "a")
by simp_all
lemma aux6: "twoNetsDistinct a1 a2 a b \<Longrightarrow>
dom (deny_all_from_to a1 a2) \<inter> dom (deny_all_from_to a b) = {}"
by (auto simp: twoNetsDistinct_def netsDistinct_def src_def dest_def
@ -106,7 +97,6 @@ lemma aux6: "twoNetsDistinct a1 a2 a b \<Longrightarrow>
lemma aux5[rule_format]: "(DenyAllFromTo a b) \<in> set p \<longrightarrow> a \<in> set (net_list p)"
by (rule net_list_aux.induct,simp_all)
lemma aux5a[rule_format]: "(DenyAllFromTo b a) \<in> set p \<longrightarrow> a \<in> set (net_list p)"
by (rule net_list_aux.induct,simp_all)
@ -121,25 +111,22 @@ lemma aux5d[rule_format]:
lemma aux10[rule_format]: "a \<in> set (net_list p) \<longrightarrow> a \<in> set (net_list_aux p)"
by simp
lemma srcInNetListaux[simp]: "\<lbrakk>x \<in> set p; singleCombinators[x]; x \<noteq> DenyAll\<rbrakk> \<Longrightarrow>
srcNet x \<in> set (net_list_aux p)"
lemma srcInNetListaux[simp]:
"\<lbrakk>x \<in> set p; singleCombinators[x]; x \<noteq> DenyAll\<rbrakk> \<Longrightarrow> srcNet x \<in> set (net_list_aux p)"
apply (induct p)
apply simp_all
apply (case_tac "x = a", simp_all)
apply (case_tac a, simp_all)+
done
lemma destInNetListaux[simp]: "\<lbrakk>x \<in> set p; singleCombinators[x]; x \<noteq> DenyAll\<rbrakk> \<Longrightarrow>
destNet x \<in> set (net_list_aux p)"
lemma destInNetListaux[simp]:
"\<lbrakk>x \<in> set p; singleCombinators[x]; x \<noteq> DenyAll\<rbrakk> \<Longrightarrow> destNet x \<in> set (net_list_aux p)"
apply (induct p)
apply simp_all
apply (case_tac "x = a", simp_all)
apply (case_tac a, simp_all)+
done
lemma tND1: "\<lbrakk>allNetsDistinct p; x \<in> set p; y \<in> set p; a = srcNet x;
b = destNet x; c = srcNet y; d = destNet y; a \<noteq> c;
singleCombinators[x]; x \<noteq> DenyAll; singleCombinators[y];
@ -189,8 +176,6 @@ lemma aux[rule_format]: "a \<in> set (removeShadowRules2 p) \<longrightarrow> a
lemma aux12: "\<lbrakk>a \<in> x; b \<notin> x\<rbrakk> \<Longrightarrow> a \<noteq> b"
by auto
lemma ND0aux1[rule_format]: "DenyAllFromTo x y \<in> set b \<Longrightarrow>
x \<in> set (net_list_aux b)"
by (metis aux5 net_list.simps set_remdups)
@ -233,8 +218,7 @@ lemma aNDSubset: "\<lbrakk>singleCombinators a;set a \<subseteq> set b; allNetsD
apply (rule allI)+
apply (rule impI)+
apply (drule_tac x = "aa" in spec, drule_tac x = "ba" in spec)
apply (metis subsetD aNDSubsetaux)
done
using aNDSubsetaux by blast
lemma aNDSetsEq: "\<lbrakk>singleCombinators a; singleCombinators b; set a = set b;
allNetsDistinct b\<rbrakk> \<Longrightarrow> allNetsDistinct a"
@ -242,21 +226,15 @@ lemma aNDSetsEq: "\<lbrakk>singleCombinators a; singleCombinators b; set a = set
apply (rule allI)+
apply (rule impI)+
apply (drule_tac x = "aa" in spec, drule_tac x = "ba" in spec)
apply (metis aNDSetsEqaux )
done
using aNDSetsEqaux by auto
lemma SCConca: "\<lbrakk>singleCombinators p; singleCombinators [a]\<rbrakk> \<Longrightarrow>
singleCombinators (a#p)"
by (case_tac "a",simp_all)
by(metis aux0_1)
lemma aux3[simp]: "\<lbrakk>singleCombinators p; singleCombinators [a];
allNetsDistinct (a#p)\<rbrakk> \<Longrightarrow> allNetsDistinct (a#a#p)"
apply (insert aNDSubset[of "(a#a#p)" "(a#p)"])
by (simp add: SCConca)
by (metis aNDSetsEq aux0_1 insert_absorb2 list.set(2))
lemma wp1_aux1a[rule_format]: "xs \<noteq> [] \<longrightarrow> wellformed_policy1_strong (xs @ [x]) \<longrightarrow>
wellformed_policy1_strong xs"
@ -307,10 +285,9 @@ lemma waux3[rule_format]: "\<lbrakk>x \<noteq> a; x \<notin> set p\<rbrakk> \<L
by (metis aux79)
lemma wellformed1_sorted_aux[rule_format]: "wellformed_policy1 (x#p) \<Longrightarrow>
wellformed_policy1 (insort x p l)"
apply (case_tac x,simp_all)
by (rule waux2,rule waux3, simp_all)+
wellformed_policy1 (insort x p l)" try
by (metis NormalisationGenericProofs.set_insort list.set(2) saux waux2 wellformed_eq
wellformed_policy1_strong.simps(2))
lemma wellformed1_sorted_auxQ[rule_format]: "wellformed_policy1 (p) \<Longrightarrow>
wellformed_policy1 (qsort p l)"
@ -323,9 +300,6 @@ next
by (metis Combinators.simps append_Cons append_Nil qsort.simps(2) set_ConsD set_qsort waux2)+
qed
lemma SR1Subset: "set (removeShadowRules1 p) \<subseteq> set p"
apply (induct_tac p, simp_all)
apply (case_tac a, simp_all)
@ -363,7 +337,6 @@ lemma denyAllDom[simp]: "x \<in> dom (deny_all)"
lemma lCdom2: "(list2FWpolicy (a @ (b @ c))) = (list2FWpolicy ((a@b)@c))"
by auto
lemma SCConcEnd: "singleCombinators (xs @ [xa]) \<Longrightarrow> singleCombinators xs"
by (induct "xs", simp_all, case_tac a, simp_all)
@ -371,7 +344,6 @@ lemma list2FWpolicyconc[rule_format]: "a \<noteq> [] \<longrightarrow>
(list2FWpolicy (xa # a)) = (xa) \<oplus> (list2FWpolicy a)"
by (induct a,simp_all)
lemma wp1n_tl [rule_format]: "wellformed_policy1_strong p \<longrightarrow>
p = (DenyAll#(tl p))"
by (induct p, simp_all)
@ -424,7 +396,6 @@ lemma domInterMT[rule_format]: "\<lbrakk>dom a \<inter> dom b = {}; x \<in> dom
lemma domComm: "dom a \<inter> dom b = dom b \<inter> dom a"
by auto
lemma r_not_DA_in_tl[rule_format]:
"wellformed_policy1_strong p \<longrightarrow> a \<in> set p\<longrightarrow> a \<noteq> DenyAll \<longrightarrow> a \<in> set (tl p)"
by (induct p,simp_all)
@ -442,8 +413,6 @@ lemma l2p_aux[rule_format]: "list \<noteq> [] \<longrightarrow>
lemma l2p_aux2[rule_format]: "list = [] \<Longrightarrow> list2FWpolicy (a # list) = a"
by simp
lemma aux7aa:
assumes 1 : "AllowPortFromTo a b poo \<in> set p"
and 2 : "allNetsDistinct ((AllowPortFromTo c d po) # p)"
@ -459,10 +428,9 @@ next
list.set_intros(2) twoNetsDistinct_def)
qed
lemma ANDConcEnd: "\<lbrakk> allNetsDistinct (xs @ [xa]); singleCombinators xs\<rbrakk> \<Longrightarrow>
allNetsDistinct xs"
by (rule aNDSubset) auto
by (rule aNDSubset, auto)
lemma WP1ConcEnd[rule_format]:
"wellformed_policy1 (xs@[xa]) \<longrightarrow> wellformed_policy1 xs"
@ -505,21 +473,15 @@ next
qed
qed
lemma SC1[simp]: "singleCombinators p \<Longrightarrow>singleCombinators (removeShadowRules1 p)"
by (erule SCSubset) (rule SR1Subset)
lemma SC2[simp]: "singleCombinators p \<Longrightarrow>singleCombinators (removeShadowRules2 p)"
by (erule SCSubset) (rule RS2Set)
lemma SC3[simp]: "singleCombinators p \<Longrightarrow> singleCombinators (sort p l)"
by (erule SCSubset) simp
lemma SC3Q[simp]: "singleCombinators p \<Longrightarrow> singleCombinators (qsort p l)"
by (erule SCSubset) simp
@ -549,8 +511,6 @@ lemma aND_sortQ[simp]: "\<lbrakk>singleCombinators p; allNetsDistinct p\<rbrakk>
apply (rule aNDSubset)
by (erule SC3Q, simp_all)
lemma inRS2[rule_format,simp]: "x \<notin> set p \<longrightarrow> x \<notin> set (removeShadowRules2 p)"
apply (insert RS2Set [of p])
by blast
@ -715,8 +675,6 @@ lemma sorted_ConsA:"\<lbrakk>all_in_list (x#xs) l; singleCombinators (x#xs)\<rbr
lemma is_in_insort: "y \<in> set xs \<Longrightarrow> y \<in> set (insort x xs l)"
by (simp add: NormalisationGenericProofs.set_insort)
lemma sorted_insorta[rule_format]:
assumes 1 : "sorted (insort x xs l) l"
and 2 : "all_in_list (x#xs) l"
@ -732,7 +690,7 @@ next
then show ?case
apply simp
apply (auto intro: is_in_insort sorted_ConsA set_insort singleCombinators_insort
singleCombinatorsConc sortedConcEnd all_in_list_insort)
singleCombinatorsConc sortedConcEnd all_in_list_insort) [1]
apply(cases "smaller x a l", simp_all)
by (metis NormalisationGenericProofs.set_insort NormalisationGenericProofs.sorted_Cons
all_in_list.simps(2) all_in_list_insort aux0_1 insert_iff singleCombinatorsConc
@ -750,7 +708,7 @@ next
have * : "sorted (a # xs) l \<Longrightarrow> all_in_list (x # a # xs) l \<Longrightarrow>
distinct (x # a # xs) \<Longrightarrow> singleCombinators [x] \<Longrightarrow>
singleCombinators (a # xs) \<Longrightarrow> sorted (insort x xs l) l"
apply(insert Cons.hyps)apply simp_all
apply(insert Cons.hyps, simp_all)
apply(metis sorted_Cons all_in_list.simps(2) singleCombinatorsConc)
done
show ?case
@ -789,7 +747,6 @@ next
qed
qed
lemma sorted_insort:
"\<lbrakk>all_in_list (x#xs) l; distinct(x#xs); singleCombinators [x];
singleCombinators xs\<rbrakk> \<Longrightarrow>
@ -806,52 +763,42 @@ lemma distinct_sort[simp]: "distinct (sort xs l) = distinct xs"
lemma sort_is_sorted[rule_format]:
"all_in_list p l \<longrightarrow> distinct p \<longrightarrow> singleCombinators p \<longrightarrow> sorted (sort p l) l"
apply (induct p)
apply (auto intro: SC3 all_in_listSubset singleCombinatorsConc sorted_insort)
apply (subst sorted_insort)
apply (auto intro: singleCombinatorsConc all_in_listSubset SC3)
apply (erule all_in_listSubset)
by (auto intro: SC3 singleCombinatorsConc sorted_insort)
apply simp
by (metis (no_types, lifting) NormalisationGenericProofs.distinct_sort
NormalisationGenericProofs.set_sort SC3 all_in_list.simps(2) all_in_listSubset
distinct.simps(2) set_subset_Cons singleCombinatorsConc singleCombinatorsStart
sort.simps(2) sorted_insortb)
lemma smaller_sym[rule_format]: "all_in_list [a] l \<longrightarrow> smaller a a l"
by (case_tac a,simp_all)
lemma SC_sublist[rule_format]:
"singleCombinators xs \<Longrightarrow> singleCombinators (qsort [y\<leftarrow>xs. P y] l)"
by (auto intro: SCSubset)
lemma all_in_list_sublist[rule_format]:
"singleCombinators xs \<longrightarrow> all_in_list xs l \<longrightarrow> all_in_list (qsort [y\<leftarrow>xs. P y] l) l"
by (auto intro: all_in_listSubset SC_sublist)
lemma SC_sublist2[rule_format]:
"singleCombinators xs \<longrightarrow> singleCombinators ([y\<leftarrow>xs. P y])"
by (auto intro: SCSubset)
lemma all_in_list_sublist2[rule_format]:
"singleCombinators xs \<longrightarrow> all_in_list xs l \<longrightarrow> all_in_list ( [y\<leftarrow>xs. P y]) l"
by (auto intro: all_in_listSubset SC_sublist2)
lemma all_in_listAppend[rule_format]:
"all_in_list (xs) l \<longrightarrow> all_in_list (ys) l \<longrightarrow> all_in_list (xs @ ys) l"
by (induct xs) simp_all
lemma distinct_sortQ[rule_format]:
"singleCombinators xs \<longrightarrow> all_in_list xs l \<longrightarrow> distinct xs \<longrightarrow> distinct (qsort xs l)"
apply (induct xs l rule: qsort.induct)
apply simp
apply (auto simp: SC_sublist2 singleCombinatorsConc all_in_list_sublist2)
done
lemma singleCombinatorsAppend[rule_format]:
"singleCombinators (xs) \<longrightarrow> singleCombinators (ys) \<longrightarrow> singleCombinators (xs @ ys)"
apply (induct xs, auto)
@ -889,9 +836,6 @@ lemma sorted_append[rule_format]:
apply (rule sorted_append2,simp_all)
done
lemma sort_is_sortedQ[rule_format]:
"all_in_list p l \<longrightarrow> singleCombinators p \<longrightarrow> sorted (qsort p l) l"
proof (induct p l rule: qsort.induct) print_cases
@ -941,7 +885,7 @@ next
FWNormalisationCore.sorted (x # qsort [y\<leftarrow>xs . smaller x y l] l) l \<and>
(\<forall>x'\<in>set (qsort [y\<leftarrow>xs . \<not> smaller x y l] l).
\<forall>y\<in>set (x # qsort [y\<leftarrow>xs . smaller x y l] l). smaller x' y l)"
apply auto
apply(auto)[1]
apply (metis (mono_tags, lifting) SC_sublist all_in_list.simps(2)
all_in_list_sublist aux0_1 mem_Collect_eq set_filter set_qsort
singleCombinatorsConc singleCombinatorsStart sorted_Consb)
@ -974,7 +918,6 @@ lemma RS1n_nMT[rule_format,simp]: "p \<noteq> []\<longrightarrow> removeShadowRu
lemma RS1N_DA[simp]: "removeShadowRules1_alternative (a@[DenyAll]) = [DenyAll]"
by (simp add: removeShadowRules1_alternative_def)
lemma WP1n_DA_notinSet[rule_format]: "wellformed_policy1_strong p \<longrightarrow>
DenyAll \<notin> set (tl p)"
by (induct p) (simp_all)
@ -991,11 +934,9 @@ lemma AND_tl[rule_format]: "allNetsDistinct ( p) \<longrightarrow> allNetsDistin
apply (induct p, simp_all)
by (auto intro: ANDConc)
lemma distinct_tl[rule_format]: "distinct p \<longrightarrow> distinct (tl p)"
by (induct p, simp_all)
lemma SC_tl[rule_format]: "singleCombinators ( p) \<longrightarrow> singleCombinators (tl p)"
by (induct p, simp_all) (auto intro: singleCombinatorsConc)
@ -1006,7 +947,6 @@ lemma wp1_tl[rule_format]:
"p \<noteq> [] \<and> wellformed_policy1 p \<longrightarrow> wellformed_policy1 (tl p)"
by (induct p) (auto intro: waux2)
lemma wp1_eq[rule_format]:
"wellformed_policy1_strong p \<Longrightarrow> wellformed_policy1 p"
apply (case_tac "DenyAll \<in> set p")
@ -1014,13 +954,10 @@ lemma wp1_eq[rule_format]:
apply (auto elim: waux2)
done
lemma wellformed1_alternative_sorted:
"wellformed_policy1_strong p \<Longrightarrow> wellformed_policy1_strong (sort p l)"
by (case_tac "p", simp_all)
lemma wp1n_RS2[rule_format]:
"wellformed_policy1_strong p \<longrightarrow> wellformed_policy1_strong (removeShadowRules2 p)"
by (induct p, simp_all)
@ -1031,7 +968,6 @@ lemma RS2_NMT[rule_format]: "p \<noteq> [] \<longrightarrow> removeShadowRules2
apply (case_tac "a", simp_all)+
done
lemma wp1_alternative_not_mt[simp]: "wellformed_policy1_strong p \<Longrightarrow> p \<noteq> []"
by auto
@ -1039,8 +975,6 @@ lemma AIL1[rule_format,simp]: "all_in_list p l \<longrightarrow>
all_in_list (removeShadowRules1 p) l"
by (induct_tac p, simp_all)
lemma wp1ID: "wellformed_policy1_strong (insertDeny (removeShadowRules1 p))"
by (induct p, simp_all, case_tac a, simp_all)
@ -1050,7 +984,6 @@ lemma dRD[simp]: "distinct (remdups p)"
lemma AILrd[rule_format,simp]: "all_in_list p l \<longrightarrow> all_in_list (remdups p) l"
by (induct p, simp_all)
lemma AILiD[rule_format,simp]: "all_in_list p l \<longrightarrow> all_in_list (insertDeny p) l"
apply (induct p, simp_all)
apply (rule impI, simp)
@ -1069,7 +1002,6 @@ lemma SCRiD[rule_format,simp]: "singleCombinators p \<longrightarrow>
apply (case_tac "a", simp_all)
done
lemma WP1rd[rule_format,simp]:
"wellformed_policy1_strong p \<longrightarrow> wellformed_policy1_strong (remdups p)"
by (induct p, simp_all)
@ -1089,7 +1021,6 @@ lemma ANDiD[rule_format,simp]:
apply (case_tac "a",simp_all add: allNetsDistinct_def)
done
lemma mr_iD[rule_format]:
"wellformed_policy1_strong p \<longrightarrow> matching_rule x p = matching_rule x (insertDeny p)"
by (induct p, simp_all)
@ -1098,7 +1029,6 @@ lemma WP1iD[rule_format,simp]: "wellformed_policy1_strong p \<longrightarrow>
wellformed_policy1_strong (insertDeny p)"
by (induct p, simp_all)
lemma DAiniD: "DenyAll \<in> set (insertDeny p)"
by (induct p, simp_all, case_tac a, simp_all)
@ -1109,7 +1039,6 @@ lemma AIL2[rule_format,simp]:
"all_in_list p l \<longrightarrow> all_in_list (removeShadowRules2 p) l"
by (induct_tac p, simp_all, case_tac a, simp_all)
lemma SCConc: "singleCombinators x \<Longrightarrow> singleCombinators y \<Longrightarrow> singleCombinators (x@y)"
apply (rule aux0_5)
apply (metis aux0_0 aux0_4)
@ -1118,7 +1047,6 @@ lemma SCConc: "singleCombinators x \<Longrightarrow> singleCombinators y \<Long
lemma SCp2l: "singleCombinators (policy2list p)"
by (induct_tac p) (auto intro: SCConc)
lemma subnetAux: "Dd \<inter> A \<noteq> {} \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Dd \<inter> B \<noteq> {}"
by auto
@ -1126,8 +1054,7 @@ lemma soadisj: "x \<in> subnetsOfAdr a \<Longrightarrow> y \<in> subnetsOfAdr a
by(simp add: subnetsOfAdr_def netsDistinct_def,auto)
lemma not_member: "\<not> member a (x\<oplus>y) \<Longrightarrow> \<not> member a x"
apply auto
done
by auto
lemma soadisj2: "(\<forall> a x y. x \<in> subnetsOfAdr a \<and> y \<in> subnetsOfAdr a \<longrightarrow> \<not> netsDistinct x y)"
by (simp add: subnetsOfAdr_def netsDistinct_def, auto)
@ -1139,14 +1066,12 @@ lemma ndFalse1:
apply (auto simp: soadisj)
using soadisj2 by blast
lemma ndFalse2: "(\<forall>a b c d. (a,b)\<in>A \<and> (c,d)\<in>B \<longrightarrow> netsDistinct b d) \<Longrightarrow>
\<exists>(a, b)\<in>A. b \<in> subnetsOfAdr D \<Longrightarrow>
\<exists>(a, b)\<in>B. b \<in> subnetsOfAdr D \<Longrightarrow> False"
apply (auto simp: soadisj)
using soadisj2 by blast
lemma tndFalse: "(\<forall>a b c d. (a,b)\<in>A \<and> (c,d)\<in>B \<longrightarrow> twoNetsDistinct a b c d) \<Longrightarrow>
\<exists>(a, b)\<in>A. a \<in> subnetsOfAdr (D::('a::adr)) \<and> b \<in> subnetsOfAdr (F::'a) \<Longrightarrow>
\<exists>(a, b)\<in>B. a \<in> subnetsOfAdr D\<and> b\<in> subnetsOfAdr F
@ -1162,7 +1087,6 @@ lemma sepnMT[rule_format]: "p \<noteq> [] \<longrightarrow> (separate p) \<noteq
lemma sepDA[rule_format]: "DenyAll \<notin> set p \<longrightarrow> DenyAll \<notin> set (separate p)"
by (induct p rule: separate.induct) simp_all
lemma setnMT: "set a = set b \<Longrightarrow> a \<noteq> [] \<Longrightarrow> b \<noteq> []"
by auto
@ -1185,7 +1109,6 @@ lemma OTNoTN[rule_format]: " OnlyTwoNets p \<longrightarrow> x \<noteq> DenyAll
lemma first_isIn[rule_format]: "\<not> member DenyAll x \<longrightarrow> (first_srcNet x,first_destNet x) \<in> sdnets x"
by (induct x,case_tac x, simp_all)
lemma sdnets2:
"\<exists>a b. sdnets x = {(a, b), (b, a)} \<Longrightarrow> \<not> member DenyAll x \<Longrightarrow>
sdnets x = {(first_srcNet x, first_destNet x), (first_destNet x, first_srcNet x)}"
@ -1193,7 +1116,6 @@ proof -
have * : "\<exists>a b. sdnets x = {(a, b), (b, a)} \<Longrightarrow> \<not> member DenyAll x
\<Longrightarrow> (first_srcNet x, first_destNet x) \<in> sdnets x"
by (erule first_isIn)
show "\<exists>a b. sdnets x = {(a, b), (b, a)} \<Longrightarrow> \<not> member DenyAll x \<Longrightarrow>
sdnets x = {(first_srcNet x, first_destNet x), (first_destNet x, first_srcNet x)}"
using * by auto
@ -1207,7 +1129,6 @@ lemma alternativelistconc2[rule_format]:
"a \<in> set (net_list_aux [x]) \<longrightarrow> a \<in> set (net_list_aux [y,x])"
by (induct y, simp_all)
lemma noDA[rule_format]:
"noDenyAll xs \<longrightarrow> s \<in> set xs \<longrightarrow> \<not> member DenyAll s"
by (induct xs, simp_all)
@ -1240,7 +1161,6 @@ lemma disjComm: "disjSD_2 a b \<Longrightarrow> disjSD_2 b a"
using tNDComm apply blast
by (meson tNDComm twoNetsDistinct_def)
lemma disjSD2aux:
"disjSD_2 a b \<Longrightarrow> \<not> member DenyAll a \<Longrightarrow> \<not> member DenyAll b \<Longrightarrow>
disjSD_2 (DenyAllFromTo (first_srcNet a) (first_destNet a) \<oplus>
@ -1250,11 +1170,9 @@ lemma disjSD2aux:
apply (simp add: disjSD_2_def)
using first_isIn by blast
lemma noDA1eq[rule_format]: "noDenyAll p \<longrightarrow> noDenyAll1 p"
by (induct p, simp,rename_tac a p, case_tac a, simp_all)
lemma noDA1C[rule_format]: "noDenyAll1 (a#p) \<longrightarrow> noDenyAll1 p"
by (case_tac a, simp_all,rule impI, rule noDA1eq, simp)+
@ -1267,11 +1185,9 @@ lemma disjSD_2IDa:
disjSD_2 (DenyAllFromTo a b \<oplus> DenyAllFromTo b a \<oplus> x) y"
by(simp add:disjSD2aux)
lemma noDAID[rule_format]: "noDenyAll p \<longrightarrow> noDenyAll (insertDenies p)"
by (induct p, simp_all,case_tac a, simp_all)
lemma isInIDo[rule_format]:
"noDenyAll p \<longrightarrow> s \<in> set (insertDenies p) \<longrightarrow>
(\<exists>! a. s = (DenyAllFromTo (first_srcNet a) (first_destNet a)) \<oplus>
@ -1286,7 +1202,6 @@ lemma id_aux1[rule_format]: "DenyAllFromTo (first_srcNet s) (first_destNet s) \<
\<longrightarrow> s \<in> set p"
by (induct p, simp_all, rename_tac a p, case_tac a, simp_all)
lemma id_aux2:
"noDenyAll p \<Longrightarrow>
\<forall>s. s \<in> set p \<longrightarrow> disjSD_2 a s \<Longrightarrow>
@ -1297,8 +1212,6 @@ lemma id_aux2:
DenyAllFromTo (first_destNet s) (first_srcNet s) \<oplus> s)"
by (metis disjComm disjSD2aux isInIDo noDA)
lemma id_aux4[rule_format]:
"noDenyAll p \<Longrightarrow> \<forall>s. s \<in> set p \<longrightarrow> disjSD_2 a s \<Longrightarrow>
s \<in> set (insertDenies p) \<Longrightarrow> \<not> member DenyAll a \<Longrightarrow>
@ -1310,10 +1223,6 @@ lemma id_aux4[rule_format]:
apply (drule_tac Q = "disjSD_2 a s" in exE, simp_all, rule id_aux2, simp_all)
using isInIDo by blast
(* XXX *)
lemma sepNetsID[rule_format]:
"noDenyAll1 p \<longrightarrow> separated p \<longrightarrow> separated (insertDenies p)"
apply (induct p, simp)
@ -1328,7 +1237,6 @@ lemma sepNetsID[rule_format]:
lemma aNDDA[rule_format]: "allNetsDistinct p \<longrightarrow> allNetsDistinct(DenyAll#p)"
by (case_tac p,auto simp: allNetsDistinct_def)
lemma OTNConc[rule_format]: "OnlyTwoNets (y # z) \<longrightarrow> OnlyTwoNets z"
by (case_tac y, simp_all)
@ -1353,7 +1261,6 @@ lemma setPair4: "{a,b} = {c,d} \<Longrightarrow> a \<noteq> c \<Longrightarrow>
lemma otnaux1: " {x, y, x, y} = {x,y}"
by auto
lemma OTNIDaux4: "{x,y,x} = {y,x}"
by auto
@ -1383,8 +1290,7 @@ lemma otnaux: "
apply simp
apply (subgoal_tac "first_destNet x = first_srcNet y")
apply simp
apply (rule_tac x ="first_srcNet y" in exI,
rule_tac x = "first_destNet y" in exI,simp)
apply (rule_tac x ="first_srcNet y" in exI, rule_tac x = "first_destNet y" in exI,simp)
apply (rule setPair1)
apply simp
apply (rule setPair4)
@ -1405,8 +1311,7 @@ lemma otnaux: "
apply simp
apply (subgoal_tac "first_destNet x = first_srcNet y")
apply simp
apply (rule_tac x ="first_srcNet y" in exI,
rule_tac x = "first_destNet y" in exI)
apply (rule_tac x ="first_srcNet y" in exI, rule_tac x = "first_destNet y" in exI)
apply (metis OTNIDaux4 insert_commute )
apply (rule setPair1)
apply simp
@ -1426,8 +1331,7 @@ lemma otnaux: "
apply simp_all
apply (subgoal_tac "first_destNet x = first_destNet y")
apply simp
apply (rule_tac x ="first_srcNet y" in exI,
rule_tac x = "first_destNet y" in exI)
apply (rule_tac x ="first_srcNet y" in exI, rule_tac x = "first_destNet y" in exI)
apply (metis OTNIDaux4 insert_commute )
apply (rule setPair)
apply simp
@ -1453,8 +1357,7 @@ lemma otnaux: "
apply simp_all
apply (subgoal_tac "first_destNet x = first_destNet y")
apply simp
apply (rule_tac x ="first_srcNet y" in exI,
rule_tac x = "first_destNet y" in exI)
apply (rule_tac x ="first_srcNet y" in exI, rule_tac x = "first_destNet y" in exI)
apply (rule otnaux1)
apply (rule setPair)
apply simp
@ -1462,8 +1365,7 @@ lemma otnaux: "
apply simp
apply (subgoal_tac "first_destNet x = first_srcNet y")
apply simp
apply (rule_tac x ="first_srcNet y" in exI,
rule_tac x = "first_destNet y" in exI)
apply (rule_tac x ="first_srcNet y" in exI, rule_tac x = "first_destNet y" in exI)
apply (metis OTNIDaux4 insert_commute)
apply (rule setPair1)
apply simp
@ -1503,8 +1405,7 @@ lemma nda[rule_format]:
done
lemma nDAcharn[rule_format]: "noDenyAll p = (\<forall> r \<in> set p. \<not> member DenyAll r)"
by (induct p) simp_all
by (induct p, simp_all)
lemma nDAeqSet: "set p = set s \<Longrightarrow> noDenyAll p = noDenyAll s"
by (simp add: nDAcharn)
@ -1520,40 +1421,32 @@ lemma nDASC[rule_format]:
apply (induct p, simp_all)
using nDASCaux nDAcharn nda singleCombinatorsConc by blast
lemma noDAAll[rule_format]: "noDenyAll p = (\<not> memberP DenyAll p)"
by (induct p) simp_all
lemma memberPsep[symmetric]: "memberP x p = memberP x (separate p)"
by (induct p rule: separate.induct) simp_all
lemma noDAsep[rule_format]: "noDenyAll p \<Longrightarrow> noDenyAll (separate p)"
by (simp add:noDAAll,subst memberPsep, simp)
lemma noDA1sep[rule_format]: "noDenyAll1 p \<longrightarrow> noDenyAll1 (separate p)"
by (induct p rule:separate.induct, simp_all add: noDAsep)
lemma isInAlternativeLista:
"(aa \<in> set (net_list_aux [a]))\<Longrightarrow> aa \<in> set (net_list_aux (a # p))"
by (case_tac a,auto)
lemma isInAlternativeListb:
"(aa \<in> set (net_list_aux p))\<Longrightarrow> aa \<in> set (net_list_aux (a # p))"
by (case_tac a,simp_all)
lemma ANDSepaux: "allNetsDistinct (x # y # z) \<Longrightarrow> allNetsDistinct (x \<oplus> y # z)"
apply (simp add: allNetsDistinct_def)
apply (intro allI impI, rename_tac a b)
apply (drule_tac x = a in spec, drule_tac x = b in spec)
by (meson isInAlternativeList)
lemma netlistalternativeSeparateaux:
"net_list_aux [y] @ net_list_aux z = net_list_aux (y # z)"
by (case_tac y, simp_all)
@ -1561,14 +1454,11 @@ lemma netlistalternativeSeparateaux:
lemma netlistalternativeSeparate: "net_list_aux p = net_list_aux (separate p)"
by (induct p rule:separate.induct, simp_all) (simp_all add: netlistalternativeSeparateaux)
lemma ANDSepaux2:
"allNetsDistinct(x#y#z) \<Longrightarrow> allNetsDistinct(separate(y#z)) \<Longrightarrow> allNetsDistinct(x#separate(y#z))"
apply (simp add: allNetsDistinct_def)
by (metis isInAlternativeList netlistalternativeSeparate netlistaux)
lemma ANDSep[rule_format]: "allNetsDistinct p \<longrightarrow> allNetsDistinct(separate p)"
apply (induct p rule: separate.induct, simp_all)
apply (metis ANDConc aNDDA)
@ -1595,7 +1485,6 @@ proof -
list.simps(15) nDAeqSet noDA1eq)
qed
lemma OTNSC[rule_format]: "singleCombinators p \<longrightarrow> OnlyTwoNets p"
apply (induct p,simp_all)
apply (rename_tac a p)
@ -1616,32 +1505,27 @@ lemma fl3[rule_format]: "NetsCollected p \<longrightarrow> (first_bothNet x \<no
(\<forall>a\<in>set p. first_bothNet x \<noteq> first_bothNet a))\<longrightarrow> NetsCollected (x#p)"
by (induct p) simp_all
lemma sortedConc[rule_format]: " sorted (a # p) l \<longrightarrow> sorted p l"
by (induct p) simp_all
lemma smalleraux2:
"{a,b} \<in> set l \<Longrightarrow> {c,d} \<in> set l \<Longrightarrow> {a,b} \<noteq> {c,d} \<Longrightarrow>
smaller (DenyAllFromTo a b) (DenyAllFromTo c d) l \<Longrightarrow>
\<not> smaller (DenyAllFromTo c d) (DenyAllFromTo a b) l"
by (metis bothNet.simps(2) pos_noteq smaller.simps(5))
lemma smalleraux2a:
"{a,b} \<in> set l \<Longrightarrow> {c,d} \<in> set l \<Longrightarrow> {a,b} \<noteq> {c,d} \<Longrightarrow>
smaller (DenyAllFromTo a b) (AllowPortFromTo c d p) l \<Longrightarrow>
\<not> smaller (AllowPortFromTo c d p) (DenyAllFromTo a b) l"
by (simp) (metis eq_imp_le pos_noteq)
lemma smalleraux2b:
"{a,b} \<in> set l \<Longrightarrow> {c,d} \<in> set l \<Longrightarrow> {a,b} \<noteq> {c,d} \<Longrightarrow> y = DenyAllFromTo a b \<Longrightarrow>
smaller (AllowPortFromTo c d p) y l \<Longrightarrow>
\<not> smaller y (AllowPortFromTo c d p) l"
by (simp) (metis eq_imp_le pos_noteq)
lemma smalleraux2c:
"{a,b} \<in> set l\<Longrightarrow>{c,d}\<in>set l\<Longrightarrow>{a,b} \<noteq> {c,d} \<Longrightarrow> y = AllowPortFromTo a b q \<Longrightarrow>
smaller (AllowPortFromTo c d p) y l \<Longrightarrow> \<not> smaller y (AllowPortFromTo c d p) l"
@ -1684,8 +1568,6 @@ next
case (Conc c d) thus ?thesis using assms by simp
qed
thm Combinators.split
lemma smalleraux3a:
"a \<noteq> DenyAll \<Longrightarrow> b \<noteq> DenyAll \<Longrightarrow> in_list b l \<Longrightarrow> in_list a l \<Longrightarrow>
bothNet a \<noteq> bothNet b \<Longrightarrow> smaller a b l \<Longrightarrow> singleCombinators [a] \<Longrightarrow>
@ -1696,13 +1578,12 @@ lemma smalleraux3a:
done
lemma posaux[rule_format]: "position a l < position b l \<longrightarrow> a \<noteq> b"
by (induct l) simp_all
by (induct l, simp_all)
lemma posaux6[rule_format]:
"a \<in> set l \<longrightarrow> b \<in> set l \<longrightarrow> a \<noteq> b \<longrightarrow> position a l \<noteq> position b l"
by (induct l) (simp_all add: position_positive)
lemma notSmallerTransaux[rule_format]:
"x \<noteq> DenyAll \<Longrightarrow> r \<noteq> DenyAll \<Longrightarrow>
singleCombinators [x] \<Longrightarrow> singleCombinators [y] \<Longrightarrow> singleCombinators [r] \<Longrightarrow>
@ -1710,7 +1591,6 @@ lemma notSmallerTransaux[rule_format]:
in_list x l \<Longrightarrow> in_list y l \<Longrightarrow> in_list r l \<Longrightarrow> \<not> smaller r x l"
by (metis order_trans)
lemma notSmallerTrans[rule_format]:
"x \<noteq> DenyAll \<longrightarrow> r \<noteq> DenyAll \<longrightarrow> singleCombinators (x#y#z) \<longrightarrow>
\<not> smaller y x l \<longrightarrow> sorted (x#y#z) l \<longrightarrow> r \<in> set z \<longrightarrow>
@ -1800,9 +1680,6 @@ next
case (Conc c d) thus ?thesis by simp
qed
(* a terrible proof, but I didn't get it better;
complex context dependencies into huge case-distinction cascades. bu *)
lemma NCSaux2[rule_format]:
"noDenyAll p \<longrightarrow> {a, b} \<in> set l \<longrightarrow> all_in_list p l \<longrightarrow>singleCombinators p \<longrightarrow>
sorted (DenyAllFromTo a b # p) l \<longrightarrow> {a, b} \<noteq> firstList p \<longrightarrow>
@ -1859,8 +1736,6 @@ next
using * by force
qed
lemma NCSaux3[rule_format]:
"noDenyAll p \<longrightarrow> {a, b} \<in> set l \<longrightarrow> all_in_list p l \<longrightarrow>singleCombinators p \<longrightarrow>
sorted (AllowPortFromTo a b w # p) l \<longrightarrow> {a, b} \<noteq> firstList p \<longrightarrow>
@ -1891,7 +1766,6 @@ proof -
show "{a, b} \<noteq> {u, v}" by (insert * **, simp split: if_splits)
qed
lemma NCSaux4[rule_format]:
"noDenyAll p \<longrightarrow> {a, b} \<in> set l \<longrightarrow> all_in_list p l \<longrightarrow> singleCombinators p \<longrightarrow>
sorted (AllowPortFromTo a b c # p) l \<longrightarrow> {a, b} \<noteq> firstList p \<longrightarrow>
@ -1930,7 +1804,6 @@ proof -
show "{a, b} \<noteq> {u, v}" by (insert * **, simp_all split: if_splits)
qed
lemma NetsCollectedSorted[rule_format]:
"noDenyAll1 p \<longrightarrow> all_in_list p l \<longrightarrow> singleCombinators p \<longrightarrow> sorted p l \<longrightarrow> NetsCollected p"
apply (induct p)
@ -1980,37 +1853,28 @@ proof -
qed
qed
lemma NetsCollectedSort: "distinct p \<Longrightarrow>noDenyAll1 p \<Longrightarrow> all_in_list p l \<Longrightarrow>
singleCombinators p \<Longrightarrow> NetsCollected (sort p l)"
apply (rule_tac l = l in NetsCollectedSorted,rule noDAsort, simp_all)
apply (rule_tac b=p in all_in_listSubset)
by (auto intro: sort_is_sorted)
lemma fBNsep[rule_format]: "(\<forall>a\<in>set z. {b,c} \<noteq> first_bothNet a) \<longrightarrow>
(\<forall>a\<in>set (separate z). {b,c} \<noteq> first_bothNet a)"
apply (induct z rule: separate.induct, simp)
by (rule impI, simp)+
lemma fBNsep1[rule_format]: " (\<forall>a\<in>set z. first_bothNet x \<noteq> first_bothNet a) \<longrightarrow>
(\<forall>a\<in>set (separate z). first_bothNet x \<noteq> first_bothNet a)"
apply (induct z rule: separate.induct, simp)
by (rule impI, simp)+
lemma NetsCollectedSepauxa:
"{b,c}\<noteq>firstList z \<Longrightarrow> noDenyAll1 z \<Longrightarrow> \<forall>a\<in>set z. {b,c}\<noteq>first_bothNet a \<Longrightarrow> NetsCollected z \<Longrightarrow>
NetsCollected (separate z) \<Longrightarrow> {b, c} \<noteq> firstList (separate z) \<Longrightarrow> a \<in> set (separate z) \<Longrightarrow>
{b, c} \<noteq> first_bothNet a"
by (rule fBNsep) simp_all
lemma NetsCollectedSepaux:
"first_bothNet (x::('a,'b)Combinators) \<noteq> first_bothNet y \<Longrightarrow> \<not> member DenyAll y \<and> noDenyAll z \<Longrightarrow>
(\<forall>a\<in>set z. first_bothNet x \<noteq> first_bothNet a) \<and> NetsCollected (y # z) \<Longrightarrow>
@ -2044,7 +1908,6 @@ next
by (metis NetsCollectedSepaux fl3 noDA1eq noDenyAll.simps(1))
qed
lemma OTNaux:
"onlyTwoNets a \<Longrightarrow> \<not> member DenyAll a \<Longrightarrow> (x,y) \<in> sdnets a \<Longrightarrow>
(x = first_srcNet a \<and> y = first_destNet a) \<or> (x = first_destNet a \<and> y = first_srcNet a)"
@ -2078,8 +1941,6 @@ lemma sdnets_noteq:
apply (insert first_bothNet_charn [of aa])
by(metis OTNaux first_isIn insert_absorb2 insert_commute)
lemma fbn_noteq:
"onlyTwoNets a \<Longrightarrow> onlyTwoNets aa \<Longrightarrow> first_bothNet a \<noteq> first_bothNet aa \<Longrightarrow>
\<not> member DenyAll a \<Longrightarrow> \<not> member DenyAll aa \<Longrightarrow> allNetsDistinct [a, aa] \<Longrightarrow>
@ -2089,7 +1950,6 @@ lemma fbn_noteq:
apply (insert sdnets_charn [of aa])
by (metis first_bothNet_charn)
lemma NCisSD2aux:
assumes 1: "onlyTwoNets a" and 2 : "onlyTwoNets aa" and 3 : "first_bothNet a \<noteq> first_bothNet aa"
and 4: "\<not> member DenyAll a" and 5: "\<not> member DenyAll aa" and 6:" allNetsDistinct [a, aa]"
@ -2146,7 +2006,6 @@ proof -
apply (metis 4 5 firstInNeta firstInNet alternativelistconc2)
apply (metis 3 4 5 first_bothNetsd insert_commute)
apply (metis)
proof -
assume 14 : "(ab = first_srcNet a \<and> b = first_destNet a \<or> ab = first_destNet a \<and> b = first_srcNet a) \<and> (c, d) \<in> sdnets aa "
and 15 : "sdnets a = {(first_srcNet a, first_destNet a), (first_destNet a, first_srcNet a)} "
@ -2156,7 +2015,6 @@ proof -
and 19 : "first_destNet a \<noteq> first_srcNet a"
and 20 : "c = first_srcNet aa \<longrightarrow> d \<noteq> first_destNet aa"
show " netsDistinct ab c \<or> netsDistinct b d"
apply (case_tac "(ab = first_srcNet a \<and> b = first_destNet a)",simp_all)
apply (case_tac "c = first_srcNet aa", simp_all)
apply (metis 2 5 14 20 OTNaux)
@ -2268,14 +2126,11 @@ lemma ANDaux3[rule_format]:
"y \<in> set xs \<longrightarrow> a \<in> set (net_list_aux [y]) \<longrightarrow> a \<in> set (net_list_aux xs)"
by (induct xs) (simp_all add: isInAlternativeList)
lemma ANDaux2:
"allNetsDistinct (x # xs) \<Longrightarrow> y \<in> set xs \<Longrightarrow> allNetsDistinct [x,y]"
apply (simp add: allNetsDistinct_def)
by (meson ANDaux3 isInAlternativeList netlistaux)
lemma NCisSD2[rule_format]:
"\<not> member DenyAll a \<Longrightarrow> OnlyTwoNets (a#p) \<Longrightarrow>
NetsCollected2 (a # p) \<Longrightarrow> NetsCollected (a#p) \<Longrightarrow>
@ -2343,7 +2198,6 @@ next
qed
qed
lemma NC2Sep[rule_format]: "noDenyAll1 p \<longrightarrow> NetsCollected2 (separate p)"
proof (induct p rule: separate.induct, simp_all, goal_cases)
fix x :: "('a, 'b) Combinators list"
@ -2352,15 +2206,15 @@ proof (induct p rule: separate.induct, simp_all, goal_cases)
next
fix v va fix y::" ('a, 'b) Combinators" fix z
case 2 then show ?case
by (simp add: firstList.simps(1) fl2 noDA1eq noDenyAll.simps(1))
by (simp add: fl2 noDA1eq)
next
fix v va vb fix y::" ('a, 'b) Combinators" fix z
case 3 then show ?case
by (simp add: firstList.simps(1) fl2 noDA1eq noDenyAll.simps(1))
by (simp add: fl2 noDA1eq)
next
fix v va fix y::" ('a, 'b) Combinators" fix z
case 4 then show ?case
by (simp add: firstList.simps(1) fl2 noDA1eq noDenyAll.simps(1))
by (simp add: fl2 noDA1eq)
qed
lemma separatedSep[rule_format]:
@ -2372,12 +2226,6 @@ lemma separatedSep[rule_format]:
lemma rADnMT[rule_format]: "p \<noteq> [] \<longrightarrow> removeAllDuplicates p \<noteq> []"
by (induct p) simp_all
(* TODO: Prove this lemma:
lemma all2: "all_in_list (policy2list p) (Nets_List p)"
apply (induct "policy2list p" "Nets_List p" rule: all_in_list.induct)
apply simp_all
*)
lemma remDupsNMT[rule_format]: "p \<noteq> [] \<longrightarrow> remdups p \<noteq> []"
by (metis remdups_eq_nil_iff)
@ -2392,6 +2240,5 @@ lemma sets_distinct5: "(n::int) < m \<Longrightarrow> {(a,b). a = n} \<noteq> {(
lemma sets_distinct6: "(m::int) < n \<Longrightarrow> {(a,b). a = n} \<noteq> {(a,b). a = m}"
by auto
end

View File

@ -36,21 +36,23 @@
*****************************************************************************)
subsection {* Normalisation Proofs: Integer Protocol *}
theory NormalisationIPPProofs
imports NormalisationIntegerPortProof
theory
NormalisationIPPProofs
imports
NormalisationIntegerPortProof
begin
text{* Normalisation proofs which are specific to the IntegerProtocol address representation. *}
text{*
Normalisation proofs which are specific to the IntegerProtocol address representation.
*}
lemma ConcAssoc: "Cp((A \<oplus> B) \<oplus> D) = Cp(A \<oplus> (B \<oplus> D))"
by (simp add: Cp.simps)
lemma aux26[simp]:
"twoNetsDistinct a b c d \<Longrightarrow> dom (Cp (AllowPortFromTo a b p)) \<inter> dom (Cp (DenyAllFromTo c d)) = {}"
by(auto simp:twoNetsDistinct_def netsDistinct_def PLemmas, auto)
lemma wp2_aux[rule_format]:
"wellformed_policy2Pr (xs @ [x]) \<longrightarrow> wellformed_policy2Pr xs"
by (induct xs, simp_all) (case_tac "a", simp_all)
@ -67,8 +69,6 @@ apply (simp add: applied_rule_rev_def)
apply (rule_tac xs = p in rev_induct, simp_all)
by (metis Cp.simps(1) denyAllDom)
lemma DAimplieMR[rule_format]: "DenyAll \<in> set p \<Longrightarrow> applied_rule_rev Cp x p \<noteq> None"
by (auto intro: DAimpliesMR_E)
@ -94,7 +94,6 @@ lemma CConcStart[rule_format]:
"xs \<noteq> [] \<longrightarrow> Cp a x = None \<longrightarrow> Cp (list2FWpolicy (xs @ [a])) x = Cp (list2FWpolicy xs) x"
by (rule list2FWpolicy.induct) (simp_all add: PLemmas)
lemma mrNnt[simp]: "applied_rule_rev Cp x p = Some a \<Longrightarrow> p \<noteq> []"
by (simp add: applied_rule_rev_def)(auto)
@ -106,7 +105,6 @@ apply (metis CConcEnd )
apply (metis CConcEnd)
by (metis CConcStart applied_rule_rev_def mrNnt option.exhaust)
lemma CConcStart2:
"p \<noteq> [] \<Longrightarrow> x \<notin> dom (Cp a) \<Longrightarrow> Cp(list2FWpolicy (p@[a])) x = Cp (list2FWpolicy p)x"
by (erule CConcStart,simp add: PLemmas)
@ -119,12 +117,10 @@ lemma CConcEnd2[rule_format]:
"x \<in> dom (Cp a) \<longrightarrow> Cp (list2FWpolicy (xs @ [a])) x = Cp a x" (is "?P xs")
by (rule_tac P = ?P in list2FWpolicy.induct) (auto simp:Cp.simps)
lemma bar3:
"x \<in> dom (Cp (list2FWpolicy (xs @ [xa]))) \<Longrightarrow> x \<in> dom (Cp (list2FWpolicy xs)) \<or> x \<in> dom (Cp xa)"
by auto (metis CConcStart eq_Nil_appendI l2p_aux2 option.simps(3))
lemma CeqEnd[rule_format,simp]:
"a \<noteq> [] \<longrightarrow> x \<in> dom (Cp(list2FWpolicy a)) \<longrightarrow> Cp(list2FWpolicy(b@a)) x = (Cp(list2FWpolicy a)) x"
proof (induct rule: rev_induct)print_cases
@ -142,7 +138,6 @@ lemma CConcStartA[rule_format,simp]:
"x \<in> dom (Cp a) \<longrightarrow> x \<in> dom (Cp (list2FWpolicy (a # b)))" (is "?P b")
by (rule_tac P = ?P in list2FWpolicy.induct) (simp_all add: Cp.simps)
lemma domConc:
"x \<in> dom (Cp (list2FWpolicy b)) \<Longrightarrow> b \<noteq> [] \<Longrightarrow> x \<in> dom (Cp (list2FWpolicy (a@b)))"
by (auto simp: PLemmas)
@ -152,13 +147,11 @@ lemma CeqStart[rule_format,simp]:
Cp (list2FWpolicy (b@a)) x = (Cp (list2FWpolicy b)) x"
by (rule list2FWpolicy.induct,simp_all) (auto simp: list2FWpolicyconc PLemmas)
lemma C_eq_if_mr_eq2:
"applied_rule_rev Cp x a = Some r \<Longrightarrow> applied_rule_rev Cp x b = Some r \<Longrightarrow> a\<noteq>[] \<Longrightarrow> b\<noteq>[] \<Longrightarrow>
(Cp (list2FWpolicy a)) x = (Cp (list2FWpolicy b)) x"
by (metis mr_is_C)
lemma nMRtoNone[rule_format]:
"p \<noteq> [] \<longrightarrow> applied_rule_rev Cp x p = None \<longrightarrow> Cp (list2FWpolicy p) x = None"
proof (induct rule: rev_induct) print_cases
@ -169,7 +162,6 @@ next
by (simp_all add: snoc.hyps applied_rule_rev_def dom_def)
qed
lemma C_eq_if_mr_eq:
"applied_rule_rev Cp x b = applied_rule_rev Cp x a \<Longrightarrow> a \<noteq> [] \<Longrightarrow> b \<noteq> [] \<Longrightarrow>
(Cp (list2FWpolicy a)) x = (Cp (list2FWpolicy b)) x"
@ -178,7 +170,6 @@ apply (subst nMRtoNone,simp_all)
apply (subst nMRtoNone,simp_all)
by (auto intro: C_eq_if_mr_eq2)
lemma notmatching_notdom:
"applied_rule_rev Cp x (p@[a]) \<noteq> Some a \<Longrightarrow> x \<notin> dom (Cp a)"
by (simp add: applied_rule_rev_def split: if_splits)
@ -253,7 +244,6 @@ next
by simp_all (case_tac "r = xa", simp_all)
qed
lemma mrMTNone[simp]: "applied_rule_rev Cp x [] = None"
by (simp add: applied_rule_rev_def)
@ -264,7 +254,6 @@ lemma mrSet[rule_format]: "applied_rule_rev Cp x p = Some r \<longrightarrow> r
unfolding applied_rule_rev_def
by (rule_tac xs=p in rev_induct) simp_all
lemma mr_not_Conc: "singleCombinators p \<Longrightarrow> applied_rule_rev Cp x p \<noteq> Some (a\<oplus>b)"
by (auto simp: mrSet dest: mrSet elim: SCnotConc)
@ -280,8 +269,8 @@ lemma wp3EndMT[rule_format]:
"wellformed_policy3Pr (p@[xs]) \<longrightarrow> AllowPortFromTo a b po \<in> set p \<longrightarrow>
dom (Cp (AllowPortFromTo a b po)) \<inter> dom (Cp xs) = {}"
apply (induct p,simp_all)
by (metis NormalisationIPPProofs.wp3Conc aux0_4 inf_commute list.set_intros(1) wellformed_policy3Pr.simps(2))
by (metis NormalisationIPPProofs.wp3Conc aux0_4 inf_commute list.set_intros(1)
wellformed_policy3Pr.simps(2))
lemma foo29: "dom (Cp a) \<noteq> {} \<Longrightarrow> dom (Cp a) \<inter> dom (Cp b) = {} \<Longrightarrow> a \<noteq> b"
by auto
@ -311,12 +300,10 @@ lemma foo31:
(\<forall>r. r\<in>set s \<and> x\<in>dom(Cp r) \<longrightarrow> r=AllowPortFromTo a b po \<or> r=DenyAllFromTo a b \<or> r = DenyAll)"
by auto
lemma wp1_auxa: "wellformed_policy1_strong p\<Longrightarrow>(\<exists> r. applied_rule_rev Cp x p = Some r)"
apply (rule DAimpliesMR_E)
by (erule wp1_aux1aa)
lemma deny_dom[simp]:
"twoNetsDistinct a b c d \<Longrightarrow> dom (Cp (DenyAllFromTo a b)) \<inter> dom (Cp (DenyAllFromTo c d)) = {}"
by (simp add: Cp.simps) (erule aux6)
@ -332,7 +319,6 @@ apply (metis domComm aux26 tNDComm)
apply (simp add: twoNetsDistinct_def netsDistinct_def PLemmas)
by (auto simp: prod_eqI)
lemma DomInterAllowsMT_Ports:
"p \<noteq> po \<Longrightarrow> dom (Cp (AllowPortFromTo a b p)) \<inter> dom (Cp (AllowPortFromTo c d po)) = {}"
apply (simp add: twoNetsDistinct_def netsDistinct_def PLemmas)
@ -355,7 +341,6 @@ next
done
qed
lemma DistinctNetsDenyAllow:
"DenyAllFromTo b c \<in> set p \<Longrightarrow> AllowPortFromTo a d po \<in> set p\<Longrightarrow> allNetsDistinct p \<Longrightarrow>
dom (Cp (DenyAllFromTo b c)) \<inter> dom (Cp (AllowPortFromTo a d po)) \<noteq> {}\<Longrightarrow>
@ -380,8 +365,6 @@ apply (drule_tac x = "c" in spec)
apply (metis DomInterAllowsMT DomInterAllowsMT_Ports ND0aux3 ND0aux4 NDComm twoNetsDistinct_def)
done
lemma WP2RS2[simp]:
"singleCombinators p \<Longrightarrow> distinct p \<Longrightarrow> allNetsDistinct p \<Longrightarrow>
wellformed_policy2Pr (removeShadowRules2 p)"
@ -409,15 +392,12 @@ next
qed
qed
lemma AD_aux:
"AllowPortFromTo a b po \<in> set p \<Longrightarrow> DenyAllFromTo c d \<in> set p \<Longrightarrow>
allNetsDistinct p \<Longrightarrow> singleCombinators p \<Longrightarrow> a \<noteq> c \<or> b \<noteq> d \<Longrightarrow>
dom (Cp (AllowPortFromTo a b po)) \<inter> dom (Cp (DenyAllFromTo c d)) = {}"
by (rule aux26,rule_tac x ="AllowPortFromTo a b po" and y = "DenyAllFromTo c d" in tND) auto
lemma sorted_WP2[rule_format]:
"sorted p l \<longrightarrow> all_in_list p l \<longrightarrow> distinct p \<longrightarrow> allNetsDistinct p \<longrightarrow> singleCombinators p \<longrightarrow>
wellformed_policy2Pr p"
@ -450,23 +430,19 @@ next
qed
qed
lemma wellformed2_sorted[simp]:
"all_in_list p l \<Longrightarrow> distinct p \<Longrightarrow> allNetsDistinct p \<Longrightarrow> singleCombinators p \<Longrightarrow>
wellformed_policy2Pr (sort p l)"
by (metis distinct_sort set_sort sorted_WP2 SC3 aND_sort all_in_listSubset order_refl sort_is_sorted)
lemma wellformed2_sortedQ[simp]:
"all_in_list p l \<Longrightarrow> distinct p \<Longrightarrow> allNetsDistinct p \<Longrightarrow> singleCombinators p \<Longrightarrow>
wellformed_policy2Pr (qsort p l)"
by (metis sorted_WP2 SC3Q aND_sortQ all_in_listSubset distinct_sortQ set_sortQ sort_is_sortedQ subsetI)
lemma C_DenyAll[simp]: "Cp (list2FWpolicy (xs @ [DenyAll])) x = Some (deny ())"
by (auto simp: PLemmas)
lemma C_eq_RS1n:
"Cp(list2FWpolicy (removeShadowRules1_alternative p)) = Cp(list2FWpolicy p)"
proof (cases "p")
@ -493,17 +469,14 @@ next
qed
qed
lemma C_eq_RS1[simp]:
"p \<noteq> [] \<Longrightarrow> Cp(list2FWpolicy (removeShadowRules1 p)) = Cp(list2FWpolicy p)"
by (metis rSR1_eq C_eq_RS1n)
lemma EX_MR_aux[rule_format]:
"applied_rule_rev Cp x (DenyAll # p) \<noteq> Some DenyAll \<longrightarrow> (\<exists>y. applied_rule_rev Cp x p = Some y)"
by (simp add: applied_rule_rev_def) (rule_tac xs = p in rev_induct, simp_all)
lemma EX_MR :
"applied_rule_rev Cp x p \<noteq> (Some DenyAll) \<Longrightarrow> p = DenyAll#ps \<Longrightarrow>
(applied_rule_rev Cp x p = applied_rule_rev Cp x ps)"
@ -511,8 +484,6 @@ apply (auto,subgoal_tac "applied_rule_rev Cp x (DenyAll#ps) \<noteq> None", auto
apply (metis mrConcEnd)
by (metis DAimpliesMR_E list.sel(1) hd_in_set list.simps(3) not_Some_eq)
lemma mr_not_DA:
"wellformed_policy1_strong s \<Longrightarrow> applied_rule_rev Cp x p = Some (DenyAllFromTo a ab) \<Longrightarrow>
set p = set s \<Longrightarrow> applied_rule_rev Cp x s \<noteq> Some DenyAll"
@ -520,27 +491,22 @@ apply (subst wp1n_tl, simp_all)
by (metis (mono_tags, lifting) Combinators.distinct(1) foo98
mrSet mr_in_dom WP1n_DA_notinSet set_ConsD wp1n_tl)
lemma domsMT_notND_DD:
"dom (Cp (DenyAllFromTo a b)) \<inter> dom (Cp (DenyAllFromTo c d)) \<noteq> {} \<Longrightarrow> \<not> netsDistinct a c"
by (erule contrapos_nn) (simp add: Cp.simps aux6 twoNetsDistinct_def)
lemma domsMT_notND_DD2:
"dom (Cp (DenyAllFromTo a b)) \<inter> dom (Cp (DenyAllFromTo c d)) \<noteq> {} \<Longrightarrow> \<not> netsDistinct b d"
by (erule contrapos_nn) (simp add: Cp.simps aux6 twoNetsDistinct_def)
lemma domsMT_notND_DD3:
"x \<in> dom (Cp (DenyAllFromTo a b)) \<Longrightarrow> x \<in> dom (Cp (DenyAllFromTo c d)) \<Longrightarrow> \<not> netsDistinct a c"
by (auto intro!: domsMT_notND_DD)
lemma domsMT_notND_DD4:
"x \<in> dom (Cp (DenyAllFromTo a b)) \<Longrightarrow> x \<in> dom (Cp (DenyAllFromTo c d)) \<Longrightarrow> \<not> netsDistinct b d"
by (auto intro!: domsMT_notND_DD2)
lemma NetsEq_if_sameP_DD:
"allNetsDistinct p \<Longrightarrow> u\<in> set p \<Longrightarrow> v\<in> set p \<Longrightarrow> u = (DenyAllFromTo a b) \<Longrightarrow>
v = (DenyAllFromTo c d) \<Longrightarrow> x \<in> dom (Cp (u)) \<Longrightarrow> x \<in> dom (Cp (v)) \<Longrightarrow>
@ -548,7 +514,6 @@ lemma NetsEq_if_sameP_DD:
unfolding allNetsDistinct_def
by (simp)(metis allNetsDistinct_def ND0aux1 ND0aux2 domsMT_notND_DD3 domsMT_notND_DD4 )
lemma rule_charn1:
assumes aND : "allNetsDistinct p"
and mr_is_allow : "applied_rule_rev Cp x p = Some (AllowPortFromTo a b po)"
@ -575,12 +540,9 @@ lemma none_MT_rulessubset[rule_format]:
"none_MT_rules Cp a \<longrightarrow> set b \<subseteq> set a \<longrightarrow> none_MT_rules Cp b"
by (induct b,simp_all) (metis notMTnMT)
lemma nMTSort: "none_MT_rules Cp p \<Longrightarrow> none_MT_rules Cp (sort p l)"
by (metis set_sort nMTeqSet)
lemma nMTSortQ: "none_MT_rules Cp p \<Longrightarrow> none_MT_rules Cp (qsort p l)"
by (metis set_sortQ nMTeqSet)
@ -589,8 +551,6 @@ lemma wp3char[rule_format]: "none_MT_rules Cp xs \<and> Cp (AllowPortFromTo a b
AllowPortFromTo a b po \<notin> set xs"
by (induct xs, simp_all) (metis domNMT wp3Conc)
lemma wp3charn[rule_format]:
assumes domAllow: "dom (Cp (AllowPortFromTo a b po)) \<noteq> {}"
and wp3: "wellformed_policy3Pr (xs @ [DenyAllFromTo a b])"
@ -603,7 +563,6 @@ next
by (simp,auto intro: wp3Conc) (auto simp: DenyAllowDisj domAllow)
qed
lemma rule_charn2:
assumes aND: "allNetsDistinct p"
and wp1: "wellformed_policy1 p"
@ -649,14 +608,12 @@ lemma rule_charn2:
qed
qed
lemma rule_charn3:
"wellformed_policy1 p \<Longrightarrow> allNetsDistinct p \<Longrightarrow> singleCombinators p \<Longrightarrow>
wellformed_policy3Pr p \<Longrightarrow> applied_rule_rev Cp x p = Some (DenyAllFromTo c d) \<Longrightarrow>
AllowPortFromTo a b po \<in> set p \<Longrightarrow> x \<notin> dom (Cp (AllowPortFromTo a b po))"
by (clarify) (simp add: NormalisationIPPProofs.rule_charn2 domI)
lemma rule_charn4:
assumes wp1: "wellformed_policy1 p"
and aND: "allNetsDistinct p"
@ -687,8 +644,6 @@ next
by (metis Conc aux0_0)
qed
lemma foo31a:
"(\<forall> r. r \<in> set p \<and> x \<in> dom (Cp r) \<longrightarrow>
(r = AllowPortFromTo a b po \<or> r = DenyAllFromTo a b \<or> r = DenyAll)) \<Longrightarrow>
@ -696,12 +651,10 @@ lemma foo31a:
(r = AllowPortFromTo a b po \<or> r = DenyAllFromTo a b \<or> r = DenyAll)"
by auto
lemma aux4[rule_format]:
"applied_rule_rev Cp x (a#p) = Some a \<longrightarrow> a \<notin> set (p) \<longrightarrow> applied_rule_rev Cp x p = None"
by (rule rev_induct, simp_all) (intro impI,simp add: applied_rule_rev_def split: if_splits)
lemma mrDA_tl:
assumes mr_DA: "applied_rule_rev Cp x p = Some DenyAll"
and wp1n: "wellformed_policy1_strong p"
@ -721,7 +674,6 @@ apply (subgoal_tac "p = DenyAll#(tl p)")
singleCombinatorsConc waux2)
using wp1n_tl by auto
lemma mrDenyAll_is_unique:
"wellformed_policy1_strong p \<Longrightarrow> applied_rule_rev Cp x p = Some DenyAll \<Longrightarrow> r \<in> set (tl p) \<Longrightarrow>
x \<notin> dom (Cp r)"
@ -729,7 +681,6 @@ apply (rule_tac a = "[]" and b = "DenyAll" and c = "tl p" in foo3a, simp_all)
apply (metis wp1n_tl)
by (metis WP1n_DA_notinSet)
theorem C_eq_Sets_mr:
assumes sets_eq: "set p = set s"
and SC: "singleCombinators p"
@ -812,15 +763,12 @@ next
qed
qed
lemma C_eq_Sets:
"singleCombinators p \<Longrightarrow> wellformed_policy1_strong p \<Longrightarrow> wellformed_policy1_strong s \<Longrightarrow>
wellformed_policy3Pr p \<Longrightarrow> wellformed_policy3Pr s \<Longrightarrow> allNetsDistinct p \<Longrightarrow> set p = set s \<Longrightarrow>
Cp (list2FWpolicy p) x = Cp (list2FWpolicy s) x"
by (metis C_eq_Sets_mr C_eq_if_mr_eq wellformed_policy1_strong.simps(1))
lemma C_eq_sorted:
"distinct p \<Longrightarrow> all_in_list p l \<Longrightarrow> singleCombinators p \<Longrightarrow>
wellformed_policy1_strong p\<Longrightarrow> wellformed_policy3Pr p\<Longrightarrow> allNetsDistinct p \<Longrightarrow>
@ -829,7 +777,6 @@ by (rule ext)
(meson distinct_sort set_sort C_eq_Sets wellformed2_sorted wellformed_policy3_charn SC3 aND_sort
wellformed1_alternative_sorted wp1_eq)
lemma C_eq_sortedQ:
"distinct p \<Longrightarrow> all_in_list p l \<Longrightarrow> singleCombinators p \<Longrightarrow>
wellformed_policy1_strong p \<Longrightarrow> wellformed_policy3Pr p \<Longrightarrow> allNetsDistinct p \<Longrightarrow>
@ -838,9 +785,6 @@ by (rule ext)
(metis C_eq_Sets wellformed2_sortedQ wellformed_policy3_charn SC3Q aND_sortQ distinct_sortQ
set_sortQ wellformed1_sorted_auxQ wellformed_eq wp1_aux1aa)
lemma C_eq_RS2_mr: "applied_rule_rev Cp x (removeShadowRules2 p)= applied_rule_rev Cp x p"
proof (induct p)
case Nil thus ?case by simp
@ -876,8 +820,6 @@ next
qed
qed
lemma C_eq_None[rule_format]:
"p \<noteq> [] \<longrightarrow> applied_rule_rev Cp x p = None \<longrightarrow> Cp (list2FWpolicy p) x = None"
unfolding applied_rule_rev_def
@ -902,7 +844,6 @@ lemma C_eq_RS2:
apply (rule ext)
by (metis C_eq_RS2_mr C_eq_if_mr_eq RS2_NMT wp1_alternative_not_mt)
lemma none_MT_rulesRS2: "none_MT_rules Cp p \<Longrightarrow> none_MT_rules Cp (removeShadowRules2 p)"
by (auto simp: RS2Set none_MT_rulessubset)
@ -914,7 +855,6 @@ apply (metis Cdom2 list2FWpolicyconc)
apply (metis Cp.simps(4) map_add_dom_app_simps(2) inSet_not_MT list2FWpolicyconc set_empty2)
done
lemma none_MT_rulesrd[rule_format]: "none_MT_rules Cp p \<longrightarrow> none_MT_rules Cp (remdups p)"
by (induct p, simp_all)
@ -927,9 +867,6 @@ by (simp add: dom_def Cp.simps PolicyCombinators.PolicyCombinators)
lemma DAnMT2: "Cp DenyAll \<noteq> empty"
by (metis DAAux dom_eq_empty_conv empty_iff)
lemma wp1n_RS3[rule_format,simp]:
"wellformed_policy1_strong p \<longrightarrow> wellformed_policy1_strong (rm_MT_rules Cp p)"
apply (induct p, simp_all)
@ -946,16 +883,13 @@ lemma SCRS3[rule_format,simp]:
"singleCombinators p \<longrightarrow> singleCombinators(rm_MT_rules Cp p)"
by (induct p, simp_all) (case_tac "a", simp_all)
lemma RS3subset: "set (rm_MT_rules Cp p) \<subseteq> set p "
by (induct p, auto)
lemma ANDRS3[simp]:
"singleCombinators p \<Longrightarrow> allNetsDistinct p \<Longrightarrow> allNetsDistinct (rm_MT_rules Cp p)"
by (rule_tac b = p in aNDSubset, simp_all add:RS3subset)
lemma nlpaux: "x \<notin> dom (Cp b) \<Longrightarrow> Cp (a \<oplus> b) x = Cp a x"
by (metis Cp.simps(4) map_add_dom_app_simps(3))
@ -989,8 +923,6 @@ next
qed
qed
lemma nMT_domMT:
"\<not> not_MT Cp p \<Longrightarrow> p \<noteq> [] \<Longrightarrow> r \<notin> dom (Cp (list2FWpolicy p))"
proof (induct p)
@ -1039,30 +971,23 @@ next
qed
qed
lemma C_eq_id:
"wellformed_policy1_strong p \<Longrightarrow> Cp(list2FWpolicy (insertDeny p)) = Cp (list2FWpolicy p)"
by (rule ext) (metis insertDeny.simps(1) wp1n_tl)
lemma C_eq_RS3:
"not_MT Cp p \<Longrightarrow> Cp(list2FWpolicy (rm_MT_rules Cp p)) = Cp (list2FWpolicy p)"
by (rule ext) (erule C_eq_RS3_aux[symmetric])
lemma NMPrd[rule_format]: "not_MT Cp p \<longrightarrow> not_MT Cp (remdups p)"
by (induct p, simp_all) (auto simp: NMPcharn)
lemma NMPDA[rule_format]: "DenyAll \<in> set p \<longrightarrow> not_MT Cp p"
by (induct p, simp_all add: DAnMT)
lemma NMPiD[rule_format]: "not_MT Cp (insertDeny p)"
by (insert DAiniD [of p]) (erule NMPDA)
lemma list2FWpolicy2list[rule_format]:
"Cp (list2FWpolicy(policy2list p)) = (Cp p)"
apply (rule ext)
@ -1072,13 +997,11 @@ apply (metis Cdom2 CeqEnd domIff p2lNmt)
apply (metis CeqStart domIff p2lNmt nlpaux)
done
lemmas C_eq_Lemmas = none_MT_rulesRS2 none_MT_rulesrd SCp2l wp1n_RS2 wp1ID NMPiD waux2
wp1alternative_RS1 p2lNmt list2FWpolicy2list wellformed_policy3_charn wp1_eq
lemmas C_eq_subst_Lemmas = C_eq_sorted C_eq_sortedQ C_eq_RS2 C_eq_rd C_eq_RS3 C_eq_id
lemma C_eq_All_untilSorted:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
Cp(list2FWpolicy (sort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny
@ -1091,7 +1014,6 @@ apply (subst C_eq_RS3,simp_all add: C_eq_Lemmas)
apply (subst C_eq_id,simp_all add: C_eq_Lemmas)
done
lemma C_eq_All_untilSortedQ:
"DenyAll\<in> set(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
Cp(list2FWpolicy (qsort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny
@ -1104,8 +1026,6 @@ apply (subst C_eq_RS3,simp_all add: C_eq_Lemmas)
apply (subst C_eq_id,simp_all add: C_eq_Lemmas)
done
(* or, even shorter *)
lemma C_eq_All_untilSorted_withSimps:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> all_in_list (policy2list p) l \<Longrightarrow>
allNetsDistinct (policy2list p) \<Longrightarrow>
@ -1114,7 +1034,6 @@ lemma C_eq_All_untilSorted_withSimps:
Cp p"
by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas)
lemma C_eq_All_untilSorted_withSimpsQ:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> all_in_list (policy2list p) l \<Longrightarrow>
allNetsDistinct (policy2list p) \<Longrightarrow>
@ -1123,16 +1042,13 @@ lemma C_eq_All_untilSorted_withSimpsQ:
Cp p"
by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas)
lemma InDomConc[rule_format]: "p \<noteq> [] \<longrightarrow> x \<in> dom (Cp (list2FWpolicy (p))) \<longrightarrow>
x \<in> dom (Cp (list2FWpolicy (a#p)))"
by (induct p, simp_all) (case_tac "p = []",simp_all add: dom_def Cp.simps)
lemma not_in_member[rule_format]: "member a b \<longrightarrow> x \<notin> dom (Cp b) \<longrightarrow> x \<notin> dom (Cp a)"
by (induct b)(simp_all add: dom_def Cp.simps)
lemma src_in_sdnets[rule_format]:
"\<not> member DenyAll x \<longrightarrow> p \<in> dom (Cp x) \<longrightarrow> subnetsOfAdr (src p) \<inter> (fst_set (sdnets x)) \<noteq> {}"
apply (induct rule: Combinators.induct)
@ -1154,8 +1070,6 @@ apply (rule subnetAux)
apply (auto simp: PLemmas)
done
lemma sdnets_in_subnets[rule_format]:
"p\<in> dom (Cp x) \<longrightarrow> \<not> member DenyAll x \<longrightarrow>
(\<exists> (a,b)\<in>sdnets x. a \<in> subnetsOfAdr (src p) \<and> b \<in> subnetsOfAdr (dest p))"
@ -1176,8 +1090,6 @@ lemma list2FWpolicy_eq:
"zs \<noteq> [] \<Longrightarrow> Cp (list2FWpolicy (x \<oplus> y # z)) p = Cp (x \<oplus> list2FWpolicy (y # z)) p"
by (metis ConcAssoc l2p_aux list2FWpolicy.simps(2))
lemma dom_sep[rule_format]:
"x \<in> dom (Cp (list2FWpolicy p)) \<longrightarrow> x \<in> dom (Cp (list2FWpolicy(separate p)))"
proof (induct p rule: separate.induct,simp_all, goal_cases)
@ -1232,19 +1144,16 @@ lemma domdConcStart[rule_format]:
"x \<in> dom (Cp (list2FWpolicy (a#b))) \<longrightarrow> x \<notin> dom (Cp (list2FWpolicy b)) \<longrightarrow> x \<in> dom (Cp (a))"
by (induct b, simp_all) (auto simp: PLemmas)
lemma sep_dom2_aux:
"x \<in> dom (Cp (list2FWpolicy (a \<oplus> y # z))) \<Longrightarrow> x \<in> dom (Cp (a \<oplus> list2FWpolicy (y # z)))"
by auto (metis list2FWpolicy_eq p2lNmt)
lemma sep_dom2_aux2:
"(x \<in> dom (Cp (list2FWpolicy (separate (y # z)))) \<longrightarrow> x \<in> dom (Cp (list2FWpolicy (y # z)))) \<Longrightarrow>
x \<in> dom (Cp (list2FWpolicy (a # separate (y # z)))) \<Longrightarrow>
x \<in> dom (Cp (list2FWpolicy (a \<oplus> y # z)))"
by (metis CConcStartA InDomConc domdConcStart list.simps(2) list2FWpolicy.simps(2) list2FWpolicyconc)
lemma sep_dom2[rule_format]:
"x \<in> dom (Cp (list2FWpolicy (separate p))) \<longrightarrow> x \<in> dom (Cp (list2FWpolicy( p)))"
by (rule separate.induct) (simp_all add: sep_dom2_aux sep_dom2_aux2)
@ -1292,14 +1201,9 @@ qed
lemma C_eq_s: "p \<noteq> [] \<Longrightarrow> Cp (list2FWpolicy (separate p)) = Cp (list2FWpolicy p)"
by (rule ext) (simp add: C_eq_s_ext)
(*legacy *)
lemmas sortnMTQ = NormalisationIntegerPortProof.C_eq_Lemmas_sep(14)
lemmas C_eq_Lemmas_sep = C_eq_Lemmas sortnMT sortnMTQ RS2_NMT NMPrd not_MTimpnotMT
lemma C_eq_until_separated:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
Cp (list2FWpolicy (separate (sort (removeShadowRules2 (remdups (rm_MT_rules Cp
@ -1307,7 +1211,6 @@ lemma C_eq_until_separated:
Cp p"
by (simp add: C_eq_All_untilSorted_withSimps C_eq_s wellformed1_alternative_sorted wp1ID wp1n_RS2)
lemma C_eq_until_separatedQ:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> all_in_list (policy2list p) l \<Longrightarrow>
allNetsDistinct (policy2list p) \<Longrightarrow>
@ -1317,7 +1220,6 @@ lemma C_eq_until_separatedQ:
Cp p"
by (simp add: C_eq_All_untilSorted_withSimpsQ C_eq_s setnMT wp1ID wp1n_RS2)
lemma domID[rule_format]:
"p \<noteq> [] \<and> x \<in> dom(Cp(list2FWpolicy p)) \<longrightarrow> x \<in> dom (Cp(list2FWpolicy(insertDenies p)))"
proof(induct p)
@ -1375,13 +1277,11 @@ next
qed
qed
lemma DA_is_deny:
"x \<in> dom (Cp (DenyAllFromTo a b \<oplus> DenyAllFromTo b a \<oplus> DenyAllFromTo a b)) \<Longrightarrow>
Cp (DenyAllFromTo a b\<oplus>DenyAllFromTo b a \<oplus> DenyAllFromTo a b) x = Some (deny ())"
by (case_tac "x \<in> dom (Cp (DenyAllFromTo a b))") (simp_all add: PLemmas split: if_splits)
lemma iDdomAux[rule_format]:
"p \<noteq> [] \<longrightarrow> x \<notin> dom (Cp (list2FWpolicy p)) \<longrightarrow>
x \<in> dom (Cp (list2FWpolicy (insertDenies p))) \<longrightarrow>
@ -1472,14 +1372,13 @@ lemma iD_isD[rule_format]:
Cp (DenyAll \<oplus> list2FWpolicy (insertDenies p)) x = Cp DenyAll x"
apply (case_tac "x \<in> dom (Cp (list2FWpolicy (insertDenies p)))")
apply (simp add: Cp.simps(1) Cdom2 iDdomAux deny_all_def)
using NormalisationIPPProofs.nlpaux by blast
using NormalisationIPPProofs.nlpaux
by blast
lemma inDomConc:
"x\<notin>dom (Cp a) \<Longrightarrow> x\<notin>dom (Cp (list2FWpolicy p)) \<Longrightarrow> x \<notin> dom (Cp (list2FWpolicy(a#p)))"
by (metis domdConcStart)
lemma domsdisj[rule_format]:
"p \<noteq> [] \<longrightarrow> (\<forall> x s. s \<in> set p \<and> x \<in> dom (Cp A) \<longrightarrow> x \<notin> dom (Cp s)) \<longrightarrow> y \<in> dom (Cp A) \<longrightarrow>
y \<notin> dom (Cp (list2FWpolicy p))"
@ -1493,7 +1392,6 @@ next
by (metis Cons.hyps inDomConc list.set_intros(1) list.set_intros(2))
qed
lemma isSepaux:
"p \<noteq> [] \<Longrightarrow> noDenyAll (a#p) \<Longrightarrow> separated (a # p) \<Longrightarrow>
x \<in> dom (Cp (DenyAllFromTo (first_srcNet a) (first_destNet a) \<oplus>
@ -1506,13 +1404,12 @@ apply (rule_tac p = xa and x ="(DenyAllFromTo (first_srcNet a) (first_destNet a)
DenyAllFromTo (first_destNet a) (first_srcNet a) \<oplus> a)"
and y = s in disjSD_no_p_in_both, simp_all)
using disjSD2aux noDA apply blast
using noDA by blast
using noDA
by blast
lemma none_MT_rulessep[rule_format]: "none_MT_rules Cp p \<longrightarrow> none_MT_rules Cp (separate p)"
by (induct p rule: separate.induct) (simp_all add: Cp.simps map_add_le_mapE map_le_antisym)
lemma dom_id:
"noDenyAll (a#p) \<Longrightarrow> separated (a#p) \<Longrightarrow> p \<noteq> [] \<Longrightarrow>
x \<notin> dom (Cp (list2FWpolicy p)) \<Longrightarrow> x \<in> dom (Cp (a)) \<Longrightarrow>
@ -1523,7 +1420,6 @@ apply (rule_tac a = a in isSepaux, simp_all)
using id_aux4 noDA1eq sepNetsID apply blast
by (simp add: NormalisationIPPProofs.Cdom2 domIff)
lemma C_eq_iD_aux2[rule_format]:
"noDenyAll1 p \<longrightarrow> separated p\<longrightarrow> p \<noteq> []\<longrightarrow> x \<in> dom (Cp (list2FWpolicy p))\<longrightarrow>
Cp(list2FWpolicy (insertDenies p)) x = Cp(list2FWpolicy p) x"
@ -1611,7 +1507,6 @@ lemma C_eq_iD:
Cp(list2FWpolicy (insertDenies p)) = Cp (list2FWpolicy p)"
by (rule ext) (metis CConcStartA C_eq_iD_aux2 DAAux wp1_alternative_not_mt wp1n_tl)
lemma noDAsortQ[rule_format]: "noDenyAll1 p \<longrightarrow> noDenyAll1 (qsort p l)"
proof (cases "p")
case Nil then show ?thesis by simp
@ -1636,24 +1531,18 @@ next
qed
qed
(*MOVE FORWARD*)
lemma NetsCollectedSortQ:
"distinct p \<Longrightarrow>noDenyAll1 p \<Longrightarrow> all_in_list p l \<Longrightarrow> singleCombinators p \<Longrightarrow>
NetsCollected (qsort p l)"
by(metis C_eqLemmas_id(22))
lemmas CLemmas = nMTSort nMTSortQ none_MT_rulesRS2 none_MT_rulesrd
noDAsort noDAsortQ nDASC wp1_eq wp1ID SCp2l ANDSep wp1n_RS2
OTNSEp OTNSC noDA1sep wp1_alternativesep wellformed_eq
wellformed1_alternative_sorted
lemmas C_eqLemmas_id = CLemmas NC2Sep NetsCollectedSep
NetsCollectedSort NetsCollectedSortQ separatedNC
lemma C_eq_Until_InsertDenies:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct (policy2list p)\<Longrightarrow>
Cp (list2FWpolicy((insertDenies(separate(sort(removeShadowRules2
@ -1661,7 +1550,6 @@ lemma C_eq_Until_InsertDenies:
Cp p"
by (subst C_eq_iD,simp_all add: C_eqLemmas_id) (rule C_eq_until_separated, simp_all)
lemma C_eq_Until_InsertDeniesQ:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> all_in_list (policy2list p) l \<Longrightarrow>
allNetsDistinct (policy2list p) \<Longrightarrow>
@ -1674,13 +1562,11 @@ apply (subst C_eq_iD, simp_all add: C_eqLemmas_id)
apply (rule C_eq_until_separatedQ)
by simp_all
lemma C_eq_RD_aux[rule_format]: "Cp (p) x = Cp (removeDuplicates p) x"
apply (induct p, simp_all)
apply (intro conjI impI)
by (metis Cdom2 domIff nlpaux not_in_member) (metis Cp.simps(4) CConcStartaux Cdom2 domIff)
lemma C_eq_RAD_aux[rule_format]:
"p \<noteq> [] \<longrightarrow> Cp (list2FWpolicy p) x = Cp (list2FWpolicy (removeAllDuplicates p)) x"
proof (induct p)
@ -1703,7 +1589,6 @@ lemma C_eq_RAD:
"p \<noteq> [] \<Longrightarrow> Cp (list2FWpolicy p) = Cp (list2FWpolicy (removeAllDuplicates p)) "
by (rule ext) (erule C_eq_RAD_aux)
lemma C_eq_compile:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> all_in_list (policy2list p) l \<Longrightarrow>
allNetsDistinct (policy2list p) \<Longrightarrow>
@ -1712,7 +1597,6 @@ lemma C_eq_compile:
(removeShadowRules1 (policy2list p)))))) l))))) = Cp p"
by (metis C_eq_RAD C_eq_Until_InsertDenies removeAllDuplicates.simps(2))
lemma C_eq_compileQ:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
Cp (list2FWpolicy (removeAllDuplicates (insertDenies (separate (qsort
@ -1724,7 +1608,6 @@ apply (metis WP1rd sepnMT sortnMTQ wellformed_policy1_strong.simps(1) wp1ID wp1n
apply (rule C_eq_Until_InsertDeniesQ, simp_all)
done
lemma C_eq_normalizePr:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> allNetsDistinct (policy2list p) \<Longrightarrow>
all_in_list (policy2list p) (Nets_List p) \<Longrightarrow>
@ -1732,7 +1615,6 @@ lemma C_eq_normalizePr:
unfolding normalizePrQ_def
by (simp add: C_eq_compile normalizePr_def)
lemma C_eq_normalizePrQ:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> allNetsDistinct (policy2list p) \<Longrightarrow>
all_in_list (policy2list p) (Nets_List p) \<Longrightarrow>
@ -1740,68 +1622,55 @@ lemma C_eq_normalizePrQ:
unfolding normalizePrQ_def
using C_eq_compileQ by auto
lemma domSubset3: "dom (Cp (DenyAll \<oplus> x)) = dom (Cp (DenyAll))"
by (simp add: PLemmas split_tupled_all split: option.splits)
lemma domSubset4:
"dom (Cp (DenyAllFromTo x y \<oplus> DenyAllFromTo y x \<oplus> AllowPortFromTo x y dn)) =
dom (Cp (DenyAllFromTo x y \<oplus> DenyAllFromTo y x))"
by (simp add: PLemmas split: option.splits decision.splits) auto
lemma domSubset5:
"dom (Cp (DenyAllFromTo x y \<oplus> DenyAllFromTo y x \<oplus> AllowPortFromTo y x dn)) =
dom (Cp (DenyAllFromTo x y \<oplus> DenyAllFromTo y x))"
by (simp add: PLemmas split: option.splits decision.splits) auto
lemma domSubset1:
"dom (Cp (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> AllowPortFromTo one two dn \<oplus> x)) =
dom (Cp (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> x))"
by (simp add: PLemmas allow_all_def deny_all_def split: option.splits decision.splits) auto
lemma domSubset2:
"dom (Cp (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> AllowPortFromTo two one dn \<oplus> x)) =
dom (Cp (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> x))"
by (simp add: PLemmas allow_all_def deny_all_def split: option.splits decision.splits) auto
lemma ConcAssoc2: "Cp (X \<oplus> Y \<oplus> ((A \<oplus> B) \<oplus> D)) = Cp (X \<oplus> Y \<oplus> A \<oplus> B \<oplus> D)"
by (simp add: Cp.simps)
lemma ConcAssoc3: "Cp (X \<oplus> ((Y \<oplus> A) \<oplus> D)) = Cp (X \<oplus> Y \<oplus> A \<oplus> D)"
by (simp add: Cp.simps)
lemma RS3_NMT[rule_format]: "DenyAll \<in> set p \<longrightarrow>
rm_MT_rules Cp p \<noteq> []"
by (induct_tac p) (simp_all add: PLemmas)
lemma norm_notMT: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalizePr p \<noteq> []"
unfolding normalizePrQ_def
by (simp add: DAiniD RS3_NMT RS2_NMT idNMT normalizePr_def rADnMT sepnMT sortnMT)
lemma norm_notMTQ: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalizePrQ p \<noteq> []"
unfolding normalizePrQ_def
by (simp add: DAiniD RS3_NMT sortnMTQ RS2_NMT idNMT rADnMT sepnMT)
lemma domDA: "dom (Cp (DenyAll \<oplus> A)) = dom (Cp (DenyAll))"
by (rule domSubset3)
lemmas domain_reasoningPr = domDA ConcAssoc2 domSubset1 domSubset2
domSubset3 domSubset4 domSubset5 domSubsetDistr1
domSubsetDistr2 domSubsetDistrA domSubsetDistrD coerc_assoc ConcAssoc
ConcAssoc3
text {* The following lemmas help with the normalisation *}
lemma list2policyR_Start[rule_format]: "p \<in> dom (Cp a) \<longrightarrow>
Cp (list2policyR (a # list)) p = Cp a p"
@ -1813,7 +1682,6 @@ lemma list2policyR_End: "p \<notin> dom (Cp a) \<Longrightarrow>
by (rule list2policyR.induct)
(simp_all add: Cp.simps dom_def map_add_def list2policy_def split: option.splits)
lemma l2polR_eq_el[rule_format]: "N \<noteq> [] \<longrightarrow>
Cp( list2policyR N) p = (list2policy (map Cp N)) p"
proof (induct N)
@ -1824,14 +1692,10 @@ next
by (metis list2policyR_End list2policyR_Start domStart list2policy_def)
qed
lemma l2polR_eq:
"N \<noteq> [] \<Longrightarrow> Cp( list2policyR N) = (list2policy (map Cp N))"
by (auto simp: list2policy_def l2polR_eq_el )
lemma list2FWpolicys_eq_el[rule_format]:
"Filter \<noteq> [] \<longrightarrow> Cp (list2policyR Filter) p = Cp (list2FWpolicy (rev Filter)) p"
apply (induct_tac Filter)
@ -1857,40 +1721,33 @@ lemma list2FWpolicys_eq:
Cp (list2policyR Filter) = Cp (list2FWpolicy (rev Filter))"
by (rule ext, erule list2FWpolicys_eq_el)
lemma list2FWpolicys_eq_sym:
"Filter \<noteq> [] \<Longrightarrow>
Cp (list2policyR (rev Filter)) = Cp (list2FWpolicy Filter)"
by (metis list2FWpolicys_eq rev_is_Nil_conv rev_rev_ident)
lemma p_eq[rule_format]: "p \<noteq> [] \<longrightarrow>
list2policy (map Cp (rev p)) = Cp (list2FWpolicy p)"
by (metis l2polR_eq list2FWpolicys_eq_sym rev.simps(1) rev_rev_ident)
lemma p_eq2[rule_format]: "normalizePr x \<noteq> [] \<longrightarrow>
Cp (list2FWpolicy (normalizePr x)) = Cp x \<longrightarrow>
list2policy (map Cp (rev (normalizePr x))) = Cp x"
by (simp add: p_eq)
lemma p_eq2Q[rule_format]: "normalizePrQ x \<noteq> [] \<longrightarrow>
Cp (list2FWpolicy (normalizePrQ x)) = Cp x \<longrightarrow>
list2policy (map Cp (rev (normalizePrQ x))) = Cp x"
by (simp add: p_eq)
lemma list2listNMT[rule_format]: "x \<noteq> [] \<longrightarrow>map sem x \<noteq> []"
by (case_tac x) (simp_all)
lemma Norm_Distr2:
"r o_f ((P \<Otimes>\<^sub>2 (list2policy Q)) o d) =
(list2policy ((P \<Otimes>\<^sub>L Q) (op \<Otimes>\<^sub>2) r d))"
by (rule ext, rule Norm_Distr_2)
lemma NATDistr:
"N \<noteq> [] \<Longrightarrow> F = Cp (list2policyR N) \<Longrightarrow>
((\<lambda> (x,y). x) o_f ((NAT \<Otimes>\<^sub>2 F) o (\<lambda> x. (x,x))) =
@ -1898,7 +1755,6 @@ lemma NATDistr:
(\<lambda> (x,y). x) (\<lambda> x. (x,x))))))"
by (simp add: l2polR_eq) (rule ext,rule Norm_Distr_2)
lemma C_eq_normalize_manual:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> allNetsDistinct (policy2list p) \<Longrightarrow>
all_in_list (policy2list p) l \<Longrightarrow>
@ -1906,18 +1762,15 @@ lemma C_eq_normalize_manual:
unfolding normalize_manual_orderPr_def
by(simp_all add:C_eq_compile)
lemma p_eq2_manualQ[rule_format]:
"normalize_manual_orderPrQ x l \<noteq> [] \<longrightarrow>
Cp (list2FWpolicy (normalize_manual_orderPrQ x l)) = Cp x \<longrightarrow>
list2policy (map Cp (rev (normalize_manual_orderPrQ x l))) = Cp x"
by (simp add: p_eq)
lemma norm_notMT_manualQ: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalize_manual_orderPrQ p l \<noteq> []"
by (simp add: DAiniD RS3_NMT sortnMTQ RS2_NMT idNMT normalize_manual_orderPrQ_def rADnMT sepnMT)
lemma C_eq_normalizePr_manualQ:
"DenyAll \<in> set (policy2list p) \<Longrightarrow>
allNetsDistinct (policy2list p) \<Longrightarrow>
@ -1925,21 +1778,18 @@ lemma C_eq_normalizePr_manualQ:
Cp (list2FWpolicy (normalize_manual_orderPrQ p l)) = Cp p"
by (simp add: normalize_manual_orderPrQ_def C_eq_compileQ)
lemma p_eq2_manual[rule_format]: "normalize_manual_orderPr x l \<noteq> [] \<longrightarrow>
Cp (list2FWpolicy (normalize_manual_orderPr x l)) = Cp x \<longrightarrow>
list2policy (map Cp (rev (normalize_manual_orderPr x l))) = Cp x"
by (simp add: p_eq)
lemma norm_notMT_manual: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalize_manual_orderPr p l \<noteq> []"
unfolding normalize_manual_orderPr_def
by (simp add: idNMT rADnMT wellformed1_alternative_sorted wp1ID wp1_alternativesep wp1n_RS2)
text{* As an example, how this theorems can be used for a concrete
normalisation instantiation. *}
text{*
As an example, how this theorems can be used for a concrete normalisation instantiation.
*}
lemma normalizePrNAT:
"DenyAll \<in> set (policy2list Filter) \<Longrightarrow>
@ -1949,11 +1799,7 @@ lemma normalizePrNAT:
list2policy ((NAT \<Otimes>\<^sub>L (map Cp (rev (normalizePr Filter)))) (op \<Otimes>\<^sub>2) (\<lambda> (x,y). x) (\<lambda> x. (x,x)))"
by (simp add: C_eq_normalizePr NATDistr list2FWpolicys_eq_sym norm_notMT)
lemma domSimpl[simp]: "dom (Cp (A \<oplus> DenyAll)) = dom (Cp (DenyAll))"
by (simp add: PLemmas)
end

View File

@ -36,26 +36,28 @@
*****************************************************************************)
subsection {* Normalisation Proofs: Integer Port *}
theory NormalisationIntegerPortProof
imports NormalisationGenericProofs
theory
NormalisationIntegerPortProof
imports
NormalisationGenericProofs
begin
text{* Normalisation proofs which are specific to the IntegerPort address representation. *}
text{*
Normalisation proofs which are specific to the IntegerPort address representation.
*}
lemma ConcAssoc: "C((A \<oplus> B) \<oplus> D) = C(A \<oplus> (B \<oplus> D))"
by (simp add: C.simps)
lemma aux26[simp]: "twoNetsDistinct a b c d \<Longrightarrow>
dom (C (AllowPortFromTo a b p)) \<inter> dom (C (DenyAllFromTo c d)) = {}"
by (auto simp: PLemmas twoNetsDistinct_def netsDistinct_def) auto
apply (auto simp: PLemmas twoNetsDistinct_def netsDistinct_def)[1]
by auto
lemma wp2_aux[rule_format]: "wellformed_policy2 (xs @ [x]) \<longrightarrow>
wellformed_policy2 xs"
by (induct xs, simp_all) (case_tac "a", simp_all)
lemma Cdom2: "x \<in> dom(C b) \<Longrightarrow> C (a \<oplus> b) x = (C b) x"
by (auto simp: C.simps)
@ -123,12 +125,10 @@ lemma CConcEnd2[rule_format]:
apply (rule_tac P = ?P in list2FWpolicy.induct)
by (auto simp:C.simps)
lemma bar3:
"x \<in> dom (C (list2FWpolicy (xs @ [xa]))) \<Longrightarrow> x \<in> dom (C (list2FWpolicy xs)) \<or> x \<in> dom (C xa)"
by auto (metis CConcStart eq_Nil_appendI l2p_aux2 option.simps(3))
lemma CeqEnd[rule_format,simp]:
"a \<noteq> [] \<longrightarrow> x \<in> dom (C (list2FWpolicy a)) \<longrightarrow> C (list2FWpolicy(b@a)) x = (C (list2FWpolicy a)) x"
apply (rule rev_induct,simp_all)
@ -145,7 +145,6 @@ apply (rule_tac P = ?P in list2FWpolicy.induct)
apply (simp_all add: C.simps)
done
lemma domConc:
"x \<in> dom (C (list2FWpolicy b)) \<Longrightarrow> b \<noteq> [] \<Longrightarrow> x \<in> dom (C (list2FWpolicy (a @ b)))"
by (auto simp: PLemmas)
@ -176,8 +175,6 @@ apply (cases "applied_rule_rev C x a = None", simp_all)
apply (subst nMRtoNone, simp_all)
by (auto intro: C_eq_if_mr_eq2)
lemma notmatching_notdom: "applied_rule_rev C x (p@[a]) \<noteq> Some a \<Longrightarrow> x \<notin> dom (C a)"
by (simp add: applied_rule_rev_def split: if_splits)
@ -189,7 +186,6 @@ apply (intro impI conjI, simp)
apply (rule_tac p = "a @ b # xs" in notmatching_notdom,simp_all)
by (metis MRList2 MRList3 append_Cons option.inject)
lemma foo3D:
"wellformed_policy1 p \<Longrightarrow> p = DenyAll # ps \<Longrightarrow>
applied_rule_rev C x p = \<lfloor>DenyAll\<rfloor> \<Longrightarrow> r \<in> set ps \<Longrightarrow> x \<notin> dom (C r)"
@ -301,15 +297,11 @@ lemma foo31:
\<forall>r. r \<in> set s \<and> x \<in> dom (C r) \<longrightarrow> r=AllowPortFromTo a b po \<or> r = DenyAllFromTo a b \<or> r = DenyAll"
by auto
lemma wp1_auxa:
"wellformed_policy1_strong p\<Longrightarrow>(\<exists> r. applied_rule_rev C x p = Some r)"
apply (rule DAimpliesMR_E)
by (erule wp1_aux1aa)
lemma deny_dom[simp]:
"twoNetsDistinct a b c d \<Longrightarrow> dom (C (DenyAllFromTo a b)) \<inter> dom (C (DenyAllFromTo c d)) = {}"
apply (simp add: C.simps)
@ -344,8 +336,6 @@ apply (metis DomInterAllowsMT aux7aa DomInterAllowsMT_Ports)
apply (metis aux0_0 )
done
lemma DistinctNetsDenyAllow:
"DenyAllFromTo b c \<in> set p \<Longrightarrow>
AllowPortFromTo a d po \<in> set p \<Longrightarrow>
@ -370,11 +360,10 @@ apply (frule_tac x = "b" in spec)
apply (drule_tac x = "d" in spec)
apply (drule_tac x = "a" in spec)
apply (drule_tac x = "c" in spec)
apply (simp,metis DomInterAllowsMT DomInterAllowsMT_Ports ND0aux3 ND0aux4 NDComm twoNetsDistinct_def)
apply (simp,metis DomInterAllowsMT DomInterAllowsMT_Ports ND0aux3 ND0aux4 NDComm
twoNetsDistinct_def)
done
lemma WP2RS2[simp]:
"singleCombinators p \<Longrightarrow> distinct p \<Longrightarrow> allNetsDistinct p \<Longrightarrow>
wellformed_policy2 (removeShadowRules2 p)"
@ -399,16 +388,12 @@ next
qed
qed
lemma AD_aux:
"AllowPortFromTo a b po \<in> set p \<Longrightarrow> DenyAllFromTo c d \<in> set p \<Longrightarrow>
allNetsDistinct p \<Longrightarrow> singleCombinators p \<Longrightarrow> a \<noteq> c \<or> b \<noteq> d \<Longrightarrow>
dom (C (AllowPortFromTo a b po)) \<inter> dom (C (DenyAllFromTo c d)) = {}"
by (rule aux26,rule_tac x ="AllowPortFromTo a b po" and y = "DenyAllFromTo c d" in tND, auto)
lemma sorted_WP2[rule_format]: "sorted p l \<longrightarrow> all_in_list p l \<longrightarrow> distinct p \<longrightarrow>
allNetsDistinct p \<longrightarrow> singleCombinators p \<longrightarrow> wellformed_policy2 p"
proof (induct p)
@ -442,9 +427,6 @@ next
qed
qed
lemma wellformed2_sorted[simp]:
"all_in_list p l \<Longrightarrow> distinct p \<Longrightarrow> allNetsDistinct p \<Longrightarrow>
singleCombinators p \<Longrightarrow> wellformed_policy2 (sort p l)"
@ -452,18 +434,15 @@ apply (rule sorted_WP2,erule sort_is_sorted, simp_all)
apply (auto elim: all_in_listSubset intro: SC3 singleCombinatorsConc sorted_insort)
done
lemma wellformed2_sortedQ[simp]: "\<lbrakk>all_in_list p l; distinct p; allNetsDistinct p;
singleCombinators p\<rbrakk> \<Longrightarrow> wellformed_policy2 (qsort p l)"
apply (rule sorted_WP2,erule sort_is_sortedQ, simp_all)
apply (auto elim: all_in_listSubset intro: SC3Q singleCombinatorsConc distinct_sortQ)
done
lemma C_DenyAll[simp]: "C (list2FWpolicy (xs @ [DenyAll])) x = Some (deny ())"
by (auto simp: PLemmas)
lemma C_eq_RS1n:
"C(list2FWpolicy (removeShadowRules1_alternative p)) = C(list2FWpolicy p)"
proof (cases "p")print_cases
@ -481,13 +460,10 @@ next
removeShadowRules1_alternative_rev.simps(2) rev.simps(2))
qed
lemma C_eq_RS1[simp]:
"p \<noteq> [] \<Longrightarrow> C(list2FWpolicy (removeShadowRules1 p)) = C(list2FWpolicy p)"
by (metis rSR1_eq C_eq_RS1n)
lemma EX_MR_aux[rule_format]:
"applied_rule_rev C x (DenyAll # p) \<noteq> Some DenyAll \<longrightarrow> (\<exists>y. applied_rule_rev C x p = Some y)"
apply (simp add: applied_rule_rev_def)
@ -502,8 +478,6 @@ apply auto
apply (metis mrConcEnd)
by (metis DAimpliesMR_E list.sel(1) hd_in_set list.simps(3) not_Some_eq)
lemma mr_not_DA:
"wellformed_policy1_strong s \<Longrightarrow>
applied_rule_rev C x p = \<lfloor>DenyAllFromTo a ab\<rfloor> \<Longrightarrow> set p = set s \<Longrightarrow>
@ -515,34 +489,28 @@ apply (subst wp1n_tl, simp_all)
using mrSet r_not_DA_in_tl apply blast
by (simp add: mr_in_dom)
lemma domsMT_notND_DD:
"dom (C (DenyAllFromTo a b)) \<inter> dom (C (DenyAllFromTo c d)) \<noteq> {} \<Longrightarrow> \<not> netsDistinct a c"
using deny_dom twoNetsDistinct_def by blast
lemma domsMT_notND_DD2:
"dom (C (DenyAllFromTo a b)) \<inter> dom (C (DenyAllFromTo c d)) \<noteq> {} \<Longrightarrow> \<not> netsDistinct b d"
using deny_dom twoNetsDistinct_def by blast
lemma domsMT_notND_DD3:
"x \<in> dom (C (DenyAllFromTo a b)) \<Longrightarrow> x \<in> dom (C (DenyAllFromTo c d)) \<Longrightarrow> \<not> netsDistinct a c"
by(auto intro!:domsMT_notND_DD)
lemma domsMT_notND_DD4:
"x \<in> dom (C (DenyAllFromTo a b)) \<Longrightarrow> x \<in> dom (C (DenyAllFromTo c d)) \<Longrightarrow> \<not> netsDistinct b d"
by(auto intro!:domsMT_notND_DD2)
lemma NetsEq_if_sameP_DD:
"allNetsDistinct p \<Longrightarrow> u \<in> set p \<Longrightarrow> v \<in> set p \<Longrightarrow> u = DenyAllFromTo a b \<Longrightarrow>
v = DenyAllFromTo c d \<Longrightarrow> x \<in> dom (C u) \<Longrightarrow> x \<in> dom (C v) \<Longrightarrow> a = c \<and> b = d"
apply (simp add: allNetsDistinct_def)
by (metis ND0aux1 ND0aux2 domsMT_notND_DD3 domsMT_notND_DD4 )
lemma rule_charn1:
assumes aND: "allNetsDistinct p"
and mr_is_allow: "applied_rule_rev C x p = Some (AllowPortFromTo a b po)"
@ -567,12 +535,9 @@ lemma none_MT_rulessubset[rule_format]:
"none_MT_rules C a \<longrightarrow> set b \<subseteq> set a \<longrightarrow> none_MT_rules C b"
by (induct b,simp_all) (metis notMTnMT)
lemma nMTSort: "none_MT_rules C p \<Longrightarrow> none_MT_rules C (sort p l)"
by (metis set_sort nMTeqSet)
lemma nMTSortQ: "none_MT_rules C p \<Longrightarrow> none_MT_rules C (qsort p l)"
by (metis set_sortQ nMTeqSet)
@ -582,8 +547,6 @@ lemma wp3char[rule_format]:
apply (induct xs,simp_all)
by (metis domNMT wp3Conc)
lemma wp3charn[rule_format]:
assumes domAllow: "dom (C (AllowPortFromTo a b po)) \<noteq> {}"
and wp3: "wellformed_policy3 (xs @ [DenyAllFromTo a b])"
@ -596,7 +559,6 @@ next
by (simp,auto intro: wp3Conc) (auto simp: DenyAllowDisj domAllow)
qed
lemma rule_charn2:
assumes aND: "allNetsDistinct p"
and wp1: "wellformed_policy1 p"
@ -677,21 +639,17 @@ next
by (metis Conc aux0_0)
qed
lemma foo31a:
"\<forall>r. r \<in> set p \<and> x \<in> dom (C r) \<longrightarrow> r=AllowPortFromTo a b po \<or> r=DenyAllFromTo a b \<or> r=DenyAll \<Longrightarrow>
set p = set s \<Longrightarrow> r \<in> set s \<Longrightarrow> x \<in> dom (C r) \<Longrightarrow>
r = AllowPortFromTo a b po \<or> r = DenyAllFromTo a b \<or> r = DenyAll"
by auto
lemma aux4[rule_format]:
"applied_rule_rev C x (a#p) = Some a \<longrightarrow> a \<notin> set (p) \<longrightarrow> applied_rule_rev C x p = None"
apply (rule rev_induct,simp_all)
by (metis aux0_4 empty_iff empty_set insert_iff list.simps(15) mrSet mreq_end3)
lemma mrDA_tl:
assumes mr_DA: "applied_rule_rev C x p = Some DenyAll"
and wp1n: "wellformed_policy1_strong p"
@ -708,7 +666,6 @@ apply (subgoal_tac "p = DenyAll#(tl p)")
apply (metis AND_tl Combinators.distinct(1) SC_tl list.sel(3) mrConcEnd rule_charn4 waux2 wellformed_policy1_charn wp1_aux1aa wp1_eq wp3tl)
using wp1n_tl by blast
lemma mrDenyAll_is_unique:
"\<lbrakk>wellformed_policy1_strong p; applied_rule_rev C x p = Some DenyAll;
r \<in> set (tl p)\<rbrakk> \<Longrightarrow> x \<notin> dom (C r)"
@ -716,7 +673,6 @@ apply (rule_tac a = "[]" and b = "DenyAll" and c = "tl p" in foo3a, simp_all)
apply (metis wp1n_tl)
by (metis WP1n_DA_notinSet)
theorem C_eq_Sets_mr:
assumes sets_eq: "set p = set s"
and SC: "singleCombinators p"
@ -801,7 +757,6 @@ lemma C_eq_Sets:
C (list2FWpolicy p) x = C (list2FWpolicy s) x"
by(auto intro: C_eq_if_mr_eq C_eq_Sets_mr [symmetric])
lemma C_eq_sorted:
"distinct p \<Longrightarrow> all_in_list p l \<Longrightarrow> singleCombinators p \<Longrightarrow> wellformed_policy1_strong p \<Longrightarrow>
wellformed_policy3 p \<Longrightarrow> allNetsDistinct p \<Longrightarrow>
@ -810,8 +765,6 @@ apply (rule ext)
by (auto intro: C_eq_Sets simp: nMTSort wellformed1_alternative_sorted
wellformed_policy3_charn wp1_eq)
lemma C_eq_sortedQ:
"distinct p \<Longrightarrow> all_in_list p l \<Longrightarrow> singleCombinators p \<Longrightarrow> wellformed_policy1_strong p \<Longrightarrow>
wellformed_policy3 p \<Longrightarrow> allNetsDistinct p \<Longrightarrow>
@ -821,8 +774,6 @@ apply (auto intro!: C_eq_Sets simp: nMTSortQ wellformed1_alternative_sorted dist
wellformed_policy3_charn wp1_eq)
by (metis set_qsort wellformed1_sortedQ wellformed_eq wp1_aux1aa)
lemma C_eq_RS2_mr: "applied_rule_rev C x (removeShadowRules2 p)= applied_rule_rev C x p"
proof (induct p)
case Nil thus ?case by simp
@ -859,8 +810,6 @@ next
qed
qed
lemma C_eq_None[rule_format]:
"p \<noteq> [] --> applied_rule_rev C x p = None \<longrightarrow> C (list2FWpolicy p) x = None"
apply (simp add: applied_rule_rev_def)
@ -880,7 +829,6 @@ lemma C_eq_RS2:
apply (rule ext)
by (metis C_eq_RS2_mr C_eq_if_mr_eq wellformed_policy1_strong.simps(1) wp1n_RS2)
lemma none_MT_rulesRS2:
"none_MT_rules C p \<Longrightarrow> none_MT_rules C (removeShadowRules2 p)"
by (auto simp: RS2Set none_MT_rulessubset)
@ -893,7 +841,6 @@ apply (metis Cdom2 list2FWpolicyconc)
apply (metis C.simps(4) map_add_dom_app_simps(2) inSet_not_MT list2FWpolicyconc set_empty2)
done
lemma none_MT_rulesrd[rule_format]:
"none_MT_rules C p \<longrightarrow> none_MT_rules C (remdups p)"
by (induct p, simp_all)
@ -908,9 +855,6 @@ by (simp add: dom_def C.simps PolicyCombinators.PolicyCombinators)
lemma DAnMT2: "C DenyAll \<noteq> empty"
by (metis DAAux dom_eq_empty_conv empty_iff)
lemma wp1n_RS3[rule_format,simp]:
"wellformed_policy1_strong p \<longrightarrow> wellformed_policy1_strong (rm_MT_rules C p)"
by (induct p, simp_all add: DARS3 DAnMT)
@ -926,7 +870,6 @@ by (induct p, simp_all, case_tac "a", simp_all)
lemma RS3subset: "set (rm_MT_rules C p) \<subseteq> set p "
by (induct p, auto)
lemma ANDRS3[simp]:
"singleCombinators p \<Longrightarrow> allNetsDistinct p \<Longrightarrow> allNetsDistinct (rm_MT_rules C p)"
using RS3subset SCRS3 aNDSubset by blast
@ -965,8 +908,6 @@ next
qed
qed
lemma nMT_domMT:
"\<not> not_MT C p \<Longrightarrow> p \<noteq> [] \<Longrightarrow> r \<notin> dom (C (list2FWpolicy p))"
proof (induct p)
@ -1010,8 +951,6 @@ next
qed
qed
lemma C_eq_id:
"wellformed_policy1_strong p \<Longrightarrow> C(list2FWpolicy (insertDeny p)) = C (list2FWpolicy p)"
by (rule ext) (auto intro: C_eq_if_mr_eq elim: mr_iD)
@ -1020,25 +959,20 @@ lemma C_eq_RS3:
"not_MT C p \<Longrightarrow> C(list2FWpolicy (rm_MT_rules C p)) = C (list2FWpolicy p)"
by (rule ext) (erule C_eq_RS3_aux[symmetric])
lemma NMPrd[rule_format]: "not_MT C p \<longrightarrow> not_MT C (remdups p)"
by (induct p) (auto simp: NMPcharn)
lemma NMPDA[rule_format]: "DenyAll \<in> set p \<longrightarrow> not_MT C p"
by (induct p, simp_all add: DAnMT)
lemma NMPiD[rule_format]: "not_MT C (insertDeny p)"
by (simp add: DAiniD NMPDA)
lemma list2FWpolicy2list[rule_format]: "C (list2FWpolicy(policy2list p)) = (C p)"
apply (rule ext)
apply (induct_tac p, simp_all)
by (metis (no_types, lifting) Cdom2 CeqEnd CeqStart domIff nlpaux p2lNmt)
lemmas C_eq_Lemmas = none_MT_rulesRS2 none_MT_rulesrd SCp2l wp1n_RS2 wp1ID NMPiD wp1_eq
wp1alternative_RS1 p2lNmt list2FWpolicy2list wellformed_policy3_charn waux2
@ -1058,7 +992,6 @@ apply (subst C_eq_RS3,simp_all add: C_eq_Lemmas)
apply (subst C_eq_id,simp_all add: C_eq_Lemmas)
done
lemma C_eq_All_untilSortedQ:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
C (list2FWpolicy
@ -1072,9 +1005,6 @@ apply (subst C_eq_RS3,simp_all add: C_eq_Lemmas)
apply (subst C_eq_id,simp_all add: C_eq_Lemmas)
done
(* or, even shorter *)
lemma C_eq_All_untilSorted_withSimps:
"DenyAll\<in>set(policy2list p) \<Longrightarrow>all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct (policy2list p) \<Longrightarrow>
C (list2FWpolicy
@ -1084,7 +1014,6 @@ lemma C_eq_All_untilSorted_withSimps:
C p"
by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas)
lemma C_eq_All_untilSorted_withSimpsQ:
" DenyAll\<in>set(policy2list p)\<Longrightarrow>all_in_list(policy2list p) l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
C (list2FWpolicy
@ -1098,11 +1027,9 @@ lemma InDomConc[rule_format]:
"p \<noteq> [] \<longrightarrow> x \<in> dom (C (list2FWpolicy (p))) \<longrightarrow> x \<in> dom (C (list2FWpolicy (a#p)))"
by (induct p, simp_all) (case_tac "p = []", simp_all add: dom_def C.simps)
lemma not_in_member[rule_format]: "member a b \<longrightarrow> x \<notin> dom (C b) \<longrightarrow> x \<notin> dom (C a)"
by (induct b) (simp_all add: dom_def C.simps)
lemma src_in_sdnets[rule_format]:
"\<not> member DenyAll x \<longrightarrow> p \<in> dom (C x) \<longrightarrow> subnetsOfAdr (src p) \<inter> (fst_set (sdnets x)) \<noteq> {}"
apply (induct rule: Combinators.induct)
@ -1124,8 +1051,6 @@ apply (rule subnetAux)
apply (auto simp: PLemmas)
done
lemma sdnets_in_subnets[rule_format]:
"p\<in> dom (C x) \<longrightarrow> \<not> member DenyAll x \<longrightarrow>
(\<exists> (a,b)\<in>sdnets x. a \<in> subnetsOfAdr (src p) \<and> b \<in> subnetsOfAdr (dest p))"
@ -1146,7 +1071,6 @@ lemma list2FWpolicy_eq:
"zs \<noteq> [] \<Longrightarrow> C (list2FWpolicy (x \<oplus> y # z)) p = C (x \<oplus> list2FWpolicy (y # z)) p"
by (metis ConcAssoc l2p_aux list2FWpolicy.simps(2))
lemma dom_sep[rule_format]:
"x \<in> dom (C (list2FWpolicy p)) \<longrightarrow> x \<in> dom (C (list2FWpolicy(separate p)))"
proof (induct p rule: separate.induct, simp_all, goal_cases)
@ -1199,13 +1123,11 @@ lemma sep_dom2_aux:
" x \<in> dom (C (list2FWpolicy (a \<oplus> y # z))) \<Longrightarrow> x \<in> dom (C (a \<oplus> list2FWpolicy (y # z)))"
by (auto) (metis list2FWpolicy_eq p2lNmt)
lemma sep_dom2_aux2:
"x \<in> dom (C (list2FWpolicy (separate (y # z)))) \<longrightarrow> x \<in> dom (C (list2FWpolicy (y # z))) \<Longrightarrow>
x \<in> dom (C (list2FWpolicy (a # separate (y # z)))) \<Longrightarrow> x \<in> dom (C (list2FWpolicy (a \<oplus> y # z)))"
by (metis CConcStartA InDomConc domdConcStart list.simps(2) list2FWpolicy.simps(2) list2FWpolicyconc)
lemma sep_dom2[rule_format]:
"x \<in> dom (C (list2FWpolicy (separate p))) \<longrightarrow> x \<in> dom (C (list2FWpolicy( p)))"
by (rule separate.induct) (simp_all add: sep_dom2_aux sep_dom2_aux2)
@ -1281,21 +1203,16 @@ next
case 8 thus ?case by simp
qed
lemma C_eq_s:
"p \<noteq> [] \<Longrightarrow> C (list2FWpolicy (separate p)) = C (list2FWpolicy p)"
apply (rule ext) using C_eq_s_ext by blast
(*MOVE FORWARD*)
lemma sortnMTQ: "p \<noteq> [] \<Longrightarrow> qsort p l \<noteq> []"
by (metis set_sortQ setnMT)
lemmas C_eq_Lemmas_sep =
C_eq_Lemmas sortnMT sortnMTQ RS2_NMT NMPrd not_MTimpnotMT
lemma C_eq_until_separated:
" DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p)l \<Longrightarrow> allNetsDistinct (policy2list p) \<Longrightarrow>
C (list2FWpolicy
@ -1306,19 +1223,14 @@ lemma C_eq_until_separated:
C p"
by (simp add: C_eq_All_untilSorted_withSimps C_eq_s wellformed1_alternative_sorted wp1ID wp1n_RS2)
lemma C_eq_until_separatedQ:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p)l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
C (list2FWpolicy
(separate (qsort (removeShadowRules2 (remdups (rm_MT_rules C
(insertDeny (removeShadowRules1 (policy2list p)))))) l))) =
C p"
by (simp add: C_eq_All_untilSorted_withSimpsQ C_eq_s sortnMTQ wp1ID wp1n_RS2)
lemma domID[rule_format]: "p \<noteq> [] \<and> x \<in> dom(C(list2FWpolicy p)) \<longrightarrow>
x \<in> dom (C(list2FWpolicy(insertDenies p)))"
proof(induct p)
@ -1376,7 +1288,6 @@ next
qed
qed
lemma DA_is_deny:
"x \<in> dom (C (DenyAllFromTo a b \<oplus> DenyAllFromTo b a \<oplus> DenyAllFromTo a b)) \<Longrightarrow>
C (DenyAllFromTo a b\<oplus>DenyAllFromTo b a \<oplus> DenyAllFromTo a b) x = Some (deny ())"
@ -1520,7 +1431,6 @@ apply simp_all
by (metis Combinators.distinct(1) FWNormalisationCore.member.simps(1)
FWNormalisationCore.member.simps(3) disjSD2aux disjSD_no_p_in_both noDA)
lemma none_MT_rulessep[rule_format]: "none_MT_rules C p \<longrightarrow> none_MT_rules C (separate p)"
apply(induct p rule: separate.induct)
by (simp_all add: C.simps map_add_le_mapE map_le_antisym)
@ -1534,7 +1444,6 @@ using noDAID apply blast
using id_aux4 noDA1eq sepNetsID apply blast
by (metis list.set_intros(1) list.set_intros(2) list2FWpolicy.simps(2) list2FWpolicy.simps(3) notindom)
lemma C_eq_iD_aux2[rule_format]:
"noDenyAll1 p \<longrightarrow> separated p\<longrightarrow> p \<noteq> []\<longrightarrow> x \<in> dom (C (list2FWpolicy p))\<longrightarrow>
C(list2FWpolicy (insertDenies p)) x = C(list2FWpolicy p) x"
@ -1592,8 +1501,6 @@ lemma C_eq_iD:
C (list2FWpolicy (insertDenies p)) = C (list2FWpolicy p)"
by (rule ext) (metis CConcStartA C_eq_iD_aux2 DAAux wp1_alternative_not_mt wp1n_tl)
(*MOVE FORWARD*)
lemma noDAsortQ[rule_format]: "noDenyAll1 p \<longrightarrow> noDenyAll1 (qsort p l)"
apply (case_tac "p",simp_all, rename_tac a list)
apply (case_tac "a = DenyAll",simp_all)
@ -1603,9 +1510,6 @@ apply (rule impI,rule noDA1eq)
apply (metis append_Cons append_Nil nDAeqSet qsort.simps(2) set_sortQ)
by (case_tac a, simp_all)
(*MOVE FORWARD*)
lemma NetsCollectedSortQ:
"distinct p \<Longrightarrow>noDenyAll1 p \<Longrightarrow> all_in_list p l \<Longrightarrow> singleCombinators p \<Longrightarrow>
NetsCollected (qsort p l)"
@ -1624,8 +1528,6 @@ lemmas CLemmas = nMTSort nMTSortQ none_MT_rulesRS2 none_MT_rulesrd
lemmas C_eqLemmas_id = CLemmas NC2Sep NetsCollectedSep
NetsCollectedSort NetsCollectedSortQ separatedNC
lemma C_eq_Until_InsertDenies:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p)l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
C (list2FWpolicy
@ -1650,12 +1552,10 @@ apply (subst C_eq_iD,simp_all add: C_eqLemmas_id)
apply (metis WP1rd set_qsort wellformed1_sortedQ wellformed_eq wp1ID wp1_alternativesep wp1_aux1aa wp1n_RS2 wp1n_RS3)
by (rule C_eq_until_separatedQ, simp_all)
lemma C_eq_RD_aux[rule_format]: "C (p) x = C (removeDuplicates p) x"
apply (induct p,simp_all)
by (metis Cdom2 domIff nlpaux not_in_member)
lemma C_eq_RAD_aux[rule_format]:
"p \<noteq> [] \<longrightarrow> C (list2FWpolicy p) x = C (list2FWpolicy (removeAllDuplicates p)) x"
proof (induct p)
@ -1670,12 +1570,9 @@ next
by (metis C_eq_RD_aux Cons.hyps domIff list2FWpolicyconc nlpaux rADnMT)
qed
lemma C_eq_RAD:
"p \<noteq> [] \<Longrightarrow> C (list2FWpolicy p) = C (list2FWpolicy (removeAllDuplicates p)) "
by (rule ext,erule C_eq_RAD_aux)
lemma C_eq_compile:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p)l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
C (list2FWpolicy
@ -1690,7 +1587,6 @@ apply (subst C_eq_RAD[symmetric])
apply (rule idNMT,simp add: C_eqLemmas_id)
by (rule C_eq_Until_InsertDenies, simp_all)
lemma C_eq_compileQ:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> all_in_list(policy2list p)l \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow>
C (list2FWpolicy
@ -1704,7 +1600,6 @@ apply (subst C_eq_RAD[symmetric],rule idNMT)
apply (metis WP1rd sepnMT sortnMTQ wellformed_policy1_strong.simps(1) wp1ID wp1n_RS2 wp1n_RS3)
by (rule C_eq_Until_InsertDeniesQ, simp_all)
lemma C_eq_normalize:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> allNetsDistinct (policy2list p) \<Longrightarrow>
all_in_list(policy2list p)(Nets_List p) \<Longrightarrow>
@ -1712,85 +1607,68 @@ lemma C_eq_normalize:
unfolding normalize_def
by (simp add: C_eq_compile)
lemma C_eq_normalizeQ:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> allNetsDistinct (policy2list p) \<Longrightarrow>
all_in_list (policy2list p) (Nets_List p) \<Longrightarrow>
C (list2FWpolicy (normalizeQ p)) = C p"
by (simp add: normalizeQ_def C_eq_compileQ)
lemma domSubset3: "dom (C (DenyAll \<oplus> x)) = dom (C (DenyAll))"
by (simp add: PLemmas split_tupled_all split: option.splits)
lemma domSubset4:
"dom (C (DenyAllFromTo x y \<oplus> DenyAllFromTo y x \<oplus> AllowPortFromTo x y dn)) =
dom (C (DenyAllFromTo x y \<oplus> DenyAllFromTo y x))"
by (auto simp: PLemmas split: option.splits decision.splits )
lemma domSubset5:
"dom (C (DenyAllFromTo x y \<oplus> DenyAllFromTo y x \<oplus> AllowPortFromTo y x dn)) =
dom (C (DenyAllFromTo x y \<oplus> DenyAllFromTo y x))"
by (auto simp: PLemmas split: option.splits decision.splits )
lemma domSubset1:
"dom (C (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> AllowPortFromTo one two dn \<oplus> x)) =
dom (C (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> x))"
by (simp add: PLemmas split: option.splits decision.splits) (auto simp: allow_all_def deny_all_def)
lemma domSubset2:
"dom (C (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> AllowPortFromTo two one dn \<oplus> x)) =
dom (C (DenyAllFromTo one two \<oplus> DenyAllFromTo two one \<oplus> x))"
by (simp add: PLemmas split: option.splits decision.splits) (auto simp: allow_all_def deny_all_def)
lemma ConcAssoc2: "C (X \<oplus> Y \<oplus> ((A \<oplus> B) \<oplus> D)) = C (X \<oplus> Y \<oplus> A \<oplus> B \<oplus> D)"
by (simp add: C.simps)
lemma ConcAssoc3: "C (X \<oplus> ((Y \<oplus> A) \<oplus> D)) = C (X \<oplus> Y \<oplus> A \<oplus> D)"
by (simp add: C.simps)
lemma RS3_NMT[rule_format]:
"DenyAll \<in> set p \<longrightarrow> rm_MT_rules C p \<noteq> []"
by (induct_tac p) (simp_all add: PLemmas)
lemma norm_notMT: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalize p \<noteq> []"
by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalize_def rADnMT sepnMT sortnMT)
lemma norm_notMTQ: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalizeQ p \<noteq> []"
by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalizeQ_def rADnMT sepnMT sortnMTQ)
lemmas domDA = NormalisationIntegerPortProof.domSubset3 (* legacy *)
lemmas domain_reasoning = domDA ConcAssoc2 domSubset1 domSubset2
domSubset3 domSubset4 domSubset5 domSubsetDistr1
domSubsetDistr2 domSubsetDistrA domSubsetDistrD coerc_assoc ConcAssoc
ConcAssoc3
text {* The following lemmas help with the normalisation *}
lemma list2policyR_Start[rule_format]: "p \<in> dom (C a) \<longrightarrow>
C (list2policyR (a # list)) p = C a p"
by (induct "a # list" rule:list2policyR.induct) (auto simp: C.simps dom_def map_add_def)
lemma list2policyR_End: "p \<notin> dom (C a) \<Longrightarrow>
C (list2policyR (a # list)) p = (C a \<Oplus> list2policy (map C list)) p"
by (rule list2policyR.induct)
(simp_all add: C.simps dom_def map_add_def list2policy_def split: option.splits)
lemma l2polR_eq_el[rule_format]:
"N \<noteq> [] \<longrightarrow> C(list2policyR N) p = (list2policy (map C N)) p"
proof (induct N)
@ -1805,12 +1683,10 @@ next
done
qed
lemma l2polR_eq:
"N \<noteq> [] \<Longrightarrow> C( list2policyR N) = (list2policy (map C N))"
by (auto simp: list2policy_def l2polR_eq_el )
lemma list2FWpolicys_eq_el[rule_format]:
"Filter \<noteq> [] \<longrightarrow> C (list2policyR Filter) p = C (list2FWpolicy (rev Filter)) p"
proof (induct Filter) print_cases
@ -1828,38 +1704,31 @@ lemma list2FWpolicys_eq:
"Filter \<noteq> [] \<Longrightarrow> C (list2policyR Filter) = C (list2FWpolicy (rev Filter))"
by (rule ext, erule list2FWpolicys_eq_el)
lemma list2FWpolicys_eq_sym:
"Filter \<noteq> [] \<Longrightarrow>C (list2policyR (rev Filter)) = C (list2FWpolicy Filter)"
by (metis list2FWpolicys_eq rev_is_Nil_conv rev_rev_ident)
lemma p_eq[rule_format]:
"p \<noteq> [] \<longrightarrow> list2policy (map C (rev p)) = C (list2FWpolicy p)"
by (metis l2polR_eq list2FWpolicys_eq_sym rev.simps(1) rev_rev_ident)
lemma p_eq2[rule_format]:
"normalize x \<noteq> [] \<longrightarrow> C(list2FWpolicy(normalize x)) = C x \<longrightarrow>
list2policy(map C (rev(normalize x))) = C x"
by (simp add: p_eq)
lemma p_eq2Q[rule_format]:
"normalizeQ x \<noteq> [] \<longrightarrow> C (list2FWpolicy (normalizeQ x)) = C x \<longrightarrow>
list2policy (map C (rev (normalizeQ x))) = C x"
by (simp add: p_eq)
lemma list2listNMT[rule_format]: "x \<noteq> [] \<longrightarrow>map sem x \<noteq> []"
by (case_tac x) simp_all
lemma Norm_Distr2:
"r o_f ((P \<Otimes>\<^sub>2 (list2policy Q)) o d) = (list2policy ((P \<Otimes>\<^sub>L Q) (op \<Otimes>\<^sub>2) r d))"
by (rule ext, rule Norm_Distr_2)
lemma NATDistr:
"N \<noteq> [] \<Longrightarrow> F = C (list2policyR N) \<Longrightarrow>
(\<lambda>(x, y). x) o\<^sub>f (NAT \<Otimes>\<^sub>2 F \<circ> (\<lambda>x. (x, x))) =
@ -1869,41 +1738,35 @@ apply (rule ext)
apply (rule Norm_Distr_2)
done
lemma C_eq_normalize_manual:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow>
C (list2FWpolicy (normalize_manual_order p l)) = C p"
by (simp add: normalize_manual_order_def C_eq_compile)
lemma p_eq2_manualQ[rule_format]:
"normalize_manual_orderQ x l \<noteq> [] \<longrightarrow> C(list2FWpolicy (normalize_manual_orderQ x l)) = C x \<longrightarrow>
list2policy (map C (rev (normalize_manual_orderQ x l))) = C x"
by (simp add: p_eq)
lemma norm_notMT_manualQ: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalize_manual_orderQ p l \<noteq> []"
by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalize_manual_orderQ_def rADnMT sepnMT sortnMTQ)
lemma C_eq_normalize_manualQ:
"DenyAll\<in>set(policy2list p) \<Longrightarrow> allNetsDistinct(policy2list p) \<Longrightarrow> all_in_list(policy2list p) l \<Longrightarrow>
C (list2FWpolicy (normalize_manual_orderQ p l)) = C p"
by (simp add: normalize_manual_orderQ_def C_eq_compileQ)
lemma p_eq2_manual[rule_format]:
"normalize_manual_order x l \<noteq> [] \<longrightarrow> C (list2FWpolicy (normalize_manual_order x l)) = C x \<longrightarrow>
list2policy (map C (rev (normalize_manual_order x l))) = C x"
by (simp add: p_eq)
lemma norm_notMT_manual: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalize_manual_order p l \<noteq> []"
by (simp add: RS2_NMT idNMT normalize_manual_order_def rADnMT sepnMT sortnMT wp1ID)
text{* As an example, how this theorems can be used for a concrete
normalisation instantiation. *}
text{*
As an example, how this theorems can be used for a concrete normalisation instantiation.
*}
lemma normalizeNAT:
"DenyAll \<in> set (policy2list Filter) \<Longrightarrow> allNetsDistinct (policy2list Filter) \<Longrightarrow>
@ -1913,35 +1776,30 @@ lemma normalizeNAT:
(\<lambda>(x, y). x) (\<lambda>x. (x, x)))"
by (simp add: C_eq_normalize NATDistr list2FWpolicys_eq_sym norm_notMT)
lemma domSimpl[simp]: "dom (C (A \<oplus> DenyAll)) = dom (C (DenyAll))"
by (simp add: PLemmas)
text {* The followin theorems can be applied when prepending the usual normalisation with an
text {*
The followin theorems can be applied when prepending the usual normalisation with an
additional step and using another semantical interpretation function. This is a general recipe
which can be applied whenever one nees to combine several normalisation strategies. *}
which can be applied whenever one nees to combine several normalisation strategies.
*}
lemma CRotate_eq_rotateC: "CRotate p = C (rotatePolicy p)"
by (induct p rule: rotatePolicy.induct) (simp_all add: C.simps map_add_def)
lemma DAinRotate:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> DenyAll \<in> set (policy2list (rotatePolicy p))"
by (induct p,simp_all) (case_tac "DenyAll \<in> set (policy2list p1)",simp_all)
lemma DAUniv: "dom (CRotate (P \<oplus> DenyAll)) = UNIV"
by (metis CRotate.simps(1) CRotate.simps(4) CRotate_eq_rotateC DAAux PLemmas(4) UNIV_eq_I domSubset3)
lemma p_eq2R[rule_format]:
"normalize (rotatePolicy x) \<noteq> [] \<longrightarrow> C(list2FWpolicy(normalize (rotatePolicy x))) = CRotate x \<longrightarrow>
list2policy (map C (rev (normalize (rotatePolicy x)))) = CRotate x"
by (simp add: p_eq)
lemma C_eq_normalizeRotate:
"DenyAll \<in> set (policy2list p) \<Longrightarrow> allNetsDistinct (policy2list (rotatePolicy p)) \<Longrightarrow>
all_in_list (policy2list (rotatePolicy p)) (Nets_List (rotatePolicy p)) \<Longrightarrow>
@ -1955,7 +1813,6 @@ lemma C_eq_normalizeRotate:
CRotate p"
by (simp add: CRotate_eq_rotateC C_eq_compile DAinRotate)
lemma C_eq_normalizeRotate2:
"DenyAll \<in> set (policy2list p) \<Longrightarrow>
allNetsDistinct (policy2list (rotatePolicy p)) \<Longrightarrow>
@ -1963,7 +1820,4 @@ lemma C_eq_normalizeRotate2:
C (list2FWpolicy (FWNormalisationCore.normalize (rotatePolicy p))) = CRotate p"
by (simp add: normalize_def, erule C_eq_normalizeRotate,simp_all)
end