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

@ -652,13 +652,13 @@ lemma in_set_in_list[rule_format]: "a \<in> set p \<longrightarrow> all_in_list
lemma sorted_Consb[rule_format]:
"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 (auto simp: order_trans)
done
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 (rule sorted_Consb, simp_all)
apply (metis singleCombinatorsConc singleCombinatorsStart sortedConcEnd)
@ -702,7 +702,7 @@ lemma all_in_list_insort: "\<lbrakk>all_in_list xs l; singleCombinators (x#xs);
done
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)
lemma is_in_insort: "y \<in> set xs \<Longrightarrow> y \<in> set (insort x xs l)"

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)"
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>
AllowPortFromTo a b po \<notin> set xs"
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> {}"
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)
lemma wp1n_RS3[rule_format,simp]:
@ -1782,13 +1782,13 @@ lemma list2listNMT[rule_format]: "x \<noteq> [] \<longrightarrow>map sem x \<not
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))"
(list2policy ((P \<Otimes>\<^sub>L Q) (\<Otimes>\<^sub>2) r d))"
by (rule ext, rule Norm_Distr_2)
lemma NATDistr:
"N \<noteq> [] \<Longrightarrow> F = Cp (list2policyR N) \<Longrightarrow>
((\<lambda> (x,y). x) o_f ((NAT \<Otimes>\<^sub>2 F) o (\<lambda> x. (x,x))) =
(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))))))"
by (simp add: l2polR_eq) (rule ext,rule Norm_Distr_2)
@ -1833,7 +1833,7 @@ lemma normalizePrNAT:
allNetsDistinct (policy2list 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))))) =
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)
lemma domSimpl[simp]: "dom (Cp (A \<oplus> DenyAll)) = dom (Cp (DenyAll))"

View File

@ -870,7 +870,7 @@ lemma DARS3[rule_format]:
lemma DAnMT: "dom (C DenyAll) \<noteq> {}"
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)
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
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)
lemma NATDistr:
"N \<noteq> [] \<Longrightarrow> F = C (list2policyR N) \<Longrightarrow>
(\<lambda>(x, y). x) o\<^sub>f (NAT \<Otimes>\<^sub>2 F \<circ> (\<lambda>x. (x, x))) =
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 (rule ext)
apply (rule Norm_Distr_2)
@ -1806,7 +1806,7 @@ lemma normalizeNAT:
"DenyAll \<in> set (policy2list Filter) \<Longrightarrow> allNetsDistinct (policy2list 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))) =
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)))"
by (simp add: C_eq_normalize NATDistr list2FWpolicys_eq_sym norm_notMT)

View File

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