Import of official AFP entry for Isabelle 2018.

This commit is contained in:
Achim D. Brucker 2018-12-22 13:12:43 +00:00
parent 28e090fdde
commit 60415c8f6f
16 changed files with 28 additions and 28 deletions

View File

@ -222,4 +222,4 @@ definition policy_10 where
lemmas UnfoldPolicy10 = UnfoldPolicy9 nat_10_def Adr110_def Adr210_def policy_10_def lemmas UnfoldPolicy10 = UnfoldPolicy9 nat_10_def Adr110_def Adr210_def policy_10_def
end end

View File

@ -94,4 +94,4 @@ definition
lemmas policyLemmas = strictPolicy_def PortPolicy_def PC_def Internet_def PortPolicyBig_def src_def lemmas policyLemmas = strictPolicy_def PortPolicy_def PC_def Internet_def PortPolicyBig_def src_def
PolicyCombinators PortCombinators in_subnet_def PolicyCombinators PortCombinators in_subnet_def
end end

View File

@ -95,4 +95,4 @@ declare Ports [simp add]
definition wellformed_packet::"(adr\<^sub>i\<^sub>p,DummyContent) packet \<Rightarrow> bool" where definition wellformed_packet::"(adr\<^sub>i\<^sub>p,DummyContent) packet \<Rightarrow> bool" where
"wellformed_packet p = (content p = data)" "wellformed_packet p = (content p = data)"
end end

View File

@ -43,4 +43,4 @@ theory
Transformation02 Transformation02
begin begin
end end

View File

@ -192,4 +192,4 @@ lemma all_in_list: "all_in_list (policy2list Policy) (Nets_List Policy)"
lemmas normalizeUnfold = normalize_def Policy_def Nets_List_def bothNets_def aux aux2 bothNets_def lemmas normalizeUnfold = normalize_def Policy_def Nets_List_def bothNets_def aux aux2 bothNets_def
end end

View File

@ -177,4 +177,4 @@ lemma all_in_list: "all_in_list (policy2list Policy) (Nets_List Policy)"
lemmas normalizeUnfold = normalize_def PolicyL Nets_List_def bothNets_def aux aux2 bothNets_def sets_distinct1 sets_distinct2 sets_distinct3 sets_distinct4 sets_distinct5 sets_distinct6 aux5 aux2 lemmas normalizeUnfold = normalize_def PolicyL Nets_List_def bothNets_def aux aux2 bothNets_def sets_distinct1 sets_distinct2 sets_distinct3 sets_distinct4 sets_distinct5 sets_distinct6 aux5 aux2
end end

View File

