diff --git a/FWNormalisation/ElementaryRules.thy b/FWNormalisation/ElementaryRules.thy index 32523f2..5a849ca 100644 --- a/FWNormalisation/ElementaryRules.thy +++ b/FWNormalisation/ElementaryRules.thy @@ -36,41 +36,43 @@ *****************************************************************************) subsection {* Elementary Firewall Policy Transformation Rules *} -theory ElementaryRules -imports FWNormalisationCore +theory + ElementaryRules + imports + FWNormalisationCore begin - - + + text{* This theory contains those elementary transformation rules which are presented in the ICST 2010 paper~\cite{brucker.ea:firewall:2010}. They are not used elsewhere. *} - + lemma elem1: - "C (AllowPortFromTo x y p \ DenyAllFromTo x y) = C (DenyAllFromTo x y)" -by (rule ext, auto simp: PLemmas) - - + "C (AllowPortFromTo x y p \ DenyAllFromTo x y) = C (DenyAllFromTo x y)" + by (rule ext, auto simp: PLemmas) + + lemma elem2: - "C ((a \ b) \ c) = C (a \ (b \ c))" -by (simp add: C.simps) - + "C ((a \ b) \ c) = C (a \ (b \ c))" + by (simp add: C.simps) + lemma elem3: - "C (AllowPortFromTo x y a \ AllowPortFromTo x y b) = + "C (AllowPortFromTo x y a \ AllowPortFromTo x y b) = C (AllowPortFromTo x y b \ AllowPortFromTo x y a)" -by (rule ext, auto simp: PLemmas) - + by (rule ext, auto simp: PLemmas) + lemma elem4: - "C (a \ DenyAll) = C DenyAll" -by (rule ext, auto simp: PLemmas) - + "C (a \ DenyAll) = C DenyAll" + by (rule ext, auto simp: PLemmas) + lemma elem5: - "C (DenyAllFromTo x y \ DenyAllFromTo u v) = C (DenyAllFromTo u v \ DenyAllFromTo x y)" -by (rule ext, auto simp: PLemmas) - - + "C (DenyAllFromTo x y \ DenyAllFromTo u v) = C (DenyAllFromTo u v \ DenyAllFromTo x y)" + by (rule ext, auto simp: PLemmas) + + lemma elem6: - "dom (C a) \ dom (C b) = {} \ C (a \ b) = C (b \ a)" -by (rule ext, metis C.simps(4) map_add_comm) - + "dom (C a) \ dom (C b) = {} \ C (a \ b) = C (b \ a)" + by (rule ext, metis C.simps(4) map_add_comm) + end diff --git a/FWNormalisation/FWNormalisation.thy b/FWNormalisation/FWNormalisation.thy index 9ad2469..1f38f9b 100644 --- a/FWNormalisation/FWNormalisation.thy +++ b/FWNormalisation/FWNormalisation.thy @@ -35,8 +35,11 @@ *****************************************************************************) chapter {* Firewall Policy Normalisation *} -theory FWNormalisation -imports NormalisationIPPProofs +theory + FWNormalisation + imports + NormalisationIPPProofs + ElementaryRules begin - + end diff --git a/FWNormalisation/FWNormalisationCore.thy b/FWNormalisation/FWNormalisationCore.thy index 7cabc60..1d34ec6 100644 --- a/FWNormalisation/FWNormalisationCore.thy +++ b/FWNormalisation/FWNormalisationCore.thy @@ -85,21 +85,23 @@ text{* We define a very simple policy language: *} datatype ('\,'\) Combinators = DenyAll -| DenyAllFromTo '\ '\ -| AllowPortFromTo '\ '\ '\ -| Conc "(('\,'\) Combinators)" "(('\,'\) Combinators)" (infixr "\" 80) + | DenyAllFromTo '\ '\ + | AllowPortFromTo '\ '\ '\ + | Conc "(('\,'\) Combinators)" "(('\,'\) Combinators)" (infixr "\" 80) 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::"('\, '\) Combinators \ @@ -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:: "(('\, '\) Combinators) list \ (('\, '\) 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 :: "('\, '\) Combinators \ ('\, '\) Combinators \ @@ -346,16 +352,11 @@ 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 - -lemma set_qsort [simp]: "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" + by (simp add: qsort_permutes) + fun insort where "insort a [] l = [a]" | "insort a (x#xs) l = (if (smaller a x l) then a#x#xs else x#(insort a xs l))" @@ -364,12 +365,10 @@ fun sort where "sort [] l = []" | "sort (x#xs) l = insort x (sort xs l) l" - - fun sorted where -"sorted [] l \ True" | -"sorted [x] l \ True" | -"sorted (x#y#zs) l \ smaller x y l \ sorted (y#zs) l" + "sorted [] l = True" +| "sorted [x] l = True" +| "sorted (x#y#zs) l = (smaller x y l \ 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::"(('\, '\) Combinators) list \ @@ -425,7 +423,9 @@ fun list2policyR::"(('\, '\) Combinators) list \ |"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 \ y) = C x ++ C y" - - - fun CRotate :: "(adr\<^sub>i\<^sub>p net, port) Combinators \ (adr\<^sub>i\<^sub>p,DummyContent) packet \ unit" where " CRotate DenyAll = C DenyAll" @@ -447,24 +444,20 @@ where |"CRotate (AllowPortFromTo x y p) = C (AllowPortFromTo x y p)" |"CRotate (x \ 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\b) = (rotatePolicy b) \ (rotatePolicy a)" - lemma check: "rev (policy2list (rotatePolicy p)) = policy2list p" -apply (induct p) -apply simp -apply simp_all -done - + apply (induct p) + 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) = ((\ c a b. c = DenyAllFromTo a b \ c \ set xs \ Map.dom (C x) \ Map.dom (C c) = {}) \ 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 \ 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 \ (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 \ (adr\<^sub>i\<^sub>p net, port) Combinators list" @@ -524,11 +515,11 @@ 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) - - + 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 \ y) = Cp (y \ 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) = ((\ c a b. c = DenyAllFromTo a b \ c \ set xs \ Map.dom (Cp x) \ Map.dom (Cp c) = {}) \ 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 \ port) Combinators) list \ 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) \ dom (Cp (AllowPortFromTo a b p)) = {}) \ wellformed_policy3Pr xs)" | "wellformed_policy3Pr (x#xs) = wellformed_policy3Pr xs" - - - definition normalizePr' :: "(adr\<^sub>i\<^sub>p\<^sub>p net, protocol \ port) Combinators \ (adr\<^sub>i\<^sub>p\<^sub>p net, protocol \ 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 \ port) Combinators \ (adr\<^sub>i\<^sub>p\<^sub>p net, protocol \ 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 \ port) Combinators \ (adr\<^sub>i\<^sub>p\<^sub>p net, protocol \ 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) + 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)) \ 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: "\x \ a; y\b; (x \ y \ x \ b) \ (a \ b \ a \ y)\ \ {x,a} \ {y,b}" by (auto) - - lemma aux2: "{a,b} = {b,a}" by auto - - end diff --git a/FWNormalisation/NormalisationGenericProofs.thy b/FWNormalisation/NormalisationGenericProofs.thy index 3d01156..7cfe1f1 100644 --- a/FWNormalisation/NormalisationGenericProofs.thy +++ b/FWNormalisation/NormalisationGenericProofs.thy @@ -36,31 +36,28 @@ *****************************************************************************) subsection{* Normalisation Proofs (Generic) *} -theory NormalisationGenericProofs - imports - FWNormalisationCore + 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 \ {} \ X \ \" - apply auto - done + by auto lemma denyNMT: "deny_all \ \" 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 \ -DenyAll \ set p \ (\ p'. p = DenyAll # p' \ DenyAll \ set p')" +lemma wellformed_policy1_charn[rule_format]: +"wellformed_policy1 p \ DenyAll \ set p \ (\ p'. p = DenyAll # p' \ DenyAll \ set p')" by(induct p,simp_all) - lemma singleCombinatorsConc: "singleCombinators (x#xs) \ singleCombinators xs" by (case_tac x,simp_all) @@ -69,7 +66,6 @@ lemma aux0_0: "singleCombinators x \ \ (\ a b. (a\< apply (rule allI)+ by (case_tac a,simp_all) - lemma aux0_4: "(a \ set x \ a \ set y) = (a \ set (x@y))" by auto @@ -77,27 +73,22 @@ lemma aux0_1: "\singleCombinators xs; singleCombinators [x]\ \singleCombinators xs; \ (\ a b. x = a \ b)\ \ singleCombinators(x#xs)" apply (rule aux0_1,simp_all) apply (case_tac x,simp_all) - apply auto - done + by auto lemma aux0_5: " \ (\ a b. (a\b) \ set x) \ singleCombinators x" apply (induct x) apply simp_all by (metis aux0_6) - - lemma ANDConc[rule_format]: "allNetsDistinct (a#p) \ allNetsDistinct (p)" apply (simp add: allNetsDistinct_def) apply (case_tac "a") by simp_all - lemma aux6: "twoNetsDistinct a1 a2 a b \ dom (deny_all_from_to a1 a2) \ 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 \ lemma aux5[rule_format]: "(DenyAllFromTo a b) \ set p \ a \ set (net_list p)" by (rule net_list_aux.induct,simp_all) - lemma aux5a[rule_format]: "(DenyAllFromTo b a) \ set p \ a \ 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 \ set (net_list p) \ a \ set (net_list_aux p)" by simp - -lemma srcInNetListaux[simp]: "\x \ set p; singleCombinators[x]; x \ DenyAll\ \ - srcNet x \ set (net_list_aux p)" +lemma srcInNetListaux[simp]: + "\x \ set p; singleCombinators[x]; x \ DenyAll\ \ srcNet x \ 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]: "\x \ set p; singleCombinators[x]; x \ DenyAll\ \ - destNet x \ set (net_list_aux p)" +lemma destInNetListaux[simp]: + "\x \ set p; singleCombinators[x]; x \ DenyAll\ \ destNet x \ 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: "\allNetsDistinct p; x \ set p; y \ set p; a = srcNet x; b = destNet x; c = srcNet y; d = destNet y; a \ c; singleCombinators[x]; x \ DenyAll; singleCombinators[y]; @@ -189,8 +176,6 @@ lemma aux[rule_format]: "a \ set (removeShadowRules2 p) \ a lemma aux12: "\a \ x; b \ x\ \ a \ b" by auto - - lemma ND0aux1[rule_format]: "DenyAllFromTo x y \ set b \ x \ set (net_list_aux b)" by (metis aux5 net_list.simps set_remdups) @@ -233,8 +218,7 @@ lemma aNDSubset: "\singleCombinators a;set a \ 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: "\singleCombinators a; singleCombinators b; set a = set b; allNetsDistinct b\ \ allNetsDistinct a" @@ -242,21 +226,15 @@ lemma aNDSetsEq: "\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: "\singleCombinators p; singleCombinators [a]\ \ - singleCombinators (a#p)" - by (case_tac "a",simp_all) + singleCombinators (a#p)" + by(metis aux0_1) lemma aux3[simp]: "\singleCombinators p; singleCombinators [a]; - allNetsDistinct (a#p)\ \ allNetsDistinct (a#a#p)" - apply (insert aNDSubset[of "(a#a#p)" "(a#p)"]) - by (simp add: SCConca) - - - - + allNetsDistinct (a#p)\ \ allNetsDistinct (a#a#p)" + by (metis aNDSetsEq aux0_1 insert_absorb2 list.set(2)) lemma wp1_aux1a[rule_format]: "xs \ [] \ wellformed_policy1_strong (xs @ [x]) \ wellformed_policy1_strong xs" @@ -307,10 +285,9 @@ lemma waux3[rule_format]: "\x \ a; x \ set p\ \ - 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) \ wellformed_policy1 (qsort p l)" @@ -318,14 +295,11 @@ proof (induct p) case Nil show ?case by simp next case (Cons a S) then show ?case - apply simp_all + apply simp_all apply (cases a,simp_all) - by (metis Combinators.simps append_Cons append_Nil qsort.simps(2) set_ConsD set_qsort waux2)+ + by (metis Combinators.simps append_Cons append_Nil qsort.simps(2) set_ConsD set_qsort waux2)+ qed - - - lemma SR1Subset: "set (removeShadowRules1 p) \ set p" apply (induct_tac p, simp_all) apply (case_tac a, simp_all) @@ -346,7 +320,7 @@ lemma setInsert[simp]: "set list \ insert a (set list)" lemma SC_RS1[rule_format,simp]: "singleCombinators p \ allNetsDistinct p \ singleCombinators (removeShadowRules1 p)" apply (induct_tac p) - apply simp_all + apply simp_all using ANDConc singleCombinatorsConc by blast lemma RS2Set[rule_format]: "set (removeShadowRules2 p) \ set p" @@ -363,7 +337,6 @@ lemma denyAllDom[simp]: "x \ dom (deny_all)" lemma lCdom2: "(list2FWpolicy (a @ (b @ c))) = (list2FWpolicy ((a@b)@c))" by auto - lemma SCConcEnd: "singleCombinators (xs @ [xa]) \ singleCombinators xs" by (induct "xs", simp_all, case_tac a, simp_all) @@ -371,7 +344,6 @@ lemma list2FWpolicyconc[rule_format]: "a \ [] \ (list2FWpolicy (xa # a)) = (xa) \ (list2FWpolicy a)" by (induct a,simp_all) - lemma wp1n_tl [rule_format]: "wellformed_policy1_strong p \ p = (DenyAll#(tl p))" by (induct p, simp_all) @@ -423,8 +395,7 @@ lemma domInterMT[rule_format]: "\dom a \ dom b = {}; x \ dom lemma domComm: "dom a \ dom b = dom b \ dom a" by auto - - + lemma r_not_DA_in_tl[rule_format]: "wellformed_policy1_strong p \ a \ set p\ a \ DenyAll \ a \ set (tl p)" by (induct p,simp_all) @@ -441,9 +412,7 @@ lemma l2p_aux[rule_format]: "list \ [] \ lemma l2p_aux2[rule_format]: "list = [] \ list2FWpolicy (a # list) = a" by simp - - - + lemma aux7aa: assumes 1 : "AllowPortFromTo a b poo \ set p" and 2 : "allNetsDistinct ((AllowPortFromTo c d po) # p)" @@ -459,10 +428,9 @@ next list.set_intros(2) twoNetsDistinct_def) qed - lemma ANDConcEnd: "\ allNetsDistinct (xs @ [xa]); singleCombinators xs\ \ allNetsDistinct xs" - by (rule aNDSubset) auto + by (rule aNDSubset, auto) lemma WP1ConcEnd[rule_format]: "wellformed_policy1 (xs@[xa]) \ wellformed_policy1 xs" @@ -505,21 +473,15 @@ next qed qed - - - lemma SC1[simp]: "singleCombinators p \singleCombinators (removeShadowRules1 p)" by (erule SCSubset) (rule SR1Subset) lemma SC2[simp]: "singleCombinators p \singleCombinators (removeShadowRules2 p)" by (erule SCSubset) (rule RS2Set) - lemma SC3[simp]: "singleCombinators p \ singleCombinators (sort p l)" by (erule SCSubset) simp - - lemma SC3Q[simp]: "singleCombinators p \ singleCombinators (qsort p l)" by (erule SCSubset) simp @@ -549,8 +511,6 @@ lemma aND_sortQ[simp]: "\singleCombinators p; allNetsDistinct p\ apply (rule aNDSubset) by (erule SC3Q, simp_all) - - lemma inRS2[rule_format,simp]: "x \ set p \ x \ set (removeShadowRules2 p)" apply (insert RS2Set [of p]) by blast @@ -627,11 +587,11 @@ lemma sortedConcStart[rule_format]: singleCombinators [a] \ singleCombinators [aa] \ singleCombinators p \ sorted (a#p) l" apply (induct p) - apply simp_all + apply simp_all apply (rule impI)+ apply simp apply (rule_tac y = "aa" in order_trans) - apply simp_all + apply simp_all apply (case_tac ab, simp_all) done @@ -644,14 +604,14 @@ lemma sorted_is_smaller[rule_format]: "sorted (a # p) l \ in_list a l \ in_list b l \ all_in_list p l \ singleCombinators [a] \ singleCombinators p \ b \ set p \ smaller a b l" apply (induct p) - apply (auto intro: singleCombinatorsConc sortedConcStart) + apply (auto intro: singleCombinatorsConc sortedConcStart) done lemma sortedConcEnd[rule_format]: "sorted (a # p) l \ in_list a l \ all_in_list p l \ singleCombinators [a] \ singleCombinators p \ sorted p l" apply (induct p) - apply (auto intro: singleCombinatorsConc sortedConcStart) + apply (auto intro: singleCombinatorsConc sortedConcStart) done lemma in_set_in_list[rule_format]: "a \ set p \ all_in_list p l\ in_list a l" @@ -661,14 +621,14 @@ lemma sorted_Consb[rule_format]: "all_in_list (x#xs) l \ singleCombinators (x#xs) \ (sorted xs l & (ALL y:set xs. smaller x y l)) \ (sorted (x#xs) l) " apply(induct xs arbitrary: x) - apply (auto simp: order_trans) + apply (auto simp: order_trans) done lemma sorted_Cons: "\all_in_list (x#xs) l; singleCombinators (x#xs)\ \ (sorted xs l & (ALL y:set xs. smaller x y l)) = (sorted (x#xs) l)" apply auto - apply (rule sorted_Consb, simp_all) - apply (metis singleCombinatorsConc singleCombinatorsStart sortedConcEnd) + apply (rule sorted_Consb, simp_all) + apply (metis singleCombinatorsConc singleCombinatorsStart sortedConcEnd) apply (erule sorted_is_smaller) apply (auto intro: singleCombinatorsConc singleCombinatorsStart in_set_in_list) done @@ -677,15 +637,15 @@ lemma smaller_antisym: "\\ smaller a b l; in_list a l; in_list b l; singleCombinators[a]; singleCombinators [b]\ \ smaller b a l" apply (case_tac a) - apply simp_all + apply simp_all + apply (case_tac b) + apply simp_all + apply (simp_all split: if_splits) + apply (rule setPaireq) + apply simp apply (case_tac b) - apply simp_all - apply (simp_all split: if_splits) - apply (rule setPaireq) - apply simp - apply (case_tac b) - apply simp_all - apply (simp_all split: if_splits) + apply simp_all + apply (simp_all split: if_splits) done lemma set_insort_insert: "set (insort x xs l) \ insert x (set xs)" @@ -702,8 +662,8 @@ lemma singleCombinators_insort: "\singleCombinators [x]; singleCombinato lemma all_in_list_insort: "\all_in_list xs l; singleCombinators (x#xs); in_list x l\ \ all_in_list (insort x xs l) l" apply (rule_tac b = "x#xs" in all_in_listSubset) - apply simp_all - apply (metis singleCombinatorsConc singleCombinatorsStart + apply simp_all + apply (metis singleCombinatorsConc singleCombinatorsStart singleCombinators_insort) apply (rule set_insort_insert) done @@ -715,8 +675,6 @@ lemma sorted_ConsA:"\all_in_list (x#xs) l; singleCombinators (x#xs)\ set xs \ y \ 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,11 +708,11 @@ next have * : "sorted (a # xs) l \ all_in_list (x # a # xs) l \ distinct (x # a # xs) \ singleCombinators [x] \ singleCombinators (a # xs) \ sorted (insort x xs l) l" - apply(insert Cons.hyps)apply simp_all - apply (metis sorted_Cons all_in_list.simps(2) singleCombinatorsConc) + apply(insert Cons.hyps, simp_all) + apply(metis sorted_Cons all_in_list.simps(2) singleCombinatorsConc) done show ?case - apply(insert Cons.hyps) + apply (insert Cons.hyps) apply (rule impI)+ apply (insert *, auto intro!: sorted_Consb) proof (rule_tac b = "x#xs" in all_in_listSubset) @@ -789,7 +747,6 @@ next qed qed - lemma sorted_insort: "\all_in_list (x#xs) l; distinct(x#xs); singleCombinators [x]; singleCombinators xs\ \ @@ -806,56 +763,46 @@ lemma distinct_sort[simp]: "distinct (sort xs l) = distinct xs" lemma sort_is_sorted[rule_format]: "all_in_list p l \ distinct p \ singleCombinators p \ 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 \ smaller a a l" by (case_tac a,simp_all) - lemma SC_sublist[rule_format]: "singleCombinators xs \ singleCombinators (qsort [y\xs. P y] l)" by (auto intro: SCSubset) - lemma all_in_list_sublist[rule_format]: "singleCombinators xs \ all_in_list xs l \ all_in_list (qsort [y\xs. P y] l) l" by (auto intro: all_in_listSubset SC_sublist) - lemma SC_sublist2[rule_format]: "singleCombinators xs \ singleCombinators ([y\xs. P y])" by (auto intro: SCSubset) - lemma all_in_list_sublist2[rule_format]: "singleCombinators xs \ all_in_list xs l \ all_in_list ( [y\xs. P y]) l" by (auto intro: all_in_listSubset SC_sublist2) - lemma all_in_listAppend[rule_format]: "all_in_list (xs) l \ all_in_list (ys) l \ all_in_list (xs @ ys) l" by (induct xs) simp_all - - lemma distinct_sortQ[rule_format]: "singleCombinators xs \ all_in_list xs l \ distinct xs \ 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) \ singleCombinators (ys) \ singleCombinators (xs @ ys)" apply (induct xs, auto) - apply (case_tac a,simp_all)+ + apply (case_tac a,simp_all)+ done lemma sorted_append1[rule_format]: @@ -864,7 +811,7 @@ lemma sorted_append1[rule_format]: (sorted (xs@ys) l \ (sorted xs l & sorted ys l & (\x \ set xs. \y \ set ys. smaller x y l)))" apply(induct xs) - apply(simp_all) + apply(simp_all) by (metis NormalisationGenericProofs.sorted_Cons all_in_list.simps(2) all_in_listAppend aux0_1 aux0_4 singleCombinatorsAppend singleCombinatorsConc singleCombinatorsStart) @@ -874,7 +821,7 @@ lemma sorted_append2[rule_format]: (sorted xs l & sorted ys l & (\x \ set xs. \y \ set ys. smaller x y l)) \ (sorted (xs@ys) l)" apply (induct xs) - apply simp_all + apply simp_all by (metis NormalisationGenericProofs.sorted_Cons all_in_list.simps(2) all_in_listAppend aux0_1 aux0_4 singleCombinatorsAppend singleCombinatorsConc singleCombinatorsStart) @@ -885,13 +832,10 @@ lemma sorted_append[rule_format]: (sorted xs l & sorted ys l & (\x \ set xs. \y \ set ys. smaller x y l))" apply (rule impI)+ apply (rule iffI) - apply (rule sorted_append1,simp_all) + apply (rule sorted_append1,simp_all) apply (rule sorted_append2,simp_all) done - - - lemma sort_is_sortedQ[rule_format]: "all_in_list p l \ singleCombinators p \ sorted (qsort p l) l" proof (induct p l rule: qsort.induct) print_cases @@ -941,7 +885,7 @@ next FWNormalisationCore.sorted (x # qsort [y\xs . smaller x y l] l) l \ (\x'\set (qsort [y\xs . \ smaller x y l] l). \y\set (x # qsort [y\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) @@ -966,15 +910,14 @@ lemma RS1n_nMT[rule_format,simp]: "p \ []\ removeShadowRu apply (simp add: removeShadowRules1_alternative_def) apply (rule_tac xs = p in rev_induct, simp_all) apply (case_tac "xs = []", simp_all) - apply (case_tac x, simp_all) + apply (case_tac x, simp_all) apply (rule_tac xs = "xs" in rev_induct, simp_all) - apply (case_tac x, simp_all)+ + apply (case_tac x, simp_all)+ done 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 \ DenyAll \ set (tl p)" by (induct p) (simp_all) @@ -991,11 +934,9 @@ lemma AND_tl[rule_format]: "allNetsDistinct ( p) \ allNetsDistin apply (induct p, simp_all) by (auto intro: ANDConc) - lemma distinct_tl[rule_format]: "distinct p \ distinct (tl p)" by (induct p, simp_all) - lemma SC_tl[rule_format]: "singleCombinators ( p) \ singleCombinators (tl p)" by (induct p, simp_all) (auto intro: singleCombinatorsConc) @@ -1006,21 +947,17 @@ lemma wp1_tl[rule_format]: "p \ [] \ wellformed_policy1 p \ wellformed_policy1 (tl p)" by (induct p) (auto intro: waux2) - lemma wp1_eq[rule_format]: "wellformed_policy1_strong p \ wellformed_policy1 p" apply (case_tac "DenyAll \ set p") - apply (subst wellformed_eq) - apply (auto elim: waux2) + apply (subst wellformed_eq) + apply (auto elim: waux2) done - lemma wellformed1_alternative_sorted: "wellformed_policy1_strong p \ wellformed_policy1_strong (sort p l)" by (case_tac "p", simp_all) - - lemma wp1n_RS2[rule_format]: "wellformed_policy1_strong p \ wellformed_policy1_strong (removeShadowRules2 p)" by (induct p, simp_all) @@ -1028,10 +965,9 @@ lemma wp1n_RS2[rule_format]: lemma RS2_NMT[rule_format]: "p \ [] \ removeShadowRules2 p \ []" apply (induct p, simp_all) apply (case_tac "p \ []", simp_all) - apply (case_tac "a", simp_all)+ + apply (case_tac "a", simp_all)+ done - lemma wp1_alternative_not_mt[simp]: "wellformed_policy1_strong p \ p \ []" by auto @@ -1039,8 +975,6 @@ lemma AIL1[rule_format,simp]: "all_in_list p l \ 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 \ all_in_list (remdups p) l" by (induct p, simp_all) - lemma AILiD[rule_format,simp]: "all_in_list p l \ all_in_list (insertDeny p) l" apply (induct p, simp_all) apply (rule impI, simp) @@ -1060,7 +993,7 @@ lemma AILiD[rule_format,simp]: "all_in_list p l \ all_in_list (i lemma SCrd[rule_format,simp]:"singleCombinators p\ singleCombinators(remdups p)" apply (induct p, simp_all) apply (case_tac a) - apply simp_all + apply simp_all done lemma SCRiD[rule_format,simp]: "singleCombinators p \ @@ -1069,7 +1002,6 @@ lemma SCRiD[rule_format,simp]: "singleCombinators p \ apply (case_tac "a", simp_all) done - lemma WP1rd[rule_format,simp]: "wellformed_policy1_strong p \ wellformed_policy1_strong (remdups p)" by (induct p, simp_all) @@ -1078,18 +1010,17 @@ lemma ANDrd[rule_format,simp]: "singleCombinators p \ allNetsDistinct p \ allNetsDistinct (remdups p)" apply (rule impI)+ apply (rule_tac b = p in aNDSubset) - apply simp_all + apply simp_all done lemma ANDiD[rule_format,simp]: "allNetsDistinct p \ allNetsDistinct (insertDeny p)" apply (induct p, simp_all) - apply (simp add: allNetsDistinct_def) + apply (simp add: allNetsDistinct_def) apply (auto intro: ANDConc) apply (case_tac "a",simp_all add: allNetsDistinct_def) done - lemma mr_iD[rule_format]: "wellformed_policy1_strong p \ 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 \ wellformed_policy1_strong (insertDeny p)" by (induct p, simp_all) - lemma DAiniD: "DenyAll \ 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 \ all_in_list (removeShadowRules2 p) l" by (induct_tac p, simp_all, case_tac a, simp_all) - lemma SCConc: "singleCombinators x \ singleCombinators y \ singleCombinators (x@y)" apply (rule aux0_5) apply (metis aux0_0 aux0_4) @@ -1118,7 +1047,6 @@ lemma SCConc: "singleCombinators x \ singleCombinators y \ A \ {} \ A \ B \ Dd \ B \ {}" by auto @@ -1126,8 +1054,7 @@ lemma soadisj: "x \ subnetsOfAdr a \ y \ subnetsOfAdr a by(simp add: subnetsOfAdr_def netsDistinct_def,auto) lemma not_member: "\ member a (x\y) \ \ member a x" - apply auto - done + by auto lemma soadisj2: "(\ a x y. x \ subnetsOfAdr a \ y \ subnetsOfAdr a \ \ 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: "(\a b c d. (a,b)\A \ (c,d)\B \ netsDistinct b d) \ \(a, b)\A. b \ subnetsOfAdr D \ \(a, b)\B. b \ subnetsOfAdr D \ False" apply (auto simp: soadisj) using soadisj2 by blast - lemma tndFalse: "(\a b c d. (a,b)\A \ (c,d)\B \ twoNetsDistinct a b c d) \ \(a, b)\A. a \ subnetsOfAdr (D::('a::adr)) \ b \ subnetsOfAdr (F::'a) \ \(a, b)\B. a \ subnetsOfAdr D\ b\ subnetsOfAdr F @@ -1162,7 +1087,6 @@ lemma sepnMT[rule_format]: "p \ [] \ (separate p) \ set p \ DenyAll \ set (separate p)" by (induct p rule: separate.induct) simp_all - lemma setnMT: "set a = set b \ a \ [] \ b \ []" by auto @@ -1177,7 +1101,7 @@ lemma idNMT[rule_format]: "p \ [] \ insertDenies p \ x \ DenyAll \ x \ set p \ onlyTwoNets x" apply (induct p, simp_all, rename_tac a p) apply (intro impI conjI, simp) - apply (case_tac a, simp_all) + apply (case_tac a, simp_all) apply (drule mp, simp_all) apply (case_tac a, simp_all) done @@ -1185,7 +1109,6 @@ lemma OTNoTN[rule_format]: " OnlyTwoNets p \ x \ DenyAll lemma first_isIn[rule_format]: "\ member DenyAll x \ (first_srcNet x,first_destNet x) \ sdnets x" by (induct x,case_tac x, simp_all) - lemma sdnets2: "\a b. sdnets x = {(a, b), (b, a)} \ \ member DenyAll x \ sdnets x = {(first_srcNet x, first_destNet x), (first_destNet x, first_srcNet x)}" @@ -1193,7 +1116,6 @@ proof - have * : "\a b. sdnets x = {(a, b), (b, a)} \ \ member DenyAll x \ (first_srcNet x, first_destNet x) \ sdnets x" by (erule first_isIn) - show "\a b. sdnets x = {(a, b), (b, a)} \ \ member DenyAll x \ 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 \ set (net_list_aux [x]) \ a \ set (net_list_aux [y,x])" by (induct y, simp_all) - lemma noDA[rule_format]: "noDenyAll xs \ s \ set xs \ \ member DenyAll s" by (induct xs, simp_all) @@ -1240,7 +1161,6 @@ lemma disjComm: "disjSD_2 a b \ disjSD_2 b a" using tNDComm apply blast by (meson tNDComm twoNetsDistinct_def) - lemma disjSD2aux: "disjSD_2 a b \ \ member DenyAll a \ \ member DenyAll b \ disjSD_2 (DenyAllFromTo (first_srcNet a) (first_destNet a) \ @@ -1250,11 +1170,9 @@ lemma disjSD2aux: apply (simp add: disjSD_2_def) using first_isIn by blast - lemma noDA1eq[rule_format]: "noDenyAll p \ noDenyAll1 p" by (induct p, simp,rename_tac a p, case_tac a, simp_all) - lemma noDA1C[rule_format]: "noDenyAll1 (a#p) \ 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 \ DenyAllFromTo b a \ x) y" by(simp add:disjSD2aux) - lemma noDAID[rule_format]: "noDenyAll p \ noDenyAll (insertDenies p)" by (induct p, simp_all,case_tac a, simp_all) - lemma isInIDo[rule_format]: "noDenyAll p \ s \ set (insertDenies p) \ (\! a. s = (DenyAllFromTo (first_srcNet a) (first_destNet a)) \ @@ -1286,7 +1202,6 @@ lemma id_aux1[rule_format]: "DenyAllFromTo (first_srcNet s) (first_destNet s) \< \ s \ set p" by (induct p, simp_all, rename_tac a p, case_tac a, simp_all) - lemma id_aux2: "noDenyAll p \ \s. s \ set p \ disjSD_2 a s \ @@ -1297,8 +1212,6 @@ lemma id_aux2: DenyAllFromTo (first_destNet s) (first_srcNet s) \ s)" by (metis disjComm disjSD2aux isInIDo noDA) - - lemma id_aux4[rule_format]: "noDenyAll p \ \s. s \ set p \ disjSD_2 a s \ s \ set (insertDenies p) \ \ member DenyAll a \ @@ -1307,28 +1220,23 @@ lemma id_aux4[rule_format]: DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a \ a \ set p") - apply (drule_tac Q = "disjSD_2 a s" in exE, simp_all, rule id_aux2, simp_all) + 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 \ separated p \ separated (insertDenies p)" apply (induct p, simp) apply (rename_tac a p, auto) using noDA1C apply blast apply (case_tac "a = DenyAll", auto) - apply (simp add: disjSD_2_def) + apply (simp add: disjSD_2_def) apply (case_tac a,auto) - apply (rule disjSD_2IDa, simp_all, rule id_aux4, simp_all, metis noDA noDAID)+ + apply (rule disjSD_2IDa, simp_all, rule id_aux4, simp_all, metis noDA noDAID)+ done lemma aNDDA[rule_format]: "allNetsDistinct p \ allNetsDistinct(DenyAll#p)" by (case_tac p,auto simp: allNetsDistinct_def) - lemma OTNConc[rule_format]: "OnlyTwoNets (y # z) \ OnlyTwoNets z" by (case_tac y, simp_all) @@ -1352,8 +1260,7 @@ lemma setPair4: "{a,b} = {c,d} \ a \ c \ lemma otnaux1: " {x, y, x, y} = {x,y}" by auto - - + lemma OTNIDaux4: "{x,y,x} = {y,x}" by auto @@ -1367,110 +1274,105 @@ lemma otnaux: " apply (simp add: onlyTwoNets_def) apply (subgoal_tac "{first_srcNet x, first_destNet x} = {first_srcNet y, first_destNet y}") - apply (case_tac "(\a b. sdnets y = {(a, b)})") - apply simp_all - apply (case_tac "(\a b. sdnets x = {(a, b)})") - apply simp_all - apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x)}") - apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y)}") - apply simp - apply (case_tac "first_srcNet x = first_srcNet y") - apply simp_all - apply (rule disjI1) - apply (rule setPair) - apply simp - apply (subgoal_tac "first_srcNet x = first_destNet y") - 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 setPair1) - apply simp - apply (rule setPair4) - apply simp_all - apply (metis first_isIn singletonE) - apply (metis first_isIn singletonE) - apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x), - (first_destNet x, first_srcNet x)}") - apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y)}") - apply simp - apply (case_tac "first_srcNet x = first_srcNet y") - apply simp_all - apply (subgoal_tac "first_destNet x = first_destNet y") - apply simp - apply (rule setPair) - apply simp - apply (subgoal_tac "first_srcNet x = first_destNet y") - 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 (metis OTNIDaux4 insert_commute ) - apply (rule setPair1) - apply simp - apply (rule setPair5) - apply assumption - apply simp - apply (metis first_isIn singletonE) - apply (rule sdnets2) - apply simp_all - apply (case_tac "(\a b. sdnets x = {(a, b)})") - apply simp_all - apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x)}") - apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y), + apply (case_tac "(\a b. sdnets y = {(a, b)})") + apply simp_all + apply (case_tac "(\a b. sdnets x = {(a, b)})") + apply simp_all + apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x)}") + apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y)}") + apply simp + apply (case_tac "first_srcNet x = first_srcNet y") + apply simp_all + apply (rule disjI1) + apply (rule setPair) + apply simp + apply (subgoal_tac "first_srcNet x = first_destNet y") + 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 setPair1) + apply simp + apply (rule setPair4) + apply simp_all + apply (metis first_isIn singletonE) + apply (metis first_isIn singletonE) + apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x), + (first_destNet x, first_srcNet x)}") + apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y)}") + apply simp + apply (case_tac "first_srcNet x = first_srcNet y") + apply simp_all + apply (subgoal_tac "first_destNet x = first_destNet y") + apply simp + apply (rule setPair) + apply simp + apply (subgoal_tac "first_srcNet x = first_destNet y") + 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 (metis OTNIDaux4 insert_commute ) + apply (rule setPair1) + apply simp + apply (rule setPair5) + apply assumption + apply simp + apply (metis first_isIn singletonE) + apply (rule sdnets2) + apply simp_all + apply (case_tac "(\a b. sdnets x = {(a, b)})") + apply simp_all + apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x)}") + apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y), (first_destNet y, first_srcNet y)}") - apply simp - apply (case_tac "first_srcNet x = first_srcNet y") - 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 (metis OTNIDaux4 insert_commute ) - apply (rule setPair) - apply simp - apply (subgoal_tac "first_srcNet x = first_destNet y") - apply simp - apply (subgoal_tac "first_destNet x = first_srcNet y") - apply simp - apply (rule setPair1) - apply simp - apply (rule setPair4) - apply assumption - apply simp - apply (rule sdnets2) - apply simp - apply simp - apply (metis singletonE first_isIn) - apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x), + apply simp + apply (case_tac "first_srcNet x = first_srcNet y") + 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 (metis OTNIDaux4 insert_commute ) + apply (rule setPair) + apply simp + apply (subgoal_tac "first_srcNet x = first_destNet y") + apply simp + apply (subgoal_tac "first_destNet x = first_srcNet y") + apply simp + apply (rule setPair1) + apply simp + apply (rule setPair4) + apply assumption + apply simp + apply (rule sdnets2) + apply simp + apply simp + apply (metis singletonE first_isIn) + apply (subgoal_tac "sdnets x = {(first_srcNet x, first_destNet x), (first_destNet x, first_srcNet x)}") - apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y), + apply (subgoal_tac "sdnets y = {(first_srcNet y, first_destNet y), (first_destNet y, first_srcNet y)}") - apply simp - apply (case_tac "first_srcNet x = first_srcNet y") - 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 otnaux1) - apply (rule setPair) - apply simp - apply (subgoal_tac "first_srcNet x = first_destNet y") - 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 (metis OTNIDaux4 insert_commute) - apply (rule setPair1) - apply simp - apply (rule setPair4) - apply assumption - apply simp - apply (rule sdnets2,simp_all)+ + apply simp + apply (case_tac "first_srcNet x = first_srcNet y") + 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 otnaux1) + apply (rule setPair) + apply simp + apply (subgoal_tac "first_srcNet x = first_destNet y") + 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 (metis OTNIDaux4 insert_commute) + apply (rule setPair1) + apply simp + apply (rule setPair4) + apply assumption + apply simp + apply (rule sdnets2,simp_all)+ apply (rule bNaux, simp_all) done @@ -1480,13 +1382,13 @@ lemma OTNSepaux: noDenyAll z \ onlyTwoNets a \ OnlyTwoNets (y # z) \ first_bothNet a = first_bothNet y \ OnlyTwoNets (separate (a \ y # z))" apply (drule mp) - apply simp_all + apply simp_all apply (rule conjI) - apply (rule otnaux) - apply simp_all - apply (rule_tac p = "(y # z)" in OTNoTN) - apply simp_all - apply (metis member.simps(2)) + apply (rule otnaux) + apply simp_all + apply (rule_tac p = "(y # z)" in OTNoTN) + apply simp_all + apply (metis member.simps(2)) apply (simp add: onlyTwoNets_def) apply (rule_tac y = y in OTNConc,simp) done @@ -1499,12 +1401,11 @@ lemma OTNSEp[rule_format]: lemma nda[rule_format]: "singleCombinators (a#p) \ noDenyAll p \ noDenyAll1 (a # p)" apply (induct p,simp_all) - apply (case_tac a, simp_all)+ + apply (case_tac a, simp_all)+ done lemma nDAcharn[rule_format]: "noDenyAll p = (\ r \ set p. \ member DenyAll r)" - by (induct p) simp_all - + by (induct p, simp_all) lemma nDAeqSet: "set p = set s \ 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 = (\ 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 \ noDenyAll (separate p)" by (simp add:noDAAll,subst memberPsep, simp) - lemma noDA1sep[rule_format]: "noDenyAll1 p \ noDenyAll1 (separate p)" by (induct p rule:separate.induct, simp_all add: noDAsep) - lemma isInAlternativeLista: "(aa \ set (net_list_aux [a]))\ aa \ set (net_list_aux (a # p))" by (case_tac a,auto) - lemma isInAlternativeListb: "(aa \ set (net_list_aux p))\ aa \ set (net_list_aux (a # p))" by (case_tac a,simp_all) - lemma ANDSepaux: "allNetsDistinct (x # y # z) \ allNetsDistinct (x \ 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,19 +1454,16 @@ 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) \ allNetsDistinct(separate(y#z)) \ allNetsDistinct(x#separate(y#z))" apply (simp add: allNetsDistinct_def) by (metis isInAlternativeList netlistalternativeSeparate netlistaux) - - lemma ANDSep[rule_format]: "allNetsDistinct p \ allNetsDistinct(separate p)" apply (induct p rule: separate.induct, simp_all) - apply (metis ANDConc aNDDA) - apply (metis ANDConc ANDSepaux ANDSepaux2) - apply (metis ANDConc ANDSepaux ANDSepaux2) + apply (metis ANDConc aNDDA) + apply (metis ANDConc ANDSepaux ANDSepaux2) + apply (metis ANDConc ANDSepaux ANDSepaux2) apply (metis ANDConc ANDSepaux ANDSepaux2) done @@ -1595,14 +1485,13 @@ proof - list.simps(15) nDAeqSet noDA1eq) qed - lemma OTNSC[rule_format]: "singleCombinators p \ OnlyTwoNets p" apply (induct p,simp_all) apply (rename_tac a p) apply (rule impI,drule mp) - apply (erule singleCombinatorsConc) + apply (erule singleCombinatorsConc) apply (case_tac a, simp_all) - apply (simp add: onlyTwoNets_def)+ + apply (simp add: onlyTwoNets_def)+ done lemma fMTaux: "\ member DenyAll x \ first_bothNet x \ {}" @@ -1615,36 +1504,31 @@ lemma fl2[rule_format]: "firstList (separate p) = firstList p" lemma fl3[rule_format]: "NetsCollected p \ (first_bothNet x \ firstList p \ (\a\set p. first_bothNet x \ first_bothNet a))\ NetsCollected (x#p)" by (induct p) simp_all - - + lemma sortedConc[rule_format]: " sorted (a # p) l \ sorted p l" by (induct p) simp_all - lemma smalleraux2: "{a,b} \ set l \ {c,d} \ set l \ {a,b} \ {c,d} \ smaller (DenyAllFromTo a b) (DenyAllFromTo c d) l \ \ smaller (DenyAllFromTo c d) (DenyAllFromTo a b) l" by (metis bothNet.simps(2) pos_noteq smaller.simps(5)) - lemma smalleraux2a: "{a,b} \ set l \ {c,d} \ set l \ {a,b} \ {c,d} \ smaller (DenyAllFromTo a b) (AllowPortFromTo c d p) l \ \ smaller (AllowPortFromTo c d p) (DenyAllFromTo a b) l" by (simp) (metis eq_imp_le pos_noteq) - lemma smalleraux2b: "{a,b} \ set l \ {c,d} \ set l \ {a,b} \ {c,d} \ y = DenyAllFromTo a b \ smaller (AllowPortFromTo c d p) y l \ \ smaller y (AllowPortFromTo c d p) l" by (simp) (metis eq_imp_le pos_noteq) - lemma smalleraux2c: "{a,b} \ set l\{c,d}\set l\{a,b} \ {c,d} \ y = AllowPortFromTo a b q \ -smaller (AllowPortFromTo c d p) y l \ \ smaller y (AllowPortFromTo c d p) l" + smaller (AllowPortFromTo c d p) y l \ \ smaller y (AllowPortFromTo c d p) l" by (simp) (metis pos_noteq) lemma smalleraux3: @@ -1684,25 +1568,22 @@ next case (Conc c d) thus ?thesis using assms by simp qed -thm Combinators.split - lemma smalleraux3a: "a \ DenyAll \ b \ DenyAll \ in_list b l \ in_list a l \ bothNet a \ bothNet b \ smaller a b l \ singleCombinators [a] \ singleCombinators [b] \ \ smaller b a l" apply (rule smalleraux3,simp_all) - apply (case_tac a, simp_all) + apply (case_tac a, simp_all) apply (case_tac b, simp_all) done lemma posaux[rule_format]: "position a l < position b l \ a \ b" - by (induct l) simp_all + by (induct l, simp_all) lemma posaux6[rule_format]: "a \ set l \ b \ set l \ a \ b \ position a l \ position b l" by (induct l) (simp_all add: position_positive) - lemma notSmallerTransaux[rule_format]: "x \ DenyAll \ r \ DenyAll \ singleCombinators [x] \ singleCombinators [y] \ singleCombinators [r] \ @@ -1710,22 +1591,21 @@ lemma notSmallerTransaux[rule_format]: in_list x l \ in_list y l \ in_list r l \ \ smaller r x l" by (metis order_trans) - lemma notSmallerTrans[rule_format]: "x \ DenyAll \ r \ DenyAll \ singleCombinators (x#y#z) \ \ smaller y x l \ sorted (x#y#z) l \ r \ set z \ all_in_list (x#y#z) l \ \ smaller r x l" apply (rule impI)+ apply (rule notSmallerTransaux, simp_all) - apply (metis singleCombinatorsConc singleCombinatorsStart) - apply (metis SCSubset equalityE remdups.simps(2) set_remdups - singleCombinatorsConc singleCombinatorsStart) - apply metis - apply (metis sorted.simps(3) in_set_in_list singleCombinatorsConc - singleCombinatorsStart sortedConcStart sorted_is_smaller) - apply (metis sorted_Cons all_in_list.simps(2) - singleCombinatorsConc) - apply (metis,metis in_set_in_list) + apply (metis singleCombinatorsConc singleCombinatorsStart) + apply (metis SCSubset equalityE remdups.simps(2) set_remdups + singleCombinatorsConc singleCombinatorsStart) + apply metis + apply (metis sorted.simps(3) in_set_in_list singleCombinatorsConc + singleCombinatorsStart sortedConcStart sorted_is_smaller) + apply (metis sorted_Cons all_in_list.simps(2) + singleCombinatorsConc) + apply (metis,metis in_set_in_list) done lemma NCSaux1[rule_format]: @@ -1738,7 +1618,7 @@ next case (Cons a list) then show ?thesis apply simp apply (intro impI conjI) - apply (metis bothNet.simps(2) first_bothNet.simps(3)) + apply (metis bothNet.simps(2) first_bothNet.simps(3)) proof - assume 1: "{x, y} \ set l" and 2: "in_list a l \ all_in_list list l" and 3 : "singleCombinators (a # list)" @@ -1756,7 +1636,7 @@ next rule_tac y = "a" and z = "list" in notSmallerTrans, simp_all del: smaller.simps) apply (rule smalleraux3a,simp_all del: smaller.simps) - apply (case_tac a, simp_all del: smaller.simps) + apply (case_tac a, simp_all del: smaller.simps) by (metis aux0_0 first_bothNet.elims list.set_intros(1)) show " {x, y} \ first_bothNet a \ {x, y} \ {u, v}" using 3 "*" "**" by force @@ -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 \ {a, b} \ set l \ all_in_list p l \singleCombinators p \ sorted (DenyAllFromTo a b # p) l \ {a, b} \ firstList p \ @@ -1822,45 +1699,43 @@ next in_list aa l \ all_in_list list l \ in_list (AllowPortFromTo u v w) l" apply (rule_tac p = list in in_set_in_list) - apply simp_all + apply simp_all done assume "p = aa # list" then show ?thesis apply simp apply (intro impI conjI,hypsubst, simp) apply (subgoal_tac "smaller (DenyAllFromTo a b) (AllowPortFromTo u v w) l") - apply (subgoal_tac "\ smaller (AllowPortFromTo u v w) (DenyAllFromTo a b) l") - apply (rule_tac l = l in posaux) - apply (rule_tac y = "position (first_bothNet aa) l" in basic_trans_rules(22)) - apply (simp_all split: if_splits) - apply (case_tac aa, simp_all) - apply (case_tac "a = x21 \ b = x22", simp_all) - apply (case_tac "a = x21", simp_all) - apply (simp add: order.not_eq_order_implies_strict posaux6) - apply (simp add: order.not_eq_order_implies_strict posaux6) - apply (simp add: order.not_eq_order_implies_strict posaux6) - apply (rule basic_trans_rules(18)) - apply (rule_tac a = "DenyAllFromTo a b" and b = aa in posaux4, simp_all) - apply (case_tac aa,simp_all) - apply (case_tac aa, simp_all) - apply (rule posaux3, simp_all) - apply (case_tac aa, simp_all) - apply (rule_tac a = aa and b = "AllowPortFromTo u v w" in posaux4, simp_all) - apply (case_tac aa,simp_all) - apply (rule_tac p = list in sorted_is_smaller, simp_all) - apply (case_tac aa, simp_all) - apply (case_tac aa, simp_all) - apply (rule_tac a = aa and b = "AllowPortFromTo u v w" in posaux4, simp_all) - apply (case_tac aa,simp_all) + apply (subgoal_tac "\ smaller (AllowPortFromTo u v w) (DenyAllFromTo a b) l") + apply (rule_tac l = l in posaux) + apply (rule_tac y = "position (first_bothNet aa) l" in basic_trans_rules(22)) + apply (simp_all split: if_splits) + apply (case_tac aa, simp_all) + apply (case_tac "a = x21 \ b = x22", simp_all) + apply (case_tac "a = x21", simp_all) + apply (simp add: order.not_eq_order_implies_strict posaux6) + apply (simp add: order.not_eq_order_implies_strict posaux6) + apply (simp add: order.not_eq_order_implies_strict posaux6) + apply (rule basic_trans_rules(18)) + apply (rule_tac a = "DenyAllFromTo a b" and b = aa in posaux4, simp_all) + apply (case_tac aa,simp_all) + apply (case_tac aa, simp_all) + apply (rule posaux3, simp_all) + apply (case_tac aa, simp_all) + apply (rule_tac a = aa and b = "AllowPortFromTo u v w" in posaux4, simp_all) + apply (case_tac aa,simp_all) + apply (rule_tac p = list in sorted_is_smaller, simp_all) + apply (case_tac aa, simp_all) + apply (case_tac aa, simp_all) + apply (rule_tac a = aa and b = "AllowPortFromTo u v w" in posaux4, simp_all) + apply (case_tac aa,simp_all) using ** apply auto[1] - apply (metis all_in_list.simps(2) sorted_Cons) - apply (case_tac aa, simp_all) - apply (metis ** bothNet.simps(3) in_list.simps(3) posaux6) + apply (metis all_in_list.simps(2) sorted_Cons) + apply (case_tac aa, simp_all) + apply (metis ** bothNet.simps(3) in_list.simps(3) posaux6) using * by force qed - - lemma NCSaux3[rule_format]: "noDenyAll p \ {a, b} \ set l \ all_in_list p l \singleCombinators p \ sorted (AllowPortFromTo a b w # p) l \ {a, b} \ firstList p \ @@ -1874,31 +1749,30 @@ proof - and 6 : "{a, b} \ first_bothNet aa" and 7: "DenyAllFromTo u v \ set list" have *: "\ smaller (DenyAllFromTo u v) (AllowPortFromTo a b w) l" apply (insert 1 2 3 4 5 6 7, rule_tac y = aa and z = list in notSmallerTrans) - apply (simp_all del: smaller.simps) + apply (simp_all del: smaller.simps) apply (rule smalleraux3a,simp_all del: smaller.simps) - apply (case_tac aa, simp_all del: smaller.simps) + apply (case_tac aa, simp_all del: smaller.simps) apply (case_tac aa, simp_all del: smaller.simps) done have **: "smaller (AllowPortFromTo a b w) (DenyAllFromTo u v) l" apply (insert 1 2 3 4 5 6 7,rule_tac y = aa in order_trans,simp_all del: smaller.simps) - apply (subgoal_tac "in_list (DenyAllFromTo u v) l", simp) - apply (rule_tac p = list in in_set_in_list, simp_all) + apply (subgoal_tac "in_list (DenyAllFromTo u v) l", simp) + apply (rule_tac p = list in in_set_in_list, simp_all) apply (rule_tac p = list in sorted_is_smaller,simp_all del: smaller.simps) - apply (subgoal_tac "in_list (DenyAllFromTo u v) l", simp) + apply (subgoal_tac "in_list (DenyAllFromTo u v) l", simp) apply (rule_tac p = list in in_set_in_list, simp_all) apply (erule singleCombinatorsConc) done show "{a, b} \ {u, v}" by (insert * **, simp split: if_splits) qed - lemma NCSaux4[rule_format]: "noDenyAll p \ {a, b} \ set l \ all_in_list p l \ singleCombinators p \ sorted (AllowPortFromTo a b c # p) l \ {a, b} \ firstList p \ AllowPortFromTo u v w \ set p \ {a, b} \ {u, v}" apply (cases p, simp_all) apply (intro impI conjI) - apply (hypsubst,simp_all) + apply (hypsubst,simp_all) proof - fix aa::"('a, 'b) Combinators" fix list::"('a, 'b) Combinators list" assume 1 : "\ member DenyAll aa \ noDenyAll list" and 2: "{a, b} \ set l " @@ -1907,34 +1781,33 @@ proof - and 6 : "{a, b} \ first_bothNet aa" and 7: "AllowPortFromTo u v w \ set list" have *: "\ smaller (AllowPortFromTo u v w) (AllowPortFromTo a b c) l" apply (insert 1 2 3 4 5 6 7, rule_tac y = aa and z = list in notSmallerTrans) - apply (simp_all del: smaller.simps) + apply (simp_all del: smaller.simps) apply (rule smalleraux3a,simp_all del: smaller.simps) - apply (case_tac aa, simp_all del: smaller.simps) + apply (case_tac aa, simp_all del: smaller.simps) apply (case_tac aa, simp_all del: smaller.simps) done have **: "smaller (AllowPortFromTo a b c) (AllowPortFromTo u v w) l" apply(insert 1 2 3 4 5 6 7) apply (case_tac aa, simp_all del: smaller.simps) + apply (rule_tac y = aa in order_trans,simp_all del: smaller.simps) + apply (subgoal_tac "in_list (AllowPortFromTo u v w) l", simp) + apply (rule_tac p = list in in_set_in_list, simp) + apply (case_tac aa, simp_all del: smaller.simps) + apply (rule_tac p = list in sorted_is_smaller,simp_all del: smaller.simps) + apply (subgoal_tac "in_list (AllowPortFromTo u v w) l", simp) + apply (rule_tac p = list in in_set_in_list, simp, simp) apply (rule_tac y = aa in order_trans,simp_all del: smaller.simps) - apply (subgoal_tac "in_list (AllowPortFromTo u v w) l", simp) - apply (rule_tac p = list in in_set_in_list, simp) - apply (case_tac aa, simp_all del: smaller.simps) - apply (rule_tac p = list in sorted_is_smaller,simp_all del: smaller.simps) - apply (subgoal_tac "in_list (AllowPortFromTo u v w) l", simp) - apply (rule_tac p = list in in_set_in_list, simp, simp) - apply (rule_tac y = aa in order_trans,simp_all del: smaller.simps) - apply (subgoal_tac "in_list (AllowPortFromTo u v w) l", simp) + apply (subgoal_tac "in_list (AllowPortFromTo u v w) l", simp) using in_set_in_list apply blast by (metis all_in_list.simps(2) bothNet.simps(3) in_list.simps(3) singleCombinators.simps(5) sorted_ConsA) show "{a, b} \ {u, v}" by (insert * **, simp_all split: if_splits) qed - - + lemma NetsCollectedSorted[rule_format]: "noDenyAll1 p \ all_in_list p l \ singleCombinators p \ sorted p l \ NetsCollected p" apply (induct p) - apply simp + apply simp apply (intro impI,drule mp,erule noDA1C,drule mp,simp) apply (drule mp,erule singleCombinatorsConc) apply (drule mp,erule sortedConc) @@ -1944,7 +1817,7 @@ proof - and 3: "singleCombinators (a # p)" and 4: "sorted (a # p) l" and 5: "NetsCollected p" show "NetsCollected (a # p)" apply(insert 1 2 3 4 5, rule fl3) - apply(simp, rename_tac "aa") + apply(simp, rename_tac "aa") proof (cases a) case DenyAll fix aa::"('a, 'b) Combinators" @@ -1959,8 +1832,8 @@ proof - show "first_bothNet a \ first_bothNet aa" apply(insert 1 2 3 4 5 6 7 `a = DenyAllFromTo x21 x22`) apply(case_tac aa, simp_all) - apply (meson NCSaux1) - apply (meson NCSaux2) + apply (meson NCSaux1) + apply (meson NCSaux2) using SCnotConc by auto[1] next case (AllowPortFromTo x31 x32 x33) @@ -1969,8 +1842,8 @@ proof - show "first_bothNet a \ first_bothNet aa" apply(insert 1 2 3 4 6 7 `a = AllowPortFromTo x31 x32 x33`) apply(case_tac aa, simp_all) - apply (meson NCSaux3) - apply (meson NCSaux4) + apply (meson NCSaux3) + apply (meson NCSaux4) using SCnotConc by auto next case (Conc x41 x42) @@ -1980,37 +1853,28 @@ proof - qed qed - - lemma NetsCollectedSort: "distinct p \noDenyAll1 p \ all_in_list p l \ singleCombinators p \ NetsCollected (sort p l)" apply (rule_tac l = l in NetsCollectedSorted,rule noDAsort, simp_all) - apply (rule_tac b=p in all_in_listSubset) + apply (rule_tac b=p in all_in_listSubset) by (auto intro: sort_is_sorted) - - lemma fBNsep[rule_format]: "(\a\set z. {b,c} \ first_bothNet a) \ (\a\set (separate z). {b,c} \ first_bothNet a)" apply (induct z rule: separate.induct, simp) by (rule impI, simp)+ - - lemma fBNsep1[rule_format]: " (\a\set z. first_bothNet x \ first_bothNet a) \ (\a\set (separate z). first_bothNet x \ first_bothNet a)" apply (induct z rule: separate.induct, simp) by (rule impI, simp)+ - - lemma NetsCollectedSepauxa: "{b,c}\firstList z \ noDenyAll1 z \ \a\set z. {b,c}\first_bothNet a \ NetsCollected z \ NetsCollected (separate z) \ {b, c} \ firstList (separate z) \ a \ set (separate z) \ {b, c} \ first_bothNet a" by (rule fBNsep) simp_all - lemma NetsCollectedSepaux: "first_bothNet (x::('a,'b)Combinators) \ first_bothNet y \ \ member DenyAll y \ noDenyAll z \ (\a\set z. first_bothNet x \ first_bothNet a) \ NetsCollected (y # z) \ @@ -2030,13 +1894,13 @@ next fix v va y fix z::"('a, 'b) Combinators list" case 2 then show ?case apply (intro conjI impI, simp) - apply (metis NetsCollectedSepaux fl3 noDA1eq noDenyAll.simps(1)) + apply (metis NetsCollectedSepaux fl3 noDA1eq noDenyAll.simps(1)) by (metis noDA1eq noDenyAll.simps(1)) next fix v va vb y fix z::"('a, 'b) Combinators list" case 3 then show ?case apply (intro conjI impI) - apply (metis NetsCollectedSepaux fl3 noDA1eq noDenyAll.simps(1)) + apply (metis NetsCollectedSepaux fl3 noDA1eq noDenyAll.simps(1)) by (metis noDA1eq noDenyAll.simps(1)) next fix v va y fix z::"('a, 'b) Combinators list" @@ -2044,14 +1908,13 @@ next by (metis NetsCollectedSepaux fl3 noDA1eq noDenyAll.simps(1)) qed - lemma OTNaux: "onlyTwoNets a \ \ member DenyAll a \ (x,y) \ sdnets a \ (x = first_srcNet a \ y = first_destNet a) \ (x = first_destNet a \ y = first_srcNet a)" apply (case_tac "(x = first_srcNet a \ y = first_destNet a)",simp_all add: onlyTwoNets_def) apply (case_tac "(\aa b. sdnets a = {(aa, b)})", simp_all) - apply (subgoal_tac "sdnets a = {(first_srcNet a,first_destNet a)}", simp_all) - apply (metis singletonE first_isIn) + apply (subgoal_tac "sdnets a = {(first_srcNet a,first_destNet a)}", simp_all) + apply (metis singletonE first_isIn) apply (subgoal_tac"sdnets a = {(first_srcNet a,first_destNet a),(first_destNet a, first_srcNet a)}") by(auto intro!: sdnets2) @@ -2060,7 +1923,7 @@ sdnets a = {(first_srcNet a,first_destNet a)} \ sdnets a = {(first_srcNet a, first_destNet a),(first_destNet a, first_srcNet a)}" apply (case_tac "sdnets a = {(first_srcNet a, first_destNet a)}", simp_all add: onlyTwoNets_def) apply (case_tac "(\aa b. sdnets a = {(aa, b)})", simp_all) - apply (metis singletonE first_isIn) + apply (metis singletonE first_isIn) apply (subgoal_tac "sdnets a = {(first_srcNet a,first_destNet a),(first_destNet a,first_srcNet a)}") by( auto intro!: sdnets2) @@ -2077,9 +1940,7 @@ lemma sdnets_noteq: apply (insert first_bothNet_charn [of a]) apply (insert first_bothNet_charn [of aa]) by(metis OTNaux first_isIn insert_absorb2 insert_commute) - - - + lemma fbn_noteq: "onlyTwoNets a \ onlyTwoNets aa \ first_bothNet a \ first_bothNet aa \ \ member DenyAll a \ \ member DenyAll aa \ allNetsDistinct [a, aa] \ @@ -2087,8 +1948,7 @@ lemma fbn_noteq: first_destNet a \ first_srcNet aa \ first_destNet a \ first_destNet aa" apply (insert sdnets_charn [of a]) apply (insert sdnets_charn [of aa]) - by (metis first_bothNet_charn) - + by (metis first_bothNet_charn) lemma NCisSD2aux: assumes 1: "onlyTwoNets a" and 2 : "onlyTwoNets aa" and 3 : "first_bothNet a \ first_bothNet aa" @@ -2117,36 +1977,35 @@ proof - show "netsDistinct ab c \ netsDistinct b d" apply(insert 7 8 9 10 11 12) apply (cases "sdnets a = {(first_srcNet a, first_destNet a)}") - apply (cases "sdnets aa = {(first_srcNet aa, first_destNet aa)}", simp_all) - apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) + apply (cases "sdnets aa = {(first_srcNet aa, first_destNet aa)}", simp_all) + apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) + apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) + apply (case_tac "(first_srcNet a) \ (first_srcNet aa)",simp_all) + apply (metis 4 5 firstInNeta alternativelistconc2) + apply (subgoal_tac "first_destNet a \ first_destNet aa") + apply (metis 4 5 firstInNet alternativelistconc2) + apply (metis 3 4 5 first_bothNetsd) + apply (case_tac "(first_destNet aa) \ (first_srcNet a)",simp_all) + apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) + apply (case_tac "first_destNet aa \ first_destNet a",simp_all) + apply (subgoal_tac "first_srcNet aa \ first_destNet a") + apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) + apply (metis 3 4 5 first_bothNetsd insert_commute) + apply (metis 5 firstInNeta firstInNet alternativelistconc2) apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) - apply (case_tac "(first_srcNet a) \ (first_srcNet aa)",simp_all) - apply (metis 4 5 firstInNeta alternativelistconc2) - apply (subgoal_tac "first_destNet a \ first_destNet aa") - apply (metis 4 5 firstInNet alternativelistconc2) - apply (metis 3 4 5 first_bothNetsd) - apply (case_tac "(first_destNet aa) \ (first_srcNet a)",simp_all) - apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) - apply (case_tac "first_destNet aa \ first_destNet a",simp_all) - apply (subgoal_tac "first_srcNet aa \ first_destNet a") - apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) - apply (metis 3 4 5 first_bothNetsd insert_commute) - apply (metis 5 firstInNeta firstInNet alternativelistconc2) - apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) - apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) - apply (case_tac "(first_srcNet a) \ (first_srcNet aa)",simp_all) - apply (metis 4 5 firstInNeta alternativelistconc2) - apply (subgoal_tac "first_destNet a \ first_destNet aa") - apply (metis 4 5 firstInNet alternativelistconc2) - apply (metis 3 4 5 first_bothNetsd ) - apply (case_tac "(first_destNet aa) \ (first_srcNet a)",simp_all) - apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) - apply (case_tac "first_destNet aa \ first_destNet a", simp) - apply (subgoal_tac "first_srcNet aa \ first_destNet a") - apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) - apply (metis 3 4 5 first_bothNetsd insert_commute) - apply (metis) - + apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) + apply (case_tac "(first_srcNet a) \ (first_srcNet aa)",simp_all) + apply (metis 4 5 firstInNeta alternativelistconc2) + apply (subgoal_tac "first_destNet a \ first_destNet aa") + apply (metis 4 5 firstInNet alternativelistconc2) + apply (metis 3 4 5 first_bothNetsd ) + apply (case_tac "(first_destNet aa) \ (first_srcNet a)",simp_all) + apply (metis 4 5 firstInNeta firstInNet alternativelistconc2) + apply (case_tac "first_destNet aa \ first_destNet a", simp) + apply (subgoal_tac "first_srcNet aa \ first_destNet a") + 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 \ b = first_destNet a \ ab = first_destNet a \ b = first_srcNet a) \ (c, d) \ sdnets aa " and 15 : "sdnets a = {(first_srcNet a, first_destNet a), (first_destNet a, first_srcNet a)} " @@ -2156,29 +2015,28 @@ proof - and 19 : "first_destNet a \ first_srcNet a" and 20 : "c = first_srcNet aa \ d \ first_destNet aa" show " netsDistinct ab c \ netsDistinct b d" - apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)",simp_all) + apply (case_tac "c = first_srcNet aa", simp_all) + apply (metis 2 5 14 20 OTNaux) + apply (subgoal_tac "c = first_destNet aa", simp) + apply (subgoal_tac "d = first_srcNet aa", simp) + apply (case_tac "(first_srcNet a) \ (first_destNet aa)",simp_all) + apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) + apply (subgoal_tac "first_destNet a \ first_srcNet aa") + apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) + apply (metis 3 4 5 first_bothNetsd insert_commute) + apply (metis 2 5 14 OTNaux) + apply (metis 2 5 14 OTNaux) apply (case_tac "c = first_srcNet aa", simp_all) - apply (metis 2 5 14 20 OTNaux) + apply (metis 2 5 14 20 OTNaux) apply (subgoal_tac "c = first_destNet aa", simp) - apply (subgoal_tac "d = first_srcNet aa", simp) - apply (case_tac "(first_srcNet a) \ (first_destNet aa)",simp_all) - apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) - apply (subgoal_tac "first_destNet a \ first_srcNet aa") - apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) - apply (metis 3 4 5 first_bothNetsd insert_commute) - apply (metis 2 5 14 OTNaux) - apply (metis 2 5 14 OTNaux) - apply (case_tac "c = first_srcNet aa", simp_all) - apply (metis 2 5 14 20 OTNaux) - apply (subgoal_tac "c = first_destNet aa", simp) - apply (subgoal_tac "d = first_srcNet aa", simp) - apply (case_tac "(first_destNet a) \ (first_destNet aa)",simp_all) - apply (metis 4 5 7 14 firstInNet alternativelistconc2) - apply (subgoal_tac "first_srcNet a \ first_srcNet aa") - apply (metis 4 5 7 14 firstInNeta alternativelistconc2) - apply (metis 3 4 5 first_bothNetsd insert_commute) - apply (metis 2 5 14 OTNaux) + apply (subgoal_tac "d = first_srcNet aa", simp) + apply (case_tac "(first_destNet a) \ (first_destNet aa)",simp_all) + apply (metis 4 5 7 14 firstInNet alternativelistconc2) + apply (subgoal_tac "first_srcNet a \ first_srcNet aa") + apply (metis 4 5 7 14 firstInNeta alternativelistconc2) + apply (metis 3 4 5 first_bothNetsd insert_commute) + apply (metis 2 5 14 OTNaux) apply (metis 2 5 14 OTNaux) done qed @@ -2186,25 +2044,25 @@ proof - show "netsDistinct ab d \ netsDistinct b c" apply (insert 1 2 3 4 5 6 7 8 9 10 11 12) apply (cases "sdnets a = {(first_srcNet a, first_destNet a)}") - apply (cases "sdnets aa = {(first_srcNet aa, first_destNet aa)}", simp_all) - apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) - apply (case_tac "(first_srcNet a) \ (first_destNet aa)", simp_all) - apply (metis firstInNeta firstInNet alternativelistconc2) - apply (subgoal_tac "first_destNet a \ first_srcNet aa") - apply (metis firstInNeta firstInNet alternativelistconc2) - apply (metis first_bothNetsd insert_commute) - apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) - apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) - apply (case_tac "(first_destNet a) \ (first_srcNet aa)",simp_all) - apply (metis firstInNeta firstInNet alternativelistconc2) - apply (subgoal_tac "first_srcNet a \ first_destNet aa") - apply (metis firstInNeta firstInNet alternativelistconc2) - apply (metis first_bothNetsd insert_commute) - apply (case_tac "(first_srcNet aa) \ (first_srcNet a)",simp_all) - apply (metis firstInNeta alternativelistconc2) - apply (case_tac "first_destNet aa \ first_destNet a",simp_all) - apply (metis firstInNet alternativelistconc2) - apply (metis first_bothNetsd) + apply (cases "sdnets aa = {(first_srcNet aa, first_destNet aa)}", simp_all) + apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) + apply (case_tac "(first_srcNet a) \ (first_destNet aa)", simp_all) + apply (metis firstInNeta firstInNet alternativelistconc2) + apply (subgoal_tac "first_destNet a \ first_srcNet aa") + apply (metis firstInNeta firstInNet alternativelistconc2) + apply (metis first_bothNetsd insert_commute) + apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) + apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) + apply (case_tac "(first_destNet a) \ (first_srcNet aa)",simp_all) + apply (metis firstInNeta firstInNet alternativelistconc2) + apply (subgoal_tac "first_srcNet a \ first_destNet aa") + apply (metis firstInNeta firstInNet alternativelistconc2) + apply (metis first_bothNetsd insert_commute) + apply (case_tac "(first_srcNet aa) \ (first_srcNet a)",simp_all) + apply (metis firstInNeta alternativelistconc2) + apply (case_tac "first_destNet aa \ first_destNet a",simp_all) + apply (metis firstInNet alternativelistconc2) + apply (metis first_bothNetsd) proof - assume 13:" \ab b. ab \ b \ ab\set(net_list_aux[a,aa]) \ b \ set(net_list_aux[a,aa]) \ netsDistinct ab b " @@ -2220,42 +2078,42 @@ proof - show "first_destNet a \ first_srcNet a \ netsDistinct ab d \ netsDistinct b c" apply (insert 1 2 3 4 5 6 13 14 15 16 17) apply (cases "sdnets aa = {(first_srcNet aa, first_destNet aa)}", simp_all) - apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) - apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) - apply (case_tac "(first_destNet a) \ (first_srcNet aa)",simp_all) - apply (metis firstInNeta firstInNet alternativelistconc2) - apply (subgoal_tac "first_srcNet a \ first_destNet aa") - apply (metis firstInNeta firstInNet alternativelistconc2) - apply (metis first_bothNetsd insert_commute) - apply (case_tac "(first_srcNet aa) \ (first_srcNet a)",simp_all) - apply (metis firstInNeta alternativelistconc2) - apply (case_tac "first_destNet aa \ first_destNet a",simp_all) - apply (metis firstInNet alternativelistconc2) - apply (metis first_bothNetsd) + apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) + apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) + apply (case_tac "(first_destNet a) \ (first_srcNet aa)",simp_all) + apply (metis firstInNeta firstInNet alternativelistconc2) + apply (subgoal_tac "first_srcNet a \ first_destNet aa") + apply (metis firstInNeta firstInNet alternativelistconc2) + apply (metis first_bothNetsd insert_commute) + apply (case_tac "(first_srcNet aa) \ (first_srcNet a)",simp_all) + apply (metis firstInNeta alternativelistconc2) + apply (case_tac "first_destNet aa \ first_destNet a",simp_all) + apply (metis firstInNet alternativelistconc2) + apply (metis first_bothNetsd) proof - assume 20: "{(first_srcNet a, first_destNet a), (first_destNet a, first_srcNet a)} \ {(first_srcNet aa, first_destNet aa), (first_destNet aa, first_srcNet aa)}" and 21: "first_destNet a \ first_srcNet a" show "netsDistinct ab d \ netsDistinct b c" apply (case_tac "(c = first_srcNet aa \ d = first_destNet aa)", simp_all) - apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) - apply (case_tac "(first_destNet a) \ (first_srcNet aa)", simp_all) - apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) - apply (subgoal_tac "first_srcNet a \ first_destNet aa") - apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) - apply (metis 20 insert_commute) - apply (case_tac "(first_srcNet aa) \ (first_srcNet a)", simp_all) - apply (metis 4 5 13 14 firstInNeta alternativelistconc2) - apply (case_tac "first_destNet aa \ first_destNet a", simp_all) - apply (metis 4 5 13 14 firstInNet alternativelistconc2) - apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) - apply (case_tac "(first_destNet a) \ (first_srcNet aa)", simp_all) - apply (metis 20) + apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) + apply (case_tac "(first_destNet a) \ (first_srcNet aa)", simp_all) + apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) + apply (subgoal_tac "first_srcNet a \ first_destNet aa") + apply (metis 4 5 7 firstInNeta firstInNet alternativelistconc2) + apply (metis 20 insert_commute) + apply (case_tac "(first_srcNet aa) \ (first_srcNet a)", simp_all) + apply (metis 4 5 13 14 firstInNeta alternativelistconc2) + apply (case_tac "first_destNet aa \ first_destNet a", simp_all) + apply (metis 4 5 13 14 firstInNet alternativelistconc2) + apply (case_tac "(ab = first_srcNet a \ b = first_destNet a)", simp_all) + apply (case_tac "(first_destNet a) \ (first_srcNet aa)", simp_all) + apply (metis 20) apply (subgoal_tac "first_srcNet a \ first_srcNet aa") - apply (metis 20) - apply (metis 21) + apply (metis 20) + apply (metis 21) apply (case_tac "(first_srcNet aa) \ (first_destNet a)") - apply (metis (no_types, lifting) 2 3 4 5 7 14 OTNaux + apply (metis (no_types, lifting) 2 3 4 5 7 14 OTNaux firstInNet firstInNeta first_bothNetsd isInAlternativeList) by (metis 2 4 5 7 20 14 OTNaux doubleton_eq_iff firstInNet firstInNeta isInAlternativeList) @@ -2268,14 +2126,11 @@ lemma ANDaux3[rule_format]: "y \ set xs \ a \ set (net_list_aux [y]) \ a \ set (net_list_aux xs)" by (induct xs) (simp_all add: isInAlternativeList) - lemma ANDaux2: "allNetsDistinct (x # xs) \ y \ set xs \ allNetsDistinct [x,y]" apply (simp add: allNetsDistinct_def) by (meson ANDaux3 isInAlternativeList netlistaux) - - lemma NCisSD2[rule_format]: "\ member DenyAll a \ OnlyTwoNets (a#p) \ NetsCollected2 (a # p) \ NetsCollected (a#p) \ @@ -2297,10 +2152,10 @@ proof (induct p, simp_all, case_tac "a = DenyAll", simp_all, goal_cases) noDenyAll1 (a # p) \ allNetsDistinct (a # p) \ (\s. s \ set p \ disjSD_2 a s) \ separated p" apply (intro impI,drule mp, erule OTNConc,drule mp) - apply (case_tac p, simp_all) + apply (case_tac p, simp_all) apply (drule mp,erule noDA1C, intro conjI allI impI NCisSD2, simp_all) - apply (case_tac a, simp_all) - apply (case_tac a, simp_all) + apply (case_tac a, simp_all) + apply (case_tac a, simp_all) using ANDConc by auto next fix a::"('a set set,'b) Combinators " fix p ::"('a set set,'b) Combinators list" @@ -2328,7 +2183,7 @@ next (\s. s \ set p \ disjSD_2 a s) \ separated p" apply(insert Cons.hyps `a = DenyAll`) apply (intro impI,drule mp, erule OTNConc,drule mp) - apply (case_tac p, simp_all) + apply (case_tac p, simp_all) apply (case_tac a, simp_all) apply (case_tac a, simp_all) by (simp add: ANDConc disjSD_2_def noDA1eq) @@ -2343,7 +2198,6 @@ next qed qed - lemma NC2Sep[rule_format]: "noDenyAll1 p \ 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]: @@ -2371,13 +2225,7 @@ lemma separatedSep[rule_format]: lemma rADnMT[rule_format]: "p \ [] \ removeAllDuplicates p \ []" 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 \ [] \ remdups p \ []" by (metis remdups_eq_nil_iff) @@ -2392,6 +2240,5 @@ lemma sets_distinct5: "(n::int) < m \ {(a,b). a = n} \ {( lemma sets_distinct6: "(m::int) < n \ {(a,b). a = n} \ {(a,b). a = m}" by auto - end \ No newline at end of file diff --git a/FWNormalisation/NormalisationIPPProofs.thy b/FWNormalisation/NormalisationIPPProofs.thy index 98b4abe..5aed764 100644 --- a/FWNormalisation/NormalisationIPPProofs.thy +++ b/FWNormalisation/NormalisationIPPProofs.thy @@ -36,131 +36,124 @@ *****************************************************************************) 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 \ B) \ D) = Cp(A \ (B \ D))" -by (simp add: Cp.simps) - + by (simp add: Cp.simps) lemma aux26[simp]: "twoNetsDistinct a b c d \ dom (Cp (AllowPortFromTo a b p)) \ dom (Cp (DenyAllFromTo c d)) = {}" -by(auto simp:twoNetsDistinct_def netsDistinct_def PLemmas, auto) - + by(auto simp:twoNetsDistinct_def netsDistinct_def PLemmas, auto) lemma wp2_aux[rule_format]: "wellformed_policy2Pr (xs @ [x]) \ wellformed_policy2Pr xs" -by (induct xs, simp_all) (case_tac "a", simp_all) + by (induct xs, simp_all) (case_tac "a", simp_all) lemma Cdom2: "x \ dom(Cp b) \ Cp (a \ b) x = (Cp b) x" -by (auto simp: Cp.simps) + by (auto simp: Cp.simps) lemma wp2Conc[rule_format]: "wellformed_policy2Pr (x#xs) \ wellformed_policy2Pr xs" -by (case_tac "x",simp_all) - + by (case_tac "x",simp_all) + lemma DAimpliesMR_E[rule_format]: "DenyAll \ set p \ (\ r. applied_rule_rev Cp x p = Some r)" -apply (simp add: applied_rule_rev_def) -apply (rule_tac xs = p in rev_induct, simp_all) -by (metis Cp.simps(1) denyAllDom) - - - + 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 \ set p \ applied_rule_rev Cp x p \ None" -by (auto intro: DAimpliesMR_E) - + by (auto intro: DAimpliesMR_E) + lemma MRList1[rule_format]: "x \ dom (Cp a) \ applied_rule_rev Cp x (b@[a]) = Some a" -by (simp add: applied_rule_rev_def) - + by (simp add: applied_rule_rev_def) + lemma MRList2: "x \ dom (Cp a) \ applied_rule_rev Cp x (c@b@[a]) = Some a" -by (simp add: applied_rule_rev_def) - + by (simp add: applied_rule_rev_def) + lemma MRList3: "x \ dom(Cp xa) \ applied_rule_rev Cp x (a@b#xs@[xa]) = applied_rule_rev Cp x (a @ b # xs)" -by (simp add: applied_rule_rev_def) - + by (simp add: applied_rule_rev_def) + lemma CConcEnd[rule_format]: "Cp a x = Some y \ Cp (list2FWpolicy (xs @ [a])) x = Some y" (is "?P xs") -apply (rule_tac P = ?P in list2FWpolicy.induct) -by (simp_all add:Cp.simps) - + apply (rule_tac P = ?P in list2FWpolicy.induct) + by (simp_all add:Cp.simps) + lemma CConcStartaux: "Cp a x = None \ (Cp aa ++ Cp a) x = Cp aa x" -by (simp add: PLemmas) - + by (simp add: PLemmas) + lemma CConcStart[rule_format]: "xs \ [] \ Cp a x = None \ Cp (list2FWpolicy (xs @ [a])) x = Cp (list2FWpolicy xs) x" -by (rule list2FWpolicy.induct) (simp_all add: PLemmas) - - + by (rule list2FWpolicy.induct) (simp_all add: PLemmas) + lemma mrNnt[simp]: "applied_rule_rev Cp x p = Some a \ p \ []" -by (simp add: applied_rule_rev_def)(auto) - + by (simp add: applied_rule_rev_def)(auto) + lemma mr_is_C[rule_format]: "applied_rule_rev Cp x p = Some a \ Cp (list2FWpolicy (p)) x = Cp a x" -apply (simp add: applied_rule_rev_def) -apply (rule rev_induct, simp_all, safe) -apply (metis CConcEnd ) -apply (metis CConcEnd) -by (metis CConcStart applied_rule_rev_def mrNnt option.exhaust) - - + apply (simp add: applied_rule_rev_def) + apply (rule rev_induct, simp_all, safe) + apply (metis CConcEnd ) + apply (metis CConcEnd) + by (metis CConcStart applied_rule_rev_def mrNnt option.exhaust) + lemma CConcStart2: "p \ [] \ x \ dom (Cp a) \ Cp(list2FWpolicy (p@[a])) x = Cp (list2FWpolicy p)x" -by (erule CConcStart,simp add: PLemmas) - + by (erule CConcStart,simp add: PLemmas) + lemma CConcEnd1: "q@p \ [] \ x \ dom (Cp a) \ Cp(list2FWpolicy(q@p@[a])) x = Cp (list2FWpolicy (q@p))x" -by (subst lCdom2) (rule CConcStart2, simp_all) - + by (subst lCdom2) (rule CConcStart2, simp_all) + lemma CConcEnd2[rule_format]: "x \ dom (Cp a) \ Cp (list2FWpolicy (xs @ [a])) x = Cp a x" (is "?P xs") -by (rule_tac P = ?P in list2FWpolicy.induct) (auto simp:Cp.simps) - - + by (rule_tac P = ?P in list2FWpolicy.induct) (auto simp:Cp.simps) + lemma bar3: "x \ dom (Cp (list2FWpolicy (xs @ [xa]))) \ x \ dom (Cp (list2FWpolicy xs)) \ x \ dom (Cp xa)" -by auto (metis CConcStart eq_Nil_appendI l2p_aux2 option.simps(3)) - - + by auto (metis CConcStart eq_Nil_appendI l2p_aux2 option.simps(3)) + lemma CeqEnd[rule_format,simp]: "a \ [] \ x \ dom (Cp(list2FWpolicy a)) \ Cp(list2FWpolicy(b@a)) x = (Cp(list2FWpolicy a)) x" proof (induct rule: rev_induct)print_cases - case Nil show ?case by simp + case Nil show ?case by simp next - case (snoc xa xs) show ?case - apply (case_tac "xs \ []", simp_all) + case (snoc xa xs) show ?case + apply (case_tac "xs \ []", simp_all) apply (case_tac "x \ dom (Cp xa)") - apply (metis CConcEnd2 MRList2 mr_is_C ) + apply (metis CConcEnd2 MRList2 mr_is_C ) apply (metis snoc.hyps CConcEnd1 CConcStart2 Nil_is_append_conv bar3 ) - by (metis MRList2 eq_Nil_appendI mr_is_C ) + by (metis MRList2 eq_Nil_appendI mr_is_C ) qed - + lemma CConcStartA[rule_format,simp]: "x \ dom (Cp a) \ x \ dom (Cp (list2FWpolicy (a # b)))" (is "?P b") -by (rule_tac P = ?P in list2FWpolicy.induct) (simp_all add: Cp.simps) - - + by (rule_tac P = ?P in list2FWpolicy.induct) (simp_all add: Cp.simps) + lemma domConc: "x \ dom (Cp (list2FWpolicy b)) \ b \ [] \ x \ dom (Cp (list2FWpolicy (a@b)))" -by (auto simp: PLemmas) - + by (auto simp: PLemmas) + lemma CeqStart[rule_format,simp]: "x \ dom (Cp (list2FWpolicy a)) \ a \ [] \ b \ [] \ Cp (list2FWpolicy (b@a)) x = (Cp (list2FWpolicy b)) x" -by (rule list2FWpolicy.induct,simp_all) (auto simp: list2FWpolicyconc PLemmas) - - + by (rule list2FWpolicy.induct,simp_all) (auto simp: list2FWpolicyconc PLemmas) + lemma C_eq_if_mr_eq2: "applied_rule_rev Cp x a = Some r \ applied_rule_rev Cp x b = Some r \ a\[] \ b\[] \ (Cp (list2FWpolicy a)) x = (Cp (list2FWpolicy b)) x" -by (metis mr_is_C) - - + by (metis mr_is_C) + lemma nMRtoNone[rule_format]: - "p \ [] \ applied_rule_rev Cp x p = None \ Cp (list2FWpolicy p) x = None" + "p \ [] \ applied_rule_rev Cp x p = None \ Cp (list2FWpolicy p) x = None" proof (induct rule: rev_induct) print_cases case Nil show ?case by simp next @@ -169,19 +162,17 @@ 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 \ a \ [] \ b \ [] \ + "applied_rule_rev Cp x b = applied_rule_rev Cp x a \ a \ [] \ b \ [] \ (Cp (list2FWpolicy a)) x = (Cp (list2FWpolicy b)) x" -apply (cases "applied_rule_rev Cp x a = None", simp_all) -apply (subst nMRtoNone,simp_all) -apply (subst nMRtoNone,simp_all) -by (auto intro: C_eq_if_mr_eq2) - - + apply (cases "applied_rule_rev Cp x a = None", simp_all) + 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]) \ Some a \ x \ dom (Cp a)" -by (simp add: applied_rule_rev_def split: if_splits) + by (simp add: applied_rule_rev_def split: if_splits) lemma foo3a[rule_format]: "applied_rule_rev Cp x (a@[b]@c) = Some b \ r \ set c \ b \ set c \ x \ dom (Cp r)" @@ -191,197 +182,189 @@ next case (snoc xa xs) show ?case apply simp_all apply (rule impI|rule conjI|simp)+ - apply (rule_tac p = "a @ b # xs" in notmatching_notdom,simp_all) + apply (rule_tac p = "a @ b # xs" in notmatching_notdom,simp_all) by (metis Cons_eq_appendI NormalisationIPPProofs.MRList2 NormalisationIPPProofs.MRList3 - append_Nil option.inject snoc.hyps) + append_Nil option.inject snoc.hyps) qed lemma foo3D: "wellformed_policy1 p \ p=DenyAll#ps \ applied_rule_rev Cp x p = Some DenyAll \ r\set ps \ x \ dom (Cp r)" -by (rule_tac a = "[]" and b = "DenyAll" and c = "ps" in foo3a, simp_all) - + by (rule_tac a = "[]" and b = "DenyAll" and c = "ps" in foo3a, simp_all) + lemma foo4[rule_format]: "set p = set s \ (\ r. r \ set p \ x \ dom (Cp r)) \ (\ r .r \ set s \ x \ dom (Cp r))" -by simp - + by simp + lemma foo5b[rule_format]: "x \ dom (Cp b) \ (\ r. r \ set c \ x \ dom (Cp r))\ applied_rule_rev Cp x (b#c) = Some b" -apply (simp add: applied_rule_rev_def) -apply (rule_tac xs = c in rev_induct, simp_all) -done + apply (simp add: applied_rule_rev_def) + apply (rule_tac xs = c in rev_induct, simp_all) + done lemma mr_first: "x \ dom (Cp b) \ (\ r. r \ set c \ x \ dom (Cp r)) \ s = b#c \ applied_rule_rev Cp x s = Some b" -by (simp add: foo5b) - + by (simp add: foo5b) + lemma mr_charn[rule_format]: "a \ set p \ (x \ dom (Cp a)) \(\ r. r \ set p \ x \ dom (Cp r) \ r = a) \ applied_rule_rev Cp x p = Some a" -apply(rule_tac xs = p in rev_induct) -apply(simp_all only:applied_rule_rev_def) -apply(simp,safe,simp_all) -by(auto) + apply(rule_tac xs = p in rev_induct) + apply(simp_all only:applied_rule_rev_def) + apply(simp,safe,simp_all) + by(auto) lemma foo8: "\ r. r \ set p \ x \ dom (Cp r) \ r = a \ set p = set s \ \ r. r \ set s \ x \ dom (Cp r) \ r = a" -by auto - + by auto + lemma mrConcEnd[rule_format]: "applied_rule_rev Cp x (b # p) = Some a \ a \ b \ applied_rule_rev Cp x p = Some a" -apply (simp add: applied_rule_rev_def) -apply (rule_tac xs = p in rev_induct,simp_all) -by auto - - + apply (simp add: applied_rule_rev_def) + apply (rule_tac xs = p in rev_induct,simp_all) + by auto + + lemma wp3tl[rule_format]: "wellformed_policy3Pr p \ wellformed_policy3Pr (tl p)" -by (induct p, simp_all, case_tac a, simp_all) - + by (induct p, simp_all, case_tac a, simp_all) + lemma wp3Conc[rule_format]: "wellformed_policy3Pr (a#p) \ wellformed_policy3Pr p" -by (induct p, simp_all, case_tac a, simp_all) + by (induct p, simp_all, case_tac a, simp_all) lemma foo98[rule_format]: "applied_rule_rev Cp x (aa # p) = Some a \ x \ dom (Cp r) \ r \ set p \ a \ set p" -unfolding applied_rule_rev_def + unfolding applied_rule_rev_def proof (induct rule: rev_induct) case Nil show ?case by simp next case (snoc xa xs) then show ?case 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) - + by (simp add: applied_rule_rev_def) + lemma DAAux[simp]: "x \ dom (Cp DenyAll)" -by (simp add: dom_def PolicyCombinators.PolicyCombinators Cp.simps) - + by (simp add: dom_def PolicyCombinators.PolicyCombinators Cp.simps) + lemma mrSet[rule_format]: "applied_rule_rev Cp x p = Some r \ r \ set p" -unfolding applied_rule_rev_def -by (rule_tac xs=p in rev_induct) simp_all - - + unfolding applied_rule_rev_def + by (rule_tac xs=p in rev_induct) simp_all + lemma mr_not_Conc: "singleCombinators p \ applied_rule_rev Cp x p \ Some (a\b)" -by (auto simp: mrSet dest: mrSet elim: SCnotConc) - - + by (auto simp: mrSet dest: mrSet elim: SCnotConc) + + lemma foo25[rule_format]: "wellformed_policy3Pr (p@[x]) \ wellformed_policy3Pr p" -by (induct p, simp_all, case_tac a, simp_all) - + by (induct p, simp_all, case_tac a, simp_all) + lemma mr_in_dom[rule_format]: "applied_rule_rev Cp x p = Some a \ x \ dom (Cp a)" -by (rule_tac xs = p in rev_induct) (auto simp: applied_rule_rev_def) - - + by (rule_tac xs = p in rev_induct) (auto simp: applied_rule_rev_def) + + lemma wp3EndMT[rule_format]: "wellformed_policy3Pr (p@[xs]) \ AllowPortFromTo a b po \ set p \ dom (Cp (AllowPortFromTo a b po)) \ dom (Cp xs) = {}" -apply (induct p,simp_all) -by (metis NormalisationIPPProofs.wp3Conc aux0_4 inf_commute list.set_intros(1) wellformed_policy3Pr.simps(2)) - - + apply (induct p,simp_all) + by (metis NormalisationIPPProofs.wp3Conc aux0_4 inf_commute list.set_intros(1) + wellformed_policy3Pr.simps(2)) + lemma foo29: "dom (Cp a) \ {} \ dom (Cp a) \ dom (Cp b) = {} \ a \ b" -by auto - + by auto + lemma foo28: "AllowPortFromTo a b po\set p \ dom(Cp(AllowPortFromTo a b po))\{} \ (wellformed_policy3Pr(p@[x])) \ x \ AllowPortFromTo a b po" -by (metis foo29 Cp.simps(3) wp3EndMT) - + by (metis foo29 Cp.simps(3) wp3EndMT) + lemma foo28a[rule_format]: "x \ dom (Cp a) \ dom (Cp a) \ {}" -by auto - + by auto + lemma allow_deny_dom[simp]: "dom (Cp (AllowPortFromTo a b po)) \ dom (Cp (DenyAllFromTo a b))" -by (simp_all add: twoNetsDistinct_def netsDistinct_def PLemmas) auto - + by (simp_all add: twoNetsDistinct_def netsDistinct_def PLemmas) auto + lemma DenyAllowDisj: "dom (Cp (AllowPortFromTo a b p)) \ {} \ dom (Cp (DenyAllFromTo a b)) \ dom (Cp (AllowPortFromTo a b p)) \ {}" -by (metis Int_absorb1 allow_deny_dom) - + by (metis Int_absorb1 allow_deny_dom) + lemma foo31: "\ r. r \ set p \ x \ dom (Cp r) \ (r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll) \ set p = set s \ (\r. r\set s \ x\dom(Cp r) \ r=AllowPortFromTo a b po \ r=DenyAllFromTo a b \ r = DenyAll)" -by auto - - + by auto + lemma wp1_auxa: "wellformed_policy1_strong p\(\ r. applied_rule_rev Cp x p = Some r)" -apply (rule DAimpliesMR_E) -by (erule wp1_aux1aa) - - + apply (rule DAimpliesMR_E) + by (erule wp1_aux1aa) + lemma deny_dom[simp]: "twoNetsDistinct a b c d \ dom (Cp (DenyAllFromTo a b)) \ dom (Cp (DenyAllFromTo c d)) = {}" -by (simp add: Cp.simps) (erule aux6) - + by (simp add: Cp.simps) (erule aux6) + lemma domTrans: "\dom a \ dom b; dom(b) \ dom (c) = {}\ \ dom(a) \ dom(c) = {}" -by auto - + by auto + lemma DomInterAllowsMT: " twoNetsDistinct a b c d \ dom (Cp(AllowPortFromTo a b p)) \ dom(Cp(AllowPortFromTo c d po))={}" -apply (case_tac "p = po", simp_all) -apply (rule_tac b = "Cp (DenyAllFromTo a b)" in domTrans, simp_all) -apply (metis domComm aux26 tNDComm) -apply (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) -by (auto simp: prod_eqI) - - + apply (case_tac "p = po", simp_all) + apply (rule_tac b = "Cp (DenyAllFromTo a b)" in domTrans, simp_all) + apply (metis domComm aux26 tNDComm) + apply (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) + by (auto simp: prod_eqI) + lemma DomInterAllowsMT_Ports: "p \ po \ dom (Cp (AllowPortFromTo a b p)) \ dom (Cp (AllowPortFromTo c d po)) = {}" -apply (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) -by (auto simp: prod_eqI) - - + apply (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) + by (auto simp: prod_eqI) + + lemma wellformed_policy3_charn[rule_format]: "singleCombinators p \ distinct p \ allNetsDistinct p \ wellformed_policy1 p \ wellformed_policy2Pr p \ wellformed_policy3Pr p" proof (induct p) - case Nil show ?case by simp + case Nil show ?case by simp next - case (Cons a p) then show ?case - apply (auto intro: singleCombinatorsConc ANDConc waux2 wp2Conc) - apply (case_tac a,simp_all, clarify) - apply (case_tac r,simp_all) + case (Cons a p) then show ?case + apply (auto intro: singleCombinatorsConc ANDConc waux2 wp2Conc) + apply (case_tac a,simp_all, clarify) + apply (case_tac r,simp_all) apply (metis Int_commute) apply (metis DomInterAllowsMT aux7aa DomInterAllowsMT_Ports) - apply (metis aux0_0 ) - done + apply (metis aux0_0 ) + done qed - - + lemma DistinctNetsDenyAllow: - "DenyAllFromTo b c \ set p \ AllowPortFromTo a d po \ set p\ allNetsDistinct p \ + "DenyAllFromTo b c \ set p \ AllowPortFromTo a d po \ set p\ allNetsDistinct p \ dom (Cp (DenyAllFromTo b c)) \ dom (Cp (AllowPortFromTo a d po)) \ {}\ b = a \ c = d" -apply (simp add: allNetsDistinct_def) -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 (metis Int_commute ND0aux1 ND0aux3 NDComm aux26 twoNetsDistinct_def ND0aux2 ND0aux4) -done - + apply (simp add: allNetsDistinct_def) + 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 (metis Int_commute ND0aux1 ND0aux3 NDComm aux26 twoNetsDistinct_def ND0aux2 ND0aux4) + done + lemma DistinctNetsAllowAllow: - "AllowPortFromTo b c poo \ set p \ AllowPortFromTo a d po \ set p \ + "AllowPortFromTo b c poo \ set p \ AllowPortFromTo a d po \ set p \ allNetsDistinct p \ dom(Cp(AllowPortFromTo b c poo)) \ dom(Cp(AllowPortFromTo a d po)) \ {} \ b = a \ c = d \ poo = po" -apply (simp add: allNetsDistinct_def) -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 (metis DomInterAllowsMT DomInterAllowsMT_Ports ND0aux3 ND0aux4 NDComm twoNetsDistinct_def) -done - - - + apply (simp add: allNetsDistinct_def) + 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 (metis DomInterAllowsMT DomInterAllowsMT_Ports ND0aux3 ND0aux4 NDComm twoNetsDistinct_def) + done + lemma WP2RS2[simp]: "singleCombinators p \ distinct p \ allNetsDistinct p \ wellformed_policy2Pr (removeShadowRules2 p)" @@ -390,34 +373,31 @@ proof (induct p) then show ?case by simp next case (Cons x xs) - have wp_xs: "wellformed_policy2Pr (removeShadowRules2 xs)" - by (metis Cons ANDConc distinct.simps(2) singleCombinatorsConc) - show ?case - proof (cases x) - case DenyAll thus ?thesis using wp_xs by simp - next - case (DenyAllFromTo a b) thus ?thesis - using wp_xs Cons - by (simp,metis DenyAllFromTo aux aux7 tNDComm deny_dom) - next - case (AllowPortFromTo a b p) thus ?thesis - using wp_xs - by (simp, metis aux26 AllowPortFromTo Cons(4) aux aux7a tNDComm) - next - case (Conc a b) thus ?thesis - by (metis Conc Cons(2) singleCombinators.simps(2)) - qed + have wp_xs: "wellformed_policy2Pr (removeShadowRules2 xs)" + by (metis Cons ANDConc distinct.simps(2) singleCombinatorsConc) + show ?case + proof (cases x) + case DenyAll thus ?thesis using wp_xs by simp + next + case (DenyAllFromTo a b) thus ?thesis + using wp_xs Cons + by (simp,metis DenyAllFromTo aux aux7 tNDComm deny_dom) + next + case (AllowPortFromTo a b p) thus ?thesis + using wp_xs + by (simp, metis aux26 AllowPortFromTo Cons(4) aux aux7a tNDComm) + next + case (Conc a b) thus ?thesis + by (metis Conc Cons(2) singleCombinators.simps(2)) + qed qed - - - + lemma AD_aux: "AllowPortFromTo a b po \ set p \ DenyAllFromTo c d \ set p \ allNetsDistinct p \ singleCombinators p \ a \ c \ b \ d \ dom (Cp (AllowPortFromTo a b po)) \ dom (Cp (DenyAllFromTo c d)) = {}" -by (rule aux26,rule_tac x ="AllowPortFromTo a b po" and y = "DenyAllFromTo c d" in tND) auto - - + 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 \ all_in_list p l \ distinct p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy2Pr p" @@ -432,43 +412,39 @@ next case (DenyAllFromTo c d) thus ?thesis using Cons apply (simp, intro impI conjI allI impI deny_dom) by (auto intro: aux7 tNDComm ANDConc singleCombinatorsConc sortedConcEnd) - next - case (AllowPortFromTo c d e) thus ?thesis - using Cons apply simp - apply (intro impI conjI allI, rename_tac "aa" "b") - apply (rule aux26) - apply (rule_tac x = "AllowPortFromTo c d e" and y = "DenyAllFromTo aa b" in tND, - assumption,simp_all) - apply (subgoal_tac "smaller (AllowPortFromTo c d e) (DenyAllFromTo aa b) l") - apply (simp split: if_splits) - apply metis - apply (erule sorted_is_smaller, simp_all) - apply (metis bothNet.simps(2) in_list.simps(2) in_set_in_list) - by (auto intro: aux7 tNDComm ANDConc singleCombinatorsConc sortedConcEnd) - next - case (Conc a b) thus ?thesis using Cons by simp - qed + next + case (AllowPortFromTo c d e) thus ?thesis + using Cons apply simp + apply (intro impI conjI allI, rename_tac "aa" "b") + apply (rule aux26) + apply (rule_tac x = "AllowPortFromTo c d e" and y = "DenyAllFromTo aa b" in tND, + assumption,simp_all) + apply (subgoal_tac "smaller (AllowPortFromTo c d e) (DenyAllFromTo aa b) l") + apply (simp split: if_splits) + apply metis + apply (erule sorted_is_smaller, simp_all) + apply (metis bothNet.simps(2) in_list.simps(2) in_set_in_list) + by (auto intro: aux7 tNDComm ANDConc singleCombinatorsConc sortedConcEnd) + next + case (Conc a b) thus ?thesis using Cons by simp + qed qed - - + lemma wellformed2_sorted[simp]: "all_in_list p l \ distinct p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy2Pr (sort p l)" -by (metis distinct_sort set_sort sorted_WP2 SC3 aND_sort all_in_listSubset order_refl sort_is_sorted) - - + 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 \ distinct p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy2Pr (qsort p l)" -by (metis sorted_WP2 SC3Q aND_sortQ all_in_listSubset distinct_sortQ set_sortQ sort_is_sortedQ subsetI) - - + 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) - - + by (auto simp: PLemmas) + lemma C_eq_RS1n: - "Cp(list2FWpolicy (removeShadowRules1_alternative p)) = Cp(list2FWpolicy p)" + "Cp(list2FWpolicy (removeShadowRules1_alternative p)) = Cp(list2FWpolicy p)" proof (cases "p") case Nil then show ?thesis by (simp, metis list2FWpolicy.simps(1) rSR1_eq removeShadowRules1.simps(2)) @@ -476,205 +452,186 @@ next case (Cons a list) then show ?thesis apply (hypsubst, simp) apply (thin_tac "p = a # list") - proof (induct rule: rev_induct) - case Nil show ?case by (metis rSR1_eq removeShadowRules1.simps(2)) - next - case (snoc x xs) show ?case - apply (case_tac "xs = []", simp_all) - apply (simp add: removeShadowRules1_alternative_def) - apply (insert snoc.hyps, case_tac x, simp_all) - apply (rule ext, rename_tac xa) - apply (case_tac "x = DenyAll",simp_all add: PLemmas) - apply (rule_tac t = "removeShadowRules1_alternative (xs @ [x])" and - s = "(removeShadowRules1_alternative xs)@[x]" in subst) - apply (erule RS1n_assoc) - apply (case_tac "xa \ dom (Cp x)", simp_all) - done - qed + proof (induct rule: rev_induct) + case Nil show ?case by (metis rSR1_eq removeShadowRules1.simps(2)) + next + case (snoc x xs) show ?case + apply (case_tac "xs = []", simp_all) + apply (simp add: removeShadowRules1_alternative_def) + apply (insert snoc.hyps, case_tac x, simp_all) + apply (rule ext, rename_tac xa) + apply (case_tac "x = DenyAll",simp_all add: PLemmas) + apply (rule_tac t = "removeShadowRules1_alternative (xs @ [x])" and + s = "(removeShadowRules1_alternative xs)@[x]" in subst) + apply (erule RS1n_assoc) + apply (case_tac "xa \ dom (Cp x)", simp_all) + done + qed qed - - + lemma C_eq_RS1[simp]: -"p \ [] \ Cp(list2FWpolicy (removeShadowRules1 p)) = Cp(list2FWpolicy p)" -by (metis rSR1_eq C_eq_RS1n) - - + "p \ [] \ 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) \ Some DenyAll \ (\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) - - + 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 \ (Some DenyAll) \ p = DenyAll#ps \ (applied_rule_rev Cp x p = applied_rule_rev Cp x ps)" -apply (auto,subgoal_tac "applied_rule_rev Cp x (DenyAll#ps) \ None", auto) - apply (metis mrConcEnd) -by (metis DAimpliesMR_E list.sel(1) hd_in_set list.simps(3) not_Some_eq) - - + apply (auto,subgoal_tac "applied_rule_rev Cp x (DenyAll#ps) \ 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 \ applied_rule_rev Cp x p = Some (DenyAllFromTo a ab) \ set p = set s \ applied_rule_rev Cp x s \ Some DenyAll" -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) - + 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)) \ dom (Cp (DenyAllFromTo c d)) \ {} \ \ netsDistinct a c" -by (erule contrapos_nn) (simp add: Cp.simps aux6 twoNetsDistinct_def) - - + by (erule contrapos_nn) (simp add: Cp.simps aux6 twoNetsDistinct_def) + lemma domsMT_notND_DD2: "dom (Cp (DenyAllFromTo a b)) \ dom (Cp (DenyAllFromTo c d)) \ {} \ \ netsDistinct b d" -by (erule contrapos_nn) (simp add: Cp.simps aux6 twoNetsDistinct_def) - - + by (erule contrapos_nn) (simp add: Cp.simps aux6 twoNetsDistinct_def) + lemma domsMT_notND_DD3: "x \ dom (Cp (DenyAllFromTo a b)) \ x \ dom (Cp (DenyAllFromTo c d)) \ \ netsDistinct a c" -by (auto intro!: domsMT_notND_DD) - - + by (auto intro!: domsMT_notND_DD) + lemma domsMT_notND_DD4: "x \ dom (Cp (DenyAllFromTo a b)) \ x \ dom (Cp (DenyAllFromTo c d)) \ \ netsDistinct b d" -by (auto intro!: domsMT_notND_DD2) - - + by (auto intro!: domsMT_notND_DD2) + lemma NetsEq_if_sameP_DD: "allNetsDistinct p \ u\ set p \ v\ set p \ u = (DenyAllFromTo a b) \ v = (DenyAllFromTo c d) \ x \ dom (Cp (u)) \ x \ dom (Cp (v)) \ a = c \ b = d" -unfolding allNetsDistinct_def -by (simp)(metis allNetsDistinct_def ND0aux1 ND0aux2 domsMT_notND_DD3 domsMT_notND_DD4 ) - + 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)" - and SC : "singleCombinators p" - and inp : "r \ set p" - and inDom : "x \ dom (Cp r)" - shows "(r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll)" + assumes aND : "allNetsDistinct p" + and mr_is_allow : "applied_rule_rev Cp x p = Some (AllowPortFromTo a b po)" + and SC : "singleCombinators p" + and inp : "r \ set p" + and inDom : "x \ dom (Cp r)" + shows "(r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll)" proof (cases r) case DenyAll show ?thesis by (metis DenyAll) next case (DenyAllFromTo x y) show ?thesis by (metis DenyAllFromTo NormalisationIPPProofs.AD_aux NormalisationIPPProofs.mrSet - NormalisationIPPProofs.mr_in_dom SC aND domInterMT inDom inp mr_is_allow) + NormalisationIPPProofs.mr_in_dom SC aND domInterMT inDom inp mr_is_allow) next case (AllowPortFromTo x y b) show ?thesis by (metis (mono_tags, lifting) AllowPortFromTo NormalisationIPPProofs.DistinctNetsAllowAllow - NormalisationIPPProofs.mrSet NormalisationIPPProofs.mr_in_dom aND domInterMT inDom - inp mr_is_allow) + NormalisationIPPProofs.mrSet NormalisationIPPProofs.mr_in_dom aND domInterMT inDom + inp mr_is_allow) next case (Conc x y) thus ?thesis using assms by (metis aux0_0) qed lemma none_MT_rulessubset[rule_format]: -"none_MT_rules Cp a \ set b \ set a \ none_MT_rules Cp b" -by (induct b,simp_all) (metis notMTnMT) - - + "none_MT_rules Cp a \ set b \ set a \ none_MT_rules Cp b" + by (induct b,simp_all) (metis notMTnMT) + lemma nMTSort: "none_MT_rules Cp p \ none_MT_rules Cp (sort p l)" -by (metis set_sort nMTeqSet) - - - + by (metis set_sort nMTeqSet) + lemma nMTSortQ: "none_MT_rules Cp p \ none_MT_rules Cp (qsort p l)" -by (metis set_sortQ nMTeqSet) + by (metis set_sortQ nMTeqSet) lemma wp3char[rule_format]: "none_MT_rules Cp xs \ Cp (AllowPortFromTo a b po) = empty \ wellformed_policy3Pr (xs @ [DenyAllFromTo a b]) \ AllowPortFromTo a b po \ set xs" -by (induct xs, simp_all) (metis domNMT wp3Conc) - - + by (induct xs, simp_all) (metis domNMT wp3Conc) lemma wp3charn[rule_format]: -assumes domAllow: "dom (Cp (AllowPortFromTo a b po)) \ {}" -and wp3: "wellformed_policy3Pr (xs @ [DenyAllFromTo a b])" -shows allowNotInList: "AllowPortFromTo a b po \ set xs" -apply (insert assms) + assumes domAllow: "dom (Cp (AllowPortFromTo a b po)) \ {}" + and wp3: "wellformed_policy3Pr (xs @ [DenyAllFromTo a b])" + shows allowNotInList: "AllowPortFromTo a b po \ set xs" + apply (insert assms) proof (induct xs) - case Nil show ?case by simp + case Nil show ?case by simp next - case (Cons x xs) show ?case using Cons - by (simp,auto intro: wp3Conc) (auto simp: DenyAllowDisj domAllow) + case (Cons x xs) show ?case using Cons + by (simp,auto intro: wp3Conc) (auto simp: DenyAllowDisj domAllow) qed - lemma rule_charn2: - assumes aND: "allNetsDistinct p" - and wp1: "wellformed_policy1 p" - and SC: "singleCombinators p" - and wp3: "wellformed_policy3Pr p" - and allow_in_list: "AllowPortFromTo c d po \ set p" - and x_in_dom_allow: "x \ dom (Cp (AllowPortFromTo c d po))" - shows "applied_rule_rev Cp x p = Some (AllowPortFromTo c d po)" - proof (insert assms, induct p rule: rev_induct) - case Nil show ?case using Nil by simp - next - case (snoc y ys) show ?case using snoc - apply simp - apply (case_tac "y = (AllowPortFromTo c d po)") + assumes aND: "allNetsDistinct p" + and wp1: "wellformed_policy1 p" + and SC: "singleCombinators p" + and wp3: "wellformed_policy3Pr p" + and allow_in_list: "AllowPortFromTo c d po \ set p" + and x_in_dom_allow: "x \ dom (Cp (AllowPortFromTo c d po))" + shows "applied_rule_rev Cp x p = Some (AllowPortFromTo c d po)" +proof (insert assms, induct p rule: rev_induct) + case Nil show ?case using Nil by simp +next + case (snoc y ys) show ?case using snoc + apply simp + apply (case_tac "y = (AllowPortFromTo c d po)") apply (simp add: applied_rule_rev_def) - apply simp_all - apply (subgoal_tac "ys \ []") + apply simp_all + apply (subgoal_tac "ys \ []") apply (subgoal_tac "applied_rule_rev Cp x ys = Some (AllowPortFromTo c d po)") - defer 1 - apply (metis ANDConcEnd SCConcEnd WP1ConcEnd foo25) + defer 1 + apply (metis ANDConcEnd SCConcEnd WP1ConcEnd foo25) apply (metis inSet_not_MT) - proof (cases y) - case DenyAll thus ?thesis using DenyAll snoc - apply simp - by (metis DAnotTL DenyAll inSet_not_MT policy2list.simps(2)) - next - case (DenyAllFromTo a b) thus ?thesis using snoc apply simp - apply (simp_all add: applied_rule_rev_def) - apply (rule conjI) - apply (metis domInterMT wp3EndMT) - apply (rule impI) - by (metis ANDConcEnd DenyAllFromTo SCConcEnd WP1ConcEnd foo25) - next - case (AllowPortFromTo a1 a2 b) thus ?thesis using AllowPortFromTo snoc apply simp - apply (simp_all add: applied_rule_rev_def) - apply (rule conjI) - apply (metis domInterMT wp3EndMT) - by (metis ANDConcEnd AllowPortFromTo SCConcEnd WP1ConcEnd foo25 x_in_dom_allow) - next - case (Conc a b) thus ?thesis - using Conc snoc apply simp - by (metis Conc aux0_0 in_set_conv_decomp) - qed + proof (cases y) + case DenyAll thus ?thesis using DenyAll snoc + apply simp + by (metis DAnotTL DenyAll inSet_not_MT policy2list.simps(2)) + next + case (DenyAllFromTo a b) thus ?thesis using snoc apply simp + apply (simp_all add: applied_rule_rev_def) + apply (rule conjI) + apply (metis domInterMT wp3EndMT) + apply (rule impI) + by (metis ANDConcEnd DenyAllFromTo SCConcEnd WP1ConcEnd foo25) + next + case (AllowPortFromTo a1 a2 b) thus ?thesis using AllowPortFromTo snoc apply simp + apply (simp_all add: applied_rule_rev_def) + apply (rule conjI) + apply (metis domInterMT wp3EndMT) + by (metis ANDConcEnd AllowPortFromTo SCConcEnd WP1ConcEnd foo25 x_in_dom_allow) + next + case (Conc a b) thus ?thesis + using Conc snoc apply simp + by (metis Conc aux0_0 in_set_conv_decomp) + qed qed - lemma rule_charn3: "wellformed_policy1 p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy3Pr p \ applied_rule_rev Cp x p = Some (DenyAllFromTo c d) \ AllowPortFromTo a b po \ set p \ x \ dom (Cp (AllowPortFromTo a b po))" -by (clarify) (simp add: NormalisationIPPProofs.rule_charn2 domI) - + by (clarify) (simp add: NormalisationIPPProofs.rule_charn2 domI) lemma rule_charn4: -assumes wp1: "wellformed_policy1 p" -and aND: "allNetsDistinct p" -and SC: "singleCombinators p" -and wp3: "wellformed_policy3Pr p" -and DA: "DenyAll \ set p" -and mr: "applied_rule_rev Cp x p = Some (DenyAllFromTo a b)" -and rinp: "r \ set p" -and xindom: "x \ dom (Cp r)" -shows "r = DenyAllFromTo a b" + assumes wp1: "wellformed_policy1 p" + and aND: "allNetsDistinct p" + and SC: "singleCombinators p" + and wp3: "wellformed_policy3Pr p" + and DA: "DenyAll \ set p" + and mr: "applied_rule_rev Cp x p = Some (DenyAllFromTo a b)" + and rinp: "r \ set p" + and xindom: "x \ dom (Cp r)" + shows "r = DenyAllFromTo a b" proof (cases r) case DenyAll thus ?thesis using DenyAll assms by simp next case (DenyAllFromTo c d) thus ?thesis using assms apply simp apply (erule_tac x = x and p = p and v = "(DenyAllFromTo a b)" and - u = "(DenyAllFromTo c d)" in NetsEq_if_sameP_DD, simp_all) - apply (erule mrSet) + u = "(DenyAllFromTo c d)" in NetsEq_if_sameP_DD, simp_all) + apply (erule mrSet) by (erule mr_in_dom) next case (AllowPortFromTo c d e) thus ?thesis @@ -687,200 +644,185 @@ next by (metis Conc aux0_0) qed - - lemma foo31a: "(\ r. r \ set p \ x \ dom (Cp r) \ (r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll)) \ set p = set s \ r \ set s \ x \ dom (Cp r) \ (r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll)" -by auto - + by auto lemma aux4[rule_format]: "applied_rule_rev Cp x (a#p) = Some a \ a \ set (p) \ applied_rule_rev Cp x p = None" -by (rule rev_induct, simp_all) (intro impI,simp add: applied_rule_rev_def split: if_splits) - + 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" -shows "applied_rule_rev Cp x (tl p) = None" -apply (rule aux4 [where a = DenyAll]) - apply (metis wp1n_tl mr_DA wp1n) -by (metis WP1n_DA_notinSet wp1n) + assumes mr_DA: "applied_rule_rev Cp x p = Some DenyAll" + and wp1n: "wellformed_policy1_strong p" + shows "applied_rule_rev Cp x (tl p) = None" + apply (rule aux4 [where a = DenyAll]) + apply (metis wp1n_tl mr_DA wp1n) + by (metis WP1n_DA_notinSet wp1n) lemma rule_charnDAFT: "wellformed_policy1_strong p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy3Pr p \ applied_rule_rev Cp x p = Some (DenyAllFromTo a b) \ r \ set (tl p) \ x \ dom (Cp r) \ r = DenyAllFromTo a b" -apply (subgoal_tac "p = DenyAll#(tl p)") - apply (metis (no_types, lifting) ANDConc Combinators.distinct(1) NormalisationIPPProofs.mrConcEnd - NormalisationIPPProofs.rule_charn4 NormalisationIPPProofs.wp3Conc WP1n_DA_notinSet - singleCombinatorsConc waux2) -using wp1n_tl by auto - + apply (subgoal_tac "p = DenyAll#(tl p)") + apply (metis (no_types, lifting) ANDConc Combinators.distinct(1) NormalisationIPPProofs.mrConcEnd + NormalisationIPPProofs.rule_charn4 NormalisationIPPProofs.wp3Conc WP1n_DA_notinSet + singleCombinatorsConc waux2) + using wp1n_tl by auto lemma mrDenyAll_is_unique: "wellformed_policy1_strong p \ applied_rule_rev Cp x p = Some DenyAll \ r \ set (tl p) \ x \ dom (Cp r)" -apply (rule_tac a = "[]" and b = "DenyAll" and c = "tl p" in foo3a, simp_all) - apply (metis wp1n_tl) -by (metis WP1n_DA_notinSet) - + 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" - and wp1_p: "wellformed_policy1_strong p" - and wp1_s: "wellformed_policy1_strong s" - and wp3_p: "wellformed_policy3Pr p" - and wp3_s: "wellformed_policy3Pr s" - and aND: "allNetsDistinct p" - shows "applied_rule_rev Cp x p = applied_rule_rev Cp x s" + assumes sets_eq: "set p = set s" + and SC: "singleCombinators p" + and wp1_p: "wellformed_policy1_strong p" + and wp1_s: "wellformed_policy1_strong s" + and wp3_p: "wellformed_policy3Pr p" + and wp3_s: "wellformed_policy3Pr s" + and aND: "allNetsDistinct p" + shows "applied_rule_rev Cp x p = applied_rule_rev Cp x s" proof (cases "applied_rule_rev Cp x p") case None - have DA: "DenyAll \ set p" using wp1_p by (auto simp: wp1_aux1aa) - have notDA: "DenyAll \ set p" using None by (auto simp: DAimplieMR) - thus ?thesis using DA by (contradiction) + have DA: "DenyAll \ set p" using wp1_p by (auto simp: wp1_aux1aa) + have notDA: "DenyAll \ set p" using None by (auto simp: DAimplieMR) + thus ?thesis using DA by (contradiction) next case (Some y) thus ?thesis - proof (cases y) - have tl_p: "p = DenyAll#(tl p)" by (metis wp1_p wp1n_tl) - have tl_s: "s = DenyAll#(tl s)" by (metis wp1_s wp1n_tl) - have tl_eq: "set (tl p) = set (tl s)" - by (metis list.sel(3) WP1n_DA_notinSet sets_eq foo2 - wellformed_policy1_charn wp1_aux1aa wp1_eq wp1_p wp1_s) + proof (cases y) + have tl_p: "p = DenyAll#(tl p)" by (metis wp1_p wp1n_tl) + have tl_s: "s = DenyAll#(tl s)" by (metis wp1_s wp1n_tl) + have tl_eq: "set (tl p) = set (tl s)" + by (metis list.sel(3) WP1n_DA_notinSet sets_eq foo2 + wellformed_policy1_charn wp1_aux1aa wp1_eq wp1_p wp1_s) { case DenyAll - have mr_p_is_DenyAll: "applied_rule_rev Cp x p = Some DenyAll" - by (simp add: DenyAll Some) - hence x_notin_tl_p: "\ r. r \ set (tl p) \ x \ dom (Cp r)" using wp1_p - by (auto simp: mrDenyAll_is_unique) - hence x_notin_tl_s: "\ r. r \ set (tl s) \ x \ dom (Cp r)" using tl_eq - by auto - hence mr_s_is_DenyAll: "applied_rule_rev Cp x s = Some DenyAll" using tl_s - by (auto simp: mr_first) - thus ?thesis using mr_p_is_DenyAll by simp + have mr_p_is_DenyAll: "applied_rule_rev Cp x p = Some DenyAll" + by (simp add: DenyAll Some) + hence x_notin_tl_p: "\ r. r \ set (tl p) \ x \ dom (Cp r)" using wp1_p + by (auto simp: mrDenyAll_is_unique) + hence x_notin_tl_s: "\ r. r \ set (tl s) \ x \ dom (Cp r)" using tl_eq + by auto + hence mr_s_is_DenyAll: "applied_rule_rev Cp x s = Some DenyAll" using tl_s + by (auto simp: mr_first) + thus ?thesis using mr_p_is_DenyAll by simp next case (DenyAllFromTo a b) - have mr_p_is_DAFT: "applied_rule_rev Cp x p = Some (DenyAllFromTo a b)" - by (simp add: DenyAllFromTo Some) - have DA_notin_tl: "DenyAll \ set (tl p)" - by (metis WP1n_DA_notinSet wp1_p) - have mr_tl_p: "applied_rule_rev Cp x p = applied_rule_rev Cp x (tl p)" - by (metis Combinators.simps(4) DenyAllFromTo Some mrConcEnd tl_p) - have dom_tl_p: "\ r. r \ set (tl p) \ x \ dom (Cp r) \ + have mr_p_is_DAFT: "applied_rule_rev Cp x p = Some (DenyAllFromTo a b)" + by (simp add: DenyAllFromTo Some) + have DA_notin_tl: "DenyAll \ set (tl p)" + by (metis WP1n_DA_notinSet wp1_p) + have mr_tl_p: "applied_rule_rev Cp x p = applied_rule_rev Cp x (tl p)" + by (metis Combinators.simps(4) DenyAllFromTo Some mrConcEnd tl_p) + have dom_tl_p: "\ r. r \ set (tl p) \ x \ dom (Cp r) \ r = (DenyAllFromTo a b)" - using wp1_p aND SC wp3_p mr_p_is_DAFT - by (auto simp: rule_charnDAFT) - hence dom_tl_s: "\ r. r \ set (tl s) \ x \ dom (Cp r) \ + using wp1_p aND SC wp3_p mr_p_is_DAFT + by (auto simp: rule_charnDAFT) + hence dom_tl_s: "\ r. r \ set (tl s) \ x \ dom (Cp r) \ r = (DenyAllFromTo a b)" - using tl_eq by auto - have DAFT_in_tl_s: "DenyAllFromTo a b \ set (tl s)" using mr_tl_p - by (metis DenyAllFromTo mrSet mr_p_is_DAFT tl_eq) - have x_in_dom_DAFT: "x \ dom (Cp (DenyAllFromTo a b))" - by (metis mr_p_is_DAFT DenyAllFromTo mr_in_dom) - hence mr_tl_s_is_DAFT: "applied_rule_rev Cp x (tl s) = Some (DenyAllFromTo a b)" - using DAFT_in_tl_s dom_tl_s by (metis mr_charn) - hence mr_s_is_DAFT: "applied_rule_rev Cp x s = Some (DenyAllFromTo a b)" - using tl_s - by (metis DA_notin_tl DenyAllFromTo EX_MR mrDA_tl - not_Some_eq tl_eq wellformed_policy1_strong.simps(2)) - thus ?thesis using mr_p_is_DAFT by simp + using tl_eq by auto + have DAFT_in_tl_s: "DenyAllFromTo a b \ set (tl s)" using mr_tl_p + by (metis DenyAllFromTo mrSet mr_p_is_DAFT tl_eq) + have x_in_dom_DAFT: "x \ dom (Cp (DenyAllFromTo a b))" + by (metis mr_p_is_DAFT DenyAllFromTo mr_in_dom) + hence mr_tl_s_is_DAFT: "applied_rule_rev Cp x (tl s) = Some (DenyAllFromTo a b)" + using DAFT_in_tl_s dom_tl_s by (metis mr_charn) + hence mr_s_is_DAFT: "applied_rule_rev Cp x s = Some (DenyAllFromTo a b)" + using tl_s + by (metis DA_notin_tl DenyAllFromTo EX_MR mrDA_tl + not_Some_eq tl_eq wellformed_policy1_strong.simps(2)) + thus ?thesis using mr_p_is_DAFT by simp next case (AllowPortFromTo a b c) - have wp1s: "wellformed_policy1 s" by (metis wp1_eq wp1_s) - have mr_p_is_A: "applied_rule_rev Cp x p = Some (AllowPortFromTo a b c)" - by (simp add: AllowPortFromTo Some) - hence A_in_s: "AllowPortFromTo a b c \ set s" using sets_eq - by (auto intro: mrSet) - have x_in_dom_A: "x \ dom (Cp (AllowPortFromTo a b c))" - by (metis mr_p_is_A AllowPortFromTo mr_in_dom) - have SCs: "singleCombinators s" using SC sets_eq - by (auto intro: SCSubset) - hence ANDs: "allNetsDistinct s" using aND sets_eq SC - by (auto intro: aNDSetsEq) - hence mr_s_is_A: "applied_rule_rev Cp x s = Some (AllowPortFromTo a b c)" - using A_in_s wp1s mr_p_is_A aND SCs wp3_s x_in_dom_A - by (simp add: rule_charn2) - thus ?thesis using mr_p_is_A by simp + have wp1s: "wellformed_policy1 s" by (metis wp1_eq wp1_s) + have mr_p_is_A: "applied_rule_rev Cp x p = Some (AllowPortFromTo a b c)" + by (simp add: AllowPortFromTo Some) + hence A_in_s: "AllowPortFromTo a b c \ set s" using sets_eq + by (auto intro: mrSet) + have x_in_dom_A: "x \ dom (Cp (AllowPortFromTo a b c))" + by (metis mr_p_is_A AllowPortFromTo mr_in_dom) + have SCs: "singleCombinators s" using SC sets_eq + by (auto intro: SCSubset) + hence ANDs: "allNetsDistinct s" using aND sets_eq SC + by (auto intro: aNDSetsEq) + hence mr_s_is_A: "applied_rule_rev Cp x s = Some (AllowPortFromTo a b c)" + using A_in_s wp1s mr_p_is_A aND SCs wp3_s x_in_dom_A + by (simp add: rule_charn2) + thus ?thesis using mr_p_is_A by simp } -next - case (Conc a b) thus ?thesis by (metis Some mr_not_Conc SC) -qed + next + case (Conc a b) thus ?thesis by (metis Some mr_not_Conc SC) + qed qed - lemma C_eq_Sets: "singleCombinators p \ wellformed_policy1_strong p \ wellformed_policy1_strong s \ wellformed_policy3Pr p \ wellformed_policy3Pr s \ allNetsDistinct p \ set p = set s \ Cp (list2FWpolicy p) x = Cp (list2FWpolicy s) x" -by (metis C_eq_Sets_mr C_eq_if_mr_eq wellformed_policy1_strong.simps(1)) - - + by (metis C_eq_Sets_mr C_eq_if_mr_eq wellformed_policy1_strong.simps(1)) lemma C_eq_sorted: "distinct p \ all_in_list p l \ singleCombinators p \ wellformed_policy1_strong p\ wellformed_policy3Pr p\ allNetsDistinct p \ Cp (list2FWpolicy (sort p l))= Cp (list2FWpolicy p)" -by (rule ext) - (meson distinct_sort set_sort C_eq_Sets wellformed2_sorted wellformed_policy3_charn SC3 aND_sort - wellformed1_alternative_sorted wp1_eq) - - + 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 \ all_in_list p l \ singleCombinators p \ wellformed_policy1_strong p \ wellformed_policy3Pr p \ allNetsDistinct p \ Cp (list2FWpolicy (qsort p l))= Cp (list2FWpolicy p)" -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) - - - + 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 next case (Cons y ys) thus ?case - proof (cases "ys = []") - case True thus ?thesis by (cases y, simp_all) + proof (cases "ys = []") + case True thus ?thesis by (cases y, simp_all) + next + case False thus ?thesis + proof (cases y) + case DenyAll thus ?thesis by (simp, metis Cons DenyAll mreq_end2) next - case False thus ?thesis - proof (cases y) - case DenyAll thus ?thesis by (simp, metis Cons DenyAll mreq_end2) - next - case (DenyAllFromTo a b) thus ?thesis by (simp, metis Cons DenyAllFromTo mreq_end2) - next - case (AllowPortFromTo a b p) thus ?thesis - proof (cases "DenyAllFromTo a b \ set ys") - case True thus ?thesis using AllowPortFromTo Cons - apply (cases "applied_rule_rev Cp x ys = None", simp_all) - apply (subgoal_tac "x \ dom (Cp (AllowPortFromTo a b p))") - apply (subst mrconcNone, simp_all) - apply (simp add: applied_rule_rev_def ) - apply (rule contra_subsetD [OF allow_deny_dom]) - apply (erule mrNoneMT,simp) - apply (metis AllowPortFromTo mrconc) - done - next - case False thus ?thesis using False Cons AllowPortFromTo - by (simp, metis AllowPortFromTo Cons mreq_end2) qed - next - case (Conc a b) thus ?thesis - by (metis Cons mreq_end2 removeShadowRules2.simps(4)) - qed + case (DenyAllFromTo a b) thus ?thesis by (simp, metis Cons DenyAllFromTo mreq_end2) + next + case (AllowPortFromTo a b p) thus ?thesis + proof (cases "DenyAllFromTo a b \ set ys") + case True thus ?thesis using AllowPortFromTo Cons + apply (cases "applied_rule_rev Cp x ys = None", simp_all) + apply (subgoal_tac "x \ dom (Cp (AllowPortFromTo a b p))") + apply (subst mrconcNone, simp_all) + apply (simp add: applied_rule_rev_def ) + apply (rule contra_subsetD [OF allow_deny_dom]) + apply (erule mrNoneMT,simp) + apply (metis AllowPortFromTo mrconc) + done + next + case False thus ?thesis using False Cons AllowPortFromTo + by (simp, metis AllowPortFromTo Cons mreq_end2) qed + next + case (Conc a b) thus ?thesis + by (metis Cons mreq_end2 removeShadowRules2.simps(4)) qed + qed qed - - lemma C_eq_None[rule_format]: "p \ [] \ applied_rule_rev Cp x p = None \ Cp (list2FWpolicy p) x = None" -unfolding applied_rule_rev_def + unfolding applied_rule_rev_def proof(induct rule: rev_induct) case Nil show ?case by simp next @@ -890,86 +832,78 @@ next apply (metis CConcStart2 option.simps(3)) by (metis append_Nil domIff l2p_aux2 option.distinct(1)) qed - + lemma C_eq_None2: "a \ [] \ b \ [] \ applied_rule_rev Cp x a = None \ applied_rule_rev Cp x b = None \ (Cp (list2FWpolicy a)) x = (Cp (list2FWpolicy b)) x" -by (auto simp: C_eq_None) - + by (auto simp: C_eq_None) + lemma C_eq_RS2: - "wellformed_policy1_strong p \ + "wellformed_policy1_strong p \ Cp (list2FWpolicy (removeShadowRules2 p))= Cp (list2FWpolicy p)" -apply (rule ext) -by (metis C_eq_RS2_mr C_eq_if_mr_eq RS2_NMT wp1_alternative_not_mt) - + 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 \ none_MT_rules Cp (removeShadowRules2 p)" -by (auto simp: RS2Set none_MT_rulessubset) + by (auto simp: RS2Set none_MT_rulessubset) lemma CconcNone: "dom (Cp a) = {} \ p \ [] \ Cp (list2FWpolicy (a # p)) x = Cp (list2FWpolicy p) x" -apply (case_tac "p = []", simp_all) - apply (case_tac "x\ dom (Cp (list2FWpolicy(p)))") -apply (metis Cdom2 list2FWpolicyconc) -apply (metis Cp.simps(4) map_add_dom_app_simps(2) inSet_not_MT list2FWpolicyconc set_empty2) -done - + apply (case_tac "p = []", simp_all) + apply (case_tac "x\ dom (Cp (list2FWpolicy(p)))") + 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 \ none_MT_rules Cp (remdups p)" -by (induct p, simp_all) - + by (induct p, simp_all) + lemma DARS3[rule_format]:"DenyAll \ set p\DenyAll \ set (rm_MT_rules Cp p)" -by (induct p, simp_all) - + by (induct p, simp_all) + lemma DAnMT: "dom (Cp DenyAll) \ {}" -by (simp add: dom_def Cp.simps PolicyCombinators.PolicyCombinators) - + by (simp add: dom_def Cp.simps PolicyCombinators.PolicyCombinators) + lemma DAnMT2: "Cp DenyAll \ empty" -by (metis DAAux dom_eq_empty_conv empty_iff) - - - + by (metis DAAux dom_eq_empty_conv empty_iff) lemma wp1n_RS3[rule_format,simp]: "wellformed_policy1_strong p \ wellformed_policy1_strong (rm_MT_rules Cp p)" -apply (induct p, simp_all) - apply (rule conjI| rule impI|simp)+ -apply (metis DAnMT) -apply (metis DARS3) -done + apply (induct p, simp_all) + apply (rule conjI| rule impI|simp)+ + apply (metis DAnMT) + apply (metis DARS3) + done lemma AILRS3[rule_format,simp]: "all_in_list p l \ all_in_list (rm_MT_rules Cp p) l" -by (induct p, simp_all) - + by (induct p, simp_all) + lemma SCRS3[rule_format,simp]: "singleCombinators p \ singleCombinators(rm_MT_rules Cp p)" -by (induct p, simp_all) (case_tac "a", simp_all) - - + by (induct p, simp_all) (case_tac "a", simp_all) + lemma RS3subset: "set (rm_MT_rules Cp p) \ set p " -by (induct p, auto) - - + by (induct p, auto) + lemma ANDRS3[simp]: "singleCombinators p \ allNetsDistinct p \ allNetsDistinct (rm_MT_rules Cp p)" -by (rule_tac b = p in aNDSubset, simp_all add:RS3subset) - - + by (rule_tac b = p in aNDSubset, simp_all add:RS3subset) + lemma nlpaux: "x \ dom (Cp b) \ Cp (a \ b) x = Cp a x" -by (metis Cp.simps(4) map_add_dom_app_simps(3)) - + by (metis Cp.simps(4) map_add_dom_app_simps(3)) + lemma notindom[rule_format]: "a \ set p \ x \ dom (Cp (list2FWpolicy p)) \ x \ dom (Cp a)" proof (induct p) - case Nil show ?case by simp + case Nil show ?case by simp next - case (Cons a p) then show ?case - apply (simp_all,intro conjI impI) + case (Cons a p) then show ?case + apply (simp_all,intro conjI impI) apply (metis CConcStartA) - apply simp - apply (metis Cdom2 List.set_simps(2) domIff insert_absorb list.simps(2) list2FWpolicyconc set_empty) - done + apply simp + apply (metis Cdom2 List.set_simps(2) domIff insert_absorb list.simps(2) list2FWpolicyconc set_empty) + done qed lemma C_eq_rd[rule_format]: @@ -978,205 +912,183 @@ proof (rule ext, induct p) case Nil thus ?case by simp next case (Cons y ys) thus ?case - proof (cases "ys = []") - case True thus ?thesis by simp - next - case False thus ?thesis - using Cons apply simp - apply (intro conjI impI) + proof (cases "ys = []") + case True thus ?thesis by simp + next + case False thus ?thesis + using Cons apply simp + apply (intro conjI impI) apply (metis Cdom2 nlpaux notindom domIff l2p_aux) - by (metis (no_types, lifting) Cdom2 nlpaux domIff l2p_aux remDupsNMT) - qed + by (metis (no_types, lifting) Cdom2 nlpaux domIff l2p_aux remDupsNMT) + qed qed - - - + lemma nMT_domMT: -"\ not_MT Cp p \ p \ [] \ r \ dom (Cp (list2FWpolicy p))" + "\ not_MT Cp p \ p \ [] \ r \ dom (Cp (list2FWpolicy p))" proof (induct p) - case Nil thus ?case by simp + case Nil thus ?case by simp next - case (Cons x xs) thus ?case - apply (simp split: if_splits) - apply (cases "xs = []",simp_all ) - by (metis CconcNone domIff) + case (Cons x xs) thus ?case + apply (simp split: if_splits) + apply (cases "xs = []",simp_all ) + by (metis CconcNone domIff) qed lemma C_eq_RS3_aux[rule_format]: "not_MT Cp p \ Cp (list2FWpolicy p) x = Cp (list2FWpolicy (rm_MT_rules Cp p)) x" proof (induct p) - case Nil thus ?case by simp + case Nil thus ?case by simp next - case (Cons y ys) thus ?case - proof (cases "not_MT Cp ys") - case True thus ?thesis - using Cons apply simp - apply (intro conjI impI, simp) - apply (metis CconcNone True not_MTimpnotMT) - apply (cases "x \ dom (Cp (list2FWpolicy ys))") - apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (rm_MT_rules Cp ys)))") + case (Cons y ys) thus ?case + proof (cases "not_MT Cp ys") + case True thus ?thesis + using Cons apply simp + apply (intro conjI impI, simp) + apply (metis CconcNone True not_MTimpnotMT) + apply (cases "x \ dom (Cp (list2FWpolicy ys))") + apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (rm_MT_rules Cp ys)))") apply (metis (mono_tags) Cons_eq_appendI NMPrm CeqEnd append_Nil not_MTimpnotMT) apply (simp add: domIff) - apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (rm_MT_rules Cp ys)))") + apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (rm_MT_rules Cp ys)))") apply (metis l2p_aux l2p_aux2 nlpaux) - by (metis domIff) - next - case False thus ?thesis - using Cons False - proof (cases "ys = []") - case True thus ?thesis using Cons by (simp) (rule impI, simp) - next - case False thus ?thesis - using Cons False `\ not_MT Cp ys` apply (simp) - apply (intro conjI impI| simp)+ - apply (subgoal_tac "rm_MT_rules Cp ys = []") - apply (subgoal_tac "x \ dom (Cp (list2FWpolicy ys))") - apply simp_all - apply (metis l2p_aux nlpaux) - apply (erule nMT_domMT, simp_all) - by (metis SR3nMT) - qed - qed + by (metis domIff) + next + case False thus ?thesis + using Cons False + proof (cases "ys = []") + case True thus ?thesis using Cons by (simp) (rule impI, simp) + next + case False thus ?thesis + using Cons False `\ not_MT Cp ys` apply (simp) + apply (intro conjI impI| simp)+ + apply (subgoal_tac "rm_MT_rules Cp ys = []") + apply (subgoal_tac "x \ dom (Cp (list2FWpolicy ys))") + apply simp_all + apply (metis l2p_aux nlpaux) + apply (erule nMT_domMT, simp_all) + by (metis SR3nMT) + qed + qed qed - - lemma C_eq_id: "wellformed_policy1_strong p \ Cp(list2FWpolicy (insertDeny p)) = Cp (list2FWpolicy p)" -by (rule ext) (metis insertDeny.simps(1) wp1n_tl) - + by (rule ext) (metis insertDeny.simps(1) wp1n_tl) lemma C_eq_RS3: -"not_MT Cp p \ Cp(list2FWpolicy (rm_MT_rules Cp p)) = Cp (list2FWpolicy p)" -by (rule ext) (erule C_eq_RS3_aux[symmetric]) - + "not_MT Cp p \ 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 \ not_MT Cp (remdups p)" -by (induct p, simp_all) (auto simp: NMPcharn) - + by (induct p, simp_all) (auto simp: NMPcharn) lemma NMPDA[rule_format]: "DenyAll \ set p \ not_MT Cp p" -by (induct p, simp_all add: DAnMT) - + by (induct p, simp_all add: DAnMT) lemma NMPiD[rule_format]: "not_MT Cp (insertDeny p)" -by (insert DAiniD [of p]) (erule NMPDA) - + by (insert DAiniD [of p]) (erule NMPDA) lemma list2FWpolicy2list[rule_format]: "Cp (list2FWpolicy(policy2list p)) = (Cp p)" -apply (rule ext) -apply (induct_tac p, simp_all) -apply (case_tac "x \ dom (Cp (x2))") -apply (metis Cdom2 CeqEnd domIff p2lNmt) -apply (metis CeqStart domIff p2lNmt nlpaux) -done - + apply (rule ext) + apply (induct_tac p, simp_all) + apply (case_tac "x \ dom (Cp (x2))") + 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\set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ Cp(list2FWpolicy (sort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l)) = Cp p" -apply (subst C_eq_sorted,simp_all add: C_eq_Lemmas) -apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) -apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) -apply (subst C_eq_RS3,simp_all add: C_eq_Lemmas) -apply (subst C_eq_id,simp_all add: C_eq_Lemmas) -done - + apply (subst C_eq_sorted,simp_all add: C_eq_Lemmas) + apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) + apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) + 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\ set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ Cp(list2FWpolicy (qsort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l)) = Cp p" -apply (subst C_eq_sortedQ,simp_all add: C_eq_Lemmas) -apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) -apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) -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 *) + apply (subst C_eq_sortedQ,simp_all add: C_eq_Lemmas) + apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) + apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) + 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_untilSorted_withSimps: -"DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ + "DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ allNetsDistinct (policy2list p) \ Cp(list2FWpolicy(sort(removeShadowRules2(remdups(rm_MT_rules Cp (insertDeny (removeShadowRules1(policy2list p)))))) l)) = Cp p" -by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) - - + by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) + lemma C_eq_All_untilSorted_withSimpsQ: -"DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ + "DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ allNetsDistinct (policy2list p) \ Cp(list2FWpolicy(qsort(removeShadowRules2(remdups(rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l)) = Cp p" -by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) - + by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) lemma InDomConc[rule_format]: "p \ [] \ x \ dom (Cp (list2FWpolicy (p))) \ x \ dom (Cp (list2FWpolicy (a#p)))" -by (induct p, simp_all) (case_tac "p = []",simp_all add: dom_def Cp.simps) - + by (induct p, simp_all) (case_tac "p = []",simp_all add: dom_def Cp.simps) lemma not_in_member[rule_format]: "member a b \ x \ dom (Cp b) \ x \ dom (Cp a)" -by (induct b)(simp_all add: dom_def Cp.simps) - + by (induct b)(simp_all add: dom_def Cp.simps) lemma src_in_sdnets[rule_format]: "\ member DenyAll x \ p \ dom (Cp x) \ subnetsOfAdr (src p) \ (fst_set (sdnets x)) \ {}" -apply (induct rule: Combinators.induct) -apply (simp_all add: fst_set_def subnetsOfAdr_def PLemmas, rename_tac x1 x2) -apply (intro impI) -apply (simp add: fst_set_def) -apply (case_tac "p \ dom (Cp x2)") -apply (rule subnetAux) -apply (auto simp: PLemmas) -done - + apply (induct rule: Combinators.induct) + apply (simp_all add: fst_set_def subnetsOfAdr_def PLemmas, rename_tac x1 x2) + apply (intro impI) + apply (simp add: fst_set_def) + apply (case_tac "p \ dom (Cp x2)") + apply (rule subnetAux) + apply (auto simp: PLemmas) + done + lemma dest_in_sdnets[rule_format]: "\ member DenyAll x \ p \ dom (Cp x) \ subnetsOfAdr (dest p) \ (snd_set (sdnets x)) \ {}" -apply (induct rule: Combinators.induct) -apply (simp_all add: snd_set_def subnetsOfAdr_def PLemmas, rename_tac x1 x2) -apply (intro impI,simp add: snd_set_def) -apply (case_tac "p \ dom (Cp x2)") -apply (rule subnetAux) -apply (auto simp: PLemmas) -done - - + apply (induct rule: Combinators.induct) + apply (simp_all add: snd_set_def subnetsOfAdr_def PLemmas, rename_tac x1 x2) + apply (intro impI,simp add: snd_set_def) + apply (case_tac "p \ dom (Cp x2)") + apply (rule subnetAux) + apply (auto simp: PLemmas) + done lemma sdnets_in_subnets[rule_format]: "p\ dom (Cp x) \ \ member DenyAll x \ (\ (a,b)\sdnets x. a \ subnetsOfAdr (src p) \ b \ subnetsOfAdr (dest p))" -apply (rule Combinators.induct) - apply (simp_all add: PLemmas subnetsOfAdr_def) -apply (intro impI, simp) - apply (case_tac "p \ dom (Cp (x2))") -apply (auto simp: PLemmas subnetsOfAdr_def) -done + apply (rule Combinators.induct) + apply (simp_all add: PLemmas subnetsOfAdr_def) + apply (intro impI, simp) + apply (case_tac "p \ dom (Cp (x2))") + apply (auto simp: PLemmas subnetsOfAdr_def) + done lemma disjSD_no_p_in_both[rule_format]: - "\disjSD_2 x y; \ member DenyAll x; \ member DenyAll y; + "\disjSD_2 x y; \ member DenyAll x; \ member DenyAll y; p \ dom (Cp x); p \ dom (Cp y)\ \ False" -apply (rule_tac A = "sdnets x" and B = "sdnets y" and D = "src p" and F = "dest p" in tndFalse) -by (auto simp: dest_in_sdnets src_in_sdnets sdnets_in_subnets disjSD_2_def) + apply (rule_tac A = "sdnets x" and B = "sdnets y" and D = "src p" and F = "dest p" in tndFalse) + by (auto simp: dest_in_sdnets src_in_sdnets sdnets_in_subnets disjSD_2_def) lemma list2FWpolicy_eq: "zs \ [] \ Cp (list2FWpolicy (x \ y # z)) p = Cp (x \ list2FWpolicy (y # z)) p" -by (metis ConcAssoc l2p_aux list2FWpolicy.simps(2)) - - + by (metis ConcAssoc l2p_aux list2FWpolicy.simps(2)) lemma dom_sep[rule_format]: "x \ dom (Cp (list2FWpolicy p)) \ x \ dom (Cp (list2FWpolicy(separate p)))" @@ -1198,87 +1110,84 @@ proof (induct p rule: separate.induct,simp_all, goal_cases) next case (2 v va vb y z) then show ?case apply (intro impI conjI,simp) - apply (case_tac "x \ dom (Cp (AllowPortFromTo v va vb))") - apply (metis CConcStartA domIff l2p_aux2 list2FWpolicyconc not_Cons_self ) - apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (y #z)))") - apply (metis CConcStartA Cdom2 InDomConc domIff l2p_aux2 list2FWpolicyconc nlpaux) - apply (simp add: dom_def Cp.simps, simp_all) + apply (case_tac "x \ dom (Cp (AllowPortFromTo v va vb))") + apply (metis CConcStartA domIff l2p_aux2 list2FWpolicyconc not_Cons_self ) + apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (y #z)))") + apply (metis CConcStartA Cdom2 InDomConc domIff l2p_aux2 list2FWpolicyconc nlpaux) + apply (simp add: dom_def Cp.simps, simp_all) apply (case_tac "x \ dom (Cp (AllowPortFromTo v va vb))", simp_all) apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (y #z)))",simp) - apply (metis Conc_not_MT InDomConc sepnMT) + apply (metis Conc_not_MT InDomConc sepnMT) apply (metis domIff nlpaux) done next case (3 v va y z) then show ?case apply (intro conjI impI,simp) - apply (drule mp) - apply (case_tac "x \ dom (Cp ((v \ va)))") - apply (metis Cp.simps(4) CConcStartA ConcAssoc domIff list2FWpolicy2list list2FWpolicyconc p2lNmt) - defer 1 - apply simp_all - apply (case_tac "x \ dom (Cp ((v \ va)))",simp_all) - apply (drule mp) - apply (simp add: Cp.simps dom_def) - apply (metis InDomConc list.simps(2)sepnMT) + apply (drule mp) + apply (case_tac "x \ dom (Cp ((v \ va)))") + apply (metis Cp.simps(4) CConcStartA ConcAssoc domIff list2FWpolicy2list list2FWpolicyconc p2lNmt) + defer 1 + apply simp_all + apply (case_tac "x \ dom (Cp ((v \ va)))",simp_all) + apply (drule mp) + apply (simp add: Cp.simps dom_def) + apply (metis InDomConc list.simps(2)sepnMT) apply (subgoal_tac "x \ dom (Cp (list2FWpolicy (y#z)))") - apply (case_tac "x \ dom (Cp y)",simp_all) - apply (metis CConcStartA Cdom2 ConcAssoc domIff) - apply (metis InDomConc domIff l2p_aux2 list2FWpolicyconc nlpaux) + apply (case_tac "x \ dom (Cp y)",simp_all) + apply (metis CConcStartA Cdom2 ConcAssoc domIff) + apply (metis InDomConc domIff l2p_aux2 list2FWpolicyconc nlpaux) apply (case_tac "x \ dom (Cp y)",simp_all) by (metis domIff nlpaux) qed - + lemma domdConcStart[rule_format]: "x \ dom (Cp (list2FWpolicy (a#b))) \ x \ dom (Cp (list2FWpolicy b)) \ x \ dom (Cp (a))" -by (induct b, simp_all) (auto simp: PLemmas) - - + by (induct b, simp_all) (auto simp: PLemmas) + lemma sep_dom2_aux: "x \ dom (Cp (list2FWpolicy (a \ y # z))) \ x \ dom (Cp (a \ list2FWpolicy (y # z)))" -by auto (metis list2FWpolicy_eq p2lNmt) - - + by auto (metis list2FWpolicy_eq p2lNmt) + lemma sep_dom2_aux2: -"(x \ dom (Cp (list2FWpolicy (separate (y # z)))) \ x \ dom (Cp (list2FWpolicy (y # z)))) \ + "(x \ dom (Cp (list2FWpolicy (separate (y # z)))) \ x \ dom (Cp (list2FWpolicy (y # z)))) \ x \ dom (Cp (list2FWpolicy (a # separate (y # z)))) \ x \ dom (Cp (list2FWpolicy (a \ y # z)))" -by (metis CConcStartA InDomConc domdConcStart list.simps(2) list2FWpolicy.simps(2) list2FWpolicyconc) - - + by (metis CConcStartA InDomConc domdConcStart list.simps(2) list2FWpolicy.simps(2) list2FWpolicyconc) + lemma sep_dom2[rule_format]: - "x \ dom (Cp (list2FWpolicy (separate p))) \ x \ dom (Cp (list2FWpolicy( p)))" -by (rule separate.induct) (simp_all add: sep_dom2_aux sep_dom2_aux2) - + "x \ dom (Cp (list2FWpolicy (separate p))) \ x \ dom (Cp (list2FWpolicy( p)))" + by (rule separate.induct) (simp_all add: sep_dom2_aux sep_dom2_aux2) + lemma sepDom: "dom (Cp (list2FWpolicy p)) = dom (Cp (list2FWpolicy (separate p)))" -by (rule equalityI) (rule subsetI, (erule dom_sep|erule sep_dom2))+ + by (rule equalityI) (rule subsetI, (erule dom_sep|erule sep_dom2))+ lemma C_eq_s_ext[rule_format]: "p \ [] \ Cp (list2FWpolicy (separate p)) a = Cp (list2FWpolicy p) a " proof (induct rule: separate.induct, goal_cases) case (1 x) thus ?case - apply (cases "x = []",simp_all) - apply (cases "a \ dom (Cp (list2FWpolicy x))") - apply (subgoal_tac "a \ dom (Cp (list2FWpolicy (separate x)))") - apply (metis Cdom2 list2FWpolicyconc sepDom sepnMT) - apply (metis sepDom) - by (metis nlpaux sepDom list2FWpolicyconc sepnMT) + apply (cases "x = []",simp_all) + apply (cases "a \ dom (Cp (list2FWpolicy x))") + apply (subgoal_tac "a \ dom (Cp (list2FWpolicy (separate x)))") + apply (metis Cdom2 list2FWpolicyconc sepDom sepnMT) + apply (metis sepDom) + by (metis nlpaux sepDom list2FWpolicyconc sepnMT) next case (2 v va y z) thus ?case - apply (cases "z = []",simp_all) - apply (intro conjI impI|simp)+ - apply (simp add: PLemmas(8) UPFDefs(8) list2FWpolicyconc sepnMT) - by (metis (mono_tags, lifting) Conc_not_MT Cdom2 list2FWpolicy_eq nlpaux sepDom l2p_aux sepnMT) + apply (cases "z = []",simp_all) + apply (intro conjI impI|simp)+ + apply (simp add: PLemmas(8) UPFDefs(8) list2FWpolicyconc sepnMT) + by (metis (mono_tags, lifting) Conc_not_MT Cdom2 list2FWpolicy_eq nlpaux sepDom l2p_aux sepnMT) next case (3 v va vb y z) thus ?case - apply (cases "z = []", simp_all) + apply (cases "z = []", simp_all) apply (simp add: PLemmas(8) UPFDefs(8) list2FWpolicyconc sepnMT) - by (metis (no_types, hide_lams) Conc_not_MT Cdom2 nlpaux domIff l2p_aux sepnMT) + by (metis (no_types, hide_lams) Conc_not_MT Cdom2 nlpaux domIff l2p_aux sepnMT) next case (4 v va y z) thus ?case - apply (cases "z = []", simp_all) - apply (simp add: PLemmas(8) UPFDefs(8) l2p_aux sepnMT) - by (metis (no_types, lifting) ConcAssoc PLemmas(8) UPFDefs(8) list.distinct(1) - list2FWpolicyconc sepnMT) + apply (cases "z = []", simp_all) + apply (simp add: PLemmas(8) UPFDefs(8) l2p_aux sepnMT) + by (metis (no_types, lifting) ConcAssoc PLemmas(8) UPFDefs(8) list.distinct(1) + list2FWpolicyconc sepnMT) next case 5 thus ?case by simp next @@ -1290,23 +1199,17 @@ next qed lemma C_eq_s: "p \ [] \ Cp (list2FWpolicy (separate p)) = Cp (list2FWpolicy p)" -by (rule ext) (simp add: C_eq_s_ext) + 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\set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ + "DenyAll\set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ Cp (list2FWpolicy (separate (sort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l))) = Cp p" -by (simp add: C_eq_All_untilSorted_withSimps C_eq_s wellformed1_alternative_sorted wp1ID wp1n_RS2) - + by (simp add: C_eq_All_untilSorted_withSimps C_eq_s wellformed1_alternative_sorted wp1ID wp1n_RS2) lemma C_eq_until_separatedQ: "DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ @@ -1315,8 +1218,7 @@ lemma C_eq_until_separatedQ: removeShadowRules2(remdups (rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l))) = Cp p" -by (simp add: C_eq_All_untilSorted_withSimpsQ C_eq_s setnMT wp1ID wp1n_RS2) - + by (simp add: C_eq_All_untilSorted_withSimpsQ C_eq_s setnMT wp1ID wp1n_RS2) lemma domID[rule_format]: "p \ [] \ x \ dom(Cp(list2FWpolicy p)) \ x \ dom (Cp(list2FWpolicy(insertDenies p)))" @@ -1324,63 +1226,61 @@ proof(induct p) case Nil then show ?case by simp next case (Cons a p) then show ?case - proof(cases "p=[]", goal_cases) - case 1 then show ?case - apply(simp) apply(rule impI) - apply (cases a, simp_all) - apply (simp_all add: Cp.simps dom_def)+ - by auto - next - case 2 then show ?case - proof(cases "x \ dom(Cp(list2FWpolicy p))", goal_cases) - case 1 then show ?case - apply simp apply (rule impI) - apply (cases a, simp_all) - apply (metis InDomConc idNMT) - apply (rule InDomConc, simp_all add: idNMT)+ - done - next - case 2 then show ?case - apply simp apply (rule impI) - proof(cases "x \ dom (Cp (list2FWpolicy (insertDenies p)))", goal_cases) - case 1 then show ?case - proof(induct a) - case DenyAll then show ?case by simp - next - case (DenyAllFromTo src dest) then show ?case - by simp (rule InDomConc, simp add: idNMT) - next - case (AllowPortFromTo src dest port) then show ?case - by simp (rule InDomConc, simp add: idNMT) - next - case (Conc _ _) then show ?case - by simp(rule InDomConc, simp add: idNMT) - qed - next - case 2 then show ?case - proof (induct a) - case DenyAll then show ?case by simp - next - case (DenyAllFromTo src dest) then show ?case - by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) - next - case (AllowPortFromTo src dest port) then show ?case - by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) - next - case (Conc _ _) then show ?case - by simp (metis CConcStartA Cdom2 Conc(5) ConcAssoc domIff domdConcStart) - qed - qed - qed + proof(cases "p=[]", goal_cases) + case 1 then show ?case + apply(simp) apply(rule impI) + apply (cases a, simp_all) + apply (simp_all add: Cp.simps dom_def)+ + by auto + next + case 2 then show ?case + proof(cases "x \ dom(Cp(list2FWpolicy p))", goal_cases) + case 1 then show ?case + apply simp apply (rule impI) + apply (cases a, simp_all) + apply (metis InDomConc idNMT) + apply (rule InDomConc, simp_all add: idNMT)+ + done + next + case 2 then show ?case + apply simp apply (rule impI) + proof(cases "x \ dom (Cp (list2FWpolicy (insertDenies p)))", goal_cases) + case 1 then show ?case + proof(induct a) + case DenyAll then show ?case by simp + next + case (DenyAllFromTo src dest) then show ?case + by simp (rule InDomConc, simp add: idNMT) + next + case (AllowPortFromTo src dest port) then show ?case + by simp (rule InDomConc, simp add: idNMT) + next + case (Conc _ _) then show ?case + by simp(rule InDomConc, simp add: idNMT) + qed + next + case 2 then show ?case + proof (induct a) + case DenyAll then show ?case by simp + next + case (DenyAllFromTo src dest) then show ?case + by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) + next + case (AllowPortFromTo src dest port) then show ?case + by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) + next + case (Conc _ _) then show ?case + by simp (metis CConcStartA Cdom2 Conc(5) ConcAssoc domIff domdConcStart) + qed + qed qed + qed qed - - + lemma DA_is_deny: "x \ dom (Cp (DenyAllFromTo a b \ DenyAllFromTo b a \ DenyAllFromTo a b)) \ Cp (DenyAllFromTo a b\DenyAllFromTo b a \ DenyAllFromTo a b) x = Some (deny ())" -by (case_tac "x \ dom (Cp (DenyAllFromTo a b))") (simp_all add: PLemmas split: if_splits) - + by (case_tac "x \ dom (Cp (DenyAllFromTo a b))") (simp_all add: PLemmas split: if_splits) lemma iDdomAux[rule_format]: "p \ [] \ x \ dom (Cp (list2FWpolicy p)) \ @@ -1390,95 +1290,94 @@ proof (induct p) case Nil thus ?case by simp next case (Cons y ys) thus ?case - proof (cases y) - case DenyAll then show ?thesis by simp + proof (cases y) + case DenyAll then show ?thesis by simp + next + case (DenyAllFromTo a b) then show ?thesis + using DenyAllFromTo Cons apply simp + apply (rule impI)+ + proof (cases "ys = []", goal_cases) + case 1 then show ?case by (simp add: DA_is_deny) next - case (DenyAllFromTo a b) then show ?thesis - using DenyAllFromTo Cons apply simp - apply (rule impI)+ - proof (cases "ys = []", goal_cases) - case 1 then show ?case by (simp add: DA_is_deny) - next - case 2 then show ?case - apply simp - apply (drule mp) - apply (metis DenyAllFromTo InDomConc ) - apply (cases "x \ dom (Cp (list2FWpolicy (insertDenies ys)))",simp_all) - apply (metis Cdom2 DenyAllFromTo idNMT list2FWpolicyconc) - apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b \ + case 2 then show ?case + apply simp + apply (drule mp) + apply (metis DenyAllFromTo InDomConc ) + apply (cases "x \ dom (Cp (list2FWpolicy (insertDenies ys)))",simp_all) + apply (metis Cdom2 DenyAllFromTo idNMT list2FWpolicyconc) + apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b \ DenyAllFromTo b a \ DenyAllFromTo a b#insertDenies ys)) x = Cp ((DenyAllFromTo a b \ DenyAllFromTo b a \ DenyAllFromTo a b)) x ") - apply (metis DA_is_deny DenyAllFromTo domdConcStart) - apply (metis DenyAllFromTo l2p_aux2 list2FWpolicyconc nlpaux) - done - qed + apply (metis DA_is_deny DenyAllFromTo domdConcStart) + apply (metis DenyAllFromTo l2p_aux2 list2FWpolicyconc nlpaux) + done + qed + next + case (AllowPortFromTo a b c) then show ?thesis using Cons AllowPortFromTo + proof (cases "ys = []", goal_cases) + case 1 then show ?case + apply (simp,intro impI) + apply (subgoal_tac "x \ dom (Cp (DenyAllFromTo a b \ DenyAllFromTo b a))") + apply (auto simp: PLemmas split: if_splits) + done next - case (AllowPortFromTo a b c) then show ?thesis using Cons AllowPortFromTo - proof (cases "ys = []", goal_cases) - case 1 then show ?case - apply (simp,intro impI) - apply (subgoal_tac "x \ dom (Cp (DenyAllFromTo a b \ DenyAllFromTo b a))") - apply (auto simp: PLemmas split: if_splits) - done - next - case 2 then show ?case - apply (simp, intro impI) - apply (drule mp) - apply (metis AllowPortFromTo InDomConc) - apply (cases "x \ dom (Cp (list2FWpolicy (insertDenies ys)))",simp_all) - apply (metis AllowPortFromTo Cdom2 idNMT list2FWpolicyconc) - apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b \ + case 2 then show ?case + apply (simp, intro impI) + apply (drule mp) + apply (metis AllowPortFromTo InDomConc) + apply (cases "x \ dom (Cp (list2FWpolicy (insertDenies ys)))",simp_all) + apply (metis AllowPortFromTo Cdom2 idNMT list2FWpolicyconc) + apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b \ DenyAllFromTo b a \ AllowPortFromTo a b c#insertDenies ys)) x = Cp ((DenyAllFromTo a b \ DenyAllFromTo b a)) x ") - apply (auto simp: PLemmas split: if_splits)[1] - by (metis AllowPortFromTo CConcStartA ConcAssoc idNMT list2FWpolicyconc nlpaux) - qed - next - case (Conc a b) then show ?thesis - proof (cases "ys = []", goal_cases) - case 1 then show ?case - apply(simp,intro impI) - apply (subgoal_tac "x \ dom (Cp (DenyAllFromTo (first_srcNet a) (first_destNet a) \ + apply (auto simp: PLemmas split: if_splits)[1] + by (metis AllowPortFromTo CConcStartA ConcAssoc idNMT list2FWpolicyconc nlpaux) + qed + next + case (Conc a b) then show ?thesis + proof (cases "ys = []", goal_cases) + case 1 then show ?case + apply(simp,intro impI) + apply (subgoal_tac "x \ dom (Cp (DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a)))") - by (auto simp: PLemmas split: if_splits) - next - case 2 then show ?case - apply(simp,intro impI) - apply(cases "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") - apply (metis Cdom2 Conc Cons InDomConc idNMT list2FWpolicyconc) - apply (subgoal_tac "Cp (list2FWpolicy(DenyAllFromTo (first_srcNet a)(first_destNet a) \ + by (auto simp: PLemmas split: if_splits) + next + case 2 then show ?case + apply(simp,intro impI) + apply(cases "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") + apply (metis Cdom2 Conc Cons InDomConc idNMT list2FWpolicyconc) + apply (subgoal_tac "Cp (list2FWpolicy(DenyAllFromTo (first_srcNet a)(first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a)\ a \ b#insertDenies ys)) x = Cp ((DenyAllFromTo(first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a)(first_srcNet a) \ a \ b)) x") - apply simp - defer 1 - apply (metis Conc l2p_aux2 list2FWpolicyconc nlpaux) - apply (subgoal_tac "Cp((DenyAllFromTo(first_srcNet a)(first_destNet a) \ + apply simp + defer 1 + apply (metis Conc l2p_aux2 list2FWpolicyconc nlpaux) + apply (subgoal_tac "Cp((DenyAllFromTo(first_srcNet a)(first_destNet a) \ DenyAllFromTo (first_destNet a)(first_srcNet a)\ a \ b)) x = Cp((DenyAllFromTo (first_srcNet a)(first_destNet a)\ DenyAllFromTo (first_destNet a) (first_srcNet a))) x ") - apply simp - defer 1 - apply (metis CConcStartA Conc ConcAssoc nlpaux) - by (auto simp: PLemmas split: if_splits) - qed + apply simp + defer 1 + apply (metis CConcStartA Conc ConcAssoc nlpaux) + by (auto simp: PLemmas split: if_splits) qed + qed qed lemma iD_isD[rule_format]: "p \ [] \ x \ dom (Cp (list2FWpolicy p)) \ Cp (DenyAll \ list2FWpolicy (insertDenies p)) x = Cp DenyAll x" -apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies p)))") -apply (simp add: Cp.simps(1) Cdom2 iDdomAux deny_all_def) -using NormalisationIPPProofs.nlpaux by blast - + apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies p)))") + apply (simp add: Cp.simps(1) Cdom2 iDdomAux deny_all_def) + using NormalisationIPPProofs.nlpaux + by blast lemma inDomConc: "x\dom (Cp a) \ x\dom (Cp (list2FWpolicy p)) \ x \ dom (Cp (list2FWpolicy(a#p)))" -by (metis domdConcStart) - + by (metis domdConcStart) lemma domsdisj[rule_format]: "p \ [] \ (\ x s. s \ set p \ x \ dom (Cp A) \ x \ dom (Cp s)) \ y \ dom (Cp A) \ @@ -1489,128 +1388,124 @@ next case (Cons a p) show ?case apply (case_tac "p = []", simp) apply (rule_tac x = y in spec) - apply (simp add: split_tupled_all) + apply (simp add: split_tupled_all) by (metis Cons.hyps inDomConc list.set_intros(1) list.set_intros(2)) qed - - + lemma isSepaux: "p \ [] \ noDenyAll (a#p) \ separated (a # p) \ x \ dom (Cp (DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a)) \ x \ dom (Cp (list2FWpolicy p))" -apply (rule_tac A = "(DenyAllFromTo (first_srcNet a) (first_destNet a) \ + apply (rule_tac A = "(DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a)" in domsdisj, simp_all) -apply (rule notI) -apply (rule_tac p = xa and x ="(DenyAllFromTo (first_srcNet a) (first_destNet a) \ + apply (rule notI) + apply (rule_tac p = xa and x ="(DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a)" - and y = s in disjSD_no_p_in_both, simp_all) -using disjSD2aux noDA apply blast -using noDA by blast - + and y = s in disjSD_no_p_in_both, simp_all) + using disjSD2aux noDA apply blast + using noDA + by blast lemma none_MT_rulessep[rule_format]: "none_MT_rules Cp p \ none_MT_rules Cp (separate p)" -by (induct p rule: separate.induct) (simp_all add: Cp.simps map_add_le_mapE map_le_antisym) - - + by (induct p rule: separate.induct) (simp_all add: Cp.simps map_add_le_mapE map_le_antisym) + lemma dom_id: "noDenyAll (a#p) \ separated (a#p) \ p \ [] \ x \ dom (Cp (list2FWpolicy p)) \ x \ dom (Cp (a)) \ x \ dom (Cp (list2FWpolicy (insertDenies p)))" -apply (rule_tac a = a in isSepaux, simp_all) - using idNMT apply blast + apply (rule_tac a = a in isSepaux, simp_all) + using idNMT apply blast using noDAID apply blast - using id_aux4 noDA1eq sepNetsID apply blast -by (simp add: NormalisationIPPProofs.Cdom2 domIff) - + using id_aux4 noDA1eq sepNetsID apply blast + by (simp add: NormalisationIPPProofs.Cdom2 domIff) lemma C_eq_iD_aux2[rule_format]: - "noDenyAll1 p \ separated p\ p \ []\ x \ dom (Cp (list2FWpolicy p))\ + "noDenyAll1 p \ separated p\ p \ []\ x \ dom (Cp (list2FWpolicy p))\ Cp(list2FWpolicy (insertDenies p)) x = Cp(list2FWpolicy p) x" proof (induct p) case Nil thus ?case by simp next case (Cons y ys) thus ?case using Cons proof (cases y) - case DenyAll thus ?thesis using Cons DenyAll apply simp + case DenyAll thus ?thesis using Cons DenyAll apply simp apply (case_tac "ys = []", simp_all) - apply (case_tac "x \ dom (Cp (list2FWpolicy ys))",simp_all) + apply (case_tac "x \ dom (Cp (list2FWpolicy ys))",simp_all) apply (metis Cdom2 domID idNMT list2FWpolicyconc noDA1eq) apply (metis DenyAll iD_isD idNMT list2FWpolicyconc nlpaux) done -next - case (DenyAllFromTo a b) thus ?thesis - using Cons apply simp - apply (rule impI|rule allI|rule conjI|simp)+ + next + case (DenyAllFromTo a b) thus ?thesis + using Cons apply simp + apply (rule impI|rule allI|rule conjI|simp)+ apply (case_tac "ys = []", simp_all) - apply (metis Cdom2 ConcAssoc DenyAllFromTo) - apply (case_tac "x \ dom (Cp (list2FWpolicy ys))", simp_all) - apply (drule mp) - apply (metis noDA1eq) - apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") - apply (metis Cdom2 DenyAllFromTo idNMT list2FWpolicyconc) - apply (metis domID) - apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") - apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b \ DenyAllFromTo b a \ - DenyAllFromTo a b # insertDenies ys)) x = Some (deny ())") - apply simp_all - apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b # ys)) x = - Cp ((DenyAllFromTo a b)) x") - apply (simp add: PLemmas, simp split: if_splits) - apply (metis list2FWpolicyconc nlpaux) - apply (metis Cdom2 DenyAllFromTo iD_isD iDdomAux idNMT list2FWpolicyconc) - apply (metis Cdom2 DenyAllFromTo domIff idNMT list2FWpolicyconc nlpaux) - done -next - case (AllowPortFromTo a b c) thus ?thesis - using AllowPortFromTo Cons apply simp - apply (rule impI|rule allI|rule conjI|simp)+ - apply (case_tac "ys = []", simp_all) - apply (metis Cdom2 ConcAssoc AllowPortFromTo) - apply (case_tac "x \ dom (Cp (list2FWpolicy ys))",simp_all) - apply (drule mp) - apply (metis noDA1eq) - apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") - apply (metis Cdom2 AllowPortFromTo idNMT list2FWpolicyconc) - apply (metis domID) - apply (subgoal_tac "x \ dom (Cp (AllowPortFromTo a b c))") - apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))", simp_all) - apply (metis AllowPortFromTo Cdom2 ConcAssoc l2p_aux2 list2FWpolicyconc nlpaux) - apply (meson Combinators.distinct(3) FWNormalisationCore.member.simps(4) NormalisationIPPProofs.dom_id noDenyAll.simps(1) separated.simps(1)) - apply (metis AllowPortFromTo domdConcStart) - done -next - case (Conc a b) thus ?thesis - using Cons Conc apply simp - apply (intro impI allI conjI|simp)+ - apply (case_tac "ys = []",simp_all) - apply (metis Cdom2 ConcAssoc Conc) - apply (case_tac "x \ dom (Cp (list2FWpolicy ys))",simp_all) - apply (drule mp) - apply (metis noDA1eq) - apply (case_tac "x \ dom (Cp (a \ b))") - apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))", simp_all) - apply (subst list2FWpolicyconc) - apply (rule idNMT, simp) + apply (metis Cdom2 ConcAssoc DenyAllFromTo) + apply (case_tac "x \ dom (Cp (list2FWpolicy ys))", simp_all) + apply (drule mp) + apply (metis noDA1eq) + apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") + apply (metis Cdom2 DenyAllFromTo idNMT list2FWpolicyconc) apply (metis domID) - apply (metis Cdom2 Conc idNMT list2FWpolicyconc) - apply (metis Cdom2 Conc domIff idNMT list2FWpolicyconc ) - apply (case_tac "x \ dom (Cp (a \ b))") - apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))", simp_all) - apply (subst list2FWpolicyconc) - apply (rule idNMT, simp) - apply (metis Cdom2 Conc ConcAssoc list2FWpolicyconc nlpaux) - apply (metis (lifting, no_types) FWNormalisationCore.member.simps(1) NormalisationIPPProofs.dom_id noDenyAll.simps(1) separated.simps(1)) - apply (metis Conc domdConcStart) - done + apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") + apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b \ DenyAllFromTo b a \ + DenyAllFromTo a b # insertDenies ys)) x = Some (deny ())") + apply simp_all + apply (subgoal_tac "Cp (list2FWpolicy (DenyAllFromTo a b # ys)) x = + Cp ((DenyAllFromTo a b)) x") + apply (simp add: PLemmas, simp split: if_splits) + apply (metis list2FWpolicyconc nlpaux) + apply (metis Cdom2 DenyAllFromTo iD_isD iDdomAux idNMT list2FWpolicyconc) + apply (metis Cdom2 DenyAllFromTo domIff idNMT list2FWpolicyconc nlpaux) + done + next + case (AllowPortFromTo a b c) thus ?thesis + using AllowPortFromTo Cons apply simp + apply (rule impI|rule allI|rule conjI|simp)+ + apply (case_tac "ys = []", simp_all) + apply (metis Cdom2 ConcAssoc AllowPortFromTo) + apply (case_tac "x \ dom (Cp (list2FWpolicy ys))",simp_all) + apply (drule mp) + apply (metis noDA1eq) + apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))") + apply (metis Cdom2 AllowPortFromTo idNMT list2FWpolicyconc) + apply (metis domID) + apply (subgoal_tac "x \ dom (Cp (AllowPortFromTo a b c))") + apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))", simp_all) + apply (metis AllowPortFromTo Cdom2 ConcAssoc l2p_aux2 list2FWpolicyconc nlpaux) + apply (meson Combinators.distinct(3) FWNormalisationCore.member.simps(4) NormalisationIPPProofs.dom_id noDenyAll.simps(1) separated.simps(1)) + apply (metis AllowPortFromTo domdConcStart) + done + next + case (Conc a b) thus ?thesis + using Cons Conc apply simp + apply (intro impI allI conjI|simp)+ + apply (case_tac "ys = []",simp_all) + apply (metis Cdom2 ConcAssoc Conc) + apply (case_tac "x \ dom (Cp (list2FWpolicy ys))",simp_all) + apply (drule mp) + apply (metis noDA1eq) + apply (case_tac "x \ dom (Cp (a \ b))") + apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))", simp_all) + apply (subst list2FWpolicyconc) + apply (rule idNMT, simp) + apply (metis domID) + apply (metis Cdom2 Conc idNMT list2FWpolicyconc) + apply (metis Cdom2 Conc domIff idNMT list2FWpolicyconc ) + apply (case_tac "x \ dom (Cp (a \ b))") + apply (case_tac "x \ dom (Cp (list2FWpolicy (insertDenies ys)))", simp_all) + apply (subst list2FWpolicyconc) + apply (rule idNMT, simp) + apply (metis Cdom2 Conc ConcAssoc list2FWpolicyconc nlpaux) + apply (metis (lifting, no_types) FWNormalisationCore.member.simps(1) NormalisationIPPProofs.dom_id noDenyAll.simps(1) separated.simps(1)) + apply (metis Conc domdConcStart) + done qed qed lemma C_eq_iD: "separated p \ noDenyAll1 p \ wellformed_policy1_strong p \ Cp(list2FWpolicy (insertDenies p)) = Cp (list2FWpolicy p)" -by (rule ext) (metis CConcStartA C_eq_iD_aux2 DAAux wp1_alternative_not_mt wp1n_tl) - + by (rule ext) (metis CConcStartA C_eq_iD_aux2 DAAux wp1_alternative_not_mt wp1n_tl) lemma noDAsortQ[rule_format]: "noDenyAll1 p \ noDenyAll1 (qsort p l)" proof (cases "p") @@ -1618,49 +1513,42 @@ proof (cases "p") next case (Cons a list) show ?thesis apply (insert `p = a # list`, simp_all) - proof (cases "a = DenyAll") - case True - assume * : "a = DenyAll" - show "noDenyAll1(a # list) \ + proof (cases "a = DenyAll") + case True + assume * : "a = DenyAll" + show "noDenyAll1(a # list) \ noDenyAll1(qsort[y\list . \ smaller a y l] l @ a # qsort [y\list . smaller a y l] l)" - using noDAsortQ by fastforce - next - case False - assume * : "a \ DenyAll" - have ** : "noDenyAll1 (a # list) \ noDenyAll (a # list)" by(case_tac a,simp_all add:*) - show "noDenyAll1(a # list) \ + using noDAsortQ by fastforce + next + case False + assume * : "a \ DenyAll" + have ** : "noDenyAll1 (a # list) \ noDenyAll (a # list)" by(case_tac a,simp_all add:*) + show "noDenyAll1(a # list) \ noDenyAll1(qsort[y\list . \ smaller a y l] l @ a # qsort [y\list . smaller a y l] l)" - apply (insert *,rule impI) - apply (rule noDA1eq, frule **) - by (metis append_Cons append_Nil nDAeqSet qsort.simps(2) set_sortQ) - qed + apply (insert *,rule impI) + apply (rule noDA1eq, frule **) + by (metis append_Cons append_Nil nDAeqSet qsort.simps(2) set_sortQ) + qed qed -(*MOVE FORWARD*) - lemma NetsCollectedSortQ: "distinct p \noDenyAll1 p \ all_in_list p l \ singleCombinators p \ NetsCollected (qsort p l)" -by(metis C_eqLemmas_id(22)) - + 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\set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct (policy2list p)\ Cp (list2FWpolicy((insertDenies(separate(sort(removeShadowRules2 (remdups(rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l))))) = Cp p" -by (subst C_eq_iD,simp_all add: C_eqLemmas_id) (rule C_eq_until_separated, simp_all) - + by (subst C_eq_iD,simp_all add: C_eqLemmas_id) (rule C_eq_until_separated, simp_all) lemma C_eq_Until_InsertDeniesQ: "DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ @@ -1668,19 +1556,17 @@ lemma C_eq_Until_InsertDeniesQ: Cp (list2FWpolicy ((insertDenies (separate (qsort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l))))) = Cp p" -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) + 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) apply (rule C_eq_until_separatedQ) -by simp_all - + 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) - - + 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 \ [] \ Cp (list2FWpolicy p) x = Cp (list2FWpolicy (removeAllDuplicates p)) x" proof (induct p) @@ -1695,125 +1581,107 @@ next apply (subst list2FWpolicyconc) apply (rule rADnMT, simp) apply (subst Cdom2,simp) - apply (simp add: NormalisationIPPProofs.Cdom2 domIff) + apply (simp add: NormalisationIPPProofs.Cdom2 domIff) by (metis C_eq_RD_aux nlpaux domIff list2FWpolicyconc rADnMT) qed lemma C_eq_RAD: "p \ [] \ Cp (list2FWpolicy p) = Cp (list2FWpolicy (removeAllDuplicates p)) " -by (rule ext) (erule C_eq_RAD_aux) - + by (rule ext) (erule C_eq_RAD_aux) lemma C_eq_compile: -"DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ + "DenyAll \ set (policy2list p) \ all_in_list (policy2list p) l \ allNetsDistinct (policy2list p) \ Cp (list2FWpolicy (removeAllDuplicates (insertDenies (separate (sort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l))))) = Cp p" -by (metis C_eq_RAD C_eq_Until_InsertDenies removeAllDuplicates.simps(2)) - + by (metis C_eq_RAD C_eq_Until_InsertDenies removeAllDuplicates.simps(2)) lemma C_eq_compileQ: - "DenyAll\set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ + "DenyAll\set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ Cp (list2FWpolicy (removeAllDuplicates (insertDenies (separate (qsort (removeShadowRules2 (remdups (rm_MT_rules Cp (insertDeny (removeShadowRules1 (policy2list p)))))) l))))) = Cp p" -apply (subst C_eq_RAD[symmetric]) - apply (rule idNMT) -apply (metis WP1rd sepnMT sortnMTQ wellformed_policy1_strong.simps(1) wp1ID wp1n_RS2 wp1n_RS3) -apply (rule C_eq_Until_InsertDeniesQ, simp_all) -done - + apply (subst C_eq_RAD[symmetric]) + apply (rule idNMT) + apply (metis WP1rd sepnMT sortnMTQ wellformed_policy1_strong.simps(1) wp1ID wp1n_RS2 wp1n_RS3) + apply (rule C_eq_Until_InsertDeniesQ, simp_all) + done lemma C_eq_normalizePr: -"DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ + "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ all_in_list (policy2list p) (Nets_List p) \ Cp (list2FWpolicy (normalizePr p)) = Cp p" -unfolding normalizePrQ_def -by (simp add: C_eq_compile normalizePr_def) - - + unfolding normalizePrQ_def + by (simp add: C_eq_compile normalizePr_def) + lemma C_eq_normalizePrQ: -"DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ + "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ all_in_list (policy2list p) (Nets_List p) \ Cp (list2FWpolicy (normalizePrQ p)) = Cp p" -unfolding normalizePrQ_def -using C_eq_compileQ by auto - + unfolding normalizePrQ_def + using C_eq_compileQ by auto lemma domSubset3: "dom (Cp (DenyAll \ x)) = dom (Cp (DenyAll))" -by (simp add: PLemmas split_tupled_all split: option.splits) - - + by (simp add: PLemmas split_tupled_all split: option.splits) + lemma domSubset4: "dom (Cp (DenyAllFromTo x y \ DenyAllFromTo y x \ AllowPortFromTo x y dn)) = dom (Cp (DenyAllFromTo x y \ DenyAllFromTo y x))" -by (simp add: PLemmas split: option.splits decision.splits) auto - - + by (simp add: PLemmas split: option.splits decision.splits) auto + lemma domSubset5: - "dom (Cp (DenyAllFromTo x y \ DenyAllFromTo y x \ AllowPortFromTo y x dn)) = + "dom (Cp (DenyAllFromTo x y \ DenyAllFromTo y x \ AllowPortFromTo y x dn)) = dom (Cp (DenyAllFromTo x y \ DenyAllFromTo y x))" -by (simp add: PLemmas split: option.splits decision.splits) auto - - + by (simp add: PLemmas split: option.splits decision.splits) auto + lemma domSubset1: "dom (Cp (DenyAllFromTo one two \ DenyAllFromTo two one \ AllowPortFromTo one two dn \ x)) = dom (Cp (DenyAllFromTo one two \ DenyAllFromTo two one \ x))" -by (simp add: PLemmas allow_all_def deny_all_def split: option.splits decision.splits) auto - - + by (simp add: PLemmas allow_all_def deny_all_def split: option.splits decision.splits) auto + lemma domSubset2: - "dom (Cp (DenyAllFromTo one two \ DenyAllFromTo two one \ AllowPortFromTo two one dn \ x)) = + "dom (Cp (DenyAllFromTo one two \ DenyAllFromTo two one \ AllowPortFromTo two one dn \ x)) = dom (Cp (DenyAllFromTo one two \ DenyAllFromTo two one \ x))" -by (simp add: PLemmas allow_all_def deny_all_def split: option.splits decision.splits) auto - - + by (simp add: PLemmas allow_all_def deny_all_def split: option.splits decision.splits) auto + lemma ConcAssoc2: "Cp (X \ Y \ ((A \ B) \ D)) = Cp (X \ Y \ A \ B \ D)" -by (simp add: Cp.simps) - - + by (simp add: Cp.simps) + lemma ConcAssoc3: "Cp (X \ ((Y \ A) \ D)) = Cp (X \ Y \ A \ D)" -by (simp add: Cp.simps) - - + by (simp add: Cp.simps) + lemma RS3_NMT[rule_format]: "DenyAll \ set p \ rm_MT_rules Cp p \ []" -by (induct_tac p) (simp_all add: PLemmas) - - + by (induct_tac p) (simp_all add: PLemmas) + lemma norm_notMT: "DenyAll \ set (policy2list p) \ normalizePr p \ []" -unfolding normalizePrQ_def -by (simp add: DAiniD RS3_NMT RS2_NMT idNMT normalizePr_def rADnMT sepnMT sortnMT) - - + unfolding normalizePrQ_def + by (simp add: DAiniD RS3_NMT RS2_NMT idNMT normalizePr_def rADnMT sepnMT sortnMT) + lemma norm_notMTQ: "DenyAll \ set (policy2list p) \ normalizePrQ p \ []" -unfolding normalizePrQ_def -by (simp add: DAiniD RS3_NMT sortnMTQ RS2_NMT idNMT rADnMT sepnMT) - - + unfolding normalizePrQ_def + by (simp add: DAiniD RS3_NMT sortnMTQ RS2_NMT idNMT rADnMT sepnMT) + lemma domDA: "dom (Cp (DenyAll \ A)) = dom (Cp (DenyAll))" -by (rule domSubset3) - - + by (rule domSubset3) + lemmas domain_reasoningPr = domDA ConcAssoc2 domSubset1 domSubset2 - domSubset3 domSubset4 domSubset5 domSubsetDistr1 - domSubsetDistr2 domSubsetDistrA domSubsetDistrD coerc_assoc ConcAssoc - ConcAssoc3 - - + 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 \ dom (Cp a) \ Cp (list2policyR (a # list)) p = Cp a p" -by (induct "a # list" rule:list2policyR.induct) - (auto simp: Cp.simps dom_def map_add_def) - + by (induct "a # list" rule:list2policyR.induct) + (auto simp: Cp.simps dom_def map_add_def) + lemma list2policyR_End: "p \ dom (Cp a) \ Cp (list2policyR (a # list)) p = (Cp a \ list2policy (map Cp list)) p" -by (rule list2policyR.induct) - (simp_all add: Cp.simps dom_def map_add_def list2policy_def split: option.splits) - - + 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 \ [] \ Cp( list2policyR N) p = (list2policy (map Cp N)) p" proof (induct N) @@ -1823,137 +1691,115 @@ next apply (insert Cons.hyps, simp_all add: list2policy_def) by (metis list2policyR_End list2policyR_Start domStart list2policy_def) qed - - + lemma l2polR_eq: "N \ [] \ Cp( list2policyR N) = (list2policy (map Cp N))" -by (auto simp: list2policy_def l2polR_eq_el ) - - - - + by (auto simp: list2policy_def l2polR_eq_el ) + lemma list2FWpolicys_eq_el[rule_format]: - "Filter \ [] \ Cp (list2policyR Filter) p = Cp (list2FWpolicy (rev Filter)) p" -apply (induct_tac Filter) -apply simp_all -apply (case_tac "list = []") -apply simp_all -apply (case_tac "p \ dom (Cp a)") -apply simp_all -apply (rule list2policyR_Start) -apply simp_all -apply (subgoal_tac "Cp (list2policyR (a # list)) p = Cp (list2policyR list) p") -apply (subgoal_tac "Cp (list2FWpolicy (rev list @ [a])) p = Cp (list2FWpolicy (rev list)) p") -apply simp -apply (rule CConcStart2) -apply simp -apply simp -apply (case_tac list,simp_all) -apply (simp_all add: Cp.simps dom_def map_add_def) -done - + "Filter \ [] \ Cp (list2policyR Filter) p = Cp (list2FWpolicy (rev Filter)) p" + apply (induct_tac Filter) + apply simp_all + apply (case_tac "list = []") + apply simp_all + apply (case_tac "p \ dom (Cp a)") + apply simp_all + apply (rule list2policyR_Start) + apply simp_all + apply (subgoal_tac "Cp (list2policyR (a # list)) p = Cp (list2policyR list) p") + apply (subgoal_tac "Cp (list2FWpolicy (rev list @ [a])) p = Cp (list2FWpolicy (rev list)) p") + apply simp + apply (rule CConcStart2) + apply simp + apply simp + apply (case_tac list,simp_all) + apply (simp_all add: Cp.simps dom_def map_add_def) + done + lemma list2FWpolicys_eq: - "Filter \ [] \ + "Filter \ [] \ Cp (list2policyR Filter) = Cp (list2FWpolicy (rev Filter))" -by (rule ext, erule list2FWpolicys_eq_el) - - + by (rule ext, erule list2FWpolicys_eq_el) + lemma list2FWpolicys_eq_sym: - "Filter \ [] \ + "Filter \ [] \ Cp (list2policyR (rev Filter)) = Cp (list2FWpolicy Filter)" -by (metis list2FWpolicys_eq rev_is_Nil_conv rev_rev_ident) - - + by (metis list2FWpolicys_eq rev_is_Nil_conv rev_rev_ident) + lemma p_eq[rule_format]: "p \ [] \ list2policy (map Cp (rev p)) = Cp (list2FWpolicy p)" -by (metis l2polR_eq list2FWpolicys_eq_sym rev.simps(1) rev_rev_ident) - - + by (metis l2polR_eq list2FWpolicys_eq_sym rev.simps(1) rev_rev_ident) + lemma p_eq2[rule_format]: "normalizePr x \ [] \ Cp (list2FWpolicy (normalizePr x)) = Cp x \ list2policy (map Cp (rev (normalizePr x))) = Cp x" -by (simp add: p_eq) - - + by (simp add: p_eq) + lemma p_eq2Q[rule_format]: "normalizePrQ x \ [] \ Cp (list2FWpolicy (normalizePrQ x)) = Cp x \ list2policy (map Cp (rev (normalizePrQ x))) = Cp x" -by (simp add: p_eq) - - + by (simp add: p_eq) + lemma list2listNMT[rule_format]: "x \ [] \map sem x \ []" -by (case_tac x) (simp_all) - - + by (case_tac x) (simp_all) + lemma Norm_Distr2: - "r o_f ((P \\<^sub>2 (list2policy Q)) o d) = + "r o_f ((P \\<^sub>2 (list2policy Q)) o d) = (list2policy ((P \\<^sub>L Q) (op \\<^sub>2) r d))" -by (rule ext, rule Norm_Distr_2) - - + by (rule ext, rule Norm_Distr_2) + lemma NATDistr: "N \ [] \ F = Cp (list2policyR N) \ ((\ (x,y). x) o_f ((NAT \\<^sub>2 F) o (\ x. (x,x))) = (list2policy ( ((NAT \\<^sub>L (map Cp N)) (op \\<^sub>2) (\ (x,y). x) (\ x. (x,x))))))" -by (simp add: l2polR_eq) (rule ext,rule Norm_Distr_2) - - + by (simp add: l2polR_eq) (rule ext,rule Norm_Distr_2) + lemma C_eq_normalize_manual: "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ all_in_list (policy2list p) l \ Cp (list2FWpolicy (normalize_manual_orderPr p l)) = Cp p" -unfolding normalize_manual_orderPr_def -by(simp_all add:C_eq_compile) - - + unfolding normalize_manual_orderPr_def + by(simp_all add:C_eq_compile) + lemma p_eq2_manualQ[rule_format]: "normalize_manual_orderPrQ x l \ [] \ Cp (list2FWpolicy (normalize_manual_orderPrQ x l)) = Cp x \ list2policy (map Cp (rev (normalize_manual_orderPrQ x l))) = Cp x" -by (simp add: p_eq) - - + by (simp add: p_eq) + lemma norm_notMT_manualQ: "DenyAll \ set (policy2list p) \ normalize_manual_orderPrQ p l \ []" -by (simp add: DAiniD RS3_NMT sortnMTQ RS2_NMT idNMT normalize_manual_orderPrQ_def rADnMT sepnMT) - - + by (simp add: DAiniD RS3_NMT sortnMTQ RS2_NMT idNMT normalize_manual_orderPrQ_def rADnMT sepnMT) + lemma C_eq_normalizePr_manualQ: "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ all_in_list (policy2list p) l \ Cp (list2FWpolicy (normalize_manual_orderPrQ p l)) = Cp p" -by (simp add: normalize_manual_orderPrQ_def C_eq_compileQ) - - + by (simp add: normalize_manual_orderPrQ_def C_eq_compileQ) + lemma p_eq2_manual[rule_format]: "normalize_manual_orderPr x l \ [] \ Cp (list2FWpolicy (normalize_manual_orderPr x l)) = Cp x \ list2policy (map Cp (rev (normalize_manual_orderPr x l))) = Cp x" -by (simp add: p_eq) - - + by (simp add: p_eq) + lemma norm_notMT_manual: "DenyAll \ set (policy2list p) \ normalize_manual_orderPr p l \ []" -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. *} - + 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. +*} + lemma normalizePrNAT: "DenyAll \ set (policy2list Filter) \ allNetsDistinct (policy2list Filter) \ all_in_list (policy2list Filter) (Nets_List Filter) \ ((\ (x,y). x) o_f (((NAT \\<^sub>2 Cp Filter) o (\x. (x,x))))) = list2policy ((NAT \\<^sub>L (map Cp (rev (normalizePr Filter)))) (op \\<^sub>2) (\ (x,y). x) (\ x. (x,x)))" -by (simp add: C_eq_normalizePr NATDistr list2FWpolicys_eq_sym norm_notMT) - - + by (simp add: C_eq_normalizePr NATDistr list2FWpolicys_eq_sym norm_notMT) + lemma domSimpl[simp]: "dom (Cp (A \ DenyAll)) = dom (Cp (DenyAll))" -by (simp add: PLemmas) - - -end - + by (simp add: PLemmas) +end \ No newline at end of file diff --git a/FWNormalisation/NormalisationIntegerPortProof.thy b/FWNormalisation/NormalisationIntegerPortProof.thy index 8c104a7..b4675fc 100644 --- a/FWNormalisation/NormalisationIntegerPortProof.thy +++ b/FWNormalisation/NormalisationIntegerPortProof.thy @@ -36,125 +36,124 @@ *****************************************************************************) 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 \ B) \ D) = C(A \ (B \ D))" -by (simp add: C.simps) - + by (simp add: C.simps) lemma aux26[simp]: "twoNetsDistinct a b c d \ dom (C (AllowPortFromTo a b p)) \ 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]) \ wellformed_policy2 xs" -by (induct xs, simp_all) (case_tac "a", simp_all) + by (induct xs, simp_all) (case_tac "a", simp_all) - lemma Cdom2: "x \ dom(C b) \ C (a \ b) x = (C b) x" -by (auto simp: C.simps) + by (auto simp: C.simps) lemma wp2Conc[rule_format]: "wellformed_policy2 (x#xs) \ wellformed_policy2 xs" -by (case_tac "x",simp_all) + by (case_tac "x",simp_all) lemma DAimpliesMR_E[rule_format]: "DenyAll \ set p \ (\ r. applied_rule_rev C x p = Some r)" -apply (simp add: applied_rule_rev_def) -apply (rule_tac xs = p in rev_induct, simp_all) -by (metis C.simps(1) denyAllDom) + apply (simp add: applied_rule_rev_def) + apply (rule_tac xs = p in rev_induct, simp_all) + by (metis C.simps(1) denyAllDom) lemma DAimplieMR[rule_format]: "DenyAll \ set p \ applied_rule_rev C x p \ None" -by (auto intro: DAimpliesMR_E) + by (auto intro: DAimpliesMR_E) lemma MRList1[rule_format]: "x \ dom (C a) \ applied_rule_rev C x (b@[a]) = Some a" -by (simp add: applied_rule_rev_def) + by (simp add: applied_rule_rev_def) lemma MRList2: "x \ dom (C a) \ applied_rule_rev C x (c@b@[a]) = Some a" -by (simp add: applied_rule_rev_def) + by (simp add: applied_rule_rev_def) lemma MRList3: "x \ dom (C xa) \ applied_rule_rev C x (a @ b # xs @ [xa]) = applied_rule_rev C x (a @ b # xs)" -by (simp add: applied_rule_rev_def) + by (simp add: applied_rule_rev_def) lemma CConcEnd[rule_format]: "C a x = Some y \ C (list2FWpolicy (xs @ [a])) x = Some y" (is "?P xs") -apply (rule_tac P = ?P in list2FWpolicy.induct) -by (simp_all add:C.simps) + apply (rule_tac P = ?P in list2FWpolicy.induct) + by (simp_all add:C.simps) lemma CConcStartaux: " C a x = None \ (C aa ++ C a) x = C aa x" -by (simp add: PLemmas) + by (simp add: PLemmas) lemma CConcStart[rule_format]: "xs \ [] \ C a x = None \ C (list2FWpolicy (xs @ [a])) x = C (list2FWpolicy xs) x" -apply (rule list2FWpolicy.induct) -by (simp_all add: PLemmas) + apply (rule list2FWpolicy.induct) + by (simp_all add: PLemmas) lemma mrNnt[simp]: "applied_rule_rev C x p = Some a \ p \ []" -apply (simp add: applied_rule_rev_def) -by auto + apply (simp add: applied_rule_rev_def) + by auto lemma mr_is_C[rule_format]: "applied_rule_rev C x p = Some a \ C (list2FWpolicy (p)) x = C a x" -apply (simp add: applied_rule_rev_def) + apply (simp add: applied_rule_rev_def) apply (rule rev_induct,auto) - apply (metis CConcEnd) -apply (metis CConcEnd) -by (metis CConcStart applied_rule_rev_def mrNnt option.exhaust) + apply (metis CConcEnd) + apply (metis CConcEnd) + by (metis CConcStart applied_rule_rev_def mrNnt option.exhaust) lemma CConcStart2: "p \ [] \ x \ dom (C a) \ C (list2FWpolicy (p @ [a])) x = C (list2FWpolicy p) x" -by (erule CConcStart,simp add: PLemmas) + by (erule CConcStart,simp add: PLemmas) lemma CConcEnd1: "q @ p \ [] \ x \ dom (C a) \ C (list2FWpolicy (q @ p @ [a])) x = C (list2FWpolicy (q @ p)) x" -apply (subst lCdom2) -by (rule CConcStart2, simp_all) + apply (subst lCdom2) + by (rule CConcStart2, simp_all) lemma CConcEnd2[rule_format]: "x \ dom (C a) \ C (list2FWpolicy (xs @ [a])) x = C a x" (is "?P xs") -apply (rule_tac P = ?P in list2FWpolicy.induct) -by (auto simp:C.simps) - + apply (rule_tac P = ?P in list2FWpolicy.induct) + by (auto simp:C.simps) lemma bar3: "x \ dom (C (list2FWpolicy (xs @ [xa]))) \ x \ dom (C (list2FWpolicy xs)) \ x \ dom (C xa)" -by auto (metis CConcStart eq_Nil_appendI l2p_aux2 option.simps(3)) - + by auto (metis CConcStart eq_Nil_appendI l2p_aux2 option.simps(3)) lemma CeqEnd[rule_format,simp]: "a \ [] \ x \ dom (C (list2FWpolicy a)) \ C (list2FWpolicy(b@a)) x = (C (list2FWpolicy a)) x" -apply (rule rev_induct,simp_all) - apply (case_tac "xs \ []", simp_all) - apply (case_tac "x \ dom (C xa)") - apply (metis CConcEnd2 MRList2 mr_is_C ) -apply (metis CConcEnd1 CConcStart2 Nil_is_append_conv bar3 ) -apply (metis MRList2 eq_Nil_appendI mr_is_C ) -done + apply (rule rev_induct,simp_all) + apply (case_tac "xs \ []", simp_all) + apply (case_tac "x \ dom (C xa)") + apply (metis CConcEnd2 MRList2 mr_is_C ) + apply (metis CConcEnd1 CConcStart2 Nil_is_append_conv bar3 ) + apply (metis MRList2 eq_Nil_appendI mr_is_C ) + done lemma CConcStartA[rule_format,simp]: "x \ dom (C a) \ x \ dom (C (list2FWpolicy (a # b)))" (is "?P b") -apply (rule_tac P = ?P in list2FWpolicy.induct) - apply (simp_all add: C.simps) -done - + apply (rule_tac P = ?P in list2FWpolicy.induct) + apply (simp_all add: C.simps) + done lemma domConc: "x \ dom (C (list2FWpolicy b)) \ b \ [] \ x \ dom (C (list2FWpolicy (a @ b)))" -by (auto simp: PLemmas) + by (auto simp: PLemmas) lemma CeqStart[rule_format,simp]: "x\dom(C(list2FWpolicy a)) \ a\[] \ b\[] \ C(list2FWpolicy(b@a)) x = (C(list2FWpolicy b)) x" -apply (rule list2FWpolicy.induct,simp_all) - apply (auto simp: list2FWpolicyconc PLemmas) -done + apply (rule list2FWpolicy.induct,simp_all) + apply (auto simp: list2FWpolicyconc PLemmas) + done lemma C_eq_if_mr_eq2: "applied_rule_rev C x a = \r\ \ @@ -164,200 +163,191 @@ by (metis mr_is_C) lemma nMRtoNone[rule_format]: "p \ [] \ applied_rule_rev C x p = None \ C (list2FWpolicy p) x = None" -apply (rule rev_induct, simp_all) - apply (case_tac "xs = []", simp_all) -by (simp_all add: applied_rule_rev_def dom_def) + apply (rule rev_induct, simp_all) + apply (case_tac "xs = []", simp_all) + by (simp_all add: applied_rule_rev_def dom_def) lemma C_eq_if_mr_eq: "applied_rule_rev C x b = applied_rule_rev C x a \ a \ [] \ b \ [] \ C (list2FWpolicy a) x = C (list2FWpolicy b) x" -apply (cases "applied_rule_rev C x a = None", simp_all) - apply (subst nMRtoNone,simp_all) -apply (subst nMRtoNone, simp_all) -by (auto intro: C_eq_if_mr_eq2) - - + apply (cases "applied_rule_rev C x a = None", simp_all) + apply (subst nMRtoNone,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]) \ Some a \ x \ dom (C a)" -by (simp add: applied_rule_rev_def split: if_splits) + by (simp add: applied_rule_rev_def split: if_splits) lemma foo3a[rule_format]: "applied_rule_rev C x (a@[b]@c) = Some b \ r \ set c \ b \ set c \ x \ dom (C r)" -apply (rule rev_induct) - apply simp_all -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) - + apply (rule rev_induct) + apply simp_all + 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 \ p = DenyAll # ps \ applied_rule_rev C x p = \DenyAll\ \ r \ set ps \ x \ dom (C r)" -by (rule_tac a = "[]" and b = "DenyAll" and c = "ps" in foo3a, simp_all) + by (rule_tac a = "[]" and b = "DenyAll" and c = "ps" in foo3a, simp_all) lemma foo4[rule_format]: "set p = set s \ (\ r. r \ set p \ x \ dom (C r)) \ (\ r .r \ set s \ x \ dom (C r))" -by simp + by simp lemma foo5b[rule_format]: "x \ dom (C b) \ (\ r. r \ set c \ x \ dom (C r)) \ applied_rule_rev C x (b#c) = Some b" -apply (simp add: applied_rule_rev_def) -apply (rule_tac xs = c in rev_induct, simp_all) -done + apply (simp add: applied_rule_rev_def) + apply (rule_tac xs = c in rev_induct, simp_all) + done lemma mr_first: "x \ dom (C b) \ \r. r \ set c \ x \ dom (C r) \ s = b # c \ applied_rule_rev C x s = \b\" -by (simp add: foo5b) + by (simp add: foo5b) lemma mr_charn[rule_format]: "a \ set p \ (x \ dom (C a)) \ (\ r. r \ set p \ x \ dom (C r) \ r = a) \ applied_rule_rev C x p = Some a" unfolding applied_rule_rev_def -apply (rule_tac xs = p in rev_induct) - apply(simp) -by(safe,auto) + apply (rule_tac xs = p in rev_induct) + apply(simp) + by(safe,auto) lemma foo8: "\r. r \ set p \ x \ dom (C r) \ r = a \ set p = set s \ \r. r \ set s \ x \ dom (C r) \ r = a" -by auto + by auto lemma mrConcEnd[rule_format]: "applied_rule_rev C x (b # p) = Some a \ a \ b \ applied_rule_rev C x p = Some a" -apply (simp add: applied_rule_rev_def) -by (rule_tac xs = p in rev_induct,auto) + apply (simp add: applied_rule_rev_def) + by (rule_tac xs = p in rev_induct,auto) lemma wp3tl[rule_format]: "wellformed_policy3 p \ wellformed_policy3 (tl p)" -by (induct p, simp_all, case_tac a, simp_all) + by (induct p, simp_all, case_tac a, simp_all) lemma wp3Conc[rule_format]: "wellformed_policy3 (a#p) \ wellformed_policy3 p" -by (induct p, simp_all, case_tac a, simp_all) + by (induct p, simp_all, case_tac a, simp_all) lemma foo98[rule_format]: "applied_rule_rev C x (aa # p) = Some a \ x \ dom (C r) \ r \ set p \ a \ set p" -unfolding applied_rule_rev_def -apply (rule rev_induct, simp_all) -apply (case_tac "r = xa", simp_all) -done + unfolding applied_rule_rev_def + apply (rule rev_induct, simp_all) + apply (case_tac "r = xa", simp_all) + done lemma mrMTNone[simp]: "applied_rule_rev C x [] = None" -by (simp add: applied_rule_rev_def) + by (simp add: applied_rule_rev_def) lemma DAAux[simp]: "x \ dom (C DenyAll)" -by (simp add: dom_def PolicyCombinators.PolicyCombinators C.simps) + by (simp add: dom_def PolicyCombinators.PolicyCombinators C.simps) lemma mrSet[rule_format]: "applied_rule_rev C x p = Some r \ r \ set p" -unfolding applied_rule_rev_def -by (rule_tac xs=p in rev_induct, simp_all) + unfolding applied_rule_rev_def + by (rule_tac xs=p in rev_induct, simp_all) lemma mr_not_Conc: "singleCombinators p \ applied_rule_rev C x p \ Some (a\b)" -apply (auto simp: mrSet) -apply (drule mrSet) -apply (erule SCnotConc,simp) -done + apply (auto simp: mrSet) + apply (drule mrSet) + apply (erule SCnotConc,simp) + done lemma foo25[rule_format]: "wellformed_policy3 (p@[x]) \ wellformed_policy3 p" -by (induct p, simp_all, case_tac a, simp_all) + by (induct p, simp_all, case_tac a, simp_all) lemma mr_in_dom[rule_format]: "applied_rule_rev C x p = Some a \ x \ dom (C a)" -apply (rule_tac xs = p in rev_induct) -by (auto simp: applied_rule_rev_def) + apply (rule_tac xs = p in rev_induct) + by (auto simp: applied_rule_rev_def) lemma wp3EndMT[rule_format]: "wellformed_policy3 (p@[xs]) \ AllowPortFromTo a b po \ set p \ dom (C (AllowPortFromTo a b po)) \ dom (C xs) = {}" -apply (induct p,simp_all) -apply (intro impI,drule mp,erule wp3Conc) -by clarify auto + apply (induct p,simp_all) + apply (intro impI,drule mp,erule wp3Conc) + by clarify auto lemma foo29: "\dom (C a) \ {}; dom (C a) \ dom (C b) = {}\ \ a \ b" by auto lemma foo28: "AllowPortFromTo a b po \ set p \ dom (C (AllowPortFromTo a b po)) \ {} \ wellformed_policy3 (p @ [x]) \ x \ AllowPortFromTo a b po" -by (metis foo29 C.simps(3) wp3EndMT) + by (metis foo29 C.simps(3) wp3EndMT) lemma foo28a[rule_format]: "x \ dom (C a) \ dom (C a) \ {}" by auto lemma allow_deny_dom[simp]: "dom (C (AllowPortFromTo a b po)) \ dom (C (DenyAllFromTo a b))" -by (simp_all add: twoNetsDistinct_def netsDistinct_def PLemmas) auto + by (simp_all add: twoNetsDistinct_def netsDistinct_def PLemmas) auto lemma DenyAllowDisj: "dom (C (AllowPortFromTo a b p)) \ {} \ dom (C (DenyAllFromTo a b)) \ dom (C (AllowPortFromTo a b p)) \ {}" -by (metis Int_absorb1 allow_deny_dom) + by (metis Int_absorb1 allow_deny_dom) lemma foo31: "\r. r \ set p \ x \ dom (C r) \ r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll \ set p = set s \ \r. r \ set s \ x \ dom (C r) \ r=AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll" -by auto - - + by auto lemma wp1_auxa: "wellformed_policy1_strong p\(\ r. applied_rule_rev C x p = Some r)" -apply (rule DAimpliesMR_E) -by (erule wp1_aux1aa) - - + apply (rule DAimpliesMR_E) + by (erule wp1_aux1aa) lemma deny_dom[simp]: "twoNetsDistinct a b c d \ dom (C (DenyAllFromTo a b)) \ dom (C (DenyAllFromTo c d)) = {}" -apply (simp add: C.simps) -by (erule aux6) + apply (simp add: C.simps) + by (erule aux6) lemma domTrans: "dom a \ dom b \ dom b \ dom c = {} \ dom a \ dom c = {}" by auto lemma DomInterAllowsMT: "twoNetsDistinct a b c d \ dom (C (AllowPortFromTo a b p)) \ dom (C (AllowPortFromTo c d po)) = {}" -apply (case_tac "p = po", simp_all) -apply (rule_tac b = "C (DenyAllFromTo a b)" in domTrans, simp_all) -apply (metis domComm aux26 tNDComm) -by (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) auto + apply (case_tac "p = po", simp_all) + apply (rule_tac b = "C (DenyAllFromTo a b)" in domTrans, simp_all) + apply (metis domComm aux26 tNDComm) + by (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) auto lemma DomInterAllowsMT_Ports: "p \ po \ dom (C (AllowPortFromTo a b p)) \ dom (C (AllowPortFromTo c d po)) = {}" -by (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) auto + by (simp add: twoNetsDistinct_def netsDistinct_def PLemmas) auto lemma wellformed_policy3_charn[rule_format]: "singleCombinators p \ distinct p \ allNetsDistinct p \ wellformed_policy1 p \ wellformed_policy2 p \ wellformed_policy3 p" -apply (induct_tac p) -apply simp_all -apply (auto intro: singleCombinatorsConc ANDConc waux2 wp2Conc) -apply (case_tac a, simp_all, clarify) + apply (induct_tac p) + apply simp_all + apply (auto intro: singleCombinatorsConc ANDConc waux2 wp2Conc) + apply (case_tac a, simp_all, clarify) apply (case_tac r, simp_all) - apply (metis Int_commute) -apply (metis DomInterAllowsMT aux7aa DomInterAllowsMT_Ports) -apply (metis aux0_0 ) -done - - + apply (metis Int_commute) + apply (metis DomInterAllowsMT aux7aa DomInterAllowsMT_Ports) + apply (metis aux0_0 ) + done lemma DistinctNetsDenyAllow: "DenyAllFromTo b c \ set p \ AllowPortFromTo a d po \ set p \ allNetsDistinct p \ dom (C (DenyAllFromTo b c)) \ dom (C (AllowPortFromTo a d po)) \ {} \ b = a \ c = d" -unfolding allNetsDistinct_def -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 Int_commute ND0aux1 ND0aux3 NDComm aux26 twoNetsDistinct_def ND0aux2 ND0aux4) -done + unfolding allNetsDistinct_def + 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 Int_commute ND0aux1 ND0aux3 NDComm aux26 twoNetsDistinct_def ND0aux2 ND0aux4) + done lemma DistinctNetsAllowAllow: "AllowPortFromTo b c poo \ set p \ @@ -366,433 +356,399 @@ lemma DistinctNetsAllowAllow: dom (C (AllowPortFromTo b c poo)) \ dom (C (AllowPortFromTo a d po)) \ {} \ b = a \ c = d \ poo = po" unfolding allNetsDistinct_def -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) -done - - + 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) + done lemma WP2RS2[simp]: "singleCombinators p \ distinct p \ allNetsDistinct p \ wellformed_policy2 (removeShadowRules2 p)" proof (induct p) - case Nil thus ?case by simp + case Nil thus ?case by simp next - case (Cons x xs) - have wp_xs: "wellformed_policy2 (removeShadowRules2 xs)" - by (metis Cons ANDConc distinct.simps(2) singleCombinatorsConc) - show ?case - proof (cases x) - case DenyAll thus ?thesis using wp_xs by simp - next - case (DenyAllFromTo a b) thus ?thesis - using wp_xs Cons by (simp,metis DenyAllFromTo aux aux7 tNDComm deny_dom) - next - case (AllowPortFromTo a b p) thus ?thesis - using wp_xs by (simp, metis aux26 AllowPortFromTo Cons(4) aux aux7a tNDComm) - next - case (Conc a b) thus ?thesis - by (metis Conc Cons(2) singleCombinators.simps(2)) - qed + case (Cons x xs) + have wp_xs: "wellformed_policy2 (removeShadowRules2 xs)" + by (metis Cons ANDConc distinct.simps(2) singleCombinatorsConc) + show ?case + proof (cases x) + case DenyAll thus ?thesis using wp_xs by simp + next + case (DenyAllFromTo a b) thus ?thesis + using wp_xs Cons by (simp,metis DenyAllFromTo aux aux7 tNDComm deny_dom) + next + case (AllowPortFromTo a b p) thus ?thesis + using wp_xs by (simp, metis aux26 AllowPortFromTo Cons(4) aux aux7a tNDComm) + next + case (Conc a b) thus ?thesis + by (metis Conc Cons(2) singleCombinators.simps(2)) + qed qed - - lemma AD_aux: "AllowPortFromTo a b po \ set p \ DenyAllFromTo c d \ set p \ allNetsDistinct p \ singleCombinators p \ a \ c \ b \ d \ dom (C (AllowPortFromTo a b po)) \ dom (C (DenyAllFromTo c d)) = {}" -by (rule aux26,rule_tac x ="AllowPortFromTo a b po" and y = "DenyAllFromTo c d" in tND, auto) - - - + 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 \ all_in_list p l \ distinct p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy2 p" proof (induct p) - case Nil thus ?case by simp + case Nil thus ?case by simp next - case (Cons a p) thus ?case - proof (cases a) - case DenyAll thus ?thesis using Cons - by (auto intro: ANDConc singleCombinatorsConc sortedConcEnd) - next - case (DenyAllFromTo c d) thus ?thesis using Cons - apply simp - apply (intro impI conjI allI) - apply (rule deny_dom) - apply (auto intro: aux7 tNDComm ANDConc singleCombinatorsConc sortedConcEnd) - done - next - case (AllowPortFromTo c d e) thus ?thesis using Cons - apply simp - apply (intro impI conjI allI aux26) - apply (rule_tac x = "AllowPortFromTo c d e" and y = "DenyAllFromTo aa b" in tND) - apply (assumption,simp_all) - apply (subgoal_tac "smaller (AllowPortFromTo c d e) (DenyAllFromTo aa b) l") - apply (simp split: if_splits) - apply metis - apply (erule sorted_is_smaller, simp_all) - apply (metis bothNet.simps(2) in_list.simps(2) in_set_in_list) - by (auto intro: aux7 tNDComm ANDConc singleCombinatorsConc sortedConcEnd) - next - case (Conc a b) thus ?thesis using Cons by simp - qed + case (Cons a p) thus ?case + proof (cases a) + case DenyAll thus ?thesis using Cons + by (auto intro: ANDConc singleCombinatorsConc sortedConcEnd) + next + case (DenyAllFromTo c d) thus ?thesis using Cons + apply simp + apply (intro impI conjI allI) + apply (rule deny_dom) + apply (auto intro: aux7 tNDComm ANDConc singleCombinatorsConc sortedConcEnd) + done + next + case (AllowPortFromTo c d e) thus ?thesis using Cons + apply simp + apply (intro impI conjI allI aux26) + apply (rule_tac x = "AllowPortFromTo c d e" and y = "DenyAllFromTo aa b" in tND) + apply (assumption,simp_all) + apply (subgoal_tac "smaller (AllowPortFromTo c d e) (DenyAllFromTo aa b) l") + apply (simp split: if_splits) + apply metis + apply (erule sorted_is_smaller, simp_all) + apply (metis bothNet.simps(2) in_list.simps(2) in_set_in_list) + by (auto intro: aux7 tNDComm ANDConc singleCombinatorsConc sortedConcEnd) + next + case (Conc a b) thus ?thesis using Cons by simp + qed qed - - - lemma wellformed2_sorted[simp]: "all_in_list p l \ distinct p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy2 (sort p l)" -apply (rule sorted_WP2,erule sort_is_sorted, simp_all) -apply (auto elim: all_in_listSubset intro: SC3 singleCombinatorsConc sorted_insort) -done - + 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]: "\all_in_list p l; distinct p; allNetsDistinct p; singleCombinators p\ \ 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 - + 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) - + by (auto simp: PLemmas) lemma C_eq_RS1n: "C(list2FWpolicy (removeShadowRules1_alternative p)) = C(list2FWpolicy p)" proof (cases "p")print_cases - case Nil then show ?thesis apply(simp_all) - by (metis list2FWpolicy.simps(1) rSR1_eq removeShadowRules1.simps(2)) + case Nil then show ?thesis apply(simp_all) + by (metis list2FWpolicy.simps(1) rSR1_eq removeShadowRules1.simps(2)) next - case (Cons x list) show ?thesis - apply (rule rev_induct) - apply (metis rSR1_eq removeShadowRules1.simps(2)) - apply (case_tac "xs = []", simp_all) - unfolding removeShadowRules1_alternative_def - apply (case_tac x, simp_all) - by (metis (no_types, hide_lams) CConcEnd2 CConcStart C_DenyAll RS1n_nMT aux114 - domIff removeShadowRules1_alternative_def - removeShadowRules1_alternative_rev.simps(2) rev.simps(2)) + case (Cons x list) show ?thesis + apply (rule rev_induct) + apply (metis rSR1_eq removeShadowRules1.simps(2)) + apply (case_tac "xs = []", simp_all) + unfolding removeShadowRules1_alternative_def + apply (case_tac x, simp_all) + by (metis (no_types, hide_lams) CConcEnd2 CConcStart C_DenyAll RS1n_nMT aux114 + domIff removeShadowRules1_alternative_def + removeShadowRules1_alternative_rev.simps(2) rev.simps(2)) qed - - lemma C_eq_RS1[simp]: "p \ [] \ C(list2FWpolicy (removeShadowRules1 p)) = C(list2FWpolicy p)" -by (metis rSR1_eq C_eq_RS1n) - - + by (metis rSR1_eq C_eq_RS1n) + lemma EX_MR_aux[rule_format]: "applied_rule_rev C x (DenyAll # p) \ Some DenyAll \ (\y. applied_rule_rev C x p = Some y)" -apply (simp add: applied_rule_rev_def) -apply (rule_tac xs = p in rev_induct, simp_all) -done - + apply (simp add: applied_rule_rev_def) + apply (rule_tac xs = p in rev_induct, simp_all) + done + lemma EX_MR : -"applied_rule_rev C x p \ \DenyAll\ \ p = DenyAll # ps \ + "applied_rule_rev C x p \ \DenyAll\ \ p = DenyAll # ps \ applied_rule_rev C x p = applied_rule_rev C x ps" -apply auto - apply (subgoal_tac "applied_rule_rev C x (DenyAll#ps) \ None", auto) -apply (metis mrConcEnd) -by (metis DAimpliesMR_E list.sel(1) hd_in_set list.simps(3) not_Some_eq) - - + apply auto + apply (subgoal_tac "applied_rule_rev C x (DenyAll#ps) \ 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 \ applied_rule_rev C x p = \DenyAllFromTo a ab\ \ set p = set s \ applied_rule_rev C x s \ \DenyAll\" -apply (subst wp1n_tl, simp_all) - apply (subgoal_tac "x \ dom (C (DenyAllFromTo a ab))") - apply (subgoal_tac "DenyAllFromTo a ab \ set (tl s)") - apply (metis wp1n_tl foo98 wellformed_policy1_strong.simps(2)) - using mrSet r_not_DA_in_tl apply blast -by (simp add: mr_in_dom) - + apply (subst wp1n_tl, simp_all) + apply (subgoal_tac "x \ dom (C (DenyAllFromTo a ab))") + apply (subgoal_tac "DenyAllFromTo a ab \ set (tl s)") + apply (metis wp1n_tl foo98 wellformed_policy1_strong.simps(2)) + using mrSet r_not_DA_in_tl apply blast + by (simp add: mr_in_dom) lemma domsMT_notND_DD: - "dom (C (DenyAllFromTo a b)) \ dom (C (DenyAllFromTo c d)) \ {} \ \ netsDistinct a c" -using deny_dom twoNetsDistinct_def by blast - + "dom (C (DenyAllFromTo a b)) \ dom (C (DenyAllFromTo c d)) \ {} \ \ netsDistinct a c" + using deny_dom twoNetsDistinct_def by blast lemma domsMT_notND_DD2: "dom (C (DenyAllFromTo a b)) \ dom (C (DenyAllFromTo c d)) \ {} \ \ netsDistinct b d" -using deny_dom twoNetsDistinct_def by blast - + using deny_dom twoNetsDistinct_def by blast lemma domsMT_notND_DD3: "x \ dom (C (DenyAllFromTo a b)) \ x \ dom (C (DenyAllFromTo c d)) \ \ netsDistinct a c" -by(auto intro!:domsMT_notND_DD) - + by(auto intro!:domsMT_notND_DD) lemma domsMT_notND_DD4: "x \ dom (C (DenyAllFromTo a b)) \ x \ dom (C (DenyAllFromTo c d)) \ \ netsDistinct b d" -by(auto intro!:domsMT_notND_DD2) - + by(auto intro!:domsMT_notND_DD2) lemma NetsEq_if_sameP_DD: "allNetsDistinct p \ u \ set p \ v \ set p \ u = DenyAllFromTo a b \ v = DenyAllFromTo c d \ x \ dom (C u) \ x \ dom (C v) \ a = c \ b = d" -apply (simp add: allNetsDistinct_def) -by (metis ND0aux1 ND0aux2 domsMT_notND_DD3 domsMT_notND_DD4 ) - + 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)" - and SC: "singleCombinators p" - and inp: "r \ set p" - and inDom: "x \ dom (C r)" - shows "(r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll)" + assumes aND: "allNetsDistinct p" + and mr_is_allow: "applied_rule_rev C x p = Some (AllowPortFromTo a b po)" + and SC: "singleCombinators p" + and inp: "r \ set p" + and inDom: "x \ dom (C r)" + shows "(r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll)" proof (cases r) - case DenyAll show ?thesis by (metis DenyAll) + case DenyAll show ?thesis by (metis DenyAll) next - case (DenyAllFromTo x y) show ?thesis - by (metis AD_aux DenyAllFromTo SC aND domInterMT inDom inp mrSet mr_in_dom mr_is_allow) + case (DenyAllFromTo x y) show ?thesis + by (metis AD_aux DenyAllFromTo SC aND domInterMT inDom inp mrSet mr_in_dom mr_is_allow) next - case (AllowPortFromTo x y b) show ?thesis - by (metis (no_types, lifting) AllowPortFromTo DistinctNetsAllowAllow aND domInterMT - inDom inp mrSet mr_in_dom mr_is_allow) + case (AllowPortFromTo x y b) show ?thesis + by (metis (no_types, lifting) AllowPortFromTo DistinctNetsAllowAllow aND domInterMT + inDom inp mrSet mr_in_dom mr_is_allow) next - case (Conc x y) thus ?thesis using assms by (metis aux0_0) + case (Conc x y) thus ?thesis using assms by (metis aux0_0) qed - + lemma none_MT_rulessubset[rule_format]: "none_MT_rules C a \ set b \ set a \ none_MT_rules C b" -by (induct b,simp_all) (metis notMTnMT) - + by (induct b,simp_all) (metis notMTnMT) lemma nMTSort: "none_MT_rules C p \ none_MT_rules C (sort p l)" -by (metis set_sort nMTeqSet) - - + by (metis set_sort nMTeqSet) lemma nMTSortQ: "none_MT_rules C p \ none_MT_rules C (qsort p l)" -by (metis set_sortQ nMTeqSet) - + by (metis set_sortQ nMTeqSet) + lemma wp3char[rule_format]: -"none_MT_rules C xs \ C (AllowPortFromTo a b po)=\ \ wellformed_policy3(xs@[DenyAllFromTo a b]) \ + "none_MT_rules C xs \ C (AllowPortFromTo a b po)=\ \ wellformed_policy3(xs@[DenyAllFromTo a b]) \ AllowPortFromTo a b po \ set xs" -apply (induct xs,simp_all) -by (metis domNMT wp3Conc) - - + apply (induct xs,simp_all) + by (metis domNMT wp3Conc) lemma wp3charn[rule_format]: -assumes domAllow: "dom (C (AllowPortFromTo a b po)) \ {}" -and wp3: "wellformed_policy3 (xs @ [DenyAllFromTo a b])" -shows "AllowPortFromTo a b po \ set xs" -apply (insert assms) + assumes domAllow: "dom (C (AllowPortFromTo a b po)) \ {}" + and wp3: "wellformed_policy3 (xs @ [DenyAllFromTo a b])" + shows "AllowPortFromTo a b po \ set xs" + apply (insert assms) proof (induct xs) - case Nil show ?case by simp + case Nil show ?case by simp next - case (Cons x xs) show ?case using Cons - by (simp,auto intro: wp3Conc) (auto simp: DenyAllowDisj domAllow) + case (Cons x xs) show ?case using Cons + by (simp,auto intro: wp3Conc) (auto simp: DenyAllowDisj domAllow) qed - lemma rule_charn2: - assumes aND: "allNetsDistinct p" - and wp1: "wellformed_policy1 p" - and SC: "singleCombinators p" - and wp3: "wellformed_policy3 p" - and allow_in_list: "AllowPortFromTo c d po \ set p" - and x_in_dom_allow: "x \ dom (C (AllowPortFromTo c d po))" - shows "applied_rule_rev C x p = Some (AllowPortFromTo c d po)" - proof (insert assms, induct p rule: rev_induct) - case Nil show ?case using Nil by simp - next - case (snoc y ys) - show ?case using snoc - apply (case_tac "y = (AllowPortFromTo c d po)", simp_all ) + assumes aND: "allNetsDistinct p" + and wp1: "wellformed_policy1 p" + and SC: "singleCombinators p" + and wp3: "wellformed_policy3 p" + and allow_in_list: "AllowPortFromTo c d po \ set p" + and x_in_dom_allow: "x \ dom (C (AllowPortFromTo c d po))" + shows "applied_rule_rev C x p = Some (AllowPortFromTo c d po)" +proof (insert assms, induct p rule: rev_induct) + case Nil show ?case using Nil by simp +next + case (snoc y ys) + show ?case using snoc + apply (case_tac "y = (AllowPortFromTo c d po)", simp_all ) apply (simp add: applied_rule_rev_def) - apply (subgoal_tac "ys \ []") + apply (subgoal_tac "ys \ []") apply (subgoal_tac "applied_rule_rev C x ys = Some (AllowPortFromTo c d po)") - defer 1 - apply (metis ANDConcEnd SCConcEnd WP1ConcEnd foo25) + defer 1 + apply (metis ANDConcEnd SCConcEnd WP1ConcEnd foo25) apply (metis inSet_not_MT) - proof (cases y) - case DenyAll thus ?thesis using DenyAll snoc - apply simp - by (metis DAnotTL DenyAll inSet_not_MT policy2list.simps(2)) - next - case (DenyAllFromTo a b) thus ?thesis using snoc apply simp - apply (simp_all add: applied_rule_rev_def) - apply (rule conjI) - apply (metis domInterMT wp3EndMT) - apply (rule impI) - by (metis ANDConcEnd DenyAllFromTo SCConcEnd WP1ConcEnd foo25) - next - case (AllowPortFromTo a1 a2 b) thus ?thesis - using AllowPortFromTo snoc apply simp - apply (simp_all add: applied_rule_rev_def) - apply (rule conjI) - apply (metis domInterMT wp3EndMT) - by (metis ANDConcEnd AllowPortFromTo SCConcEnd WP1ConcEnd foo25 x_in_dom_allow) - next - case (Conc a b) thus ?thesis using Conc snoc apply simp - by (metis Conc aux0_0 in_set_conv_decomp) - qed + proof (cases y) + case DenyAll thus ?thesis using DenyAll snoc + apply simp + by (metis DAnotTL DenyAll inSet_not_MT policy2list.simps(2)) + next + case (DenyAllFromTo a b) thus ?thesis using snoc apply simp + apply (simp_all add: applied_rule_rev_def) + apply (rule conjI) + apply (metis domInterMT wp3EndMT) + apply (rule impI) + by (metis ANDConcEnd DenyAllFromTo SCConcEnd WP1ConcEnd foo25) + next + case (AllowPortFromTo a1 a2 b) thus ?thesis + using AllowPortFromTo snoc apply simp + apply (simp_all add: applied_rule_rev_def) + apply (rule conjI) + apply (metis domInterMT wp3EndMT) + by (metis ANDConcEnd AllowPortFromTo SCConcEnd WP1ConcEnd foo25 x_in_dom_allow) + next + case (Conc a b) thus ?thesis using Conc snoc apply simp + by (metis Conc aux0_0 in_set_conv_decomp) + qed qed - + lemma rule_charn3: - " wellformed_policy1 p \ allNetsDistinct p \ singleCombinators p \ + " wellformed_policy1 p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy3 p \ applied_rule_rev C x p = \DenyAllFromTo c d\ \ AllowPortFromTo a b po \ set p \ x \ dom (C (AllowPortFromTo a b po))" -by (clarify, auto simp: rule_charn2 dom_def) - + by (clarify, auto simp: rule_charn2 dom_def) + lemma rule_charn4: -assumes wp1: "wellformed_policy1 p" -and aND: "allNetsDistinct p" -and SC: "singleCombinators p" -and wp3: "wellformed_policy3 p" -and DA: "DenyAll \ set p" -and mr: "applied_rule_rev C x p = Some (DenyAllFromTo a b)" -and rinp: "r \ set p" -and xindom: "x \ dom (C r)" -shows "r = DenyAllFromTo a b" + assumes wp1: "wellformed_policy1 p" + and aND: "allNetsDistinct p" + and SC: "singleCombinators p" + and wp3: "wellformed_policy3 p" + and DA: "DenyAll \ set p" + and mr: "applied_rule_rev C x p = Some (DenyAllFromTo a b)" + and rinp: "r \ set p" + and xindom: "x \ dom (C r)" + shows "r = DenyAllFromTo a b" proof (cases r) case DenyAll thus ?thesis using DenyAll assms by simp next case (DenyAllFromTo c d) thus ?thesis using assms apply simp apply (erule_tac x = x and p = p and v = "(DenyAllFromTo a b)" and - u = "(DenyAllFromTo c d)" in NetsEq_if_sameP_DD) - apply simp_all - apply (erule mrSet) + u = "(DenyAllFromTo c d)" in NetsEq_if_sameP_DD) + apply simp_all + apply (erule mrSet) by (erule mr_in_dom) next case (AllowPortFromTo c d e) thus ?thesis using assms apply simp apply (subgoal_tac "x \ dom (C (AllowPortFromTo c d e))") - apply simp + apply simp apply (rule_tac p = p in rule_charn3) by (auto intro: SCnotConc) next case (Conc a b) thus ?thesis using assms apply simp - by (metis Conc aux0_0) + by (metis Conc aux0_0) qed - - lemma foo31a: "\r. r \ set p \ x \ dom (C r) \ r=AllowPortFromTo a b po \ r=DenyAllFromTo a b \ r=DenyAll \ set p = set s \ r \ set s \ x \ dom (C r) \ r = AllowPortFromTo a b po \ r = DenyAllFromTo a b \ r = DenyAll" -by auto - + by auto lemma aux4[rule_format]: "applied_rule_rev C x (a#p) = Some a \ a \ set (p) \ 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) - + 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" - shows "applied_rule_rev C x (tl p) = None" - apply (rule aux4 [where a = DenyAll]) - apply (metis wp1n_tl mr_DA wp1n) - by (metis WP1n_DA_notinSet wp1n) - + assumes mr_DA: "applied_rule_rev C x p = Some DenyAll" + and wp1n: "wellformed_policy1_strong p" + shows "applied_rule_rev C x (tl p) = None" + apply (rule aux4 [where a = DenyAll]) + apply (metis wp1n_tl mr_DA wp1n) + by (metis WP1n_DA_notinSet wp1n) + lemma rule_charnDAFT: -"wellformed_policy1_strong p \ allNetsDistinct p \ singleCombinators p \ + "wellformed_policy1_strong p \ allNetsDistinct p \ singleCombinators p \ wellformed_policy3 p \ applied_rule_rev C x p = \DenyAllFromTo a b\ \ r \ set (tl p) \ x \ dom (C r) \ r = DenyAllFromTo a b" -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 - + 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: "\wellformed_policy1_strong p; applied_rule_rev C x p = Some DenyAll; r \ set (tl p)\ \ x \ dom (C r)" -apply (rule_tac a = "[]" and b = "DenyAll" and c = "tl p" in foo3a, simp_all) - apply (metis wp1n_tl) -by (metis WP1n_DA_notinSet) - + 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" - and wp1_p: "wellformed_policy1_strong p" - and wp1_s: "wellformed_policy1_strong s" - and wp3_p: "wellformed_policy3 p" - and wp3_s: "wellformed_policy3 s" - and aND: "allNetsDistinct p" - shows "applied_rule_rev C x p = applied_rule_rev C x s" + assumes sets_eq: "set p = set s" + and SC: "singleCombinators p" + and wp1_p: "wellformed_policy1_strong p" + and wp1_s: "wellformed_policy1_strong s" + and wp3_p: "wellformed_policy3 p" + and wp3_s: "wellformed_policy3 s" + and aND: "allNetsDistinct p" + shows "applied_rule_rev C x p = applied_rule_rev C x s" proof (cases "applied_rule_rev C x p") - case None - have DA: "DenyAll \ set p" using wp1_p by (auto simp: wp1_aux1aa) - have notDA: "DenyAll \ set p" using None by (auto simp: DAimplieMR) - thus ?thesis using DA by (contradiction) + case None + have DA: "DenyAll \ set p" using wp1_p by (auto simp: wp1_aux1aa) + have notDA: "DenyAll \ set p" using None by (auto simp: DAimplieMR) + thus ?thesis using DA by (contradiction) next - case (Some y) thus ?thesis - proof (cases y) - have tl_p: "p = DenyAll#(tl p)" by (metis wp1_p wp1n_tl) - have tl_s: "s = DenyAll#(tl s)" by (metis wp1_s wp1n_tl) - have tl_eq: "set (tl p) = set (tl s)" - by (metis list.sel(3) WP1n_DA_notinSet sets_eq foo2 - wellformed_policy1_charn wp1_aux1aa wp1_eq wp1_p wp1_s) - { case DenyAll - have mr_p_is_DenyAll: "applied_rule_rev C x p = Some DenyAll" - by (simp add: DenyAll Some) - hence x_notin_tl_p: "\ r. r \ set (tl p) \ x \ dom (C r)" using wp1_p - by (auto simp: mrDenyAll_is_unique) - hence x_notin_tl_s: "\ r. r \ set (tl s) \ x \ dom (C r)" using tl_eq - by auto - hence mr_s_is_DenyAll: "applied_rule_rev C x s = Some DenyAll" using tl_s - by (auto simp: mr_first) - thus ?thesis using mr_p_is_DenyAll by simp - } - {case (DenyAllFromTo a b) - have mr_p_is_DAFT: "applied_rule_rev C x p = Some (DenyAllFromTo a b)" - by (simp add: DenyAllFromTo Some) - have DA_notin_tl: "DenyAll \ set (tl p)" - by (metis WP1n_DA_notinSet wp1_p) - have mr_tl_p: "applied_rule_rev C x p = applied_rule_rev C x (tl p)" - by (metis Combinators.simps(4) DenyAllFromTo Some mrConcEnd tl_p) - have dom_tl_p: "\ r. r \ set (tl p) \ x \ dom (C r) \ r = (DenyAllFromTo a b)" - using wp1_p aND SC wp3_p mr_p_is_DAFT - by (auto simp: rule_charnDAFT) - hence dom_tl_s: "\ r. r \ set (tl s) \ x \ dom (C r) \ r = (DenyAllFromTo a b)" - using tl_eq by auto - have DAFT_in_tl_s: "DenyAllFromTo a b \ set (tl s)" using mr_tl_p - by (metis DenyAllFromTo mrSet mr_p_is_DAFT tl_eq) - have x_in_dom_DAFT: "x \ dom (C (DenyAllFromTo a b))" - by (metis mr_p_is_DAFT DenyAllFromTo mr_in_dom) - hence mr_tl_s_is_DAFT: "applied_rule_rev C x (tl s) = Some (DenyAllFromTo a b)" - using DAFT_in_tl_s dom_tl_s by (metis mr_charn) - hence mr_s_is_DAFT: "applied_rule_rev C x s = Some (DenyAllFromTo a b)" - using tl_s - by (metis DA_notin_tl DenyAllFromTo EX_MR mrDA_tl - not_Some_eq tl_eq wellformed_policy1_strong.simps(2)) - thus ?thesis using mr_p_is_DAFT by simp - } - {case (AllowPortFromTo a b c) - have wp1s: "wellformed_policy1 s" by (metis wp1_eq wp1_s) - have mr_p_is_A: "applied_rule_rev C x p = Some (AllowPortFromTo a b c)" - by (simp add: AllowPortFromTo Some) - hence A_in_s: "AllowPortFromTo a b c \ set s" using sets_eq - by (auto intro: mrSet) - have x_in_dom_A: "x \ dom (C (AllowPortFromTo a b c))" - by (metis mr_p_is_A AllowPortFromTo mr_in_dom) - have SCs: "singleCombinators s" using SC sets_eq - by (auto intro: SCSubset) - hence ANDs: "allNetsDistinct s" using aND sets_eq SC - by (auto intro: aNDSetsEq) - hence mr_s_is_A: "applied_rule_rev C x s = Some (AllowPortFromTo a b c)" - using A_in_s wp1s mr_p_is_A aND SCs wp3_s x_in_dom_A - by (simp add: rule_charn2) - thus ?thesis using mr_p_is_A by simp - } - case (Conc a b) thus ?thesis by (metis Some mr_not_Conc SC) - qed + case (Some y) thus ?thesis + proof (cases y) + have tl_p: "p = DenyAll#(tl p)" by (metis wp1_p wp1n_tl) + have tl_s: "s = DenyAll#(tl s)" by (metis wp1_s wp1n_tl) + have tl_eq: "set (tl p) = set (tl s)" + by (metis list.sel(3) WP1n_DA_notinSet sets_eq foo2 + wellformed_policy1_charn wp1_aux1aa wp1_eq wp1_p wp1_s) + { case DenyAll + have mr_p_is_DenyAll: "applied_rule_rev C x p = Some DenyAll" + by (simp add: DenyAll Some) + hence x_notin_tl_p: "\ r. r \ set (tl p) \ x \ dom (C r)" using wp1_p + by (auto simp: mrDenyAll_is_unique) + hence x_notin_tl_s: "\ r. r \ set (tl s) \ x \ dom (C r)" using tl_eq + by auto + hence mr_s_is_DenyAll: "applied_rule_rev C x s = Some DenyAll" using tl_s + by (auto simp: mr_first) + thus ?thesis using mr_p_is_DenyAll by simp + } + {case (DenyAllFromTo a b) + have mr_p_is_DAFT: "applied_rule_rev C x p = Some (DenyAllFromTo a b)" + by (simp add: DenyAllFromTo Some) + have DA_notin_tl: "DenyAll \ set (tl p)" + by (metis WP1n_DA_notinSet wp1_p) + have mr_tl_p: "applied_rule_rev C x p = applied_rule_rev C x (tl p)" + by (metis Combinators.simps(4) DenyAllFromTo Some mrConcEnd tl_p) + have dom_tl_p: "\ r. r \ set (tl p) \ x \ dom (C r) \ r = (DenyAllFromTo a b)" + using wp1_p aND SC wp3_p mr_p_is_DAFT + by (auto simp: rule_charnDAFT) + hence dom_tl_s: "\ r. r \ set (tl s) \ x \ dom (C r) \ r = (DenyAllFromTo a b)" + using tl_eq by auto + have DAFT_in_tl_s: "DenyAllFromTo a b \ set (tl s)" using mr_tl_p + by (metis DenyAllFromTo mrSet mr_p_is_DAFT tl_eq) + have x_in_dom_DAFT: "x \ dom (C (DenyAllFromTo a b))" + by (metis mr_p_is_DAFT DenyAllFromTo mr_in_dom) + hence mr_tl_s_is_DAFT: "applied_rule_rev C x (tl s) = Some (DenyAllFromTo a b)" + using DAFT_in_tl_s dom_tl_s by (metis mr_charn) + hence mr_s_is_DAFT: "applied_rule_rev C x s = Some (DenyAllFromTo a b)" + using tl_s + by (metis DA_notin_tl DenyAllFromTo EX_MR mrDA_tl + not_Some_eq tl_eq wellformed_policy1_strong.simps(2)) + thus ?thesis using mr_p_is_DAFT by simp + } + {case (AllowPortFromTo a b c) + have wp1s: "wellformed_policy1 s" by (metis wp1_eq wp1_s) + have mr_p_is_A: "applied_rule_rev C x p = Some (AllowPortFromTo a b c)" + by (simp add: AllowPortFromTo Some) + hence A_in_s: "AllowPortFromTo a b c \ set s" using sets_eq + by (auto intro: mrSet) + have x_in_dom_A: "x \ dom (C (AllowPortFromTo a b c))" + by (metis mr_p_is_A AllowPortFromTo mr_in_dom) + have SCs: "singleCombinators s" using SC sets_eq + by (auto intro: SCSubset) + hence ANDs: "allNetsDistinct s" using aND sets_eq SC + by (auto intro: aNDSetsEq) + hence mr_s_is_A: "applied_rule_rev C x s = Some (AllowPortFromTo a b c)" + using A_in_s wp1s mr_p_is_A aND SCs wp3_s x_in_dom_A + by (simp add: rule_charn2) + thus ?thesis using mr_p_is_A by simp + } + case (Conc a b) thus ?thesis by (metis Some mr_not_Conc SC) + qed qed lemma C_eq_Sets: @@ -801,243 +757,221 @@ 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 \ all_in_list p l \ singleCombinators p \ wellformed_policy1_strong p \ wellformed_policy3 p \ allNetsDistinct p \ C (list2FWpolicy (FWNormalisationCore.sort p l)) = C (list2FWpolicy p)" -apply (rule ext) -by (auto intro: C_eq_Sets simp: nMTSort wellformed1_alternative_sorted - wellformed_policy3_charn wp1_eq) - - + 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 \ all_in_list p l \ singleCombinators p \ wellformed_policy1_strong p \ wellformed_policy3 p \ allNetsDistinct p \ C (list2FWpolicy (qsort p l)) = C (list2FWpolicy p)" -apply (rule ext) -apply (auto intro!: C_eq_Sets simp: nMTSortQ wellformed1_alternative_sorted distinct_sortQ - wellformed_policy3_charn wp1_eq) -by (metis set_qsort wellformed1_sortedQ wellformed_eq wp1_aux1aa) - - + apply (rule ext) + apply (auto intro!: C_eq_Sets simp: nMTSortQ wellformed1_alternative_sorted distinct_sortQ + 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 + case Nil thus ?case by simp next - case (Cons y ys) thus ?case - proof (cases "ys = []") - case True thus ?thesis by (cases y, simp_all) - next - case False thus ?thesis - proof (cases y) - case DenyAll thus ?thesis by (simp, metis Cons DenyAll mreq_end2) - next - case (DenyAllFromTo a b) thus ?thesis - by (simp, metis Cons DenyAllFromTo mreq_end2) - next - case (AllowPortFromTo a b p) thus ?thesis - proof (cases "DenyAllFromTo a b \ set ys") - case True thus ?thesis using AllowPortFromTo Cons - apply (cases "applied_rule_rev C x ys = None", simp_all) - apply (subgoal_tac "x \ dom (C (AllowPortFromTo a b p))") - apply (subst mrconcNone, simp_all) - apply (simp add: applied_rule_rev_def ) - apply (rule contra_subsetD [OF allow_deny_dom]) - apply (erule mrNoneMT,simp) - apply (metis AllowPortFromTo mrconc) - done - next - case False thus ?thesis using False Cons AllowPortFromTo - by (simp, metis AllowPortFromTo Cons mreq_end2) qed - next - case (Conc a b) thus ?thesis - by (metis Cons mreq_end2 removeShadowRules2.simps(4)) - qed - qed + case (Cons y ys) thus ?case + proof (cases "ys = []") + case True thus ?thesis by (cases y, simp_all) + next + case False thus ?thesis + proof (cases y) + case DenyAll thus ?thesis by (simp, metis Cons DenyAll mreq_end2) + next + case (DenyAllFromTo a b) thus ?thesis + by (simp, metis Cons DenyAllFromTo mreq_end2) + next + case (AllowPortFromTo a b p) thus ?thesis + proof (cases "DenyAllFromTo a b \ set ys") + case True thus ?thesis using AllowPortFromTo Cons + apply (cases "applied_rule_rev C x ys = None", simp_all) + apply (subgoal_tac "x \ dom (C (AllowPortFromTo a b p))") + apply (subst mrconcNone, simp_all) + apply (simp add: applied_rule_rev_def ) + apply (rule contra_subsetD [OF allow_deny_dom]) + apply (erule mrNoneMT,simp) + apply (metis AllowPortFromTo mrconc) + done + next + case False thus ?thesis using False Cons AllowPortFromTo + by (simp, metis AllowPortFromTo Cons mreq_end2) qed + next + case (Conc a b) thus ?thesis + by (metis Cons mreq_end2 removeShadowRules2.simps(4)) + qed + qed qed - - lemma C_eq_None[rule_format]: "p \ [] --> applied_rule_rev C x p = None \ C (list2FWpolicy p) x = None" -apply (simp add: applied_rule_rev_def) -apply (rule rev_induct, simp_all) -apply (intro impI, simp) - apply (case_tac "xs \ []") -apply (simp_all add: dom_def) -done - + apply (simp add: applied_rule_rev_def) + apply (rule rev_induct, simp_all) + apply (intro impI, simp) + apply (case_tac "xs \ []") + apply (simp_all add: dom_def) + done + lemma C_eq_None2: "a \ [] \ b \ [] \ applied_rule_rev C x a = \ \ applied_rule_rev C x b = \ \ C (list2FWpolicy a) x = C (list2FWpolicy b) x" -by (auto simp: C_eq_None) - + by (auto simp: C_eq_None) + lemma C_eq_RS2: "wellformed_policy1_strong p \ C (list2FWpolicy (removeShadowRules2 p))= C (list2FWpolicy p)" -apply (rule ext) -by (metis C_eq_RS2_mr C_eq_if_mr_eq wellformed_policy1_strong.simps(1) wp1n_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 \ none_MT_rules C (removeShadowRules2 p)" -by (auto simp: RS2Set none_MT_rulessubset) - + by (auto simp: RS2Set none_MT_rulessubset) + lemma CconcNone: "dom (C a) = {} \ p \ [] \ C (list2FWpolicy (a # p)) x = C (list2FWpolicy p) x" -apply (case_tac "p = []", simp_all) - apply (case_tac "x\ dom (C (list2FWpolicy(p)))") -apply (metis Cdom2 list2FWpolicyconc) -apply (metis C.simps(4) map_add_dom_app_simps(2) inSet_not_MT list2FWpolicyconc set_empty2) -done - + apply (case_tac "p = []", simp_all) + apply (case_tac "x\ dom (C (list2FWpolicy(p)))") + 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 \ none_MT_rules C (remdups p)" -by (induct p, simp_all) - + by (induct p, simp_all) + lemma DARS3[rule_format]: "DenyAll \ set p\DenyAll \ set (rm_MT_rules C p)" -by (induct p, simp_all) - + by (induct p, simp_all) + lemma DAnMT: "dom (C DenyAll) \ {}" -by (simp add: dom_def C.simps PolicyCombinators.PolicyCombinators) - + by (simp add: dom_def C.simps PolicyCombinators.PolicyCombinators) + lemma DAnMT2: "C DenyAll \ empty" -by (metis DAAux dom_eq_empty_conv empty_iff) - - - + by (metis DAAux dom_eq_empty_conv empty_iff) lemma wp1n_RS3[rule_format,simp]: "wellformed_policy1_strong p \ wellformed_policy1_strong (rm_MT_rules C p)" -by (induct p, simp_all add: DARS3 DAnMT) - + by (induct p, simp_all add: DARS3 DAnMT) + lemma AILRS3[rule_format,simp]: "all_in_list p l \ all_in_list (rm_MT_rules C p) l" -by (induct p, simp_all) - + by (induct p, simp_all) + lemma SCRS3[rule_format,simp]: "singleCombinators p \ singleCombinators(rm_MT_rules C p)" -by (induct p, simp_all, case_tac "a", simp_all) - + by (induct p, simp_all, case_tac "a", simp_all) + lemma RS3subset: "set (rm_MT_rules C p) \ set p " -by (induct p, auto) - + by (induct p, auto) lemma ANDRS3[simp]: "singleCombinators p \ allNetsDistinct p \ allNetsDistinct (rm_MT_rules C p)" -using RS3subset SCRS3 aNDSubset by blast + using RS3subset SCRS3 aNDSubset by blast lemma nlpaux: "x \ dom (C b) \ C (a \ b) x = C a x" -by (metis C.simps(4) map_add_dom_app_simps(3)) + by (metis C.simps(4) map_add_dom_app_simps(3)) lemma notindom[rule_format]: "a \ set p \ x \ dom (C (list2FWpolicy p)) \ x \ dom (C a)" -apply (induct p, simp_all) -by (metis CConcStartA Cdom2 domIff empty_iff empty_set l2p_aux) + apply (induct p, simp_all) + by (metis CConcStartA Cdom2 domIff empty_iff empty_set l2p_aux) lemma C_eq_rd[rule_format]: - "p \ [] \ C (list2FWpolicy (remdups p)) = C (list2FWpolicy p)" + "p \ [] \ C (list2FWpolicy (remdups p)) = C (list2FWpolicy p)" proof (rule ext,induct p) case Nil thus ?case by simp next case (Cons y ys) thus ?case - proof (cases "ys = []") - case True thus ?thesis by simp - next - case False thus ?thesis using Cons - apply (simp) apply (rule conjI, rule impI) - apply (cases "x \ dom (C (list2FWpolicy ys))") - apply (metis Cdom2 False list2FWpolicyconc) - apply (metis False domIff list2FWpolicyconc nlpaux notindom) - apply (rule impI) - apply (cases "x \ dom (C (list2FWpolicy ys))") - apply (subgoal_tac "x \ dom (C (list2FWpolicy (remdups ys)))") - apply (metis Cdom2 False list2FWpolicyconc remdups_eq_nil_iff) - apply (metis domIff) - apply (subgoal_tac "x \ dom (C (list2FWpolicy (remdups ys)))") - apply (metis False list2FWpolicyconc nlpaux remdups_eq_nil_iff) - apply (metis domIff) - done - qed + proof (cases "ys = []") + case True thus ?thesis by simp + next + case False thus ?thesis using Cons + apply (simp) apply (rule conjI, rule impI) + apply (cases "x \ dom (C (list2FWpolicy ys))") + apply (metis Cdom2 False list2FWpolicyconc) + apply (metis False domIff list2FWpolicyconc nlpaux notindom) + apply (rule impI) + apply (cases "x \ dom (C (list2FWpolicy ys))") + apply (subgoal_tac "x \ dom (C (list2FWpolicy (remdups ys)))") + apply (metis Cdom2 False list2FWpolicyconc remdups_eq_nil_iff) + apply (metis domIff) + apply (subgoal_tac "x \ dom (C (list2FWpolicy (remdups ys)))") + apply (metis False list2FWpolicyconc nlpaux remdups_eq_nil_iff) + apply (metis domIff) + done + qed qed - - lemma nMT_domMT: - "\ not_MT C p \ p \ [] \ r \ dom (C (list2FWpolicy p))" + "\ not_MT C p \ p \ [] \ r \ dom (C (list2FWpolicy p))" proof (induct p) - case Nil thus ?case by simp + case Nil thus ?case by simp next - case (Cons x xs) thus ?case - apply (simp split: if_splits) - apply (cases "xs = []",simp_all ) - by (metis CconcNone domIff) + case (Cons x xs) thus ?case + apply (simp split: if_splits) + apply (cases "xs = []",simp_all ) + by (metis CconcNone domIff) qed - + lemma C_eq_RS3_aux[rule_format]: - "not_MT C p \ C (list2FWpolicy p) x = C (list2FWpolicy (rm_MT_rules C p)) x" + "not_MT C p \ C (list2FWpolicy p) x = C (list2FWpolicy (rm_MT_rules C p)) x" proof (induct p) - case Nil thus ?case by simp + case Nil thus ?case by simp next - case (Cons y ys) - thus ?case - proof (cases "not_MT C ys") - case True thus ?thesis using Cons - apply (simp) apply(rule conjI, rule impI, simp) - apply (metis CconcNone True not_MTimpnotMT) - apply (rule impI, simp) - apply (cases "x \ dom (C (list2FWpolicy ys))") - apply (subgoal_tac "x \ dom (C (list2FWpolicy (rm_MT_rules C ys)))") - apply (metis Cdom2 NMPrm l2p_aux not_MTimpnotMT) - apply (simp add: domIff) - apply (subgoal_tac "x \ dom (C (list2FWpolicy (rm_MT_rules C ys)))") - apply (metis l2p_aux l2p_aux2 nlpaux) - apply (metis domIff) - done - next - case False thus ?thesis using Cons False - proof (cases "ys = []") - case True thus ?thesis using Cons by (simp) (rule impI, simp) - next - case False thus ?thesis - using Cons False `\ not_MT C ys` apply (simp) - by (metis SR3nMT l2p_aux list2FWpolicy.simps(2) nMT_domMT nlpaux) - qed - qed + case (Cons y ys) + thus ?case + proof (cases "not_MT C ys") + case True thus ?thesis using Cons + apply (simp) apply(rule conjI, rule impI, simp) + apply (metis CconcNone True not_MTimpnotMT) + apply (rule impI, simp) + apply (cases "x \ dom (C (list2FWpolicy ys))") + apply (subgoal_tac "x \ dom (C (list2FWpolicy (rm_MT_rules C ys)))") + apply (metis Cdom2 NMPrm l2p_aux not_MTimpnotMT) + apply (simp add: domIff) + apply (subgoal_tac "x \ dom (C (list2FWpolicy (rm_MT_rules C ys)))") + apply (metis l2p_aux l2p_aux2 nlpaux) + apply (metis domIff) + done + next + case False thus ?thesis using Cons False + proof (cases "ys = []") + case True thus ?thesis using Cons by (simp) (rule impI, simp) + next + case False thus ?thesis + using Cons False `\ not_MT C ys` apply (simp) + by (metis SR3nMT l2p_aux list2FWpolicy.simps(2) nMT_domMT nlpaux) + qed + qed qed - - lemma C_eq_id: "wellformed_policy1_strong p \ C(list2FWpolicy (insertDeny p)) = C (list2FWpolicy p)" -by (rule ext) (auto intro: C_eq_if_mr_eq elim: mr_iD) - + by (rule ext) (auto intro: C_eq_if_mr_eq elim: mr_iD) + lemma C_eq_RS3: "not_MT C p \ C(list2FWpolicy (rm_MT_rules C p)) = C (list2FWpolicy p)" -by (rule ext) (erule C_eq_RS3_aux[symmetric]) - + by (rule ext) (erule C_eq_RS3_aux[symmetric]) lemma NMPrd[rule_format]: "not_MT C p \ not_MT C (remdups p)" -by (induct p) (auto simp: NMPcharn) - + by (induct p) (auto simp: NMPcharn) lemma NMPDA[rule_format]: "DenyAll \ set p \ not_MT C p" -by (induct p, simp_all add: DAnMT) - + by (induct p, simp_all add: DAnMT) lemma NMPiD[rule_format]: "not_MT C (insertDeny p)" -by (simp add: DAiniD NMPDA) - + 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) - + 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 @@ -1051,13 +985,12 @@ lemma C_eq_All_untilSorted: (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l)) = C p" -apply (subst C_eq_sorted,simp_all add: C_eq_Lemmas) -apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) -apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) -apply (subst C_eq_RS3,simp_all add: C_eq_Lemmas) -apply (subst C_eq_id,simp_all add: C_eq_Lemmas) -done - + apply (subst C_eq_sorted,simp_all add: C_eq_Lemmas) + apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) + apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) + 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\set(policy2list p) \ all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ @@ -1065,16 +998,13 @@ lemma C_eq_All_untilSortedQ: (qsort (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l)) = C p" -apply (subst C_eq_sortedQ,simp_all add: C_eq_Lemmas) -apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) -apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) -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 *) - + apply (subst C_eq_sortedQ,simp_all add: C_eq_Lemmas) + apply (subst C_eq_RS2,simp_all add: C_eq_Lemmas) + apply (subst C_eq_rd,simp_all add: C_eq_Lemmas) + 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_untilSorted_withSimps: "DenyAll\set(policy2list p) \all_in_list(policy2list p) l \ allNetsDistinct (policy2list p) \ C (list2FWpolicy @@ -1082,76 +1012,70 @@ lemma C_eq_All_untilSorted_withSimps: (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l)) = C p" -by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) - - + by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) + lemma C_eq_All_untilSorted_withSimpsQ: " DenyAll\set(policy2list p)\all_in_list(policy2list p) l \ allNetsDistinct(policy2list p) \ C (list2FWpolicy (qsort (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l)) = C p" -by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) - + by (simp_all add: C_eq_Lemmas C_eq_subst_Lemmas) + lemma InDomConc[rule_format]: "p \ [] \ x \ dom (C (list2FWpolicy (p))) \ x \ dom (C (list2FWpolicy (a#p)))" -by (induct p, simp_all) (case_tac "p = []", simp_all add: dom_def C.simps) - - + by (induct p, simp_all) (case_tac "p = []", simp_all add: dom_def C.simps) + lemma not_in_member[rule_format]: "member a b \ x \ dom (C b) \ x \ dom (C a)" -by (induct b) (simp_all add: dom_def C.simps) - - + by (induct b) (simp_all add: dom_def C.simps) + lemma src_in_sdnets[rule_format]: "\ member DenyAll x \ p \ dom (C x) \ subnetsOfAdr (src p) \ (fst_set (sdnets x)) \ {}" -apply (induct rule: Combinators.induct) -apply (simp_all add: fst_set_def subnetsOfAdr_def PLemmas fst_set_def) -apply (intro impI) -apply (case_tac "p \ dom (C x2)") -apply (rule subnetAux) -apply (auto simp: PLemmas fst_set_def) -done - + apply (induct rule: Combinators.induct) + apply (simp_all add: fst_set_def subnetsOfAdr_def PLemmas fst_set_def) + apply (intro impI) + apply (case_tac "p \ dom (C x2)") + apply (rule subnetAux) + apply (auto simp: PLemmas fst_set_def) + done + lemma dest_in_sdnets[rule_format]: "\ member DenyAll x \ p \ dom (C x) \ subnetsOfAdr (dest p) \ (snd_set (sdnets x)) \ {}" -apply (induct rule: Combinators.induct) -apply (simp_all add: snd_set_def subnetsOfAdr_def PLemmas) -apply (intro impI) -apply (simp add: snd_set_def) -apply (case_tac "p \ dom (C x2)") -apply (rule subnetAux) -apply (auto simp: PLemmas) -done - - + apply (induct rule: Combinators.induct) + apply (simp_all add: snd_set_def subnetsOfAdr_def PLemmas) + apply (intro impI) + apply (simp add: snd_set_def) + apply (case_tac "p \ dom (C x2)") + apply (rule subnetAux) + apply (auto simp: PLemmas) + done lemma sdnets_in_subnets[rule_format]: "p\ dom (C x) \ \ member DenyAll x \ (\ (a,b)\sdnets x. a \ subnetsOfAdr (src p) \ b \ subnetsOfAdr (dest p))" -apply (rule Combinators.induct) -apply (simp_all add: PLemmas subnetsOfAdr_def) -apply (intro impI, simp) -apply (case_tac "p \ dom (C (x2))") -apply (auto simp: PLemmas subnetsOfAdr_def) -done + apply (rule Combinators.induct) + apply (simp_all add: PLemmas subnetsOfAdr_def) + apply (intro impI, simp) + apply (case_tac "p \ dom (C (x2))") + apply (auto simp: PLemmas subnetsOfAdr_def) + done lemma disjSD_no_p_in_both[rule_format]: - "disjSD_2 x y \ \ member DenyAll x \ \ member DenyAll y \ p \ dom(C x) \ p \ dom(C y) \ + "disjSD_2 x y \ \ member DenyAll x \ \ member DenyAll y \ p \ dom(C x) \ p \ dom(C y) \ False" -apply (rule_tac A = "sdnets x" and B = "sdnets y" and D = "src p" and F = "dest p" in tndFalse) -by (auto simp: dest_in_sdnets src_in_sdnets sdnets_in_subnets disjSD_2_def) - + apply (rule_tac A = "sdnets x" and B = "sdnets y" and D = "src p" and F = "dest p" in tndFalse) + by (auto simp: dest_in_sdnets src_in_sdnets sdnets_in_subnets disjSD_2_def) + lemma list2FWpolicy_eq: "zs \ [] \ C (list2FWpolicy (x \ y # z)) p = C (x \ list2FWpolicy (y # z)) p" -by (metis ConcAssoc l2p_aux list2FWpolicy.simps(2)) - - + by (metis ConcAssoc l2p_aux list2FWpolicy.simps(2)) + lemma dom_sep[rule_format]: "x \ dom (C (list2FWpolicy p)) \ x \ dom (C (list2FWpolicy(separate p)))" proof (induct p rule: separate.induct, simp_all, goal_cases) case (1 v va y z) then show ?case - apply (intro conjI impI) + apply (intro conjI impI) apply (simp,drule mp) apply (case_tac "x \ dom (C (DenyAllFromTo v va))") apply (metis CConcStartA domIff l2p_aux2 list2FWpolicyconc not_Cons_self ) @@ -1162,7 +1086,7 @@ next assume * : "{v, va} = first_bothNet y \ x \ dom (C (list2FWpolicy (AllowPortFromTo v va vb \ y # z))) \ x \ dom (C (list2FWpolicy (separate (AllowPortFromTo v va vb \ y # z))))" - and **: "{v, va} \ first_bothNet y \ + and **: "{v, va} \ first_bothNet y \ x \ dom(C(list2FWpolicy(y#z))) \ x \ dom (C(list2FWpolicy(separate(y#z))))" show ?case apply (insert * **, rule impI | rule conjI)+ @@ -1179,7 +1103,7 @@ next assume * : "(first_bothNet v = first_bothNet y \ x \ dom (C (list2FWpolicy ((v \ va) \ y # z))) \ x \ dom (C (list2FWpolicy (separate ((v \ va) \ y # z)))))" - and ** : "(first_bothNet v \ first_bothNet y \ + and ** : "(first_bothNet v \ first_bothNet y \ x \ dom(C(list2FWpolicy(y#z))) \ x \ dom (C (list2FWpolicy (separate (y # z)))))" show ?case apply (insert * **, rule conjI | rule impI)+ @@ -1190,134 +1114,122 @@ next apply (metis Conc_not_MT domIff list2FWpolicy_eq) by (metis CConcStartA Conc_not_MT InDomConc domIff nlpaux sepnMT) qed - + lemma domdConcStart[rule_format]: "x \ dom (C (list2FWpolicy (a#b))) \ x \ dom (C (list2FWpolicy b)) \ x \ dom (C (a))" -by (induct b, simp_all) (auto simp: PLemmas) - + by (induct b, simp_all) (auto simp: PLemmas) + lemma sep_dom2_aux: " x \ dom (C (list2FWpolicy (a \ y # z))) \ x \ dom (C (a \ list2FWpolicy (y # z)))" -by (auto) (metis list2FWpolicy_eq p2lNmt) - + by (auto) (metis list2FWpolicy_eq p2lNmt) lemma sep_dom2_aux2: -"x \ dom (C (list2FWpolicy (separate (y # z)))) \ x \ dom (C (list2FWpolicy (y # z))) \ + "x \ dom (C (list2FWpolicy (separate (y # z)))) \ x \ dom (C (list2FWpolicy (y # z))) \ x \ dom (C (list2FWpolicy (a # separate (y # z)))) \ x \ dom (C (list2FWpolicy (a \ y # z)))" -by (metis CConcStartA InDomConc domdConcStart list.simps(2) list2FWpolicy.simps(2) list2FWpolicyconc) - + by (metis CConcStartA InDomConc domdConcStart list.simps(2) list2FWpolicy.simps(2) list2FWpolicyconc) lemma sep_dom2[rule_format]: - "x \ dom (C (list2FWpolicy (separate p))) \ x \ dom (C (list2FWpolicy( p)))" -by (rule separate.induct) (simp_all add: sep_dom2_aux sep_dom2_aux2) + "x \ dom (C (list2FWpolicy (separate p))) \ x \ dom (C (list2FWpolicy( p)))" + by (rule separate.induct) (simp_all add: sep_dom2_aux sep_dom2_aux2) lemma sepDom: "dom (C (list2FWpolicy p)) = dom (C (list2FWpolicy (separate p)))" -apply (rule equalityI) -by (rule subsetI, (erule dom_sep|erule sep_dom2))+ + apply (rule equalityI) + by (rule subsetI, (erule dom_sep|erule sep_dom2))+ lemma C_eq_s_ext[rule_format]: "p \ [] \ C (list2FWpolicy (separate p)) a = C (list2FWpolicy p) a " proof (induct rule: separate.induct, goal_cases) print_cases - case (1 x) thus ?case - apply simp - apply (cases "x = []") - apply (metis l2p_aux2 separate.simps(5)) - apply simp - apply (cases "a \ dom (C (list2FWpolicy x))") - apply (subgoal_tac "a \ dom (C (list2FWpolicy (separate x)))") - apply (metis Cdom2 list2FWpolicyconc sepDom sepnMT) - apply (metis sepDom) - apply (subgoal_tac "a \ dom (C (list2FWpolicy (separate x)))") - apply (subst list2FWpolicyconc,simp add: sepnMT) - apply (subst list2FWpolicyconc,simp add: sepnMT) - apply (metis nlpaux sepDom) - apply (metis sepDom) - done + case (1 x) thus ?case + apply simp + apply (cases "x = []") + apply (metis l2p_aux2 separate.simps(5)) + apply simp + apply (cases "a \ dom (C (list2FWpolicy x))") + apply (subgoal_tac "a \ dom (C (list2FWpolicy (separate x)))") + apply (metis Cdom2 list2FWpolicyconc sepDom sepnMT) + apply (metis sepDom) + apply (subgoal_tac "a \ dom (C (list2FWpolicy (separate x)))") + apply (subst list2FWpolicyconc,simp add: sepnMT) + apply (subst list2FWpolicyconc,simp add: sepnMT) + apply (metis nlpaux sepDom) + apply (metis sepDom) + done next - case (2 v va y z) thus ?case - apply (cases "z = []", simp_all) - apply (rule conjI|rule impI|simp)+ - apply (subst list2FWpolicyconc) - apply (metis not_Cons_self sepnMT) - apply (metis C.simps(4) CConcStartaux Cdom2 domIff) - apply (rule conjI|rule impI|simp)+ - apply (erule list2FWpolicy_eq) - apply (rule impI, simp) - apply (subst list2FWpolicyconc) - apply (metis list.simps(2) sepnMT) - by (metis C.simps(4) CConcStartaux Cdom2 domIff) + case (2 v va y z) thus ?case + apply (cases "z = []", simp_all) + apply (rule conjI|rule impI|simp)+ + apply (subst list2FWpolicyconc) + apply (metis not_Cons_self sepnMT) + apply (metis C.simps(4) CConcStartaux Cdom2 domIff) + apply (rule conjI|rule impI|simp)+ + apply (erule list2FWpolicy_eq) + apply (rule impI, simp) + apply (subst list2FWpolicyconc) + apply (metis list.simps(2) sepnMT) + by (metis C.simps(4) CConcStartaux Cdom2 domIff) next - case (3 v va vb y z) thus ?case - apply (cases "z = []", simp_all) - apply (rule conjI|rule impI|simp)+ - apply (subst list2FWpolicyconc) - apply (metis not_Cons_self sepnMT) - apply (metis C.simps(4) CConcStartaux Cdom2 domIff) - apply (rule conjI|rule impI|simp)+ - apply (erule list2FWpolicy_eq) - apply (rule impI, simp) - apply (subst list2FWpolicyconc) - apply (metis list.simps(2) sepnMT) - by (metis C.simps(4) CConcStartaux Cdom2 domIff) + case (3 v va vb y z) thus ?case + apply (cases "z = []", simp_all) + apply (rule conjI|rule impI|simp)+ + apply (subst list2FWpolicyconc) + apply (metis not_Cons_self sepnMT) + apply (metis C.simps(4) CConcStartaux Cdom2 domIff) + apply (rule conjI|rule impI|simp)+ + apply (erule list2FWpolicy_eq) + apply (rule impI, simp) + apply (subst list2FWpolicyconc) + apply (metis list.simps(2) sepnMT) + by (metis C.simps(4) CConcStartaux Cdom2 domIff) next - case (4 v va y z) thus ?case - apply (cases "z = []", simp_all) - apply (rule conjI|rule impI|simp)+ - apply (subst list2FWpolicyconc) - apply (metis not_Cons_self sepnMT) - apply (metis C.simps(4) CConcStartaux Cdom2 domIff) - apply (rule conjI|rule impI|simp)+ - apply (erule list2FWpolicy_eq) - apply (rule impI, simp) - apply (subst list2FWpolicyconc) - apply (metis list.simps(2) sepnMT) - by (metis C.simps(4) CConcStartaux Cdom2 domIff) + case (4 v va y z) thus ?case + apply (cases "z = []", simp_all) + apply (rule conjI|rule impI|simp)+ + apply (subst list2FWpolicyconc) + apply (metis not_Cons_self sepnMT) + apply (metis C.simps(4) CConcStartaux Cdom2 domIff) + apply (rule conjI|rule impI|simp)+ + apply (erule list2FWpolicy_eq) + apply (rule impI, simp) + apply (subst list2FWpolicyconc) + apply (metis list.simps(2) sepnMT) + by (metis C.simps(4) CConcStartaux Cdom2 domIff) next - case 5 thus ?case by simp + case 5 thus ?case by simp next - case 6 thus ?case by simp + case 6 thus ?case by simp next - case 7 thus ?case by simp + case 7 thus ?case by simp next - case 8 thus ?case by simp + case 8 thus ?case by simp qed - lemma C_eq_s: -"p \ [] \ C (list2FWpolicy (separate p)) = C (list2FWpolicy p)" -apply (rule ext) using C_eq_s_ext by blast - - -(*MOVE FORWARD*) + "p \ [] \ C (list2FWpolicy (separate p)) = C (list2FWpolicy p)" + apply (rule ext) using C_eq_s_ext by blast + lemma sortnMTQ: "p \ [] \ qsort p l \ []" -by (metis set_sortQ setnMT) - + 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\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct (policy2list p) \ + " DenyAll\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct (policy2list p) \ C (list2FWpolicy (separate (FWNormalisationCore.sort (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l))) = C p" -by (simp add: C_eq_All_untilSorted_withSimps C_eq_s wellformed1_alternative_sorted wp1ID wp1n_RS2) - + by (simp add: C_eq_All_untilSorted_withSimps C_eq_s wellformed1_alternative_sorted wp1ID wp1n_RS2) lemma C_eq_until_separatedQ: -"DenyAll\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct(policy2list p) \ + "DenyAll\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct(policy2list p) \ 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) - - - + by (simp add: C_eq_All_untilSorted_withSimpsQ C_eq_s sortnMTQ wp1ID wp1n_RS2) lemma domID[rule_format]: "p \ [] \ x \ dom(C(list2FWpolicy p)) \ x \ dom (C(list2FWpolicy(insertDenies p)))" @@ -1325,65 +1237,64 @@ proof(induct p) case Nil then show ?case by simp next case (Cons a p) then show ?case - proof(cases "p=[]",goal_cases) - case 1 then show ?case - apply(simp) apply(rule impI) - apply (cases a, simp_all) - apply (simp_all add: C.simps dom_def)+ - by auto - next - case 2 then show ?case - proof(cases "x \ dom(C(list2FWpolicy p))", goal_cases) - case 1 then show ?case - apply simp apply (rule impI) - apply (cases a, simp_all) - using InDomConc idNMT apply blast - apply (rule InDomConc, simp_all add: idNMT)+ - done - next - case 2 then show ?case - apply simp apply (rule impI) - proof(cases "x \ dom (C (list2FWpolicy (insertDenies p)))", goal_cases) - case 1 then show ?case - proof(induct a) - case DenyAll then show ?case by simp - next - case (DenyAllFromTo src dest) then show ?case - apply simp by( rule InDomConc, simp add: idNMT) - next - case (AllowPortFromTo src dest port) then show ?case - apply simp by(rule InDomConc, simp add: idNMT) - next - case (Conc _ _) then show ?case - apply simp by(rule InDomConc, simp add: idNMT) - qed - next - case 2 then show ?case - proof (induct a) - case DenyAll then show ?case by simp - next - case (DenyAllFromTo src dest) then show ?case - by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) - next - case (AllowPortFromTo src dest port) then show ?case - by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) - next - case (Conc _ _) then show ?case - by (simp,metis CConcStartA Cdom2 domIff domdConcStart) - qed - qed - qed + proof(cases "p=[]",goal_cases) + case 1 then show ?case + apply(simp) apply(rule impI) + apply (cases a, simp_all) + apply (simp_all add: C.simps dom_def)+ + by auto + next + case 2 then show ?case + proof(cases "x \ dom(C(list2FWpolicy p))", goal_cases) + case 1 then show ?case + apply simp apply (rule impI) + apply (cases a, simp_all) + using InDomConc idNMT apply blast + apply (rule InDomConc, simp_all add: idNMT)+ + done + next + case 2 then show ?case + apply simp apply (rule impI) + proof(cases "x \ dom (C (list2FWpolicy (insertDenies p)))", goal_cases) + case 1 then show ?case + proof(induct a) + case DenyAll then show ?case by simp + next + case (DenyAllFromTo src dest) then show ?case + apply simp by( rule InDomConc, simp add: idNMT) + next + case (AllowPortFromTo src dest port) then show ?case + apply simp by(rule InDomConc, simp add: idNMT) + next + case (Conc _ _) then show ?case + apply simp by(rule InDomConc, simp add: idNMT) + qed + next + case 2 then show ?case + proof (induct a) + case DenyAll then show ?case by simp + next + case (DenyAllFromTo src dest) then show ?case + by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) + next + case (AllowPortFromTo src dest port) then show ?case + by(simp,metis domIff CConcStartA list2FWpolicyconc nlpaux Cdom2) + next + case (Conc _ _) then show ?case + by (simp,metis CConcStartA Cdom2 domIff domdConcStart) + qed + qed qed - qed + qed +qed - lemma DA_is_deny: "x \ dom (C (DenyAllFromTo a b \ DenyAllFromTo b a \ DenyAllFromTo a b)) \ C (DenyAllFromTo a b\DenyAllFromTo b a \ DenyAllFromTo a b) x = Some (deny ())" -apply (case_tac "x \ dom (C (DenyAllFromTo a b))") -apply (simp_all add: PLemmas) -apply (simp_all split: if_splits) -done + apply (case_tac "x \ dom (C (DenyAllFromTo a b))") + apply (simp_all add: PLemmas) + apply (simp_all split: if_splits) + done lemma iDdomAux[rule_format]: "p \ [] \ x \ dom (C (list2FWpolicy p)) \ @@ -1393,120 +1304,120 @@ proof (induct p) case Nil thus ?case by simp next case (Cons y ys) thus ?case - proof (cases y) - case DenyAll then show ?thesis by simp + proof (cases y) + case DenyAll then show ?thesis by simp + next + case (DenyAllFromTo a b) then show ?thesis using DenyAllFromTo Cons + apply simp + apply (intro impI) + proof (cases "ys = []", goal_cases) + case 1 then show ?case by (simp add: DA_is_deny) next - case (DenyAllFromTo a b) then show ?thesis using DenyAllFromTo Cons + case 2 then show ?case apply simp - apply (intro impI) - proof (cases "ys = []", goal_cases) - case 1 then show ?case by (simp add: DA_is_deny) - next - case 2 then show ?case - apply simp - apply (drule mp) - apply (metis DenyAllFromTo InDomConc ) - apply (cases "x \ dom (C (list2FWpolicy (insertDenies ys)))", simp_all) - apply (metis Cdom2 DenyAllFromTo idNMT list2FWpolicyconc) - apply (subgoal_tac "C (list2FWpolicy (DenyAllFromTo a b \ + apply (drule mp) + apply (metis DenyAllFromTo InDomConc ) + apply (cases "x \ dom (C (list2FWpolicy (insertDenies ys)))", simp_all) + apply (metis Cdom2 DenyAllFromTo idNMT list2FWpolicyconc) + apply (subgoal_tac "C (list2FWpolicy (DenyAllFromTo a b \ DenyAllFromTo b a \ DenyAllFromTo a b#insertDenies ys)) x = C ((DenyAllFromTo a b \ DenyAllFromTo b a \ DenyAllFromTo a b)) x ") - apply simp - apply (rule DA_is_deny) - apply (metis DenyAllFromTo domdConcStart) - apply (metis DenyAllFromTo l2p_aux2 list2FWpolicyconc nlpaux) - done - qed + apply simp + apply (rule DA_is_deny) + apply (metis DenyAllFromTo domdConcStart) + apply (metis DenyAllFromTo l2p_aux2 list2FWpolicyconc nlpaux) + done + qed + next + case (AllowPortFromTo a b c) then show ?thesis using Cons AllowPortFromTo + proof (cases "ys = []", goal_cases) + case 1 then show ?case + apply simp + apply (intro impI) + apply (subgoal_tac "x \ dom (C (DenyAllFromTo a b \ DenyAllFromTo b a))") + apply (simp_all add: PLemmas) + apply (simp split: if_splits, auto) + done next - case (AllowPortFromTo a b c) then show ?thesis using Cons AllowPortFromTo - proof (cases "ys = []", goal_cases) - case 1 then show ?case - apply simp - apply (intro impI) - apply (subgoal_tac "x \ dom (C (DenyAllFromTo a b \ DenyAllFromTo b a))") - apply (simp_all add: PLemmas) - apply (simp split: if_splits, auto) - done - next - case 2 then show ?case - apply simp - apply (intro impI) - apply (drule mp) - apply (metis AllowPortFromTo InDomConc) - apply (cases "x \ dom (C (list2FWpolicy (insertDenies ys)))") - apply simp_all - apply (metis AllowPortFromTo Cdom2 idNMT list2FWpolicyconc) - apply (subgoal_tac "C (list2FWpolicy (DenyAllFromTo a b \ + case 2 then show ?case + apply simp + apply (intro impI) + apply (drule mp) + apply (metis AllowPortFromTo InDomConc) + apply (cases "x \ dom (C (list2FWpolicy (insertDenies ys)))") + apply simp_all + apply (metis AllowPortFromTo Cdom2 idNMT list2FWpolicyconc) + apply (subgoal_tac "C (list2FWpolicy (DenyAllFromTo a b \ DenyAllFromTo b a \ AllowPortFromTo a b c#insertDenies ys)) x = C ((DenyAllFromTo a b \ DenyAllFromTo b a)) x ") - apply simp - defer 1 - apply (metis AllowPortFromTo CConcStartA ConcAssoc idNMT list2FWpolicyconc nlpaux) - apply (simp add: PLemmas, simp split: if_splits, auto) - done - qed - next - case (Conc a b) then show ?thesis - proof (cases "ys = []", goal_cases) - case 1 then show ?case - apply simp - apply (rule impI)+ - apply (subgoal_tac "x \ dom (C (DenyAllFromTo (first_srcNet a) + apply simp + defer 1 + apply (metis AllowPortFromTo CConcStartA ConcAssoc idNMT list2FWpolicyconc nlpaux) + apply (simp add: PLemmas, simp split: if_splits, auto) + done + qed + next + case (Conc a b) then show ?thesis + proof (cases "ys = []", goal_cases) + case 1 then show ?case + apply simp + apply (rule impI)+ + apply (subgoal_tac "x \ dom (C (DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a)))") - apply (simp_all add: PLemmas) - apply (simp split: if_splits, auto) - done - next - case 2 then show ?case - apply simp - apply (intro impI) - apply (cases "x \ dom (C (list2FWpolicy (insertDenies ys)))") - apply (metis Cdom2 Conc Cons InDomConc idNMT list2FWpolicyconc) - apply (subgoal_tac "C (list2FWpolicy (DenyAllFromTo (first_srcNet a) + apply (simp_all add: PLemmas) + apply (simp split: if_splits, auto) + done + next + case 2 then show ?case + apply simp + apply (intro impI) + apply (cases "x \ dom (C (list2FWpolicy (insertDenies ys)))") + apply (metis Cdom2 Conc Cons InDomConc idNMT list2FWpolicyconc) + apply (subgoal_tac "C (list2FWpolicy (DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo(first_destNet a)(first_srcNet a) \ a \ b#insertDenies ys)) x = C ((DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a \ b)) x") - apply simp - defer 1 - apply (metis Conc l2p_aux2 list2FWpolicyconc nlpaux) - apply (subgoal_tac "C((DenyAllFromTo (first_srcNet a) + apply simp + defer 1 + apply (metis Conc l2p_aux2 list2FWpolicyconc nlpaux) + apply (subgoal_tac "C((DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a \ b)) x = C((DenyAllFromTo (first_srcNet a)(first_destNet a) \ DenyAllFromTo(first_destNet a)(first_srcNet a))) x ") - apply simp - defer 1 - apply (metis CConcStartA Conc ConcAssoc nlpaux) - apply (simp add: PLemmas, simp split: if_splits, auto) - done - qed + apply simp + defer 1 + apply (metis CConcStartA Conc ConcAssoc nlpaux) + apply (simp add: PLemmas, simp split: if_splits, auto) + done qed + qed qed - + lemma iD_isD[rule_format]: "p \ [] \ x \ dom (C (list2FWpolicy p)) \ C (DenyAll \ list2FWpolicy (insertDenies p)) x = C DenyAll x" -apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies p)))") -apply (simp add: Cdom2 PLemmas(1) deny_all_def iDdomAux) -by (simp add: nlpaux) + apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies p)))") + apply (simp add: Cdom2 PLemmas(1) deny_all_def iDdomAux) + by (simp add: nlpaux) lemma inDomConc:"\ x\dom (C a); x\dom (C (list2FWpolicy p))\ \ x \ dom (C (list2FWpolicy(a#p)))" -by (metis domdConcStart) - + by (metis domdConcStart) + lemma domsdisj[rule_format]: "p \ [] \ (\ x s. s \ set p \ x \ dom (C A) \ x \ dom (C s)) \ y \ dom (C A) \ y \ dom (C (list2FWpolicy p))" proof (induct p) - case Nil show ?case by simp + case Nil show ?case by simp next - case (Cons a p) then show ?case - apply (case_tac "p = []") - apply fastforce - by (meson domdConcStart list.set_intros(1) list.set_intros(2)) + case (Cons a p) then show ?case + apply (case_tac "p = []") + apply fastforce + by (meson domdConcStart list.set_intros(1) list.set_intros(2)) qed lemma isSepaux: @@ -1514,104 +1425,97 @@ lemma isSepaux: x \ dom (C (DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a)) \ x \ dom (C (list2FWpolicy p))" -apply (rule_tac A = "(DenyAllFromTo (first_srcNet a) (first_destNet a) \ + apply (rule_tac A = "(DenyAllFromTo (first_srcNet a) (first_destNet a) \ DenyAllFromTo (first_destNet a) (first_srcNet a) \ a)" in domsdisj) -apply simp_all -by (metis Combinators.distinct(1) FWNormalisationCore.member.simps(1) - FWNormalisationCore.member.simps(3) disjSD2aux disjSD_no_p_in_both noDA) - + 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 \ 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) + apply(induct p rule: separate.induct) + by (simp_all add: C.simps map_add_le_mapE map_le_antisym) lemma dom_id: "noDenyAll(a#p) \ separated(a#p) \ p \ [] \ x\dom(C(list2FWpolicy p)) \ x\dom (C a) \ x \ dom (C (list2FWpolicy (insertDenies p)))" -apply (rule_tac a = a in isSepaux, simp_all) -using idNMT apply blast -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) - + apply (rule_tac a = a in isSepaux, simp_all) + using idNMT apply blast + 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 \ separated p\ p \ []\ x \ dom (C (list2FWpolicy p))\ + "noDenyAll1 p \ separated p\ p \ []\ x \ dom (C (list2FWpolicy p))\ C(list2FWpolicy (insertDenies p)) x = C(list2FWpolicy p) x" proof (induct p) case Nil thus ?case by simp next case (Cons y ys) thus ?case using Cons - proof (cases y) - case DenyAll thus ?thesis using Cons DenyAll apply simp - apply (case_tac "ys = []", simp_all) - apply (case_tac "x \ dom (C (list2FWpolicy ys))",simp_all) - apply (metis Cdom2 domID idNMT list2FWpolicyconc noDA1eq) - apply (metis DenyAll iD_isD idNMT list2FWpolicyconc nlpaux) - done - next - case (DenyAllFromTo a b) thus ?thesis using Cons apply simp - apply (rule impI|rule allI|rule conjI|simp)+ - apply (case_tac "ys = []", simp_all) - apply (metis Cdom2 ConcAssoc DenyAllFromTo) - apply (case_tac "x \ dom (C (list2FWpolicy ys))", simp_all) - apply (simp add: Cdom2 domID idNMT l2p_aux noDA1eq) - apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies ys)))") - apply (meson Combinators.distinct(1) FWNormalisationCore.member.simps(3) dom_id domdConcStart - noDenyAll.simps(1) separated.simps(1)) - by (metis Cdom2 DenyAllFromTo domIff dom_def domdConcStart l2p_aux l2p_aux2 nlpaux) - next - case (AllowPortFromTo a b c) thus ?thesis - using AllowPortFromTo Cons apply simp - apply (rule impI|rule allI|rule conjI|simp)+ - apply (case_tac "ys = []", simp_all) - apply (metis Cdom2 ConcAssoc AllowPortFromTo) - apply (case_tac "x \ dom (C (list2FWpolicy ys))", simp_all) - apply (simp add: Cdom2 domID idNMT list2FWpolicyconc noDA1eq) - apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies ys)))") - apply (meson Combinators.distinct(3) FWNormalisationCore.member.simps(4) dom_id domdConcStart noDenyAll.simps(1) separated.simps(1)) - by (metis Cdom2 ConcAssoc l2p_aux list2FWpolicy.simps(2) nlpaux) - next - case (Conc a b) thus ?thesis using Cons Conc - apply simp - apply (rule impI|rule allI|rule conjI|simp)+ - apply (case_tac "ys = []", simp_all) - apply (metis Cdom2 ConcAssoc Conc) - apply (case_tac "x \ dom (C (list2FWpolicy ys))",simp_all) - apply (simp add: Cdom2 domID idNMT list2FWpolicyconc noDA1eq) - apply (case_tac "x \ dom (C (a \ b))") - apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies ys)))",simp_all) - apply (simp add: Cdom2 domIff idNMT list2FWpolicyconc nlpaux) - apply (metis FWNormalisationCore.member.simps(1) dom_id noDenyAll.simps(1) separated.simps(1)) - by (simp add: inDomConc) - qed + proof (cases y) + case DenyAll thus ?thesis using Cons DenyAll apply simp + apply (case_tac "ys = []", simp_all) + apply (case_tac "x \ dom (C (list2FWpolicy ys))",simp_all) + apply (metis Cdom2 domID idNMT list2FWpolicyconc noDA1eq) + apply (metis DenyAll iD_isD idNMT list2FWpolicyconc nlpaux) + done + next + case (DenyAllFromTo a b) thus ?thesis using Cons apply simp + apply (rule impI|rule allI|rule conjI|simp)+ + apply (case_tac "ys = []", simp_all) + apply (metis Cdom2 ConcAssoc DenyAllFromTo) + apply (case_tac "x \ dom (C (list2FWpolicy ys))", simp_all) + apply (simp add: Cdom2 domID idNMT l2p_aux noDA1eq) + apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies ys)))") + apply (meson Combinators.distinct(1) FWNormalisationCore.member.simps(3) dom_id domdConcStart + noDenyAll.simps(1) separated.simps(1)) + by (metis Cdom2 DenyAllFromTo domIff dom_def domdConcStart l2p_aux l2p_aux2 nlpaux) + next + case (AllowPortFromTo a b c) thus ?thesis + using AllowPortFromTo Cons apply simp + apply (rule impI|rule allI|rule conjI|simp)+ + apply (case_tac "ys = []", simp_all) + apply (metis Cdom2 ConcAssoc AllowPortFromTo) + apply (case_tac "x \ dom (C (list2FWpolicy ys))", simp_all) + apply (simp add: Cdom2 domID idNMT list2FWpolicyconc noDA1eq) + apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies ys)))") + apply (meson Combinators.distinct(3) FWNormalisationCore.member.simps(4) dom_id domdConcStart noDenyAll.simps(1) separated.simps(1)) + by (metis Cdom2 ConcAssoc l2p_aux list2FWpolicy.simps(2) nlpaux) + next + case (Conc a b) thus ?thesis using Cons Conc + apply simp + apply (rule impI|rule allI|rule conjI|simp)+ + apply (case_tac "ys = []", simp_all) + apply (metis Cdom2 ConcAssoc Conc) + apply (case_tac "x \ dom (C (list2FWpolicy ys))",simp_all) + apply (simp add: Cdom2 domID idNMT list2FWpolicyconc noDA1eq) + apply (case_tac "x \ dom (C (a \ b))") + apply (case_tac "x \ dom (C (list2FWpolicy (insertDenies ys)))",simp_all) + apply (simp add: Cdom2 domIff idNMT list2FWpolicyconc nlpaux) + apply (metis FWNormalisationCore.member.simps(1) dom_id noDenyAll.simps(1) separated.simps(1)) + by (simp add: inDomConc) + qed qed lemma C_eq_iD: "separated p \ noDenyAll1 p \ wellformed_policy1_strong p \ 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*) + by (rule ext) (metis CConcStartA C_eq_iD_aux2 DAAux wp1_alternative_not_mt wp1n_tl) lemma noDAsortQ[rule_format]: "noDenyAll1 p \ noDenyAll1 (qsort p l)" -apply (case_tac "p",simp_all, rename_tac a list) - apply (case_tac "a = DenyAll",simp_all) - using nDAeqSet set_sortQ apply blast -apply (rule impI,rule noDA1eq) - apply (subgoal_tac "noDenyAll (a#list)") -apply (metis append_Cons append_Nil nDAeqSet qsort.simps(2) set_sortQ) -by (case_tac a, simp_all) - - -(*MOVE FORWARD*) - + apply (case_tac "p",simp_all, rename_tac a list) + apply (case_tac "a = DenyAll",simp_all) + using nDAeqSet set_sortQ apply blast + apply (rule impI,rule noDA1eq) + apply (subgoal_tac "noDenyAll (a#list)") + apply (metis append_Cons append_Nil nDAeqSet qsort.simps(2) set_sortQ) + by (case_tac a, simp_all) + lemma NetsCollectedSortQ: "distinct p \noDenyAll1 p \ all_in_list p l \ singleCombinators p \ NetsCollected (qsort p l)" -by (metis NetsCollectedSorted SC3Q all_in_list.elims(2) all_in_list.simps(1) all_in_list.simps(2) - all_in_listAppend all_in_list_sublist noDAsortQ qsort.simps(1) qsort.simps(2) - singleCombinatorsConc sort_is_sortedQ) + by (metis NetsCollectedSorted SC3Q all_in_list.elims(2) all_in_list.simps(1) all_in_list.simps(2) + all_in_listAppend all_in_list_sublist noDAsortQ qsort.simps(1) qsort.simps(2) + singleCombinatorsConc sort_is_sortedQ) lemmas CLemmas = nMTSort nMTSortQ none_MT_rulesRS2 none_MT_rulesrd @@ -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\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct(policy2list p) \ C (list2FWpolicy @@ -1635,9 +1537,9 @@ lemma C_eq_Until_InsertDenies: (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l)))) = C p" -apply (subst C_eq_iD,simp_all add: C_eqLemmas_id) -apply (rule C_eq_until_separated, simp_all) -done + apply (subst C_eq_iD,simp_all add: C_eqLemmas_id) + apply (rule C_eq_until_separated, simp_all) + done lemma C_eq_Until_InsertDeniesQ: "DenyAll\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct(policy2list p) \ @@ -1646,16 +1548,14 @@ lemma C_eq_Until_InsertDeniesQ: (separate (qsort (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l)))) = C p" -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) - - + 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) - - + apply (induct p,simp_all) + by (metis Cdom2 domIff nlpaux not_in_member) + lemma C_eq_RAD_aux[rule_format]: "p \ [] \ 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 \ [] \ C (list2FWpolicy p) = C (list2FWpolicy (removeAllDuplicates p)) " -by (rule ext,erule C_eq_RAD_aux) - - + by (rule ext,erule C_eq_RAD_aux) lemma C_eq_compile: "DenyAll\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct(policy2list p) \ C (list2FWpolicy @@ -1686,10 +1583,9 @@ lemma C_eq_compile: (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l))))) = C p" -apply (subst C_eq_RAD[symmetric]) -apply (rule idNMT,simp add: C_eqLemmas_id) -by (rule C_eq_Until_InsertDenies, simp_all) - + 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\set(policy2list p) \ all_in_list(policy2list p)l \ allNetsDistinct(policy2list p) \ @@ -1700,96 +1596,78 @@ lemma C_eq_compileQ: (qsort (removeShadowRules2 (remdups (rm_MT_rules C (insertDeny (removeShadowRules1 (policy2list p)))))) l))))) = C p" -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) - + 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 \ set (policy2list p) \ allNetsDistinct (policy2list p) \ + "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ all_in_list(policy2list p)(Nets_List p) \ C (list2FWpolicy (normalize p)) = C p" -unfolding normalize_def -by (simp add: C_eq_compile) - + unfolding normalize_def + by (simp add: C_eq_compile) lemma C_eq_normalizeQ: "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list p) \ all_in_list (policy2list p) (Nets_List p) \ C (list2FWpolicy (normalizeQ p)) = C p" -by (simp add: normalizeQ_def C_eq_compileQ) - + by (simp add: normalizeQ_def C_eq_compileQ) lemma domSubset3: "dom (C (DenyAll \ x)) = dom (C (DenyAll))" -by (simp add: PLemmas split_tupled_all split: option.splits) - + by (simp add: PLemmas split_tupled_all split: option.splits) lemma domSubset4: "dom (C (DenyAllFromTo x y \ DenyAllFromTo y x \ AllowPortFromTo x y dn)) = dom (C (DenyAllFromTo x y \ DenyAllFromTo y x))" -by (auto simp: PLemmas split: option.splits decision.splits ) - + by (auto simp: PLemmas split: option.splits decision.splits ) lemma domSubset5: - "dom (C (DenyAllFromTo x y \ DenyAllFromTo y x \ AllowPortFromTo y x dn)) = + "dom (C (DenyAllFromTo x y \ DenyAllFromTo y x \ AllowPortFromTo y x dn)) = dom (C (DenyAllFromTo x y \ DenyAllFromTo y x))" -by (auto simp: PLemmas split: option.splits decision.splits ) - - + by (auto simp: PLemmas split: option.splits decision.splits ) lemma domSubset1: "dom (C (DenyAllFromTo one two \ DenyAllFromTo two one \ AllowPortFromTo one two dn \ x)) = dom (C (DenyAllFromTo one two \ DenyAllFromTo two one \ x))" -by (simp add: PLemmas split: option.splits decision.splits) (auto simp: allow_all_def deny_all_def) - + by (simp add: PLemmas split: option.splits decision.splits) (auto simp: allow_all_def deny_all_def) lemma domSubset2: - "dom (C (DenyAllFromTo one two \ DenyAllFromTo two one \ AllowPortFromTo two one dn \ x)) = + "dom (C (DenyAllFromTo one two \ DenyAllFromTo two one \ AllowPortFromTo two one dn \ x)) = dom (C (DenyAllFromTo one two \ DenyAllFromTo two one \ x))" -by (simp add: PLemmas split: option.splits decision.splits) (auto simp: allow_all_def deny_all_def) - + by (simp add: PLemmas split: option.splits decision.splits) (auto simp: allow_all_def deny_all_def) lemma ConcAssoc2: "C (X \ Y \ ((A \ B) \ D)) = C (X \ Y \ A \ B \ D)" -by (simp add: C.simps) - + by (simp add: C.simps) lemma ConcAssoc3: "C (X \ ((Y \ A) \ D)) = C (X \ Y \ A \ D)" -by (simp add: C.simps) - + by (simp add: C.simps) lemma RS3_NMT[rule_format]: "DenyAll \ set p \ rm_MT_rules C p \ []" -by (induct_tac p) (simp_all add: PLemmas) - + by (induct_tac p) (simp_all add: PLemmas) lemma norm_notMT: "DenyAll \ set (policy2list p) \ normalize p \ []" -by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalize_def rADnMT sepnMT sortnMT) - + by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalize_def rADnMT sepnMT sortnMT) lemma norm_notMTQ: "DenyAll \ set (policy2list p) \ normalizeQ p \ []" -by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalizeQ_def rADnMT sepnMT sortnMTQ) - + 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 \ dom (C a) \ C (list2policyR (a # list)) p = C a p" -by (induct "a # list" rule:list2policyR.induct) (auto simp: C.simps dom_def map_add_def) - + by (induct "a # list" rule:list2policyR.induct) (auto simp: C.simps dom_def map_add_def) lemma list2policyR_End: "p \ dom (C a) \ C (list2policyR (a # list)) p = (C a \ list2policy (map C list)) p" -by (rule list2policyR.induct) - (simp_all add: C.simps dom_def map_add_def list2policy_def split: option.splits) - + 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 \ [] \ C(list2policyR N) p = (list2policy (map C N)) p" @@ -1797,113 +1675,98 @@ proof (induct N) case Nil show ?case by (simp_all add: list2policy_def) next case (Cons a N) then show ?case - apply (case_tac "p \ dom (C a)",simp_all add: domStart list2policy_def) - apply (rule list2policyR_Start, simp_all) - apply (rule list2policyR.induct, simp_all) - apply (simp_all add: C.simps dom_def map_add_def) + apply (case_tac "p \ dom (C a)",simp_all add: domStart list2policy_def) + apply (rule list2policyR_Start, simp_all) + apply (rule list2policyR.induct, simp_all) + apply (simp_all add: C.simps dom_def map_add_def) apply (simp split: option.splits) done qed - lemma l2polR_eq: "N \ [] \ C( list2policyR N) = (list2policy (map C N))" -by (auto simp: list2policy_def l2polR_eq_el ) - + by (auto simp: list2policy_def l2polR_eq_el ) lemma list2FWpolicys_eq_el[rule_format]: - "Filter \ [] \ C (list2policyR Filter) p = C (list2FWpolicy (rev Filter)) p" + "Filter \ [] \ C (list2policyR Filter) p = C (list2FWpolicy (rev Filter)) p" proof (induct Filter) print_cases case Nil show ?case by (simp) next - case (Cons a list) then show ?case - apply simp_all - apply (case_tac "list = []", simp_all) - apply (case_tac "p \ dom (C a)", simp_all) + case (Cons a list) then show ?case + apply simp_all + apply (case_tac "list = []", simp_all) + apply (case_tac "p \ dom (C a)", simp_all) apply (rule list2policyR_Start, simp_all) - by (metis C.simps(4) l2polR_eq list2policyR_End nlpaux) + by (metis C.simps(4) l2polR_eq list2policyR_End nlpaux) qed lemma list2FWpolicys_eq: - "Filter \ [] \ C (list2policyR Filter) = C (list2FWpolicy (rev Filter))" -by (rule ext, erule list2FWpolicys_eq_el) - + "Filter \ [] \ C (list2policyR Filter) = C (list2FWpolicy (rev Filter))" + by (rule ext, erule list2FWpolicys_eq_el) lemma list2FWpolicys_eq_sym: - "Filter \ [] \C (list2policyR (rev Filter)) = C (list2FWpolicy Filter)" -by (metis list2FWpolicys_eq rev_is_Nil_conv rev_rev_ident) - + "Filter \ [] \C (list2policyR (rev Filter)) = C (list2FWpolicy Filter)" + by (metis list2FWpolicys_eq rev_is_Nil_conv rev_rev_ident) lemma p_eq[rule_format]: "p \ [] \ list2policy (map C (rev p)) = C (list2FWpolicy p)" -by (metis l2polR_eq list2FWpolicys_eq_sym rev.simps(1) rev_rev_ident) - + by (metis l2polR_eq list2FWpolicys_eq_sym rev.simps(1) rev_rev_ident) lemma p_eq2[rule_format]: "normalize x \ [] \ C(list2FWpolicy(normalize x)) = C x \ list2policy(map C (rev(normalize x))) = C x" -by (simp add: p_eq) - + by (simp add: p_eq) lemma p_eq2Q[rule_format]: "normalizeQ x \ [] \ C (list2FWpolicy (normalizeQ x)) = C x \ list2policy (map C (rev (normalizeQ x))) = C x" -by (simp add: p_eq) - + by (simp add: p_eq) lemma list2listNMT[rule_format]: "x \ [] \map sem x \ []" -by (case_tac x) simp_all - + by (case_tac x) simp_all lemma Norm_Distr2: - "r o_f ((P \\<^sub>2 (list2policy Q)) o d) = (list2policy ((P \\<^sub>L Q) (op \\<^sub>2) r d))" -by (rule ext, rule Norm_Distr_2) - + "r o_f ((P \\<^sub>2 (list2policy Q)) o d) = (list2policy ((P \\<^sub>L Q) (op \\<^sub>2) r d))" + by (rule ext, rule Norm_Distr_2) lemma NATDistr: "N \ [] \ F = C (list2policyR N) \ (\(x, y). x) o\<^sub>f (NAT \\<^sub>2 F \ (\x. (x, x))) = list2policy ((NAT \\<^sub>L map C N) op \\<^sub>2 (\(x, y). x) (\x. (x, x)))" -apply (simp add: l2polR_eq) -apply (rule ext) -apply (rule Norm_Distr_2) -done - + apply (simp add: l2polR_eq) + apply (rule ext) + apply (rule Norm_Distr_2) + done lemma C_eq_normalize_manual: -"DenyAll\set(policy2list p) \ allNetsDistinct(policy2list p) \ all_in_list(policy2list p) l \ + "DenyAll\set(policy2list p) \ allNetsDistinct(policy2list p) \ all_in_list(policy2list p) l \ C (list2FWpolicy (normalize_manual_order p l)) = C p" -by (simp add: normalize_manual_order_def C_eq_compile) - + by (simp add: normalize_manual_order_def C_eq_compile) lemma p_eq2_manualQ[rule_format]: "normalize_manual_orderQ x l \ [] \ C(list2FWpolicy (normalize_manual_orderQ x l)) = C x \ list2policy (map C (rev (normalize_manual_orderQ x l))) = C x" -by (simp add: p_eq) - + by (simp add: p_eq) lemma norm_notMT_manualQ: "DenyAll \ set (policy2list p) \ normalize_manual_orderQ p l \ []" -by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalize_manual_orderQ_def rADnMT sepnMT sortnMTQ) - + by (simp add: DAiniD RS2_NMT RS3_NMT idNMT normalize_manual_orderQ_def rADnMT sepnMT sortnMTQ) lemma C_eq_normalize_manualQ: "DenyAll\set(policy2list p) \ allNetsDistinct(policy2list p) \ all_in_list(policy2list p) l \ C (list2FWpolicy (normalize_manual_orderQ p l)) = C p" -by (simp add: normalize_manual_orderQ_def C_eq_compileQ) - + by (simp add: normalize_manual_orderQ_def C_eq_compileQ) lemma p_eq2_manual[rule_format]: "normalize_manual_order x l \ [] \ C (list2FWpolicy (normalize_manual_order x l)) = C x \ list2policy (map C (rev (normalize_manual_order x l))) = C x" -by (simp add: p_eq) - + by (simp add: p_eq) lemma norm_notMT_manual: "DenyAll \ set (policy2list p) \ normalize_manual_order p l \ []" -by (simp add: RS2_NMT idNMT normalize_manual_order_def rADnMT sepnMT sortnMT wp1ID) + 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 \ set (policy2list Filter) \ allNetsDistinct (policy2list Filter) \ @@ -1911,36 +1774,31 @@ lemma normalizeNAT: (\(x, y). x) o\<^sub>f (NAT \\<^sub>2 C Filter \ (\x. (x, x))) = list2policy ((NAT \\<^sub>L map C (rev (FWNormalisationCore.normalize Filter))) op \\<^sub>2 (\(x, y). x) (\x. (x, x)))" -by (simp add: C_eq_normalize NATDistr list2FWpolicys_eq_sym norm_notMT) - + by (simp add: C_eq_normalize NATDistr list2FWpolicys_eq_sym norm_notMT) lemma domSimpl[simp]: "dom (C (A \ DenyAll)) = dom (C (DenyAll))" -by (simp add: PLemmas) - - -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. *} + by (simp add: PLemmas) +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. +*} lemma CRotate_eq_rotateC: "CRotate p = C (rotatePolicy p)" -by (induct p rule: rotatePolicy.induct) (simp_all add: C.simps map_add_def) - + by (induct p rule: rotatePolicy.induct) (simp_all add: C.simps map_add_def) lemma DAinRotate: "DenyAll \ set (policy2list p) \ DenyAll \ set (policy2list (rotatePolicy p))" -by (induct p,simp_all) (case_tac "DenyAll \ set (policy2list p1)",simp_all) - + by (induct p,simp_all) (case_tac "DenyAll \ set (policy2list p1)",simp_all) lemma DAUniv: "dom (CRotate (P \ DenyAll)) = UNIV" -by (metis CRotate.simps(1) CRotate.simps(4) CRotate_eq_rotateC DAAux PLemmas(4) UNIV_eq_I domSubset3) - + 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) \ [] \ C(list2FWpolicy(normalize (rotatePolicy x))) = CRotate x \ list2policy (map C (rev (normalize (rotatePolicy x)))) = CRotate x" -by (simp add: p_eq) - + by (simp add: p_eq) lemma C_eq_normalizeRotate: "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list (rotatePolicy p)) \ @@ -1953,17 +1811,13 @@ lemma C_eq_normalizeRotate: (insertDeny(removeShadowRules1(policy2list(rotatePolicy p))))))) (Nets_List (rotatePolicy p))))))) = CRotate p" -by (simp add: CRotate_eq_rotateC C_eq_compile DAinRotate) - + by (simp add: CRotate_eq_rotateC C_eq_compile DAinRotate) lemma C_eq_normalizeRotate2: "DenyAll \ set (policy2list p) \ allNetsDistinct (policy2list (rotatePolicy p)) \ all_in_list (policy2list (rotatePolicy p)) (Nets_List (rotatePolicy p)) \ C (list2FWpolicy (FWNormalisationCore.normalize (rotatePolicy p))) = CRotate p" -by (simp add: normalize_def, erule C_eq_normalizeRotate,simp_all) - - -end - - + by (simp add: normalize_def, erule C_eq_normalizeRotate,simp_all) + +end \ No newline at end of file