@ -652,13 +652,13 @@ lemma in_set_in_list[rule_format]: "a \<in> set p \<longrightarrow> all_in_list
lemma sorted_Consb[rule_format]: lemma sorted_Consb[rule_format]:
"all_in_list (x#xs) l \<longrightarrow> singleCombinators (x#xs) \<longrightarrow> "all_in_list (x#xs) l \<longrightarrow> singleCombinators (x#xs) \<longrightarrow>
(sorted xs l & (ALL y:set xs. smaller x y l)) \<longrightarrow> (sorted (x#xs) l) " (sorted xs l & (\<forall>y\<in>set xs. smaller x y l)) \<longrightarrow> (sorted (x#xs) l) "
apply(induct xs arbitrary: x) apply(induct xs arbitrary: x)
apply (auto simp: order_trans) apply (auto simp: order_trans)
done done
lemma sorted_Cons: "\<lbrakk>all_in_list (x#xs) l; singleCombinators (x#xs)\<rbrakk> \<Longrightarrow> lemma sorted_Cons: "\<lbrakk>all_in_list (x#xs) l; singleCombinators (x#xs)\<rbrakk> \<Longrightarrow>
(sorted xs l & (ALL y:set xs. smaller x y l)) = (sorted (x#xs) l)" (sorted xs l & (\<forall>y\<in>set xs. smaller x y l)) = (sorted (x#xs) l)"
apply auto apply auto
apply (rule sorted_Consb, simp_all) apply (rule sorted_Consb, simp_all)
apply (metis singleCombinatorsConc singleCombinatorsStart sortedConcEnd) apply (metis singleCombinatorsConc singleCombinatorsStart sortedConcEnd)
@ -702,7 +702,7 @@ lemma all_in_list_insort: "\<lbrakk>all_in_list xs l; singleCombinators (x#xs);
done done
lemma sorted_ConsA:"\<lbrakk>all_in_list (x#xs) l; singleCombinators (x#xs)\<rbrakk> \<Longrightarrow> lemma sorted_ConsA:"\<lbrakk>all_in_list (x#xs) l; singleCombinators (x#xs)\<rbrakk> \<Longrightarrow>
(sorted (x#xs) l) = (sorted xs l & (ALL y:set xs. smaller x y l))" (sorted (x#xs) l) = (sorted xs l & (\<forall>y\<in>set xs. smaller x y l))"
by (metis sorted_Cons) by (metis sorted_Cons)
lemma is_in_insort: "y \<in> set xs \<Longrightarrow> y \<in> set (insort x xs l)" lemma is_in_insort: "y \<in> set xs \<Longrightarrow> y \<in> set (insort x xs l)"
@ -2331,4 +2331,4 @@ lemma sets_distinct5: "(n::int) < m \<Longrightarrow> {(a,b). a = n} \<noteq> {(
lemma sets_distinct6: "(m::int) < n \<Longrightarrow> {(a,b). a = n} \<noteq> {(a,b). a = m}" lemma sets_distinct6: "(m::int) < n \<Longrightarrow> {(a,b). a = n} \<noteq> {(a,b). a = m}"
by auto by auto
end end

View File

@ -563,7 +563,7 @@ lemma nMTSort: "none_MT_rules Cp p \<Longrightarrow> none_MT_rules Cp (sort p l)
lemma nMTSortQ: "none_MT_rules Cp p \<Longrightarrow> none_MT_rules Cp (qsort p l)" lemma nMTSortQ: "none_MT_rules Cp p \<Longrightarrow> 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 \<and> Cp (AllowPortFromTo a b po) = empty \<and> lemma wp3char[rule_format]: "none_MT_rules Cp xs \<and> Cp (AllowPortFromTo a b po) = Map.empty \<and>
wellformed_policy3Pr (xs @ [DenyAllFromTo a b]) \<longrightarrow> wellformed_policy3Pr (xs @ [DenyAllFromTo a b]) \<longrightarrow>
AllowPortFromTo a b po \<notin> set xs" AllowPortFromTo a b po \<notin> set xs"
by (induct xs, simp_all) (metis domNMT wp3Conc) by (induct xs, simp_all) (metis domNMT wp3Conc)
@ -881,7 +881,7 @@ lemma DARS3[rule_format]:"DenyAll \<notin> set p\<longrightarrow>DenyAll \<notin
lemma DAnMT: "dom (Cp DenyAll) \<noteq> {}" lemma DAnMT: "dom (Cp DenyAll) \<noteq> {}"
by (simp add: dom_def Cp.simps PolicyCombinators.PolicyCombinators) by (simp add: dom_def Cp.simps PolicyCombinators.PolicyCombinators)
lemma DAnMT2: "Cp DenyAll \<noteq> empty" lemma DAnMT2: "Cp DenyAll \<noteq> Map.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]: lemma wp1n_RS3[rule_format,simp]:
@ -1782,13 +1782,13 @@ lemma list2listNMT[rule_format]: "x \<noteq> [] \<longrightarrow>map sem x \<not
lemma Norm_Distr2: lemma Norm_Distr2:
"r o_f ((P \<Otimes>\<^sub>2 (list2policy Q)) o d) = "r o_f ((P \<Otimes>\<^sub>2 (list2policy Q)) o d) =
(list2policy ((P \<Otimes>\<^sub>L Q) (op \<Otimes>\<^sub>2) r d))" (list2policy ((P \<Otimes>\<^sub>L Q) (\<Otimes>\<^sub>2) r d))"
by (rule ext, rule Norm_Distr_2) by (rule ext, rule Norm_Distr_2)
lemma NATDistr: lemma NATDistr:
"N \<noteq> [] \<Longrightarrow> F = Cp (list2policyR N) \<Longrightarrow> "N \<noteq> [] \<Longrightarrow> F = Cp (list2policyR N) \<Longrightarrow>
((\<lambda> (x,y). x) o_f ((NAT \<Otimes>\<^sub>2 F) o (\<lambda> x. (x,x))) = ((\<lambda> (x,y). x) o_f ((NAT \<Otimes>\<^sub>2 F) o (\<lambda> x. (x,x))) =
(list2policy ( ((NAT \<Otimes>\<^sub>L (map Cp N)) (op \<Otimes>\<^sub>2) (list2policy ( ((NAT \<Otimes>\<^sub>L (map Cp N)) (\<Otimes>\<^sub>2)
(\<lambda> (x,y). x) (\<lambda> x. (x,x))))))" (\<lambda> (x,y). x) (\<lambda> 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)
@ -1833,10 +1833,10 @@ lemma normalizePrNAT:
allNetsDistinct (policy2list Filter) \<Longrightarrow> allNetsDistinct (policy2list Filter) \<Longrightarrow>
all_in_list (policy2list Filter) (Nets_List Filter) \<Longrightarrow> all_in_list (policy2list Filter) (Nets_List Filter) \<Longrightarrow>
((\<lambda> (x,y). x) o_f (((NAT \<Otimes>\<^sub>2 Cp Filter) o (\<lambda>x. (x,x))))) = ((\<lambda> (x,y). x) o_f (((NAT \<Otimes>\<^sub>2 Cp Filter) o (\<lambda>x. (x,x))))) =
list2policy ((NAT \<Otimes>\<^sub>L (map Cp (rev (normalizePr Filter)))) (op \<Otimes>\<^sub>2) (\<lambda> (x,y). x) (\<lambda> x. (x,x)))" list2policy ((NAT \<Otimes>\<^sub>L (map Cp (rev (normalizePr Filter)))) (\<Otimes>\<^sub>2) (\<lambda> (x,y). x) (\<lambda> 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 \<oplus> DenyAll)) = dom (Cp (DenyAll))" lemma domSimpl[simp]: "dom (Cp (A \<oplus> DenyAll)) = dom (Cp (DenyAll))"
by (simp add: PLemmas) by (simp add: PLemmas)
end end

View File

@ -870,7 +870,7 @@ lemma DARS3[rule_format]:
lemma DAnMT: "dom (C DenyAll) \<noteq> {}" lemma DAnMT: "dom (C DenyAll) \<noteq> {}"
by (simp add: dom_def C.simps PolicyCombinators.PolicyCombinators) by (simp add: dom_def C.simps PolicyCombinators.PolicyCombinators)
lemma DAnMT2: "C DenyAll \<noteq> empty" lemma DAnMT2: "C DenyAll \<noteq> Map.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]: lemma wp1n_RS3[rule_format,simp]:
@ -1760,13 +1760,13 @@ lemma list2listNMT[rule_format]: "x \<noteq> [] \<longrightarrow>map sem x \<not
by (case_tac x) simp_all by (case_tac x) simp_all
lemma Norm_Distr2: lemma Norm_Distr2:
"r o_f ((P \<Otimes>\<^sub>2 (list2policy Q)) o d) = (list2policy ((P \<Otimes>\<^sub>L Q) (op \<Otimes>\<^sub>2) r d))" "r o_f ((P \<Otimes>\<^sub>2 (list2policy Q)) o d) = (list2policy ((P \<Otimes>\<^sub>L Q) (\<Otimes>\<^sub>2) r d))"
by (rule ext, rule Norm_Distr_2) by (rule ext, rule Norm_Distr_2)
lemma NATDistr: lemma NATDistr:
"N \<noteq> [] \<Longrightarrow> F = C (list2policyR N) \<Longrightarrow> "N \<noteq> [] \<Longrightarrow> F = C (list2policyR N) \<Longrightarrow>
(\<lambda>(x, y). x) o\<^sub>f (NAT \<Otimes>\<^sub>2 F \<circ> (\<lambda>x. (x, x))) = (\<lambda>(x, y). x) o\<^sub>f (NAT \<Otimes>\<^sub>2 F \<circ> (\<lambda>x. (x, x))) =
list2policy ((NAT \<Otimes>\<^sub>L map C N) op \<Otimes>\<^sub>2 (\<lambda>(x, y). x) (\<lambda>x. (x, x)))" list2policy ((NAT \<Otimes>\<^sub>L map C N) (\<Otimes>\<^sub>2) (\<lambda>(x, y). x) (\<lambda>x. (x, x)))"
apply (simp add: l2polR_eq) apply (simp add: l2polR_eq)
apply (rule ext) apply (rule ext)
apply (rule Norm_Distr_2) apply (rule Norm_Distr_2)
@ -1806,7 +1806,7 @@ lemma normalizeNAT:
"DenyAll \<in> set (policy2list Filter) \<Longrightarrow> allNetsDistinct (policy2list Filter) \<Longrightarrow> "DenyAll \<in> set (policy2list Filter) \<Longrightarrow> allNetsDistinct (policy2list Filter) \<Longrightarrow>
all_in_list (policy2list Filter) (Nets_List Filter) \<Longrightarrow> all_in_list (policy2list Filter) (Nets_List Filter) \<Longrightarrow>
(\<lambda>(x, y). x) o\<^sub>f (NAT \<Otimes>\<^sub>2 C Filter \<circ> (\<lambda>x. (x, x))) = (\<lambda>(x, y). x) o\<^sub>f (NAT \<Otimes>\<^sub>2 C Filter \<circ> (\<lambda>x. (x, x))) =
list2policy ((NAT \<Otimes>\<^sub>L map C (rev (FWNormalisationCore.normalize Filter))) op \<Otimes>\<^sub>2 list2policy ((NAT \<Otimes>\<^sub>L map C (rev (FWNormalisationCore.normalize Filter))) (\<Otimes>\<^sub>2)
(\<lambda>(x, y). x) (\<lambda>x. (x, x)))" (\<lambda>(x, y). x) (\<lambda>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)
@ -1858,4 +1858,4 @@ lemma C_eq_normalizeRotate2:
C (list2FWpolicy (FWNormalisationCore.normalize (rotatePolicy p))) = CRotate p" C (list2FWpolicy (FWNormalisationCore.normalize (rotatePolicy p))) = CRotate p"
by (simp add: normalize_def, erule C_eq_normalizeRotate,simp_all) by (simp add: normalize_def, erule C_eq_normalizeRotate,simp_all)
end end

View File

@ -166,4 +166,4 @@ lemmas NATLemmas = src2pool_def src2poolPort_def
srcNat2pool_IntProtocolPort_def srcNat2pool_IntPort_def srcNat2pool_IntProtocolPort_def srcNat2pool_IntPort_def
srcPat2poolPort_t_def srcPat2poolPort_Protocol_t_def srcPat2poolPort_t_def srcPat2poolPort_Protocol_t_def
srcPat2pool_IntPort_def srcPat2pool_IntProtocol_def srcPat2pool_IntPort_def srcPat2pool_IntProtocol_def
end end

View File

@ -82,4 +82,4 @@ lemmas PolicyCombinators = allow_all_from_def deny_all_from_def
allow_all_to_def deny_all_to_def allow_all_from_to_def allow_all_to_def deny_all_to_def allow_all_from_to_def
deny_all_from_to_def UPFDefs deny_all_from_to_def UPFDefs
end end

View File

@ -176,4 +176,4 @@ lemmas PortCombinatorsCore =
lemmas PortCombinators = PortCombinatorsCore PolicyCombinators lemmas PortCombinators = PortCombinatorsCore PolicyCombinators
end end

View File

@ -173,4 +173,4 @@ lemmas ProtocolCombinatorsCore =
lemmas ProtocolCombinators = PortCombinators.PortCombinators ProtocolCombinatorsCore lemmas ProtocolCombinators = PortCombinators.PortCombinators ProtocolCombinatorsCore
end end

View File

@ -2,7 +2,7 @@ chapter AFP
session "UPF_Firewall" (AFP) = UPF + session "UPF_Firewall" (AFP) = UPF +
description {* Formal Network Models and Their Application to Firewall Policies *} description {* Formal Network Models and Their Application to Firewall Policies *}
options [timeout=600] options [timeout = 600]
theories theories
"Examples/Examples" "Examples/Examples"
document_files document_files

View File

@ -235,4 +235,4 @@ lemmas FTPLemmas = TRPolicy_def applyPolicy_def policy2MON_def
exI subnetOf_lemma subnetOf_lemma2 subnetOf_lemma3 subnetOf_lemma4 exI subnetOf_lemma subnetOf_lemma2 subnetOf_lemma3 subnetOf_lemma4
NetworkCore.id_def adr\<^sub>i\<^sub>pLemmas port_open_lemma NetworkCore.id_def adr\<^sub>i\<^sub>pLemmas port_open_lemma
bind_SE_def unit_SE_def valid_SE_def bind_SE_def unit_SE_def valid_SE_def
end end

View File

@ -72,4 +72,4 @@ fun FTP_STP ::
|"FTP_STP p x = (if p_accept p (snd x) |"FTP_STP p x = (if p_accept p (snd x)
then Some (allow (POL (snd x)),((fst x),snd x)) then Some (allow (POL (snd x)),((fst x),snd x))
else Some (deny (POL (snd x)),(fst x,snd x)))" else Some (deny (POL (snd x)),(fst x,snd x)))"
end end