Theory Stateful_Protocol_Verification
section‹Stateful Protocol Verification›
theory Stateful_Protocol_Verification
imports Stateful_Protocol_Model Term_Implication
begin
subsection ‹Fixed-Point Intruder Deduction Lemma›
context stateful_protocol_model
begin
abbreviation pubval_terms::"('fun,'atom,'sets,'lbl) prot_terms" where
"pubval_terms ≡ {t. ∃f ∈ funs_term t. is_PubConstValue f}"
abbreviation abs_terms::"('fun,'atom,'sets,'lbl) prot_terms" where
"abs_terms ≡ {t. ∃f ∈ funs_term t. is_Abs f}"
definition intruder_deduct_GSMP::
"[('fun,'atom,'sets,'lbl) prot_terms,
('fun,'atom,'sets,'lbl) prot_terms,
('fun,'atom,'sets,'lbl) prot_term]
⇒ bool" (‹⟨_;_⟩ ⊢⇩G⇩S⇩M⇩P _› 50)
where
"⟨M; T⟩ ⊢⇩G⇩S⇩M⇩P t ≡ intruder_deduct_restricted M (λt. t ∈ GSMP T - (pubval_terms ∪ abs_terms)) t"
lemma intruder_deduct_GSMP_induct[consumes 1, case_names AxiomH ComposeH DecomposeH]:
assumes "⟨M; T⟩ ⊢⇩G⇩S⇩M⇩P t" "⋀t. t ∈ M ⟹ P M t"
"⋀S f. ⟦length S = arity f; public f;
⋀s. s ∈ set S ⟹ ⟨M; T⟩ ⊢⇩G⇩S⇩M⇩P s;
⋀s. s ∈ set S ⟹ P M s;
Fun f S ∈ GSMP T - (pubval_terms ∪ abs_terms)
⟧ ⟹ P M (Fun f S)"
"⋀t K T' t⇩i. ⟦⟨M; T⟩ ⊢⇩G⇩S⇩M⇩P t; P M t; Ana t = (K, T'); ⋀k. k ∈ set K ⟹ ⟨M; T⟩ ⊢⇩G⇩S⇩M⇩P k;
⋀k. k ∈ set K ⟹ P M k; t⇩i ∈ set T'⟧ ⟹ P M t⇩i"
shows "P M t"
proof -
let ?Q = "λt. t ∈ GSMP T - (pubval_terms ∪ abs_terms)"
show ?thesis
using intruder_deduct_restricted_induct[of M ?Q t "λM Q t. P M t"] assms
unfolding intruder_deduct_GSMP_def
by blast
qed
lemma pubval_terms_subst:
assumes "t ⋅ θ ∈ pubval_terms" "θ ` fv t ∩ pubval_terms = {}"
shows "t ∈ pubval_terms"
using assms(1,2)
proof (induction t)
case (Fun f T)
let ?P = "λf. is_PubConstValue f"
from Fun show ?case
proof (cases "?P f")
case False
then obtain t where t: "t ∈ set T" "t ⋅ θ ∈ pubval_terms"
using Fun.prems by auto
hence "θ ` fv t ∩ pubval_terms = {}" using Fun.prems(2) by auto
thus ?thesis using Fun.IH[OF t] t(1) by auto
qed force
qed simp
lemma abs_terms_subst:
assumes "t ⋅ θ ∈ abs_terms" "θ ` fv t ∩ abs_terms = {}"
shows "t ∈ abs_terms"
using assms(1,2)
proof (induction t)
case (Fun f T)
let ?P = "λf. is_Abs f"
from Fun show ?case
proof (cases "?P f")
case False
then obtain t where t: "t ∈ set T" "t ⋅ θ ∈ abs_terms"
using Fun.prems by auto
hence "θ ` fv t ∩ abs_terms = {}" using Fun.prems(2) by auto
thus ?thesis using Fun.IH[OF t] t(1) by auto
qed force
qed simp
lemma pubval_terms_subst':
assumes "t ⋅ θ ∈ pubval_terms" "∀n. PubConst Value n ∉ ⋃(funs_term ` (θ ` fv t))"
shows "t ∈ pubval_terms"
proof -
have False
when fs: "f ∈ funs_term s" "s ∈ subterms⇩s⇩e⇩t (θ ` fv t)" "is_PubConstValue f"
for f s
proof -
obtain T where T: "Fun f T ∈ subterms s" using funs_term_Fun_subterm[OF fs(1)] by force
hence "Fun f T ∈ subterms⇩s⇩e⇩t (θ ` fv t)" using fs(2) in_subterms_subset_Union by blast
thus ?thesis
using assms(2) funs_term_Fun_subterm'[of f T] fs(3)
unfolding is_PubConstValue_def
by (cases f) force+
qed
thus ?thesis using pubval_terms_subst[OF assms(1)] by auto
qed
lemma abs_terms_subst':
assumes "t ⋅ θ ∈ abs_terms" "∀n. Abs n ∉ ⋃(funs_term ` (θ ` fv t))"
shows "t ∈ abs_terms"
proof -
have "¬is_Abs f" when fs: "f ∈ funs_term s" "s ∈ subterms⇩s⇩e⇩t (θ ` fv t)" for f s
proof -
obtain T where T: "Fun f T ∈ subterms s" using funs_term_Fun_subterm[OF fs(1)] by force
hence "Fun f T ∈ subterms⇩s⇩e⇩t (θ ` fv t)" using fs(2) in_subterms_subset_Union by blast
thus ?thesis using assms(2) funs_term_Fun_subterm'[of f T] by (cases f) auto
qed
thus ?thesis using abs_terms_subst[OF assms(1)] by force
qed
lemma pubval_terms_subst_range_disj:
"subst_range θ ∩ pubval_terms = {} ⟹ θ ` fv t ∩ pubval_terms = {}"
proof (induction t)
case (Var x) thus ?case by (cases "x ∈ subst_domain θ") auto
qed auto
lemma abs_terms_subst_range_disj:
"subst_range θ ∩ abs_terms = {} ⟹ θ ` fv t ∩ abs_terms = {}"
proof (induction t)
case (Var x) thus ?case by (cases "x ∈ subst_domain θ") auto
qed auto
lemma pubval_terms_subst_range_comp:
assumes "subst_range θ ∩ pubval_terms = {}" "subst_range δ ∩ pubval_terms = {}"
shows "subst_range (θ ∘⇩s δ) ∩ pubval_terms = {}"
proof -
{ fix t f assume t:
"t ∈ subst_range (θ ∘⇩s δ)" "f ∈ funs_term t" "is_PubConstValue f"
then obtain x where x: "(θ ∘⇩s δ) x = t" by auto
have "θ x ∉ pubval_terms" using assms(1) by (cases "θ x ∈ subst_range θ") force+
hence "(θ ∘⇩s δ) x ∉ pubval_terms"
using assms(2) pubval_terms_subst[of "θ x" δ] pubval_terms_subst_range_disj
by (metis (mono_tags, lifting) subst_compose_def)
hence False using t(2,3) x by blast
} thus ?thesis by fast
qed
lemma pubval_terms_subst_range_comp':
assumes "(θ ` X) ∩ pubval_terms = {}" "(δ ` fv⇩s⇩e⇩t (θ ` X)) ∩ pubval_terms = {}"
shows "((θ ∘⇩s δ) ` X) ∩ pubval_terms = {}"
proof -
{ fix t f assume t:
"t ∈ (θ ∘⇩s δ) ` X" "f ∈ funs_term t" "is_PubConstValue f"
then obtain x where x: "(θ ∘⇩s δ) x = t" "x ∈ X" by auto
have "θ x ∉ pubval_terms" using assms(1) x(2) by force
moreover have "fv (θ x) ⊆ fv⇩s⇩e⇩t (θ ` X)" using x(2) by (auto simp add: fv_subset)
hence "δ ` fv (θ x) ∩ pubval_terms = {}" using assms(2) by auto
ultimately have "(θ ∘⇩s δ) x ∉ pubval_terms"
using pubval_terms_subst[of "θ x" δ]
by (metis (mono_tags, lifting) subst_compose_def)
hence False using t(2,3) x by blast
} thus ?thesis by fast
qed
lemma abs_terms_subst_range_comp:
assumes "subst_range θ ∩ abs_terms = {}" "subst_range δ ∩ abs_terms = {}"
shows "subst_range (θ ∘⇩s δ) ∩ abs_terms = {}"
proof -
{ fix t f assume t: "t ∈ subst_range (θ ∘⇩s δ)" "f ∈ funs_term t" "is_Abs f"
then obtain x where x: "(θ ∘⇩s δ) x = t" by auto
have "θ x ∉ abs_terms" using assms(1) by (cases "θ x ∈ subst_range θ") force+
hence "(θ ∘⇩s δ) x ∉ abs_terms"
using assms(2) abs_terms_subst[of "θ x" δ] abs_terms_subst_range_disj
by (metis (mono_tags, lifting) subst_compose_def)
hence False using t(2,3) x by blast
} thus ?thesis by fast
qed
lemma abs_terms_subst_range_comp':
assumes "(θ ` X) ∩ abs_terms = {}" "(δ ` fv⇩s⇩e⇩t (θ ` X)) ∩ abs_terms = {}"
shows "((θ ∘⇩s δ) ` X) ∩ abs_terms = {}"
proof -
{ fix t f assume t:
"t ∈ (θ ∘⇩s δ) ` X" "f ∈ funs_term t" "is_Abs f"
then obtain x where x: "(θ ∘⇩s δ) x = t" "x ∈ X" by auto
have "θ x ∉ abs_terms" using assms(1) x(2) by force
moreover have "fv (θ x) ⊆ fv⇩s⇩e⇩t (θ ` X)" using x(2) by (auto simp add: fv_subset)
hence "δ ` fv (θ x) ∩ abs_terms = {}" using assms(2) by auto
ultimately have "(θ ∘⇩s δ) x ∉ abs_terms"
using abs_terms_subst[of "θ x" δ]
by (metis (mono_tags, lifting) subst_compose_def)
hence False using t(2,3) x by blast
} thus ?thesis by fast
qed
context
begin
private lemma Ana_abs_aux1:
fixes δ::"(('fun,'atom,'sets,'lbl) prot_fun, nat, ('fun,'atom,'sets,'lbl) prot_var) gsubst"
and α::"nat ⇒ 'sets set"
assumes "Ana⇩f f = (K,T)"
shows "(K ⋅⇩l⇩i⇩s⇩t δ) ⋅⇩α⇩l⇩i⇩s⇩t α = K ⋅⇩l⇩i⇩s⇩t (λn. δ n ⋅⇩α α)"
proof -
{ fix k assume "k ∈ set K"
hence "k ∈ subterms⇩s⇩e⇩t (set K)" by force
hence "k ⋅ δ ⋅⇩α α = k ⋅ (λn. δ n ⋅⇩α α)"
proof (induction k)
case (Fun g S)
have "⋀s. s ∈ set S ⟹ s ⋅ δ ⋅⇩α α = s ⋅ (λn. δ n ⋅⇩α α)"
using Fun.IH in_subterms_subset_Union[OF Fun.prems] Fun_param_in_subterms[of _ S g]
by (meson contra_subsetD)
thus ?case using Ana⇩f_assm1_alt[OF assms Fun.prems] by (cases g) auto
qed simp
} thus ?thesis unfolding abs_apply_list_def by force
qed
private lemma Ana_abs_aux2:
fixes α::"nat ⇒ 'sets set"
and K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) term list"
and M::"nat list"
and T::"('fun,'atom,'sets,'lbl) prot_term list"
assumes "∀i ∈ fv⇩s⇩e⇩t (set K) ∪ set M. i < length T"
and "(K ⋅⇩l⇩i⇩s⇩t (!) T) ⋅⇩α⇩l⇩i⇩s⇩t α = K ⋅⇩l⇩i⇩s⇩t (λn. T ! n ⋅⇩α α)"
shows "(K ⋅⇩l⇩i⇩s⇩t (!) T) ⋅⇩α⇩l⇩i⇩s⇩t α = K ⋅⇩l⇩i⇩s⇩t (!) (map (λs. s ⋅⇩α α) T)" (is "?A1 = ?A2")
and "(map ((!) T) M) ⋅⇩α⇩l⇩i⇩s⇩t α = map ((!) (map (λs. s ⋅⇩α α) T)) M" (is "?B1 = ?B2")
proof -
have "T ! i ⋅⇩α α = (map (λs. s ⋅⇩α α) T) ! i" when "i ∈ fv⇩s⇩e⇩t (set K)" for i
using that assms(1) by auto
hence "k ⋅ (λi. T ! i ⋅⇩α α) = k ⋅ (λi. (map (λs. s ⋅⇩α α) T) ! i)" when "k ∈ set K" for k
using that term_subst_eq_conv[of k "λi. T ! i ⋅⇩α α" "λi. (map (λs. s ⋅⇩α α) T) ! i"]
by auto
thus "?A1 = ?A2" using assms(2) by (force simp add: abs_apply_terms_def)
have "T ! i ⋅⇩α α = map (λs. s ⋅⇩α α) T ! i" when "i ∈ set M" for i
using that assms(1) by auto
thus "?B1 = ?B2" by (force simp add: abs_apply_list_def)
qed
private lemma Ana_abs_aux1_set:
fixes δ::"(('fun,'atom,'sets,'lbl) prot_fun, nat, ('fun,'atom,'sets,'lbl) prot_var) gsubst"
and α::"nat ⇒ 'sets set"
assumes "Ana⇩f f = (K,T)"
shows "(set K ⋅⇩s⇩e⇩t δ) ⋅⇩α⇩s⇩e⇩t α = set K ⋅⇩s⇩e⇩t (λn. δ n ⋅⇩α α)"
proof -
{ fix k assume "k ∈ set K"
hence "k ∈ subterms⇩s⇩e⇩t (set K)" by force
hence "k ⋅ δ ⋅⇩α α = k ⋅ (λn. δ n ⋅⇩α α)"
proof (induction k)
case (Fun g S)
have "⋀s. s ∈ set S ⟹ s ⋅ δ ⋅⇩α α = s ⋅ (λn. δ n ⋅⇩α α)"
using Fun.IH in_subterms_subset_Union[OF Fun.prems] Fun_param_in_subterms[of _ S g]
by (meson contra_subsetD)
thus ?case using Ana⇩f_assm1_alt[OF assms Fun.prems] by (cases g) auto
qed simp
} thus ?thesis unfolding abs_apply_terms_def by force
qed
private lemma Ana_abs_aux2_set:
fixes α::"nat ⇒ 'sets set"
and K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) terms"
and M::"nat set"
and T::"('fun,'atom,'sets,'lbl) prot_term list"
assumes "∀i ∈ fv⇩s⇩e⇩t K ∪ M. i < length T"
and "(K ⋅⇩s⇩e⇩t (!) T) ⋅⇩α⇩s⇩e⇩t α = K ⋅⇩s⇩e⇩t (λn. T ! n ⋅⇩α α)"
shows "(K ⋅⇩s⇩e⇩t (!) T) ⋅⇩α⇩s⇩e⇩t α = K ⋅⇩s⇩e⇩t (!) (map (λs. s ⋅⇩α α) T)" (is "?A1 = ?A2")
and "((!) T ` M) ⋅⇩α⇩s⇩e⇩t α = (!) (map (λs. s ⋅⇩α α) T) ` M" (is "?B1 = ?B2")
proof -
have "T ! i ⋅⇩α α = (map (λs. s ⋅⇩α α) T) ! i" when "i ∈ fv⇩s⇩e⇩t K" for i
using that assms(1) by auto
hence "k ⋅ (λi. T ! i ⋅⇩α α) = k ⋅ (λi. (map (λs. s ⋅⇩α α) T) ! i)" when "k ∈ K" for k
using that term_subst_eq_conv[of k "λi. T ! i ⋅⇩α α" "λi. (map (λs. s ⋅⇩α α) T) ! i"]
by auto
thus "?A1 = ?A2" using assms(2) by (force simp add: abs_apply_terms_def)
have "T ! i ⋅⇩α α = map (λs. s ⋅⇩α α) T ! i" when "i ∈ M" for i
using that assms(1) by auto
thus "?B1 = ?B2" by (force simp add: abs_apply_terms_def)
qed
lemma Ana_abs:
fixes t::"('fun,'atom,'sets,'lbl) prot_term"
assumes "Ana t = (K, T)"
shows "Ana (t ⋅⇩α α) = (K ⋅⇩α⇩l⇩i⇩s⇩t α, T ⋅⇩α⇩l⇩i⇩s⇩t α)"
using assms
proof (induction t rule: Ana.induct)
case (1 f S)
obtain K' T' where *: "Ana⇩f f = (K',T')" by force
show ?case using 1
proof (cases "arity⇩f f = length S ∧ arity⇩f f > 0")
case True
hence "K = K' ⋅⇩l⇩i⇩s⇩t (!) S" "T = map ((!) S) T'"
and **: "arity⇩f f = length (map (λs. s ⋅⇩α α) S)" "arity⇩f f > 0"
using 1 * by auto
hence "K ⋅⇩α⇩l⇩i⇩s⇩t α = K' ⋅⇩l⇩i⇩s⇩t (!) (map (λs. s ⋅⇩α α) S)"
"T ⋅⇩α⇩l⇩i⇩s⇩t α = map ((!) (map (λs. s ⋅⇩α α) S)) T'"
using Ana⇩f_assm2_alt[OF *] Ana_abs_aux2[OF _ Ana_abs_aux1[OF *], of T' S α]
unfolding abs_apply_list_def
by auto
moreover have "Fun (Fu f) S ⋅⇩α α = Fun (Fu f) (map (λs. s ⋅⇩α α) S)" by simp
ultimately show ?thesis using Ana_Fu_intro[OF ** *] by metis
qed (auto simp add: abs_apply_list_def)
qed (simp_all add: abs_apply_list_def)
end
lemma deduct_FP_if_deduct:
fixes M IK FP::"('fun,'atom,'sets,'lbl) prot_terms"
assumes IK: "IK ⊆ GSMP M - (pubval_terms ∪ abs_terms)" "∀t ∈ IK ⋅⇩α⇩s⇩e⇩t α. FP ⊢⇩c t"
and t: "IK ⊢ t" "t ∈ GSMP M - (pubval_terms ∪ abs_terms)"
shows "FP ⊢ t ⋅⇩α α"
proof -
let ?P = "λf. ¬is_PubConstValue f"
let ?GSMP = "GSMP M - (pubval_terms ∪ abs_terms)"
have 1: "∀m ∈ IK. m ∈ ?GSMP"
using IK(1) by blast
have 2: "∀t t'. t ∈ ?GSMP ⟶ t' ⊑ t ⟶ t' ∈ ?GSMP"
proof (intro allI impI)
fix t t' assume t: "t ∈ ?GSMP" "t' ⊑ t"
hence "t' ∈ GSMP M" using ground_subterm unfolding GSMP_def by auto
moreover have "¬is_PubConstValue f"
when "f ∈ funs_term t" for f
using t(1) that by auto
hence "¬is_PubConstValue f"
when "f ∈ funs_term t'" for f
using that subtermeq_imp_funs_term_subset[OF t(2)] by auto
moreover have "¬is_Abs f" when "f ∈ funs_term t" for f using t(1) that by auto
hence "¬is_Abs f" when "f ∈ funs_term t'" for f
using that subtermeq_imp_funs_term_subset[OF t(2)] by auto
ultimately show "t' ∈ ?GSMP" by simp
qed
have 3: "∀t K T k. t ∈ ?GSMP ⟶ Ana t = (K, T) ⟶ k ∈ set K ⟶ k ∈ ?GSMP"
proof (intro allI impI)
fix t K T k assume t: "t ∈ ?GSMP" "Ana t = (K, T)" "k ∈ set K"
hence "k ∈ GSMP M" using GSMP_Ana_key by blast
moreover have "∀f ∈ funs_term t. ?P f" using t(1) by auto
with t(2,3) have "∀f ∈ funs_term k. ?P f"
proof (induction t arbitrary: k rule: Ana.induct)
case 1 thus ?case by (metis Ana_Fu_keys_not_pubval_terms surj_pair)
qed auto
moreover have "∀f ∈ funs_term t. ¬is_Abs f" using t(1) by auto
with t(2,3) have "∀f ∈ funs_term k. ¬is_Abs f"
proof (induction t arbitrary: k rule: Ana.induct)
case 1 thus ?case by (metis Ana_Fu_keys_not_abs_terms surj_pair)
qed auto
ultimately show "k ∈ ?GSMP" by simp
qed
have "⟨IK; M⟩ ⊢⇩G⇩S⇩M⇩P t"
unfolding intruder_deduct_GSMP_def
by (rule restricted_deduct_if_deduct'[OF 1 2 3 t])
thus ?thesis
proof (induction t rule: intruder_deduct_GSMP_induct)
case (AxiomH t)
show ?case using IK(2) abs_in[OF AxiomH.hyps] by force
next
case (ComposeH T f)
have *: "Fun f T ⋅⇩α α = Fun f (map (λt. t ⋅⇩α α) T)"
using ComposeH.hyps(2,4)
by (cases f) auto
have **: "length (map (λt. t ⋅⇩α α) T) = arity f"
using ComposeH.hyps(1)
by auto
show ?case
using intruder_deduct.Compose[OF ** ComposeH.hyps(2)] ComposeH.IH(1) *
by auto
next
case (DecomposeH t K T' t⇩i)
have *: "Ana (t ⋅⇩α α) = (K ⋅⇩α⇩l⇩i⇩s⇩t α, T' ⋅⇩α⇩l⇩i⇩s⇩t α)"
using Ana_abs[OF DecomposeH.hyps(2)]
by metis
have **: "t⇩i ⋅⇩α α ∈ set (T' ⋅⇩α⇩l⇩i⇩s⇩t α)"
using DecomposeH.hyps(4) abs_in abs_list_set_is_set_abs_set[of T']
by auto
have ***: "FP ⊢ k"
when k: "k ∈ set (K ⋅⇩α⇩l⇩i⇩s⇩t α)" for k
proof -
obtain k' where k': "k' ∈ set K" "k = k' ⋅⇩α α"
by (metis (no_types) k abs_apply_terms_def imageE abs_list_set_is_set_abs_set)
show "FP ⊢ k"
using DecomposeH.IH k' by blast
qed
show ?case
using intruder_deduct.Decompose[OF _ * _ **]
DecomposeH.IH(1) ***(1)
by blast
qed
qed
end
subsection ‹Computing and Checking Term Implications and Messages›
context stateful_protocol_model
begin
abbreviation (input) "absc s ≡ (Fun (Abs s) []::('fun,'atom,'sets,'lbl) prot_term)"
fun absdbupd where
"absdbupd [] _ a = a"
| "absdbupd (insert⟨Var y, Fun (Set s) T⟩#D) x a = (
if x = y then absdbupd D x (insert s a) else absdbupd D x a)"
| "absdbupd (delete⟨Var y, Fun (Set s) T⟩#D) x a = (
if x = y then absdbupd D x (a - {s}) else absdbupd D x a)"
| "absdbupd (_#D) x a = absdbupd D x a"
lemma absdbupd_cons_cases:
"absdbupd (insert⟨Var x, Fun (Set s) T⟩#D) x d = absdbupd D x (insert s d)"
"absdbupd (delete⟨Var x, Fun (Set s) T⟩#D) x d = absdbupd D x (d - {s})"
"t ≠ Var x ∨ (∄s T. u = Fun (Set s) T) ⟹ absdbupd (insert⟨t,u⟩#D) x d = absdbupd D x d"
"t ≠ Var x ∨ (∄s T. u = Fun (Set s) T) ⟹ absdbupd (delete⟨t,u⟩#D) x d = absdbupd D x d"
proof -
assume *: "t ≠ Var x ∨ (∄s T. u = Fun (Set s) T)"
let ?P = "absdbupd (insert⟨t,u⟩#D) x d = absdbupd D x d"
let ?Q = "absdbupd (delete⟨t,u⟩#D) x d = absdbupd D x d"
{ fix y f T assume "t = Fun f T ∨ u = Var y" hence ?P ?Q by auto
} moreover {
fix y f T assume "t = Var y" "u = Fun f T" hence ?P using * by (cases f) auto
} moreover {
fix y f T assume "t = Var y" "u = Fun f T" hence ?Q using * by (cases f) auto
} ultimately show ?P ?Q by (metis term.exhaust)+
qed simp_all
lemma absdbupd_filter: "absdbupd S x d = absdbupd (filter is_Update S) x d"
by (induction S x d rule: absdbupd.induct) simp_all
lemma absdbupd_append:
"absdbupd (A@B) x d = absdbupd B x (absdbupd A x d)"
proof (induction A arbitrary: d)
case (Cons a A) thus ?case
proof (cases a)
case (Insert t u) thus ?thesis
proof (cases "t ≠ Var x ∨ (∄s T. u = Fun (Set s) T)")
case False
then obtain s T where "t = Var x" "u = Fun (Set s) T" by force
thus ?thesis by (simp add: Insert Cons.IH absdbupd_cons_cases(1))
qed (simp_all add: Cons.IH absdbupd_cons_cases(3))
next
case (Delete t u) thus ?thesis
proof (cases "t ≠ Var x ∨ (∄s T. u = Fun (Set s) T)")
case False
then obtain s T where "t = Var x" "u = Fun (Set s) T" by force
thus ?thesis by (simp add: Delete Cons.IH absdbupd_cons_cases(2))
qed (simp_all add: Cons.IH absdbupd_cons_cases(4))
qed simp_all
qed simp
lemma absdbupd_wellformed_transaction:
assumes T: "wellformed_transaction T"
shows "absdbupd (unlabel (transaction_strand T)) = absdbupd (unlabel (transaction_updates T))"
proof -
define S0 where "S0 ≡ unlabel (transaction_strand T)"
define S1 where "S1 ≡ unlabel (transaction_receive T)"
define S2 where "S2 ≡ unlabel (transaction_checks T)"
define S3 where "S3 ≡ unlabel (transaction_updates T)"
define S4 where "S4 ≡ unlabel (transaction_send T)"
note S_defs = S0_def S1_def S2_def S3_def S4_def
have 0: "list_all is_Receive S1"
"list_all is_Check_or_Assignment S2"
"list_all is_Update S3"
"list_all is_Send S4"
using T unfolding wellformed_transaction_def S_defs by metis+
have "filter is_Update S1 = []"
"filter is_Update S2 = []"
"filter is_Update S3 = S3"
"filter is_Update S4 = []"
using list_all_filter_nil[OF 0(1), of is_Update]
list_all_filter_nil[OF 0(2), of is_Update]
list_all_filter_eq[OF 0(3)]
list_all_filter_nil[OF 0(4), of is_Update]
by blast+
moreover have "S0 = S1@S2@S3@S4"
unfolding S_defs transaction_strand_def unlabel_def by auto
ultimately have "filter is_Update S0 = S3"
using filter_append[of is_Update] list_all_append[of is_Update]
by simp
thus ?thesis
using absdbupd_filter[of S0]
unfolding S_defs by presburger
qed
fun abs_substs_set::
"[('fun,'atom,'sets,'lbl) prot_var list,
'sets set list,
('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set,
('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set,
('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set ⇒ bool]
⇒ ((('fun,'atom,'sets,'lbl) prot_var × 'sets set) list) list"
where
"abs_substs_set [] _ _ _ _ = [[]]"
| "abs_substs_set (x#xs) as posconstrs negconstrs msgconstrs = (
let bs = filter (λa. posconstrs x ⊆ a ∧ a ∩ negconstrs x = {} ∧ msgconstrs x a) as;
Δ = abs_substs_set xs as posconstrs negconstrs msgconstrs
in concat (map (λb. map (λδ. (x, b)#δ) Δ) bs))"
definition abs_substs_fun::
"[(('fun,'atom,'sets,'lbl) prot_var × 'sets set) list,
('fun,'atom,'sets,'lbl) prot_var]
⇒ 'sets set"
where
"abs_substs_fun δ x = (case find (λb. fst b = x) δ of Some (_,a) ⇒ a | None ⇒ {})"
lemmas abs_substs_set_induct = abs_substs_set.induct[case_names Nil Cons]
fun transaction_poschecks_comp::
"(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand
⇒ (('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set)"
where
"transaction_poschecks_comp [] = (λ_. {})"
| "transaction_poschecks_comp (⟨_: Var x ∈ Fun (Set s) []⟩#T) = (
let f = transaction_poschecks_comp T in f(x := insert s (f x)))"
| "transaction_poschecks_comp (_#T) = transaction_poschecks_comp T"
fun transaction_negchecks_comp::
"(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand
⇒ (('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set)"
where
"transaction_negchecks_comp [] = (λ_. {})"
| "transaction_negchecks_comp (⟨Var x not in Fun (Set s) []⟩#T) = (
let f = transaction_negchecks_comp T in f(x := insert s (f x)))"
| "transaction_negchecks_comp (_#T) = transaction_negchecks_comp T"
definition transaction_check_pre where
"transaction_check_pre FPT T δ ≡
let (FP, _, TI) = FPT;
C = set (unlabel (transaction_checks T));
xs = fv_list⇩s⇩s⇩t (unlabel (transaction_strand T));
θ = λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x
in (∀x ∈ set (transaction_fresh T). δ x = {}) ∧
(∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). intruder_synth_mod_timpls FP TI (t ⋅ θ δ)) ∧
(∀u ∈ C.
(is_InSet u ⟶ (
let x = the_elem_term u; s = the_set_term u
in (is_Var x ∧ is_Fun_Set s) ⟶ the_Set (the_Fun s) ∈ δ (the_Var x))) ∧
((is_NegChecks u ∧ bvars⇩s⇩s⇩t⇩p u = [] ∧ the_eqs u = [] ∧ length (the_ins u) = 1) ⟶ (
let x = fst (hd (the_ins u)); s = snd (hd (the_ins u))
in (is_Var x ∧ is_Fun_Set s) ⟶ the_Set (the_Fun s) ∉ δ (the_Var x))))"
definition transaction_check_post where
"transaction_check_post FPT T δ ≡
let (FP, _, TI) = FPT;
xs = fv_list⇩s⇩s⇩t (unlabel (transaction_strand T));
θ = λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x;
u = λδ x. absdbupd (unlabel (transaction_updates T)) x (δ x)
in (∀x ∈ set xs - set (transaction_fresh T). δ x ≠ u δ x ⟶ List.member TI (δ x, u δ x)) ∧
(∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T). intruder_synth_mod_timpls FP TI (t ⋅ θ (u δ)))"
definition fun_point_inter where "fun_point_inter f g ≡ λx. f x ∩ g x"
definition fun_point_union where "fun_point_union f g ≡ λx. f x ∪ g x"
definition fun_point_Inter where "fun_point_Inter fs ≡ λx. ⋂f ∈ fs. f x"
definition fun_point_Union where "fun_point_Union fs ≡ λx. ⋃f ∈ fs. f x"
definition fun_point_Inter_list where "fun_point_Inter_list fs ≡ λx. ⋂(set (map (λf. f x) fs))"
definition fun_point_Union_list where "fun_point_Union_list fs ≡ λx. ⋃(set (map (λf. f x) fs))"
definition ticl_abs where "ticl_abs TI a ≡ set (a#map snd (filter (λp. fst p = a) TI))"
definition ticl_abss where "ticl_abss TI as ≡ ⋃a ∈ as. ticl_abs TI a"
lemma fun_point_Inter_set_eq:
"fun_point_Inter (set fs) = fun_point_Inter_list fs"
unfolding fun_point_Inter_def fun_point_Inter_list_def by simp
lemma fun_point_Union_set_eq:
"fun_point_Union (set fs) = fun_point_Union_list fs"
unfolding fun_point_Union_def fun_point_Union_list_def by simp
lemma ticl_abs_refl_in: "x ∈ ticl_abs TI x"
unfolding ticl_abs_def by simp
lemma ticl_abs_iff:
assumes TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
shows "ticl_abs TI a = {b. (a,b) ∈ (set TI)⇧*}"
proof (intro order_antisym subsetI)
fix x assume x: "x ∈ {b. (a, b) ∈ (set TI)⇧*}"
hence "x = a ∨ (x ≠ a ∧ (a,x) ∈ (set TI)⇧+)" by (metis mem_Collect_eq rtranclD)
moreover have "ticl_abs TI a = {a} ∪ {b. (a,b) ∈ set TI}" unfolding ticl_abs_def by force
ultimately show "x ∈ ticl_abs TI a" using TI by blast
qed (fastforce simp add: ticl_abs_def)
lemma ticl_abs_Inter:
assumes xs: "⋂(ticl_abs TI ` xs) ≠ {}"
and TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
shows "⋂(ticl_abs TI ` ⋂(ticl_abs TI ` xs)) ⊆ ⋂(ticl_abs TI ` xs)"
proof
fix x assume x: "x ∈ ⋂(ticl_abs TI ` ⋂(ticl_abs TI ` xs))"
have *: "⋂(ticl_abs TI ` xs) = {b. ∀a ∈ xs. (a,b) ∈ (set TI)⇧*}"
unfolding ticl_abs_iff[OF TI] by blast
have "(b,x) ∈ (set TI)⇧*" when b: "∀a ∈ xs. (a,b) ∈ (set TI)⇧*" for b
using x b unfolding ticl_abs_iff[OF TI] by blast
hence "(a,x) ∈ (set TI)⇧*" when "a ∈ xs" for a
using that xs rtrancl.rtrancl_into_rtrancl[of a _ "(set TI)⇧*" x]
unfolding * rtrancl_idemp[of "set TI"] by blast
thus "x ∈ ⋂(ticl_abs TI ` xs)" unfolding * by blast
qed
function (sequential) match_abss'
::"(('a,'b,'c,'d) prot_fun, 'e) term ⇒
(('a,'b,'c,'d) prot_fun, 'e) term ⇒
('e ⇒ 'c set set) option"
where
"match_abss' (Var x) (Fun (Abs a) _) = Some ((λ_. {})(x := {a}))"
| "match_abss' (Fun f ts) (Fun g ss) = (
if f = g ∧ length ts = length ss
then map_option fun_point_Union_list (those (map2 match_abss' ts ss))
else None)"
| "match_abss' _ _ = None"
by pat_completeness auto
termination
proof -
let ?m = "measures [size ∘ fst]"
have 0: "wf ?m" by simp
show ?thesis
apply (standard, use 0 in fast)
by (metis (no_types) comp_def fst_conv measures_less Fun_zip_size_lt(1))
qed
definition match_abss where
"match_abss OCC TI t s ≡ (
let xs = fv t;
OCC' = set OCC;
f = λδ x. if x ∈ xs then δ x else OCC';
g = λδ x. ⋂(ticl_abs TI ` δ x)
in case match_abss' t s of
Some δ ⇒
let δ' = g δ
in if ∀x ∈ xs. δ' x ≠ {} then Some (f δ') else None
| None ⇒ None)"
lemma match_abss'_Var_inv:
assumes δ: "match_abss' (Var x) t = Some δ"
shows "∃a ts. t = Fun (Abs a) ts ∧ δ = (λ_. {})(x := {a})"
proof -
obtain f ts where t: "t = Fun f ts" using δ by (cases t) auto
then obtain a where a: "f = Abs a" using δ by (cases f) auto
show ?thesis using δ unfolding t a by simp
qed
lemma match_abss'_Fun_inv:
assumes "match_abss' (Fun f ts) (Fun g ss) = Some δ"
shows "f = g" (is ?A)
and "length ts = length ss" (is ?B)
and "∃θ. Some θ = those (map2 match_abss' ts ss) ∧ δ = fun_point_Union_list θ" (is ?C)
and "∀(t,s) ∈ set (zip ts ss). ∃σ. match_abss' t s = Some σ" (is ?D)
proof -
note 0 = assms match_abss'.simps(2)[of f ts g ss] option.distinct(1)
show ?A by (metis 0)
show ?B by (metis 0)
show ?C by (metis (no_types, opaque_lifting) 0 map_option_eq_Some)
thus ?D using map2_those_Some_case[of match_abss' ts ss] by fastforce
qed
lemma match_abss'_FunI:
assumes Δ: "⋀i. i < length T ⟹ match_abss' (U ! i) (T ! i) = Some (Δ i)"
and T: "length T = length U"
shows "match_abss' (Fun f U) (Fun f T) = Some (fun_point_Union_list (map Δ [0..<length T]))"
proof -
have "match_abss' (Fun f U) (Fun f T) =
map_option fun_point_Union_list (those (map2 match_abss' U T))"
using T match_abss'.simps(2)[of f U f T] by presburger
moreover have "those (map2 match_abss' U T) = Some (map Δ [0..<length T])"
using Δ T those_map2_SomeI by metis
ultimately show ?thesis by simp
qed
lemma match_abss'_Fun_param_subset:
assumes "match_abss' (Fun f ts) (Fun g ss) = Some δ"
and "(t,s) ∈ set (zip ts ss)"
and "match_abss' t s = Some σ"
shows "σ x ⊆ δ x"
proof -
obtain θ where θ:
"those (map2 match_abss' ts ss) = Some θ"
"δ = fun_point_Union_list θ"
using match_abss'_Fun_inv[OF assms(1)] by metis
have "σ ∈ set θ" using θ(1) assms(2-) those_Some_iff[of "map2 match_abss' ts ss" θ] by force
thus ?thesis using θ(2) unfolding fun_point_Union_list_def by auto
qed
lemma match_abss'_fv_is_nonempty:
assumes "match_abss' t s = Some δ"
and "x ∈ fv t"
shows "δ x ≠ {}" (is "?P δ")
using assms
proof (induction t s arbitrary: δ rule: match_abss'.induct)
case (2 f ts g ss)
note prems = "2.prems"
note IH = "2.IH"
have 0: "∀(t,s) ∈ set (zip ts ss). ∃σ. match_abss' t s = Some σ" "f = g" "length ts = length ss"
using match_abss'_Fun_inv[OF prems(1)] by simp_all
obtain t where t: "t ∈ set ts" "x ∈ fv t" using prems(2) by auto
then obtain s where s: "s ∈ set ss" "(t,s) ∈ set (zip ts ss)"
by (meson 0(3) in_set_impl_in_set_zip1 in_set_zipE)
then obtain σ where σ: "match_abss' t s = Some σ" using 0(1) by fast
show ?case
using IH[OF conjI[OF 0(2,3)] s(2) _ σ] t(2) match_abss'_Fun_param_subset[OF prems(1) s(2) σ]
by auto
qed auto
lemma match_abss'_nonempty_is_fv:
fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term"
assumes "match_abss' s t = Some δ"
and "δ x ≠ {}"
shows "x ∈ fv s"
using assms
proof (induction s t arbitrary: δ rule: match_abss'.induct)
case (2 f ts g ss)
note prems = "2.prems"
note IH = "2.IH"
obtain θ where θ: "Some θ = those (map2 match_abss' ts ss)" "δ = fun_point_Union_list θ"
and fg: "f = g" "length ts = length ss"
using match_abss'_Fun_inv[OF prems(1)] by fast
have "∃σ ∈ set θ. σ x ≠ {}"
using fg(2) prems θ unfolding fun_point_Union_list_def by auto
then obtain t' s' σ where ts':
"(t',s') ∈ set (zip ts ss)" "match_abss' t' s' = Some σ" "σ x ≠ {}"
using those_map2_SomeD[OF θ(1)[symmetric]] by blast
show ?case
using ts'(3) IH[OF conjI[OF fg] ts'(1) _ ts'(2)] set_zip_leftD[OF ts'(1)] by force
qed auto
lemma match_abss'_Abs_in_funs_term:
fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term"
assumes "match_abss' s t = Some δ"
and "a ∈ δ x"
shows "Abs a ∈ funs_term t"
using assms
proof (induction s t arbitrary: a δ rule: match_abss'.induct)
case (1 y b ts) show ?case
using match_abss'_Var_inv[OF "1.prems"(1)] "1.prems"(2)
by (cases "x = y") simp_all
next
case (2 f ts g ss)
note prems = "2.prems"
note IH = "2.IH"
obtain θ where θ: "Some θ = those (map2 match_abss' ts ss)" "δ = fun_point_Union_list θ"
and fg: "f = g" "length ts = length ss"
using match_abss'_Fun_inv[OF prems(1)] by fast
obtain t' s' σ where ts': "(t',s') ∈ set (zip ts ss)" "match_abss' t' s' = Some σ" "a ∈ σ x"
using fg(2) prems θ those_map2_SomeD[OF θ(1)[symmetric]]
unfolding fun_point_Union_list_def by fastforce
show ?case
using ts'(1) IH[OF conjI[OF fg] ts'(1) _ ts'(2,3)]
by (meson set_zip_rightD term.set_intros(2))
qed auto
lemma match_abss'_subst_fv_ex_abs:
assumes "match_abss' s (s ⋅ δ) = Some σ"
and TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
shows "∀x ∈ fv s. ∃a ts. δ x = Fun (Abs a) ts ∧ σ x = {a}" (is "?P s σ")
using assms(1)
proof (induction s "s ⋅ δ" arbitrary: σ rule: match_abss'.induct)
case (2 f ts g ss)
note prems = "2.prems"
note hyps = "2.hyps"
obtain θ where θ: "Some θ = those (map2 match_abss' ts ss)" "σ = fun_point_Union_list θ"
and fg: "f = g" "length ts = length ss" "ss = ts ⋅⇩l⇩i⇩s⇩t δ"
and ts: "∀(t,s) ∈ set (zip ts ss). ∃σ. match_abss' t s = Some σ"
using match_abss'_Fun_inv[OF prems(1)[unfolded hyps(2)[symmetric]]] hyps(2) by fastforce
have 0: "those (map (λt. match_abss' t (t ⋅ δ)) ts) = Some θ"
using θ(1) map2_map_subst unfolding fg(3) by metis
have 1: "∀t ∈ set ts. ∃σ. match_abss' t (t ⋅ δ) = Some σ"
using ts zip_map_subst[of ts δ] unfolding fg(3) by simp
have 2: "σ' ∈ set θ"
when t: "t ∈ set ts" "match_abss' t (t ⋅ δ) = Some σ'" for t σ'
using t 0 those_Some_iff[of "map (λt. match_abss' t (t ⋅ δ)) ts" θ] by force
have 3: "?P t σ'" "σ' x ≠ {}"
when t: "t ∈ set ts" "x ∈ fv t" "match_abss' t (t ⋅ δ) = Some σ'" for t σ' x
using t hyps(1)[OF conjI[OF fg(1,2)], of "(t, t ⋅ δ)" t σ'] zip_map_subst[of ts δ]
match_abss'_fv_is_nonempty[of t "t ⋅ δ" σ' x]
unfolding fg(3) by auto
have 4: "σ' x = {}"
when t: "x ∉ fv t" "match_abss' t (t ⋅ δ) = Some σ'" for t σ' x
by (meson t match_abss'_nonempty_is_fv)
show ?case
proof
fix x assume "x ∈ fv (Fun f ts)"
then obtain t σ' where t: "t ∈ set ts" "x ∈ fv t" and σ': "match_abss' t (t ⋅ δ) = Some σ'"
using 1 by auto
then obtain a tsa where a: "δ x = Fun (Abs a) tsa"
using 3[OF t σ'] by fast
have "σ'' x = {a} ∨ σ'' x = {}"
when "σ'' ∈ set θ" for σ''
using that a 0 3[of _ x] 4[of x]
unfolding those_Some_iff by fastforce
thus "∃a ts. δ x = Fun (Abs a) ts ∧ σ x = {a}"
using a 2[OF t(1) σ'] 3[OF t σ'] unfolding θ(2) fun_point_Union_list_def by auto
qed
qed auto
lemma match_abss'_subst_disj_nonempty:
assumes TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and "match_abss' s (s ⋅ δ) = Some σ"
and "x ∈ fv s"
shows "⋂(ticl_abs TI ` σ x) ≠ {} ∧ (∃a tsa. δ x = Fun (Abs a) tsa ∧ σ x = {a})" (is "?P σ")
using assms(2,3)
proof (induction s "s ⋅ δ" arbitrary: σ rule: match_abss'.induct)
case (1 x a ts) thus ?case unfolding ticl_abs_def by force
next
case (2 f ts g ss)
note prems = "2.prems"
note hyps = "2.hyps"
obtain θ where θ: "Some θ = those (map2 match_abss' ts ss)" "σ = fun_point_Union_list θ"
and fg: "f = g" "length ts = length ss" "ss = ts ⋅⇩l⇩i⇩s⇩t δ"
and ts: "∀(t,s) ∈ set (zip ts ss). ∃σ. match_abss' t s = Some σ"
using match_abss'_Fun_inv[OF prems(1)[unfolded hyps(2)[symmetric]]] hyps(2) by fastforce
define ts' where "ts' ≡ filter (λt. x ∈ fv t) ts"
define θ' where "θ' ≡ map (λt. (t, the (match_abss' t (t ⋅ δ)))) ts"
define θ'' where "θ'' ≡ map (λt. the (match_abss' t (t ⋅ δ))) ts'"
have 0: "those (map (λt. match_abss' t (t ⋅ δ)) ts) = Some θ"
using θ(1) map2_map_subst unfolding fg(3) by metis
have 1: "∀t ∈ set ts. ∃σ. match_abss' t (t ⋅ δ) = Some σ"
using ts zip_map_subst[of ts δ] unfolding fg(3) by simp
have ts_not_nil: "ts ≠ []"
using prems(2) by fastforce
hence "∃t ∈ set ts. x ∈ fv t" using prems(2) by simp
then obtain a tsa where a: "δ x = Fun (Abs a) tsa"
using 1 match_abss'_subst_fv_ex_abs[OF _ TI, of _ δ]
by metis
hence a': "σ' x = {a}"
when "t ∈ set ts" "x ∈ fv t" "match_abss' t (t ⋅ δ) = Some σ'"
for t σ'
using that match_abss'_subst_fv_ex_abs[OF _ TI, of _ δ]
by fastforce
have "ts' ≠ []" using prems(2) unfolding ts'_def by (simp add: filter_empty_conv)
hence θ''_not_nil: "θ'' ≠ []" unfolding θ''_def by simp
have 2: "σ' ∈ set θ"
when t: "t ∈ set ts" "match_abss' t (t ⋅ δ) = Some σ'" for t σ'
using t 0 those_Some_iff[of "map (λt. match_abss' t (t ⋅ δ)) ts" θ] by force
have 3: "?P σ'" "σ' x ≠ {}"
when t: "t ∈ set ts'" "match_abss' t (t ⋅ δ) = Some σ'" for t σ'
using t hyps(1)[OF conjI[OF fg(1,2)], of "(t, t ⋅ δ)" t σ'] zip_map_subst[of ts δ]
match_abss'_fv_is_nonempty[of t "t ⋅ δ" σ' x]
unfolding fg(3) ts'_def by (force, force)
have 4: "σ' x = {}"
when t: "x ∉ fv t" "match_abss' t (t ⋅ δ) = Some σ'" for t σ'
by (meson t match_abss'_nonempty_is_fv)
have 5: "θ = map snd θ'"
using 0 1 unfolding θ'_def by (induct ts arbitrary: θ) auto
have "fun_point_Union_list (map snd θ') x =
fun_point_Union_list (map snd (filter (λ(t,_). x ∈ fv t) θ')) x"
using 1 4 unfolding θ'_def fun_point_Union_list_def by fastforce
hence 6: "fun_point_Union_list θ x = fun_point_Union_list θ'' x"
using 0 1 4 unfolding 5 θ'_def θ''_def fun_point_Union_list_def ts'_def by auto
have 7: "?P σ'" "σ' x ≠ {}"
when σ': "σ' ∈ set θ''" for σ'
using that 1 3 unfolding θ''_def ts'_def by auto
have "σ' x = {a}"
when σ': "σ' ∈ set θ''" for σ'
using σ' a' 1 unfolding θ''_def ts'_def by fastforce
hence "fun_point_Union_list θ'' x = {b | b σ'. σ' ∈ set θ'' ∧ b ∈ {a}}"
using θ''_not_nil unfolding fun_point_Union_list_def by auto
hence 8: "fun_point_Union_list θ'' x = {a}"
using θ''_not_nil by auto
show ?case
using 8 a
unfolding θ(2) 6 ticl_abs_iff[OF TI] by auto
qed simp_all
lemma match_abssD:
fixes OCC TI s
defines "f ≡ (λδ x. if x ∈ fv s then δ x else set OCC)"
and "g ≡ (λδ x. ⋂(ticl_abs TI ` δ x))"
assumes δ': "match_abss OCC TI s t = Some δ'"
shows "∃δ. match_abss' s t = Some δ ∧ δ' = f (g δ) ∧ (∀x ∈ fv s. δ x ≠ {} ∧ f (g δ) x ≠ {}) ∧
(set OCC ≠ {} ⟶ (∀x. f (g δ) x ≠ {}))"
proof -
obtain δ where δ: "match_abss' s t = Some δ"
using δ' unfolding match_abss_def by force
hence "Some δ' = (if ∀x ∈ fv s. g δ x ≠ {} then Some (f (g δ)) else None)"
using δ' unfolding match_abss_def f_def g_def Let_def by simp
hence "δ' = f (g δ)" "∀x ∈ fv s. δ x ≠ {} ∧ f (g δ) x ≠ {}"
by (metis (no_types, lifting) option.inject option.distinct(1),
metis (no_types, lifting) f_def option.distinct(1) match_abss'_fv_is_nonempty[OF δ])
thus ?thesis using δ unfolding f_def by force
qed
lemma match_abss_ticl_abs_Inter_subset:
assumes TI: "set TI = {(a,b). (a,b) ∈ (set TI)⇧+ ∧ a ≠ b}"
and δ: "match_abss OCC TI s t = Some δ"
and x: "x ∈ fv s"
shows "⋂(ticl_abs TI ` δ x) ⊆ δ x"
proof -
let ?h1 = "λδ x. if x ∈ fv s then δ x else set OCC"
let ?h2 = "λδ x. ⋂(ticl_abs TI ` δ x)"
obtain δ' where δ':
"match_abss' s t = Some δ'" "δ = ?h1 (?h2 δ')"
"∀x ∈ fv s. δ' x ≠ {} ∧ δ x ≠ {}"
using match_abssD[OF δ] by blast
have "δ x = ⋂(ticl_abs TI ` δ' x)" "δ' x ≠ {}" "δ x ≠ {}"
using x δ'(2,3) by auto
thus ?thesis
using ticl_abs_Inter TI by simp
qed
lemma match_abss_fv_has_abs:
assumes "match_abss OCC TI s t = Some δ"
and "x ∈ fv s"
shows "δ x ≠ {}"
using assms match_abssD by fast
lemma match_abss_OCC_if_not_fv:
fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term"
assumes δ': "match_abss OCC TI s t = Some δ'"
and x: "x ∉ fv s"
shows "δ' x = set OCC"
proof -
define f where "f ≡ λs::(('a,'b,'c,'d) prot_fun, 'v) term. λδ x. if x ∈ fv s then δ x else set OCC"
define g where "g ≡ λδ. λx::'v. ⋂(ticl_abs TI ` δ x)"
obtain δ where δ: "match_abss' s t = Some δ" "δ' = f s (g δ)"
using match_abssD[OF δ'] unfolding f_def g_def by blast
show ?thesis
using x δ(2) unfolding f_def by presburger
qed
inductive synth_abs_substs_constrs_rel for FP OCC TI where
SolveNil:
"synth_abs_substs_constrs_rel FP OCC TI [] (λ_. set OCC)"
| SolveCons:
"ts ≠ [] ⟹ ∀t ∈ set ts. synth_abs_substs_constrs_rel FP OCC TI [t] (θ t)
⟹ synth_abs_substs_constrs_rel FP OCC TI ts (fun_point_Inter (θ ` set ts))"
| SolvePubConst:
"arity c = 0 ⟹ public c
⟹ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (λ_. set OCC)"
| SolvePrivConstIn:
"arity c = 0 ⟹ ¬public c ⟹ Fun c [] ∈ set FP
⟹ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (λ_. set OCC)"
| SolvePrivConstNotin:
"arity c = 0 ⟹ ¬public c ⟹ Fun c [] ∉ set FP
⟹ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (λ_. {})"
| SolveValueVar:
"θ = ((λ_. set OCC)(x := ticl_abss TI {a ∈ set OCC. ⟨a⟩⇩a⇩b⇩s ∈ set FP}))
⟹ synth_abs_substs_constrs_rel FP OCC TI [Var x] θ"
| SolvePubComposed:
"arity f > 0 ⟹ public f ⟹ length ts = arity f
⟹ ∀δ. δ ∈ Δ ⟷ (∃s ∈ set FP. match_abss OCC TI (Fun f ts) s = Some δ)
⟹ θ1 = fun_point_Union Δ
⟹ synth_abs_substs_constrs_rel FP OCC TI ts θ2
⟹ synth_abs_substs_constrs_rel FP OCC TI [Fun f ts] (fun_point_union θ1 θ2)"
| SolvePrivComposed:
"arity f > 0 ⟹ ¬public f ⟹ length ts = arity f
⟹ ∀δ. δ ∈ Δ ⟷ (∃s ∈ set FP. match_abss OCC TI (Fun f ts) s = Some δ)
⟹ θ = fun_point_Union Δ
⟹ synth_abs_substs_constrs_rel FP OCC TI [Fun f ts] θ"
fun synth_abs_substs_constrs_aux where
"synth_abs_substs_constrs_aux FP OCC TI (Var x) = (
(λ_. set OCC)(x := ticl_abss TI (set (filter (λa. ⟨a⟩⇩a⇩b⇩s ∈ set FP) OCC))))"
| "synth_abs_substs_constrs_aux FP OCC TI (Fun f ts) = (
if ts = []
then (if ¬public f ∧ Fun f ts ∉ set FP then (λ_. {}) else (λ_. set OCC))
else (let Δ = map the (filter (λδ. δ ≠ None) (map (match_abss OCC TI (Fun f ts)) FP));
θ1 = fun_point_Union_list Δ;
θ2 = fun_point_Inter_list (
case ts of t#ts' ⇒
if ¬is_Var t ∧ args t = [] ∧ ¬public (the_Fun t)
then (if t ∉ set FP then [λ_. {}]
else (λ_. set OCC)#map (synth_abs_substs_constrs_aux FP OCC TI) ts')
else map (synth_abs_substs_constrs_aux FP OCC TI) ts
)
in fun_point_union θ1 θ2))"
lemma synth_abs_substs_constrs_aux_fun_case:
assumes ts: "ts ≠ []"
shows "synth_abs_substs_constrs_aux FP OCC TI (Fun f ts) = (
let Δ = map the (filter (λδ. δ ≠ None) (map (match_abss OCC TI (Fun f ts)) FP));
θ1 = fun_point_Union_list Δ;
θ2 = fun_point_Inter_list (map (synth_abs_substs_constrs_aux FP OCC TI) ts)
in fun_point_union θ1 θ2)"
proof -
let ?s = "synth_abs_substs_constrs_aux FP OCC TI"
let ?P = "λt. ¬is_Var t ∧ args t = [] ∧ ¬public (the_Fun t)"
obtain t ts' where ts': "ts = t#ts'" using ts by (cases ts) auto
have "fun_point_Inter_list ((λ_. {})#map ?s ts') = fun_point_Inter_list [λ_. {}]"
unfolding fun_point_Inter_list_def by simp
thus ?thesis unfolding ts' by (cases "?P t") auto
qed
definition synth_abs_substs_constrs where
"synth_abs_substs_constrs FPT T ≡
let (FP,OCC,TI) = FPT;
ts = trms_list⇩s⇩s⇩t (unlabel (transaction_receive T));
f = fun_point_Inter_list ∘ map (synth_abs_substs_constrs_aux FP OCC TI)
in if ts = [] then (λ_. set OCC) else f ts"
definition transaction_check_comp::
"[('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set ⇒ bool,
('fun,'atom,'sets,'lbl) prot_term list ×
'sets set list ×
('sets set × 'sets set) list,
('fun,'atom,'sets,'lbl) prot_transaction]
⇒ ((('fun,'atom,'sets,'lbl) prot_var × 'sets set) list) list"
where
"transaction_check_comp msgcs FPT T ≡
let (_, OCC, _) = FPT;
S = unlabel (transaction_strand T);
C = unlabel (transaction_checks T);
xs = filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) (fv_list⇩s⇩s⇩t S);
posconstrs = transaction_poschecks_comp C;
negconstrs = transaction_negchecks_comp C;
pre_check = transaction_check_pre FPT T;
Δ = abs_substs_set xs OCC posconstrs negconstrs msgcs
in filter (λδ. pre_check (abs_substs_fun δ)) Δ"
definition transaction_check'::
"[('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set ⇒ bool,
('fun,'atom,'sets,'lbl) prot_term list ×
'sets set list ×
('sets set × 'sets set) list,
('fun,'atom,'sets,'lbl) prot_transaction]
⇒ bool"
where
"transaction_check' msgcs FPT T ≡
list_all (λδ. transaction_check_post FPT T (abs_substs_fun δ))
(transaction_check_comp msgcs FPT T)"
definition transaction_check::
"[('fun,'atom,'sets,'lbl) prot_term list ×
'sets set list ×
('sets set × 'sets set) list,
('fun,'atom,'sets,'lbl) prot_transaction]
⇒ bool"
where
"transaction_check ≡ transaction_check' (λ_ _. True)"
definition transaction_check_coverage_rcv::
"[('fun,'atom,'sets,'lbl) prot_term list ×
'sets set list ×
('sets set × 'sets set) list,
('fun,'atom,'sets,'lbl) prot_transaction]
⇒ bool"
where
"transaction_check_coverage_rcv FPT T ≡
let msgcs = synth_abs_substs_constrs FPT T
in transaction_check' (λx a. a ∈ msgcs x) FPT T"
lemma abs_subst_fun_cons:
"abs_substs_fun ((x,b)#δ) = (abs_substs_fun δ)(x := b)"
unfolding abs_substs_fun_def by fastforce
lemma abs_substs_cons:
assumes "δ ∈ set (abs_substs_set xs as poss negs msgcs)"
"b ∈ set as" "poss x ⊆ b" "b ∩ negs x = {}" "msgcs x b"
shows "(x,b)#δ ∈ set (abs_substs_set (x#xs) as poss negs msgcs)"
using assms by auto
lemma abs_substs_cons':
assumes δ: "δ ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)"
and b: "b ∈ set as" "poss x ⊆ b" "b ∩ negs x = {}" "msgcs x b"
shows "δ(x := b) ∈ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)"
proof -
obtain θ where θ: "δ = abs_substs_fun θ" "θ ∈ set (abs_substs_set xs as poss negs msgcs)"
using δ by force
have "abs_substs_fun ((x, b)#θ) ∈ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)"
using abs_substs_cons[OF θ(2) b] by blast
thus ?thesis
using θ(1) abs_subst_fun_cons[of x b θ] by argo
qed
lemma abs_substs_has_abs:
assumes "∀x. x ∈ set xs ⟶ δ x ∈ set as"
and "∀x. x ∈ set xs ⟶ poss x ⊆ δ x"
and "∀x. x ∈ set xs ⟶ δ x ∩ negs x = {}"
and "∀x. x ∈ set xs ⟶ msgcs x (δ x)"
and "∀x. x ∉ set xs ⟶ δ x = {}"
shows "δ ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)"
using assms
proof (induction xs arbitrary: δ)
case (Cons x xs)
define θ where "θ ≡ λy. if y ∈ set xs then δ y else {}"
have "θ ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)"
using Cons.prems Cons.IH by (simp add: θ_def)
moreover have "δ x ∈ set as" "poss x ⊆ δ x" "δ x ∩ negs x = {}" "msgcs x (δ x)"
by (simp_all add: Cons.prems(1,2,3,4))
ultimately have 0: "θ(x := δ x) ∈ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)"
by (metis abs_substs_cons')
have "δ = θ(x := δ x)"
proof
fix y show "δ y = (θ(x := δ x)) y"
proof (cases "y ∈ set (x#xs)")
case False thus ?thesis using Cons.prems(5) by (fastforce simp add: θ_def)
qed (auto simp add: θ_def)
qed
thus ?case by (metis 0)
qed (auto simp add: abs_substs_fun_def)
lemma abs_substs_abss_bounded:
assumes "δ ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)"
and "x ∈ set xs"
shows "δ x ∈ set as"
and "poss x ⊆ δ x"
and "δ x ∩ negs x = {}"
and "msgcs x (δ x)"
using assms
proof (induct xs as poss negs msgcs arbitrary: δ rule: abs_substs_set_induct)
case (Cons y xs as poss negs msgcs)
{ case 1 thus ?case using Cons.hyps(1) unfolding abs_substs_fun_def by fastforce }
{ case 2 thus ?case
proof (cases "x = y")
case False
then obtain δ' where δ':
"δ' ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "δ' x = δ x"
using 2 unfolding abs_substs_fun_def by force
moreover have "x ∈ set xs" using 2(2) False by simp
moreover have "∃b. b ∈ set as ∧ poss y ⊆ b ∧ b ∩ negs y = {}"
using 2 False by auto
ultimately show ?thesis using Cons.hyps(2) by fastforce
qed (auto simp add: abs_substs_fun_def)
}
{ case 3 thus ?case
proof (cases "x = y")
case False
then obtain δ' where δ':
"δ' ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "δ' x = δ x"
using 3 unfolding abs_substs_fun_def by force
moreover have "x ∈ set xs" using 3(2) False by simp
moreover have "∃b. b ∈ set as ∧ poss y ⊆ b ∧ b ∩ negs y = {}"
using 3 False by auto
ultimately show ?thesis using Cons.hyps(3) by fastforce
qed (auto simp add: abs_substs_fun_def)
}
{ case 4 thus ?case
proof (cases "x = y")
case False
then obtain δ' where δ':
"δ' ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "δ' x = δ x"
using 4 unfolding abs_substs_fun_def by force
moreover have "x ∈ set xs" using 4(2) False by simp
moreover have "∃b. b ∈ set as ∧ poss y ⊆ b ∧ b ∩ negs y = {}"
using 4 False by auto
ultimately show ?thesis using Cons.hyps(4) by fastforce
qed (auto simp add: abs_substs_fun_def)
}
qed (simp_all add: abs_substs_fun_def)
lemma abs_substs_abss_bounded':
assumes "δ ∈ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)"
and "x ∉ set xs"
shows "δ x = {}"
using assms unfolding abs_substs_fun_def
by (induct xs as poss negs msgcs arbitrary: δ rule: abs_substs_set_induct) (force, fastforce)
lemma transaction_poschecks_comp_unfold:
"transaction_poschecks_comp C x = {s. ∃a. ⟨a: Var x ∈ Fun (Set s) []⟩ ∈ set C}"
proof (induction C)
case (Cons c C) thus ?case
proof (cases "∃a y s. c = ⟨a: Var y ∈ Fun (Set s) []⟩")
case True
then obtain a y s where c: "c = ⟨a: Var y ∈ Fun (Set s) []⟩" by force
define f where "f ≡ transaction_poschecks_comp C"
have "transaction_poschecks_comp (c#C) = f(y := insert s (f y))"
using c by (simp add: f_def Let_def)
moreover have "f x = {s. ∃a. ⟨a: Var x ∈ Fun (Set s) []⟩ ∈ set C}"
using Cons.IH unfolding f_def by blast
ultimately show ?thesis using c by auto
next
case False
hence "transaction_poschecks_comp (c#C) = transaction_poschecks_comp C" (is ?P)
using transaction_poschecks_comp.cases[of "c#C" ?P] by force
thus ?thesis using False Cons.IH by auto
qed
qed simp
lemma transaction_poschecks_comp_notin_fv_empty:
assumes "x ∉ fv⇩s⇩s⇩t C"
shows "transaction_poschecks_comp C x = {}"
using assms transaction_poschecks_comp_unfold[of C x] by fastforce
lemma transaction_negchecks_comp_unfold:
"transaction_negchecks_comp C x = {s. ⟨Var x not in Fun (Set s) []⟩ ∈ set C}"
proof (induction C)
case (Cons c C) thus ?case
proof (cases "∃y s. c = ⟨Var y not in Fun (Set s) []⟩")
case True
then obtain y s where c: "c = ⟨Var y not in Fun (Set s) []⟩" by force
define f where "f ≡ transaction_negchecks_comp C"
have "transaction_negchecks_comp (c#C) = f(y := insert s (f y))"
using c by (simp add: f_def Let_def)
moreover have "f x = {s. ⟨Var x not in Fun (Set s) []⟩ ∈ set C}"
using Cons.IH unfolding f_def by blast
ultimately show ?thesis using c by auto
next
case False
hence "transaction_negchecks_comp (c#C) = transaction_negchecks_comp C" (is ?P)
using transaction_negchecks_comp.cases[of "c#C" ?P]
by force
thus ?thesis using False Cons.IH by fastforce
qed
qed simp
lemma transaction_negchecks_comp_notin_fv_empty:
assumes "x ∉ fv⇩s⇩s⇩t C"
shows "transaction_negchecks_comp C x = {}"
using assms transaction_negchecks_comp_unfold[of C x] by fastforce
lemma transaction_check_preI[intro]:
fixes T
defines "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
and "C ≡ set (unlabel (transaction_checks T))"
assumes a0: "∀x ∈ set (transaction_fresh T). δ x = {}"
and a1: "∀x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value ⟶ δ x ∈ set OCC"
and a2: "∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). intruder_synth_mod_timpls FP TI (t ⋅ θ δ)"
and a3: "∀a x s. ⟨a: Var x ∈ Fun (Set s) []⟩ ∈ C ⟶ s ∈ δ x"
and a4: "∀x s. ⟨Var x not in Fun (Set s) []⟩ ∈ C ⟶ s ∉ δ x"
shows "transaction_check_pre (FP, OCC, TI) T δ"
proof -
let ?P = "λu. is_InSet u ⟶ (
let x = the_elem_term u; s = the_set_term u
in (is_Var x ∧ is_Fun_Set s) ⟶ the_Set (the_Fun s) ∈ δ (the_Var x))"
let ?Q = "λu. (is_NegChecks u ∧ bvars⇩s⇩s⇩t⇩p u = [] ∧ the_eqs u = [] ∧ length (the_ins u) = 1) ⟶ (
let x = fst (hd (the_ins u)); s = snd (hd (the_ins u))
in (is_Var x ∧ is_Fun_Set s) ⟶ the_Set (the_Fun s) ∉ δ (the_Var x))"
have 1: "?P u" when u: "u ∈ C" for u
apply (unfold Let_def, intro impI, elim conjE)
using u a3 Fun_Set_InSet_iff[of u] by metis
have 2: "?Q u" when u: "u ∈ C" for u
apply (unfold Let_def, intro impI, elim conjE)
using u a4 Fun_Set_NotInSet_iff[of u] by metis
show ?thesis
using a0 a1 a2 1 2 fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"]
unfolding transaction_check_pre_def θ_def C_def Let_def
by blast
qed
lemma transaction_check_pre_InSetE:
assumes T: "transaction_check_pre FPT T δ"
and u: "u = ⟨a: Var x ∈ Fun (Set s) []⟩"
"u ∈ set (unlabel (transaction_checks T))"
shows "s ∈ δ x"
proof -
have "is_InSet u ⟶ is_Var (the_elem_term u) ∧ is_Fun_Set (the_set_term u) ⟶
the_Set (the_Fun (the_set_term u)) ∈ δ (the_Var (the_elem_term u))"
using T u unfolding transaction_check_pre_def Let_def by blast
thus ?thesis using Fun_Set_InSet_iff[of u a x s] u by argo
qed
lemma transaction_check_pre_NotInSetE:
assumes T: "transaction_check_pre FPT T δ"
and u: "u = ⟨Var x not in Fun (Set s) []⟩"
"u ∈ set (unlabel (transaction_checks T))"
shows "s ∉ δ x"
proof -
have "is_NegChecks u ∧ bvars⇩s⇩s⇩t⇩p u = [] ∧ the_eqs u = [] ∧ length (the_ins u) = 1 ⟶
is_Var (fst (hd (the_ins u))) ∧ is_Fun_Set (snd (hd (the_ins u))) ⟶
the_Set (the_Fun (snd (hd (the_ins u)))) ∉ δ (the_Var (fst (hd (the_ins u))))"
using T u unfolding transaction_check_pre_def Let_def by blast
thus ?thesis using Fun_Set_NotInSet_iff[of u x s] u by argo
qed
lemma transaction_check_pre_ReceiveE:
defines "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
assumes T: "transaction_check_pre (FP, OCC, TI) T δ"
and t: "t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T)"
shows "intruder_synth_mod_timpls FP TI (t ⋅ θ δ)"
using T t unfolding transaction_check_pre_def Let_def θ_def by blast
lemma transaction_check_compI[intro]:
assumes T: "transaction_check_pre (FP, OCC, TI) T δ"
and T_adm: "admissible_transaction' T"
and x1: "∀x. (x ∈ fv_transaction T - set (transaction_fresh T) ∧ fst x = TAtom Value)
⟶ δ x ∈ set OCC ∧ msgcs x (δ x)"
and x2: "∀x. (x ∉ fv_transaction T - set (transaction_fresh T) ∨ fst x ≠ TAtom Value)
⟶ δ x = {}"
shows "δ ∈ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)"
proof -
define S where "S ≡ unlabel (transaction_strand T)"
define C where "C ≡ unlabel (transaction_checks T)"
let ?xs = "fv_list⇩s⇩s⇩t S"
define poss where "poss ≡ transaction_poschecks_comp C"
define negs where "negs ≡ transaction_negchecks_comp C"
define ys where "ys ≡ filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) ?xs"
have ys: "{x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value} = set ys"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of S]
unfolding ys_def S_def by force
have "δ x ∈ set OCC" "msgcs x (δ x)"
when x: "x ∈ set ys" for x
using x1 x ys by (blast, blast)
moreover have "δ x = {}"
when x: "x ∉ set ys" for x
using x2 x ys by blast
moreover have "poss x ⊆ δ x" when x: "x ∈ set ys" for x
proof -
have "s ∈ δ x" when u: "u = ⟨a: Var x ∈ Fun (Set s) []⟩" "u ∈ set C" for u a s
using T u transaction_check_pre_InSetE[of "(FP, OCC, TI)" T δ]
unfolding C_def by blast
thus ?thesis
using transaction_poschecks_comp_unfold[of C x]
unfolding poss_def by blast
qed
moreover have "δ x ∩ negs x = {}" when x: "x ∈ set ys" for x
proof (cases "x ∈ fv⇩s⇩s⇩t C")
case True
hence "s ∉ δ x" when u: "u = ⟨Var x not in Fun (Set s) []⟩" "u ∈ set C" for u s
using T u transaction_check_pre_NotInSetE[of "(FP, OCC, TI)" T δ]
unfolding C_def by blast
thus ?thesis
using transaction_negchecks_comp_unfold[of C x]
unfolding negs_def by blast
next
case False
hence "negs x = {}"
using x transaction_negchecks_comp_notin_fv_empty
unfolding negs_def by blast
thus ?thesis by blast
qed
ultimately have "δ ∈ abs_substs_fun ` set (abs_substs_set ys OCC poss negs msgcs)"
using abs_substs_has_abs[of ys δ OCC poss negs msgcs]
by fast
thus ?thesis
using T
unfolding transaction_check_comp_def Let_def S_def C_def ys_def poss_def negs_def
by fastforce
qed
context
begin
private lemma transaction_check_comp_in_aux:
fixes T
defines "C ≡ set (unlabel (transaction_checks T))"
assumes T_adm: "admissible_transaction' T"
and a1: "∀x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value ⟶ (∀s.
select⟨Var x, Fun (Set s) []⟩ ∈ C ⟶ s ∈ α x)"
and a2: "∀x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value ⟶ (∀s.
⟨Var x in Fun (Set s) []⟩ ∈ C ⟶ s ∈ α x)"
and a3: "∀x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value ⟶ (∀s.
⟨Var x not in Fun (Set s) []⟩ ∈ C ⟶ s ∉ α x)"
shows "∀a x s. ⟨a: Var x ∈ Fun (Set s) []⟩ ∈ C ⟶ s ∈ α x" (is ?A)
and "∀x s. ⟨Var x not in Fun (Set s) []⟩ ∈ C ⟶ s ∉ α x" (is ?B)
proof -
note * = admissible_transaction_strand_step_cases(2,3)[OF T_adm]
have 1: "fst x = TAtom Value" "x ∈ fv_transaction T - set (transaction_fresh T)"
when x: "⟨a: Var x ∈ Fun (Set s) []⟩ ∈ C" for a x s
using * x unfolding C_def by fast+
have 2: "fst x = TAtom Value" "x ∈ fv_transaction T - set (transaction_fresh T)"
when x: "⟨Var x not in Fun (Set s) []⟩ ∈ C" for x s
using * x unfolding C_def by fast+
show ?A
proof (intro allI impI)
fix a x s assume u: "⟨a: Var x ∈ Fun (Set s) []⟩ ∈ C"
thus "s ∈ α x" using 1 a1 a2 by (cases a) metis+
qed
show ?B
proof (intro allI impI)
fix x s assume u: "⟨Var x not in Fun (Set s) []⟩ ∈ C"
thus "s ∉ α x" using 2 a3 by meson
qed
qed
lemma transaction_check_comp_in:
fixes T
defines "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
and "C ≡ set (unlabel (transaction_checks T))"
assumes T_adm: "admissible_transaction' T"
and a1: "∀x ∈ set (transaction_fresh T). α x = {}"
and a2: "∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). intruder_synth_mod_timpls FP TI (t ⋅ θ α)"
and a3: "∀x ∈ fv_transaction T - set (transaction_fresh T). ∀s.
select⟨Var x, Fun (Set s) []⟩ ∈ C ⟶ s ∈ α x"
and a4: "∀x ∈ fv_transaction T - set (transaction_fresh T). ∀s.
⟨Var x in Fun (Set s) []⟩ ∈ C ⟶ s ∈ α x"
and a5: "∀x ∈ fv_transaction T - set (transaction_fresh T). ∀s.
⟨Var x not in Fun (Set s) []⟩ ∈ C ⟶ s ∉ α x"
and a6: "∀x ∈ fv_transaction T - set (transaction_fresh T).
fst x = TAtom Value ⟶ α x ∈ set OCC"
and a7: "∀x ∈ fv_transaction T - set (transaction_fresh T).
fst x = TAtom Value ⟶ msgcs x (α x)"
shows "∃δ ∈ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T).
∀x ∈ fv_transaction T. fst x = TAtom Value ⟶ α x = δ x"
proof -
let ?xs = "fv_list⇩s⇩s⇩t (unlabel (transaction_strand T))"
let ?ys = "filter (λx. x ∉ set (transaction_fresh T)) ?xs"
define α' where "α' ≡ λx.
if x ∈ fv_transaction T - set (transaction_fresh T) ∧ fst x = TAtom Value
then α x
else {}"
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
have θα_Fun: "is_Fun (t ⋅ θ α) ⟷ is_Fun (t ⋅ θ α')" for t
unfolding α'_def θ_def
by (induct t) auto
have "∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). intruder_synth_mod_timpls FP TI (t ⋅ θ α')"
proof (intro ballI impI)
fix t assume t: "t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T)"
have 1: "intruder_synth_mod_timpls FP TI (t ⋅ θ α)"
using t a2
by auto
obtain r where r:
"r ∈ set (unlabel (transaction_receive T))"
"t ∈ trms⇩s⇩s⇩t⇩p r"
using t by auto
hence "∃ts. r = receive⟨ts⟩ ∧ t ∈ set ts"
using wellformed_transaction_unlabel_cases(1)[OF T_wf]
by fastforce
hence 2: "fv t ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T)" using r by force
have "fv t ⊆ fv_transaction T"
by (metis (no_types, lifting) 2 transaction_strand_def sst_vars_append_subset(1)
unlabel_append subset_Un_eq sup.bounded_iff)
moreover have "fv t ∩ set (transaction_fresh T) = {}"
using 2 T_wf vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel (transaction_receive T)"]
unfolding wellformed_transaction_def
by fast
ultimately have "θ α x = θ α' x" when "x ∈ fv t" for x
using that unfolding α'_def θ_def by fastforce
hence 3: "t ⋅ θ α = t ⋅ θ α'"
using term_subst_eq by blast
show "intruder_synth_mod_timpls FP TI (t ⋅ θ α')" using 1 3 by simp
qed
moreover have
"∀x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value ⟶ (∀s.
select⟨Var x, Fun (Set s) []⟩ ∈ C ⟶ s ∈ α' x)"
"∀x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value ⟶ (∀s.
⟨Var x in Fun (Set s) []⟩ ∈ C ⟶ s ∈ α' x)"
"∀x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value ⟶ (∀s.
⟨Var x not in Fun (Set s) []⟩ ∈ C ⟶ s ∉ α' x)"
using a3 a4 a5
unfolding α'_def θ_def C_def
by meson+
hence "∀a x s. ⟨a: Var x ∈ Fun (Set s) []⟩ ∈ C ⟶ s ∈ α' x"
"∀x s. ⟨Var x not in Fun (Set s) []⟩ ∈ C ⟶ s ∉ α' x"
using transaction_check_comp_in_aux[OF T_adm, of α']
unfolding C_def
by (fast, fast)
ultimately have 4: "transaction_check_pre (FP, OCC, TI) T α'"
using a6 transaction_check_preI[of T α' OCC FP TI]
unfolding α'_def θ_def C_def
by simp
have 5: "∀x ∈ fv_transaction T. fst x = TAtom Value ⟶ α x = α' x"
using a1 by (auto simp add: α'_def)
have 6: "α' ∈ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)"
using transaction_check_compI[OF 4 T_adm, of msgcs] a6 a7
unfolding α'_def
by auto
show ?thesis using 5 6 by blast
qed
end
lemma transaction_check_trivial_case:
assumes "transaction_updates T = []"
and "transaction_send T = []"
shows "transaction_check FPT T"
using assms
by (simp add: list_all_iff transaction_check_def transaction_check'_def transaction_check_post_def)
end
subsection ‹Soundness of the Occurs-Message Transformation›
context stateful_protocol_model
begin
context
begin
text ‹The occurs-message transformation, ‹add_occurs_msgs›, extends a transaction ‹T› with
additional message-transmission steps such that the following holds:
1. for each fresh variable ‹x› of ‹T› the message ‹occurs (Var x)› now occurs in a send-step,
2. for each of the remaining free variables ‹x› of ‹T› the message ‹occurs (Var x)› now occurs in a
receive-step.›
definition add_occurs_msgs where
"add_occurs_msgs T ≡
let frsh = transaction_fresh T;
xs = filter (λx. x ∉ set frsh) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)));
f = map (λx. occurs (Var x));
g = λC. if xs = [] then C else ⟨⋆, receive⟨f xs⟩⟩#C;
h = λF. if frsh = [] then F
else if F ≠ [] ∧ fst (hd F) = ⋆ ∧ is_Send (snd (hd F))
then ⟨⋆,send⟨f frsh@the_msgs (snd (hd F))⟩⟩#tl F
else ⟨⋆,send⟨f frsh⟩⟩#F
in case T of Transaction A B C D E F ⇒ Transaction A B (g C) D E (h F)"
private fun rm_occurs_msgs_constr where
"rm_occurs_msgs_constr [] = []"
| "rm_occurs_msgs_constr ((l,receive⟨ts⟩)#A) = (
if ∃t. occurs t ∈ set ts
then if ∃t ∈ set ts. ∀s. t ≠ occurs s
then (l,receive⟨filter (λt. ∀s. t ≠ occurs s) ts⟩)#rm_occurs_msgs_constr A
else rm_occurs_msgs_constr A
else (l,receive⟨ts⟩)#rm_occurs_msgs_constr A)"
| "rm_occurs_msgs_constr ((l,send⟨ts⟩)#A) = (
if ∃t. occurs t ∈ set ts
then if ∃t ∈ set ts. ∀s. t ≠ occurs s
then (l,send⟨filter (λt. ∀s. t ≠ occurs s) ts⟩)#rm_occurs_msgs_constr A
else rm_occurs_msgs_constr A
else (l,send⟨ts⟩)#rm_occurs_msgs_constr A)"
| "rm_occurs_msgs_constr (a#A) = a#rm_occurs_msgs_constr A"
private lemma add_occurs_msgs_cases:
fixes T C frsh xs f
defines "T' ≡ add_occurs_msgs T"
and "frsh ≡ transaction_fresh T"
and "xs ≡ filter (λx. x ∉ set frsh) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
and "xs' ≡ fv_transaction T - set frsh"
and "f ≡ map (λx. occurs (Var x))"
and "C' ≡ if xs = [] then C else ⟨⋆, receive⟨f xs⟩⟩#C"
and "ts' ≡ f frsh"
assumes T: "T = Transaction A B C D E F"
shows "F = ⟨⋆,send⟨ts⟩⟩#F' ⟹ T' = Transaction A B C' D E (⟨⋆,send⟨ts'@ts⟩⟩#F')"
(is "?A ts F' ⟹ ?A' ts F'")
and "∄ts' F'. F = ⟨⋆,send⟨ts'⟩⟩#F' ⟹ frsh ≠ [] ⟹ T' = Transaction A B C' D E (⟨⋆,send⟨ts'⟩⟩#F)"
(is "?B ⟹ ?B' ⟹ ?B''")
and "frsh = [] ⟹ T' = Transaction A B C' D E F" (is "?C ⟹ ?C'")
and "transaction_decl T' = transaction_decl T"
and "transaction_fresh T' = transaction_fresh T"
and "xs = [] ⟹ transaction_receive T' = transaction_receive T"
and "xs ≠ [] ⟹ transaction_receive T' = ⟨⋆,receive⟨f xs⟩⟩#transaction_receive T"
and "transaction_checks T' = transaction_checks T"
and "transaction_updates T' = transaction_updates T"
and "transaction_send T = ⟨⋆,send⟨ts⟩⟩#F' ⟹
transaction_send T' = ⟨⋆,send⟨ts'@ts⟩⟩#F'" (is "?D ts F' ⟹ ?D' ts F'")
and "∄ts' F'. transaction_send T = ⟨⋆,send⟨ts'⟩⟩#F' ⟹ frsh ≠ [] ⟹
transaction_send T' = ⟨⋆,send⟨ts'⟩⟩#transaction_send T" (is "?E ⟹ ?E' ⟹ ?E''")
and "frsh = [] ⟹ transaction_send T' = transaction_send T" (is "?F ⟹ ?F'")
and "(xs' ≠ {} ∧ transaction_receive T' = ⟨⋆, receive⟨f xs⟩⟩#transaction_receive T) ∨
(xs' = {} ∧ transaction_receive T' = transaction_receive T)" (is ?G)
and "(frsh ≠ [] ∧ (∃ts F'.
transaction_send T = ⟨⋆,send⟨ts⟩⟩#F' ∧ transaction_send T' = ⟨⋆,send⟨ts'@ts⟩⟩#F')) ∨
(frsh ≠ [] ∧ transaction_send T' = ⟨⋆,send⟨ts'⟩⟩#transaction_send T) ∨
(frsh = [] ∧ transaction_send T' = transaction_send T)" (is ?H)
proof -
note defs = T'_def T frsh_def xs_def xs'_def f_def C'_def ts'_def add_occurs_msgs_def Let_def
show 0: "?A ts F' ⟹ ?A' ts F'" for ts F' unfolding defs by simp
have "F = [] ∨ fst (hd F) ≠ ⋆ ∨ ¬is_Send (snd (hd F))" when ?B
using that unfolding is_Send_def by (cases F) auto
thus 1: "?B ⟹ ?B' ⟹ ?B''" unfolding defs by force
show "?C ⟹ ?C'" unfolding defs by auto
show "transaction_decl T' = transaction_decl T"
"transaction_fresh T' = transaction_fresh T"
"transaction_checks T' = transaction_checks T"
"transaction_updates T' = transaction_updates T"
unfolding defs by simp_all
show "xs = [] ⟹ transaction_receive T' = transaction_receive T"
"xs ≠ [] ⟹ transaction_receive T' = ⟨⋆, receive⟨f xs⟩⟩#transaction_receive T"
unfolding defs by simp_all
moreover have "xs = [] ⟷ xs' = {}"
using filter_empty_conv[of "λx. x ∉ set frsh"]
fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"]
unfolding xs_def xs'_def by blast
ultimately show ?G by blast
show 2: "?D ts F' ⟹ ?D' ts F'" for ts F' using 0 unfolding T by simp
show 3: "?E ⟹ ?E' ⟹ ?E''" using 1 unfolding T by force
show 4: "?F ⟹ ?F'" unfolding defs by simp
show ?H
proof (cases "frsh = []")
case False thus ?thesis
using 2 3[OF _ False] by (cases "∃ts F'. transaction_send T = ⟨⋆,send⟨ts⟩⟩#F'") (blast,blast)
qed (simp add: 4)
qed
private lemma add_occurs_msgs_transaction_strand_set:
fixes T C frsh xs f
defines "frsh ≡ transaction_fresh T"
and "xs ≡ filter (λx. x ∉ set frsh) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
and "f ≡ map (λx. occurs (Var x))"
assumes T: "T = Transaction A B C D E F"
shows "F = ⟨⋆,send⟨ts⟩⟩#F' ⟹
set (transaction_strand (add_occurs_msgs T)) ⊆
set (transaction_strand T) ∪ {⟨⋆,receive⟨f xs⟩⟩,⟨⋆,send⟨f frsh@ts⟩⟩}"
(is "?A ⟹ ?A'")
and "F = ⟨⋆,send⟨ts⟩⟩#F' ⟹
set (unlabel (transaction_strand (add_occurs_msgs T))) ⊆
set (unlabel (transaction_strand T)) ∪ {receive⟨f xs⟩,send⟨f frsh@ts⟩}"
(is "?B ⟹ ?B'")
and "∄ts' F'. F = ⟨⋆,send⟨ts'⟩⟩#F' ⟹
set (transaction_strand (add_occurs_msgs T)) ⊆
set (transaction_strand T) ∪ {⟨⋆,receive⟨f xs⟩⟩,⟨⋆,send⟨f frsh⟩⟩}"
(is "?C ⟹ ?C'")
and "∄ts' F'. F = ⟨⋆,send⟨ts'⟩⟩#F' ⟹
set (unlabel (transaction_strand (add_occurs_msgs T))) ⊆
set (unlabel (transaction_strand T)) ∪ {receive⟨f xs⟩,send⟨f frsh⟩}"
(is "?D ⟹ ?D'")
proof -
note 0 = add_occurs_msgs_cases[
OF T, unfolded frsh_def[symmetric] xs_def[symmetric] f_def[symmetric]]
show "?A ⟹ ?A'" using 0(1,3) unfolding T transaction_strand_def by (cases "frsh = []") auto
thus "?B ⟹ ?B'" unfolding unlabel_def by force
show "?C ⟹ ?C'" using 0(2,3) unfolding T transaction_strand_def by (cases "frsh = []") auto
thus "?D ⟹ ?D'" unfolding unlabel_def by auto
qed
private lemma add_occurs_msgs_transaction_strand_cases:
fixes T T'::"('a,'b,'c,'d) prot_transaction" and C frsh xs f θ
defines "T' ≡ add_occurs_msgs T"
and "S ≡ transaction_strand T"
and "S' ≡ transaction_strand T'"
and "frsh ≡ transaction_fresh T"
and "xs ≡ filter (λx. x ∉ set frsh) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
and "f ≡ map (λx. occurs (Var x))"
and "C ≡ transaction_receive T"
and "D ≡ transaction_checks T"
and "E ≡ transaction_updates T"
and "F ≡ transaction_send T"
and "C' ≡ if xs = [] then C else ⟨⋆,receive⟨f xs⟩⟩#C"
and "C'' ≡ if xs = [] then dual⇩l⇩s⇩s⇩t C else ⟨⋆,send⟨f xs⟩⟩#dual⇩l⇩s⇩s⇩t C"
and "C''' ≡ if xs = [] then dual⇩l⇩s⇩s⇩t (C ⋅⇩l⇩s⇩s⇩t θ) else ⟨⋆,send⟨f xs ⋅⇩l⇩i⇩s⇩t θ⟩⟩#dual⇩l⇩s⇩s⇩t (C ⋅⇩l⇩s⇩s⇩t θ)"
shows "frsh = [] ⟹ S' = C'@D@E@F"
(is "?A ⟹ ?A'")
and "frsh ≠ [] ⟹ ∄ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ⟹ S' = C'@D@E@(⟨⋆,send⟨f frsh⟩⟩#F)"
(is "?B ⟹ ?B' ⟹ ?B''")
and "frsh ≠ [] ⟹ ∃ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ⟹
∃ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ∧ S' = C'@D@E@(⟨⋆,send⟨f frsh@ts⟩⟩#F')"
(is "?C ⟹ ?C' ⟹ ?C''")
and "frsh = [] ⟹ dual⇩l⇩s⇩s⇩t S' = C''@dual⇩l⇩s⇩s⇩t D@dual⇩l⇩s⇩s⇩t E@dual⇩l⇩s⇩s⇩t F"
(is "?D ⟹ ?D'")
and "frsh ≠ [] ⟹ ∄ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ⟹
dual⇩l⇩s⇩s⇩t S' = C''@dual⇩l⇩s⇩s⇩t D@dual⇩l⇩s⇩s⇩t E@(⟨⋆,receive⟨f frsh⟩⟩#dual⇩l⇩s⇩s⇩t F)"
(is "?E ⟹ ?E' ⟹ ?E''")
and "frsh ≠ [] ⟹ ∃ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ⟹
∃ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ∧
dual⇩l⇩s⇩s⇩t S' = C''@dual⇩l⇩s⇩s⇩t D@dual⇩l⇩s⇩s⇩t E@(⟨⋆,receive⟨f frsh@ts⟩⟩#dual⇩l⇩s⇩s⇩t F')"
(is "?F ⟹ ?F' ⟹ ?F''")
and "frsh = [] ⟹
dual⇩l⇩s⇩s⇩t (S' ⋅⇩l⇩s⇩s⇩t θ) = C'''@dual⇩l⇩s⇩s⇩t (D ⋅⇩l⇩s⇩s⇩t θ)@dual⇩l⇩s⇩s⇩t (E ⋅⇩l⇩s⇩s⇩t θ)@dual⇩l⇩s⇩s⇩t (F ⋅⇩l⇩s⇩s⇩t θ)"
(is "?G ⟹ ?G'")
and "frsh ≠ [] ⟹ ∄ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ⟹
dual⇩l⇩s⇩s⇩t (S' ⋅⇩l⇩s⇩s⇩t θ) = C'''@dual⇩l⇩s⇩s⇩t (D ⋅⇩l⇩s⇩s⇩t θ)@dual⇩l⇩s⇩s⇩t (E ⋅⇩l⇩s⇩s⇩t θ)@
(⟨⋆,receive⟨f frsh ⋅⇩l⇩i⇩s⇩t θ⟩⟩#dual⇩l⇩s⇩s⇩t (F ⋅⇩l⇩s⇩s⇩t θ))"
(is "?H ⟹ ?H' ⟹ ?H''")
and "frsh ≠ [] ⟹ ∃ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ⟹
∃ts F'. F = ⟨⋆,send⟨ts⟩⟩#F' ∧
dual⇩l⇩s⇩s⇩t (S' ⋅⇩l⇩s⇩s⇩t θ) = C'''@dual⇩l⇩s⇩s⇩t (D ⋅⇩l⇩s⇩s⇩t θ)@dual⇩l⇩s⇩s⇩t (E ⋅⇩l⇩s⇩s⇩t θ)@
(⟨⋆,receive⟨f frsh@ts ⋅⇩l⇩i⇩s⇩t θ⟩⟩#dual⇩l⇩s⇩s⇩t (F' ⋅⇩l⇩s⇩s⇩t θ))"
(is "?I ⟹ ?I' ⟹ ?I''")
proof -
obtain A' B' CC' D' E' F' where T: "T = Transaction A' B' CC' D' E' F'" by (cases T) simp
note 0 = add_occurs_msgs_cases[
OF T, unfolded frsh_def[symmetric] xs_def[symmetric] f_def[symmetric] T'_def[symmetric]]
note defs = S'_def C_def D_def E_def F_def C'_def C''_def T transaction_strand_def
show A: "?A ⟹ ?A'" using 0(3) unfolding defs by simp
show B: "?B ⟹ ?B' ⟹ ?B''" using 0(2) unfolding defs by simp
show C: "?C ⟹ ?C' ⟹ ?C''" using 0(1) unfolding defs by force
have 1: "C''' = C'' ⋅⇩l⇩s⇩s⇩t θ"
using subst_lsst_cons[of "⟨⋆, send⟨f xs⟩⟩" "dual⇩l⇩s⇩s⇩t C" θ] dual⇩l⇩s⇩s⇩t_subst[of C θ]
unfolding C'''_def C''_def by (cases "xs = []") auto
have 2: "(⟨⋆, receive⟨ts⟩⟩#dual⇩l⇩s⇩s⇩t G) ⋅⇩l⇩s⇩s⇩t θ = ⟨⋆, receive⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩⟩#dual⇩l⇩s⇩s⇩t (G ⋅⇩l⇩s⇩s⇩t θ)"
for ts and G::"('a,'b,'c,'d) prot_strand"
using dual⇩l⇩s⇩s⇩t_subst[of G θ] subst_lsst_cons[of "⟨⋆, receive⟨ts⟩⟩" "dual⇩l⇩s⇩s⇩t G" θ]
by simp
note 3 = subst_lsst_append[of _ _ θ] dual⇩l⇩s⇩s⇩t_subst[of _ θ]
show "?D ⟹ ?D'" using A unfolding defs by fastforce
thus "?G ⟹ ?G'" unfolding 1 by (metis 3)
show "?E ⟹ ?E' ⟹ ?E''" using B unfolding defs by fastforce
thus "?H ⟹ ?H' ⟹ ?H''" unfolding 1 by (metis 2 3)
show "?F ⟹ ?F' ⟹ ?F''" using C unfolding defs by fastforce
thus "?I ⟹ ?I' ⟹ ?I''" unfolding 1 by (metis 2 3)
qed
private lemma add_occurs_msgs_trms_transaction:
fixes T::"('a,'b,'c,'d) prot_transaction"
shows "trms_transaction (add_occurs_msgs T) =
trms_transaction T ∪ (λx. occurs (Var x))`(fv_transaction T ∪ set (transaction_fresh T))"
(is "?A = ?B")
proof
let ?occs = "(λx. occurs (Var x)) ` (fv_transaction T ∪ set (transaction_fresh T))"
define frsh where "frsh ≡ transaction_fresh T"
define xs where "xs ≡ filter (λx. x ∉ set frsh) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
define f where "f ≡ map (λx. occurs (Var x)::('a,'b,'c,'d) prot_term)"
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp
note 0 = add_occurs_msgs_transaction_strand_set(2,4)[
OF T, unfolded f_def[symmetric] frsh_def[symmetric] xs_def[symmetric]]
note 1 = add_occurs_msgs_transaction_strand_cases(1,2,3)[
of T, unfolded f_def[symmetric] frsh_def[symmetric] xs_def[symmetric]]
have 2: "set (f xs) ∪ set (f frsh) = ?occs"
proof -
define ys where "ys ≡ fv_list⇩s⇩s⇩t (unlabel (transaction_strand T))"
let ?ys' = "fv_transaction T - set frsh"
define g where "g ≡ filter (λx. x ∉ set frsh)"
have "set (g ys) = ?ys'"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"] unfolding ys_def g_def by auto
hence "set (f (g ys)) = (λx. occurs (Var x)) ` ?ys'" unfolding f_def by force
moreover have "set (f frsh) = (λx. occurs (Var x)) ` set frsh" unfolding f_def by force
ultimately show ?thesis
unfolding xs_def frsh_def[symmetric] ys_def[symmetric] g_def[symmetric] by blast
qed
have 3: "set (f []) = {}" unfolding f_def by blast
have "trms_transaction (add_occurs_msgs T) ⊆ trms_transaction T ∪ set (f xs) ∪ set (f frsh)"
proof (cases "∃ts F'. F = ⟨⋆, send⟨ts⟩⟩#F'")
case True
then obtain ts F' where F: "F = ⟨⋆, send⟨ts⟩⟩#F'" by blast
have "set ts ⊆ trms_transaction T" unfolding T F trms_transaction_unfold by auto
thus ?thesis using 0(1)[OF F] by force
next
case False show ?thesis using 0(2)[OF False] by force
qed
thus "?A ⊆ ?B" using 2 by blast
have "trms_transaction T ∪ set (f xs) ∪ set (f frsh) ⊆ trms_transaction (add_occurs_msgs T)"
proof (cases "frsh = []")
case True show ?thesis using 1(1)[OF True] 3 unfolding True by (cases xs) (fastforce,force)
next
case False
note * = 1(2-)[OF False]
show ?thesis
proof (cases "∃ts F'. transaction_send T = ⟨⋆, send⟨ts⟩⟩#F'")
case True show ?thesis using *(2)[OF True] 3 by force
next
case False show ?thesis using *(1)[OF False] 3 by force
qed
qed
thus "?B ⊆ ?A" using 2 by blast
qed
private lemma add_occurs_msgs_vars_eq:
fixes T::"('fun,'var,'sets,'lbl) prot_transaction"
assumes T_adm: "admissible_transaction' T"
shows "fv⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)) =
fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪ fv⇩l⇩s⇩s⇩t (transaction_checks T)" (is ?A)
and "fv⇩l⇩s⇩s⇩t (transaction_send (add_occurs_msgs T)) =
fv⇩l⇩s⇩s⇩t (transaction_send T) ∪ set (transaction_fresh T)" (is ?B)
and "fv_transaction (add_occurs_msgs T) = fv_transaction T" (is ?C)
and "bvars_transaction (add_occurs_msgs T) = bvars_transaction T" (is ?D)
and "vars_transaction (add_occurs_msgs T) = vars_transaction T" (is ?E)
and "fv⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t θ) =
fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)" (is ?F)
and "bvars⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t θ) =
bvars⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)" (is ?G)
and "vars⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t θ) =
vars⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)" (is ?H)
and "set (transaction_fresh (add_occurs_msgs T)) = set (transaction_fresh T)" (is ?I)
proof -
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp
have T_fresh: "set (transaction_fresh T) ⊆ fv_transaction T"
using admissible_transactionE(7)[OF T_adm] unfolding fv_transaction_unfold by blast
note 0 = add_occurs_msgs_cases[OF T]
define xs where "xs ≡
filter (λx. x ∉ set (transaction_fresh T)) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
show D: ?D
proof -
have "bvars⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)) = bvars⇩l⇩s⇩s⇩t (transaction_receive T)"
using 0(6,7) by (cases "xs = []") (auto simp add: xs_def)
moreover have "bvars⇩l⇩s⇩s⇩t (transaction_send (add_occurs_msgs T)) = bvars⇩l⇩s⇩s⇩t (transaction_send T)"
proof (cases "∃ts' F'. F = ⟨⋆, send⟨ts'⟩⟩#F'")
case True thus ?thesis using 0(1) unfolding T by force
next
case False show ?thesis using 0(2)[OF False] 0(3) unfolding T by (cases "B = []") auto
qed
ultimately show ?thesis using 0(8,9) unfolding bvars_transaction_unfold by argo
qed
have T_no_bvars:
"bvars_transaction T = {}"
"bvars⇩l⇩s⇩s⇩t (transaction_receive T) = {}"
"bvars⇩l⇩s⇩s⇩t (transaction_checks T) = {}"
"bvars⇩l⇩s⇩s⇩t (transaction_send T) = {}"
"bvars_transaction (add_occurs_msgs T) = {}"
using admissible_transactionE(4)[OF T_adm] D
unfolding bvars_transaction_unfold by (blast,blast,blast,blast,blast)
have T_fv_subst:
"fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t δ) = fv⇩s⇩e⇩t (δ ` fv_transaction T)" (is ?Q1)
"fv⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t δ) = fv⇩s⇩e⇩t (δ ` fv⇩l⇩s⇩s⇩t (transaction_receive T))" (is ?Q2)
"fv⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t δ) = fv⇩s⇩e⇩t (δ ` fv⇩l⇩s⇩s⇩t (transaction_checks T))" (is ?Q3)
"fv⇩l⇩s⇩s⇩t (transaction_send T ⋅⇩l⇩s⇩s⇩t δ) = fv⇩s⇩e⇩t (δ ` fv⇩l⇩s⇩s⇩t (transaction_send T))" (is ?Q4)
"fv⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t δ) =
fv⇩s⇩e⇩t (δ ` fv⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T)))" (is ?Q5)
"fv⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t δ) =
fv⇩s⇩e⇩t (δ ` fv⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)))" (is ?Q6)
for δ
proof -
note * = fv⇩s⇩s⇩t_subst_if_no_bvars
have **: "bvars⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)) = {}"
using T_no_bvars(5) unfolding bvars_transaction_unfold by fast
show ?Q1 using *[OF T_no_bvars(1)] unfolding unlabel_subst by blast
show ?Q2 using *[OF T_no_bvars(2)] unfolding unlabel_subst by blast
show ?Q3 using *[OF T_no_bvars(3)] unfolding unlabel_subst by blast
show ?Q4 using *[OF T_no_bvars(4)] unfolding unlabel_subst by blast
show ?Q5 using *[OF T_no_bvars(5)] unfolding unlabel_subst by blast
show ?Q6 using *[OF **] unfolding unlabel_subst by blast
qed
have A: "fv⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t δ) =
fv⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t δ) ∪ fv⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t δ)"
for δ
proof -
define rcv_trms where
"rcv_trms ≡ map (λx. occurs (Var x)::('fun,'var,'sets,'lbl) prot_term) xs"
have "fv⇩s⇩e⇩t (set rcv_trms) = fv_transaction T - set (transaction_fresh T)"
"rcv_trms = [] ⟷ xs = []"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"]
unfolding rcv_trms_def xs_def by auto
hence 1: "fv⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)) =
(fv_transaction T - set (transaction_fresh T)) ∪ fv⇩l⇩s⇩s⇩t (transaction_receive T)"
using 0(6,7)[unfolded rcv_trms_def[symmetric] xs_def[symmetric]] by (cases "xs = []") auto
have 2: "fv⇩l⇩s⇩s⇩t (transaction_receive T) ⊆ fv_transaction T - set (transaction_fresh T)"
using admissible_transactionE(12)[OF T_adm] unfolding fv_transaction_unfold by fast
have 3: "fv_transaction T - set (transaction_fresh T) =
fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪ fv⇩l⇩s⇩s⇩t (transaction_checks T)"
using admissible_transactionE(7,10,12,13)[OF T_adm]
unfolding fv_transaction_unfold by blast
show ?thesis using 1 2 3 T_fv_subst(2,3,6)[of δ] by force
qed
show ?A using A[of Var] unfolding subst_lsst_id_subst by blast
show B: ?B using 0(14) by fastforce
have B': "fv⇩l⇩s⇩s⇩t (transaction_send (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t δ) =
fv⇩l⇩s⇩s⇩t (transaction_send T ⋅⇩l⇩s⇩s⇩t δ) ∪ fv⇩s⇩e⇩t (δ ` set (transaction_fresh T))"
for δ
proof -
note * = fv⇩s⇩s⇩t_subst_if_no_bvars[of _ δ]
have **: "bvars⇩l⇩s⇩s⇩t (transaction_send (add_occurs_msgs T)) = {}"
using T_no_bvars(5) unfolding bvars_transaction_unfold by fast
show ?thesis
using B *[OF T_no_bvars(4)] *[OF **]
unfolding unlabel_subst by simp
qed
show C: ?C
using A[of Var] B T_fresh
unfolding fv_transaction_unfold 0(8,9) subst_lsst_id_subst by blast
show ?E using C D vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t by metis
have "fv⇩s⇩e⇩t (θ ` set (transaction_fresh T)) ⊆ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"
using T_fresh
unfolding fv⇩s⇩s⇩t_subst_if_no_bvars[OF T_no_bvars(1), of θ, unfolded unlabel_subst]
by auto
thus F: ?F
using A[of θ] B'[of θ] fv⇩s⇩s⇩t_append
fv⇩s⇩s⇩t_subst_if_no_bvars[OF T_no_bvars(1), of θ, unfolded unlabel_subst]
fv⇩s⇩s⇩t_subst_if_no_bvars[OF T_no_bvars(5), of θ, unfolded unlabel_subst C]
unfolding transaction_strand_def by argo
show G: ?G using D bvars⇩s⇩s⇩t_subst unlabel_subst by metis
show ?H using F G vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t by metis
show ?I using 0(5) by argo
qed
private lemma add_occurs_msgs_trms:
"trms_transaction (add_occurs_msgs T) =
trms_transaction T ∪ (λx. occurs (Var x)) ` (set (transaction_fresh T) ∪ fv_transaction T)"
proof -
let ?f = "λx. occurs (Var x)"
let ?xs = "filter (λx. x ∉ set (transaction_fresh T))
(fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp
note 0 = add_occurs_msgs_cases[OF T]
have "set ?xs = fv_transaction T - set (transaction_fresh T)"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"] by auto
hence 1: "trms⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)) =
trms⇩l⇩s⇩s⇩t (transaction_receive T) ∪ ?f ` (fv_transaction T - set (transaction_fresh T))"
using 0(6,7) by (cases "?xs = []") auto
have 2: "trms⇩l⇩s⇩s⇩t (transaction_send (add_occurs_msgs T)) =
trms⇩l⇩s⇩s⇩t (transaction_send T) ∪ ?f ` set (transaction_fresh T)"
using 0(10,11,12) by (cases "transaction_fresh T = []") (simp,fastforce)
have 3: "trms⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)) ∪
trms⇩l⇩s⇩s⇩t (transaction_send (add_occurs_msgs T)) =
trms⇩l⇩s⇩s⇩t (transaction_receive T) ∪ trms⇩l⇩s⇩s⇩t (transaction_send T) ∪
?f ` (set (transaction_fresh T) ∪ fv_transaction T)"
using 1 2 by blast
show ?thesis using 3 unfolding trms_transaction_unfold 0(8,9) by blast
qed
lemma add_occurs_msgs_admissible_occurs_checks:
fixes T::"('fun,'atom,'sets,'lbl) prot_transaction"
assumes T_adm: "admissible_transaction' T"
shows "admissible_transaction' (add_occurs_msgs T)" (is ?A)
and "admissible_transaction_occurs_checks (add_occurs_msgs T)" (is ?B)
proof -
let ?T' = "add_occurs_msgs T"
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp
note defs = T add_occurs_msgs_def Let_def admissible_transaction'_def
admissible_transaction_occurs_checks_def
note defs' = admissible_transaction_terms_def wf⇩t⇩r⇩m⇩s_code[symmetric]
note 1 = add_occurs_msgs_cases[OF T]
note 2 = add_occurs_msgs_vars_eq[OF T_adm]
note 3 = add_occurs_msgs_trms[of T]
note 4 = add_occurs_msgs_transaction_strand_set[OF T]
have occurs_wf: "wf⇩t⇩r⇩m (occurs (Var x))" for x::"('fun,'atom,'sets,'lbl) prot_var" by fastforce
have occurs_funs: "funs_term (occurs (Var x)) = {OccursFact, OccursSec}"
for x::"('fun,'atom,'sets,'lbl) prot_var"
by force
have occurs_funs_not_attack: "¬(∃f ∈ ⋃(funs_term ` trms⇩s⇩s⇩t⇩p r). is_Attack f)"
when "r = receive⟨map (λx. occurs (Var x)) xs⟩ ∨ r = send⟨map (λx. occurs (Var x)) ys⟩"
for r::
"(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand_step"
and xs ys::"('fun,'atom,'sets,'lbl) prot_var list"
using that by fastforce
have occurs_funs_not_attack': "¬(∃f ∈ ⋃(funs_term ` trms⇩s⇩s⇩t⇩p r). is_Attack f)"
when "r = send⟨map (λx. occurs (Var x)) xs@ts⟩"
and "¬(∃f ∈ ⋃(funs_term ` trms⇩s⇩s⇩t⇩p (send⟨ts⟩)). is_Attack f)"
for r::
"(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand_step"
and xs::"('fun,'atom,'sets,'lbl) prot_var list"
and ts
using that by fastforce
let ?P1 = "λT. wellformed_transaction T"
let ?P2 = "λT. transaction_decl T () = []"
let ?P3 = "λT. list_all (λx. fst x = TAtom Value) (transaction_fresh T)"
let ?P4 = "λT. ∀x ∈ vars_transaction T. is_Var (fst x) ∧ (the_Var (fst x) = Value)"
let ?P5 = "λT. bvars⇩l⇩s⇩s⇩t (transaction_strand T) = {}"
let ?P6 = "λT. set (transaction_fresh T) ⊆
fv⇩l⇩s⇩s⇩t (filter (is_Insert ∘ snd) (transaction_updates T)) ∪
fv⇩l⇩s⇩s⇩t (transaction_send T)"
let ?P7 = "λT. ∀x ∈ fv_transaction T - set (transaction_fresh T).
∀y ∈ fv_transaction T - set (transaction_fresh T).
x ≠ y ⟶ ⟨Var x != Var y⟩ ∈ set (unlabel (transaction_checks T)) ∨
⟨Var y != Var x⟩ ∈ set (unlabel (transaction_checks T))"
let ?P8 = "λT. fv⇩l⇩s⇩s⇩t (transaction_updates T) ∪ fv⇩l⇩s⇩s⇩t (transaction_send T)
- set (transaction_fresh T)
⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪ fv⇩l⇩s⇩s⇩t (transaction_checks T)"
let ?P9 = "λT. ∀r ∈ set (unlabel (transaction_checks T)).
is_Equality r ⟶ fv (the_rhs r) ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T)"
let ?P10 = "λT. fv⇩l⇩s⇩s⇩t (transaction_checks T) ⊆
fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪
fv⇩l⇩s⇩s⇩t (filter (λs. is_InSet (snd s) ∧ the_check (snd s) = Assign)
(transaction_checks T))"
let ?P11 = "λT. admissible_transaction_checks T"
let ?P12 = "λT. admissible_transaction_updates T"
let ?P13 = "λT. admissible_transaction_terms T"
let ?P14 = "λT. admissible_transaction_send_occurs_form T"
let ?P15 = "λT. list_all (λa. is_Receive (snd a) ⟶ the_msgs (snd a) ≠ [])
(transaction_receive T)"
let ?P16 = "λT. list_all (λa. is_Send (snd a) ⟶ the_msgs (snd a) ≠ []) (transaction_send T)"
have T_props:
"?P1 T" "?P2 T" "?P3 T" "?P4 T" "?P5 T" "?P6 T" "?P7 T" "?P8 T" "?P9 T" "?P10 T" "?P11 T"
"?P12 T" "?P13 T" "?P14 T" "?P15 T" "?P16 T"
using T_adm unfolding defs by meson+
have 5: "wf'⇩s⇩s⇩t (X ∪ Y) (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T))))"
when X: "X = fst ` set (transaction_decl T ())"
and Y: "Y = set (transaction_fresh T)"
and T_wf: "wf'⇩s⇩s⇩t (X ∪ Y) (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T)))"
for X Y
proof -
define frsh where "frsh ≡ transaction_fresh T"
define xs where "xs ≡ fv_list⇩s⇩s⇩t (unlabel (transaction_strand T))"
define ys where "ys ≡ filter (λx. x ∉ set frsh) xs"
let ?snds = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T))"
let ?snds' = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive (add_occurs_msgs T)))"
let ?chks = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_checks T))"
let ?chks' = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_checks (add_occurs_msgs T)))"
let ?upds = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_updates T))"
let ?upds' = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_updates (add_occurs_msgs T)))"
let ?rcvs = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_send T))"
let ?rcvs' = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_send (add_occurs_msgs T)))"
have p0: "set ?snds ⊆ set ?snds'" using 1(13) by auto
have p1: "?chks = ?chks'" "?upds = ?upds'" using 1(8,9) by (argo,argo)
have p2: "wfvarsoccs⇩s⇩s⇩t ?snds ⊆ wfvarsoccs⇩s⇩s⇩t ?snds'"
"wfvarsoccs⇩s⇩s⇩t (?snds@?chks@?upds) ⊆ wfvarsoccs⇩s⇩s⇩t (?snds'@?chks'@?upds')"
"X ∪ Y ∪ wfvarsoccs⇩s⇩s⇩t (?snds@?chks@?upds) ⊆
X ∪ Y ∪ wfvarsoccs⇩s⇩s⇩t (?snds'@?chks'@?upds')"
using p0 p1 unfolding wfvarsoccs⇩s⇩s⇩t_def by auto
have "wf'⇩s⇩s⇩t (X ∪ Y ∪ wfvarsoccs⇩s⇩s⇩t (?snds@?chks@?upds)) ?rcvs"
using T_wf wf⇩s⇩s⇩t_append_exec[of "X ∪ Y" "?snds@?chks@?upds" ?rcvs]
unfolding transaction_strand_unlabel_dual_unfold by simp
hence r0: "wf'⇩s⇩s⇩t (X ∪ Y ∪ wfvarsoccs⇩s⇩s⇩t (?snds'@?chks'@?upds')) ?rcvs"
using wf⇩s⇩s⇩t_vars_mono'[OF _ p2(3)] by blast
have "list_all is_Send (unlabel (transaction_send T))"
using admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
unfolding wellformed_transaction_def by blast
hence "list_all is_Receive ?rcvs" by (metis dual⇩l⇩s⇩s⇩t_list_all(2))
hence r1: "wfrestrictedvars⇩s⇩s⇩t ?rcvs ⊆ X ∪ Y ∪ wfvarsoccs⇩s⇩s⇩t (?snds'@?chks'@?upds')"
using wfrestrictedvars⇩s⇩s⇩t_receives_only_eq wf⇩s⇩s⇩t_receives_only_fv_subset[OF r0] by blast
have "fv⇩s⇩e⇩t ((λx. occurs (Var x)) ` set (transaction_fresh T)) ⊆ Y"
unfolding Y by auto
hence r2: "wfrestrictedvars⇩s⇩s⇩t ?rcvs' ⊆ X ∪ Y ∪ wfvarsoccs⇩s⇩s⇩t (?snds'@?chks'@?upds')"
using 1(14) r1 unfolding wfrestrictedvars⇩s⇩s⇩t_def by fastforce
have r3: "wf'⇩s⇩s⇩t (X ∪ Y) (?snds'@?chks'@?upds')"
proof -
have *: "wf'⇩s⇩s⇩t (X ∪ Y) (?snds@?chks'@?upds')"
using T_wf wf⇩s⇩s⇩t_prefix[of "X ∪ Y" "?snds@?chks@?upds" ?rcvs] p1
unfolding transaction_strand_unlabel_dual_unfold by simp
have "?snds' = ?snds ∨ (∃ts. ?snds' = send⟨ts⟩#?snds)" using 1(13) by auto
thus ?thesis
proof
assume "?snds' = ?snds" thus ?thesis using * by simp
next
assume "∃ts. ?snds' = send⟨ts⟩#?snds"
then obtain ts where "?snds' = send⟨ts⟩#?snds" by blast
thus ?thesis using wf⇩s⇩s⇩t_sends_only_prepend[OF *, of "[send⟨ts⟩]"] by simp
qed
qed
have "wf'⇩s⇩s⇩t (X ∪ Y) (?snds'@?chks'@?upds'@?rcvs')"
using wf⇩s⇩s⇩t_append_suffix''[OF r3] r2 by auto
thus ?thesis
using unlabel_append dual⇩l⇩s⇩s⇩t_append
unfolding transaction_strand_def by auto
qed
have T'_props_1: "?P1 ?T'"
unfolding wellformed_transaction_def
apply (intro conjI)
subgoal using 1(13) T_props(1) unfolding wellformed_transaction_def by force
subgoal using 1(8) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using 1(9) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using 1(14) T_props(1) unfolding wellformed_transaction_def by force
subgoal using 1(4) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using 1(5) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using 1(4,5) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using T_props(1) unfolding 2(1) 1(5) wellformed_transaction_def by blast
subgoal using 1(5,8) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using 1(5) 2(4) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using 2(3,4) T_props(1) unfolding wellformed_transaction_def by simp
subgoal using 1(4,5) 5 T_props(1) unfolding wellformed_transaction_def by simp
done
have T'_props_2_12:
"?P2 ?T'" "?P3 ?T'" "?P4 ?T'" "?P5 ?T'" "?P6 ?T'" "?P7 ?T'" "?P8 ?T'" "?P9 ?T'" "?P10 ?T'"
"?P11 ?T'" "?P12 ?T'"
subgoal using T_props(2) unfolding defs by force
subgoal using T_props(3) unfolding defs by force
subgoal using T_props(4) 2(5) by argo
subgoal using T_props(5) 2(4) by argo
subgoal using T_props(6) 1(5,8) 2(2) by auto
subgoal using T_props(7) 1(5,8) 2(3) by presburger
subgoal using T_props(8) 1(5,9) 2(1,2) by auto
subgoal using T_props(9) 1(8) 2(1) by auto
subgoal using T_props(10) 1(8) 2(1) by auto
subgoal using T_props(11) 1(8) unfolding admissible_transaction_checks_def by argo
subgoal using T_props(12) 1(9) unfolding admissible_transaction_updates_def by argo
done
have T'_props_13_aux:
"transaction_fresh ?T' = []" (is ?Q1)
"is_Send r" (is ?Q2)
"length (the_msgs r) = 1" (is ?Q3)
"is_Fun_Attack (hd (the_msgs r))" (is ?Q4)
when r: "r ∈ set (unlabel (transaction_strand (add_occurs_msgs T)))"
"∃f ∈ ⋃(funs_term ` trms⇩s⇩s⇩t⇩p r). is_Attack f" (is "?Q' (trms⇩s⇩s⇩t⇩p r)")
for r
proof -
note q0 = conjunct2[OF conjunct2[OF T_props(13)[unfolded defs']]]
let ?Q'' = "λts' F'. F = ⟨⋆, send⟨ts'⟩⟩#F'"
let ?f = "map (λx. occurs (Var x))"
let ?frsh = "transaction_fresh T"
let ?xs = "fv_list⇩s⇩s⇩t (unlabel (transaction_strand T))"
have q1: "r ≠ send⟨?f ?frsh⟩" "r ≠ receive⟨?f (filter (λx. x ∉ set ?frsh) ?xs)⟩"
"∀f ∈ ⋃(funs_term ` set (?f ?frsh)). ¬is_Attack f"
using r(2) by (fastforce,fastforce,simp)
have q2: "send⟨ts'⟩ ∈ set (unlabel (transaction_strand T))"
"r = send⟨?f ?frsh@ts'⟩ ∨ r ∈ set (unlabel (transaction_strand T))"
when "?Q'' ts' F'" for ts' F'
subgoal using that unfolding T transaction_strand_def by force
subgoal using that r(1) 4(2)[OF that] q1 unfolding T transaction_strand_def by fast
done
have q3: "?Q' (set ts')"
when r': "?Q'' ts' F'" "r ∉ set (unlabel (transaction_strand T))" for ts' F'
proof -
have "r = send⟨?f ?frsh@ts'⟩" using q2(2)[OF r'(1)] r'(2) by argo
thus ?thesis using r(2) by fastforce
qed
have q4: "r ∈ set (unlabel (transaction_strand T))" when "∄ts' F'. ?Q'' ts' F'"
using 4(4)[OF that] r(1) q1(1,2) by blast
have "∃r' ∈ set (unlabel (transaction_strand T)). ?Q' (trms⇩s⇩s⇩t⇩p r')"
when "?Q'' ts' F'" for ts' F'
apply (cases "r ∈ set (unlabel (transaction_strand T))")
subgoal using q2(2)[OF that] r(2) by metis
subgoal using q2(1)[OF that] q3[OF that] trms⇩s⇩s⇩t⇩p.simps(1)[of ts'] by metis
done
hence "?frsh = []" when "?Q'' ts' F'" for ts' F' using q0 that by blast
hence "r = send⟨ts'⟩ ∨ r ∈ set (unlabel (transaction_strand T))" when "?Q'' ts' F'" for ts' F'
using q2(2)[OF that] that by blast
hence "r ∈ set (unlabel (transaction_strand T))" using q2(1) q4 by fast
thus ?Q1 ?Q2 ?Q3 ?Q4 using r(2) q0 unfolding 1(5) by auto
qed
have T'_props_13: "?P13 ?T'"
unfolding defs' 3
apply (intro conjI)
subgoal using conjunct1[OF T_props(13)[unfolded defs']] occurs_wf by fast
subgoal using conjunct1[OF conjunct2[OF T_props(13)[unfolded defs']]] occurs_funs by auto
subgoal using T'_props_13_aux by meson
done
have T'_props_14: "?P14 ?T'"
proof (cases "∃ts' F'. transaction_send T = ⟨⋆,send⟨ts'⟩⟩#F'")
case True
then obtain ts' F' where F': "transaction_send T = ⟨⋆,send⟨ts'⟩⟩#F'" by meson
show ?thesis
using T_props(14) 1(10)[OF F'] F' 1(5,12)
unfolding admissible_transaction_send_occurs_form_def Let_def
by (cases "transaction_fresh T = []") auto
next
case False show ?thesis
using T_props(14) 1(11)[OF False] 1(5,12)
unfolding admissible_transaction_send_occurs_form_def Let_def
by (cases "transaction_fresh T = []") auto
qed
let ?xs = "fv_list⇩s⇩s⇩t (unlabel (transaction_strand T))"
have T'_props_15: "?P15 ?T'"
using T_props(15) 1(6,7) unfolding Let_def
by (cases "filter (λx. x ∉ set (transaction_fresh T)) ?xs = []") (simp,fastforce)
have T'_props_16: "?P16 ?T'"
proof (cases "∃ts' F'. transaction_send T = ⟨⋆,send⟨ts'⟩⟩#F'")
case True
then obtain ts' F' where F': "transaction_send T = ⟨⋆,send⟨ts'⟩⟩#F'" by meson
show ?thesis
using T_props(16) 1(10)[OF F'] F' 1(5,12)
unfolding Let_def by (cases "transaction_fresh T = []") auto
next
case False show ?thesis
using T_props(16) 1(11)[OF False] 1(5,12)
unfolding Let_def by (cases "transaction_fresh T = []") auto
qed
note T'_props = T'_props_1 T'_props_2_12 T'_props_13 T'_props_14 T'_props_15 T'_props_16
show ?A using T'_props unfolding admissible_transaction'_def by meson
have 5: "set (filter (λx. x ∉ set (transaction_fresh T))
(fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))) =
fv_transaction T - set (transaction_fresh T)"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t by fastforce
have "transaction_receive ?T' ≠ []"
and "is_Receive (hd (unlabel (transaction_receive ?T')))"
and "∀x ∈ fv_transaction ?T' - set (transaction_fresh ?T'). fst x = TAtom Value ⟶
occurs (Var x) ∈ set (the_msgs (hd (unlabel (transaction_receive ?T'))))"
when x: "x ∈ fv_transaction ?T' - set (transaction_fresh ?T')" "fst x = TAtom Value" for x
using 1(13) 5 x unfolding 1(5) 2(3) by (force,force,force)
moreover have "transaction_send ?T' ≠ []" (is ?C)
and "is_Send (hd (unlabel (transaction_send ?T')))" (is ?D)
and "∀x ∈ set (transaction_fresh ?T').
occurs (Var x) ∈ set (the_msgs (hd (unlabel (transaction_send ?T'))))" (is ?E)
when T'_frsh: "transaction_fresh ?T' ≠ []"
using 1(14) T'_frsh unfolding 1(5) by auto
ultimately show ?B
using T'_props_14 unfolding admissible_transaction_occurs_checks_def Let_def by blast
qed
private lemma add_occurs_msgs_in_trms_subst_cases:
fixes T::"('fun,'atom,'sets,'lbl) prot_transaction"
assumes T_adm: "admissible_transaction' T"
and t: "t ∈ trms⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t θ)"
shows "t ∈ trms⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ) ∨
(∃x ∈ fv_transaction T. t = occurs (θ x))"
proof -
define frsh where "frsh ≡ transaction_fresh T"
define xs where "xs ≡ filter (λx. x ∉ set frsh) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
define f where "f ≡ map (λx. occurs (Var x)::('fun,'atom,'sets,'lbl) prot_term)"
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp
note T'_adm = add_occurs_msgs_admissible_occurs_checks(1)[OF T_adm]
have 0: "set (transaction_fresh T) ⊆ fv_transaction T"
using admissible_transactionE(7)[OF T_adm]
unfolding fv_transaction_unfold by blast
hence 00: "set (f xs) ∪ set (f frsh) = (λx. occurs (Var x)) ` fv_transaction T"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"]
unfolding f_def xs_def frsh_def by auto
note 1 = add_occurs_msgs_transaction_strand_set[OF T]
have 2: "set (transaction_strand (add_occurs_msgs T)) ⊆
set (transaction_strand T) ∪ {⟨⋆,receive⟨f xs⟩⟩,⟨⋆,send⟨f frsh⟩⟩}"
when "∄ts F'. F = ⟨⋆,send⟨ts⟩⟩#F'"
using 1(3,4)[OF that] unfolding f_def[symmetric] frsh_def[symmetric] xs_def[symmetric] by blast
have 3: "trms_transaction (add_occurs_msgs T) =
trms_transaction T ∪ (λx. occurs (Var x)) ` fv_transaction T"
using 0 add_occurs_msgs_trms_transaction[of T] by blast
have 4: "bvars_transaction T ∩ subst_domain θ = {}"
"bvars_transaction (add_occurs_msgs T) ∩ subst_domain θ = {}"
using admissible_transactionE(4)[OF T_adm] admissible_transactionE(4)[OF T'_adm]
by (blast,blast)
note 5 = trms⇩s⇩s⇩t_subst[OF 4(1), unfolded unlabel_subst]
trms⇩s⇩s⇩t_subst[OF 4(2), unfolded unlabel_subst]
note 6 = fv⇩s⇩s⇩t_is_subterm_trms⇩s⇩s⇩t_subst[
OF _ 4(1), unfolded add_occurs_msgs_admissible_occurs_checks(1)[OF T_adm] unlabel_subst]
show ?thesis
using t 6 unfolding 3 5 by fastforce
qed
private lemma add_occurs_msgs_updates_send_filter_iff:
fixes f
defines "f ≡ λT. list_ex (λa. is_Send (snd a) ∨ is_Update (snd a)) (transaction_strand T)"
and "g ≡ λT. transaction_fresh T = [] ⟶ f T"
shows "map add_occurs_msgs (filter g P) = filter g (map add_occurs_msgs P)"
proof -
have "g T ⟷ g (add_occurs_msgs T)" for T
proof -
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp_all
note 0 = add_occurs_msgs_cases[OF T]
show ?thesis using 0(6,7,12) unfolding g_def f_def transaction_strand_def 0(5,8,9) by fastforce
qed
thus ?thesis by (induct P) simp_all
qed
lemma add_occurs_msgs_updates_send_filter_iff':
fixes f
defines "f ≡ λT. list_ex (λa. is_Send (snd a) ∨ is_Update (snd a)) (transaction_strand T)"
and "g ≡ λT. transaction_fresh T = [] ⟶ transaction_updates T ≠ [] ∨ transaction_send T ≠ []"
shows "map add_occurs_msgs (filter g P) = filter g (map add_occurs_msgs P)"
proof -
have "g T ⟷ g (add_occurs_msgs T)" for T
proof -
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp_all
note 0 = add_occurs_msgs_cases[OF T]
show ?thesis using 0(6,7,12) unfolding g_def f_def transaction_strand_def 0(5,8,9) by argo
qed
thus ?thesis by (induct P) simp_all
qed
private lemma rm_occurs_msgs_constr_Cons:
defines "f ≡ rm_occurs_msgs_constr"
shows
"¬is_Receive a ⟹ ¬is_Send a ⟹ f ((l,a)#A) = (l,a)#f A"
"is_Receive a ⟹ ∄t. occurs t ∈ set (the_msgs a) ⟹ f ((l,a)#A) = (l,a)#f A"
"is_Receive a ⟹ ∃t. occurs t ∈ set (the_msgs a) ⟹
∃t ∈ set (the_msgs a). ∀s. t ≠ occurs s ⟹
f ((l,a)#A) = (l,receive⟨filter (λt. ∀s. t ≠ occurs s) (the_msgs a)⟩)#f A"
"is_Receive a ⟹ ∃t. occurs t ∈ set (the_msgs a) ⟹
∀t ∈ set (the_msgs a). ∃s. t = occurs s ⟹ f ((l,a)#A) = f A"
"is_Send a ⟹ ∄t. occurs t ∈ set (the_msgs a) ⟹ f ((l,a)#A) = (l,a)#f A"
"is_Send a ⟹ ∃t. occurs t ∈ set (the_msgs a) ⟹
∃t ∈ set (the_msgs a). ∀s. t ≠ occurs s ⟹
f ((l,a)#A) = (l,send⟨filter (λt. ∀s. t ≠ occurs s) (the_msgs a)⟩)#f A"
"is_Send a ⟹ ∃t. occurs t ∈ set (the_msgs a) ⟹
∀t ∈ set (the_msgs a). ∃s. t = occurs s ⟹ f ((l,a)#A) = f A"
unfolding f_def by (cases a; auto)+
private lemma rm_occurs_msgs_constr_Cons':
defines "f ≡ rm_occurs_msgs_constr"
and "g ≡ filter (λt. ∀s. t ≠ occurs s)"
assumes a: "is_Receive a ∨ is_Send a"
shows
"∄t. occurs t ∈ set (the_msgs a) ⟹ f ((l,a)#A) = (l,a)#f A"
"∃t. occurs t ∈ set (the_msgs a) ⟹
∃t ∈ set (the_msgs a). ∀s. t ≠ occurs s ⟹
is_Send a ⟹ f ((l,a)#A) = (l,send⟨g (the_msgs a)⟩)#f A"
"∃t. occurs t ∈ set (the_msgs a) ⟹
∃t ∈ set (the_msgs a). ∀s. t ≠ occurs s ⟹
is_Receive a ⟹ f ((l,a)#A) = (l,receive⟨g (the_msgs a)⟩)#f A"
"∃t. occurs t ∈ set (the_msgs a) ⟹
∀t ∈ set (the_msgs a). ∃s. t = occurs s ⟹ f ((l,a)#A) = f A"
using a unfolding f_def g_def by (cases a; auto)+
private lemma rm_occurs_msgs_constr_Cons'':
defines "f ≡ rm_occurs_msgs_constr"
and "g ≡ filter (λt. ∀s. t ≠ occurs s)"
assumes a: "a = receive⟨ts⟩ ∨ a = send⟨ts⟩"
shows "f ((l,a)#A) = (l,a)#f A ∨ f ((l,a)#A) = (l,receive⟨g ts⟩)#f A ∨
f ((l,a)#A) = (l,send⟨g ts⟩)#f A ∨ f ((l,a)#A) = f A"
using rm_occurs_msgs_constr_Cons(2-)[of a l A] a unfolding f_def g_def by (cases a) auto
private lemma rm_occurs_msgs_constr_ik_subset:
"ik⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A) ⊆ ik⇩l⇩s⇩s⇩t A"
proof (induction A)
case (Cons a A)
let ?f = "filter (λt. ∀s. t ≠ occurs s)"
note IH = Cons.IH
obtain l b where a: "a = (l,b)" by (metis surj_pair)
have 0: "set (unlabel A) ⊆ set (unlabel (a#A))" by auto
note 1 = rm_occurs_msgs_constr_Cons[of b l A]
note 2 = in_ik⇩l⇩s⇩s⇩t_iff
note 3 = ik⇩s⇩s⇩t_set_subset[OF 0]
note 4 = ik⇩s⇩s⇩t_append
note 5 = 4[of "unlabel [a]" "unlabel A"] 4[of "unlabel [a]" "unlabel (rm_occurs_msgs_constr A)"]
show ?case
proof (cases "is_Send b ∨ is_Receive b")
case True
note b_cases = this
define ts where "ts ≡ the_msgs b"
have ts_cases: "is_Send b ⟹ b = send⟨ts⟩" "is_Receive b ⟹ b = receive⟨ts⟩"
unfolding ts_def by simp_all
have 6:
"is_Send b ⟹ ik⇩l⇩s⇩s⇩t [(l,b)] = {}"
"is_Send b ⟹ ik⇩l⇩s⇩s⇩t [(l,send⟨the_msgs b⟩)] = {}"
"is_Send b ⟹ ik⇩l⇩s⇩s⇩t [(l,send⟨?f (the_msgs b)⟩)] = {}"
"is_Receive b ⟹ ik⇩l⇩s⇩s⇩t [(l,b)] = set ts"
"is_Receive b ⟹ ik⇩l⇩s⇩s⇩t [(l,receive⟨the_msgs b⟩)] = set ts"
"is_Receive b ⟹ ik⇩l⇩s⇩s⇩t [(l,receive⟨?f (the_msgs b)⟩)] = set (?f ts)"
using 2[of _ "[(l, send⟨the_msgs b⟩)]"]
2[of _ "[(l, send⟨?f (the_msgs b)⟩)]"]
2[of _ "[(l, receive⟨the_msgs b⟩)]"]
2[of _ "[(l, receive⟨?f (the_msgs b)⟩)]"]
b_cases ts_cases
by auto
have "ik⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr (a#A)) = ik⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A)"
when b: "is_Send b"
proof (cases "∃t. occurs t ∈ set (the_msgs b)")
case True
note 7 = 1(6,7)[OF b True]
show ?thesis
proof (cases "∃t ∈ set (the_msgs b). ∀s. t ≠ occurs s")
case True show ?thesis
using 4[of "unlabel [(l,send⟨?f (the_msgs b)⟩)]"
"unlabel (rm_occurs_msgs_constr A)"]
unfolding a 7(1)[OF True] 6(3)[OF b] by simp
next
case False
hence F: "∀t ∈ set (the_msgs b). ∃s. t = occurs s" by simp
show ?thesis
using 4[of "unlabel [(l,send⟨the_msgs b⟩)]" "unlabel (rm_occurs_msgs_constr A)"]
unfolding a 7(2)[OF F] 6(2)[OF b] by simp
qed
next
case False show ?thesis
using 4[of "unlabel [(l,b)]" "unlabel (rm_occurs_msgs_constr A)"]
unfolding a 1(5)[OF b False] 6(1)[OF b] by auto
qed
moreover have "ik⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr (a#A)) ⊆ set ts ∪ ik⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A)"
when b: "is_Receive b"
proof (cases "∃t. occurs t ∈ set (the_msgs b)")
case True
note 8 = 1(3,4)[OF b True]
show ?thesis
proof (cases "∃t ∈ set (the_msgs b). ∀s. t ≠ occurs s")
case True show ?thesis
using 4[of "unlabel [(l,receive⟨?f (the_msgs b)⟩)]"
"unlabel (rm_occurs_msgs_constr A)"]
unfolding a 8(1)[OF True] 6(6)[OF b] by auto
next
case False
hence F: "∀t ∈ set (the_msgs b). ∃s. t = occurs s" by simp
show ?thesis
using 4[of "unlabel [(l,receive⟨the_msgs b⟩)]" "unlabel (rm_occurs_msgs_constr A)"]
unfolding a 8(2)[OF F] 6(5)[OF b] by simp
qed
next
case False show ?thesis
using 4[of "unlabel [(l,b)]" "unlabel (rm_occurs_msgs_constr A)"]
unfolding a 1(2)[OF b False] 6(4)[OF b] by auto
qed
moreover have "ik⇩l⇩s⇩s⇩t (a#A) = set ts ∪ ik⇩l⇩s⇩s⇩t A" when b: "is_Receive b"
using ik⇩l⇩s⇩s⇩t_Cons(2)[of l ts A] unfolding a ts_cases(2)[OF b] by blast
ultimately show ?thesis using IH 3 b_cases by blast
qed (use 1(1) IH 5 a in auto)
qed simp
private lemma rm_occurs_msgs_constr_append:
"rm_occurs_msgs_constr (A@B) = rm_occurs_msgs_constr A@rm_occurs_msgs_constr B"
by (induction A rule: rm_occurs_msgs_constr.induct) auto
private lemma rm_occurs_msgs_constr_dual⇩l⇩s⇩s⇩t:
"rm_occurs_msgs_constr (dual⇩l⇩s⇩s⇩t A) = dual⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A)"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case using Cons.IH unfolding a by (cases b) auto
qed simp
private lemma rm_occurs_msgs_constr_dbupd⇩s⇩s⇩t_eq:
"dbupd⇩s⇩s⇩t (unlabel (rm_occurs_msgs_constr A)) I D = dbupd⇩s⇩s⇩t (unlabel A) I D"
proof (induction A arbitrary: I D)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case
proof (cases "is_Receive b ∨ is_Send b")
case True
then obtain ts where b: "b = receive⟨ts⟩ ∨ b = send⟨ts⟩" by (cases b) simp_all
show ?thesis using rm_occurs_msgs_constr_Cons''[OF b, of l A] Cons.IH b unfolding a by fastforce
next
case False thus ?thesis using Cons.IH unfolding a by (cases b) auto
qed
qed simp
private lemma rm_occurs_msgs_constr_subst:
fixes A::"('a,'b,'c,'d) prot_strand" and θ::"('a,'b,'c,'d) prot_subst"
assumes "∀x ∈ fv⇩l⇩s⇩s⇩t A. ∄t. θ x = occurs t" "∀x ∈ fv⇩l⇩s⇩s⇩t A. θ x ≠ Fun OccursSec []"
shows "rm_occurs_msgs_constr (A ⋅⇩l⇩s⇩s⇩t θ) = (rm_occurs_msgs_constr A) ⋅⇩l⇩s⇩s⇩t θ"
(is "?f (A ⋅⇩l⇩s⇩s⇩t θ) = (?f A) ⋅⇩l⇩s⇩s⇩t θ")
using assms
proof (induction A)
case (Cons a A)
note 0 = rm_occurs_msgs_constr_Cons
note 1 = rm_occurs_msgs_constr_Cons'
define f where "f ≡ ?f"
define not_occ where "not_occ ≡ λt::('a,'b,'c,'d) prot_term. ∀s. t ≠ occurs s"
define flt where "flt ≡ filter not_occ"
obtain l b where a: "a = (l,b)" by (metis surj_pair)
have 2: "∄t. θ x = occurs t" "θ x ≠ Fun OccursSec []"
when b: "is_Receive b ∨ is_Send b" and t: "t ∈ set (the_msgs b)" and x: "x ∈ fv t" for x t
using Cons.prems x t b unfolding a by (cases b; auto)+
have IH: "f (A ⋅⇩l⇩s⇩s⇩t θ) = (f A) ⋅⇩l⇩s⇩s⇩t θ"
using Cons.prems Cons.IH unfolding f_def by simp
show ?case
proof (cases "is_Receive b ∨ is_Send b")
case True
note T = this
then obtain ts where ts: "b = receive⟨ts⟩ ∨ b = send⟨ts⟩" by (cases b) simp_all
hence ts': "b ⋅⇩s⇩s⇩t⇩p θ = receive⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩ ∨ b ⋅⇩s⇩s⇩t⇩p θ = send⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩" by auto
have the_msgs_b: "the_msgs b = ts" "the_msgs (b ⋅⇩s⇩s⇩t⇩p θ) = ts ⋅⇩l⇩i⇩s⇩t θ"
using ts ts' by auto
have 4: "is_Receive (b ⋅⇩s⇩s⇩t⇩p θ) ∨ is_Send (b ⋅⇩s⇩s⇩t⇩p θ)"
using T by (cases b) simp_all
note 6 = 1[OF T, of l A, unfolded f_def[symmetric]]
note 7 = 1[OF 4, of l "A ⋅⇩l⇩s⇩s⇩t θ", unfolded f_def[symmetric]]
note 8 = ts IH subst_lsst_cons[of _ _ θ]
have 9: "t ⋅ θ ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ))" "not_occ (t ⋅ θ)"
when t: "t ∈ set (the_msgs b)" "not_occ t" for t
proof -
show "t ⋅ θ ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ))" using t ts ts' by auto
moreover have "not_occ (t ⋅ θ)" when "t = Var x" for x
using 2[OF T t(1)] t(2) unfolding that not_occ_def by simp
moreover have "not_occ (t ⋅ θ)" when "t = Fun g ss" "g ≠ OccursFact" for g ss
using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by simp
moreover have "not_occ (t ⋅ θ)"
when "t = Fun OccursFact ss" "∄s1 s2. ss = [s1,s2]" for ss
using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by auto
moreover have "not_occ (t ⋅ θ)"
when "t = Fun OccursFact [s1,s2]" for s1 s2
using 2[OF T t(1)] t(2) unfolding that not_occ_def by (cases s1) auto
ultimately show "not_occ (t ⋅ θ)" by (cases t) (metis, metis)
qed
have 10: "not_occ t"
when t: "t ∈ set (the_msgs b)" "not_occ (t ⋅ θ)" for t
proof -
have "t ⋅ θ ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ))" using t ts ts' by auto
moreover have "not_occ t" when "t = Var x" for x
using 2[OF T t(1)] t(2) unfolding that not_occ_def by simp
moreover have "not_occ t" when "t = Fun g ss" "g ≠ OccursFact" for g ss
using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by simp
moreover have "not_occ t"
when "t = Fun OccursFact ss" "∄s1 s2. ss = [s1,s2]" for ss
using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by auto
moreover have "not_occ t"
when "t = Fun OccursFact [s1,s2]" for s1 s2
using 2[OF T t(1)] t(2) unfolding that not_occ_def by (cases s1) auto
ultimately show "not_occ t" unfolding not_occ_def by force
qed
have 11: "not_occ (t ⋅ θ) ⟷ not_occ t" when "t ∈ set ts" for t
using that 9 10 unfolding the_msgs_b by blast
have 5: "(∃t. occurs t ∈ set ts) ⟷ (∃t. occurs t ∈ set ts ⋅⇩s⇩e⇩t θ)"
using 11 image_iff unfolding not_occ_def by fastforce
have 12: "flt (ts ⋅⇩l⇩i⇩s⇩t θ) = (flt ts) ⋅⇩l⇩i⇩s⇩t θ" using 11
proof (induction ts)
case (Cons t ts)
hence "not_occ (t ⋅ θ) = not_occ t" "flt (ts ⋅⇩l⇩i⇩s⇩t θ) = (flt ts) ⋅⇩l⇩i⇩s⇩t θ" by auto
thus ?case unfolding flt_def by auto
qed (metis flt_def filter.simps(1) map_is_Nil_conv)
show ?thesis
proof (cases "∃t. occurs t ∈ set (the_msgs b)")
case True
note T1 = this
hence T2: "∃t. occurs t ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ))" using 5 unfolding the_msgs_b by simp
show ?thesis
proof (cases "∃t ∈ set (the_msgs b). ∀s. t ≠ occurs s")
case True
note T1' = this
have T2': "∃t ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ)). ∀s. t ≠ occurs s"
using T1' 11 unfolding the_msgs_b not_occ_def by auto
show ?thesis using T
proof
assume b: "is_Receive b"
hence bθ: "is_Receive (b ⋅⇩s⇩s⇩t⇩p θ)" using ts by fastforce
show ?thesis
using 6(3)[OF T1 T1' b] 7(3)[OF T2 T2' bθ] IH 12
unfolding f_def[symmetric] a flt_def[symmetric] not_occ_def[symmetric] the_msgs_b
by (simp add: subst_lsst_cons)
next
assume b: "is_Send b"
hence bθ: "is_Send (b ⋅⇩s⇩s⇩t⇩p θ)" using ts by fastforce
show ?thesis
using 6(2)[OF T1 T1' b] 7(2)[OF T2 T2' bθ] IH 12
unfolding f_def[symmetric] a flt_def[symmetric] not_occ_def[symmetric] the_msgs_b
by (simp add: subst_lsst_cons)
qed
next
case False
hence F: "∀t ∈ set (the_msgs b). ∃s. t = occurs s" by blast
hence F': "∀t ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ)). ∃s. t = occurs s" unfolding the_msgs_b by auto
have *: "∃t. occurs t ∈ set (the_msgs b)" when "the_msgs b ≠ []"
using that F by (cases "the_msgs b") auto
hence **: "∃t. occurs t ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ))" when "the_msgs b ≠ []"
using that 5 unfolding the_msgs_b by simp
show ?thesis
proof (cases "ts = []")
case True
hence ***: "∄t. occurs t ∈ set (the_msgs b)" "∄t. occurs t ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ))"
unfolding the_msgs_b by simp_all
show ?thesis
using IH 6(1)[OF ***(1)] 7(1)[OF ***(2)]
unfolding a f_def[symmetric] True
by (simp add: subst_lsst_cons)
next
case False thus ?thesis
using IH 6(4)[OF * F] 7(4)[OF ** F']
unfolding f_def[symmetric] not_occ_def[symmetric] a the_msgs_b
by (simp add: subst_lsst_cons)
qed
qed
next
case False
note F = this
have F': "∄t. occurs t ∈ set (the_msgs (b ⋅⇩s⇩s⇩t⇩p θ))"
using F 11 unfolding not_occ_def the_msgs_b by fastforce
show ?thesis
using IH 6(1)[OF F] 7(1)[OF F']
unfolding a f_def[symmetric] True
by (simp add: subst_lsst_cons)
qed
next
case False
hence *: "¬is_Receive b" "¬is_Send b" "¬is_Receive (b ⋅⇩s⇩s⇩t⇩p θ)" "¬is_Send (b ⋅⇩s⇩s⇩t⇩p θ)"
by (cases b; auto)+
show ?thesis
using IH 0(1)[OF *(1,2), of l A] 0(1)[OF *(3,4), of l "A ⋅⇩l⇩s⇩s⇩t θ"] subst_lsst_cons[of a _ θ]
unfolding a f_def by simp
qed
qed simp
private lemma rm_occurs_msgs_constr_transaction_strand:
assumes T_adm: "admissible_transaction' T"
shows "rm_occurs_msgs_constr (transaction_checks T) = transaction_checks T" (is ?A)
and "rm_occurs_msgs_constr (transaction_updates T) = transaction_updates T" (is ?B)
and "admissible_transaction_no_occurs_msgs T ⟹
rm_occurs_msgs_constr (transaction_receive T) = transaction_receive T" (is "?C ⟹ ?C'")
and "admissible_transaction_no_occurs_msgs T ⟹
rm_occurs_msgs_constr (transaction_send T) = transaction_send T" (is "?D ⟹ ?D'")
proof -
note 0 = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
note 1 = wellformed_transaction_cases[OF 0]
have 2: "∃ts. b = receive⟨ts⟩ ∧ (∄t. occurs t ∈ set ts)"
when "admissible_transaction_no_occurs_msgs T" "(l,b) ∈ set (transaction_receive T)" for l b
using that 1(1)[OF that(2)]
unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by fastforce
have 3: "∃ts. b = send⟨ts⟩ ∧ (∄t. occurs t ∈ set ts)"
when "admissible_transaction_no_occurs_msgs T" "(l,b) ∈ set (transaction_send T)" for l b
using that 1(4)[OF that(2)]
unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by fastforce
define A where "A ≡ transaction_receive T"
define B where "B ≡ transaction_checks T"
define C where "C ≡ transaction_updates T"
define D where "D ≡ transaction_send T"
show ?A using 1(2) unfolding B_def[symmetric]
proof (induction B)
case (Cons a A)
hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2))
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case using Cons.prems IH unfolding a by (cases b) auto
qed simp
show ?B using 1(3) unfolding C_def[symmetric]
proof (induction C)
case (Cons a A)
hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2))
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case using Cons.prems IH unfolding a by (cases b) auto
qed simp
show ?C' when ?C using 2[OF that] unfolding A_def[symmetric]
proof (induction A)
case (Cons a A)
hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2))
obtain l b where a: "a = (l,b)" by (metis surj_pair)
obtain ts where b: "b = receive⟨ts⟩" using Cons.prems unfolding a by auto
show ?case using Cons.prems IH unfolding a b by fastforce
qed simp
show ?D' when ?D using 3[OF that] unfolding D_def[symmetric]
proof (induction D)
case (Cons a A)
hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2))
obtain l b where a: "a = (l,b)" by (metis surj_pair)
obtain ts where b: "b = send⟨ts⟩" using Cons.prems unfolding a by auto
show ?case using Cons.prems IH unfolding a b by fastforce
qed simp
qed
private lemma rm_occurs_msgs_constr_transaction_strand':
fixes T::"('fun,'atom,'sets,'lbl) prot_transaction"
assumes T_adm: "admissible_transaction' T"
and T_no_occ: "admissible_transaction_no_occurs_msgs T"
shows "rm_occurs_msgs_constr (transaction_strand (add_occurs_msgs T)) = transaction_strand T"
(is "?f (?g (?h T)) = ?g T")
proof -
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp
have B: "B = transaction_fresh T" unfolding T by simp
have F: "F = transaction_send T" unfolding T by simp
define xs where "xs ≡ filter (λx. x ∉ set B) (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
note 0 = rm_occurs_msgs_constr_transaction_strand
note 1 = add_occurs_msgs_admissible_occurs_checks[OF T_adm]
note 2 = 0(3,4)[OF T_adm T_no_occ]
note 3 = add_occurs_msgs_cases[OF T]
note 4 = 0(1,2)[OF 1(1)]
have 5: "?f (transaction_checks (?h T)) = transaction_checks T"
"?f (transaction_updates (?h T)) = transaction_updates T"
using 4 3(8,9) by (argo, argo)
have 6: "?f (transaction_receive (?h T)) = transaction_receive T"
proof (cases "xs = []")
case True show ?thesis using 3(6)[OF True[unfolded xs_def B]] 2(1) by simp
next
case False show ?thesis
using False 3(7)[OF False[unfolded xs_def B]] 2(1)
rm_occurs_msgs_constr_Cons(4)[
of "receive⟨map (λx. occurs (Var x)) xs⟩" ⋆ "transaction_receive T"]
unfolding B[symmetric] xs_def[symmetric]
by (cases xs) (blast, auto)
qed
have 7: "?f (transaction_send (?h T)) = transaction_send T"
proof (cases "∃ts' F'. F = ⟨⋆, send⟨ts'⟩⟩#F'")
case True
then obtain ts' F' where F': "F = ⟨⋆, send⟨ts'⟩⟩#F'" by blast
have *: "transaction_send (?h T) = ⟨⋆, send⟨map (λx. occurs (Var x)) B@ts'⟩⟩#F'"
using 3(1)[OF F'] unfolding T by fastforce
have **: "ts' ≠ []" using admissible_transactionE(17)[OF T_adm] unfolding T F' by auto
have ***: "∀s. t ≠ occurs s" when t: "t ∈ set ts'" for t
using that T_no_occ
unfolding T F' admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by auto
let ?ts = "map (λx. occurs (Var x)) B@ts'"
have "∃t ∈ set ?ts. ∀s. t ≠ occurs s" using ** *** by (cases ts') auto
moreover have "filter (λt. ∀s. t ≠ occurs s) ?ts = ts'" using *** by simp
moreover have "∃t. occurs t ∈ set ?ts" when "B ≠ []" using that by (cases B) auto
moreover have "?f [⟨⋆, send⟨ts'⟩⟩] = [⟨⋆, send⟨ts'⟩⟩]"
using 2(2) ** *** unfolding F[symmetric] F' by force
hence "?f F' = F'"
using 2(2) rm_occurs_msgs_constr_append[of "[⟨⋆, send⟨ts'⟩⟩]" F']
unfolding F[symmetric] F' by fastforce
ultimately have "?f (⟨⋆, send⟨?ts⟩⟩#F') = ⟨⋆, send⟨ts'⟩⟩#F'"
using 2(2) 3(10)[OF F'[unfolded F]] 3(12)
rm_occurs_msgs_constr.simps(3)[of ⋆ ts' F']
rm_occurs_msgs_constr_append[of "[⟨⋆, send⟨ts'⟩⟩]" F']
unfolding F[symmetric] F' B[symmetric] by auto
thus ?thesis using F * unfolding F' by argo
next
case False show ?thesis
using 3(2)[OF False] 3(3) 2(2)
unfolding B[symmetric] xs_def[symmetric] F[symmetric]
by (cases B) auto
qed
show ?thesis
using 5 6 7 rm_occurs_msgs_constr_append
unfolding transaction_strand_def by metis
qed
private lemma rm_occurs_msgs_constr_transaction_strand'':
fixes T::"('fun,'atom,'sets,'lbl) prot_transaction"
assumes T_adm: "admissible_transaction' T"
and T_no_occ: "admissible_transaction_no_occurs_msgs T"
and θ: "∀x ∈ fv_transaction (add_occurs_msgs T). ∄t. θ x = occurs t"
"∀x ∈ fv_transaction (add_occurs_msgs T). θ x ≠ Fun OccursSec []"
shows "rm_occurs_msgs_constr (dual⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t θ)) =
dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"
using rm_occurs_msgs_constr_dual⇩l⇩s⇩s⇩t[of "transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t θ"]
rm_occurs_msgs_constr_subst[OF θ] rm_occurs_msgs_constr_transaction_strand'[OF T_adm T_no_occ]
by argo
private lemma rm_occurs_msgs_constr_bvars_subst_eq:
"bvars⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A ⋅⇩l⇩s⇩s⇩t θ) = bvars⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)"
proof -
have "bvars⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A) = bvars⇩l⇩s⇩s⇩t A"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case using Cons.IH unfolding a by (cases b) auto
qed simp
thus ?thesis by (metis bvars⇩s⇩s⇩t_subst unlabel_subst)
qed
private lemma rm_occurs_msgs_constr_reachable_constraints_fv_eq:
assumes P: "∀T ∈ set P. admissible_transaction' T"
"∀T ∈ set P. admissible_transaction_no_occurs_msgs T"
and A: "A ∈ reachable_constraints (map add_occurs_msgs P)"
shows "fv⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A) = fv⇩l⇩s⇩s⇩t A"
using A
proof (induction A rule: reachable_constraints.induct)
case (step 𝒜 T ξ σ α)
let ?f = rm_occurs_msgs_constr
let ?B = "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
obtain T' where T': "T' ∈ set P" "T = add_occurs_msgs T'"
using step.hyps(2) by auto
have T_adm: "admissible_transaction' T"
using add_occurs_msgs_admissible_occurs_checks(1) step.hyps(2) P by auto
have T'_adm: "admissible_transaction' T'"
and T'_no_occ: "admissible_transaction_no_occurs_msgs T'"
using T'(1) P by (blast,blast)
have "?f (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)) = dual⇩l⇩s⇩s⇩t (transaction_strand T' ⋅⇩l⇩s⇩s⇩t θ)"
using rm_occurs_msgs_constr_transaction_strand''[OF T'_adm T'_no_occ, of θ]
admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)]
unfolding T'(2) θ_def[symmetric] by blast
moreover have "fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ) = fv⇩l⇩s⇩s⇩t (transaction_strand T' ⋅⇩l⇩s⇩s⇩t θ)"
using add_occurs_msgs_vars_eq(6)[OF T'_adm, of θ] unfolding T'(2) by blast
ultimately have "fv⇩l⇩s⇩s⇩t (?f ?B) = fv⇩l⇩s⇩s⇩t ?B"
using fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unfolding T'(2) θ_def[symmetric] by metis
thus ?case
using step.IH fv⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel ?B"]
rm_occurs_msgs_constr_append[of 𝒜 ?B]
by force
qed simp
private lemma rm_occurs_msgs_constr_reachable_constraints_vars_eq:
assumes P: "∀T ∈ set P. admissible_transaction' T"
"∀T ∈ set P. admissible_transaction_no_occurs_msgs T"
and A: "A ∈ reachable_constraints (map add_occurs_msgs P)"
shows "vars⇩l⇩s⇩s⇩t (rm_occurs_msgs_constr A) = vars⇩l⇩s⇩s⇩t A"
using rm_occurs_msgs_constr_bvars_subst_eq[of _ Var]
rm_occurs_msgs_constr_reachable_constraints_fv_eq[OF P A]
by (metis vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t subst_lsst_id_subst)
private lemma rm_occurs_msgs_constr_reachable_constraints_trms_cases_aux:
assumes A: "x ∈ fv⇩s⇩s⇩t A" "bvars⇩s⇩s⇩t A = {}"
and t: "t = occurs (θ x)"
and θ: "(∃y. θ x = Var y) ∨ (∃c. θ x = Fun c [])"
shows "(∃x ∈ fv⇩s⇩s⇩t (A ⋅⇩s⇩s⇩t θ). t = occurs (Var x)) ∨
(∃c. Fun c [] ⊑⇩s⇩e⇩t trms⇩s⇩s⇩t (A ⋅⇩s⇩s⇩t θ) ∧ t = occurs (Fun c []))"
using A
proof (induction A)
case (Cons a A)
have 0: "bvars⇩s⇩s⇩t A = {}" "set (bvars⇩s⇩s⇩t⇩p a) = {}" "set (bvars⇩s⇩s⇩t⇩p a) ∩ subst_domain θ = {}"
using Cons.prems(2) by auto
note 1 = fv⇩s⇩s⇩t_Cons[of a A] trms⇩s⇩s⇩t_cons[of a A] subst_sst_cons[of a A θ]
show ?case
proof (cases "x ∈ fv⇩s⇩s⇩t A")
case False
hence x: "x ∈ fv⇩s⇩s⇩t⇩p a" using Cons.prems(1) by simp
note 2 = x t θ
have 3: "θ x ⊑⇩s⇩e⇩t trms⇩s⇩s⇩t⇩p (a ⋅⇩s⇩s⇩t⇩p θ)"
using subst_subterms[OF fv⇩s⇩s⇩t⇩p_is_subterm_trms⇩s⇩s⇩t⇩p[OF x]] trms⇩s⇩s⇩t⇩p_subst[OF 0(3)] by auto
have "Fun c [] ⊑⇩s⇩e⇩t trms⇩s⇩s⇩t⇩p (a ⋅⇩s⇩s⇩t⇩p θ)" when "θ x = Fun c []" for c
using that 3 t by argo
moreover have "y ∈ fv⇩s⇩s⇩t⇩p (a ⋅⇩s⇩s⇩t⇩p θ)" when "θ x = Var y" for y
using that 3 var_subterm_trms⇩s⇩s⇩t⇩p_is_vars⇩s⇩s⇩t⇩p[of y "a ⋅⇩s⇩s⇩t⇩p θ"] 0(2)
unfolding vars⇩s⇩s⇩t⇩p_is_fv⇩s⇩s⇩t⇩p_bvars⇩s⇩s⇩t⇩p bvars⇩s⇩s⇩t⇩p_subst by simp
ultimately have
"(∃x ∈ fv⇩s⇩s⇩t⇩p (a ⋅⇩s⇩s⇩t⇩p θ). t = occurs (Var x)) ∨
(∃c. Fun c [] ⊑⇩s⇩e⇩t trms⇩s⇩s⇩t⇩p (a ⋅⇩s⇩s⇩t⇩p θ) ∧ t = occurs (Fun c []))"
using t θ by fast
thus ?thesis using 1 by auto
qed (use Cons.IH[OF _ 0(1)] 1 in force)
qed simp
private lemma rm_occurs_msgs_constr_reachable_constraints_trms_cases:
assumes P: "∀T ∈ set P. admissible_transaction' T"
"∀T ∈ set P. admissible_transaction_no_occurs_msgs T"
and A: "A = rm_occurs_msgs_constr B"
and B: "B ∈ reachable_constraints (map add_occurs_msgs P)"
and t: "t ∈ trms⇩l⇩s⇩s⇩t B"
shows "t ∈ trms⇩l⇩s⇩s⇩t A ∨ (∃x ∈ fv⇩l⇩s⇩s⇩t A. t = occurs (Var x)) ∨
(∃c. Fun c [] ⊑⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t A) ∧ t = occurs (Fun c []))"
(is "?A A ∨ ?B A ∨ ?C A")
proof -
define rm_occs where
"rm_occs ≡ λA::('fun,'atom,'sets,'lbl) prot_strand. rm_occurs_msgs_constr A"
define Q where "Q ≡ λA. ?A A ∨ ?B A ∨ ?C A"
have 0: "Q B" when "Q A" "set A ⊆ set B" for A B
using that unfolding Q_def fv⇩s⇩s⇩t_def trms⇩s⇩s⇩t_def unlabel_def by auto
have "Q A" using B t unfolding A
proof (induction rule: reachable_constraints.induct)
case (step 𝒜 T ξ σ α)
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
define ℬ where "ℬ ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"
obtain T' where T': "T' ∈ set P" "T = add_occurs_msgs T'"
using step.hyps(2) by auto
note T'_adm = bspec[OF P(1) T'(1)] bspec[OF P(2) T'(1)]
note T_adm = add_occurs_msgs_admissible_occurs_checks[OF T'_adm(1), unfolded T'(2)[symmetric]]
note 1 = θ_def[symmetric] ℬ_def[symmetric] rm_occs_def[symmetric]
note 2 = rm_occurs_msgs_constr_append[of 𝒜 ℬ, unfolded rm_occs_def[symmetric]]
note 3 = admissible_transaction_decl_fresh_renaming_subst_not_occurs[
OF T_adm(1) step.hyps(3,4,5)]
have 4: "rm_occs (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)) =
dual⇩l⇩s⇩s⇩t (transaction_strand T' ⋅⇩l⇩s⇩s⇩t θ)"
using 3 rm_occurs_msgs_constr_transaction_strand''[OF T'_adm]
unfolding T'(2) 1 by blast
have 5: "(∃y. θ x = Var y) ∨ (∃c. θ x = Fun c [])" for x
using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3,4,5)]
unfolding θ_def[symmetric] by blast
show ?case
proof (cases "t ∈ trms⇩l⇩s⇩s⇩t 𝒜")
case True show ?thesis using 0[OF step.IH[OF True]] unfolding 1 2 by simp
next
case False
hence "t ∈ trms⇩l⇩s⇩s⇩t ℬ" using step.prems unfolding ℬ_def θ_def by simp
hence "t ∈ trms⇩l⇩s⇩s⇩t (transaction_strand T' ⋅⇩l⇩s⇩s⇩t θ) ∨
(∃x ∈ fv_transaction T'. t = occurs (θ x))"
using add_occurs_msgs_in_trms_subst_cases[OF T'_adm(1), of t θ]
unfolding ℬ_def trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq T'(2) by blast
moreover have "(∃y. θ x = Var y) ∨ (∃c. θ x = Fun c [])" for x
using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3,4,5)]
unfolding θ_def[symmetric] by blast
ultimately have "Q (rm_occs ℬ)"
using rm_occurs_msgs_constr_reachable_constraints_trms_cases_aux[
of _ "unlabel (transaction_strand T')" t θ]
admissible_transactionE(4)[OF T'_adm(1)]
unfolding Q_def ℬ_def 4 trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unlabel_subst
by fast
thus ?thesis using 0[of "rm_occs ℬ"] unfolding 1 2 by auto
qed
qed simp
thus ?thesis unfolding Q_def by blast
qed
private lemma rm_occurs_msgs_constr_receive_attack_iff:
fixes A::"('a,'b,'c,'d) prot_strand"
shows "(∃ts. attack⟨n⟩ ∈ set ts ∧ receive⟨ts⟩ ∈ set (unlabel A)) ⟷
(∃ts. attack⟨n⟩ ∈ set ts ∧ receive⟨ts⟩ ∈ set (unlabel (rm_occurs_msgs_constr A)))"
(is "(∃ts. attack⟨n⟩ ∈ set ts ∧ ?A A ts) ⟷ (∃ts. attack⟨n⟩ ∈ set ts ∧ ?B A ts)")
proof
let ?att = "λts. attack⟨n⟩ ∈ set ts"
define f where "f ≡ λts::('a,'b,'c,'d) prot_term list. filter (λt. ∀s. t ≠ occurs s) ts"
have 0: "?att ts ⟷ ?att (f ts)"
"?att ts ⟹ ∃t. occurs t ∈ set ts ⟹ ∃t ∈ set ts. ∀s. t ≠ occurs s"
"∄t. occurs t ∈ set ts ⟹ f ts = ts"
for ts::"('a,'b,'c,'d) prot_term list"
unfolding f_def
subgoal by simp
subgoal by auto
subgoal by (induct ts) auto
done
have "?B A (f ts)" when A: "?A A ts" and ts: "?att ts" for ts using A
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case
proof (cases "?A A ts")
case True thus ?thesis using Cons.IH unfolding a by (cases b) simp_all
next
case False
hence b: "b = receive⟨ts⟩" using Cons.prems unfolding a by simp
show ?thesis using 0(2)[OF ts] 0(3) unfolding a b f_def by simp
qed
qed simp
thus "(∃ts. ?att ts ∧ ?A A ts) ⟹ (∃ts. ?att ts ∧ ?B A ts)" using 0(1) by fast
have "∃ts'. ts = f ts' ∧ ?A A ts'" when B: "?B A ts" and ts: "?att ts" for ts using B
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
note 1 = rm_occurs_msgs_constr_Cons
have 2: "receive⟨ts⟩ ∈ set (unlabel (rm_occurs_msgs_constr A))"
when rcv_ts: "receive⟨ts⟩ ∈ set (unlabel (rm_occurs_msgs_constr ((l,send⟨ts'⟩)#A)))"
for l ts ts' and A::"('a,'b,'c,'d) prot_strand"
proof -
have *: "is_Send (send⟨ts'⟩)" by simp
have "set (unlabel (rm_occurs_msgs_constr [(l, send⟨ts'⟩)])) ⊆ {send⟨ts'⟩, send⟨f ts'⟩}"
using 1(5-7)[OF *, of l "[]"] unfolding f_def by auto
thus ?thesis using rcv_ts rm_occurs_msgs_constr_append[of "[(l,send⟨ts'⟩)]" A] by auto
qed
show ?case
proof (cases "?B A ts")
case True thus ?thesis using Cons.IH by auto
next
case False
hence 3: "receive⟨ts⟩ ∈ set (unlabel (rm_occurs_msgs_constr [a]))"
using rm_occurs_msgs_constr_append[of "[a]" A] Cons.prems by simp
obtain ts' where b: "b = receive⟨ts'⟩" and b': "is_Receive b"
using 2[of ts l _ A] Cons.prems False
unfolding a by (cases b) auto
have ts': "the_msgs (receive⟨ts'⟩) = ts'" by simp
have "∃t ∈ set (the_msgs b). ∀s. t ≠ occurs s" when "∃t. occurs t ∈ set (the_msgs b)"
using that 3 1(4)[OF b' that, of l "[]"] unfolding a by force
hence "ts = f ts'"
using 3 0(3)[of ts'] 1(3)[OF b', of l "[]", unfolded rm_occurs_msgs_constr.simps(1)]
unfolding a b ts' f_def[symmetric] by fastforce
thus ?thesis unfolding a b by auto
qed
qed simp
thus "(∃ts. ?att ts ∧ ?B A ts) ⟹ (∃ts. ?att ts ∧ ?A A ts)" using 0 by fast
qed
private lemma add_occurs_msgs_soundness_aux1:
fixes P::"('fun,'atom,'sets,'lbl) prot"
defines "wt_attack ≡ λℐ 𝒜 l n. welltyped_constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])"
assumes P: "∀T ∈ set P. admissible_transaction' T"
and P_val: "has_initial_value_producing_transaction P"
and A: "𝒜 ∈ reachable_constraints P" "wt_attack ℐ 𝒜 l n"
shows "∃ℬ ∈ reachable_constraints P. ∃𝒥.
wt_attack 𝒥 ℬ l n ∧ (∀x ∈ fv⇩l⇩s⇩s⇩t ℬ. ∃n. 𝒥 x = Fun (Val n) [])"
proof -
let ?f = "λ(T,ξ,σ,α). dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
let ?g = "concat ∘ map ?f"
let ?rcv_att = "λA n. receive⟨[attack⟨n⟩]⟩ ∈ set (unlabel A)"
let ?wt_model = welltyped_constraint_model
define valconst_cases where "valconst_cases ≡
λI::('fun,'atom,'sets,'lbl) prot_subst. λx.
(∃n. I x = Fun (Val n) []) ∨ (∃n. I x = Fun (PubConst Value n) [])"
define valconsts_only where "valconsts_only ≡
λX. λI::('fun,'atom,'sets,'lbl) prot_subst. ∀x ∈ X. ∃n. I x = Fun (Val n) []"
define db_eq where "db_eq ≡
λA B::('fun,'atom,'sets,'lbl) prot_constr. λs. λupds::('fun,'atom,'sets,'lbl) prot_constr.
let f = filter is_Update ∘ unlabel;
g = filter (λa. ∄l t ts. a = (l,insert⟨t,Fun (Set s) ts⟩))
in (upds = [] ∧ f A = f B) ∨ (upds ≠ [] ∧ f (g A) = f (g B))"
define db_upds_consts_fresh where "db_upds_consts_fresh ≡
λA::('fun,'atom,'sets,'lbl) prot_constr. λX. λJ::('fun,'atom,'sets,'lbl) prot_subst.
∀x ∈ X. (∃n. ℐ x = Fun (PubConst Value n) []) ⟶ (∀n s.
insert⟨Fun (Val n) [],s⟩ ∈ set (unlabel A) ∨
delete⟨Fun (Val n) [],s⟩ ∈ set (unlabel A) ⟶
J x ≠ Fun (Val n) [])"
define subst_eq_on_privvals where "subst_eq_on_privvals ≡ λX 𝒥.
∀x ∈ X. (∃n. ℐ x = Fun (Val n) []) ⟶ ℐ x = 𝒥 x"
define subst_in_ik_if_subst_pubval where "subst_in_ik_if_subst_pubval ≡
λX 𝒥. λℬ::('fun,'atom,'sets,'lbl) prot_constr.
∀x ∈ X. (∃n. ℐ x = Fun (PubConst Value n) []) ⟶ 𝒥 x ∈ ik⇩l⇩s⇩s⇩t ℬ"
define subst_eq_iff where "subst_eq_iff ≡ λX. λ𝒥::('fun,'atom,'sets,'lbl) prot_subst.
∀x ∈ X. ∀y ∈ X. ℐ x = ℐ y ⟷ 𝒥 x = 𝒥 y"
obtain x_val T_val T_upds s_val ts_val l1_val l2_val where x_val:
"T_val ∈ set P" "Var x_val ∈ set ts_val" "Γ⇩v x_val = TAtom Value"
"fv⇩s⇩e⇩t (set ts_val) = {x_val}" "∀n. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t set ts_val)"
"T_val = Transaction (λ(). []) [x_val] [] [] T_upds [(l1_val,send⟨ts_val⟩)]"
"T_upds = [] ∨
(T_upds = [(l2_val,insert⟨Var x_val,⟨s_val⟩⇩s⟩)] ∧
(∀T ∈ set P. ∀(l,a) ∈ set (transaction_strand T). ∀t.
a ≠ select⟨t,⟨s_val⟩⇩s⟩ ∧ a ≠ ⟨t in ⟨s_val⟩⇩s⟩ ∧ a ≠ ⟨t not in ⟨s_val⟩⇩s⟩ ∧
a ≠ delete⟨t,⟨s_val⟩⇩s⟩))"
using has_initial_value_producing_transactionE[OF P_val P, of thesis]
by (auto simp add: disj_commute)
have 0: "∄n. Fun (PubConst Value n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ" when "ℬ ∈ reachable_constraints P" for ℬ
using reachable_constraints_val_funs_private(1)[OF that P(1)] funs_term_Fun_subterm'
unfolding is_PubConstValue_def by fastforce
have I: "?wt_model ℐ 𝒜" "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ)" "wt⇩s⇩u⇩b⇩s⇩t ℐ"
using welltyped_constraint_model_prefix[OF A(2)[unfolded wt_attack_def]]
A(2)[unfolded wt_attack_def welltyped_constraint_model_def constraint_model_def]
by blast+
have 1: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜. valconst_cases ℐ x"
using reachable_constraints_fv_Value_const_cases[OF P(1) A(1) I(1)]
unfolding valconst_cases_def by blast
have 2: "?rcv_att 𝒜 n"
using A(2) strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" "[send⟨[attack⟨n⟩]⟩]" ℐ]
reachable_constraints_receive_attack_if_attack'(2)[OF A(1) P(1) I(1)]
unfolding wt_attack_def welltyped_constraint_model_def constraint_model_def by simp
note ξ_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P(1)]]
have lmm:
"∃ℬ ∈ reachable_constraints P. ∃𝒥.
?wt_model 𝒥 ℬ ∧ valconsts_only (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X) 𝒥 ∧ (?rcv_att 𝒜 n ⟶ ?rcv_att ℬ n) ∧
subst_eq_on_privvals (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X) 𝒥 ∧
subst_in_ik_if_subst_pubval (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X) 𝒥 ℬ ∧
subst_eq_iff (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X) 𝒥 ∧
vars⇩l⇩s⇩s⇩t 𝒜 = vars⇩l⇩s⇩s⇩t ℬ ∧ fv⇩l⇩s⇩s⇩t 𝒜 = fv⇩l⇩s⇩s⇩t ℬ ∧
(∀n ∈ N. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ)) ∧
ik⇩l⇩s⇩s⇩t 𝒜 ⊆ ik⇩l⇩s⇩s⇩t ℬ ∧ trms⇩l⇩s⇩s⇩t 𝒜 ⊆ trms⇩l⇩s⇩s⇩t ℬ ∧
db_eq 𝒜 ℬ s_val T_upds ∧
db_upds_consts_fresh 𝒜 (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X) 𝒥"
when "finite N" "∀n ∈ N. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜)" "X ∩ fv⇩l⇩s⇩s⇩t 𝒜 = {}"
"finite X" "∀x ∈ X. valconst_cases ℐ x" "∀x ∈ X. Γ⇩v x = TAtom Value"
for N X
using A(1) I(1) 1 that
proof (induction arbitrary: N X rule: reachable_constraints.induct)
case init
define pubvals where "pubvals ≡ {n | n x. x ∈ X ∧ ℐ x = Fun (PubConst Value n) []}"
define X_vals where "X_vals ≡ {n | n x. x ∈ X ∧ ℐ x = Fun (Val n) []}"
have X_vals_finite: "finite X_vals"
using finite_surj[OF init.prems(6),
of X_vals "λx. THE n. ℐ x = Fun (Val n) []"]
unfolding X_vals_def by force
have pubvals_finite: "finite pubvals"
using finite_surj[OF init.prems(6),
of pubvals "λx. THE n. ℐ x = Fun (PubConst Value n) []"]
unfolding pubvals_def by force
obtain T_val_fresh_vals and δ::"nat ⇒ nat"
where T_val_fresh_vals: "T_val_fresh_vals ∩ (N ∪ X_vals) = {}"
and δ: "inj δ" "δ ` pubvals = T_val_fresh_vals"
using ex_finite_disj_nat_inj[OF pubvals_finite finite_UnI[OF init.prems(3) X_vals_finite]]
by blast
have T_val_fresh_vals_finite: "finite T_val_fresh_vals"
using pubvals_finite δ(2) by blast
obtain ℬ::"('fun,'atom,'sets,'lbl) prot_constr"
where B:
"ℬ ∈ reachable_constraints P"
"T_upds = [] ⟹ list_all is_Receive (unlabel ℬ)"
"T_upds ≠ [] ⟹ list_all (λa. is_Insert a ∨ is_Receive a) (unlabel ℬ)"
"vars⇩l⇩s⇩s⇩t ℬ = {}"
"∀n. Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ ⟶ Fun (Val n) [] ∈ ik⇩l⇩s⇩s⇩t ℬ"
"T_val_fresh_vals = {n. Fun (Val n) [] ∈ ik⇩l⇩s⇩s⇩t ℬ}"
"∀l a. (l,a) ∈ set ℬ ∧ is_Insert a ⟶
(l = l2_val ∧ (∃n. a = insert⟨Fun (Val n) [],⟨s_val⟩⇩s⟩))"
using reachable_constraints_initial_value_transaction[
OF P reachable_constraints.init T_val_fresh_vals_finite _ x_val]
by auto
define 𝒥 where "𝒥 ≡ λx.
if x ∈ X ∧ (∃n. ℐ x = Fun (PubConst Value n) [])
then Fun (Val (δ (THE n. ℐ x = Fun (PubConst Value n) []))) []
else ℐ x"
have 0: "ik⇩l⇩s⇩s⇩t [] ⊆ ik⇩l⇩s⇩s⇩t ℬ" "trms⇩l⇩s⇩s⇩t [] ⊆ trms⇩l⇩s⇩s⇩t ℬ" "?rcv_att [] n ⟶ ?rcv_att ℬ n"
"vars⇩l⇩s⇩s⇩t [] = vars⇩l⇩s⇩s⇩t ℬ" "fv⇩l⇩s⇩s⇩t [] = fv⇩l⇩s⇩s⇩t ℬ"
using B(4) vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel ℬ"] by auto
have 1: "db_eq [] ℬ s_val T_upds" using B(2,3,7)
proof (induction ℬ)
case (Cons a B)
then obtain l b where a: "a = (l,b)" by (metis surj_pair)
have IH: "db_eq [] B s_val T_upds" using Cons.prems Cons.IH by auto
show ?case
proof (cases "T_upds = []")
case True
hence "is_Receive b" using a Cons.prems(1) by simp
thus ?thesis using IH unfolding a db_eq_def Let_def by auto
next
case False
hence "is_Insert b ∨ is_Receive b" using a Cons.prems(2) by simp
hence "∃t. a = (l2_val,insert⟨t,⟨s_val⟩⇩s⟩)" when b: "¬is_Receive b"
using b Cons.prems(3) unfolding a by (metis list.set_intros(1))
thus ?thesis using IH False unfolding a db_eq_def Let_def by auto
qed
qed (simp add: db_eq_def)
have 2: "?wt_model 𝒥 ℬ"
unfolding welltyped_constraint_model_def constraint_model_def
proof (intro conjI)
show "wt⇩s⇩u⇩b⇩s⇩t 𝒥" using I(4) init.prems(8) unfolding 𝒥_def wt⇩s⇩u⇩b⇩s⇩t_def by fastforce
show "strand_sem_stateful {} {} (unlabel ℬ) 𝒥"
using B(2,3) strand_sem_stateful_if_no_send_or_check[of "unlabel ℬ" "{}" "{}" 𝒥]
unfolding list_all_iff by blast
show "subst_domain 𝒥 = UNIV" "ground (subst_range 𝒥)"
using I(2) unfolding 𝒥_def subst_domain_def by auto
show "wf⇩t⇩r⇩m⇩s (subst_range 𝒥)"
using I(3) unfolding 𝒥_def by fastforce
qed
have 3: "valconsts_only (fv⇩l⇩s⇩s⇩t [] ∪ X) 𝒥"
using init.prems(7) unfolding 𝒥_def valconsts_only_def valconst_cases_def by fastforce
have 4: "subst_eq_on_privvals (fv⇩l⇩s⇩s⇩t [] ∪ X) 𝒥"
unfolding subst_eq_on_privvals_def 𝒥_def by force
have 5: "subst_in_ik_if_subst_pubval (fv⇩l⇩s⇩s⇩t [] ∪ X) 𝒥 ℬ"
proof (unfold subst_in_ik_if_subst_pubval_def; intro ballI impI)
fix x assume x: "x ∈ fv⇩l⇩s⇩s⇩t [] ∪ X" and "∃n. ℐ x = Fun (PubConst Value n) []"
then obtain n where n: "ℐ x = Fun (PubConst Value n) []" by blast
have "n ∈ pubvals" using x n unfolding pubvals_def by fastforce
hence "δ n ∈ T_val_fresh_vals" using δ(2) by fast
hence "Fun (Val (δ n)) [] ∈ ik⇩l⇩s⇩s⇩t ℬ" using B(6) by fast
thus "𝒥 x ∈ ik⇩l⇩s⇩s⇩t ℬ" using x n unfolding 𝒥_def by simp
qed
have 6: "subst_eq_iff (fv⇩l⇩s⇩s⇩t [] ∪ X) 𝒥"
proof (unfold subst_eq_iff_def; intro ballI)
fix x y assume "x ∈ fv⇩l⇩s⇩s⇩t [] ∪ X" "y ∈ fv⇩l⇩s⇩s⇩t [] ∪ X"
hence x: "x ∈ X" and y: "y ∈ X" by auto
show "ℐ x = ℐ y ⟷ 𝒥 x = 𝒥 y"
proof
show "ℐ x = ℐ y ⟹ 𝒥 x = 𝒥 y" using x y unfolding 𝒥_def by presburger
next
assume J_eq: "𝒥 x = 𝒥 y" show "ℐ x = ℐ y"
proof (cases "∃n. ℐ x = Fun (PubConst Value n) []")
case True
then obtain n where n: "ℐ x = Fun (PubConst Value n) []" by blast
hence J_x: "𝒥 x = Fun (Val (δ n)) []" using x unfolding 𝒥_def by simp
show ?thesis
proof (cases "∃m. ℐ y = Fun (PubConst Value m) []")
case True
then obtain m where m: "ℐ y = Fun (PubConst Value m) []" by blast
have J_y: "𝒥 y = Fun (Val (δ m)) []" using y m unfolding 𝒥_def by simp
show ?thesis using J_eq J_x J_y injD[OF δ(1), of n m] n m by auto
next
case False
then obtain m where m: "ℐ y = Fun (Val m) []"
using init.prems(7) y unfolding valconst_cases_def by blast
moreover have "δ n ∈ T_val_fresh_vals" using δ(2) x n unfolding pubvals_def by blast
moreover have "m ∈ X_vals" using y m unfolding X_vals_def by blast
ultimately have "𝒥 x ≠ ℐ y" using m J_x T_val_fresh_vals by auto
moreover have "𝒥 y = ℐ y" using m unfolding 𝒥_def by simp
ultimately show ?thesis using J_eq by argo
qed
next
case False
then obtain n where n: "ℐ x = Fun (Val n) []"
using init.prems(7) x unfolding valconst_cases_def by blast
hence J_x: "𝒥 x = ℐ x" unfolding 𝒥_def by auto
show ?thesis
proof (cases "∃m. ℐ y = Fun (PubConst Value m) []")
case False
then obtain m where m: "ℐ y = Fun (Val m) []"
using init.prems(7) y unfolding valconst_cases_def by blast
have J_y: "𝒥 y = ℐ y" using y m unfolding 𝒥_def by simp
show ?thesis using J_x J_y J_eq by presburger
next
case True
then obtain m where m: "ℐ y = Fun (PubConst Value m) []" by blast
hence "𝒥 y = Fun (Val (δ m)) []" using y unfolding 𝒥_def by fastforce
moreover have "δ m ∈ T_val_fresh_vals" using δ(2) y m unfolding pubvals_def by blast
moreover have "n ∈ X_vals" using x n unfolding X_vals_def by blast
ultimately have "𝒥 y ≠ ℐ x" using n J_x T_val_fresh_vals by auto
thus ?thesis using J_x J_eq by argo
qed
qed
qed
qed
have 7: "∀n ∈ N. Fun (Val n) [] ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t ℬ)"
using B(5,6) T_val_fresh_vals by blast
have 8: "db_upds_consts_fresh [] (fv⇩l⇩s⇩s⇩t [] ∪ X) 𝒥" unfolding db_upds_consts_fresh_def by simp
show ?case using B(1) 0 1 2 3 4 5 6 7 8 by blast
next
case (step 𝒜 T ξ σ α N X')
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"
define T'_pubvals where "T'_pubvals ≡ {n. ∃x ∈ fv⇩l⇩s⇩s⇩t T'. ℐ x = Fun (PubConst Value n) []}"
define 𝒜_vals where "𝒜_vals ≡ {n. Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜}"
define ℐ_vals where "ℐ_vals ≡ {n. ∃x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X' ∪ fv⇩l⇩s⇩s⇩t T'. ℐ x = Fun (Val n) []}"
define σ_vals where "σ_vals ≡ {n. Fun (Val n) [] ∈ subst_range σ}"
have 3: "welltyped_constraint_model ℐ 𝒜"
"∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜. valconst_cases ℐ x"
"∀x ∈ fv⇩l⇩s⇩s⇩t T'. valconst_cases ℐ x"
"strand_sem_stateful (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) (dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}) (unlabel T') ℐ"
"∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'. valconst_cases ℐ x"
using step.prems(2) welltyped_constraint_model_prefix[OF step.prems(1)]
step.prems(1)[unfolded welltyped_constraint_model_def constraint_model_def]
strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" "unlabel T'" ℐ]
step.prems(7)
unfolding θ_def[symmetric] T'_def[symmetric] unlabel_append fv⇩s⇩s⇩t_append
by (blast,blast,blast,simp,blast)
note T_adm = bspec[OF P step.hyps(2)]
note T_wf = admissible_transaction_is_wellformed_transaction[OF T_adm]
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)]
note 4 = admissible_transaction_sem_iff
note 5 = iffD1[OF 4[OF T_adm I(2,3), of "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" "dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}" θ,
unfolded T'_def[symmetric]]
3(4)]
note σ_dom = transaction_fresh_subst_domain[OF step.hyps(4)]
have σ_ran: "∃n. t = Fun (Val n) []" when t: "t ∈ subst_range σ" for t
proof -
obtain x where x: "x ∈ set (transaction_fresh T)" "t = σ x"
using σ_dom t by auto
show ?thesis
using x(1) admissible_transactionE(2)[OF T_adm]
transaction_fresh_subst_sends_to_val[OF step.hyps(4) x(1)]
unfolding x(2) by meson
qed
have T'_vals_in_σ: "Fun (Val k) [] ∈ subst_range σ"
when k: "Fun (Val k) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t T'" for k
proof -
have "Fun (Val k) [] ∈ (subterms⇩s⇩e⇩t (trms_transaction T)) ⋅⇩s⇩e⇩t θ"
using k admissible_transactionE(4)[OF T_adm]
transaction_decl_fresh_renaming_substs_trms[
OF step.hyps(3,4,5), of "transaction_strand T"]
unfolding T'_def θ_def[symmetric] trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq by fast
then obtain t where t: "t ⊑⇩s⇩e⇩t trms_transaction T" "t ⋅ θ = Fun (Val k) []" by force
hence "Fun (Val k) [] ∈ subst_range θ"
using admissible_transactions_no_Value_consts(1)[OF T_adm] by (cases t) force+
thus ?thesis
using transaction_decl_fresh_renaming_substs_range'(4)[OF step.hyps(3,4,5)] ξ_empty
unfolding θ_def[symmetric] by blast
qed
have σ_vals_is_T'_vals: "k ∈ σ_vals ⟷ Fun (Val k) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t T'" for k
proof
show "k ∈ σ_vals" when "Fun (Val k) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t T'"
using that T'_vals_in_σ unfolding σ_vals_def by blast
show "Fun (Val k) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t T'" when k: "k ∈ σ_vals"
proof -
have "Fun (Val k) [] ∈ subst_range σ" using k unfolding σ_vals_def by fast
then obtain x where x: "x ∈ fv_transaction T" "σ x = Fun (Val k) []"
using admissible_transactionE(7)[OF T_adm]
transaction_fresh_subst_domain[OF step.hyps(4)]
unfolding fv_transaction_unfold by fastforce
have "θ x = Fun (Val k) []" using x(2) unfolding θ_def ξ_empty subst_compose_def by auto
thus ?thesis
using fv⇩s⇩s⇩t_is_subterm_trms⇩s⇩s⇩t_subst[OF x(1), of θ]
admissible_transactionE(4)[OF T_adm]
unfolding T'_def trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unlabel_subst by simp
qed
qed
have σ_vals_N_disj: "N ∩ σ_vals = {}"
using step.prems(4) σ_vals_is_T'_vals
unfolding θ_def[symmetric] T'_def[symmetric] unlabel_append trms⇩s⇩s⇩t_append by blast
have T'_pubvals_finite: "finite T'_pubvals"
using finite_surj[OF fv⇩s⇩s⇩t_finite[of "unlabel T'"],
of T'_pubvals "λx. THE n. ℐ x = Fun (PubConst Value n) []"]
unfolding T'_pubvals_def by force
have σ_vals_finite: "finite σ_vals"
proof -
have *: "finite (subst_range σ)" using transaction_fresh_subst_domain[OF step.hyps(4)] by simp
show ?thesis
using finite_surj[OF *, of σ_vals "λt. THE n. t = Fun (Val n) []"]
unfolding σ_vals_def by force
qed
have 𝒜_vals_finite: "finite 𝒜_vals"
proof -
have *: "𝒜_vals ⊆ (λt. THE n. t = Fun (Val n) []) ` subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
unfolding 𝒜_vals_def by force
show ?thesis
by (rule finite_surj[OF subterms_union_finite[OF trms⇩s⇩s⇩t_finite] *])
qed
have ℐ_vals_finite: "finite ℐ_vals"
proof -
define X where "X ≡ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X' ∪ fv⇩l⇩s⇩s⇩t T'"
have *: "finite X" using fv⇩s⇩s⇩t_finite step.prems(6) unfolding X_def by blast
show ?thesis
using finite_surj[OF *, of ℐ_vals "λx. THE n. ℐ x = Fun (Val n) []"]
unfolding ℐ_vals_def X_def[symmetric] by force
qed
obtain T_val_fresh_vals and δ::"nat ⇒ nat"
where T_val_fresh_vals: "T_val_fresh_vals ∩ (N ∪ σ_vals ∪ 𝒜_vals ∪ ℐ_vals) = {}"
and δ: "inj δ" "δ ` T'_pubvals = T_val_fresh_vals"
using step.prems(3) T'_pubvals_finite σ_vals_finite 𝒜_vals_finite ℐ_vals_finite
by (metis finite_UnI ex_finite_disj_nat_inj)
define N' where "N' ≡ N ∪ σ_vals ∪ T_val_fresh_vals"
have T_val_fresh_vals_finite: "finite T_val_fresh_vals"
using T'_pubvals_finite δ(2) by blast
have N'_finite: "finite N'"
using step.prems(3) T'_pubvals_finite T_val_fresh_vals_finite σ_vals_finite
unfolding N'_def by auto
have 𝒜_vals_trms_in: "n ∈ 𝒜_vals" when "Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜" for n
using that unfolding 𝒜_vals_def by blast
have N'_notin_𝒜: "¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜)" when n: "n ∈ N'" for n
proof -
have ?thesis when n': "n ∈ N"
using n' step.prems(4) unfolding N'_def unlabel_append trms⇩s⇩s⇩t_append by blast
moreover have ?thesis when n': "n ∈ σ_vals"
using n' step.hyps(4) unfolding σ_vals_def transaction_fresh_subst_def by blast
moreover have ?thesis when n': "n ∈ T_val_fresh_vals"
using n' T_val_fresh_vals 𝒜_vals_trms_in by blast
ultimately show ?thesis using n unfolding N'_def by blast
qed
have T'_fv_𝒜_disj: "fv⇩l⇩s⇩s⇩t 𝒜 ∩ fv⇩l⇩s⇩s⇩t T' = {}"
using transaction_decl_fresh_renaming_substs_vars_disj(8)[OF step.hyps(3,4,5)]
transaction_decl_fresh_renaming_substs_vars_subset(4)[OF step.hyps(3,4,5,2)]
unfolding θ_def[symmetric] T'_def fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq by blast
have X'_disj: "X' ∩ fv⇩l⇩s⇩s⇩t 𝒜 = {}" "X' ∩ fv⇩l⇩s⇩s⇩t T' = {}"
using step.prems(5)
unfolding θ_def[symmetric] T'_def[symmetric] unlabel_append fv⇩s⇩s⇩t_append
by (blast, blast)
have X'_disj': "(X' ∪ fv⇩l⇩s⇩s⇩t T') ∩ fv⇩l⇩s⇩s⇩t 𝒜 = {}"
using X'_disj(1) T'_fv_𝒜_disj by blast
have X'_finite: "finite (X' ∪ fv⇩l⇩s⇩s⇩t T')"
using step.prems(6) fv⇩s⇩s⇩t_finite by blast
have 𝒜_X'_valconstcases: "∀x ∈ X' ∪ fv⇩l⇩s⇩s⇩t T'. valconst_cases ℐ x"
using 3(3,5) by blast
have T'_value_vars: "Γ⇩v x = TAtom Value" when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" for x
using x reachable_constraints_fv_Value_typed[
OF P reachable_constraints.step[OF step.hyps]]
unfolding θ_def[symmetric] T'_def[symmetric] unlabel_append fv⇩s⇩s⇩t_append by blast
have X'_T'_value_vars: "∀x ∈ X' ∪ fv⇩l⇩s⇩s⇩t T'. Γ⇩v x = TAtom Value"
using step.prems(8) T'_value_vars by blast
have N'_not_subterms_𝒜: "∀n ∈ N'. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜)"
using N'_notin_𝒜 by blast
obtain ℬ 𝒥 where B:
"ℬ ∈ reachable_constraints P" "?wt_model 𝒥 ℬ"
"valconsts_only (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X' ∪ fv⇩l⇩s⇩s⇩t T') 𝒥" "?rcv_att 𝒜 n ⟶ ?rcv_att ℬ n"
"subst_eq_on_privvals (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X' ∪ fv⇩l⇩s⇩s⇩t T') 𝒥"
"subst_in_ik_if_subst_pubval (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X' ∪ fv⇩l⇩s⇩s⇩t T') 𝒥 ℬ"
"subst_eq_iff (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X' ∪ fv⇩l⇩s⇩s⇩t T') 𝒥"
"vars⇩l⇩s⇩s⇩t 𝒜 = vars⇩l⇩s⇩s⇩t ℬ" "fv⇩l⇩s⇩s⇩t 𝒜 = fv⇩l⇩s⇩s⇩t ℬ" "ik⇩l⇩s⇩s⇩t 𝒜 ⊆ ik⇩l⇩s⇩s⇩t ℬ" "trms⇩l⇩s⇩s⇩t 𝒜 ⊆ trms⇩l⇩s⇩s⇩t ℬ"
"∀n ∈ N'. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ)"
"db_eq 𝒜 ℬ s_val T_upds"
"db_upds_consts_fresh 𝒜 (fv⇩l⇩s⇩s⇩t 𝒜 ∪ X' ∪ fv⇩l⇩s⇩s⇩t T') 𝒥"
using step.IH[OF 3(1,2) N'_finite N'_not_subterms_𝒜 X'_disj' X'_finite
𝒜_X'_valconstcases X'_T'_value_vars]
unfolding Un_assoc by fast
have J:
"wt⇩s⇩u⇩b⇩s⇩t 𝒥" "constr_sem_stateful 𝒥 (unlabel ℬ)"
"interpretation⇩s⇩u⇩b⇩s⇩t 𝒥" "wf⇩t⇩r⇩m⇩s (subst_range 𝒥)"
using B(2) unfolding welltyped_constraint_model_def constraint_model_def by blast+
have T_val_fresh_vals_notin_ℬ: "¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ)"
when "n ∈ T_val_fresh_vals" for n
using that B(12) unfolding N'_def by blast
hence "∀n ∈ T_val_fresh_vals. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ)" by blast
then obtain T_val_constr::"('fun,'atom,'sets,'lbl) prot_constr"
where T_val_constr:
"ℬ@T_val_constr ∈ reachable_constraints P"
"T_val_constr ∈ reachable_constraints P"
"T_upds = [] ⟹ list_all is_Receive (unlabel T_val_constr)"
"T_upds ≠ [] ⟹ list_all (λa. is_Insert a ∨ is_Receive a) (unlabel T_val_constr)"
"vars⇩l⇩s⇩s⇩t T_val_constr = {}"
"∀n. Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ ⟶ Fun (Val n) [] ∉ ik⇩l⇩s⇩s⇩t T_val_constr"
"∀n. Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t T_val_constr ⟶ Fun (Val n) [] ∈ ik⇩l⇩s⇩s⇩t T_val_constr"
"T_val_fresh_vals = {n. Fun (Val n) [] ∈ ik⇩l⇩s⇩s⇩t T_val_constr}"
"∀l a. (l,a) ∈ set T_val_constr ∧ is_Insert a ⟶
(l = l2_val ∧ (∃n. a = insert⟨Fun (Val n) [],⟨s_val⟩⇩s⟩))"
using reachable_constraints_initial_value_transaction[
OF P B(1) T_val_fresh_vals_finite _ x_val]
by blast
have T_val_constr_no_upds_if_no_T_upds:
"filter is_Update (unlabel T_val_constr) = []"
when "T_upds = []"
using T_val_constr(3)[OF that] by (induct T_val_constr) auto
have T_val_fresh_vals_is_T_val_constr_vals:
"k ∈ T_val_fresh_vals ⟷ Fun (Val k) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t T_val_constr"
for k
using that T_val_constr(7,8) ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset by fast
have T_val_constr_no_fv: "fv⇩l⇩s⇩s⇩t T_val_constr = {}"
using T_val_constr(5) vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t by fast
have T_val_σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t (ℬ@T_val_constr))"
proof -
have "¬(t ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t (ℬ@T_val_constr))" when t: "t ∈ subst_range σ" for t
proof -
obtain k where k: "t = Fun (Val k) []" using t σ_ran by fast
have "k ∈ σ_vals" using t unfolding k σ_vals_def by blast
thus ?thesis
using B(12) T_val_fresh_vals T_val_constr(7,8)
unfolding N'_def k unlabel_append trms⇩s⇩s⇩t_append by blast
qed
thus ?thesis using step.hyps(4) unfolding transaction_fresh_subst_def by fast
qed
have T_val_α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t (ℬ@T_val_constr))"
using step.hyps B(8) T_val_constr(5) by auto
define ℬ' where "ℬ' ≡ ℬ@T_val_constr@T'"
define 𝒦 where "𝒦 ≡ λx.
if x ∈ fv⇩l⇩s⇩s⇩t T'
then if ∃n. ℐ x = Fun (PubConst Value n) []
then if ∃y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ y = ℐ x
then 𝒥 (SOME y. y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X' ∧ ℐ y = ℐ x)
else Fun (Val (δ (THE n. ℐ x = Fun (PubConst Value n) []))) []
else ℐ x
else 𝒥 x"
have σ_ground_ran: "ground (subst_range σ)" "range_vars σ = {}"
and ξ_ran_bvars_disj: "range_vars ξ ∩ bvars_transaction T = {}"
using transaction_fresh_subst_domain[OF step.hyps(4)]
transaction_fresh_subst_range_vars_empty[OF step.hyps(4)]
transaction_decl_subst_range_vars_empty[OF step.hyps(3)]
by (metis range_vars_alt_def, argo, blast)
have ℬ_T'_fv_disj: "fv⇩l⇩s⇩s⇩t ℬ ∩ fv⇩l⇩s⇩s⇩t T' = {}"
using T'_fv_𝒜_disj unfolding B(9) by argo
hence 𝒥_𝒦_fv_ℬ_eq: "𝒥 x = 𝒦 x" when x: "x ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'" for x
using x X'_disj unfolding 𝒦_def by auto
have B'1: "ℬ' ∈ reachable_constraints P"
using reachable_constraints.step[OF T_val_constr(1) step.hyps(2,3) T_val_σ T_val_α]
unfolding ℬ'_def T'_def θ_def by simp
have "∃n. 𝒦 x = Fun (Val n) []" when x: "x ∈ fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X'" for x
proof (cases "x ∈ fv⇩l⇩s⇩s⇩t T'")
case True
note T = this
show ?thesis
proof (cases "∃n. ℐ x = Fun (PubConst Value n) []")
case True thus ?thesis
using T B(3,9) someI_ex[of "λy. y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X' ∧ ℐ y = ℐ x"]
unfolding 𝒦_def valconsts_only_def
by (cases "∃y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ y = ℐ x") (meson, auto)
next
case False thus ?thesis
using T 3(3) unfolding 𝒦_def valconst_cases_def by fastforce
qed
next
case False thus ?thesis using x B(3) unfolding 𝒦_def valconsts_only_def by auto
qed
hence B'3: "valconsts_only (fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X') 𝒦" unfolding valconsts_only_def by blast
have B'4: "?rcv_att ℬ' n" when "?rcv_att (𝒜@T') n"
using that B(4) unfolding ℬ'_def by auto
have "ℐ x = 𝒦 x" when x: "x ∈ fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X'" "ℐ x = Fun (Val n) []" for x n
proof -
have "𝒦 x = 𝒥 x" when "x ∉ fv⇩l⇩s⇩s⇩t T'" using that unfolding 𝒦_def by meson
moreover have "𝒦 x = ℐ x" when "x ∈ fv⇩l⇩s⇩s⇩t T'" using that x unfolding 𝒦_def by simp
ultimately show ?thesis
using B(5) x
unfolding subst_eq_on_privvals_def unlabel_append fv⇩s⇩s⇩t_append
by (cases "x ∈ fv⇩l⇩s⇩s⇩t T'") auto
qed
hence B'5: "subst_eq_on_privvals (fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X') 𝒦"
unfolding subst_eq_on_privvals_def by blast
have 𝒜_fv_𝒦_eq_𝒥: "𝒦 x = 𝒥 x" when x: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" for x
proof -
have "x ∉ fv⇩l⇩s⇩s⇩t T'" using x T'_fv_𝒜_disj X'_disj by blast
thus ?thesis unfolding 𝒦_def by argo
qed
have T'_fv_ℐ_val_𝒦_eq_𝒥: "𝒦 x = ℐ x"
when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" "∄n. ℐ x = Fun (PubConst Value n) []" for x
using x B'5 3(3) unfolding unlabel_append fv⇩s⇩s⇩t_append valconst_cases_def 𝒦_def by meson
have T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val:
"𝒦 x = Fun (Val (δ n)) []" "δ n ∈ T_val_fresh_vals"
when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" "ℐ x = Fun (PubConst Value n) []" "∀y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ y ≠ ℐ x"
for x n
proof -
show "𝒦 x = Fun (Val (δ n)) []" using x unfolding 𝒦_def by auto
show "δ n ∈ T_val_fresh_vals" using δ(2) x unfolding T'_pubvals_def by blast
qed
have T'_fv_ℐ_pubval_𝒦_eq_𝒥_val:
"∃y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ∃m. ℐ y = ℐ x ∧ 𝒦 x = 𝒥 y ∧ 𝒦 x = Fun (Val m) []"
when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" "ℐ x = Fun (PubConst Value n) []" "∃y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ y = ℐ x"
for x n
proof -
have "𝒦 x = 𝒥 (SOME y. y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X' ∧ ℐ y = ℐ x)" using x unfolding 𝒦_def by meson
then obtain y where y: "y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'" "ℐ y = ℐ x" "𝒦 x = 𝒥 y"
using x(3) someI_ex[of "λy. y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X' ∧ ℐ y = ℐ x"] by blast
thus ?thesis using B(3,9) unfolding valconsts_only_def by auto
qed
have T'_fv_ℐ_pubval_𝒦_eq_val: "∃n. 𝒦 x = Fun (Val n) []"
when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" "ℐ x = Fun (PubConst Value n) []" for x n
using T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val[OF x] T'_fv_ℐ_pubval_𝒦_eq_𝒥_val[OF x] by auto
have B'6': "𝒦 x ∈ ik⇩l⇩s⇩s⇩t ℬ"
when x: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" "ℐ x = Fun (PubConst Value n) []" for x n
using x B(6) 𝒜_fv_𝒦_eq_𝒥 x(2) unfolding B(8) subst_in_ik_if_subst_pubval_def by simp
have B'6'': "𝒦 x ∈ ik⇩l⇩s⇩s⇩t ℬ ∪ ik⇩l⇩s⇩s⇩t T_val_constr"
when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" "ℐ x = Fun (PubConst Value n) []" for x n
proof (cases "∃y ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ y = ℐ x")
case True thus ?thesis
using B(6) x(2) T'_fv_ℐ_pubval_𝒦_eq_𝒥_val[OF x True]
unfolding B(9) subst_in_ik_if_subst_pubval_def by force
next
case False thus ?thesis
using T_val_constr(8) T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val[OF x] by force
qed
have "𝒦 x ∈ ik⇩l⇩s⇩s⇩t ℬ'"
when x: "x ∈ fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X'" "ℐ x = Fun (PubConst Value n) []" for x n
proof (cases "x ∈ fv⇩l⇩s⇩s⇩t T'")
case True thus ?thesis using B'6''[OF _ x(2)] unfolding ℬ'_def by auto
next
case False
hence "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'"
using x(1) unfolding unlabel_append fv⇩s⇩s⇩t_append by blast
thus ?thesis using B'6' x(2) unfolding ℬ'_def by simp
qed
hence B'6: "subst_in_ik_if_subst_pubval (fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X') 𝒦 ℬ'"
unfolding subst_in_ik_if_subst_pubval_def by blast
have B'7: "subst_eq_iff (fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X') 𝒦"
proof (unfold subst_eq_iff_def; intro ballI)
fix x y assume xy: "x ∈ fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X'" "y ∈ fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X'"
let ?Q = "λx y. ℐ x = ℐ y ⟷ 𝒦 x = 𝒦 y"
have *: "?Q x y"
when xy: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" "x ∉ fv⇩l⇩s⇩s⇩t T'" "y ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" "y ∉ fv⇩l⇩s⇩s⇩t T'" for x y
using B(7) xy unfolding 𝒦_def subst_eq_iff_def by force
have **: "?Q x y" when x: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" and y: "y ∈ fv⇩l⇩s⇩s⇩t T'" for x y
proof -
have xy_neq: "x ≠ y" using x y T'_fv_𝒜_disj X'_disj by blast
have x_eq: "𝒦 x = 𝒥 x"
using 𝒜_fv_𝒦_eq_𝒥 x by blast
have x_eq_if_val: "ℐ x = 𝒥 x" when "ℐ x = Fun (Val n) []" for n
using that x B(5) unfolding subst_eq_on_privvals_def by blast
have x_neq_if_neq_val: "ℐ x ≠ 𝒥 x" when "ℐ x = Fun (PubConst Value n) []" for n
by (metis that B(3) x UnI1 prot_fun.distinct(37) term.inject(2) valconsts_only_def)
have y_eq_if_val: "ℐ y = 𝒦 y" when "ℐ y = Fun (Val n) []" for n
using that y B'5 unfolding subst_eq_on_privvals_def by simp
have y_eq: "𝒦 y = Fun (Val (δ n)) []"
when "ℐ y = Fun (PubConst Value n) []" "∀z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z ≠ ℐ y" for n
by (rule T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val(1)[OF y that])
have y_eq': "∃z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ∃m. ℐ z = ℐ y ∧ 𝒦 y = 𝒥 z ∧ 𝒦 y = Fun (Val m) []"
when "ℐ y = Fun (PubConst Value n) []" "∃z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z = ℐ y" for n
by (rule T'_fv_ℐ_pubval_𝒦_eq_𝒥_val[OF y that])
have K_eq_if_I_eq: "ℐ x = ℐ y ⟹ 𝒦 x = 𝒦 y"
apply (cases "∃n. ℐ x = Fun (PubConst Value n) []")
subgoal using B(7,9) unfolding subst_eq_iff_def by (metis UnI1 x x_eq y_eq')
subgoal by (metis x_eq x T'_fv_ℐ_val_𝒦_eq_𝒥[OF y] 3(5) valconst_cases_def x_eq_if_val)
done
have K_neq_if_I_neq_val: "𝒦 x ≠ 𝒦 y"
when n: "ℐ y = Fun (Val n) []"
and m: "ℐ x = Fun (PubConst Value m) []"
for n m
proof -
have I_neq: "ℐ x ≠ ℐ y" using n m by simp
note y_eq'' = y_eq_if_val[OF n]
note x_neq = x_neq_if_neq_val[OF m]
have x_ex: "∃z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z = ℐ x" using x unfolding B(9) by blast
have J1: "𝒥 x ∈ ik⇩l⇩s⇩s⇩t ℬ" using B(6) x m unfolding subst_in_ik_if_subst_pubval_def by fast
have J2: "ℐ x ≠ 𝒥 x"
by (metis m B(3) x UnI1 prot_fun.distinct(37) term.inject(2) valconsts_only_def)
have J3: "ℐ y = 𝒥 y" using B(5) n y unfolding subst_eq_on_privvals_def by blast
have K_x: "𝒦 x ≠ ℐ x" using J2 x_eq by presburger
have x_notin: "x ∉ fv⇩l⇩s⇩s⇩t T'" using x T'_fv_𝒜_disj X'_disj by blast
have K_x': "𝒦 x = 𝒥 x" using x_notin unfolding 𝒦_def by argo
have K_y: "𝒦 y = 𝒥 y" using y_eq'' J3 by argo
have "𝒥 x ≠ 𝒥 y" using I_neq x y B(7) unfolding subst_eq_iff_def by blast
thus ?thesis using K_x' K_y by argo
qed
show ?thesis
proof
show "ℐ x = ℐ y ⟹ 𝒦 x = 𝒦 y" by (rule K_eq_if_I_eq)
next
assume xy_eq: "𝒦 x = 𝒦 y" show "ℐ x = ℐ y"
proof (cases "∃n. ℐ y = Fun (PubConst Value n) []")
case True
then obtain n where n: "ℐ y = Fun (PubConst Value n) []" by blast
show ?thesis
proof (cases "∃z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z = ℐ y")
case True thus ?thesis
using B(7,9) unfolding subst_eq_iff_def by (metis xy_eq UnI1 x x_eq y_eq'[OF n])
next
case False
hence F: "∀z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z ≠ ℐ y" by blast
note y_eq'' = y_eq[OF n F]
have n_in: "δ n ∈ T_val_fresh_vals"
using δ(2) x_eq xy_eq T_val_fresh_vals_notin_ℬ y n
unfolding T'_pubvals_def by blast
hence y_notin: "¬(𝒦 y ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t ℬ)"
using T_val_fresh_vals_notin_ℬ y_eq'' ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset[of "unlabel ℬ"]
by auto
show ?thesis
proof (cases "∃m. ℐ x = Fun (PubConst Value m) []")
case True thus ?thesis
using y_notin B(6) x xy_eq x_eq
unfolding B(9) subst_in_ik_if_subst_pubval_def
by fastforce
next
case False
then obtain m where m: "ℐ x = Fun (Val m) []"
using 3(5) x unfolding valconst_cases_def by fast
hence "ℐ x = 𝒥 x" using x B(5) unfolding subst_eq_on_privvals_def by blast
hence "𝒦 x = Fun (Val m) []" using m x_eq by argo
moreover have "m ∉ T_val_fresh_vals"
using m T_val_fresh_vals x unfolding ℐ_vals_def by blast
hence "m ≠ δ n" using n_in by blast
ultimately have False using xy_eq y_eq'' by force
thus ?thesis by simp
qed
qed
next
case False
then obtain n where n: "ℐ y = Fun (Val n) []"
using 3(3) y unfolding valconst_cases_def by fast
note y_eq'' = y_eq_if_val[OF n]
show ?thesis
proof (cases "∃m. ℐ x = Fun (Val m) []")
case True thus ?thesis by (metis xy_eq x_eq y_eq'' x_eq_if_val)
next
case False
then obtain m where m: "ℐ x = Fun (PubConst Value m) []"
using 3(5) x unfolding valconst_cases_def by blast
show ?thesis using K_neq_if_I_neq_val[OF n m] xy_eq by blast
qed
qed
qed
qed
have ***: "?Q x y" when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" and y: "y ∈ fv⇩l⇩s⇩s⇩t T'" for x y
proof
assume xy_eq: "ℐ x = ℐ y" show "𝒦 x = 𝒦 y"
proof (cases "∃n. ℐ x = Fun (PubConst Value n) []")
case True thus ?thesis
using xy_eq x y B(7) T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val(1) T'_fv_ℐ_pubval_𝒦_eq_𝒥_val
unfolding B(9) subst_eq_iff_def by (metis (no_types) UnI1)
qed (metis xy_eq x y T'_fv_ℐ_val_𝒦_eq_𝒥)
next
assume xy_eq: "𝒦 x = 𝒦 y"
have case1: False
when x': "x ∈ fv⇩l⇩s⇩s⇩t T'"
and y': "y ∈ fv⇩l⇩s⇩s⇩t T'"
and xy_eq': "𝒦 x = 𝒦 y"
and m: "ℐ x = Fun (PubConst Value m) []"
and n: "ℐ y = Fun (Val n) []"
for x y n m
proof -
have F: "∄n. ℐ y = Fun (PubConst Value n) []" using n by auto
note x_eq = T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val[OF x' m]
note y_eq = T'_fv_ℐ_val_𝒦_eq_𝒥[OF y' F]
have "∃z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z = ℐ x"
proof (rule ccontr)
assume no_z: "¬(∃z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z = ℐ x)"
hence "n ≠ δ m" using n y' x_eq(2) T_val_fresh_vals unfolding ℐ_vals_def by blast
thus False using xy_eq' x_eq y_eq(1) n no_z by auto
qed
then obtain z k where z:
"z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'" "ℐ z = ℐ x" "𝒦 x = 𝒥 z" "𝒦 x = Fun (Val k) []"
using T'_fv_ℐ_pubval_𝒦_eq_𝒥_val[OF x' m] by blast
have "ℐ y = 𝒥 z" using z(2,3) y_eq xy_eq' by presburger
hence "ℐ x = ℐ y" using z(1,2) ** B(9) 𝒥_𝒦_fv_ℬ_eq y' y_eq by metis
thus False using n m by simp
qed
have case2: "m = n"
when x': "x ∈ fv⇩l⇩s⇩s⇩t T'"
and y': "y ∈ fv⇩l⇩s⇩s⇩t T'"
and xy_eq': "𝒦 x = 𝒦 y"
and m: "ℐ x = Fun (PubConst Value m) []"
and n: "ℐ y = Fun (PubConst Value n) []"
for x y n m
proof (cases "∀z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z ≠ ℐ x")
case True show ?thesis
apply (cases "∀z ∈ fv⇩l⇩s⇩s⇩t ℬ ∪ X'. ℐ z ≠ ℐ y")
subgoal
using xy_eq' m n T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val[OF x' m True]
T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val[OF y' n] injD[OF δ(1), of m n]
by fastforce
subgoal by (metis x' y' xy_eq' ** B(9) True)
done
qed (metis x' y' xy_eq' m n ** B(9) prot_fun.inject(6) term.inject(2))
have case3: "m = n"
when x': "x ∈ fv⇩l⇩s⇩s⇩t T'"
and y': "y ∈ fv⇩l⇩s⇩s⇩t T'"
and xy_eq': "𝒦 x = 𝒦 y"
and m: "ℐ x = Fun (Val m) []"
and n: "ℐ y = Fun (Val n) []"
for x y n m
using x' y' xy_eq' m n T'_fv_ℐ_val_𝒦_eq_𝒥 by fastforce
show "ℐ x = ℐ y"
using x y xy_eq case1 case2 case3 3(3)
unfolding valconst_cases_def by metis
qed
have ****: "?Q x y" when xy: "x ∈ X'" "x ∉ fv⇩l⇩s⇩s⇩t T'" "y ∈ fv⇩l⇩s⇩s⇩t T'" for x y
proof -
have xy': "y ∉ X'" "y ∉ fv⇩l⇩s⇩s⇩t 𝒜" "x ∉ fv⇩l⇩s⇩s⇩t 𝒜"
using xy X'_disj T'_fv_𝒜_disj by (blast, blast, blast)
have K_x: "𝒦 x = 𝒥 x" using xy(2) unfolding 𝒦_def by argo
have I_iff_J: "ℐ x = ℐ y ⟷ 𝒥 x = 𝒥 y" using xy B(7) unfolding subst_eq_iff_def by fast
show ?thesis using K_x I_iff_J K_x injD[OF δ(1)] xy B(7,9) ** by (meson UnCI)
qed
show "?Q x y"
using xy * **[of x y] **[of y x] ***[of x y] ****[of x y] ****[of y x]
unfolding unlabel_append fv⇩s⇩s⇩t_append by (metis Un_iff)
qed
have B'8_9: "vars⇩l⇩s⇩s⇩t (𝒜@T') = vars⇩l⇩s⇩s⇩t ℬ'" "fv⇩l⇩s⇩s⇩t (𝒜@T') = fv⇩l⇩s⇩s⇩t ℬ'"
using B(8,9) T_val_constr(5) vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel T_val_constr"]
unfolding ℬ'_def unlabel_append vars⇩s⇩s⇩t_append fv⇩s⇩s⇩t_append by simp_all
have I': "∃n. ℐ x = Fun (PubConst Value n) []"
when x: "x ∈ fv⇩l⇩s⇩s⇩t T'" "∄n. ℐ x = Fun (Val n) []" for x
using x 3(3) unfolding valconst_cases_def by fast
have B'10_11: "ik⇩l⇩s⇩s⇩t (𝒜@T') ⊆ ik⇩l⇩s⇩s⇩t ℬ'" "trms⇩l⇩s⇩s⇩t (𝒜@T') ⊆ trms⇩l⇩s⇩s⇩t ℬ'"
using B(10,11) unfolding N'_def ℬ'_def unlabel_append trms⇩s⇩s⇩t_append ik⇩s⇩s⇩t_append
by (blast, blast)
have B'12: "∀n ∈ N. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ')"
using B(12) σ_vals_is_T'_vals σ_vals_N_disj T_val_fresh_vals
T_val_fresh_vals_is_T_val_constr_vals
unfolding N'_def ℬ'_def unfolding unlabel_append trms⇩s⇩s⇩t_append ik⇩s⇩s⇩t_append by fast
have B'13: "db_eq (𝒜@T') ℬ' s_val T_upds"
proof -
let ?f = "filter is_Update ∘ unlabel"
let ?g = "filter (λa. ∄l t ts. a = (l, insert⟨t,Fun (Set s_val) ts⟩))"
have "?f (?g T_val_constr) = []" using T_val_constr(3,4,9)
proof (induction T_val_constr)
case (Cons a B)
then obtain l b where a: "a = (l,b)" by (metis surj_pair)
have IH: "?f (?g B) = []" using Cons.prems Cons.IH by auto
show ?case
proof (cases "T_upds = []")
case True
hence "is_Receive b" using a Cons.prems(1,2) by simp
thus ?thesis using IH unfolding a Let_def by auto
next
case False
hence "is_Insert b ∨ is_Receive b" using a Cons.prems(1,2) by simp
hence "∃t. a = (l2_val, insert⟨t,⟨s_val⟩⇩s⟩)" when b: "¬is_Receive b"
using b Cons.prems(3) unfolding a by (metis list.set_intros(1))
thus ?thesis using IH unfolding a Let_def by auto
qed
qed simp
hence "?f (?g (𝒜@T')) = ?f (?g 𝒜)@?f (?g T')"
"?f (?g (ℬ@T_val_constr@T')) = ?f (?g ℬ)@?f (?g T')"
when "T_upds ≠ []"
by simp_all
moreover have "?f T_val_constr = []" when "T_upds = []"
using T_val_constr_no_upds_if_no_T_upds[OF that] by force
hence "?f (𝒜@T') = ?f 𝒜@?f T'"
"?f (ℬ@T_val_constr@T') = ?f ℬ@?f T'"
when "T_upds = []"
using that by auto
ultimately show ?thesis using B(13) unfolding ℬ'_def db_eq_def Let_def by presburger
qed
have B'14: "db_upds_consts_fresh (𝒜@T') (fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X') 𝒦"
proof (unfold db_upds_consts_fresh_def; intro ballI allI impI; elim exE)
fix x s n m
assume x: "x ∈ fv⇩l⇩s⇩s⇩t (𝒜@T') ∪ X'"
and n: "insert⟨Fun (Val n) [],s⟩ ∈ set (unlabel (𝒜@T')) ∨
delete⟨Fun (Val n) [],s⟩ ∈ set (unlabel (𝒜@T'))"
(is "?A (𝒜@T')")
and m: "ℐ x = Fun (PubConst Value m) []"
have A_cases: "?A 𝒜 ∨ ?A T'" using n by force
have n_in_case: "n ∈ σ_vals" when A: "?A T'"
proof -
obtain t s' where t:
"insert⟨t,s'⟩ ∈ set (unlabel (transaction_strand T)) ∨
delete⟨t,s'⟩ ∈ set (unlabel (transaction_strand T))"
"Fun (Val n) [] = t ⋅ θ"
using A dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(4,5)
stateful_strand_step_mem_substD(4,5)[of _ _ _ θ]
subst_lsst_unlabel[of _ θ]
unfolding T'_def by (metis (no_types, opaque_lifting))
have "Fun (Val n) [] ∈ subst_range θ"
using t transaction_inserts_are_Value_vars(1)[OF T_wf(1,3), of t s']
transaction_deletes_are_Value_vars(1)[OF T_wf(1,3), of t s']
by force
hence "Fun (Val n) [] ∈ subst_range σ"
using transaction_decl_fresh_renaming_substs_range'(4)[
OF step.hyps(3,4,5) _ ξ_empty]
unfolding θ_def by blast
thus ?thesis unfolding σ_vals_def by fast
qed
have in_A_case: "𝒦 x ≠ Fun (Val n) []"
when y: "y ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" "ℐ x = ℐ y" "𝒦 x = 𝒥 y" for y
using A_cases
proof
assume "?A 𝒜" thus ?thesis
using B(14) m y(1,3) unfolding db_upds_consts_fresh_def y(2) by auto
next
assume "?A T'"
hence "n ∈ N'" using n_in_case unfolding N'_def by blast
moreover have "𝒥 y ∈ trms⇩l⇩s⇩s⇩t ℬ"
using B(6) y(1) m ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset
unfolding y(2) subst_in_ik_if_subst_pubval_def by blast
ultimately show ?thesis using B(12) y(3) by fastforce
qed
show "𝒦 x ≠ Fun (Val n) []"
proof (cases "x ∈ fv⇩l⇩s⇩s⇩t T'")
case True
note 0 = T'_fv_ℐ_pubval_𝒦_eq_δ_fresh_val[OF True m, unfolded B(9)[symmetric]]
note 1 = T'_fv_ℐ_pubval_𝒦_eq_𝒥_val[OF True m, unfolded B(9)[symmetric]]
show ?thesis
proof (cases "∀y∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'. ℐ y ≠ ℐ x")
case True show ?thesis
using A_cases 0[OF True] T_val_fresh_vals n_in_case
unfolding 𝒜_vals_def by force
next
case False
then obtain y where "y ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" "ℐ y = ℐ x" "𝒦 x = 𝒥 y" using 1 by blast
thus ?thesis using in_A_case by auto
qed
next
case False
hence x_in: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ X'" using x unfolding unlabel_append fv⇩s⇩s⇩t_append by fast
hence x_eq: "𝒦 x = 𝒥 x" using 𝒜_fv_𝒦_eq_𝒥 by blast
show ?thesis using in_A_case[OF x_in _ x_eq] by blast
qed
qed
have B'2: "?wt_model 𝒦 ℬ'"
proof (unfold welltyped_constraint_model_def; intro conjI)
have "Γ (𝒦 x) = Γ⇩v x" for x
proof -
have "wt⇩s⇩u⇩b⇩s⇩t 𝒥" "wt⇩s⇩u⇩b⇩s⇩t ℐ"
using B(2) 3(1) unfolding welltyped_constraint_model_def by (blast,blast)
hence *: "⋀y. Γ (𝒥 y) = Γ⇩v y" "⋀y. Γ (ℐ y) = Γ⇩v y" unfolding wt⇩s⇩u⇩b⇩s⇩t_def by auto
show ?thesis
proof (cases "x ∈ fv⇩l⇩s⇩s⇩t T'")
case True
note x = this
show ?thesis
proof (cases "∃n. ℐ x = Fun (PubConst Value n) []")
case True thus ?thesis using T'_fv_ℐ_pubval_𝒦_eq_val[OF x] T'_value_vars[OF x] by force
next
case False thus ?thesis using x * unfolding 𝒦_def by presburger
qed
next
case False thus ?thesis using *(1) unfolding 𝒦_def by presburger
qed
qed
thus "wt⇩s⇩u⇩b⇩s⇩t 𝒦" unfolding 𝒦_def wt⇩s⇩u⇩b⇩s⇩t_def by force
show "constraint_model 𝒦 ℬ'"
proof (unfold constraint_model_def; intro conjI)
have *: "strand_sem_stateful {} {} (unlabel 𝒜) ℐ"
"strand_sem_stateful {} {} (unlabel ℬ) 𝒥"
"interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "interpretation⇩s⇩u⇩b⇩s⇩t 𝒥"
"wf⇩t⇩r⇩m⇩s (subst_range ℐ)" "wf⇩t⇩r⇩m⇩s (subst_range 𝒥)"
using B(2) 3(1) unfolding welltyped_constraint_model_def constraint_model_def by fast+
show K0: "subst_domain 𝒦 = UNIV"
proof -
have "x ∈ subst_domain 𝒦" for x
proof (cases "x ∈ fv⇩l⇩s⇩s⇩t T'")
case True thus ?thesis
using T'_fv_ℐ_pubval_𝒦_eq_val[OF True] T'_fv_ℐ_val_𝒦_eq_𝒥[OF True] *(3)
unfolding subst_domain_def by (cases "∃n. ℐ x = Fun (PubConst Value n) []") auto
next
case False thus ?thesis using *(4) unfolding 𝒦_def subst_domain_def by auto
qed
thus ?thesis by blast
qed
have "fv (𝒦 x) = {}" for x
using interpretation_grounds_all[OF *(3)]
interpretation_grounds_all[OF *(4)]
unfolding 𝒦_def by simp
thus K1: "ground (subst_range 𝒦)" by simp
have "wf⇩t⇩r⇩m (Fun (Val n) [])" for n by fastforce
moreover have "wf⇩t⇩r⇩m (ℐ x)" "wf⇩t⇩r⇩m (𝒥 x)" for x using *(5,6) by (fastforce,fastforce)
ultimately have "wf⇩t⇩r⇩m (𝒦 x)" for x unfolding 𝒦_def by auto
thus K2: "wf⇩t⇩r⇩m⇩s (subst_range 𝒦)" by simp
show "strand_sem_stateful {} {} (unlabel ℬ') 𝒦"
proof (unfold ℬ'_def unlabel_append strand_sem_append_stateful Un_empty_left; intro conjI)
let ?sem = "λM D A. strand_sem_stateful M D (unlabel A) 𝒦"
let ?M1 = "ik⇩l⇩s⇩s⇩t ℬ ⋅⇩s⇩e⇩t 𝒦"
let ?M2 = "?M1 ∪ (ik⇩l⇩s⇩s⇩t T_val_constr ⋅⇩s⇩e⇩t 𝒦)"
let ?D1 = "dbupd⇩s⇩s⇩t (unlabel ℬ) 𝒦 {}"
let ?D2 = "dbupd⇩s⇩s⇩t (unlabel T_val_constr) 𝒦 ?D1"
show "?sem {} {} ℬ"
using 𝒥_𝒦_fv_ℬ_eq strand_sem_model_swap[OF _ *(2)] by blast
show "?sem ?M1 ?D1 T_val_constr"
using T_val_constr(3,4) strand_sem_stateful_if_no_send_or_check
unfolding list_all_iff by blast
have D2: "?D2 = ?D1 ∪ {(t ⋅ 𝒦, s ⋅ 𝒦) | t s. insert⟨t,s⟩ ∈ set (unlabel T_val_constr)}"
using T_val_constr(3,4) dbupd⇩s⇩s⇩t_no_deletes
unfolding list_all_iff by blast
have K3: "interpretation⇩s⇩u⇩b⇩s⇩t 𝒦"
using K0 K1 by argo
have rcv_θ_is_α: "t ⋅ θ = t ⋅ α"
when t: "(l,receive⟨ts⟩) ∈ set (transaction_receive T)" "t ∈ set ts" for l ts t
proof -
have "fv t ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T)"
using t(2) stateful_strand_step_fv_subset_cases(2)[OF unlabel_in[OF t(1)]] by auto
hence "t ⋅ σ = t" using t σ_dom σ_ran admissible_transactionE(12,13)[OF T_adm] by blast
thus ?thesis unfolding θ_def ξ_empty by simp
qed
have eq_θ_is_α: "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α"
when t: "(l,⟨ac: t ≐ s⟩) ∈ set (transaction_checks T)" for l ac t s
proof -
have "fv t ∪ fv s ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T)"
using stateful_strand_step_fv_subset_cases(3)[OF unlabel_in[OF t]] by auto
hence "t ⋅ σ = t" "s ⋅ σ = s"
using t σ_dom σ_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast)
thus "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α" unfolding θ_def ξ_empty by simp_all
qed
have noteq_θ_is_α: "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α"
when t: "(l,⟨t != s⟩) ∈ set (transaction_checks T)" for l t s
proof -
have "fv t ∪ fv s ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T)"
using stateful_strand_step_fv_subset_cases(8)[OF unlabel_in[OF t]] by auto
hence "t ⋅ σ = t" "s ⋅ σ = s"
using t σ_dom σ_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast)
thus "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α" unfolding θ_def ξ_empty by simp_all
qed
have in_θ_is_α: "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α"
when t: "(l,⟨ac: t ∈ s⟩) ∈ set (transaction_checks T)" for l ac t s
proof -
have "fv t ∪ fv s ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T)"
using stateful_strand_step_fv_subset_cases(6)[OF unlabel_in[OF t]] by auto
hence "t ⋅ σ = t" "s ⋅ σ = s"
using t σ_dom σ_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast)
thus "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α" unfolding θ_def ξ_empty by simp_all
qed
have notin_θ_is_α: "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α"
when t: "(l,⟨t not in s⟩) ∈ set (transaction_checks T)" for l t s
proof -
have "fv t ∪ fv s ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T)"
using stateful_strand_step_fv_subset_cases(9)[OF unlabel_in[OF t]] by auto
hence "t ⋅ σ = t" "s ⋅ σ = s"
using t σ_dom σ_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast)
thus "t ⋅ θ = t ⋅ α" "s ⋅ θ = s ⋅ α" unfolding θ_def ξ_empty by simp_all
qed
have T'_trm_no_val: "∄n. s = Fun (Val n) [] ∨ s = Fun (PubConst Value n) []"
when t: "t ∈ trms_transaction T" "s ⊑ t ⋅ α" for t s
proof -
have ?thesis when "s ⊑ t"
using that t admissible_transactions_no_Value_consts'[OF T_adm]
admissible_transactions_no_PubConsts[OF T_adm]
by blast
moreover have "Fun k [] ⊑ u" when "Fun k [] ⊑ u ⋅ α" for k u using that
proof (induction u)
case (Var x) thus ?case
using transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of x] by fastforce
qed auto
ultimately show ?thesis using t by blast
qed
define flt1 where "flt1 ≡ λA::('fun,'atom,'sets,'lbl) prot_constr.
filter is_Update (unlabel A)"
define flt2 where "flt2 ≡ λA::('fun,'atom,'sets,'lbl) prot_constr.
filter (λa. ∄l t ts. a = (l, insert⟨t,⟨s_val⟨ts⟩⟩⇩s⟩)) A"
define flt3 where "flt3 ≡ λA::(('fun,'atom,'sets,'lbl) prot_fun,
('fun,'atom,'sets,'lbl) prot_var) stateful_strand.
filter (λa. ∄t ts. a = insert⟨t,⟨s_val⟨ts⟩⟩⇩s⟩) A"
have flt2_subset: "set (unlabel (flt2 A)) ⊆ set (unlabel A)" for A
unfolding flt2_def unlabel_def by auto
have flt2_unlabel: "unlabel (flt2 A) = flt3 (unlabel A)" for A
unfolding flt2_def flt3_def by (induct A) auto
have flt2_suffix:
"suffix (filter (λa. ∄t ts. a = insert⟨t,⟨s_val⟨ts⟩⟩⇩s⟩) A) (unlabel (flt2 B))"
when "suffix A (unlabel B)" for A B
using that unfolding flt2_def by (induct B arbitrary: A rule: List.rev_induct) auto
have flt_AB: "flt1 (flt2 𝒜) = flt1 (flt2 ℬ)"
proof -
have *: "flt1 (flt2 𝒜) = filter is_Update (flt3 (unlabel 𝒜))"
"flt1 (flt2 ℬ) = filter is_Update (flt3 (unlabel ℬ))"
using flt2_unlabel unfolding flt1_def by presburger+
have **: "filter is_Update (flt3 C) = flt3 (filter is_Update C)" for C
proof (induction C)
case Nil thus ?case unfolding flt3_def by force
next
case (Cons c C) thus ?case unfolding flt3_def by (cases c) auto
qed
show ?thesis
proof (cases "T_upds = []")
case True
hence "filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ)"
using B(13) unfolding db_eq_def by fastforce
thus ?thesis using ** unfolding * by presburger
next
case False thus ?thesis
using B(13) unfolding flt1_def flt2_def db_eq_def Let_def by force
qed
qed
have A_setops_Fun: "∀t s. insert⟨t,s⟩ ∈ set (unlabel 𝒜) ⟶ (∃g ts. s = Fun g ts)"
using reachable_constraints_setops_form[OF step.hyps(1) P]
unfolding setops⇩s⇩s⇩t_def by fastforce
have A_insert_delete_not_subterm:
"ℐ x = 𝒦 x ∨ (¬(ℐ x ⊑ t) ∧ ¬(ℐ x ⊑ s) ∧ ¬(𝒦 x ⊑ t) ∧ ¬(𝒦 x ⊑ s))"
when x: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv⇩l⇩s⇩s⇩t T' ∪ fv t ∪ fv s"
and x_neq: "ℐ x ≠ 𝒦 x"
and ts: "insert⟨t,s⟩ ∈ set (unlabel 𝒜) ∨ delete⟨t,s⟩ ∈ set (unlabel 𝒜)"
for x t s
proof -
have x_in: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv⇩l⇩s⇩s⇩t T'"
using ts x stateful_strand_step_fv_subset_cases(4,5) by blast
note ts' = reachable_constraints_insert_delete_form[OF step.hyps(1) P ts]
have *: "ℐ x = 𝒦 x" when n: "ℐ x = Fun (Val n) []" for n
using n B'5 x_in
unfolding subst_eq_on_privvals_def unlabel_append fv⇩s⇩s⇩t_append
by blast
have **: "¬(ℐ x ⊑ t)" "¬(ℐ x ⊑ s)" "¬(𝒦 x ⊑ t)" "¬(𝒦 x ⊑ s)"
when n: "ℐ x = Fun (PubConst Value n) []" for n
proof -
show "¬(ℐ x ⊑ s)"
using ts'(1) x_in 3(2,3) unfolding valconst_cases_def by fastforce
show "¬(𝒦 x ⊑ s)"
using ts'(1) x_in B'3
unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append by force
show "¬(ℐ x ⊑ t)" using n ts'(3) by fastforce
from ts'(3) have "𝒦 x ≠ t"
proof
assume "∃y. t = Var y" thus ?thesis
using B'3 x_in unfolding valconsts_only_def by force
next
assume "∃k. t = Fun (Val k) []" thus ?thesis
using B'14 n x_in ts unfolding db_upds_consts_fresh_def by auto
qed
thus "¬(𝒦 x ⊑ t)" using ts'(3) by auto
qed
show ?thesis using * ** 3(2,3) x_in unfolding valconst_cases_def by fast
qed
have flt2_insert_in_iff:
"insert⟨u,v⟩ ∈ set (unlabel A) ⟷ insert⟨u,v⟩ ∈ set (unlabel (flt2 A))"
(is "?A A ⟷ ?B A")
when h: "s = ⟨h⟩⇩s" "h ≠ s_val" and t: "(t ⋅ I,s ⋅ I) = (u,v) ⋅⇩p I"
for t s h u v A and I::"('fun,'atom,'sets,'lbl) prot_subst"
proof
show "?B A ⟹ ?A A" using flt2_subset by fast
show "?A A ⟹ ?B A"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case
proof (cases "b = insert⟨u,v⟩")
case True thus ?thesis using h t unfolding a flt2_def by force
next
case False thus ?thesis using Cons.prems Cons.IH unfolding a flt2_def by auto
qed
qed simp
qed
have flt2_inset_iff:
"(t ⋅ 𝒦, s ⋅ 𝒦) ∈ dbupd⇩s⇩s⇩t (unlabel (flt2 ℬ)) 𝒦 {} ⟷
(t ⋅ 𝒦, s ⋅ 𝒦) ∈ dbupd⇩s⇩s⇩t (unlabel ℬ) 𝒦 {}"
(is "?A ⟷ ?B")
when h: "s = ⟨h⟩⇩s" "h ≠ s_val"
for t s h
proof
let ?C1 = "λu v B C. suffix (delete⟨u,v⟩#B) (unlabel C)"
let ?C2 = "λt s u v. (t,s) = (u,v) ⋅⇩p 𝒦"
let ?C3 = "λt s C. ∃u v. ?C2 t s u v ∧ insert⟨u,v⟩ ∈ set C"
let ?D = "λt s C. ∀u v B. ?C1 u v B C ∧ ?C2 t s u v ⟶ ?C3 t s B"
let ?db = "λC D. dbupd⇩s⇩s⇩t C 𝒦 D"
have "?C3 t s B"
when "?D t s (flt2 ℬ)" "?C1 u v B ℬ" "?C2 t s u v" for u v B t s
using that flt2_suffix flt2_subset by fastforce
thus "?A ⟹ ?B" using flt2_subset unfolding dbupd⇩s⇩s⇩t_in_iff by blast
show ?A when ?B using that
proof (induction ℬ rule: List.rev_induct)
case (snoc a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
have *:
"?db (unlabel (A@[a])) {} = ?db [b] (?db (unlabel A) {})"
"?db (unlabel (flt2 (A@[a]))) {} =
?db (unlabel (flt2 [a])) (?db (unlabel (flt2 A)) {})"
using dbupd⇩s⇩s⇩t_append[of _ _ 𝒦 "{}"] unfolding a flt2_def by auto
show ?case
proof (cases "∃u v. b = insert⟨u,v⟩ ∧ (t ⋅ 𝒦, s ⋅ 𝒦) = (u,v) ⋅⇩p 𝒦")
case True
then obtain u v where "b = insert⟨u,v⟩" "(t ⋅ 𝒦, s ⋅ 𝒦) = (u, v) ⋅⇩p 𝒦" by force
thus ?thesis using h *(2) unfolding a flt2_def by auto
next
case False
hence IH: "(t ⋅ 𝒦, s ⋅ 𝒦) ∈ dbupd⇩s⇩s⇩t (unlabel (flt2 A)) 𝒦 {}"
using snoc.prems snoc.IH unfolding *(1) by (cases b) auto
show ?thesis
proof (cases "is_Delete b")
case True
then obtain u v where b: "b = delete⟨u,v⟩" by (cases b) auto
have b': "unlabel (flt2 [a]) = [b]"
"unlabel (flt2 (A@[a])) = unlabel (flt2 A)@[b]"
unfolding a flt2_def b by (fastforce,fastforce)
have "(t ⋅ 𝒦, s ⋅ 𝒦) ≠ (u,v) ⋅⇩p 𝒦" using *(1) snoc.prems unfolding b' b by simp
thus ?thesis using *(2) IH unfolding b' b by simp
next
case False thus ?thesis using *(2) IH unfolding a flt2_def by (cases b) auto
qed
qed
qed simp
qed
have inset_model_swap:
"(t ⋅ ℐ, s ⋅ ℐ) ∈ dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {} ⟷
(t ⋅ 𝒦, s ⋅ 𝒦) ∈ dbupd⇩s⇩s⇩t (unlabel ℬ) 𝒦 {}"
(is "?in ℐ (unlabel 𝒜) ⟷ ?in 𝒦 (unlabel ℬ)")
when h: "s = ⟨h⟩⇩s"
"h ≠ s_val ∨ filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ)"
and t: "t = Var tx"
and t_s_fv: "fv t ∪ fv s ⊆ fv⇩l⇩s⇩s⇩t T'"
and q: "∀x ∈ fv t ∪ fv s.
ℐ x = 𝒦 x ∨ (¬(ℐ x ⊑ t) ∧ ¬(ℐ x ⊑ s) ∧ ¬(𝒦 x ⊑ t) ∧ ¬(𝒦 x ⊑ s))"
"∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv t ∪ fv s. ∃c. ℐ x = Fun c []"
"∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv t ∪ fv s. ∃c. 𝒦 x = Fun c []"
"∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv t ∪ fv s. ∀y ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv t ∪ fv s.
ℐ x = ℐ y ⟷ 𝒦 x = 𝒦 y"
for t s h tx
proof -
let ?upds = "λA. filter is_Update (unlabel A)"
have flt2_fv: "fv⇩l⇩s⇩s⇩t (flt2 𝒜) ⊆ fv⇩l⇩s⇩s⇩t 𝒜"
using fv⇩s⇩s⇩t_mono[OF flt2_subset[of 𝒜]] by blast
have upds_fv: "fv⇩s⇩s⇩t (?upds 𝒜) ⊆ fv⇩l⇩s⇩s⇩t 𝒜" by auto
have flt2_upds_fv: "fv⇩s⇩s⇩t (?upds (flt2 𝒜)) ⊆ fv⇩s⇩s⇩t (?upds 𝒜)"
using flt2_subset[of 𝒜] by auto
have h_neq: "Set h ≠ (Set s_val::('fun,'atom,'sets,'lbl) prot_fun)"
when "h ≠ s_val"
using that by simp
have *: "⋃(fv⇩p⇩a⇩i⇩r ` {}) = {}" "{} ⋅⇩p⇩s⇩e⇩t ℐ = {}" "{} ⋅⇩p⇩s⇩e⇩t 𝒦 = {}" by blast+
have "?in ℐ (?upds (flt2 𝒜)) ⟷ ?in 𝒦 (?upds (flt2 𝒜))"
proof
let ?X = "fv⇩s⇩s⇩t (?upds (flt2 𝒜)) ∪ fv t ∪ fv s ∪
⋃(fv⇩p⇩a⇩i⇩r ` ({}::(('fun,'atom,'sets,'lbl) prot_term ×
('fun,'atom,'sets,'lbl) prot_term) set))"
let ?q0 = "λδ θ.
∀x ∈ ?X.
δ x = θ x ∨
(¬(δ x ⊑ t) ∧ ¬(δ x ⊑ s) ∧¬(θ x ⊑ t) ∧ ¬(θ x ⊑ s) ∧
(∀(u,v) ∈ {}. ¬(δ x ⊑ u) ∧ ¬(δ x ⊑ v) ∧ ¬(θ x ⊑ u) ∧ ¬(θ x ⊑ v)) ∧
(∀u v. insert⟨u,v⟩ ∈ set (?upds (flt2 𝒜)) ∨
delete⟨u,v⟩ ∈ set (?upds (flt2 𝒜)) ⟶
¬(δ x ⊑ u) ∧ ¬(δ x ⊑ v) ∧ ¬(θ x ⊑ u) ∧ ¬(θ x ⊑ v)))"
let ?q1 = "λδ. ∀x ∈ ?X. ∃c. δ x = Fun c []"
let ?q2 = "λδ θ. ∀x ∈ ?X. ∀y ∈ ?X. δ x = δ y ⟷ θ x = θ y"
have q0: "?q0 ℐ 𝒦" "?q0 𝒦 ℐ"
proof -
have upd_ex:
"∃u v. x ∈ fv u ∪ fv v ∧
(insert⟨u,v⟩ ∈ set (?upds A) ∨ delete⟨u,v⟩ ∈ set (?upds A))"
when "x ∈ fv⇩s⇩s⇩t (?upds A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr"
using that
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case using Cons.IH Cons.prems unfolding a by (cases b) auto
qed simp
have "¬(ℐ x ⊑ t)" "¬(ℐ x ⊑ s)" "¬(𝒦 x ⊑ t)" "¬(𝒦 x ⊑ s)"
when x: "x ∈ fv⇩s⇩s⇩t (?upds (flt2 𝒜)) ∪ fv t ∪ fv s"
and x_neq: "ℐ x ≠ 𝒦 x"
for x
proof -
have "¬(ℐ x ⊑ t) ∧ ¬(ℐ x ⊑ s) ∧ ¬(𝒦 x ⊑ t) ∧ ¬(𝒦 x ⊑ s)"
proof (cases "x ∈ fv t ∪ fv s")
case True thus ?thesis using q(1) x_neq by blast
next
case False
hence "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" using x flt2_upds_fv upds_fv by blast
hence "∃n. 𝒦 x = Fun (Val n) []"
"∃n. ℐ x = Fun (Val n) [] ∨ ℐ x = Fun (PubConst Value n) []"
using B'3 3(2)
unfolding valconst_cases_def valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append
by (blast, blast)
thus ?thesis unfolding t h(1) by auto
qed
thus "¬(ℐ x ⊑ t)" "¬(ℐ x ⊑ s)" "¬(𝒦 x ⊑ t)" "¬(𝒦 x ⊑ s)" by simp_all
qed
moreover have "¬(ℐ x ⊑ u)" "¬(ℐ x ⊑ v)" "¬(𝒦 x ⊑ u)" "¬(𝒦 x ⊑ v)"
when x: "x ∈ fv⇩s⇩s⇩t (?upds (flt2 𝒜)) ∪ fv t ∪ fv s"
and x_neq: "ℐ x ≠ 𝒦 x"
and uv: "insert⟨u,v⟩ ∈ set (?upds (flt2 𝒜)) ∨
delete⟨u,v⟩ ∈ set (?upds (flt2 𝒜))"
for x u v
proof -
have uv': "insert⟨u,v⟩ ∈ set (unlabel 𝒜) ∨ delete⟨u,v⟩ ∈ set (unlabel 𝒜)"
using uv flt2_subset by auto
have x_in: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv⇩l⇩s⇩s⇩t T' ∪ fv u ∪ fv v"
using t_s_fv x flt2_upds_fv upds_fv by blast
show "¬(ℐ x ⊑ u)" "¬(ℐ x ⊑ v)" "¬(𝒦 x ⊑ u)" "¬(𝒦 x ⊑ v)"
using x_neq A_insert_delete_not_subterm[OF x_in x_neq uv'] by simp_all
qed
ultimately show "?q0 ℐ 𝒦" unfolding upd_ex unfolding *
by (metis (no_types, lifting) empty_iff sup_bot_right)
thus "?q0 𝒦 ℐ" by (metis (lifting) empty_iff)
qed
have q1: "?q1 ℐ" "?q1 𝒦"
using q(2,3) flt2_upds_fv upds_fv by (blast,blast)
have q2: "?q2 ℐ 𝒦" "?q2 𝒦 ℐ"
using q(4) flt2_upds_fv upds_fv unfolding * by (blast,blast)
show "?in ℐ (?upds (flt2 𝒜)) ⟹ ?in 𝒦 (?upds (flt2 𝒜))"
using dbupd⇩s⇩s⇩t_subst_const_swap[OF _ q0(1) q1(1,2) q2(1)] by force
show "?in 𝒦 (?upds (flt2 𝒜)) ⟹ ?in ℐ (?upds (flt2 𝒜))"
using dbupd⇩s⇩s⇩t_subst_const_swap[OF _ q0(2) q1(2,1) q2(2)] by force
qed
hence flt2_subst_swap: "?in ℐ (unlabel (flt2 𝒜)) ⟷ ?in 𝒦 (unlabel (flt2 𝒜))"
using dbupd⇩s⇩s⇩t_filter by blast
have "?in ℐ (?upds 𝒜) ⟷ ?in 𝒦 (?upds 𝒜)"
proof
let ?X = "fv⇩s⇩s⇩t (?upds 𝒜) ∪ fv t ∪ fv s ∪
⋃(fv⇩p⇩a⇩i⇩r ` ({}::(('fun,'atom,'sets,'lbl) prot_term ×
('fun,'atom,'sets,'lbl) prot_term) set))"
let ?q0 = "λδ θ.
∀x ∈ ?X.
δ x = θ x ∨
(¬(δ x ⊑ t) ∧ ¬(δ x ⊑ s) ∧¬(θ x ⊑ t) ∧ ¬(θ x ⊑ s) ∧
(∀(u,v) ∈ {}. ¬(δ x ⊑ u) ∧ ¬(δ x ⊑ v) ∧ ¬(θ x ⊑ u) ∧ ¬(θ x ⊑ v)) ∧
(∀u v. insert⟨u,v⟩ ∈ set (?upds 𝒜) ∨
delete⟨u,v⟩ ∈ set (?upds 𝒜) ⟶
¬(δ x ⊑ u) ∧ ¬(δ x ⊑ v) ∧ ¬(θ x ⊑ u) ∧ ¬(θ x ⊑ v)))"
let ?q1 = "λδ. ∀x ∈ ?X. ∃c. δ x = Fun c []"
let ?q2 = "λδ θ. ∀x ∈ ?X. ∀y ∈ ?X. δ x = δ y ⟷ θ x = θ y"
have q0: "?q0 ℐ 𝒦" "?q0 𝒦 ℐ"
proof -
have upd_ex:
"∃u v. x ∈ fv u ∪ fv v ∧
(insert⟨u,v⟩ ∈ set (?upds A) ∨ delete⟨u,v⟩ ∈ set (?upds A))"
when "x ∈ fv⇩s⇩s⇩t (?upds A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr"
using that
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?case using Cons.IH Cons.prems unfolding a by (cases b) auto
qed simp
have "¬(ℐ x ⊑ t)" "¬(ℐ x ⊑ s)" "¬(𝒦 x ⊑ t)" "¬(𝒦 x ⊑ s)"
when x: "x ∈ fv⇩s⇩s⇩t (?upds 𝒜) ∪ fv t ∪ fv s"
and x_neq: "ℐ x ≠ 𝒦 x"
for x
proof -
have "¬(ℐ x ⊑ t) ∧ ¬(ℐ x ⊑ s) ∧ ¬(𝒦 x ⊑ t) ∧ ¬(𝒦 x ⊑ s)"
proof (cases "x ∈ fv t ∪ fv s")
case True thus ?thesis using q(1) x_neq by blast
next
case False
hence "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" using x flt2_upds_fv upds_fv by blast
hence "∃n. 𝒦 x = Fun (Val n) []"
"∃n. ℐ x = Fun (Val n) [] ∨ ℐ x = Fun (PubConst Value n) []"
using B'3 3(2)
unfolding valconst_cases_def valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append
by (blast, blast)
thus ?thesis unfolding t h(1) by auto
qed
thus "¬(ℐ x ⊑ t)" "¬(ℐ x ⊑ s)" "¬(𝒦 x ⊑ t)" "¬(𝒦 x ⊑ s)" by simp_all
qed
moreover have "¬(ℐ x ⊑ u)" "¬(ℐ x ⊑ v)" "¬(𝒦 x ⊑ u)" "¬(𝒦 x ⊑ v)"
when x: "x ∈ fv⇩s⇩s⇩t (?upds 𝒜) ∪ fv t ∪ fv s"
and x_neq: "ℐ x ≠ 𝒦 x"
and uv: "insert⟨u,v⟩ ∈ set (?upds 𝒜) ∨
delete⟨u,v⟩ ∈ set (?upds 𝒜)"
for x u v
proof -
have uv': "insert⟨u,v⟩ ∈ set (unlabel 𝒜) ∨ delete⟨u,v⟩ ∈ set (unlabel 𝒜)"
using uv flt2_subset by auto
have x_in: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv⇩l⇩s⇩s⇩t T' ∪ fv u ∪ fv v"
using t_s_fv x flt2_upds_fv upds_fv by blast
show "¬(ℐ x ⊑ u)" "¬(ℐ x ⊑ v)" "¬(𝒦 x ⊑ u)" "¬(𝒦 x ⊑ v)"
using x_neq A_insert_delete_not_subterm[OF x_in x_neq uv'] by simp_all
qed
ultimately show "?q0 ℐ 𝒦" unfolding upd_ex unfolding *
by (metis (no_types, lifting) empty_iff sup_bot_right)
thus "?q0 𝒦 ℐ" by (metis (lifting) empty_iff)
qed
have q1: "?q1 ℐ" "?q1 𝒦"
using q(2,3) flt2_upds_fv upds_fv by (blast,blast)
have q2: "?q2 ℐ 𝒦" "?q2 𝒦 ℐ"
using q(4) flt2_upds_fv upds_fv unfolding * by (blast,blast)
show "?in ℐ (?upds 𝒜) ⟹ ?in 𝒦 (?upds 𝒜)"
using dbupd⇩s⇩s⇩t_subst_const_swap[OF _ q0(1) q1(1,2) q2(1)] by force
show "?in 𝒦 (?upds 𝒜) ⟹ ?in ℐ (?upds 𝒜)"
using dbupd⇩s⇩s⇩t_subst_const_swap[OF _ q0(2) q1(2,1) q2(2)] by force
qed
hence db_subst_swap:
"?in ℐ (unlabel 𝒜) ⟷ ?in 𝒦 (unlabel 𝒜)"
using dbupd⇩s⇩s⇩t_filter by blast
have "?in 𝒦 (unlabel ℬ)" when A: "?in ℐ (unlabel 𝒜)" using h(2)
proof
assume h': "h ≠ s_val"
have "?in ℐ (unlabel (flt2 𝒜))"
using A flt2_unlabel dbupd⇩s⇩s⇩t_set_term_neq_in_iff[OF h_neq[OF h'] A_setops_Fun]
unfolding h(1) flt3_def by simp
hence "?in 𝒦 (unlabel (flt2 𝒜))" using flt2_subst_swap by blast
hence "?in 𝒦 (flt1 (flt2 𝒜))" using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
hence "?in 𝒦 (flt1 (flt2 ℬ))" using flt_AB by simp
hence "?in 𝒦 (unlabel (flt2 ℬ))" using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
thus ?thesis using flt2_inset_iff[OF h(1) h'] by fast
next
assume h': "filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ)"
have "?in 𝒦 (unlabel 𝒜)" using A db_subst_swap by blast
hence "?in 𝒦 (flt1 𝒜)" using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
hence "?in 𝒦 (flt1 ℬ)" using h' unfolding flt1_def by simp
thus ?thesis using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
qed
moreover have "¬?in 𝒦 (unlabel ℬ)" when A: "¬?in ℐ (unlabel 𝒜)" using h(2)
proof
assume h': "h ≠ s_val"
have "¬?in ℐ (unlabel (flt2 𝒜))"
using A flt2_unlabel dbupd⇩s⇩s⇩t_set_term_neq_in_iff[OF h_neq[OF h'] A_setops_Fun]
unfolding h(1) flt3_def by simp
hence "¬?in 𝒦 (unlabel (flt2 𝒜))" using flt2_subst_swap by blast
hence "¬?in 𝒦 (flt1 (flt2 𝒜))" using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
hence "¬?in 𝒦 (flt1 (flt2 ℬ))" using flt_AB by simp
hence "¬?in 𝒦 (unlabel (flt2 ℬ))" using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
thus ?thesis using flt2_inset_iff[OF h(1) h'] by fast
next
assume h': "filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ)"
have "¬?in 𝒦 (unlabel 𝒜)" using A db_subst_swap by blast
hence "¬?in 𝒦 (flt1 𝒜)" using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
hence "¬?in 𝒦 (flt1 ℬ)" using h' unfolding flt1_def by simp
thus ?thesis using dbupd⇩s⇩s⇩t_filter unfolding flt1_def by blast
qed
ultimately show ?thesis by blast
qed
have "?M2 ⊢ t ⋅ θ ⋅ 𝒦"
when ts: "(l, receive⟨ts⟩) ∈ set (transaction_receive T)" "t ∈ set ts" for l t ts
proof -
have *: "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ θ ⋅ ℐ" using 5 ts by blast
note tθα = rcv_θ_is_α[OF ts]
have t_T'_trm: "t ∈ trms_transaction T"
using trms⇩s⇩s⇩t_memI(2)[OF unlabel_in[OF ts(1)] ts(2)]
unfolding trms_transaction_unfold by blast
have t_T'_trm': "t ⋅ θ ∈ trms⇩l⇩s⇩s⇩t T'"
using trms⇩s⇩s⇩t_memI(2)[
OF stateful_strand_step_subst_inI(2)[
OF unlabel_in[OF ts(1)], unfolded unlabel_subst]]
ts(2)
unfolding T'_def trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq trms_transaction_subst_unfold by auto
note t_no_val = T'_trm_no_val[OF t_T'_trm, unfolded tθα[symmetric]]
have t_fv_T': "fv (t ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'"
using ts(2) stateful_strand_step_fv_subset_cases(2)[
OF stateful_strand_step_subst_inI(2)[OF unlabel_in[OF ts(1)], of θ]]
unfolding T'_def unlabel_subst fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq fv_transaction_subst_unfold
by auto
have ik_B_fv_subset: "fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ⊆ fv⇩l⇩s⇩s⇩t ℬ"
by (meson UnE fv_ik⇩s⇩s⇩t_is_fv⇩s⇩s⇩t subset_iff)
let ?fresh_vals = "(λn. Fun (Val n) []) ` T_val_fresh_vals"
have q0: "ik⇩l⇩s⇩s⇩t ℬ ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ θ ⋅ ℐ" using * B(10) by (blast intro: ideduct_mono)
have q1: "∀x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ∪ fv (t ⋅ θ). valconst_cases ℐ x"
using 3(2,3) t_fv_T' ik_B_fv_subset unfolding B(9) by blast
have q2: "∀x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ∪ fv (t ⋅ θ). ∃n. 𝒦 x = Fun (Val n) []"
using B'3 t_fv_T' ik_B_fv_subset
unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append B(9)
by blast
have T_val_constr_ik:
"∃M. ik⇩l⇩s⇩s⇩t T_val_constr = M ∪ ?fresh_vals"
"∃M. ik⇩l⇩s⇩s⇩t T_val_constr ⋅⇩s⇩e⇩t 𝒦 = (M ⋅⇩s⇩e⇩t 𝒦) ∪ ?fresh_vals"
proof -
obtain M where M: "ik⇩l⇩s⇩s⇩t T_val_constr = M ∪ ?fresh_vals"
using T_val_constr(8) by blast
have "?fresh_vals ⋅⇩s⇩e⇩t 𝒦 = ?fresh_vals" by fastforce
thus "∃M. ik⇩l⇩s⇩s⇩t T_val_constr = M ∪ ?fresh_vals"
"∃M. ik⇩l⇩s⇩s⇩t T_val_constr ⋅⇩s⇩e⇩t 𝒦 = (M ⋅⇩s⇩e⇩t 𝒦) ∪ ?fresh_vals"
using M by (fastforce, fastforce)
qed
have "𝒦 x ∈ ik⇩l⇩s⇩s⇩t ℬ ∪ ik⇩l⇩s⇩s⇩t T_val_constr"
when x: "x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ∪ fv (t ⋅ θ)" "ℐ x = Fun (PubConst Value n) []" for x n
using x(1) B'6'[OF _ x(2)] B'6''[OF _ x(2)] t_fv_T' ik_B_fv_subset
unfolding B(9) unlabel_append fv⇩s⇩s⇩t_append by blast
hence q3: "∀x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ∪ fv (t ⋅ θ).
(∃n. ℐ x = Fun (PubConst Value n) []) ⟶ 𝒦 x ∈ ik⇩l⇩s⇩s⇩t ℬ ∪ ?fresh_vals"
using T_val_constr_ik(1) T_val_constr(8) q2
unfolding B(9) ℬ'_def unlabel_append ik⇩s⇩s⇩t_append fv⇩s⇩s⇩t_append
by (metis (no_types, lifting) UnE UnI1 UnI2 image_iff mem_Collect_eq)
have q4: "∀x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ∪ fv (t ⋅ θ). (∃n. ℐ x = Fun (Val n) []) ⟶ ℐ x = 𝒦 x"
using B'5 t_fv_T' ik_B_fv_subset
unfolding subst_eq_on_privvals_def B(9) unlabel_append fv⇩s⇩s⇩t_append
by blast
have q5: "∀x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ∪ fv (t ⋅ θ). ∀y ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t ℬ) ∪ fv (t ⋅ θ).
ℐ x = ℐ y ⟷ 𝒦 x = 𝒦 y"
using B'7 t_fv_T' ik_B_fv_subset
unfolding subst_eq_iff_def B(9) unlabel_append fv⇩s⇩s⇩t_append
by blast
have q6: "∀n. ¬(Fun (PubConst Value n) [] ⊑⇩s⇩e⇩t insert (t ⋅ θ) (ik⇩l⇩s⇩s⇩t ℬ))"
proof -
have "∄n. s = Fun (PubConst Value n) []" when s: "s ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ'" for s
proof -
have "f ≠ PubConst Value n" when f: "f ∈ funs_term s" for f n
using f s reachable_constraints_val_funs_private(1)[OF B'1 P, of f]
unfolding is_PubConstValue_def is_PubConst_def the_PubConst_type_def
by (metis (mono_tags, lifting) UN_I funs_term_subterms_eq(2) prot_fun.simps(85))
thus ?thesis by fastforce
qed
moreover have "ik⇩l⇩s⇩s⇩t ℬ ⊆ trms⇩l⇩s⇩s⇩t ℬ'"
using ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset unfolding ℬ'_def unlabel_append trms⇩s⇩s⇩t_append by blast
ultimately show ?thesis
using t_no_val by blast
qed
show ?thesis
using deduct_val_const_swap[OF q0 q1[unfolded valconst_cases_def] q2 q3 q4 q5 q6]
T_val_constr_ik(2)
by (blast intro: ideduct_mono)
qed
moreover have "t ⋅ θ ⋅ 𝒦 = s ⋅ θ ⋅ 𝒦"
when ts: "(l, ⟨ac: t ≐ s⟩) ∈ set (transaction_checks T)" for l ac t s
proof -
have q0: "t ⋅ θ ⋅ ℐ = s ⋅ θ ⋅ ℐ" using 5 ts by blast
have "fv⇩s⇩s⇩t⇩p (⟨ac: (t ⋅ θ) ≐ (s ⋅ θ)⟩) ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)"
using stateful_strand_step_fv_subset_cases(3)[
OF stateful_strand_step_subst_inI(3)[OF unlabel_in[OF ts], of θ]]
unfolding unlabel_subst by simp
hence t_s_fv: "fv (t ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'" "fv (s ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'"
unfolding T'_def fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq fv_transaction_subst_unfold[of T θ]
by (fastforce, fastforce)
have "t ∈ trms_transaction T" "s ∈ trms_transaction T"
using trms⇩s⇩s⇩t_memI(3,4)[OF unlabel_in[OF ts]]
unfolding trms_transaction_unfold by (blast, blast)
hence "∄n. u = Fun (Val n) [] ∨ u = Fun (PubConst Value n) []"
when u: "u ⊑ t ⋅ θ ∨ u ⊑ s ⋅ θ" for u
using u T'_trm_no_val unfolding eq_θ_is_α[OF ts] by blast
hence "¬(ℐ x ⊑ t ⋅ θ)" "¬(ℐ x ⊑ s ⋅ θ)"
when x: "x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ)" for x
using x t_s_fv I' by (fast, fast)
hence q1:
"∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ℐ x = 𝒦 x ∨ (¬(ℐ x ⊑ t ⋅ θ) ∧ ¬(ℐ x ⊑ s ⋅ θ))"
by blast
have q2: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. ℐ x = Fun c []"
using t_s_fv 3(3) unfolding valconst_cases_def by blast
have q3: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. 𝒦 x = Fun c []"
using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append by blast
have q4: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∀y ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
ℐ x = ℐ y ⟷ 𝒦 x = 𝒦 y"
using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv⇩s⇩s⇩t_append by blast
show ?thesis by (rule subst_const_swap_eq'[OF q0 q1 q2 q3 q4])
qed
moreover have "t ⋅ θ ⋅ 𝒦 ≠ s ⋅ θ ⋅ 𝒦"
when ts: "(l, ⟨t != s⟩) ∈ set (transaction_checks T)" for l t s
proof -
have q0: "t ⋅ θ ⋅ ℐ ≠ s ⋅ θ ⋅ ℐ" using 5 ts by blast
have "fv⇩s⇩s⇩t⇩p (⟨(t ⋅ θ) != (s ⋅ θ)⟩) ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)"
using stateful_strand_step_fv_subset_cases(8)[
OF stateful_strand_step_subst_inI(8)[OF unlabel_in[OF ts], of θ]]
unfolding unlabel_subst by simp
hence t_s_fv: "fv (t ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'" "fv (s ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'"
unfolding T'_def fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq fv_transaction_subst_unfold[of T θ]
by (fastforce, fastforce)
have "t ∈ trms_transaction T" "s ∈ trms_transaction T"
using trms⇩s⇩s⇩t_memI(9)[OF unlabel_in[OF ts]]
unfolding trms_transaction_unfold by auto
hence "∄n. u = Fun (Val n) []" when u: "u ⊑ t ⋅ θ ∨ u ⊑ s ⋅ θ" for u
using u T'_trm_no_val unfolding noteq_θ_is_α[OF ts] by blast
hence "¬(𝒦 x ⊑ t ⋅ θ)" "¬(𝒦 x ⊑ s ⋅ θ)"
when x: "x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ)" for x
using x t_s_fv B'3
unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append
by (fast, fast)
hence q1: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). 𝒦 x = ℐ x ∨ (¬(𝒦 x ⊑ t ⋅ θ) ∧ ¬(𝒦 x ⊑ s ⋅ θ))"
by blast
have q2: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. 𝒦 x = Fun c []"
using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append by blast
have q3: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. ℐ x = Fun c []"
using t_s_fv 3(3) unfolding valconst_cases_def by blast
have q4: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∀y ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
𝒦 x = 𝒦 y ⟷ ℐ x = ℐ y"
using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv⇩s⇩s⇩t_append by blast
show ?thesis using q0 subst_const_swap_eq'[OF _ q1 q2 q3 q4] by fast
qed
moreover have "(t ⋅ θ ⋅ 𝒦, s ⋅ θ ⋅ 𝒦) ∈ ?D2"
when ts: "(l, ⟨ac: t ∈ s⟩) ∈ set (transaction_checks T)" for l ac t s
proof -
have s_neq_s_val:
"s ≠ ⟨s_val⟩⇩s ∨ filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ)"
proof (cases "T_upds = []")
case False thus ?thesis
using step.hyps(2) ts x_val(7)
unfolding transaction_strand_def
by (cases ac) fastforce+
qed (use B(13)[unfolded db_eq_def] in simp)
have ts': "⟨ac: t ∈ s⟩ ∈ set (unlabel (transaction_strand T))"
using ts unlabel_in[OF ts] unfolding transaction_strand_def by fastforce
have "fv⇩s⇩s⇩t⇩p (⟨ac: (t ⋅ θ) ∈ (s ⋅ θ)⟩) ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)"
using stateful_strand_step_fv_subset_cases(6)[
OF stateful_strand_step_subst_inI(6)[OF unlabel_in[OF ts], of θ]]
unfolding unlabel_subst by simp
hence t_s_fv: "fv (t ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'" "fv (s ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'"
unfolding T'_def fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq fv_transaction_subst_unfold[of T θ]
by (fastforce, fastforce)
have "t ∈ trms_transaction T" "s ∈ trms_transaction T"
using ts' unfolding trms⇩s⇩s⇩t_def by (force, force)
hence "∄n. u = Fun (Val n) [] ∨ u = Fun (PubConst Value n) []"
when u: "u ⊑ t ⋅ θ ∨ u ⊑ s ⋅ θ" for u
using u T'_trm_no_val unfolding in_θ_is_α[OF ts] by blast
hence "¬(𝒦 x ⊑ t ⋅ θ)" "¬(𝒦 x ⊑ s ⋅ θ)" "¬(ℐ x ⊑ t ⋅ θ)" "¬(ℐ x ⊑ s ⋅ θ)"
when x: "x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ)" for x
using x t_s_fv B'3 I'
unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append
by (fast,fast,fast,fast)
hence q1: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
ℐ x = 𝒦 x ∨
(¬(ℐ x ⊑ t ⋅ θ) ∧ ¬(ℐ x ⊑ s ⋅ θ) ∧ ¬(𝒦 x ⊑ t ⋅ θ) ∧ ¬(𝒦 x ⊑ s ⋅ θ))"
by blast
have q2: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. ℐ x = Fun c []"
using t_s_fv 3(2,3) unfolding valconst_cases_def by blast
have q3: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. 𝒦 x = Fun c []"
using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append by blast
have q4: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
∀y ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
ℐ x = ℐ y ⟷ 𝒦 x = 𝒦 y"
using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv⇩s⇩s⇩t_append by blast
obtain h tx where s: "s = ⟨h⟩⇩s" and tx: "t = Var tx"
using ts' transaction_selects_are_Value_vars[OF T_wf(1,2), of t s]
transaction_inset_checks_are_Value_vars[OF T_adm, of t s]
by (cases ac) auto
have h:
"s ⋅ θ = ⟨h⟩⇩s"
"h ≠ s_val ∨ filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ)"
using s s_neq_s_val by (simp,blast)
obtain ty where ty: "t ⋅ θ = Var ty"
using tx transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of tx]
unfolding in_θ_is_α[OF ts] by force
have "(t ⋅ θ ⋅ ℐ, s ⋅ θ ⋅ ℐ) ∈ dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}" using 5 ts by blast
hence "(t ⋅ θ ⋅ 𝒦, s ⋅ θ ⋅ 𝒦) ∈ dbupd⇩s⇩s⇩t (unlabel ℬ) 𝒦 {}"
using inset_model_swap[OF h ty _ q1 q2 q3 q4] t_s_fv by simp
thus ?thesis unfolding D2 by blast
qed
moreover have "(t ⋅ θ ⋅ 𝒦, s ⋅ θ ⋅ 𝒦) ∉ ?D2"
when ts: "(l, ⟨t not in s⟩) ∈ set (transaction_checks T)" for l t s
proof -
have s_neq_s_val:
"(T_upds ≠ [] ∧ s ≠ ⟨s_val⟩⇩s) ∨
(T_upds = [] ∧ filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ))"
proof (cases "T_upds = []")
case False thus ?thesis
using step.hyps(2) ts x_val(7) unfolding transaction_strand_def by force
qed (use B(13)[unfolded db_eq_def] in simp)
have ts': "⟨t not in s⟩ ∈ set (unlabel (transaction_strand T))"
using ts unlabel_in[OF ts] unfolding transaction_strand_def by fastforce
have "fv⇩s⇩s⇩t⇩p (⟨(t ⋅ θ) not in (s ⋅ θ)⟩) ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)"
using stateful_strand_step_fv_subset_cases(9)[
OF stateful_strand_step_subst_inI(9)[OF unlabel_in[OF ts], of θ]]
unfolding unlabel_subst by simp
hence t_s_fv: "fv (t ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'" "fv (s ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'"
unfolding T'_def fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq fv_transaction_subst_unfold[of T θ]
by (fastforce, fastforce)
have "t ∈ trms_transaction T" "s ∈ trms_transaction T"
using ts' unfolding trms⇩s⇩s⇩t_def by (force, force)
hence "∄n. u = Fun (Val n) [] ∨ u = Fun (PubConst Value n) []"
when u: "u ⊑ t ⋅ θ ∨ u ⊑ s ⋅ θ" for u
using u T'_trm_no_val unfolding notin_θ_is_α[OF ts] by blast
hence "¬(𝒦 x ⊑ t ⋅ θ)" "¬(𝒦 x ⊑ s ⋅ θ)" "¬(ℐ x ⊑ t ⋅ θ)" "¬(ℐ x ⊑ s ⋅ θ)"
when x: "x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ)" for x
using x t_s_fv B'3 I'
unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append
by (fast,fast,fast,fast)
hence q1: "∀x ∈ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
ℐ x = 𝒦 x ∨
(¬(ℐ x ⊑ t ⋅ θ) ∧ ¬(ℐ x ⊑ s ⋅ θ) ∧ ¬(𝒦 x ⊑ t ⋅ θ) ∧ ¬(𝒦 x ⊑ s ⋅ θ))"
by blast
have q2: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. ℐ x = Fun c []"
using t_s_fv 3(2,3) unfolding valconst_cases_def by blast
have q3: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ). ∃c. 𝒦 x = Fun c []"
using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv⇩s⇩s⇩t_append by blast
have q4: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
∀y ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∪ fv (t ⋅ θ) ∪ fv (s ⋅ θ).
ℐ x = ℐ y ⟷ 𝒦 x = 𝒦 y"
using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv⇩s⇩s⇩t_append by blast
obtain h tx where s: "s = ⟨h⟩⇩s" and tx: "t = Var tx"
using transaction_notinset_checks_are_Value_vars(1,2)[OF T_adm ts', of t s] by auto
have h:
"s ⋅ θ = ⟨h⟩⇩s"
"h ≠ s_val ∨ filter is_Update (unlabel 𝒜) = filter is_Update (unlabel ℬ)"
"T_upds ≠ [] ⟹ h ≠ s_val"
using s s_neq_s_val by (simp,blast,blast)
obtain ty where ty: "t ⋅ θ = Var ty"
using tx transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of tx]
unfolding notin_θ_is_α[OF ts] by force
have *: "(t ⋅ θ ⋅ 𝒦, s ⋅ θ ⋅ 𝒦) ≠ (u ⋅ 𝒦, v ⋅ 𝒦)"
when u: "insert⟨u,v⟩ ∈ set (unlabel T_val_constr)"
and h': "h ≠ s_val"
for u v
proof -
have "v = ⟨s_val⟩⇩s" using T_val_constr(9) unlabel_mem_has_label[OF u] by force
thus ?thesis using h(1) h' by simp
qed
have "(t ⋅ θ ⋅ ℐ, s ⋅ θ ⋅ ℐ) ∉ dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}" using 5 ts by blast
hence **: "(t ⋅ θ ⋅ 𝒦, s ⋅ θ ⋅ 𝒦) ∉ dbupd⇩s⇩s⇩t (unlabel ℬ) 𝒦 {}"
using inset_model_swap[OF h(1,2) ty _ q1 q2 q3 q4] t_s_fv by simp
show ?thesis
proof (cases "T_upds = []")
case True
have "dbupd⇩s⇩s⇩t (unlabel T_val_constr) I D = D" for I D
using T_val_constr_no_upds_if_no_T_upds[OF True]
dbupd⇩s⇩s⇩t_filter[of "unlabel T_val_constr"]
by force
thus ?thesis using ** by simp
next
case False thus ?thesis
using ** * h(3) T_val_constr_no_upds_if_no_T_upds unfolding D2 by blast
qed
qed
ultimately show "?sem ?M2 ?D2 T'"
unfolding T'_def 4[OF T_adm K3 K2] by blast
qed
qed
qed
show ?case
using B'1 B'2 B'3 B'4 B'5 B'6 B'7 B'8_9 B'10_11 B'12 B'13 B'14
unfolding θ_def[symmetric] T'_def[symmetric] by blast
qed
obtain ℬ 𝒥 where B:
"ℬ ∈ reachable_constraints P" "?wt_model 𝒥 ℬ"
"∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜. ∃n. 𝒥 x = Fun (Val n) []" "?rcv_att 𝒜 n ⟶ ?rcv_att ℬ n" "fv⇩l⇩s⇩s⇩t 𝒜 = fv⇩l⇩s⇩s⇩t ℬ"
using lmm[OF finite.emptyI _ _ finite.emptyI] unfolding valconsts_only_def by auto
show ?thesis
using B(1,3) welltyped_constraint_model_attack_if_receive_attack[OF B(2)] B(4) 2
unfolding wt_attack_def B(5) by (meson list.set_intros(1))
qed
private lemma add_occurs_msgs_soundness_aux2:
assumes P: "∀T ∈ set P. admissible_transaction T"
and A: "𝒜 ∈ reachable_constraints P"
shows "∃ℬ ∈ reachable_constraints (map add_occurs_msgs P). 𝒜 = rm_occurs_msgs_constr ℬ"
using A
proof (induction rule: reachable_constraints.induct)
case (step 𝒜 T ξ σ α)
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
let ?A' = "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"
let ?B' = "dual⇩l⇩s⇩s⇩t (transaction_strand (add_occurs_msgs T) ⋅⇩l⇩s⇩s⇩t θ)"
obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp
have P': "∀T ∈ set P. admissible_transaction' T"
"∀T ∈ set P. admissible_transaction_no_occurs_msgs T"
using P admissible_transactionE'(1,2) by (blast,blast)
note T_adm' = bspec[OF P step.hyps(2)]
note T_adm = bspec[OF P'(1) step.hyps(2)]
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)]
note T_fresh_val = admissible_transactionE(2)[OF T_adm]
note T_no_occ = admissible_transactionE'(2)[OF T_adm']
obtain ℬ where B:
"ℬ ∈ reachable_constraints (map add_occurs_msgs P)" "𝒜 = rm_occurs_msgs_constr ℬ"
using step.IH by blast
note 0 = add_occurs_msgs_cases[OF T]
note 1 = add_occurs_msgs_vars_eq[OF bspec[OF P'(1)]]
note 2 = add_occurs_msgs_trms[of T]
note 3 = add_occurs_msgs_transaction_strand_set[OF T]
have 4: "add_occurs_msgs T ∈ set (map add_occurs_msgs P)"
using step.hyps(2) by simp
have 5: "transaction_decl_subst ξ (add_occurs_msgs T)"
using step.hyps(3) 0(4) unfolding transaction_decl_subst_def by argo
have "t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t ℬ)"
"t ∉ subterms⇩s⇩e⇩t (trms_transaction (add_occurs_msgs T))"
when t: "t ∈ subst_range σ" for t
proof -
obtain c where c: "t = Fun (Val c) []"
using t T_fresh_val transaction_fresh_subst_domain[OF step.hyps(4)]
transaction_fresh_subst_sends_to_val[OF step.hyps(4), of _ thesis]
by fastforce
have *: "t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" "t ∉ subterms⇩s⇩e⇩t (trms_transaction T)"
using t step.hyps(4) unfolding transaction_fresh_subst_def by (fast,fast)
have "t ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜 ∨ (∃x ∈ fv⇩l⇩s⇩s⇩t 𝒜. t ⊑ occurs (Var x)) ∨
(∃c. Fun c [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜 ∧ t ⊑ occurs (Fun c []))"
when t: "t ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t ℬ"
using t rm_occurs_msgs_constr_reachable_constraints_trms_cases[OF P' B(2,1)] by fast
thus "t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t ℬ)"
using *(1) unfolding c by fastforce
show "t ∉ subterms⇩s⇩e⇩t (trms_transaction (add_occurs_msgs T))"
using *(2) unfolding 2 c by force
qed
hence 6: "transaction_fresh_subst σ (add_occurs_msgs T) (trms⇩l⇩s⇩s⇩t ℬ)"
using step.hyps(4) unfolding transaction_fresh_subst_def 0(5) 2 by fast
have 7: "transaction_renaming_subst α (map add_occurs_msgs P) (vars⇩l⇩s⇩s⇩t ℬ)"
using step.hyps(5) rm_occurs_msgs_constr_reachable_constraints_vars_eq[OF P' B(1)] B(2) 1(5)
unfolding transaction_renaming_subst_def by simp
have "?A' = rm_occurs_msgs_constr ?B'"
using admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)]
rm_occurs_msgs_constr_transaction_strand''[OF T_adm T_no_occ]
unfolding θ_def[symmetric] by metis
hence 8: "𝒜@?A' = rm_occurs_msgs_constr (ℬ@?B')"
by (metis rm_occurs_msgs_constr_append B(2))
show ?case using reachable_constraints.step[OF B(1) 4 5 6 7] 8 unfolding θ_def by blast
qed (metis reachable_constraints.init rm_occurs_msgs_constr.simps(1))
private lemma add_occurs_msgs_soundness_aux3:
assumes P: "∀T ∈ set P. admissible_transaction T"
and A: "𝒜 ∈ reachable_constraints (map add_occurs_msgs P)"
"welltyped_constraint_model ℐ (rm_occurs_msgs_constr 𝒜)"
and I: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜. ∃n. ℐ x = Fun (Val n) []" (is "?I 𝒜")
shows "welltyped_constraint_model ℐ 𝒜" (is "?Q ℐ 𝒜")
using A I
proof (induction 𝒜 rule: reachable_constraints.induct)
case (step 𝒜 T ξ σ α)
let ?f = rm_occurs_msgs_constr
let ?sem = "λB. strand_sem_stateful (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) (dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}) (unlabel B) ℐ"
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
define ℬ where "ℬ ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"
obtain T' where T': "T' ∈ set P" "T = add_occurs_msgs T'"
using step.hyps(2) by fastforce
then obtain A' B' C' D' E' F' where T'': "T' = Transaction A' B' C' D' E' F'"
using prot_transaction.exhaust by blast
have P': "∀T ∈ set (map add_occurs_msgs P). admissible_transaction' T"
"∀T ∈ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T"
"∀T ∈ set P. admissible_transaction' T"
"∀T ∈ set P. admissible_transaction_no_occurs_msgs T"
using P admissible_transactionE' add_occurs_msgs_admissible_occurs_checks
by (fastforce,fastforce,fastforce,fastforce)
note T_adm = bspec[OF P'(1) step.hyps(2)]
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
note T'_adm = bspec[OF P'(3) T'(1)]
note T'_no_occ = bspec[OF P'(4) T'(1)]
note T'_wf = admissible_transaction_is_wellformed_transaction(1)[OF T'_adm]
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)]
note T_fresh_val = admissible_transactionE(2)[OF T_adm]
have 0: "?Q ℐ (?f 𝒜)" "?I 𝒜" "?I ℬ"
by (metis step.prems(1) welltyped_constraint_model_prefix rm_occurs_msgs_constr_append,
simp_all add: step.prems(2) θ_def ℬ_def)
note IH = step.IH[OF 0(1,2)]
have I': "wt⇩s⇩u⇩b⇩s⇩t ℐ" "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ)"
using step.prems(1) unfolding welltyped_constraint_model_def constraint_model_def by blast+
have 1: "∀x ∈ fv_transaction T. ∄t. θ x = occurs t"
"∀x ∈ fv_transaction T. θ x ≠ Fun OccursSec []"
using admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)]
unfolding θ_def[symmetric] by simp_all
have "(ik⇩l⇩s⇩s⇩t (?f 𝒜) ⋅⇩s⇩e⇩t ℐ) ∪ (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) = ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ"
using rm_occurs_msgs_constr_ik_subset by fast
hence 2: "?sem (?f ℬ)"
using step.prems(1) strand_sem_append_stateful[of "{}" "{}" "unlabel (?f 𝒜)" "unlabel (?f ℬ)"]
rm_occurs_msgs_constr_dbupd⇩s⇩s⇩t_eq[of 𝒜 ℐ "{}"] rm_occurs_msgs_constr_append[of 𝒜 ℬ]
strand_sem_ik_mono_stateful[of "ik⇩l⇩s⇩s⇩t (?f 𝒜) ⋅⇩s⇩e⇩t ℐ" _ _ _ "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ"]
unfolding welltyped_constraint_model_def constraint_model_def θ_def[symmetric] ℬ_def[symmetric]
by auto
note 3 = rm_occurs_msgs_constr_transaction_strand''[
OF T'_adm T'_no_occ 1[unfolded T'(2)], unfolded ℬ_def[symmetric] T'(2)[symmetric]]
note 4 = add_occurs_msgs_cases[OF T'', unfolded T'(2)[symmetric]]
define xs where "xs ≡ fv_list⇩s⇩s⇩t (unlabel (transaction_strand T'))"
define flt where "flt ≡ filter (λx. x ∉ set (transaction_fresh T'))"
define occs where "occs ≡ map (λx. occurs (Var x)::('fun,'atom,'sets,'lbl) prot_term)"
note 6 = add_occurs_msgs_transaction_strand_cases(7,8,9)[
of T' θ, unfolded xs_def[symmetric] flt_def[symmetric] occs_def[symmetric]
T'(2)[symmetric]]
have 7: "x ∈ fv_transaction T - set (transaction_fresh T)"
when x: "x ∈ set (flt xs)" for x
using that fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t add_occurs_msgs_vars_eq(3,9)[OF T'_adm]
unfolding xs_def flt_def T'(2) by force
have 9: "∃y. θ x = Var y"
when x: "x ∈ set (flt xs)" for x
proof -
have *: "x ∉ fst ` set (transaction_decl T ())"
using admissible_transactionE(1)[OF T_adm] by simp
have **: "x ∉ set (transaction_fresh T)" using 7[OF x] by simp
show ?thesis
using transaction_decl_fresh_renaming_substs_range(4)[OF step.hyps(3,4,5) * **]
unfolding θ_def by blast
qed
have 8: "∃y ∈ fv⇩l⇩s⇩s⇩t ℬ. θ x = Var y"
when x: "x ∈ set (flt xs)" for x
proof -
note * = 7[OF x]
obtain y where y: "θ x = Var y" using 9[OF x] by blast
have "x ∈ fv⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T))" by (metis * Diff_iff fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq)
have "θ x ∈ θ ` fv⇩l⇩s⇩s⇩t (transaction_strand T)" using * by fast
hence "fv (θ x) ⊆ fv⇩s⇩e⇩t (θ ` fv_transaction T)" by force
hence "fv (θ x) ⊆ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"
using fv⇩s⇩s⇩t_subst_if_no_bvars[OF admissible_transactionE(4)[OF T_adm], of θ]
by (metis unlabel_subst)
hence "fv (θ x) ⊆ fv⇩l⇩s⇩s⇩t ℬ" by (metis fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq ℬ_def)
thus ?thesis using y by simp
qed
have ℬ_var_is_ℐ_val: "∃n. ℐ x = Fun (Val n) []" when x: "x ∈ fv⇩l⇩s⇩s⇩t ℬ" for x
using step.prems(2) x unfolding ℬ_def[symmetric] θ_def[symmetric] by auto
have T'_var_is_θℐ_val: "∃n. θ x ⋅ ℐ = Fun (Val n) []" when x: "x ∈ set (flt xs)" for x
using 8[OF x] ℬ_var_is_ℐ_val by force
have poschecks_has_occ: "occurs (Fun (Val n) []) ∈ ik⇩l⇩s⇩s⇩t 𝒜"
when x: "⟨ac: t ∈ s⟩ ∈ set (unlabel ℬ)"
and n: "t ⋅ ℐ = Fun (Val n) []"
for ac t s n
proof -
have *: "(t ⋅ ℐ, s ⋅ ℐ) ∈ dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}"
proof -
obtain t' s' where t':
"⟨ac: t' ∈ s'⟩ ∈ set (unlabel (transaction_checks T'))" "t = t' ⋅ θ" "s = s' ⋅ θ"
using 4(8) x stateful_strand_step_mem_substD(6)
wellformed_transaction_strand_unlabel_memberD(10)[OF T_wf(1)]
dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(6)
unfolding ℬ_def by (metis (no_types) unlabel_subst dual⇩l⇩s⇩s⇩t_subst)
have "(t' ⋅ θ ⋅ ℐ, s' ⋅ θ ⋅ ℐ) ∈ dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}"
using t'(1) 2
wellformed_transaction_unlabel_sem_iff[
OF T'_wf(1) I'(2,3), of "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" "dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}" θ]
unfolding 3 by blast
thus ?thesis using t'(2,3) by simp
qed
have "t' ∈ trms⇩l⇩s⇩s⇩t 𝒜"
when "insert⟨t',s'⟩ ∈ set (unlabel 𝒜)" for t' s'
using that by force
have "t ⋅ ℐ ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜"
proof -
obtain t' s' where t': "insert⟨t',s'⟩ ∈ set (unlabel 𝒜)" "t ⋅ ℐ = t' ⋅ ℐ" "s ⋅ ℐ = s' ⋅ ℐ"
using * db⇩s⇩s⇩t_in_cases[of "t ⋅ ℐ" "s ⋅ ℐ" "unlabel 𝒜" ℐ "[]"]
db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel 𝒜" ℐ "[]"]
by auto
have t'': "t' = t ⋅ ℐ ∨ (∃y ∈ fv⇩l⇩s⇩s⇩t 𝒜. t' = Var y ∧ ℐ y = t ⋅ ℐ)"
using t'(1,2) stateful_strand_step_fv_subset_cases(4)
unfolding n by (cases t') (force,force)
thus ?thesis
proof
assume "t' = t ⋅ ℐ" thus ?thesis using t'(1) by force
next
assume "∃y ∈ fv⇩l⇩s⇩s⇩t 𝒜. t' = Var y ∧ ℐ y = t ⋅ ℐ"
then obtain y where y: "y ∈ fv⇩l⇩s⇩s⇩t 𝒜" "ℐ y = t ⋅ ℐ" by blast
have "Γ⇩v y = TAtom Value"
using y(2) wt_subst_trm''[OF I'(1), of "Var y"] unfolding n by simp
hence "∃B. prefix B 𝒜 ∧ t ⋅ ℐ ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t B"
by (metis y constraint_model_Value_var_in_constr_prefix[OF step.hyps(1) IH P'(1,2)])
thus ?thesis unfolding prefix_def by auto
qed
qed
thus ?thesis
using reachable_constraints_occurs_fact_ik_case'[OF step.hyps(1) P'(1,2)]
unfolding n by blast
qed
have snds_has_occ: "occurs (Fun (Val n) []) ∈ ik⇩l⇩s⇩s⇩t 𝒜"
when ts: "send⟨ts⟩ ∈ set (unlabel ℬ)"
and n: "Fun (Val n) [] ⊑⇩s⇩e⇩t set ts ⋅⇩s⇩e⇩t ℐ"
for ts n
proof -
have "receive⟨ts⟩ ∈ set (unlabel (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))"
using ts dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(2) unfolding ℬ_def by metis
then obtain ts' where ts':
"receive⟨ts'⟩ ∈ set (unlabel (transaction_strand T))" "ts = ts' ⋅⇩l⇩i⇩s⇩t θ"
by (metis subst_lsst_memD(1) unlabel_in unlabel_mem_has_label)
have "?sem (dual⇩l⇩s⇩s⇩t (transaction_receive T' ⋅⇩l⇩s⇩s⇩t θ))"
using 2 strand_sem_append_stateful[of "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" "dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}"]
unfolding 3 transaction_dual_subst_unlabel_unfold by blast
moreover have "list_all is_Receive (unlabel (transaction_receive T'))"
using T'_wf unfolding wellformed_transaction_def by blast
hence "list_all is_Send (unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T' ⋅⇩l⇩s⇩s⇩t θ)))"
by (metis subst_lsst_unlabel subst_sst_list_all(2) dual⇩l⇩s⇩s⇩t_list_all(1))
hence "ik⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_receive T' ⋅⇩l⇩s⇩s⇩t θ)) = {}"
using in_ik⇩s⇩s⇩t_iff unfolding list_all_iff is_Send_def by fast
ultimately have *: "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ ℐ"
when "send⟨ts⟩ ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T' ⋅⇩l⇩s⇩s⇩t θ)))" "t ∈ set ts"
for t ts
using strand_sem_stateful_sends_deduct[OF _ that] by simp
hence *: "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ θ ⋅ ℐ"
when ts: "receive⟨ts⟩ ∈ set (unlabel (transaction_receive T'))" "t ∈ set ts" for t ts
using ts(2) dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(2)[of "ts ⋅⇩l⇩i⇩s⇩t θ" "transaction_receive T' ⋅⇩l⇩s⇩s⇩t θ"]
stateful_strand_step_subst_inI(2)[OF ts(1), of θ, unfolded unlabel_subst]
by auto
have **: "set (flt xs) = fv_transaction T' - set (transaction_fresh T')"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t unfolding flt_def xs_def by fastforce
have rcv_case: ?thesis
when "ts = ts' ⋅⇩l⇩i⇩s⇩t θ" "Fun (Val n) [] ⊑⇩s⇩e⇩t set ts ⋅⇩s⇩e⇩t ℐ"
"receive⟨ts'⟩ ∈ set (unlabel (transaction_receive T'))"
for ts ts'
using that * reachable_constraints_occurs_fact_ik_case''[OF step.hyps(1) IH P'(1,2)] by auto
have "receive⟨ts'⟩ ∈ set (unlabel (transaction_receive T))"
using wellformed_transaction_strand_unlabel_memberD(1)[OF T_wf] ts'(1) by blast
hence "(ts' = map (λx. occurs (Var x)) (flt xs) ∧ ts' ≠ []) ∨
receive⟨ts'⟩ ∈ set (unlabel (transaction_receive T'))"
(is "?A ∨ ?B")
using ** ts'(1) add_occurs_msgs_cases(13)[OF T'']
unfolding T'(2)[symmetric] xs_def[symmetric] flt_def[symmetric] by force
thus ?thesis
proof
assume ?A
then obtain x where x: "x ∈ set (flt xs)" "Fun (Val n) [] ⊑ θ x ⋅ ℐ"
using ts' n by fastforce
have x': "θ x ⋅ ℐ = Fun (Val n) []" "x ∈ fv_transaction T" "x ∉ set (transaction_fresh T)"
"x ∈ fv_transaction T'" "x ∉ set (transaction_fresh T')"
using x(2) T'_var_is_θℐ_val[OF x(1)] 7[OF x(1)] ** x(1) by fastforce+
let ?snds = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T' ⋅⇩l⇩s⇩s⇩t θ))"
let ?chks = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_checks T' ⋅⇩l⇩s⇩s⇩t θ))"
have B_subsets: "set ?chks ⊆ set (unlabel ℬ)"
unfolding ℬ_def transaction_dual_subst_unlabel_unfold 4(8) by fastforce
from admissible_transaction_fv_in_receives_or_selects_dual_subst[OF T'_adm x'(4,5), of θ]
show ?thesis
proof
assume "∃ts. send⟨ts⟩ ∈ set ?snds ∧ θ x ⊑⇩s⇩e⇩t set ts"
then obtain ss where ss: "send⟨ss⟩ ∈ set ?snds" "θ x ⊑⇩s⇩e⇩t set ss" by blast
obtain ss' where ss':
"ss = ss' ⋅⇩l⇩i⇩s⇩t θ" "receive⟨ss'⟩ ∈ set (unlabel (transaction_receive T'))"
by (metis ss(1) dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(2) subst_lsst_memD(1)
unlabel_in unlabel_mem_has_label)
show ?thesis
using rcv_case[OF ss'(1) _ ss'(2)] subst_subterms[OF ss(2), of ℐ] x'(1) by argo
qed (use B_subsets poschecks_has_occ[OF _ x'(1)] in blast)
qed (metis rcv_case[OF ts'(2) n])
qed
have "occurs (θ x ⋅ ℐ) ∈ ik⇩l⇩s⇩s⇩t 𝒜" when x: "x ∈ set (flt xs)" for x
proof -
have "(∃ac s. ⟨ac: θ x ∈ s⟩ ∈ set (unlabel ℬ)) ∨
(∃ts. send⟨ts⟩ ∈ set (unlabel ℬ) ∧ θ x ⊑⇩s⇩e⇩t set ts)"
(is "(∃ac s. ?A ac s) ∨ (∃ts. ?B1 ts ∧ ?B2 ts)")
using 7[OF x] admissible_transaction_fv_in_receives_or_selects_dual_subst[OF T_adm, of x θ]
unfolding ℬ_def transaction_dual_subst_unlabel_unfold by auto
thus ?thesis
proof
assume "∃ac s. ?A ac s"
then obtain ac s where s: "?A ac s" by blast
show ?thesis using poschecks_has_occ[OF s] T'_var_is_θℐ_val[OF x] by force
next
assume "∃ts. ?B1 ts ∧ ?B2 ts"
then obtain ts where ts: "?B1 ts" "?B2 ts" by meson
have ts': "θ x ⋅ ℐ ⊑⇩s⇩e⇩t set ts ⋅⇩s⇩e⇩t ℐ" by (metis ts(2) subst_subterms)
show ?thesis using snds_has_occ[OF ts(1)] ts' T'_var_is_θℐ_val[OF x] by force
qed
qed
hence "occurs (θ x ⋅ ℐ) ⋅ ℐ ∈ ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" when "x ∈ set (flt xs)" for x using that by fast
moreover have "occurs (θ x ⋅ ℐ) ⋅ ℐ = occurs (θ x ⋅ ℐ)" for x
using subst_ground_ident[OF interpretation_grounds[OF I'(2), of "θ x"], of ℐ] by simp
ultimately have "occurs (θ x ⋅ ℐ) ∈ ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" when "x ∈ set (flt xs)" for x
using that by auto
hence "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ ℐ" when "t ∈ set (occs (flt xs) ⋅⇩l⇩i⇩s⇩t θ)" for t
using that unfolding occs_def by auto
hence occs_sem: "?sem [⟨⋆, send⟨occs (flt xs) ⋅⇩l⇩i⇩s⇩t θ⟩⟩]"
by auto
have "?sem ℬ"
proof -
let ?IK = "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ"
let ?DB = "dbupd⇩s⇩s⇩t (unlabel 𝒜) ℐ {}"
let ?snds = "dual⇩l⇩s⇩s⇩t (transaction_receive T' ⋅⇩l⇩s⇩s⇩t θ)"
let ?snds_occs = "(⟨⋆, send⟨occs (flt xs) ⋅⇩l⇩i⇩s⇩t θ⟩⟩)#?snds"
let ?chks = "dual⇩l⇩s⇩s⇩t (transaction_checks T' ⋅⇩l⇩s⇩s⇩t θ)"
let ?upds = "dual⇩l⇩s⇩s⇩t (transaction_updates T' ⋅⇩l⇩s⇩s⇩t θ)"
let ?rcvs = "dual⇩l⇩s⇩s⇩t (transaction_send T' ⋅⇩l⇩s⇩s⇩t θ)"
note * = strand_sem_append_stateful[of _ _ _ _ ℐ]
note ** = transaction_dual_subst_unlabel_unfold
have ***: "⋀M. M ∪ (ik⇩s⇩s⇩t [] ⋅⇩s⇩e⇩t ℐ) = M"
"⋀D. dbupd⇩s⇩s⇩t [] ℐ D = D"
by simp_all
have snds_sem:
"?sem ?snds"
"?sem ?snds_occs"
using 2 occs_sem *[of ?IK ?DB]
unfolding 3 ** by (blast, fastforce)
have "list_all is_Receive (unlabel (transaction_receive T'))"
using T'_wf unfolding wellformed_transaction_def by blast
hence "list_all is_Send (unlabel ?snds)" "list_all is_Send (unlabel ?snds_occs)"
using subst_sst_list_all(2) unlabel_subst dual⇩l⇩s⇩s⇩t_list_all(1)
by (metis, metis (no_types) list.pred_inject(2) stateful_strand_step.disc(1) unlabel_Cons(1))
hence "∀a ∈ set (unlabel ?snds). ¬is_Receive a ∧ ¬is_Insert a ∧ ¬is_Delete a"
"∀a ∈ set (unlabel ?snds_occs). ¬is_Receive a ∧ ¬is_Insert a ∧ ¬is_Delete a"
unfolding list_all_iff by (blast,blast)
hence snds_no_upds:
"ik⇩l⇩s⇩s⇩t ?snds ⋅⇩s⇩e⇩t ℐ = {}"
"dbupd⇩s⇩s⇩t (unlabel ?snds) ℐ ?DB = ?DB"
"ik⇩l⇩s⇩s⇩t (?snds_occs) ⋅⇩s⇩e⇩t ℐ = {}"
"dbupd⇩s⇩s⇩t (unlabel ?snds_occs) ℐ ?DB = ?DB"
by (metis ik⇩s⇩s⇩t_snoc_no_receive_empty, metis dbupd⇩s⇩s⇩t_no_upd,
metis ik⇩s⇩s⇩t_snoc_no_receive_empty, metis dbupd⇩s⇩s⇩t_no_upd)
have chks_sem:
"?sem ?chks"
using 2 snds_no_upds *
unfolding 3 ** by auto
have "list_all is_Check_or_Assignment (unlabel (transaction_checks T'))"
using T'_wf unfolding wellformed_transaction_def by blast
hence "list_all is_Check_or_Assignment (unlabel ?chks)"
by (metis (no_types) subst_sst_list_all(11) unlabel_subst dual⇩l⇩s⇩s⇩t_list_all(11))
hence "∀a ∈ set (unlabel ?chks). ¬is_Receive a ∧ ¬is_Insert a ∧ ¬is_Delete a"
unfolding list_all_iff by blast
hence chks_no_upds:
"ik⇩l⇩s⇩s⇩t ?chks ⋅⇩s⇩e⇩t ℐ = {}"
"dbupd⇩s⇩s⇩t (unlabel ?chks) ℐ ?DB = ?DB"
by (metis ik⇩s⇩s⇩t_snoc_no_receive_empty, metis dbupd⇩s⇩s⇩t_no_upd)
have upds_sem:
"?sem ?upds"
using 2 snds_no_upds chks_no_upds *
unfolding 3 ** by auto
have "list_all is_Send (unlabel (transaction_send T'))"
using T'_wf unfolding wellformed_transaction_def by fast
hence "list_all is_Send (unlabel (transaction_send T' ⋅⇩l⇩s⇩s⇩t θ))"
by (metis (no_types, opaque_lifting) subst_sst_list_all(1) unlabel_subst)
hence rcvs_is_rcvs: "list_all is_Receive (unlabel ?rcvs)"
using dual⇩l⇩s⇩s⇩t_list_all(2) by blast
have rcvs_sem: "strand_sem_stateful M D (unlabel rcvs) ℐ"
when "list_all is_Receive (unlabel rcvs)"
for M D and rcvs::"('fun, 'atom, 'sets, 'lbl) prot_strand"
using rcvs_is_rcvs strand_sem_receive_prepend_stateful[of M D "[]" ℐ, OF _ that] by auto
have B_sem: "?sem (?snds@?chks@?upds@rcvs)"
"?sem (?snds_occs@?chks@?upds@rcvs)"
when "list_all is_Receive (unlabel rcvs)" for rcvs
using strand_sem_append_stateful[of _ _ _ _ ℐ]
snds_sem snds_no_upds chks_sem chks_no_upds
upds_sem rcvs_sem[OF that]
by (force, force)
show ?thesis
proof (cases "transaction_fresh T' = []")
case True
show ?thesis using B_sem[OF rcvs_is_rcvs] unfolding ℬ_def 6(1)[OF True] by force
next
case False
note F = this
show ?thesis
proof (cases "∃ts F'. transaction_send T' = ⟨⋆, send⟨ts⟩⟩#F'")
case True
obtain ts F' rcvs' where F':
"transaction_send T' = ⟨⋆, send⟨ts⟩⟩#F'"
"ℬ = (if flt xs = [] then ?snds else ?snds_occs)@?chks@?upds@rcvs'"
"rcvs' = ⟨⋆, receive⟨occs (transaction_fresh T')@ts ⋅⇩l⇩i⇩s⇩t θ⟩⟩#dual⇩l⇩s⇩s⇩t (F' ⋅⇩l⇩s⇩s⇩t θ)"
using 6(3)[OF F True] unfolding ℬ_def by blast
have *: "list_all is_Receive (unlabel rcvs')"
using rcvs_is_rcvs dual⇩l⇩s⇩s⇩t_Cons(1)[of ⋆ ts "F' ⋅⇩l⇩s⇩s⇩t θ"]
subst_lsst_cons[of "⟨⋆, send⟨ts⟩⟩" F' θ]
unfolding F'(1,3) list_all_iff by auto
show ?thesis using B_sem[OF *] unfolding F'(2) by fastforce
next
case False
have *:
"list_all is_Receive (unlabel (⟨⋆, receive⟨occs (transaction_fresh T') ⋅⇩l⇩i⇩s⇩t θ⟩⟩#?rcvs))"
using rcvs_is_rcvs by auto
show ?thesis using B_sem[OF *] unfolding ℬ_def 6(2)[OF F False] by fastforce
qed
qed
qed
thus ?case
using IH strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" "unlabel ℬ" ℐ]
unfolding welltyped_constraint_model_def constraint_model_def θ_def[symmetric] ℬ_def[symmetric]
by simp
qed simp
theorem add_occurs_msgs_soundness:
defines "wt_attack ≡ λℐ 𝒜 l n. welltyped_constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])"
assumes P: "∀T ∈ set P. admissible_transaction T"
"has_initial_value_producing_transaction P"
and A: "𝒜 ∈ reachable_constraints P" "wt_attack ℐ 𝒜 l n"
shows "∃ℬ ∈ reachable_constraints (map add_occurs_msgs P). ∃𝒥. wt_attack 𝒥 ℬ l n"
proof -
have P': "∀T ∈ set (map add_occurs_msgs P). admissible_transaction' T"
"∀T ∈ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T"
"∀T ∈ set P. admissible_transaction' T"
"∀T ∈ set P. admissible_transaction_no_occurs_msgs T"
using P admissible_transactionE' add_occurs_msgs_admissible_occurs_checks
by (fastforce,fastforce,fastforce,fastforce)
obtain 𝒜' 𝒥 where A':
"𝒜' ∈ reachable_constraints P" "wt_attack 𝒥 𝒜' l n" "∀x∈fv⇩l⇩s⇩s⇩t 𝒜'. ∃n. 𝒥 x = Fun (Val n) []"
using add_occurs_msgs_soundness_aux1[OF P'(3) P(2) A[unfolded wt_attack_def]]
unfolding wt_attack_def by blast
have J: "welltyped_constraint_model 𝒥 𝒜'"
using A'(2) welltyped_constraint_model_prefix
unfolding wt_attack_def by blast
obtain ℬ where B:
"ℬ ∈ reachable_constraints (map add_occurs_msgs P)" "𝒜' = rm_occurs_msgs_constr ℬ"
using add_occurs_msgs_soundness_aux2[OF P(1) A'(1)] by blast
have J': "welltyped_constraint_model 𝒥 ℬ"
using add_occurs_msgs_soundness_aux3[OF P(1) B(1) J[unfolded B(2)]]
A'(3) rm_occurs_msgs_constr_reachable_constraints_fv_eq[OF P'(3,4) B(1)]
unfolding wt_attack_def B(2) by blast
obtain ts where ts: "receive⟨ts⟩ ∈ set (unlabel ℬ)" "attack⟨n⟩ ∈ set ts"
using reachable_constraints_receive_attack_if_attack''[OF P'(3) A'(1,2)[unfolded wt_attack_def]]
rm_occurs_msgs_constr_receive_attack_iff[of n ℬ]
unfolding B(2)[symmetric] by auto
have J'': "wt_attack 𝒥 ℬ l n"
using welltyped_constraint_model_attack_if_receive_attack[OF J' ts]
unfolding wt_attack_def by fast
show ?thesis
using B(1) J'' by blast
qed
end
end
subsection ‹Automatically Checking Protocol Security in a Typed Model›
context stateful_protocol_model
begin
definition abs_intruder_knowledge (‹α⇩i⇩k›) where
"α⇩i⇩k S ℐ ≡ (ik⇩l⇩s⇩s⇩t S ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t α⇩0 (db⇩l⇩s⇩s⇩t S ℐ)"
definition abs_value_constants (‹α⇩v⇩a⇩l⇩s›) where
"α⇩v⇩a⇩l⇩s S ℐ ≡ {t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t S) ⋅⇩s⇩e⇩t ℐ. ∃n. t = Fun (Val n) []} ⋅⇩α⇩s⇩e⇩t α⇩0 (db⇩l⇩s⇩s⇩t S ℐ)"
definition abs_term_implications (‹α⇩t⇩i›) where
"α⇩t⇩i 𝒜 T θ ℐ ≡ {(s,t) | s t x.
s ≠ t ∧ x ∈ fv_transaction T ∧ x ∉ set (transaction_fresh T) ∧
Fun (Abs s) [] = θ x ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) ∧
Fun (Abs t) [] = θ x ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)) ℐ)}"
lemma abs_intruder_knowledge_append:
"α⇩i⇩k (A@B) ℐ =
(ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t α⇩0 (db⇩l⇩s⇩s⇩t (A@B) ℐ) ∪
(ik⇩l⇩s⇩s⇩t B ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t α⇩0 (db⇩l⇩s⇩s⇩t (A@B) ℐ)"
by (metis unlabel_append abs_set_union image_Un ik⇩s⇩s⇩t_append abs_intruder_knowledge_def)
lemma abs_value_constants_append:
fixes A B::"('a,'b,'c,'d) prot_strand"
shows "α⇩v⇩a⇩l⇩s (A@B) ℐ =
{t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t A) ⋅⇩s⇩e⇩t ℐ. ∃n. t = Fun (Val n) []} ⋅⇩α⇩s⇩e⇩t α⇩0 (db⇩l⇩s⇩s⇩t (A@B) ℐ) ∪
{t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B) ⋅⇩s⇩e⇩t ℐ. ∃n. t = Fun (Val n) []} ⋅⇩α⇩s⇩e⇩t α⇩0 (db⇩l⇩s⇩s⇩t (A@B) ℐ)"
proof -
define a0 where "a0 ≡ α⇩0 (db⇩s⇩s⇩t (unlabel (A@B)) ℐ)"
define M where "M ≡ λa::('a,'b,'c,'d) prot_strand.
{t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t a) ⋅⇩s⇩e⇩t ℐ. ∃n. t = Fun (Val n) []}"
have "M (A@B) = M A ∪ M B"
using image_Un[of "λx. x ⋅ ℐ" "subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t A)" "subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)"]
unfolding M_def unlabel_append[of A B] trms⇩s⇩s⇩t_append[of "unlabel A" "unlabel B"] by blast
hence "M (A@B) ⋅⇩α⇩s⇩e⇩t a0 = (M A ⋅⇩α⇩s⇩e⇩t a0) ∪ (M B ⋅⇩α⇩s⇩e⇩t a0)" by (simp add: abs_set_union)
thus ?thesis unfolding abs_value_constants_def a0_def M_def by force
qed
lemma transaction_renaming_subst_has_no_pubconsts_abss:
fixes α::"('fun,'atom,'sets,'lbl) prot_subst"
assumes "transaction_renaming_subst α P A"
shows "subst_range α ∩ pubval_terms = {}" (is ?A)
and "subst_range α ∩ abs_terms = {}" (is ?B)
proof -
{ fix t assume "t ∈ subst_range α"
then obtain x where "t = Var x"
using transaction_renaming_subst_is_renaming(1)[OF assms]
by force
hence "t ∉ pubval_terms" "t ∉ abs_terms" by simp_all
} thus ?A ?B by auto
qed
lemma transaction_fresh_subst_has_no_pubconsts_abss:
fixes σ::"('fun,'atom,'sets,'lbl) prot_subst"
assumes "transaction_fresh_subst σ T 𝒜" "∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value"
shows "subst_range σ ∩ pubval_terms = {}" (is ?A)
and "subst_range σ ∩ abs_terms = {}" (is ?B)
proof -
{ fix t assume "t ∈ subst_range σ"
then obtain x where "x ∈ set (transaction_fresh T)" "σ x = t"
using assms(1) unfolding transaction_fresh_subst_def by auto
then obtain n where "t = Fun (Val n) []"
using transaction_fresh_subst_sends_to_val[OF assms(1)] assms(2)
by meson
hence "t ∉ pubval_terms" "t ∉ abs_terms" unfolding is_PubConstValue_def by simp_all
} thus ?A ?B by auto
qed
lemma reachable_constraints_GSMP_no_pubvals_abss:
assumes "𝒜 ∈ reachable_constraints P"
and P: "∀T ∈ set P. admissible_transaction' T"
and ℐ: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "wt⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ)"
"∀n. PubConst Value n ∉ ⋃(funs_term ` (ℐ ` fv⇩l⇩s⇩s⇩t 𝒜))"
"∀n. Abs n ∉ ⋃(funs_term ` (ℐ ` fv⇩l⇩s⇩s⇩t 𝒜))"
shows "trms⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊆ GSMP (⋃T ∈ set P. trms_transaction T) - (pubval_terms ∪ abs_terms)"
(is "?A ⊆ ?B")
using assms(1) ℐ(4,5)
proof (induction 𝒜 rule: reachable_constraints.induct)
case (step 𝒜 T ξ σ α)
define trms_P where "trms_P ≡ (⋃T ∈ set P. trms_transaction T)"
define T' where "T' ≡ transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"
have ξ_elim: "ξ ∘⇩s σ ∘⇩s α = σ ∘⇩s α"
using admissible_transaction_decl_subst_empty[OF bspec[OF P step.hyps(2)] step.hyps(3)]
by simp
note T_fresh = admissible_transactionE(2)[OF bspec[OF P step.hyps(2)]]
note T_no_bvars = admissible_transactionE(4)[OF bspec[OF P step.hyps(2)]]
have T_no_PubVal: "∀T ∈ set P. ∀n. PubConst Value n ∉ ⋃(funs_term ` trms_transaction T)"
and T_no_Abs: "∀T ∈ set P. ∀n. Abs n ∉ ⋃(funs_term ` trms_transaction T)"
using admissible_transactions_no_Value_consts''[OF bspec[OF P]] by metis+
have ℐ': "∀n. PubConst Value n ∉ ⋃ (funs_term ` (ℐ ` fv⇩l⇩s⇩s⇩t 𝒜))"
"∀n. Abs n ∉ ⋃ (funs_term ` (ℐ ` fv⇩l⇩s⇩s⇩t 𝒜))"
using step.prems fv⇩s⇩s⇩t_append[of "unlabel 𝒜"] unlabel_append[of 𝒜]
by auto
have "wt⇩s⇩u⇩b⇩s⇩t (rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α))" for X
using wt_subst_rm_vars[of "ξ ∘⇩s σ ∘⇩s α" "set X"]
transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)]
by metis
hence wt: "wt⇩s⇩u⇩b⇩s⇩t ((rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α)) ∘⇩s ℐ)" for X
using ℐ(2) wt_subst_compose by fast
have wftrms: "wf⇩t⇩r⇩m⇩s (subst_range ((rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α)) ∘⇩s ℐ))" for X
using wf_trms_subst_compose[OF wf_trms_subst_rm_vars' ℐ(3)]
transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)]
by fast
have "trms⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t T') ⋅⇩s⇩e⇩t ℐ ⊆ ?B"
proof
fix t assume "t ∈ trms⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t T') ⋅⇩s⇩e⇩t ℐ"
hence "t ∈ trms⇩l⇩s⇩s⇩t T' ⋅⇩s⇩e⇩t ℐ" using trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq by blast
then obtain s X where s:
"s ∈ trms_transaction T"
"t = s ⋅ rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α) ∘⇩s ℐ"
"set X ⊆ bvars_transaction T"
using trms⇩s⇩s⇩t_unlabel_subst'' unfolding T'_def by blast
define θ where "θ ≡ rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α)"
note 0 = pubval_terms_subst_range_comp[OF
transaction_fresh_subst_has_no_pubconsts_abss(1)[OF step.hyps(4) T_fresh]
transaction_renaming_subst_has_no_pubconsts_abss(1)[OF step.hyps(5)]]
abs_terms_subst_range_comp[OF
transaction_fresh_subst_has_no_pubconsts_abss(2)[OF step.hyps(4) T_fresh]
transaction_renaming_subst_has_no_pubconsts_abss(2)[OF step.hyps(5)]]
have 1: "s ∈ trms_P" using step.hyps(2) s(1) unfolding trms_P_def by auto
have s_nin: "s ∉ pubval_terms" "s ∉ abs_terms"
using 1 T_no_PubVal T_no_Abs funs_term_Fun_subterm
unfolding trms_P_def is_PubConstValue_def is_Abs_def is_PubConst_def
by (fastforce, blast)
have 2: "(ℐ ` fv⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t T')) ∩ pubval_terms = {}"
"(ℐ ` fv⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t T')) ∩ abs_terms = {}"
"subst_range (ξ ∘⇩s σ ∘⇩s α) ∩ pubval_terms = {}"
"subst_range (ξ ∘⇩s σ ∘⇩s α) ∩ abs_terms = {}"
using 0 step.prems funs_term_Fun_subterm
unfolding T'_def θ_def ξ_elim
by (fastforce simp add: is_PubConstValue_def is_PubConst_def,
fastforce simp add: is_Abs_def,
argo, argo)
have "subst_range θ ⊆ subst_range (ξ ∘⇩s σ ∘⇩s α)"
using rm_vars_img_subset unfolding θ_def ξ_elim by blast
hence 3: "subst_range θ ∩ pubval_terms = {}"
"subst_range θ ∩ abs_terms = {}"
using 2(3,4) step.prems funs_term_Fun_subterm
unfolding T'_def θ_def ξ_elim by (blast,blast)
have "(ℐ ` fv (s ⋅ θ)) ∩ pubval_terms = {}"
"(ℐ ` fv (s ⋅ θ)) ∩ abs_terms = {}"
proof -
have "θ = ξ ∘⇩s σ ∘⇩s α" "bvars_transaction T = {}" "vars⇩l⇩s⇩s⇩t T' = fv⇩l⇩s⇩s⇩t T'"
using s(3) T_no_bvars step.hyps(2) rm_vars_empty
vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel T'"]
bvars⇩s⇩s⇩t_subst[of "unlabel (transaction_strand T)" "ξ ∘⇩s σ ∘⇩s α"]
unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"]
unfolding θ_def T'_def by simp_all
hence "fv (s ⋅ θ) ⊆ fv⇩l⇩s⇩s⇩t T'"
using trms⇩s⇩s⇩t_fv_subst_subset[OF s(1), of θ] unlabel_subst[of "transaction_strand T" θ]
unfolding T'_def by auto
moreover have "fv⇩l⇩s⇩s⇩t T' ⊆ fv⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t T')"
using fv⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel (dual⇩l⇩s⇩s⇩t T')"]
unlabel_append[of 𝒜 "dual⇩l⇩s⇩s⇩t T'"]
fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of T']
by simp_all
hence "ℐ ` fv⇩l⇩s⇩s⇩t T' ∩ pubval_terms = {}" "ℐ ` fv⇩l⇩s⇩s⇩t T' ∩ abs_terms = {}"
using 2(1,2) by blast+
ultimately show "(ℐ ` fv (s ⋅ θ)) ∩ pubval_terms = {}" "(ℐ ` fv (s ⋅ θ)) ∩ abs_terms = {}"
by blast+
qed
hence σαℐ_disj: "((θ ∘⇩s ℐ) ` fv s) ∩ pubval_terms = {}"
"((θ ∘⇩s ℐ) ` fv s) ∩ abs_terms = {}"
using pubval_terms_subst_range_comp'[of θ "fv s" ℐ]
abs_terms_subst_range_comp'[of θ "fv s" ℐ]
pubval_terms_subst_range_disj[OF 3(1), of s]
abs_terms_subst_range_disj[OF 3(2), of s]
by (simp_all add: subst_apply_fv_unfold)
have 4: "t ∉ pubval_terms" "t ∉ abs_terms"
using s(2) s_nin σαℐ_disj
pubval_terms_subst[of s "rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α) ∘⇩s ℐ"]
pubval_terms_subst_range_disj[of "rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α) ∘⇩s ℐ" s]
abs_terms_subst[of s "rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α) ∘⇩s ℐ"]
abs_terms_subst_range_disj[of "rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α) ∘⇩s ℐ" s]
unfolding θ_def
by blast+
have "t ∈ SMP trms_P" "fv t = {}"
by (metis s(2) SMP.Substitution[OF SMP.MP[OF 1] wt wftrms, of X],
metis s(2) subst_subst_compose[of s "rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α)" ℐ]
interpretation_grounds[OF ℐ(1), of "s ⋅ rm_vars (set X) (ξ ∘⇩s σ ∘⇩s α)"])
hence 5: "t ∈ GSMP trms_P" unfolding GSMP_def by simp
show "t ∈ ?B" using 4 5 by (auto simp add: trms_P_def)
qed
thus ?case
using step.IH[OF ℐ'] trms⇩s⇩s⇩t_append[of "unlabel 𝒜"] unlabel_append[of 𝒜]
image_Un[of "λx. x ⋅ ℐ" "trms⇩l⇩s⇩s⇩t 𝒜"]
by (simp add: T'_def)
qed simp
lemma α⇩t⇩i_covers_α⇩0_aux:
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and P: "∀T ∈ set P. admissible_transaction' T"
and P_occ: "∀T ∈ set P. admissible_transaction_occurs_checks T"
and t: "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
"t = Fun (Val n) [] ∨ t = Var x"
and neq:
"t ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) ≠
t ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ)"
shows "∃y ∈ fv_transaction T - set (transaction_fresh T).
t ⋅ ℐ = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ∧ Γ⇩v y = TAtom Value"
proof -
let ?𝒜' = "𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
let ?ℬ = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T))"
let ?ℬ' = "?ℬ ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"
let ?ℬ'' = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
have ℐ_interp: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ"
and ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t ℐ"
and ℐ_wf: "wf⇩t⇩r⇩m⇩s (subst_range ℐ)"
by (metis ℐ welltyped_constraint_model_def constraint_model_def,
metis ℐ welltyped_constraint_model_def,
metis ℐ welltyped_constraint_model_def constraint_model_def)
note T_adm = bspec[OF P(1) T]
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
note T_adm_upds = admissible_transaction_is_wellformed_transaction(3)[OF T_adm]
have T_fresh_vars_value_typed: "∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value"
using T P(1) protocol_transaction_vars_TAtom_typed(3)[of T] P(1) by simp
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm ξ]
note ξσα_wt = transaction_decl_fresh_renaming_substs_wt[OF ξ σ α]
have 𝒜_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)"
by (metis reachable_constraints_wf⇩t⇩r⇩m⇩s admissible_transactions_wf⇩t⇩r⇩m⇩s P(1) 𝒜_reach)
hence t_wf: "wf⇩t⇩r⇩m t" using t by auto
have 𝒜_no_val_bvars: "¬TAtom Value ⊑ Γ⇩v x"
when "x ∈ bvars⇩l⇩s⇩s⇩t 𝒜" for x
using P(1) reachable_constraints_no_bvars 𝒜_reach
vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel 𝒜"] that
admissible_transactionE(4)
by fast
have x': "x ∈ vars⇩l⇩s⇩s⇩t 𝒜" when "t = Var x"
using that t by (simp add: var_subterm_trms⇩s⇩s⇩t_is_vars⇩s⇩s⇩t)
have "∃f ∈ funs_term (t ⋅ ℐ). is_Val f"
using abs_eq_if_no_Val neq by metis
hence "∃n T. Fun (Val n) T ⊑ t ⋅ ℐ"
using funs_term_Fun_subterm
unfolding is_Val_def by fast
hence "TAtom Value ⊑ Γ (Var x)" when "t = Var x"
using wt_subst_trm''[OF ℐ_wt, of "Var x"] that
subtermeq_imp_subtermtypeeq[of "t ⋅ ℐ"] wf_trm_subst[OF ℐ_wf, of t] t_wf
by fastforce
hence x_val: "Γ⇩v x = TAtom Value" when "t = Var x"
using reachable_constraints_vars_TAtom_typed[OF 𝒜_reach P x'] that
by fastforce
hence x_fv: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" when "t = Var x" using x'
using reachable_constraints_Value_vars_are_fv[OF 𝒜_reach P x'] that
by blast
then obtain m where m: "t ⋅ ℐ = Fun (Val m) []"
using constraint_model_Value_term_is_Val[
OF 𝒜_reach welltyped_constraint_model_prefix[OF ℐ] P P_occ, of x]
t(2) x_val
by force
hence 0: "α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) m ≠ α⇩0 (db⇩s⇩s⇩t (unlabel 𝒜@?ℬ'') ℐ) m"
using neq by (simp add: unlabel_def)
have t_val: "Γ t = TAtom Value" using x_val t by force
obtain u s where s: "t ⋅ ℐ = u ⋅ ℐ" "insert⟨u,s⟩ ∈ set ?ℬ' ∨ delete⟨u,s⟩ ∈ set ?ℬ'"
using to_abs_neq_imp_db_update[OF 0] m
by (metis (no_types, lifting) dual⇩l⇩s⇩s⇩t_subst subst_lsst_unlabel)
then obtain u' s' where s':
"u = u' ⋅ ξ ∘⇩s σ ∘⇩s α" "s = s' ⋅ ξ ∘⇩s σ ∘⇩s α"
"insert⟨u',s'⟩ ∈ set ?ℬ ∨ delete⟨u',s'⟩ ∈ set ?ℬ"
using stateful_strand_step_mem_substD(4,5)
by blast
hence s'': "insert⟨u',s'⟩ ∈ set (unlabel (transaction_strand T)) ∨
delete⟨u',s'⟩ ∈ set (unlabel (transaction_strand T))"
using dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(4,5)[of u' s' "transaction_strand T"]
by simp_all
then obtain y where y: "y ∈ fv_transaction T" "u' = Var y"
using transaction_inserts_are_Value_vars[OF T_wf T_adm_upds, of u' s']
transaction_deletes_are_Value_vars[OF T_wf T_adm_upds, of u' s']
stateful_strand_step_fv_subset_cases(4,5)[of u' s' "unlabel (transaction_strand T)"]
by auto
hence 1: "t ⋅ ℐ = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ" using y s(1) s'(1) by (metis eval_term.simps(1))
have 2: "y ∉ set (transaction_fresh T)" when "(ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ≠ σ y"
using transaction_fresh_subst_grounds_domain[OF σ, of y] subst_compose[of σ α y] that ξ_empty
by (auto simp add: subst_ground_ident)
have 3: "y ∉ set (transaction_fresh T)" when "(ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
using 2 that σ unfolding transaction_fresh_subst_def by fastforce
have 4: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜. Γ⇩v x = TAtom Value ⟶
(∃B. prefix B 𝒜 ∧ x ∉ fv⇩l⇩s⇩s⇩t B ∧ ℐ x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B))"
by (metis welltyped_constraint_model_prefix[OF ℐ]
constraint_model_Value_var_in_constr_prefix[OF 𝒜_reach _ P P_occ])
have 5: "Γ⇩v y = TAtom Value"
using 1 t_val
wt_subst_trm''[OF ξσα_wt, of "Var y"]
wt_subst_trm''[OF ℐ_wt, of t]
wt_subst_trm''[OF ℐ_wt, of "(ξ ∘⇩s σ ∘⇩s α) y"]
by (auto simp del: subst_subst_compose)
have "y ∉ set (transaction_fresh T)"
proof (cases "t = Var x")
case True
hence *: "ℐ x = Fun (Val m) []" "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" "ℐ x = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ"
using m t(1) 1 x_fv x' by (force, blast, force)
obtain B where B: "prefix B 𝒜" "ℐ x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)"
using *(2) 4 x_val[OF True] by fastforce
hence "∀t ∈ subst_range σ. t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)"
using transaction_fresh_subst_range_fresh(1)[OF σ] trms⇩s⇩s⇩t_unlabel_prefix_subset(1)[of B]
unfolding prefix_def by fast
thus ?thesis using *(1,3) B(2) 2 by (metis subst_imgI term.distinct(1))
next
case False
hence "t ⋅ ℐ ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" using t by simp
thus ?thesis using 1 3 by argo
qed
thus ?thesis using 1 5 y(1) by fast
qed
lemma α⇩t⇩i_covers_α⇩0_Var:
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and P: "∀T ∈ set P. admissible_transaction' T"
and P_occ: "∀T ∈ set P. admissible_transaction_occurs_checks T"
and x: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜"
shows "ℐ x ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ) ∈
timpl_closure_set {ℐ x ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)} (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)"
proof -
define a0 where "a0 ≡ α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
define a0' where "a0' ≡ α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ)"
define a3 where "a3 ≡ α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"
have 𝒜_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)"
by (metis reachable_constraints_wf⇩t⇩r⇩m⇩s admissible_transactions_wf⇩t⇩r⇩m⇩s P(1) 𝒜_reach)
note T_adm = bspec[OF P(1) T]
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm ξ]
note ξσα_wt = transaction_decl_fresh_renaming_substs_wt[OF ξ σ α]
have ℐ_interp: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ"
and ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t ℐ"
and ℐ_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (subst_range ℐ)"
by (metis ℐ welltyped_constraint_model_def constraint_model_def,
metis ℐ welltyped_constraint_model_def,
metis ℐ welltyped_constraint_model_def constraint_model_def)
have "Γ⇩v x = Var Value ∨ (∃a. Γ⇩v x = Var (prot_atom.Atom a))"
using reachable_constraints_vars_TAtom_typed[OF 𝒜_reach P, of x]
x vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel 𝒜"]
by auto
hence "ℐ x ⋅⇩α a0' ∈ timpl_closure_set {ℐ x ⋅⇩α a0} a3"
proof
assume x_val: "Γ⇩v x = TAtom Value"
show "ℐ x ⋅⇩α a0' ∈ timpl_closure_set {ℐ x ⋅⇩α a0} a3"
proof (cases "ℐ x ⋅⇩α a0 = ℐ x ⋅⇩α a0'")
case False
hence "∃y ∈ fv_transaction T - set (transaction_fresh T).
ℐ x = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ∧ Γ⇩v y = TAtom Value"
using α⇩t⇩i_covers_α⇩0_aux[OF 𝒜_reach T ℐ ξ σ α P P_occ fv⇩s⇩s⇩t_is_subterm_trms⇩s⇩s⇩t[OF x], of _ x]
unfolding a0_def a0'_def
by fastforce
then obtain y where y:
"y ∈ fv_transaction T - set (transaction_fresh T)"
"ℐ x = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ"
"ℐ x ⋅⇩α a0 = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ⋅⇩α a0"
"ℐ x ⋅⇩α a0' = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ⋅⇩α a0'"
"Γ⇩v y = TAtom Value"
by metis
then obtain n where n: "(ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ = Fun (Val n) []"
using Γ⇩v_TAtom''(2)[of y] x x_val
transaction_var_becomes_Val[
OF reachable_constraints.step[OF 𝒜_reach T ξ σ α] ℐ ξ σ α P P_occ T, of y]
by force
have "a0 n ≠ a0' n"
"y ∈ fv_transaction T"
"y ∉ set (transaction_fresh T)"
"absc (a0 n) = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ⋅⇩α a0"
"absc (a0' n) = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ⋅⇩α a0'"
using y n False by force+
hence 1: "(a0 n, a0' n) ∈ a3"
unfolding a0_def a0'_def a3_def abs_term_implications_def
by blast
have 2: "ℐ x ⋅⇩α a0' ∈ set ⟨a0 n --» a0' n⟩⟨ℐ x ⋅⇩α a0⟩"
using y n timpl_apply_const by auto
show ?thesis
using timpl_closure.TI[OF timpl_closure.FP 1] 2
term_variants_pred_iff_in_term_variants[
of "(λ_. [])(Abs (a0 n) := [Abs (a0' n)])"]
unfolding timpl_closure_set_def timpl_apply_term_def
by auto
qed (auto intro: timpl_closure_setI)
next
assume "∃a. Γ⇩v x = TAtom (Atom a)"
then obtain a where x_atom: "Γ⇩v x = TAtom (Atom a)" by force
obtain f T where fT: "ℐ x = Fun f T"
using interpretation_grounds[OF ℐ_interp, of "Var x"]
by (cases "ℐ x") auto
have fT_atom: "Γ (Fun f T) = TAtom (Atom a)"
using wt_subst_trm''[OF ℐ_wt, of "Var x"] x_atom fT
by simp
have T: "T = []"
using fT wf_trm_subst[OF ℐ_wf⇩t⇩r⇩m⇩s, of "Var x"] const_type_inv_wf[OF fT_atom]
by fastforce
have f: "¬is_Val f" using fT_atom unfolding is_Val_def by auto
have "ℐ x ⋅⇩α b = ℐ x" for b
using T fT abs_term_apply_const(2)[OF f]
by auto
thus "ℐ x ⋅⇩α a0' ∈ timpl_closure_set {ℐ x ⋅⇩α a0} a3"
by (auto intro: timpl_closure_setI)
qed
thus ?thesis by (metis a0_def a0'_def a3_def)
qed
lemma α⇩t⇩i_covers_α⇩0_Val:
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and P: "∀T ∈ set P. admissible_transaction' T"
and P_occ: "∀T ∈ set P. admissible_transaction_occurs_checks T"
and n: "Fun (Val n) [] ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
shows "Fun (Val n) [] ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ) ∈
timpl_closure_set {Fun (Val n) [] ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)} (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)"
proof -
define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define a0 where "a0 ≡ α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
define a0' where "a0' ≡ α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@T') ℐ)"
define a3 where "a3 ≡ α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"
have 𝒜_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)"
by (metis reachable_constraints_wf⇩t⇩r⇩m⇩s admissible_transactions_wf⇩t⇩r⇩m⇩s P(1) 𝒜_reach)
note T_adm = bspec[OF P(1) T]
have "Fun (Abs (a0' n)) [] ∈ timpl_closure_set {Fun (Abs (a0 n)) []} a3"
proof (cases "a0 n = a0' n")
case False
then obtain x where x:
"x ∈ fv_transaction T - set (transaction_fresh T)" "Fun (Val n) [] = (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ"
using α⇩t⇩i_covers_α⇩0_aux[OF 𝒜_reach T ℐ ξ σ α P P_occ n]
by (fastforce simp add: a0_def a0'_def T'_def)
hence "absc (a0 n) = (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0" "absc (a0' n) = (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0'"
by simp_all
hence 1: "(a0 n, a0' n) ∈ a3"
using False x(1)
unfolding a0_def a0'_def a3_def abs_term_implications_def T'_def
by blast
show ?thesis
using timpl_apply_Abs[of "[]" "[]" "a0 n" "a0' n"]
timpl_closure.TI[OF timpl_closure.FP[of "Fun (Abs (a0 n)) []" a3] 1]
term_variants_pred_iff_in_term_variants[of "(λ_. [])(Abs (a0 n) := [Abs (a0' n)])"]
unfolding timpl_closure_set_def timpl_apply_term_def
by force
qed (auto intro: timpl_closure_setI)
thus ?thesis by (simp add: a0_def a0'_def a3_def T'_def)
qed
lemma α⇩t⇩i_covers_α⇩0_ik:
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and P: "∀T ∈ set P. admissible_transaction' T"
and P_occ: "∀T ∈ set P. admissible_transaction_occurs_checks T"
and t: "t ∈ ik⇩l⇩s⇩s⇩t 𝒜"
shows "t ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ) ∈
timpl_closure_set {t ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)} (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)"
proof -
define a0 where "a0 ≡ α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
define a0' where "a0' ≡ α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ)"
define a3 where "a3 ≡ α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"
let ?U = "λT a. map (λs. s ⋅ ℐ ⋅⇩α a) T"
have 𝒜_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)"
by (metis reachable_constraints_wf⇩t⇩r⇩m⇩s admissible_transactions_wf⇩t⇩r⇩m⇩s P(1) 𝒜_reach)
note T_adm = bspec[OF P(1) T]
have "t ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜)" "wf⇩t⇩r⇩m t" using 𝒜_wf⇩t⇩r⇩m⇩s t ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset by force+
hence "∀t0 ∈ subterms t. t0 ⋅ ℐ ⋅⇩α a0' ∈ timpl_closure_set {t0 ⋅ ℐ ⋅⇩α a0} a3"
proof (induction t)
case (Var x) thus ?case
using α⇩t⇩i_covers_α⇩0_Var[OF 𝒜_reach T ℐ ξ σ α P P_occ, of x]
ik⇩s⇩s⇩t_var_is_fv[of x "unlabel 𝒜"] vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel 𝒜"]
by (simp add: a0_def a0'_def a3_def)
next
case (Fun f S)
have IH: "∀t0 ∈ subterms t. t0 ⋅ ℐ ⋅⇩α a0' ∈ timpl_closure_set {t0 ⋅ ℐ ⋅⇩α a0} a3"
when "t ∈ set S" for t
using that Fun.prems(1) wf_trm_param[OF Fun.prems(2)] Fun.IH
by (meson in_subterms_subset_Union params_subterms subsetCE)
hence "t ⋅⇩α a0' ∈ timpl_closure_set {t ⋅⇩α a0} a3"
when "t ∈ set (map (λs. s ⋅ ℐ) S)" for t
using that by auto
hence "t ⋅⇩α a0' ∈ timpl_closure (t ⋅⇩α a0) a3"
when "t ∈ set (map (λs. s ⋅ ℐ) S)" for t
using that timpl_closureton_is_timpl_closure by auto
hence "(t ⋅⇩α a0, t ⋅⇩α a0') ∈ timpl_closure' a3"
when "t ∈ set (map (λs. s ⋅ ℐ) S)" for t
using that timpl_closure_is_timpl_closure' by auto
hence IH': "((?U S a0) ! i, (?U S a0') ! i) ∈ timpl_closure' a3"
when "i < length (map (λs. s ⋅ ℐ ⋅⇩α a0) S)" for i
using that by auto
show ?case
proof (cases "∃n. f = Val n")
case True
then obtain n where "Fun f S = Fun (Val n) []"
using Fun.prems(2) unfolding wf⇩t⇩r⇩m_def by force
moreover have "Fun f S ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
using ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset Fun.prems(1) by blast
ultimately show ?thesis
using α⇩t⇩i_covers_α⇩0_Val[OF 𝒜_reach T ℐ ξ σ α P P_occ]
by (simp add: a0_def a0'_def a3_def)
next
case False
hence "Fun f S ⋅ ℐ ⋅⇩α a = Fun f (map (λt. t ⋅ ℐ ⋅⇩α a) S)" for a by (cases f) simp_all
hence "(Fun f S ⋅ ℐ ⋅⇩α a0, Fun f S ⋅ ℐ ⋅⇩α a0') ∈ timpl_closure' a3"
using timpl_closure_FunI[OF IH']
by simp
hence "Fun f S ⋅ ℐ ⋅⇩α a0' ∈ timpl_closure_set {Fun f S ⋅ ℐ ⋅⇩α a0} a3"
using timpl_closureton_is_timpl_closure
timpl_closure_is_timpl_closure'
by metis
thus ?thesis using IH by simp
qed
qed
thus ?thesis by (simp add: a0_def a0'_def a3_def)
qed
lemma transaction_prop1:
assumes "δ ∈ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)"
and "x ∈ fv_transaction T"
and "x ∉ set (transaction_fresh T)"
and "δ x ≠ absdbupd (unlabel (transaction_updates T)) x (δ x)"
and "transaction_check' msgcs (FP, OCC, TI) T"
and TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
shows "(δ x, absdbupd (unlabel (transaction_updates T)) x (δ x)) ∈ (set TI)⇧+"
proof -
let ?upd = "λx. absdbupd (unlabel (transaction_updates T)) x (δ x)"
have 0: "fv_transaction T = set (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)))"
by (metis fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"])
have 1: "transaction_check_post (FP, OCC, TI) T δ"
using assms(1,5)
unfolding transaction_check_def transaction_check'_def list_all_iff
by blast
have "(δ x, ?upd x) ∈ set TI ⟷ (δ x, ?upd x) ∈ (set TI)⇧+"
using TI using assms(4) by blast
thus ?thesis
using assms(2,3,4) 0 1 in_trancl_closure_iff_in_trancl_fun[of _ _ TI]
unfolding transaction_check_post_def List.member_def Let_def by blast
qed
lemma transaction_prop2:
assumes δ: "δ ∈ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)"
and x: "x ∈ fv_transaction T" "fst x = TAtom Value"
and T_check: "transaction_check' msgcs (FP, OCC, TI) T"
and T_adm: "admissible_transaction' T"
and T_occ: "admissible_transaction_occurs_checks T"
and FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
shows "x ∉ set (transaction_fresh T) ⟹ δ x ∈ set OCC" (is "?A' ⟹ ?A")
and "absdbupd (unlabel (transaction_updates T)) x (δ x) ∈ set OCC" (is ?B)
proof -
let ?xs = "fv_list⇩s⇩s⇩t (unlabel (transaction_strand T))"
let ?ys = "filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) ?xs"
let ?C = "unlabel (transaction_checks T)"
let ?poss = "transaction_poschecks_comp ?C"
let ?negs = "transaction_negchecks_comp ?C"
let ?δupd = "λy. absdbupd (unlabel (transaction_updates T)) y (δ y)"
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
have 0: "{x ∈ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value} = set ?ys"
using fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"]
by force
have 1: "transaction_check_pre (FP, OCC, TI) T δ"
using δ unfolding transaction_check_comp_def Let_def by fastforce
have 2: "transaction_check_post (FP, OCC, TI) T δ"
using δ T_check unfolding transaction_check_def transaction_check'_def list_all_iff by auto
have 3: "δ ∈ abs_substs_fun ` set (abs_substs_set ?ys OCC ?poss ?negs msgcs)"
using δ unfolding transaction_check_comp_def Let_def by force
show A: ?A when ?A' using that 0 3 x abs_substs_abss_bounded by blast
have 4: "x ∈ fv⇩l⇩s⇩s⇩t (transaction_updates T) ∪ fv⇩l⇩s⇩s⇩t (transaction_send T)"
when x': "x ∈ set (transaction_fresh T)"
using x' admissible_transactionE(7)[OF T_adm] by blast
have "intruder_synth_mod_timpls FP TI (occurs (absc (?δupd x)))"
when x': "x ∈ set (transaction_fresh T)"
proof -
obtain l ts S where ts:
"transaction_send T = (l,send⟨ts⟩)#S" "occurs (Var x) ∈ set ts"
using admissible_transaction_occurs_checksE2[OF T_occ x'] by blast
hence "occurs (Var x) ∈ set ts" "send⟨ts⟩ ∈ set (unlabel (transaction_send T))"
using x' unfolding suffix_def by (fastforce, fastforce)
thus ?thesis using 2 x unfolding transaction_check_post_def by fastforce
qed
hence "timpl_closure_set (set FP) (set TI) ⊢⇩c occurs (absc (?δupd x))"
when x': "x ∈ set (transaction_fresh T)"
using x' intruder_synth_mod_timpls_is_synth_timpl_closure_set[
OF TI, of FP "occurs (absc (?δupd x))"]
by argo
hence "Abs (?δupd x) ∈ ⋃(funs_term ` timpl_closure_set (set FP) (set TI))"
when x': "x ∈ set (transaction_fresh T)"
using x' ideduct_synth_priv_fun_in_ik[
of "timpl_closure_set (set FP) (set TI)" "occurs (absc (?δupd x))"]
by simp
hence "∃t ∈ timpl_closure_set (set FP) (set TI). Abs (?δupd x) ∈ funs_term t"
when x': "x ∈ set (transaction_fresh T)"
using x' by force
hence 5: "?δupd x ∈ set OCC" when x': "x ∈ set (transaction_fresh T)"
using x' OCC by fastforce
have 6: "?δupd x ∈ set OCC" when x': "x ∉ set (transaction_fresh T)"
proof (cases "δ x = ?δupd x")
case False
hence "(δ x, ?δupd x) ∈ (set TI)⇧+" "δ x ∈ set OCC"
using A 2 x' x TI
unfolding transaction_check_post_def fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t Let_def
in_trancl_closure_iff_in_trancl_fun[symmetric]
List.member_def
by blast+
thus ?thesis using timpl_closure_set_absc_subset_in[OF OCC(2)] by blast
qed (simp add: A x' x(1))
show ?B by (metis 5 6)
qed
lemma transaction_prop3:
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
"∀t ∈ α⇩i⇩k 𝒜 ℐ. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
"α⇩v⇩a⇩l⇩s 𝒜 ℐ ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and P:
"∀T ∈ set P. admissible_transaction' T"
and P_occ: "∀T ∈ set P. admissible_transaction_occurs_checks T"
shows "∀x ∈ set (transaction_fresh T). (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc {}" (is ?A)
and "∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T).
intruder_synth_mod_timpls FP TI (t ⋅ (ξ ∘⇩s σ ∘⇩s α) ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ))" (is ?B)
and "∀x ∈ fv_transaction T - set (transaction_fresh T).
∀s. select⟨Var x,Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))
⟶ (∃ss. (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc ss ∧ s ∈ ss)" (is ?C)
and "∀x ∈ fv_transaction T - set (transaction_fresh T).
∀s. ⟨Var x in Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))
⟶ (∃ss. (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc ss ∧ s ∈ ss)" (is ?D)
and "∀x ∈ fv_transaction T - set (transaction_fresh T).
∀s. ⟨Var x not in Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))
⟶ (∃ss. (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc ss ∧ s ∉ ss)" (is ?E)
and "∀x ∈ fv_transaction T - set (transaction_fresh T). Γ⇩v x = TAtom Value ⟶
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) ∈ absc ` set OCC" (is ?F)
proof -
let ?T' = "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
let ?θ = "ξ ∘⇩s σ ∘⇩s α"
define a0 where "a0 ≡ α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
define a0' where "a0' ≡ α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@?T') ℐ)"
define fv_AT' where "fv_AT' ≡ fv⇩l⇩s⇩s⇩t (𝒜@?T')"
note T_adm = bspec[OF P(1) T]
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
note T_adm' = admissible_transaction_is_wellformed_transaction(2-4)[OF T_adm]
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm ξ]
hence ξ_elim: "?θ = σ ∘⇩s α" by simp
have ℐ': "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "wt⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ)"
"∀n. PubConst Value n ∉ ⋃(funs_term ` (ℐ ` fv⇩l⇩s⇩s⇩t 𝒜))"
"∀n. Abs n ∉ ⋃(funs_term ` (ℐ ` fv⇩l⇩s⇩s⇩t 𝒜))"
"∀n. PubConst Value n ∉ ⋃(funs_term ` (ℐ ` fv_AT'))"
"∀n. Abs n ∉ ⋃(funs_term ` (ℐ ` fv_AT'))"
using ℐ admissible_transaction_occurs_checks_prop'[
OF 𝒜_reach welltyped_constraint_model_prefix[OF ℐ] P P_occ]
admissible_transaction_occurs_checks_prop'[
OF reachable_constraints.step[OF 𝒜_reach T ξ σ α] ℐ P P_occ]
unfolding welltyped_constraint_model_def constraint_model_def is_Val_def is_Abs_def fv_AT'_def
by (meson,meson,meson,metis,metis,metis,metis)
have T_no_pubconsts: "∀n. PubConst Value n ∉ ⋃(funs_term ` trms_transaction T)"
and T_no_abss: "∀n. Abs n ∉ ⋃(funs_term ` trms_transaction T)"
and T_fresh_vars_value_typed: "∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value"
and T_fv_const_typed: "∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))"
using protocol_transaction_vars_TAtom_typed
protocol_transactions_no_pubconsts
protocol_transactions_no_abss
funs_term_Fun_subterm P T
by (fast, fast, fast, fast)
have wt_σαℐ: "wt⇩s⇩u⇩b⇩s⇩t (σ ∘⇩s α ∘⇩s ℐ)"
using ℐ'(2) wt_subst_compose transaction_fresh_subst_wt[OF σ]
transaction_renaming_subst_wt[OF α]
by blast
have 1: "?θ y ⋅ ℐ = σ y" when "y ∈ set (transaction_fresh T)" for y
using transaction_fresh_subst_grounds_domain[OF σ that] subst_compose[of σ α y]
unfolding ξ_elim by (simp add: subst_ground_ident)
have 2: "?θ y ⋅ ℐ ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" when "y ∈ set (transaction_fresh T)" for y
using 1[OF that] that σ unfolding transaction_fresh_subst_def by auto
have 3: "∀x ∈ fv⇩l⇩s⇩s⇩t 𝒜. Γ⇩v x = TAtom Value ⟶
(∃B. prefix B 𝒜 ∧ x ∉ fv⇩l⇩s⇩s⇩t B ∧ ℐ x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B))"
by (metis welltyped_constraint_model_prefix[OF ℐ]
constraint_model_Value_var_in_constr_prefix[OF 𝒜_reach _ P P_occ])
have 4: "∃n. ?θ y ⋅ ℐ = Fun (Val n) []"
when "y ∈ fv_transaction T" "Γ⇩v y = TAtom Value" for y
using transaction_var_becomes_Val[
OF reachable_constraints.step[OF 𝒜_reach T ξ σ α] ℐ ξ σ α P P_occ T]
that T_fv_const_typed Γ⇩v_TAtom''[of y]
by metis
have ℐ_is_T_model: "strand_sem_stateful (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) (set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)) (unlabel ?T') ℐ"
using ℐ unlabel_append[of 𝒜 ?T'] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel 𝒜" ℐ "[]"]
strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" "unlabel ?T'" ℐ]
by (simp add: welltyped_constraint_model_def constraint_model_def db⇩s⇩s⇩t_def)
have T_rcv_no_val_bvars: "bvars⇩l⇩s⇩s⇩t (transaction_receive T) ∩ subst_domain ?θ = {}"
using admissible_transaction_no_bvars[OF T_adm] bvars_transaction_unfold[of T] by blast
show ?A
proof
fix y assume y: "y ∈ set (transaction_fresh T)"
then obtain yn where yn: "(ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ = Fun (Val yn) []" "Fun (Val yn) [] ∈ subst_range σ"
by (metis ξ_elim T_fresh_vars_value_typed transaction_fresh_subst_sends_to_val'[OF σ])
{
fix t' s assume t': "insert⟨t',s⟩ ∈ set (unlabel 𝒜)" "t' ⋅ ℐ = Fun (Val yn) []"
then obtain z where t'_z: "t' = Var z" using 2[OF y] yn(1) by (cases t') auto
hence z: "z ∈ fv⇩l⇩s⇩s⇩t 𝒜" "ℐ z = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ" using t' yn(1) by force+
hence z': "Γ⇩v z = TAtom Value"
by (metis Γ.simps(1) Γ_consts_simps(2) t'(2) t'_z wt_subst_trm'' ℐ'(2))
obtain B where B: "prefix B 𝒜" "ℐ z ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)" using z z' 3 by fastforce
hence "∀t ∈ subst_range σ. t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)"
using transaction_fresh_subst_range_fresh(1)[OF σ] trms⇩s⇩s⇩t_unlabel_prefix_subset(1)[of B]
unfolding prefix_def by fast
hence False using B(2) 1[OF y] z yn(1) by (metis subst_imgI term.distinct(1))
} hence "∄s. (?θ y ⋅ ℐ, s) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using db⇩s⇩s⇩t_in_cases[of "?θ y ⋅ ℐ" _ "unlabel 𝒜" ℐ "[]"] yn(1)
by (force simp add: db⇩s⇩s⇩t_def)
thus "?θ y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc {}"
using to_abs_empty_iff_notin_db[of yn "db'⇩l⇩s⇩s⇩t 𝒜 ℐ []"] yn(1)
by (simp add: db⇩s⇩s⇩t_def)
qed
show receives_covered: ?B
proof
fix t assume t: "t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T)"
hence t_in_T: "t ∈ trms_transaction T"
using trms⇩s⇩s⇩t_unlabel_prefix_subset(1)[of "transaction_receive T"]
unfolding transaction_strand_def by fast
obtain ts where ts: "t ∈ set ts" "receive⟨ts⟩ ∈ set (unlabel (transaction_receive T))"
using t wellformed_transaction_send_receive_trm_cases(1)[OF T_wf] by blast
have t_rcv: "receive⟨ts ⋅⇩l⇩i⇩s⇩t ?θ⟩ ∈ set (unlabel (transaction_receive T ⋅⇩l⇩s⇩s⇩t ?θ))"
using subst_lsst_unlabel_member[of "receive⟨ts⟩" "transaction_receive T" ?θ]
trms⇩s⇩s⇩t_in[OF t] ts
by fastforce
have "list_all (λt. ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ ?θ ⋅ ℐ) ts"
using wellformed_transaction_sem_receives[OF T_wf ℐ_is_T_model t_rcv]
unfolding list_all_iff by fastforce
hence *: "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ ?θ ⋅ ℐ" using ts(1) unfolding list_all_iff by fast
have t_fv: "fv (t ⋅ ?θ) ⊆ fv_AT'"
using fv⇩s⇩s⇩t_append[of "unlabel 𝒜"] unlabel_append[of 𝒜]
fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ?θ"]
ts(1) t_rcv fv_transaction_subst_unfold[of T ?θ]
unfolding fv_AT'_def by force
have **: "∀t ∈ (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
using FP(3) by (auto simp add: a0_def abs_intruder_knowledge_def)
note lms1 = pubval_terms_subst[OF _ pubval_terms_subst_range_disj[
OF transaction_fresh_subst_has_no_pubconsts_abss(1)[OF σ], of t]]
pubval_terms_subst[OF _ pubval_terms_subst_range_disj[
OF transaction_renaming_subst_has_no_pubconsts_abss(1)[OF α], of "t ⋅ σ"]]
note lms2 = abs_terms_subst[OF _ abs_terms_subst_range_disj[
OF transaction_fresh_subst_has_no_pubconsts_abss(2)[OF σ], of t]]
abs_terms_subst[OF _ abs_terms_subst_range_disj[
OF transaction_renaming_subst_has_no_pubconsts_abss(2)[OF α], of "t ⋅ σ"]]
have "t ∈ (⋃T∈set P. trms_transaction T)" "fv (t ⋅ σ ∘⇩s α ⋅ ℐ) = {}"
using t_in_T T interpretation_grounds[OF ℐ'(1)] by fast+
moreover have "wf⇩t⇩r⇩m⇩s (subst_range (σ ∘⇩s α ∘⇩s ℐ))"
using wf_trm_subst_rangeI[of σ, OF transaction_fresh_subst_is_wf_trm[OF σ]]
wf_trm_subst_rangeI[of α, OF transaction_renaming_subst_is_wf_trm[OF α]]
wf_trms_subst_compose[of σ α, THEN wf_trms_subst_compose[OF _ ℐ'(3)]]
by blast
moreover
have "t ∉ pubval_terms"
using t_in_T T_no_pubconsts funs_term_Fun_subterm
unfolding is_PubConstValue_def is_PubConst_def by fastforce
hence "t ⋅ ?θ ∉ pubval_terms"
using lms1 T_fresh_vars_value_typed
unfolding ξ_elim by auto
hence "t ⋅ ?θ ⋅ ℐ ∉ pubval_terms"
using ℐ'(6) t_fv pubval_terms_subst'[of "t ⋅ ?θ" ℐ]
by auto
moreover have "t ∉ abs_terms"
using t_in_T T_no_abss funs_term_Fun_subterm
unfolding is_Abs_def by force
hence "t ⋅ ?θ ∉ abs_terms"
using lms2 T_fresh_vars_value_typed
unfolding ξ_elim by auto
hence "t ⋅ ?θ ⋅ ℐ ∉ abs_terms"
using ℐ'(7) t_fv abs_terms_subst'[of "t ⋅ ?θ" ℐ]
by auto
ultimately have ***:
"t ⋅ ξ ∘⇩s σ ∘⇩s α ⋅ ℐ ∈ GSMP (⋃T∈set P. trms_transaction T) - (pubval_terms ∪ abs_terms)"
using SMP.Substitution[OF SMP.MP[of t "⋃T∈set P. trms_transaction T"], of "σ ∘⇩s α ∘⇩s ℐ"]
subst_subst_compose[of t ?θ ℐ] wt_σαℐ
unfolding GSMP_def ξ_elim by fastforce
have ****:
"ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊆ GSMP (⋃T∈set P. trms_transaction T) - (pubval_terms ∪ abs_terms)"
using reachable_constraints_GSMP_no_pubvals_abss[OF 𝒜_reach P ℐ'(1-5)]
ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset[of "unlabel 𝒜"]
by blast
show "intruder_synth_mod_timpls FP TI (t ⋅ ?θ ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ))"
using deduct_FP_if_deduct[OF **** ** * ***] deducts_eq_if_analyzed[OF FP(1)]
intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, of FP]
unfolding a0_def by force
qed
show ?C
proof (intro ballI allI impI)
fix y s
assume y: "y ∈ fv_transaction T - set (transaction_fresh T)"
and s: "select⟨Var y, Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))"
hence "select⟨Var y, Fun (Set s) []⟩ ∈ set (unlabel (transaction_strand T))"
unfolding transaction_strand_def unlabel_def by auto
hence y_val: "Γ⇩v y = TAtom Value"
using transaction_selects_are_Value_vars[OF T_wf T_adm'(1)]
by fastforce
have "select⟨?θ y, Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t ?θ))"
using subst_lsst_unlabel_member[OF s]
by fastforce
hence "((ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ, Fun (Set s) []) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using wellformed_transaction_sem_pos_checks[
OF T_wf ℐ_is_T_model,
of Assign "(ξ ∘⇩s σ ∘⇩s α) y" "Fun (Set s) []"]
by simp
thus "∃ss. (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc ss ∧ s ∈ ss"
using to_abs_alt_def[of "db⇩l⇩s⇩s⇩t 𝒜 ℐ"] 4[of y] y y_val by auto
qed
show ?D
proof (intro ballI allI impI)
fix y s
assume y: "y ∈ fv_transaction T - set (transaction_fresh T)"
and s: "⟨Var y in Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))"
hence "⟨Var y in Fun (Set s) []⟩ ∈ set (unlabel (transaction_strand T))"
unfolding transaction_strand_def unlabel_def by auto
hence y_val: "Γ⇩v y = TAtom Value"
using transaction_inset_checks_are_Value_vars[OF T_adm]
by fastforce
have "⟨?θ y in Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t ?θ))"
using subst_lsst_unlabel_member[OF s]
by fastforce
hence "(?θ y ⋅ ℐ, Fun (Set s) []) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using wellformed_transaction_sem_pos_checks[
OF T_wf ℐ_is_T_model,
of Check "?θ y" "Fun (Set s) []"]
by simp
thus "∃ss. ?θ y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc ss ∧ s ∈ ss"
using to_abs_alt_def[of "db⇩l⇩s⇩s⇩t 𝒜 ℐ"] 4[of y] y y_val by auto
qed
show ?E
proof (intro ballI allI impI)
fix y s
assume y: "y ∈ fv_transaction T - set (transaction_fresh T)"
and s: "⟨Var y not in Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))"
hence "⟨Var y not in Fun (Set s) []⟩ ∈ set (unlabel (transaction_strand T))"
unfolding transaction_strand_def unlabel_def by auto
hence y_val: "Γ⇩v y = TAtom Value"
using transaction_notinset_checks_are_Value_vars(1)[
OF T_adm, of "[]" "[]" "[(Var y, Fun (Set s) [])]" "Var y" "Fun (Set s) []"]
by force
have "⟨?θ y not in Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t ?θ))"
using subst_lsst_unlabel_member[OF s]
by fastforce
hence "(?θ y ⋅ ℐ, Fun (Set s) []) ∉ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using wellformed_transaction_sem_neg_checks'(2)[
OF T_wf ℐ_is_T_model,
of "[]" "?θ y" "Fun (Set s) []"]
by simp
moreover have "list_all admissible_transaction_updates P"
using Ball_set[of P "admissible_transaction"] P(1)
Ball_set[of P admissible_transaction_updates]
admissible_transaction_is_wellformed_transaction(3)
by fast
moreover have "list_all wellformed_transaction P"
using P(1) Ball_set[of P "admissible_transaction"] Ball_set[of P wellformed_transaction]
admissible_transaction_is_wellformed_transaction(1)
by blast
ultimately have "((ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ, Fun (Set s) S) ∉ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)" for S
using reachable_constraints_db⇩l⇩s⇩s⇩t_set_args_empty[OF 𝒜_reach]
unfolding admissible_transaction_updates_def
by auto
thus "∃ss. (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = absc ss ∧ s ∉ ss"
using to_abs_alt_def[of "db⇩l⇩s⇩s⇩t 𝒜 ℐ"] 4[of y] y y_val by auto
qed
show ?F
proof (intro ballI impI)
fix y assume y: "y ∈ fv_transaction T - set (transaction_fresh T)" "Γ⇩v y = TAtom Value"
then obtain yn where yn: "?θ y ⋅ ℐ = Fun (Val yn) []" using 4 by blast
hence y_abs: "?θ y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) = Fun (Abs (α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) yn)) []" by simp
have "y ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∨ (y ∈ fv⇩l⇩s⇩s⇩t (transaction_checks T) ∧
(∃t s. select⟨t,s⟩ ∈ set (unlabel (transaction_checks T)) ∧ y ∈ fv t ∪ fv s))"
using admissible_transaction_fv_in_receives_or_selects[OF T_adm] y by blast
thus "?θ y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) ∈ absc ` set OCC"
proof
assume "y ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T)"
then obtain ts where ts:
"receive⟨ts⟩ ∈ set (unlabel (transaction_receive T))" "y ∈ fv⇩s⇩e⇩t (set ts)"
using wellformed_transaction_unlabel_cases(1)[OF T_wf]
by (force simp add: unlabel_def)
have *: "?θ y ⋅ ℐ ∈ subterms⇩s⇩e⇩t (set ts ⋅⇩s⇩e⇩t ?θ ∘⇩s ℐ)"
"list_all (λt. timpl_closure_set (set FP) (set TI) ⊢⇩c t ⋅ ?θ ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)) ts"
using ts fv_subterms_substI[of y _ "?θ ∘⇩s ℐ"] subst_compose[of ?θ ℐ y]
subterms_subst_subset[of "?θ ∘⇩s ℐ"] receives_covered
unfolding intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, symmetric] list_all_iff
by fastforce+
have "Abs (α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) yn) ∈ ⋃(funs_term ` (timpl_closure_set (set FP) (set TI)))"
using * y_abs abs_subterms_in[of "?θ y ⋅ ℐ" _ "α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"]
ideduct_synth_priv_fun_in_ik[
OF _ funs_term_Fun_subterm'[of "Abs (α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) yn)" "[]"]]
unfolding list_all_iff by fastforce
hence "?θ y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) ∈ subterms⇩s⇩e⇩t (timpl_closure_set (set FP) (set TI))"
using y_abs wf_trms_subterms[OF timpl_closure_set_wf_trms[OF FP(2), of "set TI"]]
funs_term_Fun_subterm[of "Abs (α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ) yn)"]
unfolding wf⇩t⇩r⇩m_def by fastforce
hence "funs_term (?θ y ⋅ ℐ ⋅⇩α α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ))
⊆ (⋃t ∈ timpl_closure_set (set FP) (set TI). funs_term t)"
using funs_term_subterms_eq(2)[of "timpl_closure_set (set FP) (set TI)"] by blast
thus ?thesis using y_abs OCC(1) by fastforce
next
assume "y ∈ fv⇩l⇩s⇩s⇩t (transaction_checks T) ∧
(∃t s. select⟨t,s⟩ ∈ set (unlabel (transaction_checks T)) ∧ y ∈ fv t ∪ fv s)"
then obtain t t' where
"select⟨t,t'⟩ ∈ set (unlabel (transaction_checks T))" "y ∈ fv t ∪ fv t'"
by blast
then obtain l s where "(l,select⟨Var y, Fun (Set s) []⟩) ∈ set (transaction_checks T)"
using admissible_transaction_strand_step_cases(2)[OF T_adm]
unfolding unlabel_def by fastforce
then obtain U where U:
"prefix (U@[(l,select⟨Var y, Fun (Set s) []⟩)]) (transaction_checks T)"
using in_set_conv_decomp[of "(l, select⟨Var y,Fun (Set s) []⟩)" "transaction_checks T"]
by (auto simp add: prefix_def)
hence "select⟨Var y, Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))"
by (force simp add: prefix_def unlabel_def)
hence "select⟨?θ y, Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t ?θ))"
using subst_lsst_unlabel_member
by fastforce
hence "(Fun (Val yn) [], Fun (Set s) []) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using yn wellformed_transaction_sem_pos_checks[
OF T_wf ℐ_is_T_model, of Assign "?θ y" "Fun (Set s) []"]
by fastforce
hence "Fun (Val yn) [] ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜) ⋅⇩s⇩e⇩t ℐ"
using db⇩s⇩s⇩t_in_cases[of "Fun (Val yn) []"]
by (fastforce simp add: db⇩s⇩s⇩t_def)
thus ?thesis
using OCC(3) yn abs_in[of "Fun (Val yn) []" _ "α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"]
unfolding abs_value_constants_def
by (metis (mono_tags, lifting) mem_Collect_eq subsetCE)
qed
qed
qed
lemma transaction_prop4:
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and P: "∀T ∈ set P. admissible_transaction' T"
and P_occ: "∀T ∈ set P. admissible_transaction_occurs_checks T"
and x: "x ∈ set (transaction_fresh T)"
and y: "y ∈ fv_transaction T - set (transaction_fresh T)" "Γ⇩v y = TAtom Value"
shows "(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ))" (is ?A)
and "(ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ))" (is ?B)
proof -
let ?T' = "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
from ℐ have ℐ': "welltyped_constraint_model ℐ 𝒜"
using welltyped_constraint_model_prefix by auto
note T_P_adm = bspec[OF P(1)]
note T_adm = bspec[OF P(1) T]
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
have be: "bvars⇩l⇩s⇩s⇩t 𝒜 = {}"
using T_P_adm 𝒜_reach reachable_constraints_no_bvars admissible_transaction_no_bvars(2)
by blast
have T_no_bvars: "fv_transaction T = vars_transaction T"
using admissible_transaction_no_bvars[OF T_adm] by simp
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm ξ]
have ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t ℐ" by (metis ℐ welltyped_constraint_model_def)
have x_val: "fst x = TAtom Value"
using x admissible_transactionE(14)[OF T_adm] unfolding list_all_iff by blast
obtain xn where xn: "σ x = Fun (Val xn) []"
using x_val transaction_fresh_subst_sends_to_val[OF σ x] Γ⇩v_TAtom''(2)[of x] by meson
hence xnxn: "(ξ ∘⇩s σ ∘⇩s α) x = Fun (Val xn) []"
unfolding subst_compose_def ξ_empty by auto
from xn xnxn have a0: "(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ = Fun (Val xn) []"
by auto
have b0: "Γ⇩v x = TAtom Value"
using P x T protocol_transaction_vars_TAtom_typed(3)
by metis
note 0 = a0 b0
have σ_x_nin_A: "σ x ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
proof -
have "σ x ∈ subst_range σ"
by (metis transaction_fresh_subst_sends_to_val[OF σ x b0])
moreover
have "∀t ∈ subst_range σ. t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
using σ transaction_fresh_subst_def[of σ T "trms⇩l⇩s⇩s⇩t 𝒜"] by blast
ultimately
show ?thesis
by auto
qed
have *: "y ∉ set (transaction_fresh T)"
using assms by auto
have **: "y ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∨ (y ∈ fv⇩l⇩s⇩s⇩t (transaction_checks T) ∧
(∃t s. select⟨t,s⟩ ∈ set (unlabel (transaction_checks T)) ∧ y ∈ fv t ∪ fv s))"
using * y admissible_transaction_fv_in_receives_or_selects[OF T_adm]
by blast
have y_fv: "y ∈ fv_transaction T" using y fv_transaction_unfold by blast
have y_val: "fst y = TAtom Value" using y(2) Γ⇩v_TAtom''(2) by blast
have "σ x ⋅ ℐ ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ))"
proof (rule ccontr)
assume "¬σ x ⋅ ℐ ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ))"
then have a: "σ x ⋅ ℐ ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ))"
by auto
then have σ_x_I_in_A: "σ x ⋅ ℐ ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜) ⋅⇩s⇩e⇩t ℐ"
using reachable_constraints_subterms_subst[OF 𝒜_reach ℐ' P] by blast
have "∃u. u ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∧ ℐ u = σ x"
proof -
from σ_x_I_in_A have "∃tu. tu ∈ ⋃ (subterms ` (trms⇩l⇩s⇩s⇩t 𝒜)) ∧ tu ⋅ ℐ = σ x ⋅ ℐ"
by force
then obtain tu where tu: "tu ∈ ⋃ (subterms ` (trms⇩l⇩s⇩s⇩t 𝒜)) ∧ tu ⋅ ℐ = σ x ⋅ ℐ"
by auto
then have "tu ≠ σ x"
using σ_x_nin_A by auto
moreover
have "tu ⋅ ℐ = σ x"
using tu by (simp add: xn)
ultimately
have "∃u. tu = Var u"
unfolding xn by (cases tu) auto
then obtain u where "tu = Var u"
by auto
have "u ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∧ ℐ u = σ x"
proof -
have "u ∈ vars⇩l⇩s⇩s⇩t 𝒜"
using ‹tu = Var u› tu var_subterm_trms⇩s⇩s⇩t_is_vars⇩s⇩s⇩t by fastforce
then have "u ∈ fv⇩l⇩s⇩s⇩t 𝒜"
using be vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel 𝒜"] by blast
moreover
have "ℐ u = σ x"
using ‹tu = Var u› ‹tu ⋅ ℐ = σ x› by auto
ultimately
show ?thesis
by auto
qed
then show "∃u. u ∈ fv⇩l⇩s⇩s⇩t 𝒜 ∧ ℐ u = σ x"
by metis
qed
then obtain u where u:
"u ∈ fv⇩l⇩s⇩s⇩t 𝒜" "ℐ u = σ x"
by auto
then have u_TA: "Γ⇩v u = TAtom Value"
using P(1) T x_val Γ⇩v_TAtom''(2)[of x]
wt_subst_trm''[OF ℐ_wt, of "Var u"] wt_subst_trm''[of σ "Var x"]
transaction_fresh_subst_wt[OF σ] protocol_transaction_vars_TAtom_typed(3)
by force
have "∃B. prefix B 𝒜 ∧ u ∉ fv⇩l⇩s⇩s⇩t B ∧ ℐ u ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)"
using u u_TA
by (metis welltyped_constraint_model_prefix[OF ℐ]
constraint_model_Value_var_in_constr_prefix[OF 𝒜_reach _ P P_occ])
then obtain B where "prefix B 𝒜 ∧ u ∉ fv⇩l⇩s⇩s⇩t B ∧ ℐ u ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)"
by blast
moreover have "⋃(subterms ` trms⇩l⇩s⇩s⇩t xs) ⊆ ⋃(subterms ` trms⇩l⇩s⇩s⇩t ys)"
when "prefix xs ys"
for xs ys::"('fun,'atom,'sets,'lbl) prot_strand"
using that subterms⇩s⇩e⇩t_mono trms⇩s⇩s⇩t_mono unlabel_mono set_mono_prefix by metis
ultimately have "ℐ u ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
by blast
then have "σ x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
using u by auto
then show "False"
using σ_x_nin_A by auto
qed
then show ?A
using eval_term.simps(1)[of _ x σ]
unfolding subst_compose_def xn ξ_empty by auto
from ** show ?B
proof
define T' where "T' ≡ transaction_receive T"
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
assume y: "y ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T)"
hence "Var y ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t T')" by (metis T'_def fv⇩s⇩s⇩t_is_subterm_trms⇩s⇩s⇩t)
then obtain z where z: "z ∈ set (unlabel T')" "Var y ∈ subterms⇩s⇩e⇩t (trms⇩s⇩s⇩t⇩p z)"
by (induct T') auto
have "is_Receive z"
using Ball_set[of "unlabel T'" is_Receive] z(1)
admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
unfolding wellformed_transaction_def T'_def
by blast
then obtain tys where "z = receive⟨tys⟩" by (cases z) auto
hence tys: "receive⟨tys ⋅⇩l⇩i⇩s⇩t θ⟩ ∈ set (unlabel (T' ⋅⇩l⇩s⇩s⇩t θ))" "θ y ∈ subterms⇩s⇩e⇩t (set tys ⋅⇩s⇩e⇩t θ)"
using z subst_mono unfolding subst_apply_labeled_stateful_strand_def unlabel_def by force+
hence y_deduct: "list_all (λt. ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ θ ⋅ ℐ) tys"
using transaction_receive_deduct[OF T_wf _ ξ σ α] ℐ
unfolding T'_def θ_def welltyped_constraint_model_def list_all_iff by auto
obtain ty where ty: "ty ∈ set tys" "θ y ⊑ ty ⋅ θ" "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ ty ⋅ θ ⋅ ℐ"
using tys y_deduct unfolding list_all_iff by blast
obtain zn where zn: "(ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ = Fun (Val zn) []"
using transaction_var_becomes_Val[
OF reachable_constraints.step[OF 𝒜_reach T ξ σ α] ℐ ξ σ α P P_occ T y_fv y_val]
by metis
have "(ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ)"
using ty tys(2) y_deduct private_fun_deduct_in_ik[of _ _ "Val zn"]
by (metis θ_def zn subst_mono public.simps(3))
thus ?B
using ik⇩s⇩s⇩t_subst[of "unlabel 𝒜" ℐ] unlabel_subst[of 𝒜 ℐ]
subterms⇩s⇩e⇩t_mono[OF ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset[of "unlabel (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ)"]]
by fastforce
next
assume y': "y ∈ fv⇩l⇩s⇩s⇩t (transaction_checks T) ∧
(∃t s. select⟨t,s⟩ ∈ set (unlabel (transaction_checks T)) ∧ y ∈ fv t ∪ fv s)"
then obtain s where s: "select⟨Var y,s⟩ ∈ set (unlabel (transaction_checks T))"
"fst y = TAtom Value"
using admissible_transaction_strand_step_cases(1,2)[OF T_adm] by fastforce
obtain z zn where zn: "(ξ ∘⇩s σ ∘⇩s α) y = Var z" "ℐ z = Fun (Val zn) []"
using transaction_var_becomes_Val[
OF reachable_constraints.step[OF 𝒜_reach T ξ σ α] ℐ ξ σ α P P_occ T y_fv s(2)]
transaction_decl_fresh_renaming_substs_range(4)[OF ξ σ α _ *]
transaction_decl_subst_empty_inv[OF ξ[unfolded ξ_empty]]
by auto
have transaction_selects_db_here:
"⋀n s. select⟨Var (TAtom Value, n), Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T))
⟹ (α (TAtom Value, n) ⋅ ℐ, Fun (Set s) []) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using transaction_selects_db[OF T_adm _ ξ σ α] ℐ
unfolding welltyped_constraint_model_def by auto
have "∃n. y = (Var Value, n)"
using T Γ⇩v_TAtom_inv(2) y_fv y(2)
by blast
moreover
have "admissible_transaction_checks T"
using admissible_transaction_is_wellformed_transaction(2)[OF T_adm]
by blast
then have "is_Fun_Set (the_set_term (select⟨Var y,s⟩))"
using s unfolding admissible_transaction_checks_def
by auto
then have "∃ss. s = Fun (Set ss) []"
using is_Fun_Set_exi
by auto
ultimately
obtain n ss where nss: "y = (TAtom Value, n)" "s = Fun (Set ss) []"
by auto
then have "select⟨Var (TAtom Value, n), Fun (Set ss) []⟩ ∈ set (unlabel (transaction_checks T))"
using s by auto
then have in_db: "(α (TAtom Value, n) ⋅ ℐ, Fun (Set ss) []) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using transaction_selects_db_here[of n ss] by auto
have "(ℐ z, s) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
proof -
have "(α y ⋅ ℐ, s) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
using in_db nss by auto
moreover
have "α y = Var z"
using zn ξ_empty * σ
by (metis (no_types, opaque_lifting) subst_compose_def subst_imgI subst_to_var_is_var
term.distinct(1) transaction_fresh_subst_def var_comp(2))
then have "α y ⋅ ℐ = ℐ z"
by auto
ultimately
show "(ℐ z, s) ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
by auto
qed
then have "∃t' s'. insert⟨t',s'⟩ ∈ set (unlabel 𝒜) ∧ ℐ z = t' ⋅ ℐ ∧ s = s' ⋅ ℐ"
using db⇩s⇩s⇩t_in_cases[of "ℐ z" s "unlabel 𝒜" ℐ "[]"] unfolding db⇩s⇩s⇩t_def by auto
then obtain t' s' where t's': "insert⟨t',s'⟩ ∈ set (unlabel 𝒜) ∧ ℐ z = t' ⋅ ℐ ∧ s = s' ⋅ ℐ"
by auto
then have "t' ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)"
by force
then have "t' ⋅ ℐ ∈ (subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)) ⋅⇩s⇩e⇩t ℐ"
by auto
then have "ℐ z ∈ (subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)) ⋅⇩s⇩e⇩t ℐ"
using t's' by auto
then have "ℐ z ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ))"
using reachable_constraints_subterms_subst[
OF 𝒜_reach welltyped_constraint_model_prefix[OF ℐ] P]
by auto
then show ?B
using zn(1) by simp
qed
qed
lemma transaction_prop5:
fixes T ξ σ α 𝒜 ℐ T' a0 a0' θ
defines "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
and "a0 ≡ α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
and "a0' ≡ α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@T') ℐ)"
and "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@T')"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
"∀t ∈ α⇩i⇩k 𝒜 ℐ. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
"α⇩v⇩a⇩l⇩s 𝒜 ℐ ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and P:
"∀T ∈ set P. admissible_transaction' T"
and P_occ:
"∀T ∈ set P. admissible_transaction_occurs_checks T"
and step: "list_all (transaction_check (FP, OCC, TI)) P"
shows "∃δ ∈ abs_substs_fun ` set (transaction_check_comp (λ_ _. True) (FP, OCC, TI) T).
∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ⟶
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0 = absc (δ x) ∧
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0' = absc (absdbupd (unlabel (transaction_updates T)) x (δ x))"
proof -
define comp0 where
"comp0 ≡ abs_substs_fun ` set (transaction_check_comp (λ_ _. True) (FP, OCC, TI) T)"
define check0 where "check0 ≡ transaction_check (FP, OCC, TI) T"
define upd where "upd ≡ λδ x. absdbupd (unlabel (transaction_updates T)) x (δ x)"
define b0 where "b0 ≡ λx. THE b. absc b = (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0"
note all_defs = comp0_def check0_def a0_def a0'_def upd_def b0_def θ_def T'_def
have 𝒜_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)"
by (metis reachable_constraints_wf⇩t⇩r⇩m⇩s admissible_transactions_wf⇩t⇩r⇩m⇩s P(1) 𝒜_reach)
have ℐ_interp: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ"
and ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t ℐ"
and ℐ_wf_trms: "wf⇩t⇩r⇩m⇩s (subst_range ℐ)"
by (metis ℐ welltyped_constraint_model_def constraint_model_def,
metis ℐ welltyped_constraint_model_def,
metis ℐ welltyped_constraint_model_def constraint_model_def)
have ℐ_is_T_model: "strand_sem_stateful (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) (set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)) (unlabel T') ℐ"
using ℐ unlabel_append[of 𝒜 T'] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel 𝒜" ℐ "[]"]
strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" "unlabel T'" ℐ]
by (simp add: welltyped_constraint_model_def constraint_model_def db⇩s⇩s⇩t_def)
note T_adm = bspec[OF P(1) T]
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
have T_no_bvars: "fv_transaction T = vars_transaction T" "bvars_transaction T = {}"
using admissible_transaction_no_bvars[OF T_adm] by simp_all
have T_vars_const_typed: "∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))"
and T_fresh_vars_value_typed: "∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value"
using T P protocol_transaction_vars_TAtom_typed(2,3)[of T] by simp_all
note ξ_empty = admissible_transaction_decl_subst_empty[OF T_adm ξ]
have wt_σαℐ: "wt⇩s⇩u⇩b⇩s⇩t (ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ)" and wt_σα: "wt⇩s⇩u⇩b⇩s⇩t (ξ ∘⇩s σ ∘⇩s α)"
using ℐ_wt wt_subst_compose transaction_decl_fresh_renaming_substs_wt[OF ξ σ α]
by (blast, blast)
have T_vars_vals: "∀x ∈ fv_transaction T. ∃n. (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ = Fun (Val n) []"
proof
fix x assume x: "x ∈ fv_transaction T"
have "∃n. (σ ∘⇩s α) x ⋅ ℐ = Fun (Val n) []"
proof (cases "x ∈ subst_domain σ")
case True
then obtain n where "σ x = Fun (Val n) []"
using transaction_fresh_subst_sends_to_val[OF σ]
transaction_fresh_subst_domain[OF σ]
T_fresh_vars_value_typed
by metis
thus ?thesis by (simp add: subst_compose_def)
next
case False
hence *: "(σ ∘⇩s α) x = α x" by (auto simp add: subst_compose_def)
obtain y where y: "Γ⇩v x = Γ⇩v y" "α x = Var y"
using transaction_renaming_subst_wt[OF α]
transaction_renaming_subst_is_renaming(1)[OF α]
by (metis Γ.simps(1) prod.exhaust wt⇩s⇩u⇩b⇩s⇩t_def)
hence "y ∈ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t σ ∘⇩s α)"
using x * T_no_bvars(2) unlabel_subst[of "transaction_strand T" "σ ∘⇩s α"]
fv⇩s⇩s⇩t_subst_fv_subset[of x "unlabel (transaction_strand T)" "σ ∘⇩s α"]
by auto
hence "y ∈ fv⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t σ ∘⇩s α))"
using fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t σ ∘⇩s α"]
fv⇩s⇩s⇩t_append[of "unlabel 𝒜"] unlabel_append[of 𝒜]
by auto
thus ?thesis
using x y * T P
constraint_model_Value_term_is_Val[
OF reachable_constraints.step[OF 𝒜_reach T ξ σ α] ℐ[unfolded T'_def] P P_occ, of y]
admissible_transaction_Value_vars[of T] ξ_empty
by simp
qed
thus "∃n. (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ = Fun (Val n) []" using ξ_empty by simp
qed
have T_vars_absc: "∀x ∈ fv_transaction T. ∃!n. (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0 = absc n"
using T_vars_vals by fastforce
hence "(absc ∘ b0) x = (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0" when "x ∈ fv_transaction T" for x
using that unfolding b0_def by fastforce
hence T_vars_absc': "t ⋅ (absc ∘ b0) = t ⋅ (ξ ∘⇩s σ ∘⇩s α) ⋅ ℐ ⋅⇩α a0"
when "fv t ⊆ fv_transaction T" "∄n T. Fun (Val n) T ∈ subterms t" for t
using that(1) abs_term_subst_eq'[OF _ that(2), of "ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ" a0 "absc ∘ b0"]
subst_compose[of "ξ ∘⇩s σ ∘⇩s α" ℐ] subst_subst_compose[of t "ξ ∘⇩s σ ∘⇩s α" ℐ]
by fastforce
have "∃δ ∈ comp0. ∀x ∈ fv_transaction T. fst x = TAtom Value ⟶ b0 x = δ x"
proof -
let ?C = "set (unlabel (transaction_checks T))"
let ?xs = "fv_transaction T - set (transaction_fresh T)"
note * = transaction_prop3[OF 𝒜_reach T ℐ[unfolded T'_def] ξ σ α FP OCC TI P P_occ]
have **:
"∀x ∈ set (transaction_fresh T). b0 x = {}"
"∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). intruder_synth_mod_timpls FP TI (t ⋅ θ b0)"
(is ?B)
proof -
show ?B
proof (intro ballI impI)
fix t assume t: "t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T)"
hence t': "fv t ⊆ fv_transaction T" "∄n T. Fun (Val n) T ∈ subterms t"
using trms_transaction_unfold[of T] vars_transaction_unfold[of T]
trms⇩s⇩s⇩t_fv_vars⇩s⇩s⇩t_subset[of t "unlabel (transaction_strand T)"]
admissible_transactions_no_Value_consts'[OF T_adm]
wellformed_transaction_send_receive_fv_subset(1)[OF T_wf t(1)]
by blast+
have "intruder_synth_mod_timpls FP TI (t ⋅ (absc ∘ b0))"
using t(1) t' *(2) T_vars_absc'
by (metis a0_def)
moreover have "(absc ∘ b0) x = (θ b0) x" when "x ∈ fv t" for x
using that T P admissible_transaction_Value_vars[of T]
‹fv t ⊆ fv_transaction T› Γ⇩v_TAtom''(2)[of x]
unfolding θ_def by fastforce
hence "t ⋅ (absc ∘ b0) = t ⋅ θ b0"
using term_subst_eq[of t "absc ∘ b0" "θ b0"] by argo
ultimately show "intruder_synth_mod_timpls FP TI (t ⋅ θ b0)"
using intruder_synth.simps[of "set FP"] by (cases "t ⋅ θ b0") metis+
qed
qed (simp add: *(1) a0_def b0_def)
have ***: "∀x ∈ ?xs. ∀s. select⟨Var x,Fun (Set s) []⟩ ∈ ?C ⟶ s ∈ b0 x"
"∀x ∈ ?xs. ∀s. ⟨Var x in Fun (Set s) []⟩ ∈ ?C ⟶ s ∈ b0 x"
"∀x ∈ ?xs. ∀s. ⟨Var x not in Fun (Set s) []⟩ ∈ ?C ⟶ s ∉ b0 x"
"∀x ∈ ?xs. fst x = TAtom Value ⟶ b0 x ∈ set OCC"
unfolding a0_def b0_def
using *(3,4) apply (force, force)
using *(5) apply force
using *(6) admissible_transaction_Value_vars[OF bspec[OF P T]] by force
show ?thesis
using transaction_check_comp_in[OF T_adm **[unfolded θ_def] ***]
unfolding comp0_def
by metis
qed
hence 1: "∃δ ∈ comp0. ∀x ∈ fv_transaction T.
fst x = TAtom Value ⟶ (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0 = absc (δ x)"
using T_vars_absc unfolding b0_def a0_def by fastforce
obtain δ where δ:
"δ ∈ comp0"
"∀x ∈ fv_transaction T. fst x = TAtom Value ⟶ (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0 = absc (δ x)"
using 1 by force
have 2: "θ x ⋅ ℐ ⋅⇩α α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) = absc (absdbupd (unlabel A) x d)"
when "θ x ⋅ ℐ ⋅⇩α α⇩0 D = absc d"
and "∀t u. insert⟨t,u⟩ ∈ set (unlabel A) ⟶ (∃y s. t = Var y ∧ u = Fun (Set s) [])"
and "∀t u. delete⟨t,u⟩ ∈ set (unlabel A) ⟶ (∃y s. t = Var y ∧ u = Fun (Set s) [])"
and "∀y ∈ fv⇩l⇩s⇩s⇩t A. θ x ⋅ ℐ = θ y ⋅ ℐ ⟶ x = y"
and "∀y ∈ fv⇩l⇩s⇩s⇩t A. ∃n. θ y ⋅ ℐ = Fun (Val n) []"
and x: "θ x ⋅ ℐ = Fun (Val n) []"
and D: "∀d ∈ set D. ∃s. snd d = Fun (Set s) []"
for A::"('fun,'atom,'sets,'lbl) prot_strand" and x θ D n d
using that(2,3,4,5)
proof (induction A rule: List.rev_induct)
case (snoc a A)
then obtain l b where a: "a = (l,b)" by (metis surj_pair)
have IH: "α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n = absdbupd (unlabel A) x d"
using snoc unlabel_append[of A "[a]"] a x
by (simp del: unlabel_append)
have b_prems: "∀y ∈ fv⇩s⇩s⇩t⇩p b. θ x ⋅ ℐ = θ y ⋅ ℐ ⟶ x = y"
"∀y ∈ fv⇩s⇩s⇩t⇩p b. ∃n. θ y ⋅ ℐ = Fun (Val n) []"
using snoc.prems(3,4) a by (simp_all add: unlabel_def)
have *: "filter is_Update (unlabel (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ))) =
filter is_Update (unlabel (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)))"
"filter is_Update (unlabel (A@[a])) = filter is_Update (unlabel A)"
when "¬is_Update b"
using that a
by (cases b, simp_all add: dual⇩l⇩s⇩s⇩t_def unlabel_def subst_apply_labeled_stateful_strand_def)+
note ** = IH a dual⇩l⇩s⇩s⇩t_subst_append[of A "[a]" θ]
note *** = * absdbupd_filter[of "unlabel (A@[a])"]
absdbupd_filter[of "unlabel A"]
db⇩s⇩s⇩t_filter[of "unlabel (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ))"]
db⇩s⇩s⇩t_filter[of "unlabel (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ))"]
note **** = **(2,3) dual⇩l⇩s⇩s⇩t_subst_snoc[of A a θ]
unlabel_append[of "dual⇩l⇩s⇩s⇩t A ⋅⇩l⇩s⇩s⇩t θ" "[dual⇩l⇩s⇩s⇩t⇩p a ⋅⇩l⇩s⇩s⇩t⇩p θ]"]
db⇩s⇩s⇩t_append[of "unlabel (dual⇩l⇩s⇩s⇩t A ⋅⇩l⇩s⇩s⇩t θ)" "unlabel [dual⇩l⇩s⇩s⇩t⇩p a ⋅⇩l⇩s⇩s⇩t⇩p θ]" ℐ D]
have "α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n = absdbupd (unlabel (A@[a])) x d" using ** ***
proof (cases b)
case (Insert t t')
then obtain y s m where y: "t = Var y" "t' = Fun (Set s) []" "θ y ⋅ ℐ = Fun (Val m) []"
using snoc.prems(1) b_prems(2) a by (fastforce simp add: unlabel_def)
hence a': "db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ)) ℐ D =
List.insert ((Fun (Val m) [], Fun (Set s) [])) (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t A ⋅⇩l⇩s⇩s⇩t θ) ℐ D)"
"unlabel [dual⇩l⇩s⇩s⇩t⇩p a ⋅⇩l⇩s⇩s⇩t⇩p θ] = [insert⟨θ y, Fun (Set s) []⟩]"
"unlabel [a] = [insert⟨Var y, Fun (Set s) []⟩]"
using **** Insert by simp_all
show ?thesis
proof (cases "x = y")
case True
hence "θ x ⋅ ℐ = θ y ⋅ ℐ" by simp
hence "α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n =
insert s (α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n)"
by (metis (no_types, lifting) y(3) a'(1) x dual⇩l⇩s⇩s⇩t_subst to_abs_list_insert')
thus ?thesis using True IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def)
next
case False
hence "θ x ⋅ ℐ ≠ θ y ⋅ ℐ" using b_prems(1) y Insert by simp
hence "α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n = α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n"
by (metis (no_types, lifting) y(3) a'(1) x dual⇩l⇩s⇩s⇩t_subst to_abs_list_insert)
thus ?thesis using False IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def)
qed
next
case (Delete t t')
then obtain y s m where y: "t = Var y" "t' = Fun (Set s) []" "θ y ⋅ ℐ = Fun (Val m) []"
using snoc.prems(2) b_prems(2) a by (fastforce simp add: unlabel_def)
hence a': "db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ)) ℐ D =
List.removeAll ((Fun (Val m) [], Fun (Set s) [])) (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t A ⋅⇩l⇩s⇩s⇩t θ) ℐ D)"
"unlabel [dual⇩l⇩s⇩s⇩t⇩p a ⋅⇩l⇩s⇩s⇩t⇩p θ] = [delete⟨θ y, Fun (Set s) []⟩]"
"unlabel [a] = [delete⟨Var y, Fun (Set s) []⟩]"
using **** Delete by simp_all
have "∃s S. snd d = Fun (Set s) []" when "d ∈ set (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t A ⋅⇩l⇩s⇩s⇩t θ) ℐ D)" for d
using snoc.prems(1,2) db⇩l⇩s⇩s⇩t_dual⇩l⇩s⇩s⇩t_set_ex[OF that _ _ D] by (simp add: unlabel_def)
moreover {
fix t::"('fun,'atom,'sets,'lbl) prot_term"
and D::"(('fun,'atom,'sets,'lbl) prot_term × ('fun,'atom,'sets,'lbl) prot_term) list"
assume "∀d ∈ set D. ∃s. snd d = Fun (Set s) []"
hence "removeAll (t, Fun (Set s) []) D = filter (λd. ∄S. d = (t, Fun (Set s) S)) D"
by (induct D) auto
} ultimately have a'':
"List.removeAll ((Fun (Val m) [], Fun (Set s) [])) (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t A ⋅⇩l⇩s⇩s⇩t θ) ℐ D) =
filter (λd. ∄S. d = (Fun (Val m) [], Fun (Set s) S)) (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t A ⋅⇩l⇩s⇩s⇩t θ) ℐ D)"
by simp
show ?thesis
proof (cases "x = y")
case True
hence "θ x ⋅ ℐ = θ y ⋅ ℐ" by simp
hence "α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n =
(α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n) - {s}"
using y(3) a'' a'(1) x by (simp add: dual⇩l⇩s⇩s⇩t_subst to_abs_list_remove_all')
thus ?thesis using True IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def)
next
case False
hence "θ x ⋅ ℐ ≠ θ y ⋅ ℐ" using b_prems(1) y Delete by simp
hence "α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A@[a] ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n = α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ)) ℐ D) n"
by (metis (no_types, lifting) y(3) a'(1) x dual⇩l⇩s⇩s⇩t_subst to_abs_list_remove_all)
thus ?thesis using False IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def)
qed
qed simp_all
thus ?case by (simp add: x)
qed (simp add: that(1))
have 3: "x = y"
when xy: "(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ" "x ∈ fv_transaction T" "y ∈ fv_transaction T"
for x y
proof -
have "x ∉ set (transaction_fresh T) ⟹ y ∉ set (transaction_fresh T) ⟹ ?thesis"
using xy admissible_transaction_strand_sem_fv_ineq[OF T_adm ℐ_is_T_model[unfolded T'_def]]
by fast
moreover {
assume *: "x ∈ set (transaction_fresh T)" "y ∈ set (transaction_fresh T)"
hence "Γ⇩v x = TAtom Value" "Γ⇩v y = TAtom Value"
using T_fresh_vars_value_typed by (blast, blast)
then obtain xn yn where "σ x = Fun (Val xn) []" "σ y = Fun (Val yn) []"
using * transaction_fresh_subst_sends_to_val[OF σ] by meson
hence "σ x = σ y" using that(1) ξ_empty by (simp add: subst_compose)
moreover have "inj_on σ (subst_domain σ)" "x ∈ subst_domain σ" "y ∈ subst_domain σ"
using * σ unfolding transaction_fresh_subst_def by auto
ultimately have ?thesis unfolding inj_on_def by blast
} moreover have False when "x ∈ set (transaction_fresh T)" "y ∉ set (transaction_fresh T)"
using that(2) xy T_no_bvars admissible_transaction_Value_vars[OF bspec[OF P T], of y]
transaction_prop4[OF 𝒜_reach T ℐ[unfolded T'_def] ξ σ α P P_occ that(1), of y]
by auto
moreover have False when "x ∉ set (transaction_fresh T)" "y ∈ set (transaction_fresh T)"
using that(1) xy T_no_bvars admissible_transaction_Value_vars[OF bspec[OF P T], of x]
transaction_prop4[OF 𝒜_reach T ℐ[unfolded T'_def] ξ σ α P P_occ that(2), of x]
by fastforce
ultimately show ?thesis by metis
qed
have 4: "∃y s. t = Var y ∧ u = Fun (Set s) []"
when "insert⟨t,u⟩ ∈ set (unlabel (transaction_strand T))" for t u
using that admissible_transaction_strand_step_cases(3)[OF T_adm] T_wf
by blast
have 5: "∃y s. t = Var y ∧ u = Fun (Set s) []"
when "delete⟨t,u⟩ ∈ set (unlabel (transaction_strand T))" for t u
using that admissible_transaction_strand_step_cases(3)[OF T_adm] T_wf
by blast
have 6: "∃n. (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ = Fun (Val n) []" when "y ∈ fv_transaction T" for y
using that by (simp add: T_vars_vals)
have "list_all wellformed_transaction P" "list_all admissible_transaction_updates P"
using P(1) Ball_set[of P admissible_transaction'] Ball_set[of P wellformed_transaction]
Ball_set[of P admissible_transaction_updates]
admissible_transaction_is_wellformed_transaction(1,3)
by fastforce+
hence 7: "∃s. snd d = Fun (Set s) []" when "d ∈ set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)" for d
using that reachable_constraints_db⇩l⇩s⇩s⇩t_set_args_empty[OF 𝒜_reach]
unfolding admissible_transaction_updates_def by (cases d) simp
have "(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0' = absc (upd δ x)"
when x: "x ∈ fv_transaction T" "fst x = TAtom Value" for x
proof -
have "(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α α⇩0 (db'⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ (db⇩l⇩s⇩s⇩t 𝒜 ℐ))
= absc (absdbupd (unlabel (transaction_strand T)) x (δ x))"
using 2[of "ξ ∘⇩s σ ∘⇩s α" x "db⇩l⇩s⇩s⇩t 𝒜 ℐ" "δ x" "transaction_strand T"]
3[OF _ x(1)] 4 5 6[OF that(1)] 6 7 x δ(2)
unfolding all_defs by blast
thus ?thesis
using x db⇩s⇩s⇩t_append[of "unlabel 𝒜"] absdbupd_wellformed_transaction[OF T_wf]
unfolding all_defs db⇩s⇩s⇩t_def by force
qed
thus ?thesis using δ Γ⇩v_TAtom''(2) unfolding all_defs by blast
qed
lemma transaction_prop6:
fixes T ξ σ α 𝒜 ℐ T' a0 a0'
defines "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
and "a0 ≡ α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
and "a0' ≡ α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@T') ℐ)"
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@T')"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
"∀t ∈ α⇩i⇩k 𝒜 ℐ. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
"α⇩v⇩a⇩l⇩s 𝒜 ℐ ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and P:
"∀T ∈ set P. admissible_transaction' T"
and P_occ:
"∀T ∈ set P. admissible_transaction_occurs_checks T"
and step: "list_all (transaction_check (FP, OCC, TI)) P"
shows "∀t ∈ timpl_closure_set (α⇩i⇩k 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ).
timpl_closure_set (set FP) (set TI) ⊢⇩c t" (is ?A)
and "timpl_closure_set (α⇩v⇩a⇩l⇩s 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ) ⊆ absc ` set OCC" (is ?B)
and "∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T). is_Fun (t ⋅ (ξ ∘⇩s σ ∘⇩s α) ⋅ ℐ ⋅⇩α a0') ⟶
timpl_closure_set (set FP) (set TI) ⊢⇩c t ⋅ (ξ ∘⇩s σ ∘⇩s α) ⋅ ℐ ⋅⇩α a0'" (is ?C)
and "∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ⟶
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0' ∈ absc ` set OCC" (is ?D)
proof -
define comp0 where
"comp0 ≡ abs_substs_fun ` set (transaction_check_comp (λ_ _. True) (FP, OCC, TI) T)"
define check0 where "check0 ≡ transaction_check (FP, OCC, TI) T"
define upd where "upd ≡ λδ x. absdbupd (unlabel (transaction_updates T)) x (δ x)"
define θ where "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
note T_adm = bspec[OF P T]
note T_occ = bspec[OF P_occ T]
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
have θ_prop: "θ σ x = absc (σ x)" when "Γ⇩v x = TAtom Value" for σ x
using that Γ⇩v_TAtom''(2)[of x] unfolding θ_def by simp
have 0: "∃δ ∈ comp0. ∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ⟶
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0 = absc (δ x) ∧
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0' = absc (upd δ x)"
using transaction_prop5[OF 𝒜_reach T ℐ[unfolded T'_def] ξ σ α FP OCC TI P P_occ step]
unfolding a0_def a0'_def T'_def upd_def comp0_def
by blast
have 1: "(δ x, upd δ x) ∈ (set TI)⇧+"
when "δ ∈ comp0" "δ x ≠ upd δ x" "x ∈ fv_transaction T" "x ∉ set (transaction_fresh T)"
for x δ
using T that step Ball_set[of P "transaction_check (FP, OCC, TI)"]
transaction_prop1[of δ "λ_ _. True" FP OCC TI T x] TI
unfolding upd_def comp0_def transaction_check_def
by blast
have 2: "upd δ x ∈ set OCC"
when "δ ∈ comp0" "x ∈ fv_transaction T" "fst x = TAtom Value" for x δ
using T that step Ball_set[of P "transaction_check (FP, OCC, TI)"]
T_adm T_occ FP OCC TI transaction_prop2[of δ "λ_ _. True" FP OCC TI T x]
unfolding upd_def comp0_def transaction_check_def
by blast
obtain δ where δ:
"δ ∈ comp0"
"∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ⟶
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0 = absc (δ x) ∧
(ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0' = absc (upd δ x)"
using 0 by force
have "∃x. ab = (δ x, upd δ x) ∧ x ∈ fv_transaction T - set (transaction_fresh T) ∧ δ x ≠ upd δ x"
when ab: "ab ∈ α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ" for ab
proof -
obtain a b where ab': "ab = (a,b)" by (metis surj_pair)
then obtain x where x:
"a ≠ b" "x ∈ fv_transaction T" "x ∉ set (transaction_fresh T)"
"absc a = (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0" "absc b = (ξ ∘⇩s σ ∘⇩s α) x ⋅ ℐ ⋅⇩α a0'"
using ab unfolding abs_term_implications_def a0_def a0'_def T'_def by blast
hence "absc a = absc (δ x)" "absc b = absc (upd δ x)"
using δ(2) admissible_transaction_Value_vars[OF bspec[OF P T] x(2)]
by metis+
thus ?thesis using x ab' by blast
qed
hence α⇩t⇩i_TI_subset: "α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ ⊆ {(a,b) ∈ (set TI)⇧+. a ≠ b}" using 1[OF δ(1)] by blast
have "timpl_closure_set (timpl_closure_set (set FP) (set TI)) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ) ⊢⇩c t"
when t: "t ∈ timpl_closure_set (α⇩i⇩k 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)" for t
using timpl_closure_set_is_timpl_closure_union[of "α⇩i⇩k 𝒜 ℐ" "α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"]
intruder_synth_timpl_closure_set FP(3) t
by blast
thus ?A
using ideduct_synth_mono[OF _ timpl_closure_set_mono[OF
subset_refl[of "timpl_closure_set (set FP) (set TI)"]
α⇩t⇩i_TI_subset]]
timpl_closure_set_timpls_trancl_eq'[of "timpl_closure_set (set FP) (set TI)" "set TI"]
unfolding timpl_closure_set_idem
by force
have "timpl_closure_set (α⇩v⇩a⇩l⇩s 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ) ⊆
timpl_closure_set (absc ` set OCC) {(a,b) ∈ (set TI)⇧+. a ≠ b}"
using timpl_closure_set_mono[OF _ α⇩t⇩i_TI_subset] OCC(3) by blast
thus ?B using OCC(2) timpl_closure_set_timpls_trancl_subset' by blast
have "transaction_check_post (FP, OCC, TI) T δ"
using T δ(1) step
unfolding transaction_check_def transaction_check'_def comp0_def list_all_iff
by fastforce
hence 3: "timpl_closure_set (set FP) (set TI) ⊢⇩c t ⋅ θ (upd δ)"
when "t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T)" "is_Fun (t ⋅ θ (upd δ))" for t
using that
unfolding transaction_check_post_def upd_def θ_def
intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, symmetric]
by fastforce
have 4: "∀x ∈ fv t. (ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ) x ⋅⇩α a0' = θ (upd δ) x"
when "t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T)" for t
using wellformed_transaction_send_receive_fv_subset(2)[OF T_wf that]
δ(2) subst_compose[of "ξ ∘⇩s σ ∘⇩s α" ℐ] θ_prop
admissible_transaction_Value_vars[OF bspec[OF P T]]
by fastforce
have 5: "∄n T. Fun (Val n) T ∈ subterms t" when "t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T)" for t
using that admissible_transactions_no_Value_consts'[OF T_adm] trms_transaction_unfold[of T]
by blast
show ?D using 2[OF δ(1)] δ(2) Γ⇩v_TAtom''(2) unfolding a0'_def T'_def by blast
show ?C using 3 abs_term_subst_eq'[OF 4 5] by simp
qed
lemma reachable_constraints_covered_step:
fixes 𝒜::"('fun,'atom,'sets,'lbl) prot_constr"
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and T: "T ∈ set P"
and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)"
and FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
"∀t ∈ α⇩i⇩k 𝒜 ℐ. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
"ground (set FP)"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
"α⇩v⇩a⇩l⇩s 𝒜 ℐ ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and P:
"∀T ∈ set P. admissible_transaction' T"
and P_occ:
"∀T ∈ set P. admissible_transaction_occurs_checks T"
and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) P"
shows "∀t ∈ α⇩i⇩k (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ.
timpl_closure_set (set FP) (set TI) ⊢⇩c t" (is ?A)
and "α⇩v⇩a⇩l⇩s (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ℐ ⊆ absc ` set OCC" (is ?B)
proof -
note step_props =
transaction_prop6[OF 𝒜_reach T ℐ ξ σ α FP(1,2,3) OCC TI P P_occ transactions_covered]
define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define a0 where "a0 ≡ α⇩0 (db⇩l⇩s⇩s⇩t 𝒜 ℐ)"
define a0' where "a0' ≡ α⇩0 (db⇩l⇩s⇩s⇩t (𝒜@T') ℐ)"
define vals where "vals ≡ λS::('fun,'atom,'sets,'lbl) prot_constr.
{t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t S) ⋅⇩s⇩e⇩t ℐ. ∃n. t = Fun (Val n) []}"
define vals_sym where "vals_sym ≡ λS::('fun,'atom,'sets,'lbl) prot_constr.
{t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t S). (∃n. t = Fun (Val n) []) ∨ (∃m. t = Var (TAtom Value,m))}"
have ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t ℐ" by (metis ℐ welltyped_constraint_model_def)
have ℐ_grounds: "fv (t ⋅ ℐ) = {}" for t
using ℐ interpretation_grounds[of ℐ]
unfolding welltyped_constraint_model_def constraint_model_def by auto
have wt_σαℐ: "wt⇩s⇩u⇩b⇩s⇩t (ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ)" and wt_σα: "wt⇩s⇩u⇩b⇩s⇩t (ξ ∘⇩s σ ∘⇩s α)"
using ℐ_wt wt_subst_compose transaction_decl_fresh_renaming_substs_wt[OF ξ σ α]
by (blast, blast)
have "∀T∈set P. bvars_transaction T = {}"
using P admissible_transactionE(4) by metis
hence 𝒜_no_bvars: "bvars⇩l⇩s⇩s⇩t 𝒜 = {}"
using reachable_constraints_no_bvars[OF 𝒜_reach] by metis
have ℐ_vals: "∃n. ℐ (TAtom Value, m) = Fun (Val n) []"
when "(TAtom Value, m) ∈ fv⇩l⇩s⇩s⇩t 𝒜" for m
using constraint_model_Value_term_is_Val'[
OF 𝒜_reach welltyped_constraint_model_prefix[OF ℐ] P P_occ]
𝒜_no_bvars vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel 𝒜"] that
by blast
have vals_sym_vals: "t ⋅ ℐ ∈ vals 𝒜" when t: "t ∈ vals_sym 𝒜" for t
proof (cases t)
case (Var x)
then obtain m where *: "x = (TAtom Value,m)" using t unfolding vals_sym_def by blast
moreover have "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" using t unfolding vals_sym_def by blast
hence "t ⋅ ℐ ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜) ⋅⇩s⇩e⇩t ℐ" "∃n. ℐ (Var Value, m) = Fun (Val n) []"
using Var * ℐ_vals[of m] var_subterm_trms⇩s⇩s⇩t_is_vars⇩s⇩s⇩t[of x "unlabel 𝒜"]
Γ⇩v_TAtom[of Value m] reachable_constraints_Value_vars_are_fv[OF 𝒜_reach P(1), of x]
by blast+
ultimately show ?thesis using Var unfolding vals_def by auto
next
case (Fun f T)
then obtain n where "f = Val n" "T = []" using t unfolding vals_sym_def by blast
moreover have "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" using t unfolding vals_sym_def by blast
hence "t ⋅ ℐ ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜) ⋅⇩s⇩e⇩t ℐ" using Fun by blast
ultimately show ?thesis using Fun unfolding vals_def by auto
qed
have vals_vals_sym: "∃s. s ∈ vals_sym 𝒜 ∧ t = s ⋅ ℐ" when "t ∈ vals 𝒜" for t
using that constraint_model_Val_is_Value_term[OF ℐ]
unfolding vals_def vals_sym_def by fast
note T_adm = bspec[OF P T]
note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm]
have 0:
"α⇩i⇩k (𝒜@T') ℐ = (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0' ∪ (ik⇩l⇩s⇩s⇩t T' ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0'"
"α⇩v⇩a⇩l⇩s (𝒜@T') ℐ = vals 𝒜 ⋅⇩α⇩s⇩e⇩t a0' ∪ vals T' ⋅⇩α⇩s⇩e⇩t a0'"
by (metis abs_intruder_knowledge_append a0'_def,
metis abs_value_constants_append[of 𝒜 T' ℐ] a0'_def vals_def)
have 1: "(ik⇩l⇩s⇩s⇩t T' ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0' =
(trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t (ξ ∘⇩s σ ∘⇩s α) ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0'"
by (metis T'_def dual_transaction_ik_is_transaction_send''[OF T_wf])
have 2: "bvars⇩l⇩s⇩s⇩t (transaction_strand T) ∩ subst_domain ξ = {}"
"bvars⇩l⇩s⇩s⇩t (transaction_strand T) ∩ subst_domain σ = {}"
"bvars⇩l⇩s⇩s⇩t (transaction_strand T) ∩ subst_domain α = {}"
using admissible_transactionE(4)[OF T_adm] by blast+
have "vals T' ⊆ (ξ ∘⇩s σ ∘⇩s α) ` fv_transaction T ⋅⇩s⇩e⇩t ℐ"
proof
fix t assume "t ∈ vals T'"
then obtain s n where s:
"s ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t T')" "t = s ⋅ ℐ" "t = Fun (Val n) []"
unfolding vals_def by fast
then obtain u where u:
"u ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_strand T))"
"s = u ⋅ (ξ ∘⇩s σ ∘⇩s α)"
using transaction_decl_fresh_renaming_substs_trms[OF ξ σ α 2]
trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"]
unfolding T'_def by blast
have *: "t = u ⋅ (ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ)" using s(2) u(2) subst_subst_compose by simp
then obtain x where x: "u = Var x"
using s(3) admissible_transactions_no_Value_consts(1)[OF T_adm u(1)] by (cases u) force+
hence **: "x ∈ vars_transaction T"
by (metis u(1) var_subterm_trms⇩s⇩s⇩t_is_vars⇩s⇩s⇩t)
have "Γ⇩v x = TAtom Value"
using * x s(3) wt_subst_trm''[OF wt_σαℐ, of u]
by simp
thus "t ∈ (ξ ∘⇩s σ ∘⇩s α) ` fv_transaction T ⋅⇩s⇩e⇩t ℐ"
using admissible_transaction_Value_vars_are_fv[OF T_adm **] x *
eval_term.simps(1)[of _ x "ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ"]
subst_comp_set_image[of "ξ ∘⇩s σ ∘⇩s α" ℐ "fv_transaction T"]
by blast
qed
hence 3: "vals T' ⋅⇩α⇩s⇩e⇩t a0' ⊆ ((ξ ∘⇩s σ ∘⇩s α) ` fv_transaction T ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0'"
by (simp add: abs_apply_terms_def image_mono)
have "t ⋅ ℐ ⋅⇩α a0' ∈ timpl_closure_set (α⇩i⇩k 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)"
when "t ∈ ik⇩l⇩s⇩s⇩t 𝒜" for t
using that abs_in[OF imageI[OF that]]
α⇩t⇩i_covers_α⇩0_ik[OF 𝒜_reach T ℐ ξ σ α P P_occ]
timpl_closure_set_mono[of "{t ⋅ ℐ ⋅⇩α a0}" "α⇩i⇩k 𝒜 ℐ" "α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"
"α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"]
unfolding a0_def a0'_def T'_def abs_intruder_knowledge_def by fast
hence A: "α⇩i⇩k (𝒜@T') ℐ ⊆
timpl_closure_set (α⇩i⇩k 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ) ∪
(trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t (ξ ∘⇩s σ ∘⇩s α) ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0'"
using 0(1) 1 by (auto simp add: abs_apply_terms_def)
have "t ⋅ ℐ ⋅⇩α a0' ∈ timpl_closure_set {t ⋅ ℐ ⋅⇩α a0} (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)"
when t: "t ∈ vals_sym 𝒜" for t
proof -
have "(∃n. t = Fun (Val n) [] ∧ t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)) ∨
(∃n. t = Var (TAtom Value,n) ∧ (TAtom Value,n) ∈ fv⇩l⇩s⇩s⇩t 𝒜)"
(is "?P ∨ ?Q")
using t var_subterm_trms⇩s⇩s⇩t_is_vars⇩s⇩s⇩t[of _ "unlabel 𝒜"]
Γ⇩v_TAtom[of Value] reachable_constraints_Value_vars_are_fv[OF 𝒜_reach P(1)]
unfolding vals_sym_def by fast
thus ?thesis
proof
assume ?P
then obtain n where n: "t = Fun (Val n) []" "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" by force
thus ?thesis
using α⇩t⇩i_covers_α⇩0_Val[OF 𝒜_reach T ℐ ξ σ α P P_occ, of n]
unfolding a0_def a0'_def T'_def by fastforce
next
assume ?Q
thus ?thesis
using α⇩t⇩i_covers_α⇩0_Var[OF 𝒜_reach T ℐ ξ σ α P P_occ]
unfolding a0_def a0'_def T'_def by fastforce
qed
qed
moreover have "t ⋅ ℐ ⋅⇩α a0 ∈ α⇩v⇩a⇩l⇩s 𝒜 ℐ"
when "t ∈ vals_sym 𝒜" for t
using that abs_in vals_sym_vals
unfolding a0_def abs_value_constants_def vals_sym_def vals_def
by (metis (mono_tags, lifting))
ultimately have "t ⋅ ℐ ⋅⇩α a0' ∈ timpl_closure_set (α⇩v⇩a⇩l⇩s 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)"
when t: "t ∈ vals_sym 𝒜" for t
using t timpl_closure_set_mono[of "{t ⋅ ℐ ⋅⇩α a0}" "α⇩v⇩a⇩l⇩s 𝒜 ℐ" "α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"
"α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ"]
by blast
hence "t ⋅⇩α a0' ∈ timpl_closure_set (α⇩v⇩a⇩l⇩s 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ)"
when t: "t ∈ vals 𝒜" for t
using vals_vals_sym[OF t] by blast
hence B: "α⇩v⇩a⇩l⇩s (𝒜@T') ℐ ⊆
timpl_closure_set (α⇩v⇩a⇩l⇩s 𝒜 ℐ) (α⇩t⇩i 𝒜 T (ξ ∘⇩s σ ∘⇩s α) ℐ) ∪
((ξ ∘⇩s σ ∘⇩s α) ` fv_transaction T ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a0'"
using 0(2) 3
by (simp add: abs_apply_terms_def image_subset_iff)
have 4: "fv (t ⋅ ξ ∘⇩s σ ∘⇩s α ⋅ ℐ ⋅⇩α a) = {}" for t a
using ℐ_grounds[of "t ⋅ ξ ∘⇩s σ ∘⇩s α"] abs_fv[of "t ⋅ ξ ∘⇩s σ ∘⇩s α ⋅ ℐ" a]
by argo
have "is_Fun (t ⋅ ξ ∘⇩s σ ∘⇩s α ⋅ ℐ ⋅⇩α a0')" for t
using 4[of t a0'] by force
thus ?A
using A step_props(1,3)
unfolding T'_def a0_def a0'_def abs_apply_terms_def
by blast
show ?B
using B step_props(2,4) admissible_transaction_Value_vars[OF bspec[OF P T]]
by (auto simp add: T'_def a0_def a0'_def abs_apply_terms_def)
qed
lemma reachable_constraints_covered:
assumes 𝒜_reach: "𝒜 ∈ reachable_constraints P"
and ℐ: "welltyped_constraint_model ℐ 𝒜"
and FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
"ground (set FP)"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and P:
"∀T ∈ set P. admissible_transaction' T"
and P_occ:
"∀T ∈ set P. admissible_transaction_occurs_checks T"
and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) P"
shows "∀t ∈ α⇩i⇩k 𝒜 ℐ. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
and "α⇩v⇩a⇩l⇩s 𝒜 ℐ ⊆ absc ` set OCC"
using 𝒜_reach ℐ
proof (induction rule: reachable_constraints.induct)
case init
{ case 1 show ?case by (simp add: abs_intruder_knowledge_def) }
{ case 2 show ?case by (simp add: abs_value_constants_def) }
next
case (step 𝒜 T ξ σ α)
{ case 1
hence "welltyped_constraint_model ℐ 𝒜"
by (metis welltyped_constraint_model_prefix)
hence IH: "∀t ∈ α⇩i⇩k 𝒜 ℐ. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
"α⇩v⇩a⇩l⇩s 𝒜 ℐ ⊆ absc ` set OCC"
using step.IH by metis+
show ?case
using reachable_constraints_covered_step[
OF step.hyps(1,2) "1.prems" step.hyps(3-5) FP(1,2) IH(1)
FP(3) OCC IH(2) TI P P_occ transactions_covered]
by metis
}
{ case 2
hence "welltyped_constraint_model ℐ 𝒜"
by (metis welltyped_constraint_model_prefix)
hence IH: "∀t ∈ α⇩i⇩k 𝒜 ℐ. timpl_closure_set (set FP) (set TI) ⊢⇩c t"
"α⇩v⇩a⇩l⇩s 𝒜 ℐ ⊆ absc ` set OCC"
using step.IH by metis+
show ?case
using reachable_constraints_covered_step[
OF step.hyps(1,2) "2.prems" step.hyps(3-5) FP(1,2) IH(1)
FP(3) OCC IH(2) TI P P_occ transactions_covered]
by metis
}
qed
lemma attack_in_fixpoint_if_attack_in_ik:
fixes FP::"('fun,'atom,'sets,'lbl) prot_terms"
assumes "∀t ∈ IK ⋅⇩α⇩s⇩e⇩t a. FP ⊢⇩c t"
and "attack⟨n⟩ ∈ IK"
shows "attack⟨n⟩ ∈ FP"
proof -
have "attack⟨n⟩ ⋅⇩α a ∈ IK ⋅⇩α⇩s⇩e⇩t a" by (rule abs_in[OF assms(2)])
hence "FP ⊢⇩c attack⟨n⟩ ⋅⇩α a" using assms(1) by blast
moreover have "attack⟨n⟩ ⋅⇩α a = attack⟨n⟩" by simp
ultimately have "FP ⊢⇩c attack⟨n⟩" by metis
thus ?thesis using ideduct_synth_priv_const_in_ik[of FP "Attack n"] by simp
qed
lemma attack_in_fixpoint_if_attack_in_timpl_closure_set:
fixes FP::"('fun,'atom,'sets,'lbl) prot_terms"
assumes "attack⟨n⟩ ∈ timpl_closure_set FP TI"
shows "attack⟨n⟩ ∈ FP"
proof -
have "∀f ∈ funs_term (attack⟨n⟩). ¬is_Abs f" by auto
thus ?thesis using timpl_closure_set_no_Abs_in_set[OF assms] by blast
qed
theorem prot_secure_if_fixpoint_covered_typed:
assumes FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
"ground (set FP)"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and P:
"∀T ∈ set P. admissible_transaction T"
"has_initial_value_producing_transaction P"
and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs P)"
and attack_notin_FP: "attack⟨n⟩ ∉ set FP"
and 𝒜: "𝒜 ∈ reachable_constraints P"
shows "∄ℐ. welltyped_constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])" (is "∄ℐ. ?Q ℐ")
proof
assume "∃ℐ. ?Q ℐ"
then obtain ℐ ℬ where
ℬ: "ℬ ∈ reachable_constraints (map add_occurs_msgs P)"
and ℐ: "welltyped_constraint_model ℐ (ℬ@[(l, send⟨[attack⟨n⟩]⟩)])"
using add_occurs_msgs_soundness[OF P 𝒜]
unfolding list_all_iff by blast
have P':
"∀T ∈ set (map add_occurs_msgs P). admissible_transaction' T"
"∀T ∈ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T"
using P add_occurs_msgs_admissible_occurs_checks[OF admissible_transactionE'(1)] by auto
have 0: "attack⟨n⟩ ∉ ik⇩l⇩s⇩s⇩t ℬ ⋅⇩s⇩e⇩t ℐ"
using welltyped_constraint_model_prefix[OF ℐ]
reachable_constraints_covered(1)[OF ℬ _ FP OCC TI P' transactions_covered]
attack_in_fixpoint_if_attack_in_ik[
of "ik⇩l⇩s⇩s⇩t ℬ ⋅⇩s⇩e⇩t ℐ" "α⇩0 (db⇩l⇩s⇩s⇩t ℬ ℐ)" "timpl_closure_set (set FP) (set TI)" n]
attack_in_fixpoint_if_attack_in_timpl_closure_set
attack_notin_FP
unfolding abs_intruder_knowledge_def by blast
have 1: "ik⇩l⇩s⇩s⇩t ℬ ⋅⇩s⇩e⇩t ℐ ⊢ attack⟨n⟩"
using ℐ strand_sem_append_stateful[of "{}" "{}" "unlabel ℬ" _ ℐ]
unfolding welltyped_constraint_model_def constraint_model_def by force
show False
using 0 private_const_deduct[OF _ 1]
reachable_constraints_receive_attack_if_attack'(1)[
OF ℬ P'(1) welltyped_constraint_model_prefix[OF ℐ] 1]
by simp
qed
end
subsection ‹Theorem: A Protocol is Secure if it is Covered by a Fixed-Point›
context stateful_protocol_model
begin
theorem prot_secure_if_fixpoint_covered:
fixes P
assumes FP:
"analyzed (timpl_closure_set (set FP) (set TI))"
"wf⇩t⇩r⇩m⇩s (set FP)"
"ground (set FP)"
and OCC:
"∀t ∈ timpl_closure_set (set FP) (set TI). ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
"timpl_closure_set (absc ` set OCC) (set TI) ⊆ absc ` set OCC"
and TI:
"set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and M:
"has_all_wt_instances_of Γ (⋃T ∈ set P. trms_transaction T) N"
"finite N"
"tfr⇩s⇩e⇩t N"
"wf⇩t⇩r⇩m⇩s N"
and P:
"∀T ∈ set P. admissible_transaction T"
"∀T ∈ set P. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))"
"has_initial_value_producing_transaction P"
and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs P)"
and attack_notin_FP: "attack⟨n⟩ ∉ set FP"
and A: "𝒜 ∈ reachable_constraints P"
shows "∄ℐ. constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])"
(is "∄ℐ. constraint_model ℐ ?A")
proof
assume "∃ℐ. constraint_model ℐ ?A"
then obtain ℐ where "constraint_model ℐ ?A" by force
then obtain ℐ⇩τ where I: "welltyped_constraint_model ℐ⇩τ ?A"
using reachable_constraints_typing_result[OF M P(1,2) A] by blast
note a = FP OCC TI P(1,3) transactions_covered attack_notin_FP A
show False
using prot_secure_if_fixpoint_covered_typed[OF a] I
by force
qed
end
subsection ‹Alternative Protocol-Coverage Check›
context stateful_protocol_model
begin
context
begin
private lemma transaction_check_variant_soundness_aux0:
assumes S: "S ≡ unlabel (transaction_strand T)"
and xs: "xs ≡ filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) (fv_list⇩s⇩s⇩t S)"
and x: "fst x = Var Value" "x ∈ fv_transaction T" "x ∉ set (transaction_fresh T)"
shows "x ∈ set xs"
using x fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"]
unfolding xs S by auto
private lemma transaction_check_variant_soundness_aux1:
fixes T FP S C xs OCC negs poss as
assumes C: "C ≡ unlabel (transaction_checks T)"
and S: "S ≡ unlabel (transaction_strand T)"
and xs: "xs ≡ filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) (fv_list⇩s⇩s⇩t S)"
and poss: "poss ≡ transaction_poschecks_comp C"
and negs: "negs ≡ transaction_negchecks_comp C"
and as: "as ≡ map (λx. (x, set (filter (λab. poss x ⊆ ab ∧ negs x ∩ ab = {}) OCC))) xs"
and f: "f ≡ λx. case List.find (λp. fst p = x) as of Some p ⇒ snd p | None ⇒ {}"
and x: "x ∈ set xs"
shows "f x = set (filter (λab. poss x ⊆ ab ∧ negs x ∩ ab = {}) OCC)"
proof -
define g where "g ≡ λx. set (filter (λab. poss x ⊆ ab ∧ negs x ∩ ab = {}) OCC)"
define gs where "gs ≡ map (λx. (x, g x)) xs"
have 1: "(x, g x) ∈ set gs" using x unfolding gs_def by simp
have 2: "distinct xs" unfolding xs fv_list⇩s⇩s⇩t_def by auto
have "∃i < length xs. xs ! i = x ∧ (∀j < i. xs ! j ≠ x)" when x: "x ∈ set xs" for x
proof (rule ex1E[OF distinct_Ex1[OF 2 x]])
fix i assume i: "i < length xs ∧ xs ! i = x" and "∀j. j < length xs ∧ xs ! j = x ⟶ j = i"
hence "∀j < length xs. xs ! j = x ⟶ j = i" by blast
hence "∀j < i. xs ! j = x ⟶ j = i" using i by auto
hence "∀j < i. xs ! j ≠ x" by blast
thus ?thesis using i by blast
qed
hence "∃i < length gs. gs ! i = (x, g x) ∧ (∀j < i. gs ! j ≠ (x, g x))"
using 1 unfolding gs_def by fastforce
hence "∃i < length gs. fst (gs ! i) = x ∧ (x, g x) = gs ! i ∧ (∀j < i. fst (gs ! j) ≠ x)"
using nth_map[of _ xs "λx. (x, g x)"] length_map[of "λx. (x, g x)" xs]
unfolding gs_def by (metis (no_types, lifting) fstI min.strict_order_iff min_less_iff_conj)
hence "List.find (λp. fst p = x) gs = Some (x, g x)"
using find_Some_iff[of "λp. fst p = x" gs "(x, g x)"] by blast
thus ?thesis
unfolding f as gs_def g_def by force
qed
private lemma transaction_check_variant_soundness_aux2:
fixes T FP S C xs OCC negs poss as
assumes C: "C ≡ unlabel (transaction_checks T)"
and S: "S ≡ unlabel (transaction_strand T)"
and xs: "xs ≡ filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) (fv_list⇩s⇩s⇩t S)"
and poss: "poss ≡ transaction_poschecks_comp C"
and negs: "negs ≡ transaction_negchecks_comp C"
and as: "as ≡ map (λx. (x, set (filter (λab. poss x ⊆ ab ∧ negs x ∩ ab = {}) OCC))) xs"
and f: "f ≡ λx. case List.find (λp. fst p = x) as of Some p ⇒ snd p | None ⇒ {}"
and x: "x ∉ set xs"
shows "f x = {}"
proof -
define g where "g ≡ λx. set (filter (λab. poss x ⊆ ab ∧ negs x ∩ ab = {}) OCC)"
define gs where "gs ≡ map (λx. (x, g x)) xs"
have "(x, y) ∉ set gs" for y using x unfolding gs_def by force
thus ?thesis
using find_None_iff[of "λp. fst p = x" gs]
unfolding f as gs_def g_def by fastforce
qed
private lemma synth_abs_substs_constrs_rel_if_synth_abs_substs_constrs:
fixes T OCC negs poss
defines "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
and "ts ≡ trms_list⇩s⇩s⇩t (unlabel (transaction_receive T))"
assumes ts_wf: "∀t ∈ set ts. wf⇩t⇩r⇩m t"
and FP_ground: "ground (set FP)"
and FP_wf: "wf⇩t⇩r⇩m⇩s (set FP)"
shows "synth_abs_substs_constrs_rel FP OCC TI ts (synth_abs_substs_constrs (FP,OCC,TI) T)"
proof -
let ?R = "synth_abs_substs_constrs_rel FP OCC TI"
let ?D = "synth_abs_substs_constrs_aux FP OCC TI"
have *: "?R [t] (?D t)" when "wf⇩t⇩r⇩m t" for t using that
proof (induction t)
case (Var x) thus ?case
using synth_abs_substs_constrs_rel.SolveValueVar[of "?D (Var x)" OCC x TI FP]
by fastforce
next
case (Fun f ss)
let ?xs = "fv (Fun f ss)"
let ?lst = "map (match_abss OCC TI (Fun f ss)) FP"
define flt where
"flt = (λδ::(('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set set) option. δ ≠ None)"
define Δ where "Δ = map the (filter flt (map (match_abss OCC TI (Fun f ss)) FP))"
define θ1 where "θ1 = fun_point_Union (set Δ)"
define θ2 where "θ2 = fun_point_Inter (?D ` set ss)"
have f: "arity f = length ss" using wf_trm_arity[OF Fun.prems] by simp
have IH: "?R [s] (?D s)" when s: "s ∈ set ss" for s
using Fun.IH[OF s wf_trm_subterm[OF Fun.prems Fun_param_is_subterm[OF s]]] s
by force
have Δ3: "∀δ. δ ∈ set Δ ⟷ (∃s ∈ set FP. match_abss OCC TI (Fun f ss) s = Some δ)"
(is "∀δ. δ ∈ set Δ ⟷ ?P δ")
proof (intro allI iffI)
fix δ assume "δ ∈ set Δ"
then obtain u where "u ∈ set FP" "match_abss OCC TI (Fun f ss) u = Some δ"
unfolding Δ_def flt_def by fastforce
thus "?P δ" by blast
next
fix δ assume "?P δ"
then obtain u where u: "u ∈ set FP" "match_abss OCC TI (Fun f ss) u = Some δ" by blast
have "Some δ ∈ set ?lst" using u unfolding flt_def by force
hence "Some δ ∈ set (filter flt ?lst)" unfolding flt_def by force
moreover have "∃θ. d = Some θ" when d: "d ∈ set (filter flt ?lst)" for d
using d unfolding flt_def by simp
ultimately have "δ ∈ set (map the (filter flt ?lst))" by force
thus "δ ∈ set Δ" unfolding Δ_def by blast
qed
show ?case
proof (cases "ss = []")
case True
note ss = this
show ?thesis
proof (cases "¬public f ∧ Fun f ss ∉ set FP")
case True thus ?thesis
using synth_abs_substs_constrs_rel.SolvePrivConstNotin[of f FP OCC TI]
unfolding ss by force
next
case False thus ?thesis
using f synth_abs_substs_constrs_rel.SolvePubConst[of f FP OCC TI]
synth_abs_substs_constrs_rel.SolvePrivConstIn[of f FP OCC TI]
unfolding ss by auto
qed
next
case False
note ss = this
hence f': "arity f > 0" using f by auto
have IH': "?R ss (fun_point_Inter (?D ` set ss))"
using IH synth_abs_substs_constrs_rel.SolveCons[OF ss, of FP OCC TI ?D] by blast
have "?D (Fun f ss) = (
fun_point_union (fun_point_Union_list Δ) (fun_point_Inter_list (map ?D ss)))"
using synth_abs_substs_constrs_aux_fun_case[OF ss, of FP OCC TI f]
unfolding Let_def Δ_def flt_def by argo
hence "?D (Fun f ss) = fun_point_union θ1 θ2"
using fun_point_Inter_set_eq[of "map ?D ss"] fun_point_Union_set_eq[of Δ]
unfolding θ1_def θ2_def by simp
thus ?thesis
using synth_abs_substs_constrs_rel.SolvePubComposed[
OF f' no_private_funs[OF f'] f[symmetric] Δ3 θ1_def IH']
unfolding θ2_def by argo
qed
qed
note l0 = synth_abs_substs_constrs_rel.SolveNil[of FP OCC TI]
note d0 = synth_abs_substs_constrs_def ts_def
note l1 = * ts_wf synth_abs_substs_constrs_rel.SolveCons[of ts FP OCC TI ?D]
note d1 = d0 Let_def fun_point_Inter_set_eq[symmetric] fun_point_Inter_def
show ?thesis
proof (cases "ts = []")
case True thus ?thesis using l0 unfolding d0 by simp
next
case False thus ?thesis using l1 unfolding d1 by auto
qed
qed
private function (sequential) match_abss'_timpls_transform
::"('c set × 'c set) list ⇒
('a,'b,'c,'d) prot_subst ⇒
('a,'b,'c,'d) prot_term ⇒
('a,'b,'c,'d) prot_term ⇒
(('a,'b,'c,'d) prot_var ⇒ 'c set set) option"
where
"match_abss'_timpls_transform TI δ (Var x) (Fun (Abs a) _) = (
if ∃b ts. δ x = Fun (Abs b) ts ∧ (a = b ∨ (a,b) ∈ set TI)
then Some ((λ_. {})(x := {a}))
else None)"
| "match_abss'_timpls_transform TI δ (Fun f ts) (Fun g ss) = (
if f = g ∧ length ts = length ss
then map_option fun_point_Union_list (those (map2 (match_abss'_timpls_transform TI δ) ts ss))
else None)"
| "match_abss'_timpls_transform _ _ _ _ = None"
by pat_completeness auto
termination
proof -
let ?m = "measures [size ∘ fst ∘ snd ∘ snd]"
have 0: "wf ?m" by simp
show ?thesis
apply (standard, use 0 in fast)
by (metis (no_types) comp_def fst_conv snd_conv measures_less Fun_zip_size_lt(1))
qed
private lemma match_abss'_timpls_transform_Var_inv:
assumes "match_abss'_timpls_transform TI δ (Var x) (Fun (Abs a) ts) = Some σ"
shows "∃b ts. δ x = Fun (Abs b) ts ∧ (a = b ∨ (a, b) ∈ set TI)"
and "σ = ((λ_. {})(x := {a}))"
using assms match_abss'_timpls_transform.simps(1)[of TI δ x a ts]
by (metis option.distinct(1), metis option.distinct(1) option.inject)
private lemma match_abss'_timpls_transform_Fun_inv:
assumes "match_abss'_timpls_transform TI δ (Fun f ts) (Fun g ss) = Some σ"
shows "f = g" (is ?A)
and "length ts = length ss" (is ?B)
and "∃θ. Some θ = those (map2 (match_abss'_timpls_transform TI δ) ts ss) ∧ σ = fun_point_Union_list θ" (is ?C)
and "∀(t,s) ∈ set (zip ts ss). ∃σ'. match_abss'_timpls_transform TI δ t s = Some σ'" (is ?D)
proof -
note 0 = assms match_abss'_timpls_transform.simps(2)[of TI δ f ts g ss] option.distinct(1)
show ?A by (metis 0)
show ?B by (metis 0)
show ?C by (metis (no_types, opaque_lifting) 0 map_option_eq_Some)
thus ?D using map2_those_Some_case[of "match_abss'_timpls_transform TI δ" ts ss] by fastforce
qed
private lemma match_abss'_timpl_transform_nonempty_is_fv:
assumes "match_abss'_timpls_transform TI δ s t = Some σ"
and "σ x ≠ {}"
shows "x ∈ fv s"
using assms
proof (induction s t arbitrary: TI δ σ rule: match_abss'_timpls_transform.induct)
case (1 TI δ y a ts) show ?case
using match_abss'_timpls_transform_Var_inv[OF "1.prems"(1)] "1.prems"(2)
by fastforce
next
case (2 TI δ f ts g ss)
note prems = "2.prems"
note IH = "2.IH"
obtain θ where θ:
"Some θ = those (map2 (match_abss'_timpls_transform TI δ) ts ss)"
"σ = fun_point_Union_list θ"
and fg: "f = g" "length ts = length ss"
using match_abss'_timpls_transform_Fun_inv[OF prems(1)] by fast
have "∃σ ∈ set θ. σ x ≠ {}"
using fg(2) prems θ unfolding fun_point_Union_list_def by auto
then obtain t' s' σ where ts':
"(t',s') ∈ set (zip ts ss)" "match_abss'_timpls_transform TI δ t' s' = Some σ" "σ x ≠ {}"
using those_map2_SomeD[OF θ(1)[symmetric]] by blast
show ?case
using ts'(3) IH[OF conjI[OF fg] ts'(1) _ ts'(2)] set_zip_leftD[OF ts'(1)] by force
qed auto
private lemma match_abss'_timpls_transformI:
fixes s t::"('a,'b,'c,'d) prot_term"
and δ::"('a,'b,'c,'d) prot_subst"
and σ::"('a,'b,'c,'d) prot_var ⇒ 'c set set"
assumes TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and δ: "timpls_transformable_to TI t (s ⋅ δ)"
and σ: "match_abss' s t = Some σ"
and t: "fv t = {}"
and s: "∀f ∈ funs_term s. ¬is_Abs f"
"∀x ∈ fv s. ∃a. δ x = ⟨a⟩⇩a⇩b⇩s"
shows "match_abss'_timpls_transform TI δ s t = Some σ"
using δ σ t s
proof (induction t arbitrary: s δ σ)
case (Fun f ts)
note prems = Fun.prems
note IH = Fun.IH
show ?case
proof (cases s)
case (Var x)
obtain a b where a: "f = Abs a" "σ = (λ_. {})(x := {a})" and b: "δ x = ⟨b⟩⇩a⇩b⇩s"
using match_abss'_Var_inv[OF prems(2)[unfolded Var]] prems(5)[unfolded Var]
by auto
thus ?thesis
using prems(1) timpls_transformable_to_inv(3)[of TI f ts "Abs b" "[]"]
unfolding Var by auto
next
case (Fun g ss)
note 0 = timpls_transformable_to_inv[OF prems(1)[unfolded Fun eval_term.simps(2)]]
note 1 = match_abss'_Fun_inv[OF prems(2)[unfolded Fun]]
obtain θ where θ: "those (map2 match_abss' ss ts) = Some θ" "σ = fun_point_Union_list θ"
using 1(3) by force
have "timpls_transformable_to TI t' (s' ⋅ δ)" "∃σ'. match_abss' s' t' = Some σ'"
when "(t',s') ∈ set (zip ts ss)" for s' t'
by (metis 0(2) nth_map[of _ ss] zip_arg_index[OF that],
use that 1(4) in_set_zip_swap[of t' s' ts ss] in fast)
hence IH': "match_abss'_timpls_transform TI δ s' t' = Some σ'"
when "(t',s') ∈ set (zip ts ss)" "match_abss' s' t' = Some σ'" for s' t' σ'
using that IH[of t' s' δ σ'] prems(3-) unfolding Fun
by (metis (no_types, lifting) set_zip_leftD set_zip_rightD subsetI subset_empty term.set_intros(2) term.set_intros(4))
have "those (map2 (match_abss'_timpls_transform TI δ) ss ts) = Some θ"
using IH' θ(1) 1(4) in_set_zip_swap[of _ _ ss ts]
those_Some_iff[of "map2 match_abss' ss ts" θ]
those_Some_iff[of "map2 (match_abss'_timpls_transform TI δ) ss ts" θ]
by auto
thus ?thesis using θ(2) 1(1,2) Fun by simp
qed
qed simp
lemma timpls_transformable_to_match_abss'_nonempty_disj':
fixes s t::"('a,'b,'c,'d) prot_term"
and δ::"('a,'b,'c,'d) prot_subst"
and σ::"('a,'b,'c,'d) prot_var ⇒ 'c set set"
assumes TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and δ: "timpls_transformable_to TI t (s ⋅ δ)"
and σ: "match_abss' s t = Some σ"
and x: "x ∈ fv s"
and t: "fv t = {}"
and s: "∀f ∈ funs_term s. ¬is_Abs f"
"∀x ∈ fv s. ∃a. δ x = ⟨a⟩⇩a⇩b⇩s"
and a: "δ x = ⟨a⟩⇩a⇩b⇩s"
shows "∀b ∈ σ x. (b,a) ∈ (set TI)⇧*" (is "?P σ x")
proof -
note 0 = match_abss'_subst_disj_nonempty[OF TI]
have 1: "s ⋅ δ ∈ timpl_closure t (set TI)"
using timpls_transformable_to_iff_in_timpl_closure[OF TI] δ by blast
have 2: "match_abss'_timpls_transform TI δ s t = Some σ"
using match_abss'_timpls_transformI[OF TI δ σ t s] by simp
show "?P σ x" using 2 TI x t s a
proof (induction TI δ s t arbitrary: σ rule: match_abss'_timpls_transform.induct)
case (1 TI δ y c ts) thus ?case
using match_abss'_timpls_transform_Var_inv[OF "1.prems"(1)] by auto
next
case (2 TI δ f ts g ss)
obtain θ where fg: "f = g" "length ts = length ss"
and θ: "Some θ = those (map2 (match_abss'_timpls_transform TI δ) ts ss)"
"σ = fun_point_Union_list θ"
"∀(t, s)∈set (zip ts ss). ∃σ'. match_abss'_timpls_transform TI δ t s = Some σ'"
using match_abss'_timpls_transform_Fun_inv[OF "2.prems"(1)] by blast
have "(b,a) ∈ (set TI)⇧*" when θ': "θ' ∈ set θ" "b ∈ θ' x" for θ' b
proof -
obtain t' s' where t':
"(t',s') ∈ set (zip ts ss)" "match_abss'_timpls_transform TI δ t' s' = Some θ'"
using those_map2_SomeD[OF θ(1)[symmetric] θ'(1)] by blast
have *: "fv s' = {}" "∀f ∈ funs_term t'. ¬is_Abs f" "∀x ∈ fv t'. ∃a. δ x = ⟨a⟩⇩a⇩b⇩s"
using "2.prems"(4-6) set_zip_leftD[OF t'(1)] set_zip_rightD[OF t'(1)]
by (fastforce, fastforce, fastforce)
have **: "x ∈ fv t'"
using θ'(2) match_abss'_timpl_transform_nonempty_is_fv[OF t'(2)] by blast
show ?thesis
using θ'(2) "2.IH"[OF conjI[OF fg] t'(1) _ t'(2) "2.prems"(2) ** * "2.prems"(7)] by blast
qed
thus ?case using θ(1) unfolding θ(2) fun_point_Union_list_def by simp
qed auto
qed
lemma timpls_transformable_to_match_abss'_nonempty_disj:
fixes s t::"('a,'b,'c,'d) prot_term"
and δ::"('a,'b,'c,'d) prot_subst"
and σ::"('a,'b,'c,'d) prot_var ⇒ 'c set set"
assumes TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and δ: "timpls_transformable_to TI t (s ⋅ δ)"
and σ: "match_abss' s t = Some σ"
and x: "x ∈ fv s"
and t: "fv t = {}"
and s: "∀f ∈ funs_term s. ¬is_Abs f"
"∀x ∈ fv s. ∃a. δ x = ⟨a⟩⇩a⇩b⇩s"
shows "⋂(ticl_abs TI ` σ x) ≠ {}"
proof -
have 0: "(a,b) ∈ (set TI)⇧*" when y: "y ∈ fv s" "a ∈ σ y" "δ y = ⟨b⟩⇩a⇩b⇩s" for a b y
using timpls_transformable_to_match_abss'_nonempty_disj'[OF TI δ σ y(1) t s y(3)] y(2) by blast
obtain b where b: "δ x = ⟨b⟩⇩a⇩b⇩s" using x s(2) by blast
have "b ∈ ticl_abs TI a" when a: "a ∈ σ x" for a
using 0[OF x a b] unfolding ticl_abs_iff[OF TI] by blast
thus ?thesis by blast
qed
lemma timpls_transformable_to_subst_subterm:
fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term"
and δ σ::"(('a,'b,'c,'d) prot_fun, 'v) subst"
assumes "timpls_transformable_to TI (t ⋅ δ) (t ⋅ σ)"
and "s ⊑ t"
shows "timpls_transformable_to TI (s ⋅ δ) (s ⋅ σ)"
using assms
proof (induction "t ⋅ δ" "t ⋅ σ" arbitrary: t rule: timpls_transformable_to.induct)
case (1 TI x y) thus ?case by (cases t) auto
next
case (2 TI f T g S)
note prems = "2.prems"
note hyps = "2.hyps"(2-)
note IH = "2.hyps"(1)
show ?case
proof (cases "s = t")
case False
then obtain h U u where t: "t = Fun h U" "u ∈ set U" "s ⊑ u"
using prems(2) by (cases t) auto
then obtain i where i: "i < length U" "U ! i = u"
by (metis in_set_conv_nth)
have "timpls_transformable_to TI (u ⋅ δ) (u ⋅ σ)"
using t i prems(1) timpls_transformable_to_inv(2)[of TI h "U ⋅⇩l⇩i⇩s⇩t δ" h "U ⋅⇩l⇩i⇩s⇩t σ" i] by simp
thus ?thesis using IH hyps t by auto
qed (use prems in auto)
qed simp_all
lemma timpls_transformable_to_subst_match_case:
assumes "timpls_transformable_to TI s (t ⋅ θ)"
and "fv s = {}"
and "∀f ∈ funs_term t. ¬is_Abs f"
and "distinct (fv_list t)"
and "∀x ∈ fv t. ∃a. θ x = ⟨a⟩⇩a⇩b⇩s"
shows "∃δ. s = t ⋅ δ"
using assms
proof (induction s "t ⋅ θ" arbitrary: t rule: timpls_transformable_to.induct)
case (2 TI f T g S)
note prems = "2.prems"
note hyps = "2.hyps"(2-)
note IH = "2.hyps"(1)
show ?case
proof (cases t)
case (Var x)
then obtain a where a: "t ⋅ θ = ⟨a⟩⇩a⇩b⇩s" using prems(5) by fastforce
show ?thesis
using hyps timpls_transformable_to_inv'[OF prems(1)[unfolded a]]
unfolding Var by force
next
case (Fun h U)
have g: "g = h" and S: "S = U ⋅⇩l⇩i⇩s⇩t θ"
using hyps unfolding Fun by simp_all
note 0 = distinct_fv_list_Fun_param[OF prems(4)[unfolded Fun]]
have 1: "∀f ∈ funs_term u. ¬is_Abs f" when u: "u ∈ set U" for u
using prems(3) u unfolding Fun by fastforce
have 2: "fv t' = {}" when t': "t' ∈ set T" for t'
using t' prems(2) by simp
have 3: "∀x ∈ fv u. ∃a. θ x = ⟨a⟩⇩a⇩b⇩s" when u: "u ∈ set U" for u
using u prems(5) unfolding Fun by simp
have "¬is_Abs f"
using prems(3) timpls_transformable_to_inv(3)[OF prems(1)[unfolded hyps[symmetric] S g]]
unfolding Fun by auto
hence f: "f = h" and T: "length T = length U"
using prems(1) timpls_transformable_to_inv(1,3)[of TI f T h "U ⋅⇩l⇩i⇩s⇩t θ"]
unfolding Fun by fastforce+
define Δ where "Δ ≡ λi. if i < length T then SOME δ. T ! i = U ! i ⋅ δ else undefined"
have "timpls_transformable_to TI (T ! i) (U ! i ⋅ θ)" when i: "i < length T" for i
using prems(1)[unfolded Fun] T i timpls_transformable_to_inv(2)[of TI f T h "U ⋅⇩l⇩i⇩s⇩t θ" i]
by auto
hence "∃δ. T ! i = U ! i ⋅ δ" when i: "i < length T" for i
using i T IH[OF _ _ _ 2 1 0 3, of "T ! i" "U ! i"]
unfolding Fun g S by simp
hence Δ: "T ! i = U ! i ⋅ Δ i" when i: "i < length T" for i
using i someI2[of "λδ. T ! i = U ! i ⋅ δ" _ "λδ. T ! i = U ! i ⋅ δ"]
unfolding Δ_def by fastforce
define δ where "δ ≡ λx. if ∃i < length T. x ∈ fv (U ! i)
then Δ (SOME i. i < length T ∧ x ∈ fv (U ! i)) x
else undefined"
have "T ! i = U ! i ⋅ δ" when i: "i < length T" for i
proof -
have "j = i"
when x: "x ∈ fv (U ! i)" and j: "j < length T" "x ∈ fv (U ! j)" for j x
using x i j T distinct_fv_list_idx_fv_disjoint[OF prems(4)[unfolded Fun], of h U]
by (metis (no_types, lifting) disjoint_iff_not_equal neqE term.dual_order.refl)
hence "δ x = Δ i x" when x: "x ∈ fv (U ! i)" for x
using x i some_equality[of "λi. i < length T ∧ x ∈ fv (U ! i)" i]
unfolding δ_def by (metis (no_types, lifting))
thus ?thesis by (metis Δ i term_subst_eq)
qed
hence "T = U ⋅⇩l⇩i⇩s⇩t δ" by (metis (no_types, lifting) T length_map nth_equalityI nth_map)
hence "Fun f T = Fun f U ⋅ δ" by simp
thus ?thesis using Fun f by fast
qed
qed simp_all
lemma timpls_transformable_to_match_abss'_case:
assumes "timpls_transformable_to TI s (t ⋅ θ)"
and "fv s = {}"
and "∀f ∈ funs_term t. ¬is_Abs f"
and "∀x ∈ fv t. ∃a. θ x = ⟨a⟩⇩a⇩b⇩s"
shows "∃δ. match_abss' t s = Some δ"
using assms
proof (induction s "t ⋅ θ" arbitrary: t rule: timpls_transformable_to.induct)
case (2 TI f T g S)
note prems = "2.prems"
note hyps = "2.hyps"(2-)
note IH = "2.hyps"(1)
show ?case
proof (cases t)
case (Var x)
then obtain a where a: "t ⋅ θ = ⟨a⟩⇩a⇩b⇩s" using prems(4) by fastforce
thus ?thesis
using timpls_transformable_to_inv'(4)[OF prems(1)[unfolded a]]
by (metis (no_types) Var is_Abs_def term.sel(2) match_abss'.simps(1))
next
case (Fun h U)
have g: "g = h" and S: "S = U ⋅⇩l⇩i⇩s⇩t θ"
using hyps unfolding Fun by simp_all
have 1: "∀f ∈ funs_term u. ¬is_Abs f" when u: "u ∈ set U" for u
using prems(3) u unfolding Fun by fastforce
have 2: "fv t' = {}" when t': "t' ∈ set T" for t'
using t' prems(2) by simp
have 3: "∀x ∈ fv u. ∃a. θ x = ⟨a⟩⇩a⇩b⇩s" when u: "u ∈ set U" for u
using u prems(4) unfolding Fun by simp
have "¬is_Abs f"
using prems(3) timpls_transformable_to_inv(3)[OF prems(1)[unfolded hyps[symmetric] S g]]
unfolding Fun by auto
hence f: "f = h" and T: "length T = length U"
using prems(1) timpls_transformable_to_inv(1,3)[of TI f T h "U ⋅⇩l⇩i⇩s⇩t θ"]
unfolding Fun by fastforce+
define Δ where "Δ ≡ λi.
if i < length T
then SOME δ. match_abss' (U ! i) (T ! i) = Some δ
else undefined"
have "timpls_transformable_to TI (T ! i) (U ! i ⋅ θ)" when i: "i < length T" for i
using prems(1)[unfolded Fun] T i timpls_transformable_to_inv(2)[of TI f T h "U ⋅⇩l⇩i⇩s⇩t θ" i]
by auto
hence "∃δ. match_abss' (U ! i) (T ! i) = Some δ" when i: "i < length T" for i
using i T IH[OF _ _ _ 2 1 3, of "T ! i" "U ! i"]
unfolding Fun g S by simp
hence "match_abss' (U ! i) (T ! i) = Some (Δ i)" when i: "i < length T" for i
using i someI2[of "λδ. match_abss' (U ! i) (T ! i) = Some δ" _
"λδ. match_abss' (U ! i) (T ! i) = Some δ"]
unfolding Δ_def by fastforce
thus ?thesis
using match_abss'_FunI[OF _ T] unfolding Fun f by auto
qed
qed simp_all
lemma timpls_transformable_to_match_abss_case:
assumes TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and "timpls_transformable_to TI s (t ⋅ θ)"
and "fv s = {}"
and "∀f ∈ funs_term t. ¬is_Abs f"
and "∀x ∈ fv t. ∃a. θ x = ⟨a⟩⇩a⇩b⇩s"
shows "∃δ. match_abss OCC TI t s = Some δ"
proof -
obtain δ where δ: "match_abss' t s = Some δ"
using timpls_transformable_to_match_abss'_case[OF assms(2-)] by blast
show ?thesis
using δ timpls_transformable_to_match_abss'_nonempty_disj[OF assms(1,2) δ _ assms(3-5)]
unfolding match_abss_def by simp
qed
lemma timpls_transformable_to_match_abss_obtain:
assumes TI: "set TI = {(a,b) ∈ (set TI)⇧+. a ≠ b}"
and s_t_timpl: "timpls_transformable_to TI s (t ⋅ θ)"
and s_ground: "fv s = {}"
and t_no_abs: "∀f ∈ funs_term t. ¬is_Abs f"
and t_θ_abs: "∀x ∈ fv t. ∃a. θ x = ⟨a⟩⇩a⇩b⇩s"
obtains δ where "match_abss OCC TI t s = Some δ"
and "∀x a. x ∈ fv t ∧ θ x = ⟨a⟩⇩a⇩b⇩s ⟶ a ∈ δ x"
and "∀x. x ∉ fv t ⟶ δ x = set OCC"
proof -
let ?f = "λδ x. if x ∈ fv t then δ x else set OCC"
let ?g = "λδ x. ⋂(ticl_abs TI ` δ x)"
obtain δ where δ: "match_abss OCC TI t s = Some δ"
using timpls_transformable_to_match_abss_case[OF TI s_t_timpl s_ground t_no_abs t_θ_abs]
by blast
obtain σ where σ:
"match_abss' t s = Some σ" "δ = ?f (?g σ)"
using match_abssD[OF δ] by blast
have "∀b ∈ σ x. a ∈ ticl_abs TI b" when x: "x ∈ fv t" and a: "θ x = ⟨a⟩⇩a⇩b⇩s" for x a
using timpls_transformable_to_match_abss'_nonempty_disj'[
OF TI s_t_timpl σ(1) x s_ground t_no_abs t_θ_abs a]
unfolding ticl_abs_iff[OF TI]
by simp
hence 0: "a ∈ δ x" when x: "x ∈ fv t" and a: "θ x = ⟨a⟩⇩a⇩b⇩s" for x a
using x a σ(2) by simp
have 1: "δ x = set OCC" when x: "x ∉ fv t" for x
using x match_abss_OCC_if_not_fv[OF δ x]
by simp
show ?thesis using δ σ 0 1 that by blast
qed
private lemma transaction_check_variant_soundness_aux3:
fixes T FP S C xs OCC negs poss as
defines "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
and "C ≡ unlabel (transaction_checks T)"
and "S ≡ unlabel (transaction_strand T)"
and "ts ≡ trms_list⇩s⇩s⇩t (unlabel (transaction_receive T))"
and "xs ≡ filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) (fv_list⇩s⇩s⇩t S)"
assumes TI0: "∀(a,b) ∈ set TI. ∀(c,d) ∈ set TI. b = c ∧ a ≠ d ⟶ (a,d) ∈ set TI"
"∀(a,b) ∈ set TI. a ≠ b"
and OCC: "∀t ∈ set FP. ∀a. Abs a ∈ funs_term t ⟶ a ∈ set OCC"
and FP_ground: "ground (set FP)"
and x: "x ∈ set xs"
and xs: "∀x. x ∈ set xs ⟶ δ x ∈ set OCC"
"∀x. x ∈ set xs ⟶ poss x ⊆ δ x"
"∀x. x ∈ set xs ⟶ δ x ∩ negs x = {}"
"∀x. x ∉ set xs ⟶ δ x = {}"
and ts: "∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). intruder_synth_mod_timpls FP TI (t ⋅ θ δ)"
"∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). ∀f ∈ funs_term t. ¬is_Abs f"
"∀x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_receive T)). fst x = TAtom Value"
and C: "∀a x s. ⟨a: Var x ∈ Fun (Set s) []⟩ ∈ set C ⟶ s ∈ δ x"
"∀x s. ⟨Var x not in Fun (Set s) []⟩ ∈ set C ⟶ s ∉ δ x"
and σ: "synth_abs_substs_constrs_rel FP OCC TI ts σ"
shows "δ x ∈ σ x"
proof -
note defs = assms(1-5)
note TI = trancl_eqI'[OF TI0]
have δx0: "δ x ∈ set OCC" "poss x ⊆ δ x" "δ x ∩ negs x = {}" using x xs by (blast,blast,blast)
have ts0: "∀t ∈ set ts. intruder_synth_mod_timpls FP TI (t ⋅ θ δ)"
using ts(1) trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t unfolding ts_def by blast
have ts1: "¬Fun (Abs n) S ⊑⇩s⇩e⇩t set ts" for n S
using ts(2) funs_term_Fun_subterm'
unfolding ts_def trms_transaction_unfold trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t[symmetric] is_Abs_def
by fastforce
have ts2: "∀x ∈ fv⇩s⇩e⇩t (set ts). fst x = TAtom Value"
using ts(3)
unfolding ts_def trms_transaction_unfold trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t[symmetric]
by fastforce
show ?thesis using σ ts0 ts1 ts2
proof (induction rule: synth_abs_substs_constrs_rel.induct)
case (SolvePrivConstNotin c)
hence "intruder_synth_mod_timpls FP TI (Fun c [])" by force
hence "list_ex (λt. timpls_transformable_to TI t (Fun c [])) FP"
using SolvePrivConstNotin.hyps(1,2) by simp
then obtain t where t: "t ∈ set FP" "timpls_transformable_to TI t (Fun c [])"
unfolding list_ex_iff by blast
have "¬is_Abs c"
using SolvePrivConstNotin.prems(2)[of _ "[]"]
by (metis in_subterms_Union is_Abs_def list.set_intros(1))
hence "t = Fun c []"
using t(2) timpls_transformable_to_inv[of TI] by (cases t) auto
thus ?case using t(1) SolvePrivConstNotin.hyps(3) by fast
next
case (SolveValueVar θ1 y)
have "list_ex (λt. timpls_transformable_to TI t ⟨δ y⟩⇩a⇩b⇩s) FP"
using SolveValueVar.prems(1-3) unfolding θ_def by simp
then obtain t where t: "t ∈ set FP" "timpls_transformable_to TI t ⟨δ y⟩⇩a⇩b⇩s"
unfolding list_ex_iff by blast
obtain a where a: "t = ⟨a⟩⇩a⇩b⇩s" "a = δ y ∨ (a, δ y) ∈ set TI"
proof -
obtain ft tst where ft: "t = Fun ft tst"
using t(2) timpls_transformable_to_inv_Var(1)[of TI _ "⟨δ y⟩⇩a⇩b⇩s"]
by (cases t) auto
have "tst = []" "is_Abs ft" "the_Abs ft = δ y ∨ (the_Abs ft, δ y) ∈ set TI"
using timpls_transformable_to_inv'(2,4,5)[OF t(2)[unfolded ft]]
by (simp, force, force)
thus thesis using that[of "the_Abs ft"] ft by force
qed
have "a ∈ set OCC"
using t(1)[unfolded a(1)] OCC by auto
thus ?case
using δx0(1) t(1)[unfolded a(1)] a(2)
unfolding SolveValueVar.hyps(1) ticl_abss_def ticl_abs_def
by force
next
case (SolvePubComposed g us Δ θ1 θ2) show ?case
proof (cases "∀t ∈ set us. intruder_synth_mod_timpls FP TI (t ⋅ θ δ)")
case True
hence "δ x ∈ θ2 x"
using SolvePubComposed.IH SolvePubComposed.prems(2,3)
distinct_fv_list_Fun_param[of g us]
by auto
thus ?thesis unfolding fun_point_union_def by simp
next
case False
hence "list_ex (λt. timpls_transformable_to TI t (Fun g us ⋅ θ δ)) FP"
using SolvePubComposed.prems(1) intruder_synth_mod_timpls.simps(2)[of FP TI g "us ⋅⇩l⇩i⇩s⇩t θ δ"]
unfolding list_all_iff by auto
then obtain t where t: "t ∈ set FP" "timpls_transformable_to TI t (Fun g us ⋅ θ δ)"
unfolding list_ex_iff by blast
have t_ground: "fv t = {}"
using t(1) FP_ground by simp
have g_no_abs: "¬is_Abs f" when f: "f ∈ funs_term (Fun g us)" for f
proof -
obtain fts where "Fun f fts ⊑ Fun g us" using funs_term_Fun_subterm[OF f] by blast
thus ?thesis using SolvePubComposed.prems(2)[of _ fts] by (cases f) auto
qed
have g_θ_abs: "∃a. θ δ y = ⟨a⟩⇩a⇩b⇩s" when y: "y ∈ fv (Fun g us)" for y
using y SolvePubComposed.prems(3) unfolding θ_def by fastforce
let ?h1 = "λδ x. if x ∈ fv (Fun g us) then δ x else set OCC"
let ?h2 = "λδ x. ⋂(ticl_abs TI ` δ x)"
obtain δ' where δ':
"match_abss OCC TI (Fun g us) t = Some δ'"
"∀x a. x ∈ fv (Fun g us) ∧ θ δ x = ⟨a⟩⇩a⇩b⇩s ⟶ a ∈ δ' x"
"∀x. x ∉ fv (Fun g us) ⟶ δ' x = set OCC"
using g_no_abs g_θ_abs
timpls_transformable_to_match_abss_obtain[OF TI t(2) t_ground, of OCC thesis]
by blast
have δ'_Δ: "δ' ∈ Δ"
using t(1) δ'(1) SolvePubComposed.hyps(4) by metis
have "δ x ∈ δ' x" when x_in_g: "x ∈ fv (Fun g us)"
proof -
have "fst x = TAtom Value" using x_in_g SolvePubComposed.prems(3) by auto
hence "θ δ x = ⟨δ x⟩⇩a⇩b⇩s" unfolding θ_def by simp
thus "δ x ∈ δ' x" using δ'(2) x_in_g by blast
qed
hence "δ x ∈ δ' x" using δ'(3) δx0(1) by blast
hence "δ x ∈ θ1 x"
using δ'_Δ δx0(1) unfolding SolvePubComposed.hyps(5) fun_point_Union_def by auto
thus ?thesis unfolding fun_point_union_def by simp
qed
qed (auto simp add: δx0 fun_point_Inter_def)
qed
private lemma transaction_check_variant_soundness_aux4:
fixes T FP S C xs OCC negs poss as
defines "θ ≡ λδ x. if fst x = TAtom Value then (absc ∘ δ) x else Var x"
and "C ≡ unlabel (transaction_checks T)"
and "S ≡ unlabel (transaction_strand T)"
and "xas ≡ (the_Abs ∘ the_Fun) ` set (filter (λt. is_Fun t ∧ is_Abs (the_Fun t)) FP)"
and "ts ≡ trms_list⇩s⇩s⇩t (unlabel (transaction_receive T))"
and "xs ≡ filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) (fv_list⇩s⇩s⇩t S)"
and "poss ≡ transaction_poschecks_comp C"
and "negs ≡ transaction_negchecks_comp C"
and "as ≡ map (λx. (x, set (filter (λab. poss x ⊆ ab ∧ negs x ∩ ab = {}) OCC))) xs"
and "f ≡ λx. case List.find (λp. fst p = x) as of Some p ⇒ snd p | None ⇒ {}"
assumes T_adm: "admissible_transaction' T"
and TI0: "∀(a,b) ∈ set TI. ∀(c,d) ∈ set TI. b = c ∧ a ≠ d ⟶ (a,d) ∈ set TI"
"∀(a,b) ∈ set TI. a ≠ b"
and OCC: "∀t ∈ set FP. ∀a. Abs a ∈ funs_term t ⟶ a ∈ set OCC"
and FP_ground: "ground (set FP)"
and FP_wf: "wf⇩t⇩r⇩m⇩s (set FP)"
and "x ∈ set xs"
and "∀x. x ∈ set xs ⟶ δ x ∈ set OCC"
and "∀x. x ∈ set xs ⟶ poss x ⊆ δ x"
and "∀x. x ∈ set xs ⟶ δ x ∩ negs x = {}"
and "∀x. x ∉ set xs ⟶ δ x = {}"
and "∀t ∈ trms⇩l⇩s⇩s⇩t (transaction_receive T). intruder_synth_mod_timpls FP TI (t ⋅ θ δ)"
and "∀a x s. ⟨a: Var x ∈ Fun (Set s) []⟩ ∈ set C ⟶ s ∈ δ x"
and "∀x s. ⟨Var x not in Fun (Set s) []⟩ ∈ set C ⟶ s ∉ δ x"
shows "δ x ∈ synth_abs_substs_constrs (FP,OCC,TI) T x"
proof -
let ?FPT = "(FP,OCC,TI)"
let ?P = "λs u. let δ = mgu s u
in δ ≠ None ⟶ (∀x ∈ fv s. is_Fun (the δ x) ⟶ is_Abs (the_Fun (the δ x)))"
define θ0 where "θ0 ≡ λx.
if fst x = TAtom Value ∧ x ∈ fv_transaction T ∧ x ∉ set (transaction_fresh T)
then {ab ∈ set OCC. poss x ⊆ ab ∧ negs x ∩ ab = {}} else {}"
define g where "g ≡ λx. set (filter (λab. poss x ⊆ ab ∧ negs x ∩ ab = {}) OCC)"
define gs where "gs ≡ map (λx. (x, g x)) xs"
note defs = assms(3-10) θ0_def
note assm = assms(11-)[unfolded defs]
have ts0: "∀t ∈ set ts. wf⇩t⇩r⇩m t"
using admissible_transaction_is_wellformed_transaction(4)[OF T_adm]
unfolding admissible_transaction_terms_def wf⇩t⇩r⇩m⇩s_code[symmetric]
ts_def trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t[symmetric]
trms_transaction_unfold
by fast
have ts1: "∀t ∈ set ts. ∀f ∈ funs_term t. ¬is_Abs f"
using protocol_transactions_no_abss[OF T_adm] funs_term_Fun_subterm
trms⇩s⇩s⇩t_unlabel_prefix_subset(1)
unfolding ts_def trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t[symmetric] is_Abs_def transaction_strand_def
by (metis (no_types, opaque_lifting) in_subterms_Union in_subterms_subset_Union subset_iff)
have ts2: "∀x ∈ fv⇩s⇩e⇩t (set ts). fst x = TAtom Value"
using admissible_transaction_Value_vars[OF T_adm]
wellformed_transaction_send_receive_fv_subset(1)[
OF admissible_transaction_is_wellformed_transaction(1)[OF T_adm]]
unfolding ts_def trms_transaction_unfold trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t[symmetric] Γ⇩v_TAtom''(2)
by fastforce
have "f x = θ0 x" for x
proof (cases "fst x = Var Value ∧ x ∈ fv_transaction T ∧ x ∉ set (transaction_fresh T)")
case True
hence "θ0 x = {ab ∈ set OCC. poss x ⊆ ab ∧ negs x ∩ ab = {}}" unfolding θ0_def by argo
thus ?thesis
using True transaction_check_variant_soundness_aux0[OF S_def xs_def, of x]
transaction_check_variant_soundness_aux1[
OF C_def S_def xs_def poss_def negs_def as_def f_def, of x]
by simp
next
case False
hence 0: "θ0 x = {}" unfolding θ0_def by argo
have "x ∉ set xs"
using False fv_list⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[of "unlabel (transaction_strand T)"]
unfolding xs_def S_def by fastforce
hence "List.find (λp. fst p = x) gs = None"
using find_None_iff[of "λp. fst p = x" gs] unfolding gs_def by simp
hence "f x = {}"
unfolding f_def as_def gs_def g_def by force
thus ?thesis using 0 by simp
qed
thus ?thesis
using synth_abs_substs_constrs_rel_if_synth_abs_substs_constrs[
OF _ FP_ground FP_wf, of T, unfolded trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t ts_def[symmetric], OF ts0]
transaction_check_variant_soundness_aux3[
OF TI0 OCC FP_ground assm(7-11),
of "synth_abs_substs_constrs ?FPT T",
unfolded trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t ts_def[symmetric],
OF assm(12)[unfolded θ_def trms_list⇩s⇩s⇩t_is_trms⇩s⇩s⇩t ts_def[symmetric]]
ts1 ts2 assm(13-)[unfolded C_def]]
unfolding defs synth_abs_substs_constrs_def Let_def by blast
qed
private lemma transaction_check_variant_soundness_aux5:
fixes FP OCC TI T S C
defines "msgcs ≡ λx a. a ∈ synth_abs_substs_constrs (FP,OCC,TI) T x"
and "S ≡ unlabel (transaction_strand T)"
and "C ≡ unlabel (transaction_checks T)"
and "xs ≡ filter (λx. x ∉ set (transaction_fresh T) ∧ fst x = TAtom Value) (fv_list⇩s⇩s⇩t S)"
and "poss ≡ transaction_poschecks_comp C"
and "negs ≡ transaction_negchecks_comp C"
assumes T_adm: "admissible_transaction' T"
and TI: "∀(a,b) ∈ set TI. ∀(c,d) ∈ set TI. b = c ∧ a ≠ d ⟶ (a,d) ∈ set TI"
"∀(a,b) ∈ set TI. a ≠ b"
and OCC: "∀t ∈ set FP. ∀a. Abs a ∈ funs_term t ⟶ a ∈ set OCC"
and FP: "ground (set FP)"
"wf⇩t⇩r⇩m⇩s (set FP)"
and δ: "δ ∈ abs_substs_fun ` set (abs_substs_set xs OCC poss negs (λ_ _. True))"
"transaction_check_pre (FP,OCC,TI) T δ"
shows "δ ∈ abs_substs_fun ` set (abs_substs_set xs OCC poss negs msgcs)"
proof -
have 0: "δ x ∈ set OCC" "poss x ⊆ δ x" "δ x ∩ negs x = {}" when x: "x ∈ set xs" for x
using abs_substs_abss_bounded[OF δ(1) x] by simp_all
have 1: "δ x = {}" when x: "x ∉ set xs" for x
by (rule abs_substs_abss_bounded'[OF δ(1) x])
have 2: "msgcs x (δ x)" when x: "x ∈ set xs" for x
using 0 1 x transaction_check_variant_soundness_aux4[OF T_adm TI OCC FP, of x δ]
transaction_check_pre_ReceiveE[OF δ(2)] transaction_check_pre_InSetE[OF δ(2)]
transaction_check_pre_NotInSetE[OF δ(2)]
unfolding msgcs_def xs_def C_def S_def negs_def poss_def by fast
show ?thesis
using abs_substs_has_abs[of xs δ OCC poss negs msgcs] 0 1 2 by blast
qed
theorem transaction_check_variant_soundness:
assumes P_adm: "∀T ∈ set P. admissible_transaction' T"
and TI: "∀(a,b) ∈ set TI. ∀(c,d) ∈ set TI. b = c ∧ a ≠ d ⟶ (a,d) ∈ set TI"
"∀(a,b) ∈ set TI. a ≠ b"
and OCC: "∀t ∈ set FP. ∀a. Abs a ∈ funs_term t ⟶ a ∈ set OCC"
and FP: "ground (set FP)"
"wf⇩t⇩r⇩m⇩s (set FP)"
and T_in: "T ∈ set P"
and T_check: "transaction_check_coverage_rcv (FP,OCC,TI) T"
shows "transaction_check (FP,OCC,TI) T"
proof -
have 0: "admissible_transaction' T"
using P_adm T_in by blast
show ?thesis
using T_check transaction_check_variant_soundness_aux5[OF 0 TI OCC FP]
unfolding transaction_check_def transaction_check'_def transaction_check_coverage_rcv_def
transaction_check_comp_def list_all_iff Let_def
by force
qed
end
end
subsection ‹Automatic Fixed-Point Computation›
context stateful_protocol_model
begin
fun reduce_fixpoint' where
"reduce_fixpoint' FP _ [] = FP"
| "reduce_fixpoint' FP TI (t#M) = (
let FP' = List.removeAll t FP
in if intruder_synth_mod_timpls FP' TI t then FP' else reduce_fixpoint' FP TI M)"
definition reduce_fixpoint where
"reduce_fixpoint FP TI ≡
let f = λFP. reduce_fixpoint' FP TI FP
in while (λM. set (f M) ≠ set M) f FP"
definition compute_fixpoint_fun' where
"compute_fixpoint_fun' P (n::nat option) enable_traces Δ S0 ≡
let P' = map add_occurs_msgs P;
sy = intruder_synth_mod_timpls;
FP' = λS. fst (fst S);
TI' = λS. snd (fst S);
OCC' = λS. remdups (
(map (λt. the_Abs (the_Fun (args t ! 1)))
(filter (λt. is_Fun t ∧ the_Fun t = OccursFact) (FP' S)))@
(map snd (TI' S)));
equal_states = λS S'. set (FP' S) = set (FP' S') ∧ set (TI' S) = set (TI' S');
trace' = λS. snd S;
close = λM f. let g = remdups ∘ f in while (λA. set (g A) ≠ set A) g M;
close' = λM f. let g = remdups ∘ f in while (λA. set (g A) ≠ set A) g M;
trancl_minus_refl = λTI.
let aux = λts p. map (λq. (fst p,snd q)) (filter ((=) (snd p) ∘ fst) ts)
in filter (λp. fst p ≠ snd p) (close' TI (λts. concat (map (aux ts) ts)@ts));
snd_Ana = λN M TI. let N' = filter (λt. ∀k ∈ set (fst (Ana t)). sy M TI k) N in
filter (λt. ¬sy M TI t)
(concat (map (λt. filter (λs. s ∈ set (snd (Ana t))) (args t)) N'));
Ana_cl = λFP TI.
close FP (λM. (M@snd_Ana M M TI));
TI_cl = λFP TI.
close FP (λM. (M@filter (λt. ¬sy M TI t)
(concat (map (λm. concat (map (λ(a,b). ⟨a --» b⟩⟨m⟩) TI)) M))));
Ana_cl' = λFP TI.
let K = λt. set (fst (Ana t));
flt = λM t. (∃k ∈ K t. ¬sy M TI k) ∧ (∃k ∈ K t. ∃f ∈ funs_term k. is_Abs f);
N = λM. comp_timpl_closure_list (filter (flt M) M) TI
in close FP (λM. M@snd_Ana (N M) M TI);
Δ' = λS. Δ (FP' S, OCC' S, TI' S);
result = λS T δ.
let not_fresh = λx. x ∉ set (transaction_fresh T);
xs = filter not_fresh (fv_list⇩s⇩s⇩t (unlabel (transaction_strand T)));
u = λδ x. absdbupd (unlabel (transaction_strand T)) x (δ x)
in (remdups (filter (λt. ¬sy (FP' S) (TI' S) t)
(concat (map (λts. the_msgs ts ⋅⇩l⇩i⇩s⇩t (absc ∘ u δ))
(filter is_Send (unlabel (transaction_send T)))))),
remdups (filter (λs. fst s ≠ snd s) (map (λx. (δ x, u δ x)) xs)));
result_tuple = λS T δ. (result S T (abs_substs_fun δ), if enable_traces then δ else []);
update_state = λS. if list_ex (λt. is_Fun t ∧ is_Attack (the_Fun t)) (FP' S) then S
else let results = map (λT. map (result_tuple S T) (Δ' S T)) P';
newtrace_flt = (λn. let x = map fst (results ! n); y = map fst x; z = map snd x
in set (concat y) - set (FP' S) ≠ {} ∨ set (concat z) - set (TI' S) ≠ {});
trace =
if enable_traces
then trace' S@[concat (map (λi. map (λa. (i,snd a)) (results ! i))
(filter newtrace_flt [0..<length results]))]
else [];
U = map fst (concat results);
V = ((remdups (concat (map fst U)@FP' S),
remdups (filter (λx. fst x ≠ snd x) (concat (map snd U)@TI' S))),
trace);
W = ((Ana_cl (TI_cl (FP' V) (TI' V)) (TI' V),
trancl_minus_refl (TI' V)),
trace' V)
in if ¬equal_states W S then W
else ((Ana_cl' (FP' W) (TI' W), TI' W), trace' W);
S = ((λh. case n of None ⇒ while (λS. ¬equal_states S (h S)) h | Some m ⇒ h ^^ m)
update_state S0)
in ((reduce_fixpoint (FP' S) (TI' S), OCC' S, TI' S), trace' S)"
definition compute_fixpoint_fun where
"compute_fixpoint_fun P ≡
let P' = (filter (λT. transaction_updates T ≠ [] ∨ transaction_send T ≠ []) (remdups P));
f = (λFPT T. let msgcs = synth_abs_substs_constrs FPT T
in transaction_check_comp (λx a. a ∈ msgcs x) FPT T)
in fst (compute_fixpoint_fun' P' None False f (([],[]),[]))"
lemmas compute_fixpoint_fun_code[code]
= compute_fixpoint_fun_def[simplified compute_fixpoint_fun'_def[of _ "None" "False" _ "(([],[]),[])"
,simplified reduce_fixpoint_def o_def Option.option.case if_False]]
definition compute_fixpoint_with_trace where
"compute_fixpoint_with_trace P ≡
compute_fixpoint_fun' P None True (transaction_check_comp (λ_ _. True)) (([],[]),[])"
definition compute_fixpoint_from_trace where
"compute_fixpoint_from_trace P trace ≡
let P' = map add_occurs_msgs P;
Δ = λFPT T.
let pre_check = transaction_check_pre FPT T;
δs = map snd (filter (λ(i,as). P' ! i = T) (concat trace))
in filter (λδ. pre_check (abs_substs_fun δ)) δs;
f = compute_fixpoint_fun' ∘ map (nth P);
g = λL FPT. fst ((f L (Some 1) False Δ ((fst FPT,snd (snd FPT)),[])))
in fold g (map (map fst) trace) ([],[],[])"
definition compute_reduced_attack_trace where
"compute_reduced_attack_trace P trace ≡
let attack_in_fixpoint = list_ex (λt. ∃f ∈ funs_term t. is_Attack f) ∘ fst;
is_attack_trace = attack_in_fixpoint ∘ compute_fixpoint_from_trace P;
trace' =
let is_attack_transaction =
list_ex is_Fun_Attack ∘ concat ∘ map the_msgs ∘
filter is_Send ∘ unlabel ∘ transaction_send;
trace' =
if trace = [] then []
else butlast trace@[filter (is_attack_transaction ∘ nth P ∘ fst) (last trace)]
in trace';
iter = λtrace_prev trace_rest elem (prev,rest).
let next =
if is_attack_trace (trace_prev@(prev@rest)#trace_rest)
then prev
else prev@[elem]
in (next, tl rest);
iter' = λtrace_part (trace_prev,trace_rest).
let updated = foldr (iter trace_prev (tl trace_rest)) trace_part ([],tl (rev trace_part))
in (trace_prev@[rev (fst updated)], tl trace_rest);
reduced_trace = fst (fold iter' trace' ([],trace'))
in concat reduced_trace"
end
subsection ‹Locales for Protocols Proven Secure through Fixed-Point Coverage›
type_synonym ('f,'a,'s,'l) fixpoint_triple =
"('f,'a,'s,'l) prot_term list × 's set list × ('s set × 's set) list"
context stateful_protocol_model
begin
definition "attack_notin_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) ≡
list_all (λt. ∀f ∈ funs_term t. ¬is_Attack f) (fst FPT)"
definition "protocol_covered_by_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) P ≡
list_all (transaction_check FPT)
(filter (λT. transaction_updates T ≠ [] ∨ transaction_send T ≠ [])
(map add_occurs_msgs P))"
definition "protocol_covered_by_fixpoint_coverage_rcv (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) P ≡
list_all (transaction_check_coverage_rcv FPT)
(filter (λT. transaction_updates T ≠ [] ∨ transaction_send T ≠ [])
(map add_occurs_msgs P))"
definition "analyzed_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) ≡
let (FP, _, TI) = FPT
in analyzed_closed_mod_timpls FP TI"
definition "wellformed_protocol_SMP_set (P::('fun,'atom,'sets,'lbl) prot) N ≡
has_all_wt_instances_of Γ (⋃T ∈ set P. trms_transaction T) (set N) ∧
comp_tfr⇩s⇩e⇩t arity Ana Γ (set N) ∧
list_all (λT. list_all (comp_tfr⇩s⇩s⇩t⇩p Γ Pair) (unlabel (transaction_strand T))) P"
definition "wellformed_protocol'' (P::('fun,'atom,'sets,'lbl) prot) N ≡
let f = λT. transaction_fresh T = [] ⟶ transaction_updates T ≠ [] ∨ transaction_send T ≠ []
in list_all (λT. list_all is_Receive (unlabel (transaction_receive T)) ∧
list_all is_Check_or_Assignment (unlabel (transaction_checks T)) ∧
list_all is_Update (unlabel (transaction_updates T)) ∧
list_all is_Send (unlabel (transaction_send T)))
P ∧
list_all admissible_transaction (filter f P) ∧
wellformed_protocol_SMP_set P N"
definition "wellformed_protocol' (P::('fun,'atom,'sets,'lbl) prot) N ≡
wellformed_protocol'' P N ∧
has_initial_value_producing_transaction P"
definition "wellformed_protocol (P::('fun,'atom,'sets,'lbl) prot) ≡
let f = λM. remdups (concat (map subterms_list M@map (fst ∘ Ana) M));
N0 = remdups (concat (map (trms_list⇩s⇩s⇩t ∘ unlabel ∘ transaction_strand) P));
N = while (λA. set (f A) ≠ set A) f N0
in wellformed_protocol' P N"
definition "wellformed_fixpoint' (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) ≡
let (FP, OCC, TI) = FPT; OCC' = set OCC
in list_all (λt. wf⇩t⇩r⇩m' arity t ∧ fv t = {}) FP ∧
list_all (λa. a ∈ OCC') (map snd TI) ∧
list_all (λt. ∀f ∈ funs_term t. is_Abs f ⟶ the_Abs f ∈ OCC') FP"
definition "wellformed_term_implication_graph (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) ≡
let (_, _, TI) = FPT
in list_all (λ(a,b). list_all (λ(c,d). b = c ∧ a ≠ d ⟶ List.member TI (a,d)) TI) TI ∧
list_all (λp. fst p ≠ snd p) TI"
definition "wellformed_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) ≡
wellformed_fixpoint' FPT ∧ wellformed_term_implication_graph FPT"
lemma wellformed_protocol_SMP_set_mono:
assumes "wellformed_protocol_SMP_set P S"
and "set P' ⊆ set P"
shows "wellformed_protocol_SMP_set P' S"
using assms
unfolding wellformed_protocol_SMP_set_def comp_tfr⇩s⇩e⇩t_def has_all_wt_instances_of_def
wf⇩t⇩r⇩m⇩s'_def list_all_iff
by fast
lemma wellformed_protocol''_mono:
assumes "wellformed_protocol'' P S"
and "set P' ⊆ set P"
shows "wellformed_protocol'' P' S"
using assms wellformed_protocol_SMP_set_mono[of P S P']
unfolding wellformed_protocol''_def list_all_iff by auto
lemma wellformed_protocol'_mono:
assumes "wellformed_protocol' P S"
and "set P' ⊆ set P"
and "has_initial_value_producing_transaction P'"
shows "wellformed_protocol' P' S"
using assms wellformed_protocol_SMP_set_mono[of P S P'] wellformed_protocol''_mono
unfolding wellformed_protocol'_def by blast
lemma protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv:
assumes P: "wellformed_protocol'' P P_SMP"
and FPT: "wellformed_fixpoint FPT"
and covered: "protocol_covered_by_fixpoint_coverage_rcv FPT P"
shows "protocol_covered_by_fixpoint FPT P"
proof -
obtain FP OCC TI where FPT': "FPT = (FP,OCC,TI)" by (metis surj_pair)
note defs = FPT' wellformed_protocol''_def wellformed_fixpoint_def wellformed_fixpoint'_def
wellformed_term_implication_graph_def Let_def
wf⇩t⇩r⇩m⇩s_code[symmetric] wf⇩t⇩r⇩m_code[symmetric]
member_def case_prod_unfold list_all_iff
let ?f = "λT. transaction_fresh T = [] ⟶ transaction_updates T ≠ [] ∨ transaction_send T ≠ []"
have TI: "∀(a,b) ∈ set TI. ∀(c,d) ∈ set TI. b = c ∧ a ≠ d ⟶ (a,d) ∈ set TI"
"∀(a,b) ∈ set TI. a ≠ b"
and OCC: "∀t ∈ set FP. ∀a. Abs a ∈ funs_term t ⟶ a ∈ set OCC"
and FP: "ground (set FP)"
"wf⇩t⇩r⇩m⇩s (set FP)"
using FPT unfolding defs by (simp, simp, fastforce, simp, simp)
have P_adm: "∀T ∈ set (filter ?f (map add_occurs_msgs P)). admissible_transaction' T"
using P add_occurs_msgs_admissible_occurs_checks(1)[OF admissible_transactionE'(1)]
unfolding defs add_occurs_msgs_updates_send_filter_iff'[of P, symmetric] by auto
show ?thesis
using covered transaction_check_variant_soundness[OF P_adm TI OCC FP]
unfolding protocol_covered_by_fixpoint_def protocol_covered_by_fixpoint_coverage_rcv_def
FPT' list_all_iff
by fastforce
qed
lemma protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv':
assumes P: "wellformed_protocol'' P P_SMP"
and P': "set P' ⊆ set P"
and FPT: "wellformed_fixpoint FPT"
and covered: "protocol_covered_by_fixpoint_coverage_rcv FPT P'"
shows "protocol_covered_by_fixpoint FPT P'"
using protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[OF _ FPT covered]
wellformed_protocol''_mono[OF P P']
by simp
lemma protocol_covered_by_fixpoint_trivial_case:
assumes "list_all (λT. transaction_updates T = [] ∧ transaction_send T = [])
(map add_occurs_msgs P)"
shows "protocol_covered_by_fixpoint FPT P"
using assms
by (simp add: list_all_iff transaction_check_trivial_case protocol_covered_by_fixpoint_def)
lemma protocol_covered_by_fixpoint_empty[simp]:
"protocol_covered_by_fixpoint FPT []"
by (simp add: protocol_covered_by_fixpoint_def)
lemma protocol_covered_by_fixpoint_Cons[simp]:
"protocol_covered_by_fixpoint FPT (T#P) ⟷
transaction_check FPT (add_occurs_msgs T) ∧ protocol_covered_by_fixpoint FPT P"
using transaction_check_trivial_case[of "add_occurs_msgs T"]
unfolding protocol_covered_by_fixpoint_def case_prod_unfold by simp
lemma protocol_covered_by_fixpoint_append[simp]:
"protocol_covered_by_fixpoint FPT (P1@P2) ⟷
protocol_covered_by_fixpoint FPT P1 ∧ protocol_covered_by_fixpoint FPT P2"
by (simp add: protocol_covered_by_fixpoint_def case_prod_unfold)
lemma protocol_covered_by_fixpoint_I1[intro]:
assumes "list_all (protocol_covered_by_fixpoint FPT) P"
shows "protocol_covered_by_fixpoint FPT (concat P)"
using assms by (auto simp add: protocol_covered_by_fixpoint_def list_all_iff)
lemma protocol_covered_by_fixpoint_I2[intro]:
assumes "protocol_covered_by_fixpoint FPT P1"
and "protocol_covered_by_fixpoint FPT P2"
shows "protocol_covered_by_fixpoint FPT (P1@P2)"
using assms by (auto simp add: protocol_covered_by_fixpoint_def)
lemma protocol_covered_by_fixpoint_I3:
assumes "∀T ∈ set P. ∀δ::('fun,'atom,'sets,'lbl) prot_var ⇒ 'sets set.
transaction_check_pre FPT (add_occurs_msgs T) δ ⟶
transaction_check_post FPT (add_occurs_msgs T) δ"
shows "protocol_covered_by_fixpoint FPT P"
using assms
unfolding protocol_covered_by_fixpoint_def transaction_check_def transaction_check'_def
transaction_check_comp_def list_all_iff Let_def case_prod_unfold
Product_Type.fst_conv Product_Type.snd_conv
by fastforce
lemmas protocol_covered_by_fixpoint_intros =
protocol_covered_by_fixpoint_I1
protocol_covered_by_fixpoint_I2
protocol_covered_by_fixpoint_I3
lemma prot_secure_if_prot_checks:
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
assumes attack_notin_fixpoint: "attack_notin_fixpoint FP_OCC_TI"
and transactions_covered: "protocol_covered_by_fixpoint FP_OCC_TI P"
and analyzed_fixpoint: "analyzed_fixpoint FP_OCC_TI"
and wellformed_protocol: "wellformed_protocol' P N"
and wellformed_fixpoint: "wellformed_fixpoint FP_OCC_TI"
shows "∀𝒜 ∈ reachable_constraints P. ∄ℐ. constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])"
(is "?secure P")
proof -
define FP where "FP ≡ let (FP,_,_) = FP_OCC_TI in FP"
define OCC where "OCC ≡ let (_,OCC,_) = FP_OCC_TI in OCC"
define TI where "TI ≡ let (_,_,TI) = FP_OCC_TI in TI"
define f where "f ≡ λT::('fun,'atom,'sets,'lbl) prot_transaction.
transaction_fresh T = [] ⟶ transaction_updates T ≠ [] ∨ transaction_send T ≠ []"
define g where "g ≡ λT::('fun,'atom,'sets,'lbl) prot_transaction.
transaction_fresh T = [] ⟶
list_ex (λa. is_Update (snd a) ∨ is_Send (snd a)) (transaction_strand T)"
note wellformed_prot_defs =
wellformed_protocol'_def wellformed_protocol''_def wellformed_protocol_SMP_set_def
have attack_notin_FP: "attack⟨n⟩ ∉ set FP"
using attack_notin_fixpoint[unfolded attack_notin_fixpoint_def]
unfolding list_all_iff FP_def by force
have 1: "∀(a,b) ∈ set TI. ∀(c,d) ∈ set TI. b = c ∧ a ≠ d ⟶ (a,d) ∈ set TI"
using wellformed_fixpoint
unfolding wellformed_fixpoint_def wf⇩t⇩r⇩m⇩s_code[symmetric] Let_def TI_def
list_all_iff member_def case_prod_unfold
wellformed_term_implication_graph_def
by auto
have 0: "wf⇩t⇩r⇩m⇩s (set FP)"
and 2: "∀(a,b) ∈ set TI. a ≠ b"
and 3: "snd ` set TI ⊆ set OCC"
and 4: "∀t ∈ set FP. ∀f ∈ funs_term t. is_Abs f ⟶ f ∈ Abs ` set OCC"
and 5: "ground (set FP)"
using wellformed_fixpoint
unfolding wellformed_fixpoint_def wf⇩t⇩r⇩m_code[symmetric] is_Abs_def the_Abs_def
list_all_iff Let_def case_prod_unfold set_map FP_def OCC_def TI_def
wellformed_fixpoint'_def wellformed_term_implication_graph_def
by (fast, fast, blast, fastforce, simp)
have 8: "finite (set N)"
and 9: "has_all_wt_instances_of Γ (⋃T ∈ set (filter g P). trms_transaction T) (set N)"
and 10: "tfr⇩s⇩e⇩t (set N)"
and 11: "∀T ∈ set (filter f P). list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))"
and 12: "∀T ∈ set (filter f P). admissible_transaction T"
using wellformed_protocol[unfolded wellformed_prot_defs]
tfr⇩s⇩e⇩t_if_comp_tfr⇩s⇩e⇩t[of "set N"]
unfolding Let_def list_all_iff wf⇩t⇩r⇩m⇩s_code[symmetric] tfr⇩s⇩s⇩t⇩p_is_comp_tfr⇩s⇩s⇩t⇩p[symmetric]
has_all_wt_instances_of_def f_def[symmetric]
by (fast, fastforce, fast, fastforce, fast)
have 13: "wf⇩t⇩r⇩m⇩s (set N)"
using wellformed_protocol[unfolded wellformed_prot_defs]
finite_SMP_representationD
unfolding wf⇩t⇩r⇩m_code[symmetric] wf⇩t⇩r⇩m⇩s'_def comp_tfr⇩s⇩e⇩t_def list_all_iff Let_def by fast
have 14: "has_initial_value_producing_transaction (filter g P)"
using wellformed_protocol has_initial_value_producing_transaction_update_send_ex_filter
unfolding wellformed_protocol'_def Let_def g_def by blast
note TI0 = trancl_eqI'[OF 1 2]
have "analyzed (timpl_closure_set (set FP) (set TI))"
using analyzed_fixpoint[unfolded analyzed_fixpoint_def]
analyzed_closed_mod_timpls_is_analyzed_timpl_closure_set[OF TI0 0]
unfolding FP_def TI_def
by force
note FP0 = this 0 5
note OCC0 = funs_term_OCC_TI_subset(1)[OF 4 3]
timpl_closure_set_supset'[OF funs_term_OCC_TI_subset(2)[OF 4 3]]
note M0 = 9 8 10 13
have "f T ⟷ g T" when T: "T ∈ set P" for T
proof -
have *: "list_all stateful_strand_step.is_Receive (unlabel (transaction_receive T))"
"list_all is_Check_or_Assignment (unlabel (transaction_checks T))"
"list_all is_Update (unlabel (transaction_updates T))"
"list_all stateful_strand_step.is_Send (unlabel (transaction_send T))"
using T wellformed_protocol
unfolding wellformed_protocol_def wellformed_prot_defs Let_def list_all_iff
by (fast, fast, fast, fast)
show ?thesis
using transaction_updates_send_ex_iff[OF *]
unfolding f_def g_def by (metis (no_types, lifting) list_ex_cong)
qed
hence 15: "∀T ∈ set (filter g P). list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))"
and 16: "∀T ∈ set (filter g P). admissible_transaction T"
using 11 12 by auto
have "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs (filter g P))"
using transactions_covered[unfolded protocol_covered_by_fixpoint_def]
transaction_check_trivial_case[of _ FP_OCC_TI]
unfolding FP_def OCC_def TI_def list_all_iff Let_def case_prod_unfold
by auto
note P0 = 16 15 14 this attack_notin_FP
show ?thesis
using prot_secure_if_fixpoint_covered[OF FP0 OCC0 TI0 M0 P0]
reachable_constraints_secure_if_filter_secure_case[unfolded g_def[symmetric]]
by fast
qed
lemma prot_secure_if_prot_checks_coverage_rcv:
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
assumes attack_notin_fixpoint: "attack_notin_fixpoint FP_OCC_TI"
and transactions_covered: "protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P"
and analyzed_fixpoint: "analyzed_fixpoint FP_OCC_TI"
and wellformed_protocol: "wellformed_protocol' P N"
and wellformed_fixpoint: "wellformed_fixpoint FP_OCC_TI"
shows "∀𝒜 ∈ reachable_constraints P. ∄ℐ. constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])"
using prot_secure_if_prot_checks[
OF attack_notin_fixpoint _
analyzed_fixpoint wellformed_protocol wellformed_fixpoint]
protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[
OF _ wellformed_fixpoint transactions_covered]
wellformed_protocol[unfolded wellformed_protocol'_def]
by blast
end
locale secure_stateful_protocol =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list"
assumes attack_notin_fixpoint: "pm.attack_notin_fixpoint FP_OCC_TI"
and transactions_covered: "pm.protocol_covered_by_fixpoint FP_OCC_TI P"
and analyzed_fixpoint: "pm.analyzed_fixpoint FP_OCC_TI"
and wellformed_protocol: "pm.wellformed_protocol' P P_SMP"
and wellformed_fixpoint: "pm.wellformed_fixpoint FP_OCC_TI"
begin
theorem protocol_secure:
"∀𝒜 ∈ pm.reachable_constraints P. ∄ℐ. pm.constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])"
by (rule pm.prot_secure_if_prot_checks[OF
attack_notin_fixpoint transactions_covered
analyzed_fixpoint wellformed_protocol wellformed_fixpoint])
corollary protocol_welltyped_secure:
"∀𝒜 ∈ pm.reachable_constraints P. ∄ℐ. pm.welltyped_constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])"
using protocol_secure unfolding pm.welltyped_constraint_model_def by fast
end
locale secure_stateful_protocol' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
assumes attack_notin_fixpoint': "pm.attack_notin_fixpoint FP_OCC_TI"
and transactions_covered': "pm.protocol_covered_by_fixpoint FP_OCC_TI P"
and analyzed_fixpoint': "pm.analyzed_fixpoint FP_OCC_TI"
and wellformed_protocol': "pm.wellformed_protocol P"
and wellformed_fixpoint': "pm.wellformed_fixpoint FP_OCC_TI"
begin
sublocale secure_stateful_protocol
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P
FP_OCC_TI
"let f = λM. remdups (concat (map subterms_list M@map (fst ∘ pm.Ana) M));
N0 = remdups (concat (map (trms_list⇩s⇩s⇩t ∘ unlabel ∘ transaction_strand) P))
in while (λA. set (f A) ≠ set A) f N0"
apply unfold_locales
using attack_notin_fixpoint' transactions_covered' analyzed_fixpoint'
wellformed_protocol'[unfolded pm.wellformed_protocol_def Let_def] wellformed_fixpoint'
unfolding Let_def by blast+
end
locale secure_stateful_protocol'' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
assumes checks: "let FPT = pm.compute_fixpoint_fun P
in pm.attack_notin_fixpoint FPT ∧ pm.protocol_covered_by_fixpoint FPT P ∧
pm.analyzed_fixpoint FPT ∧ pm.wellformed_protocol P ∧ pm.wellformed_fixpoint FPT"
begin
sublocale secure_stateful_protocol'
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P "pm.compute_fixpoint_fun P"
using checks[unfolded Let_def case_prod_unfold] by unfold_locales meson+
end
locale secure_stateful_protocol''' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list"
assumes checks': "let P' = P; FPT = FP_OCC_TI; P'_SMP = P_SMP
in pm.attack_notin_fixpoint FPT ∧
pm.protocol_covered_by_fixpoint FPT P' ∧
pm.analyzed_fixpoint FPT ∧
pm.wellformed_protocol' P' P'_SMP ∧
pm.wellformed_fixpoint FPT"
begin
sublocale secure_stateful_protocol
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P FP_OCC_TI P_SMP
using checks'[unfolded Let_def case_prod_unfold] by unfold_locales meson+
end
locale secure_stateful_protocol'''' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
assumes checks'': "let P' = P; FPT = FP_OCC_TI
in pm.attack_notin_fixpoint FPT ∧
pm.protocol_covered_by_fixpoint FPT P' ∧
pm.analyzed_fixpoint FPT ∧
pm.wellformed_protocol P' ∧
pm.wellformed_fixpoint FPT"
begin
sublocale secure_stateful_protocol'
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P FP_OCC_TI
using checks''[unfolded Let_def case_prod_unfold] by unfold_locales meson+
end
locale secure_stateful_protocol_coverage_rcv =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list"
assumes attack_notin_fixpoint_coverage_rcv: "pm.attack_notin_fixpoint FP_OCC_TI"
and transactions_covered_coverage_rcv: "pm.protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P"
and analyzed_fixpoint_coverage_rcv: "pm.analyzed_fixpoint FP_OCC_TI"
and wellformed_protocol_coverage_rcv: "pm.wellformed_protocol' P P_SMP"
and wellformed_fixpoint_coverage_rcv: "pm.wellformed_fixpoint FP_OCC_TI"
begin
sublocale secure_stateful_protocol
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P
FP_OCC_TI P_SMP
using pm.protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[
OF _ wellformed_fixpoint_coverage_rcv transactions_covered_coverage_rcv]
attack_notin_fixpoint_coverage_rcv analyzed_fixpoint_coverage_rcv
wellformed_protocol_coverage_rcv wellformed_fixpoint_coverage_rcv
wellformed_protocol_coverage_rcv[unfolded pm.wellformed_protocol'_def]
by unfold_locales meson+
end
locale secure_stateful_protocol_coverage_rcv' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
assumes attack_notin_fixpoint_coverage_rcv': "pm.attack_notin_fixpoint FP_OCC_TI"
and transactions_covered_coverage_rcv': "pm.protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P"
and analyzed_fixpoint_coverage_rcv': "pm.analyzed_fixpoint FP_OCC_TI"
and wellformed_protocol_coverage_rcv': "pm.wellformed_protocol P"
and wellformed_fixpoint_coverage_rcv': "pm.wellformed_fixpoint FP_OCC_TI"
begin
sublocale secure_stateful_protocol_coverage_rcv
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P
FP_OCC_TI
"let f = λM. remdups (concat (map subterms_list M@map (fst ∘ pm.Ana) M));
N0 = remdups (concat (map (trms_list⇩s⇩s⇩t ∘ unlabel ∘ transaction_strand) P))
in while (λA. set (f A) ≠ set A) f N0"
apply unfold_locales
using attack_notin_fixpoint_coverage_rcv' transactions_covered_coverage_rcv' analyzed_fixpoint_coverage_rcv'
wellformed_protocol_coverage_rcv'[unfolded pm.wellformed_protocol_def Let_def] wellformed_fixpoint_coverage_rcv'
unfolding Let_def by blast+
end
locale secure_stateful_protocol_coverage_rcv'' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
assumes checks_coverage_rcv: "let FPT = pm.compute_fixpoint_fun P
in pm.attack_notin_fixpoint FPT ∧ pm.protocol_covered_by_fixpoint_coverage_rcv FPT P ∧
pm.analyzed_fixpoint FPT ∧ pm.wellformed_protocol P ∧ pm.wellformed_fixpoint FPT"
begin
sublocale secure_stateful_protocol_coverage_rcv'
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P "pm.compute_fixpoint_fun P"
using checks_coverage_rcv[unfolded Let_def case_prod_unfold] by unfold_locales meson+
end
locale secure_stateful_protocol_coverage_rcv''' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list"
assumes checks_coverage_rcv': "let P' = P; FPT = FP_OCC_TI; P'_SMP = P_SMP
in pm.attack_notin_fixpoint FPT ∧
pm.protocol_covered_by_fixpoint_coverage_rcv FPT P' ∧
pm.analyzed_fixpoint FPT ∧
pm.wellformed_protocol' P' P'_SMP ∧
pm.wellformed_fixpoint FPT"
begin
sublocale secure_stateful_protocol_coverage_rcv
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P FP_OCC_TI P_SMP
using checks_coverage_rcv'[unfolded Let_def case_prod_unfold] by unfold_locales meson+
end
locale secure_stateful_protocol_coverage_rcv'''' =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"'lbl"
and label_witness2::"'lbl"
+
fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list"
and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple"
assumes checks_coverage_rcv'': "let P' = P; FPT = FP_OCC_TI
in pm.attack_notin_fixpoint FPT ∧
pm.protocol_covered_by_fixpoint_coverage_rcv FPT P' ∧
pm.analyzed_fixpoint FPT ∧
pm.wellformed_protocol P' ∧
pm.wellformed_fixpoint FPT"
begin
sublocale secure_stateful_protocol_coverage_rcv'
arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2 P FP_OCC_TI
using checks_coverage_rcv''[unfolded Let_def case_prod_unfold] by unfold_locales meson+
end
subsection ‹Automatic Protocol Composition›
context stateful_protocol_model
begin
definition welltyped_leakage_free_protocol where
"welltyped_leakage_free_protocol S P ≡
let f = λM. {t ⋅ δ | t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}};
Sec = (f (set S)) - {m. {} ⊢⇩c m}
in ∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ s.
(∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜) ∧ s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜 ℐ⇩τ ∧
welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[s]⟩⟩])"
definition wellformed_composable_protocols where
"wellformed_composable_protocols (P::('fun,'atom,'sets,'lbl) prot list) N ≡
let
Ts = concat P;
steps = remdups (concat (map transaction_strand Ts));
MP0 = ⋃T ∈ set Ts. trms_transaction T ∪ pair' Pair ` setops_transaction T
in
list_all (wf⇩t⇩r⇩m' arity) N ∧
has_all_wt_instances_of Γ MP0 (set N) ∧
comp_tfr⇩s⇩e⇩t arity Ana Γ (set N) ∧
list_all (comp_tfr⇩s⇩s⇩t⇩p Γ Pair ∘ snd) steps ∧
list_all admissible_transaction_terms Ts ∧
list_all (list_all (λx. Γ⇩v x = TAtom Value ∨ (is_Var (Γ⇩v x) ∧ is_Atom (the_Var (Γ⇩v x)))) ∘
transaction_fresh)
Ts ∧
list_all (λT. ∀x ∈ vars_transaction T. ¬TAtom AttackType ⊑ Γ⇩v x) Ts ∧
list_all (λT. ∀x ∈ vars_transaction T. ∀f ∈ funs_term (Γ⇩v x). f ≠ Pair ∧ f ≠ OccursFact)
Ts ∧
list_all (list_all (λs. is_Send (snd s) ∧ length (the_msgs (snd s)) = 1 ∧
is_Fun_Attack (hd (the_msgs (snd s))) ⟶
the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s) ∘
transaction_strand)
Ts ∧
list_all (λr. (∃f ∈ ⋃(funs_term ` (trms⇩s⇩s⇩t⇩p (snd r))). f = OccursFact ∨ f = OccursSec) ⟶
(is_Receive (snd r) ∨ is_Send (snd r)) ∧ fst r = ⋆ ∧
(∀t ∈ set (the_msgs (snd r)).
(OccursFact ∈ funs_term t ∨ OccursSec ∈ funs_term t) ⟶
is_Fun t ∧ length (args t) = 2 ∧ t = occurs (args t ! 1) ∧
is_Var (args t ! 1) ∧ (Γ (args t ! 1) = TAtom Value)))
steps"
definition wellformed_composable_protocols' where
"wellformed_composable_protocols' (P::('fun,'atom,'sets,'lbl) prot list) ≡
let
Ts = concat P
in
list_all wellformed_transaction Ts ∧
list_all (list_all
(λp. let (x,cs) = p in
is_Var (Γ⇩v x) ∧ is_Atom (the_Var (Γ⇩v x)) ∧
(∀c ∈ cs. Γ⇩v x = Γ (Fun (Fu c) []::('fun,'atom,'sets,'lbl) prot_term))) ∘
(λT. transaction_decl T ()))
Ts"
definition composable_protocols where
"composable_protocols (P::('fun,'atom,'sets,'lbl) prot list) Ms S ≡
let
steps = concat (map transaction_strand (concat P));
M_fun = (λl. case find ((=) l ∘ fst) Ms of Some M ⇒ set (snd M) | None ⇒ {})
in comp_par_comp⇩l⇩s⇩s⇩t public arity Ana Γ Pair steps M_fun (set S)"
lemma composable_protocols_par_comp_constr:
fixes S f
defines "f ≡ λM. {t ⋅ δ | t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}}"
and "Sec ≡ (f (set S)) - {m. intruder_synth {} m}"
assumes Ps_pc: "wellformed_composable_protocols Ps N"
"wellformed_composable_protocols' Ps"
"composable_protocols Ps Ms S"
shows "∀𝒜 ∈ reachable_constraints (concat Ps). ∀ℐ. constraint_model ℐ 𝒜 ⟶
(∃ℐ⇩τ. welltyped_constraint_model ℐ⇩τ 𝒜 ∧
((∀n. welltyped_constraint_model ℐ⇩τ (proj n 𝒜)) ∨
(∃𝒜' l t. prefix 𝒜' 𝒜 ∧ suffix [(l, receive⟨t⟩)] 𝒜' ∧
strand_leaks⇩l⇩s⇩s⇩t 𝒜' Sec ℐ⇩τ)))"
(is "∀𝒜 ∈ _. ∀_. _ ⟶ ?Q 𝒜 ℐ")
proof (intro allI ballI impI)
fix 𝒜 ℐ
assume 𝒜: "𝒜 ∈ reachable_constraints (concat Ps)" and ℐ: "constraint_model ℐ 𝒜"
let ?Ts = "concat Ps"
let ?steps = "concat (map transaction_strand ?Ts)"
let ?MP0 = "⋃T ∈ set ?Ts. trms_transaction T ∪ pair' Pair ` setops_transaction T"
let ?M_fun = "λl. case find ((=) l ∘ fst) Ms of Some M ⇒ set (snd M) | None ⇒ {}"
have M:
"has_all_wt_instances_of Γ ?MP0 (set N)"
"finite (set N)" "tfr⇩s⇩e⇩t (set N)" "wf⇩t⇩r⇩m⇩s (set N)"
using Ps_pc tfr⇩s⇩e⇩t_if_comp_tfr⇩s⇩e⇩t[of "set N"]
unfolding composable_protocols_def wellformed_composable_protocols_def
Let_def list_all_iff wf⇩t⇩r⇩m_code[symmetric]
by fast+
have P:
"∀T ∈ set ?Ts. wellformed_transaction T"
"∀T ∈ set ?Ts. wf⇩t⇩r⇩m⇩s' arity (trms_transaction T)"
"∀T ∈ set ?Ts. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))"
"comp_par_comp⇩l⇩s⇩s⇩t public arity Ana Γ Pair ?steps ?M_fun (set S)"
using Ps_pc tfr⇩s⇩s⇩t⇩p_is_comp_tfr⇩s⇩s⇩t⇩p
unfolding wellformed_composable_protocols_def wellformed_composable_protocols'_def
composable_protocols_def Let_def list_all_iff unlabel_def wf⇩t⇩r⇩m⇩s_code[symmetric]
admissible_transaction_terms_def
by (meson, meson, fastforce, blast)
show "?Q 𝒜 ℐ"
using reachable_constraints_par_comp_constr[OF M P 𝒜 ℐ]
unfolding Sec_def f_def by fast
qed
context
begin
private lemma reachable_constraints_no_leakage_alt_aux:
fixes P lbls L
defines "lbls ≡ λT. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))"
and "L ≡ set (remdups (concat (map lbls P)))"
assumes l: "l ∉ L"
shows "map (transaction_proj l) P = map transaction_star_proj P"
proof -
have 0: "¬list_ex (has_LabelN l) (transaction_strand T)" when "T ∈ set P" for T
using that l unfolding L_def lbls_def list_ex_iff by force
have 1: "¬list_ex (has_LabelN l) (transaction_strand T)"
when T: "T ∈ set (map (transaction_proj l) P)" for T
proof -
obtain T' where T': "T' ∈ set P" "T = transaction_proj l T'" using T by auto
show ?thesis
using T'(2) 0[OF T'(1)] proj_set_subset[of l "transaction_strand T'"]
transaction_strand_proj[of l T']
unfolding list_ex_iff by fastforce
qed
have "list_all has_LabelS (transaction_strand T)"
when "T ∈ set (map (transaction_proj l) P)" for T
using that 1[OF that] transaction_proj_idem[of l]
transaction_strand_proj[of l "transaction_proj l T"]
has_LabelS_proj_iff_not_has_LabelN[of l "transaction_strand (transaction_proj l T)"]
by (metis (no_types) ex_map_conv)
thus ?thesis
using transaction_star_proj_ident_iff transaction_proj_member
transaction_star_proj_negates_transaction_proj(1)
by (metis (mono_tags, lifting) map_eq_conv)
qed
private lemma reachable_constraints_star_no_leakage:
fixes Sec P lbls k
defines "no_leakage ≡ λ𝒜. ∄ℐ⇩τ 𝒜' s.
prefix 𝒜' 𝒜 ∧ (∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜') ∧ s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ ∧
welltyped_constraint_model ℐ⇩τ (𝒜'@[(k,send⟨[s]⟩)])"
assumes Sec: "∀s ∈ Sec. ¬{} ⊢⇩c s" "ground Sec"
shows "∀𝒜 ∈ reachable_constraints (map transaction_star_proj P). no_leakage 𝒜"
proof
fix A assume A: "A ∈ reachable_constraints (map transaction_star_proj P)"
have A': "∀(l,a) ∈ set A. l = ⋆"
using reachable_constraints_preserves_labels[OF A] transaction_star_proj_has_star_labels
unfolding list_all_iff by fastforce
show "no_leakage A"
using constr_sem_stateful_star_proj_no_leakage[OF Sec(2) A']
unlabel_append[of A] singleton_lst_proj(4)[of k]
unfolding no_leakage_def welltyped_constraint_model_def constraint_model_def by fastforce
qed
private lemma reachable_constraints_no_leakage_alt:
fixes Sec P lbls k
defines "no_leakage ≡ λ𝒜. ∄ℐ⇩τ 𝒜' s.
prefix 𝒜' 𝒜 ∧ (∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜') ∧ s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ ∧
welltyped_constraint_model ℐ⇩τ (𝒜'@[(k,send⟨[s]⟩)])"
and "lbls ≡ λT. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))"
and "L ≡ set (remdups (concat (map lbls P)))"
assumes Sec: "∀s ∈ Sec. ¬{} ⊢⇩c s" "ground Sec"
and lbl: "∀l ∈ L. ∀𝒜 ∈ reachable_constraints (map (transaction_proj l) P). no_leakage 𝒜"
shows "∀l. ∀𝒜 ∈ reachable_constraints (map (transaction_proj l) P). ∄ℐ⇩τ 𝒜'.
interpretation⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wt⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wf⇩t⇩r⇩m⇩s (subst_range ℐ⇩τ) ∧
prefix 𝒜' 𝒜 ∧ (∃l' ts. suffix [(l', receive⟨ts⟩)] 𝒜') ∧ strand_leaks⇩l⇩s⇩s⇩t 𝒜' Sec ℐ⇩τ"
proof (intro allI ballI)
fix l 𝒜
assume 𝒜: "𝒜 ∈ reachable_constraints (map (transaction_proj l) P)"
let ?Q = "λℐ⇩τ 𝒜'.
interpretation⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wt⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wf⇩t⇩r⇩m⇩s (subst_range ℐ⇩τ) ∧
prefix 𝒜' 𝒜 ∧ (∃l' t. suffix [(l', receive⟨t⟩)] 𝒜') ∧ strand_leaks⇩l⇩s⇩s⇩t 𝒜' Sec ℐ⇩τ"
show "∄ℐ⇩τ 𝒜'. ?Q ℐ⇩τ 𝒜'"
proof
assume "∃ℐ⇩τ 𝒜'. ?Q ℐ⇩τ 𝒜'"
then obtain ℐ⇩τ 𝒜' t n l' ts' where
ℐ⇩τ: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ⇩τ" "wt⇩s⇩u⇩b⇩s⇩t ℐ⇩τ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ⇩τ)" and
𝒜': "prefix 𝒜' 𝒜" "suffix [(l', receive⟨ts'⟩)] 𝒜'" and
t: "t ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ" and
n: "constr_sem_stateful ℐ⇩τ (proj_unl n 𝒜'@[send⟨[t]⟩])"
unfolding strand_leaks⇩l⇩s⇩s⇩t_def by blast
hence 0: "welltyped_constraint_model ℐ⇩τ (proj n 𝒜'@[(m,send⟨[t]⟩)])" for m
unfolding welltyped_constraint_model_def constraint_model_def by fastforce
have t_Sec: "¬{} ⊢⇩c t" "t ⋅ ℐ⇩τ = t"
using t Sec subst_ground_ident[of t ℐ⇩τ] by auto
obtain B k' s where B:
"constr_sem_stateful ℐ⇩τ (proj_unl n B@[send⟨[t]⟩])"
"prefix (proj n B) (proj n 𝒜)" "suffix [(k', receive⟨s⟩)] (proj n B)"
"t ∈ Sec - declassified⇩l⇩s⇩s⇩t (proj n B) ℐ⇩τ"
using constr_sem_stateful_proj_priv_term_prefix_obtain[OF 𝒜'(1) n t t_Sec]
by metis
hence 1: "welltyped_constraint_model ℐ⇩τ (proj n B@[(m,send⟨[t]⟩)])" for m
using 0 unfolding welltyped_constraint_model_def constraint_model_def by fastforce
note 2 = reachable_constraints_transaction_proj_proj_eq
note 3 = reachable_constraints_transaction_proj_star_proj
note 4 = reachable_constraints_no_leakage_alt_aux
note star_case = 0 t t_Sec(1) reachable_constraints_star_no_leakage[OF Sec]
𝒜'(2) 3[OF 𝒜] prefix_proj(1)[OF 𝒜'(1)]
suffix_proj(1)[OF 𝒜'(2)] declassified⇩l⇩s⇩s⇩t_proj_eq
note lbl_case = 0 t(1) 𝒜 𝒜' lbl 2(2)[OF 𝒜 𝒜'(1)]
show False
proof (cases "l = n")
case True thus ?thesis
proof (cases "l ∈ L")
case False
hence "map (transaction_proj l) P = map transaction_star_proj P"
using 4 unfolding L_def lbls_def by fast
thus ?thesis
using lbl_case(1-4,7) star_case(4,5) True by metis
qed (metis lbl_case no_leakage_def)
next
case False
hence "no_leakage (proj n 𝒜)" using star_case(4,6) unfolding no_leakage_def by fast
thus ?thesis by (metis B(2-4) 1 no_leakage_def)
qed
qed
qed
private lemma reachable_constraints_no_leakage_alt'_aux1:
fixes P::"('a,'b,'c,'d) prot_transaction list"
defines "f ≡ list_all ((list_all (Not ∘ has_LabelS)) ∘ tl ∘ transaction_send)"
assumes P: "f P"
shows "f (map (transaction_proj l) P)"
and "f (map transaction_star_proj P)"
proof -
let ?g = "λT. tl (transaction_send T)"
have "set (?g (transaction_proj l T)) ⊆ set (?g T)" (is "?A ⊆ ?C")
and "set (?g (transaction_star_proj T)) ⊆ set (?g T)" (is "?B ⊆ ?C")
for T::"('a,'b,'c,'d) prot_transaction"
proof -
obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp
have "transaction_send (transaction_proj l T) = proj l (transaction_send T)"
"transaction_send (transaction_star_proj T) = filter has_LabelS (transaction_send T)"
using transaction_proj.simps[of l T1 T2 T3 T4 T5 T6]
transaction_star_proj.simps[of T1 T2 T3 T4 T5 T6]
unfolding T proj_def Let_def by auto
hence "set (?g (transaction_proj l T)) ⊆ set (proj l (?g T))"
"set (?g (transaction_star_proj T)) ⊆ set (filter has_LabelS (?g T))"
unfolding proj_def
by (metis (no_types, lifting) filter.simps(2) list.collapse list.sel(2,3)
list.set_sel(2) subsetI)+
thus "?A ⊆ ?C" "?B ⊆ ?C" using T unfolding proj_def by auto
qed
thus "f (map (transaction_proj l) P)" "f (map transaction_star_proj P)"
using P unfolding f_def list_all_iff by fastforce+
qed
private lemma reachable_constraints_no_leakage_alt'_aux2:
fixes P
defines "f ≡ λT.
list_all is_Receive (unlabel (transaction_receive T)) ∧
list_all is_Check_or_Assignment (unlabel (transaction_checks T)) ∧
list_all is_Update (unlabel (transaction_updates T)) ∧
list_all is_Send (unlabel (transaction_send T))"
assumes P: "list_all f P"
shows "list_all f (map (transaction_proj l) P)" (is ?A)
and "list_all f (map transaction_star_proj P)" (is ?B)
proof -
have "f (transaction_proj l T)" (is ?A')
and "f (transaction_star_proj T)" (is ?B')
when T_in: "T ∈ set P" for T
proof -
obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T)
have "f T" using P T_in unfolding list_all_iff by simp
thus ?A' ?B'
unfolding f_def T unlabel_def proj_def Let_def list_all_iff
transaction_proj.simps[of l T1 T2 T3 T4 T5 T6]
transaction_star_proj.simps[of T1 T2 T3 T4 T5 T6]
by auto
qed
thus ?A ?B unfolding list_all_iff by auto
qed
private lemma reachable_constraints_no_leakage_alt':
fixes Sec P lbls k
defines "no_leakage ≡ λ𝒜. ∄ℐ⇩τ 𝒜' s.
prefix 𝒜' 𝒜 ∧ (∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜') ∧ s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ ∧
welltyped_constraint_model ℐ⇩τ (𝒜'@[(k,send⟨[s]⟩)])"
and "no_leakage' ≡ λ𝒜. ∄ℐ⇩τ s.
(∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜) ∧ s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜 ℐ⇩τ ∧
welltyped_constraint_model ℐ⇩τ (𝒜@[(k,send⟨[s]⟩)])"
assumes P: "list_all wellformed_transaction P"
"list_all ((list_all (Not ∘ has_LabelS)) ∘ tl ∘ transaction_send) P"
and Sec: "∀s ∈ Sec. ¬{} ⊢⇩c s" "ground Sec"
and lbl: "∀l ∈ L. ∀𝒜 ∈ reachable_constraints (map (transaction_proj l) P). no_leakage' 𝒜"
shows "∀l ∈ L. ∀𝒜 ∈ reachable_constraints (map (transaction_proj l) P). no_leakage 𝒜" (is ?A)
and "∀𝒜 ∈ reachable_constraints (map transaction_star_proj P). no_leakage 𝒜" (is ?B)
proof -
define f where "f ≡ λT::('fun,'atom,'sets,'lbl) prot_transaction.
list_all is_Receive (unlabel (transaction_receive T)) ∧
list_all is_Check_or_Assignment (unlabel (transaction_checks T)) ∧
list_all is_Update (unlabel (transaction_updates T)) ∧
list_all is_Send (unlabel (transaction_send T))"
define g where "(g::('fun,'atom,'sets,'lbl) prot_transaction ⇒ bool) ≡
list_all (Not ∘ has_LabelS) ∘ tl ∘ transaction_send"
have P': "list_all f P"
using P(1) unfolding wellformed_transaction_def f_def list_all_iff by fastforce
note 0 = reachable_constraints_no_leakage_alt'_aux1[OF P(2), unfolded g_def[symmetric]]
note 1 = reachable_constraints_no_leakage_alt'_aux2[
OF P'[unfolded f_def], unfolded f_def[symmetric]]
note 2 = reachable_constraints_aligned_prefix_ex[unfolded f_def[symmetric] g_def[symmetric]]
have 3: "∀𝒜 ∈ reachable_constraints (map transaction_star_proj P). no_leakage' 𝒜"
using reachable_constraints_star_no_leakage[OF Sec] unfolding no_leakage'_def by blast
show ?A
proof (intro ballI)
fix l 𝒜 assume l: "l ∈ L" and 𝒜: "𝒜 ∈ reachable_constraints (map (transaction_proj l) P)"
show "no_leakage 𝒜"
proof (rule ccontr)
assume "¬no_leakage 𝒜"
then obtain ℐ⇩τ 𝒜' s where 𝒜':
"prefix 𝒜' 𝒜" "∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜'" "s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ"
"welltyped_constraint_model ℐ⇩τ (𝒜'@[(k, send⟨[s]⟩)])"
unfolding no_leakage_def by blast
have s: "¬{} ⊢⇩c s" "fv s = {}" using 𝒜'(3) Sec by auto
have ℐ⇩τ: "constr_sem_stateful ℐ⇩τ (unlabel 𝒜'@[send⟨[s]⟩])"
"wt⇩s⇩u⇩b⇩s⇩t ℐ⇩τ" "interpretation⇩s⇩u⇩b⇩s⇩t ℐ⇩τ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ⇩τ)"
using 𝒜'(4) unfolding welltyped_constraint_model_def constraint_model_def by auto
show False
using 2[OF 1(1) 0(1) s 𝒜 𝒜'(1,2) ℐ⇩τ(1)] l lbl 𝒜'(3) ℐ⇩τ(2,3,4)
singleton_lst_proj(4)[of k "send⟨[s]⟩"] unlabel_append[of _ "[(k, send⟨[s]⟩)]"]
unfolding no_leakage'_def welltyped_constraint_model_def constraint_model_def by metis
qed
qed
show ?B
proof (intro ballI)
fix 𝒜 assume 𝒜: "𝒜 ∈ reachable_constraints (map transaction_star_proj P)"
show "no_leakage 𝒜"
proof (rule ccontr)
assume "¬no_leakage 𝒜"
then obtain ℐ⇩τ 𝒜' s where 𝒜':
"prefix 𝒜' 𝒜" "∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜'" "s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ"
"welltyped_constraint_model ℐ⇩τ (𝒜'@[(k, send⟨[s]⟩)])"
unfolding no_leakage_def by blast
have s: "¬{} ⊢⇩c s" "fv s = {}" using 𝒜'(3) Sec by auto
have ℐ⇩τ: "constr_sem_stateful ℐ⇩τ (unlabel 𝒜'@[send⟨[s]⟩])"
"wt⇩s⇩u⇩b⇩s⇩t ℐ⇩τ" "interpretation⇩s⇩u⇩b⇩s⇩t ℐ⇩τ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ⇩τ)"
using 𝒜'(4) unfolding welltyped_constraint_model_def constraint_model_def by auto
show False
using 2[OF 1(2) 0(2) s 𝒜 𝒜'(1,2) ℐ⇩τ(1)] 3 𝒜'(3) ℐ⇩τ(2,3,4)
singleton_lst_proj(4)[of k "send⟨[s]⟩"] unlabel_append[of _ "[(k, send⟨[s]⟩)]"]
unfolding no_leakage'_def welltyped_constraint_model_def constraint_model_def by metis
qed
qed
qed
lemma composable_protocols_par_comp_prot_alt:
fixes S f Sec lbls Ps
defines "f ≡ λM. {t ⋅ δ | t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}}"
and "Sec ≡ (f (set S)) - {m. {} ⊢⇩c m}"
and "lbls ≡ λT. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))"
and "L ≡ set (remdups (concat (map lbls (concat Ps))))"
and "no_leakage ≡ λ𝒜. ∄ℐ⇩τ 𝒜' s.
prefix 𝒜' 𝒜 ∧ (∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜') ∧ s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ ∧
welltyped_constraint_model ℐ⇩τ (𝒜'@[⟨⋆, send⟨[s]⟩⟩])"
assumes Ps_pc: "wellformed_composable_protocols Ps N"
"wellformed_composable_protocols' Ps"
"composable_protocols Ps Ms S"
and component_secure:
"∀𝒜 ∈ reachable_constraints (map (transaction_proj l) (concat Ps)). ∄ℐ.
welltyped_constraint_model ℐ (𝒜@[⟨l, send⟨[attack⟨ln l⟩]⟩⟩])"
and no_leakage:
"∀l ∈ L. ∀𝒜 ∈ reachable_constraints (map (transaction_proj l) (concat Ps)). no_leakage 𝒜"
shows "∀𝒜 ∈ reachable_constraints (concat Ps). ∄ℐ.
constraint_model ℐ (𝒜@[⟨l, send⟨[attack⟨ln l⟩]⟩⟩])"
proof
fix 𝒜
assume 𝒜: "𝒜 ∈ reachable_constraints (concat Ps)"
let ?att = "[⟨l, send⟨[attack⟨ln l⟩]⟩⟩]"
define Q where "Q ≡ λℐ⇩τ. interpretation⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wt⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wf⇩t⇩r⇩m⇩s (subst_range ℐ⇩τ)"
define R where "R ≡ λ𝒜 ℐ⇩τ.
∃𝒜' l t. prefix 𝒜' 𝒜 ∧ suffix [(l, receive⟨t⟩)] 𝒜' ∧ strand_leaks⇩l⇩s⇩s⇩t 𝒜' Sec ℐ⇩τ"
define M where "M ≡ ⋃T∈set (concat Ps). trms_transaction T ∪ pair' Pair ` setops_transaction T"
have Sec: "∀s ∈ Sec. ¬{} ⊢⇩c s" "ground Sec" unfolding Sec_def f_def by auto
have par_comp':
"∀𝒜 ∈ reachable_constraints (concat Ps). ∀ℐ. constraint_model ℐ 𝒜 ⟶
(∃ℐ⇩τ. welltyped_constraint_model ℐ⇩τ 𝒜 ∧
((∀n. welltyped_constraint_model ℐ⇩τ (proj n 𝒜)) ∨ R 𝒜 ℐ⇩τ))"
using 𝒜 composable_protocols_par_comp_constr[OF Ps_pc] unfolding Sec_def f_def R_def by fast
have "∀l. ∀𝒜 ∈ reachable_constraints (map (transaction_proj l) (concat Ps)). ∄ℐ⇩τ. Q ℐ⇩τ ∧ R 𝒜 ℐ⇩τ"
using reachable_constraints_no_leakage_alt[OF
Sec no_leakage[unfolded no_leakage_def L_def lbls_def]]
unfolding Q_def R_def by blast
hence no_leakage':
"∀𝒜 ∈ reachable_constraints (concat Ps). ∄ℐ⇩τ. Q ℐ⇩τ ∧ R 𝒜 ℐ⇩τ"
using reachable_constraints_component_leaks_if_composed_leaks[OF Sec, of "concat Ps"
"λℐ⇩τ. interpretation⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wt⇩s⇩u⇩b⇩s⇩t ℐ⇩τ ∧ wf⇩t⇩r⇩m⇩s (subst_range ℐ⇩τ)"]
unfolding Q_def R_def by blast
have M: "has_all_wt_instances_of Γ M (set N)" "finite (set N)" "tfr⇩s⇩e⇩t (set N)" "wf⇩t⇩r⇩m⇩s (set N)"
and P: "∀T ∈ set (concat Ps). wellformed_transaction T"
"∀T ∈ set (concat Ps). admissible_transaction_terms T"
"∀T ∈ set (concat Ps). ∀x ∈ vars_transaction T. ¬TAtom AttackType ⊑ Γ⇩v x"
"∀T ∈ set (concat Ps). ∀s ∈ set (transaction_strand T).
is_Send (snd s) ∧ length (the_msgs (snd s)) = 1 ∧
is_Fun_Attack (hd (the_msgs (snd s))) ⟶
the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s"
"∀T ∈ set (concat Ps). list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))"
using Ps_pc(1,2) tfr⇩s⇩e⇩t_if_comp_tfr⇩s⇩e⇩t tfr⇩s⇩s⇩t⇩p_is_comp_tfr⇩s⇩s⇩t⇩p
unfolding wellformed_composable_protocols_def wellformed_composable_protocols'_def
list_all_iff Let_def M_def wf⇩t⇩r⇩m⇩s'_def wf⇩t⇩r⇩m⇩s_code unlabel_def Γ⇩v_TAtom''(1,2)
by (force, force, fast, fast, fast, fast, fast, simp, simp)
have P_fresh: "∀T ∈ set (concat Ps). ∀x ∈ set (transaction_fresh T).
Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))"
(is "∀T ∈ ?P. ∀x ∈ ?frsh T. ?Q x")
proof (intro ballI)
fix T x assume T: "T ∈ ?P" "x ∈ ?frsh T"
hence "Γ⇩v x = TAtom Value ∨ (is_Var (Γ⇩v x) ∧ is_Atom (the_Var (Γ⇩v x)))"
using Ps_pc(1) unfolding wellformed_composable_protocols_def list_all_iff Let_def by fastforce
thus "?Q x" by (metis prot_atom.is_Atom_def term.collapse(1))
qed
have P': "∀T ∈ set (concat Ps). wf⇩t⇩r⇩m⇩s' arity (trms_transaction T)"
using P(2) admissible_transaction_terms_def by fast
have "¬welltyped_constraint_model ℐ (𝒜@?att)" for ℐ
proof
assume "welltyped_constraint_model ℐ (𝒜@?att)"
hence ℐ: "welltyped_constraint_model ℐ 𝒜" "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ attack⟨ln l⟩"
using strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" "unlabel ?att"]
unlabel_append[of 𝒜 ?att]
unfolding welltyped_constraint_model_def constraint_model_def by auto
obtain ℐ⇩τ where ℐ⇩τ:
"welltyped_constraint_model ℐ⇩τ 𝒜"
"welltyped_constraint_model ℐ⇩τ (proj l 𝒜)"
using 𝒜 ℐ no_leakage' par_comp'
unfolding Q_def welltyped_constraint_model_def constraint_model_def by metis
have "⟨l, receive⟨[attack⟨ln l⟩]⟩⟩ ∈ set 𝒜"
using reachable_constraints_receive_attack_if_attack(3)[OF 𝒜 P(1-2) P_fresh P(3) ℐ P(4)]
by auto
hence "ik⇩l⇩s⇩s⇩t (proj l 𝒜) ⋅⇩s⇩e⇩t ℐ⇩τ ⊢ attack⟨ln l⟩"
using in_proj_set[of l "receive⟨[attack⟨ln l⟩]⟩" 𝒜] in_ik⇩l⇩s⇩s⇩t_iff[of "attack⟨ln l⟩" "proj l 𝒜"]
intruder_deduct.Axiom[of "attack⟨ln l⟩" "ik⇩l⇩s⇩s⇩t (proj l 𝒜) ⋅⇩s⇩e⇩t ℐ⇩τ"]
by fastforce
hence "welltyped_constraint_model ℐ⇩τ (proj l 𝒜@?att)"
using ℐ⇩τ strand_sem_append_stateful[of "{}" "{}" "unlabel (proj l 𝒜)" "unlabel ?att" ℐ⇩τ]
unfolding welltyped_constraint_model_def constraint_model_def by auto
thus False
using component_secure reachable_constraints_transaction_proj[OF 𝒜, of l] by simp
qed
thus "∄ℐ. constraint_model ℐ (𝒜@?att)"
using reachable_constraints_typing_result'[OF M_def M P(1) P' P(5) 𝒜] by blast
qed
lemma composable_protocols_par_comp_prot:
fixes S f Sec lbls Ps
defines "f ≡ λM. {t ⋅ δ | t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}}"
and "Sec ≡ (f (set S)) - {m. {} ⊢⇩c m}"
and "lbls ≡ λT. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))"
and "L ≡ set (remdups (concat (map lbls (concat Ps))))"
and "no_leakage ≡ λ𝒜. ∄ℐ⇩τ s.
(∃l ts. suffix [(l, receive⟨ts⟩)] 𝒜) ∧ s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜 ℐ⇩τ ∧
welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[s]⟩⟩])"
assumes Ps_pc: "wellformed_composable_protocols Ps N"
"wellformed_composable_protocols' Ps"
"composable_protocols Ps Ms S"
"list_all ((list_all (Not ∘ has_LabelS)) ∘ tl ∘ transaction_send) (concat Ps)"
and component_secure:
"∀𝒜 ∈ reachable_constraints (map (transaction_proj l) (concat Ps)). ∄ℐ.
welltyped_constraint_model ℐ (𝒜@[⟨l, send⟨[attack⟨ln l⟩]⟩⟩])"
and no_leakage:
"∀l ∈ L. ∀𝒜 ∈ reachable_constraints (map (transaction_proj l) (concat Ps)). no_leakage 𝒜"
shows "∀𝒜 ∈ reachable_constraints (concat Ps). ∄ℐ.
constraint_model ℐ (𝒜@[⟨l, send⟨[attack⟨ln l⟩]⟩⟩])"
proof -
have P': "list_all wellformed_transaction (concat Ps)"
using Ps_pc(2) unfolding wellformed_composable_protocols'_def by meson
have Sec: "∀s ∈ Sec. ¬{} ⊢⇩c s" "ground Sec" unfolding Sec_def f_def by auto
note 0 = composable_protocols_par_comp_prot_alt[
OF Ps_pc(1-3) component_secure, unfolded lbls_def[symmetric] L_def[symmetric]]
note 1 = reachable_constraints_no_leakage_alt'[
OF P' Ps_pc(4) Sec no_leakage[unfolded no_leakage_def]]
show ?thesis using 0 1 unfolding f_def Sec_def by argo
qed
lemma composable_protocols_par_comp_prot':
assumes P_defs:
"Pc = concat Ps"
"set Ps_with_stars =
(λn. map (transaction_proj n) Pc) `
set (remdups (concat
(map (λT. map (the_LabelN ∘ fst) (filter (Not ∘ has_LabelS) (transaction_strand T)))
Pc)))"
and Ps_wellformed:
"list_all (list_all (Not ∘ has_LabelS) ∘ tl ∘ transaction_send) Pc"
"wellformed_composable_protocols Ps N"
"wellformed_composable_protocols' Ps"
"composable_protocols Ps Ms S"
and Ps_no_leakage:
"list_all (welltyped_leakage_free_protocol S) Ps_with_stars"
and P_def:
"P = map (transaction_proj n) Pc"
and P_wt_secure:
"∀𝒜 ∈ reachable_constraints P. ∄ℐ.
welltyped_constraint_model ℐ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
shows "∀𝒜 ∈ reachable_constraints Pc. ∄ℐ.
constraint_model ℐ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
by (rule composable_protocols_par_comp_prot[
OF Ps_wellformed(2,3,4,1)[unfolded P_defs(1)]
P_wt_secure[unfolded P_def[unfolded P_defs(1)]]
transaction_proj_ball_subst[
OF P_defs(2)[unfolded P_defs(1)]
Ps_no_leakage(1)[
unfolded list_all_iff welltyped_leakage_free_protocol_def Let_def]],
unfolded P_defs(1)[symmetric]])
end
context
begin
lemma welltyped_constraint_model_leakage_model_swap:
fixes I α δ::"('fun,'atom,'sets,'lbl) prot_subst" and s
assumes A: "welltyped_constraint_model I (A@[⟨⋆, send⟨[s ⋅ δ]⟩⟩])"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)"
and δ: "wt⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" "subst_domain δ = fv s" "ground (subst_range δ)"
obtains J
where "welltyped_constraint_model J (A@[⟨⋆, send⟨[s ⋅ δ]⟩⟩])"
and "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J ⊢ s ⋅ α ⋅ J"
proof
note defs = welltyped_constraint_model_def constraint_model_def
note δ_s = subst_fv_dom_ground_if_ground_img[OF equalityD2[OF δ(3)] δ(4)]
note α' = transaction_renaming_subst_is_renaming(2)[OF α]
inj_on_subset[OF transaction_renaming_subst_is_injective[OF α]
subset_UNIV[of "fv s"]]
transaction_renaming_subst_var_obtain(2)[OF α, of _ s]
transaction_renaming_subst_is_renaming(6)[OF α, of s]
transaction_renaming_subst_vars_disj(8)[OF α]
transaction_renaming_subst_wt[OF α]
define αinv where "αinv ≡ subst_var_inv α (fv s)"
define δ' where "δ' ≡ rm_vars (UNIV - fv (s ⋅ α)) (αinv ∘⇩s δ)"
define J where "J ≡ λx. if x ∈ fv (s ⋅ α) then δ' x else I x"
have α_invertible: "s = s ⋅ α ∘⇩s αinv"
using α'(1) inj_var_ran_subst_is_invertible'[of α s] inj_on_subset[OF α'(2)]
unfolding αinv_def by blast
have δ'_domain: "subst_domain δ' = fv (s ⋅ α)"
proof -
have "x ∈ subst_domain (αinv ∘⇩s δ)" when x: "x ∈ fv (s ⋅ α)" for x
proof -
obtain y where y: "y ∈ fv s" "α y = Var x"
using α'(3)[OF x] by blast
have "y ∈ subst_domain δ" using y(1) δ(3) by blast
moreover have "(αinv ∘⇩s δ) x = δ y"
using y α'(3)[OF x] α_invertible
vars_term_subset_subst_eq[of "Var y" s "α ∘⇩s αinv" Var]
unfolding δ'_def αinv_def
by (metis (no_types, lifting) fv_subst_subset eval_term.simps(1)
subst_apply_term_empty subst_compose)
ultimately show ?thesis using δ(4) by fastforce
qed
thus ?thesis using rm_vars_dom[of "UNIV - fv (s ⋅ α)" "αinv ∘⇩s δ"] unfolding δ'_def by blast
qed
have δ'_range: "fv t = {}" when t: "t ∈ (subst_range δ')" for t
proof -
obtain x where "x ∈ fv (s ⋅ α)" "δ' x = t" using t δ'_domain by auto
thus ?thesis
by (metis (no_types, lifting) δ'_def subst_compose_def δ(3,4) α_invertible fv_subst_subset
subst_fv_dom_ground_if_ground_img subst_subst_compose Diff_iff)
qed
have J0: "x ∈ fv (s ⋅ α) ⟹ J x = δ' x"
"x ∉ fv (s ⋅ α) ⟹ J x = I x" for x
unfolding J_def by (cases "x ∈ fv (s ⋅ α)") (simp_all add: subst_compose)
have J1: "subst_range J ⊆ subst_range δ' ∪ subst_range I"
proof
fix t assume "t ∈ subst_range J"
then obtain x where x: "x ∈ subst_domain J" "J x = t" by auto
hence "t = δ' x ⟹ x ∈ subst_domain δ'" "t = I x ⟹ x ∈ subst_domain I"
by (metis subst_domI subst_dom_vars_in_subst)+
thus "t ∈ subst_range δ' ∪ subst_range I" using x(2) J0[of x] by auto
qed
have "x ∉ fv (s ⋅ α)" when x: "x ∈ fv⇩l⇩s⇩s⇩t (A@[⟨⋆, send⟨[s ⋅ δ]⟩⟩])" for x
using x δ_s α'(4) α'(5) by auto
hence "I x = J x" when x: "x ∈ fv⇩l⇩s⇩s⇩t (A@[⟨⋆, send⟨[s ⋅ δ]⟩⟩])" for x
using x unfolding J_def δ'_def by auto
hence "constr_sem_stateful J (unlabel (A@[⟨⋆, send⟨[s ⋅ δ]⟩⟩]))"
using A strand_sem_model_swap[of "unlabel (A@[⟨⋆, send⟨[s ⋅ δ]⟩⟩])" I J "{}" "{}"]
unfolding defs by blast
moreover have "wt⇩s⇩u⇩b⇩s⇩t J"
using A subst_var_inv_wt[OF α'(6), of "fv s"]
wt_subst_trm''[OF δ(1)] subst_compose[of "subst_var_inv α (fv s)" δ]
unfolding defs J_def δ'_def αinv_def wt⇩s⇩u⇩b⇩s⇩t_def by presburger
moreover have "interpretation⇩s⇩u⇩b⇩s⇩t J"
proof -
have "fv t = {}" when t: "t ∈ (subst_range J)" for t
using t A J1 δ'_range unfolding defs by auto
moreover have "x ∈ subst_domain J" for x
proof (cases "x ∈ fv (s ⋅ α)")
case True thus ?thesis using J0(1)[of x] δ'_domain unfolding subst_domain_def by auto
next
case False
have "subst_domain I = UNIV" using A unfolding defs by fast
thus ?thesis using J0(2)[OF False] unfolding subst_domain_def by auto
qed
ultimately show ?thesis by auto
qed
moreover have "wf⇩t⇩r⇩m⇩s (subst_range δ')"
using wf_trms_subst_compose[OF subst_var_inv_wf_trms[of α "fv s"] δ(2)]
unfolding δ'_def αinv_def by force
hence "wf⇩t⇩r⇩m⇩s (subst_range J)" using A J1 unfolding defs by fast
ultimately show "welltyped_constraint_model J (A@[⟨⋆, send⟨[s ⋅ δ]⟩⟩])" unfolding defs by blast
hence "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J ⊢ s ⋅ δ"
using δ_s strand_sem_append_stateful[of "{}" "{}" "unlabel A" "[send⟨[s ⋅ δ]⟩]" J]
unfolding defs by (simp add: subst_ground_ident)
moreover have "s ⋅ α ⋅ J = s ⋅ δ"
proof -
have "J x = δ' x" when x: "x ∈ fv (s ⋅ α)" for x using x unfolding J_def by argo
hence "s ⋅ α ⋅ J = s ⋅ α ⋅ δ'" using subst_agreement[of "s ⋅ α" J δ'] by force
thus ?thesis
using α_invertible unfolding δ'_def rm_vars_subst_eq'[symmetric] by (metis subst_subst_compose)
qed
hence "s ⋅ α ⋅ J = s ⋅ δ" by auto
ultimately show "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J ⊢ s ⋅ α ⋅ J" by argo
qed
lemma welltyped_leakage_free_protocol_pointwise:
"welltyped_leakage_free_protocol S P ⟷ list_all (λs. welltyped_leakage_free_protocol [s] P) S"
unfolding welltyped_leakage_free_protocol_def list_all_iff Let_def by fastforce
lemma welltyped_leakage_free_no_deduct_constI:
fixes c
defines "s ≡ Fun c []::('fun,'atom,'sets,'lbl) prot_term"
assumes s: "∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[s]⟩⟩])"
shows "welltyped_leakage_free_protocol [s] P"
using s unfolding welltyped_leakage_free_protocol_def s_def by auto
lemma welltyped_leakage_free_pub_termI:
assumes s: "{} ⊢⇩c s"
shows "welltyped_leakage_free_protocol [s] P"
proof -
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define f where "f ≡ λM. {t ⋅ δ | t δ. Q M t δ}"
define Sec where "Sec ≡ f (set [s]) - {m. {} ⊢⇩c m}"
have 0: "fv s = {}" using s pgwt_ground pgwt_is_empty_synth by blast
have 1: "s ⋅ δ = s" for δ by (rule subst_ground_ident[OF 0])
have 2: "wt⇩s⇩u⇩b⇩s⇩t Var" "wf⇩t⇩r⇩m⇩s (subst_range Var)"
using wt_subst_Var wf_trm_subst_range_Var by (blast,blast)
have "f (set [s]) = {s}"
proof
show "f (set [s]) ⊆ {s}" using 0 1 unfolding f_def Q_def by auto
have "Q {s} s Var" using 0 2 unfolding Q_def by auto
thus "{s} ⊆ f (set [s])" using 1[of Var] unfolding f_def by force
qed
hence "Sec = {}" using s unfolding Sec_def by simp
thus ?thesis unfolding welltyped_leakage_free_protocol_def Let_def Sec_def f_def Q_def by blast
qed
lemma welltyped_leakage_free_pub_constI:
assumes c: "public⇩f c" "arity⇩f c = 0"
shows "welltyped_leakage_free_protocol [⟨c⟩⇩c] P"
using c welltyped_leakage_free_pub_termI[OF intruder_synth.ComposeC[of "[]" "Fu c" "{}"]] by simp
lemma welltyped_leakage_free_long_term_secretI:
fixes n
defines
"Tatt ≡ λs'. Transaction (λ(). []) [] [⟨n, receive⟨[s']⟩⟩] [] [] [⟨n, send⟨[attack⟨ln n⟩]⟩⟩]"
assumes P_wt_secure:
"∀𝒜 ∈ reachable_constraints P. ∄ℐ.
welltyped_constraint_model ℐ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
and s_long_term_secret:
"∃θ. wt⇩s⇩u⇩b⇩s⇩t θ ∧ inj_on θ (fv s) ∧ θ ` fv s ⊆ range Var ∧ Tatt (s ⋅ θ) ∈ set P"
shows "welltyped_leakage_free_protocol [s] P"
proof (rule ccontr)
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define f where "f ≡ λM. {t ⋅ δ | t δ. Q M t δ}"
define Sec where "Sec ≡ f (set [s]) - {m. {} ⊢⇩c m}"
note defs = Sec_def f_def Q_def
note defs' = welltyped_constraint_model_def constraint_model_def
assume "¬welltyped_leakage_free_protocol [s] P"
then obtain A I s' where A:
"A ∈ reachable_constraints P" "s' ∈ Sec - declassified⇩l⇩s⇩s⇩t A I"
"welltyped_constraint_model I (A@[⟨⋆, send⟨[s']⟩⟩])"
unfolding welltyped_leakage_free_protocol_def defs by fastforce
obtain θ where θ: "wt⇩s⇩u⇩b⇩s⇩t θ" "θ ` fv s ⊆ range Var" "inj_on θ (fv s)" "Tatt (s ⋅ θ) ∈ set P"
using s_long_term_secret by blast
obtain δ where δ:
"wt⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" "subst_domain δ = fv (s ⋅ θ)" "ground (subst_range δ)"
"s' = s ⋅ θ ⋅ δ"
proof -
obtain δ where *: "wt⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" "fv s' = {}" "s' = s ⋅ δ"
using A(2) unfolding defs by auto
define σ where "σ ≡ subst_var_inv θ (fv s) ∘⇩s δ"
define δ' where "δ' ≡ rm_vars (UNIV - fv (s ⋅ θ)) σ"
have **: "s' = s ⋅ θ ⋅ σ"
using *(4) inj_var_ran_subst_is_invertible[OF θ(3,2)]
unfolding σ_def by simp
have "s' = s ⋅ θ ⋅ δ'"
using ** rm_vars_subst_eq'[of "s ⋅ θ" σ]
unfolding δ'_def by simp
moreover have "wt⇩s⇩u⇩b⇩s⇩t σ"
using θ(1) *(1) subst_var_inv_wt wt_subst_compose
unfolding σ_def by presburger
hence "wt⇩s⇩u⇩b⇩s⇩t δ'" using wt_subst_rm_vars unfolding δ'_def by blast
moreover have "wf⇩t⇩r⇩m⇩s (subst_range σ)"
using wf_trms_subst_compose[OF subst_var_inv_wf_trms *(2)] unfolding σ_def by blast
hence "wf⇩t⇩r⇩m⇩s (subst_range δ')" using wf_trms_subst_rm_vars'[of σ] unfolding δ'_def by blast
moreover have "fv (s ⋅ θ) ⊆ subst_domain σ"
using *(3) ** ground_term_subst_domain_fv_subset unfolding σ_def by blast
hence "subst_domain δ' = fv (s ⋅ θ)"
using rm_vars_dom[of "UNIV - fv (s ⋅ θ)" σ] unfolding δ'_def by blast
moreover have "ground (subst_range δ')"
proof -
{ fix t assume "t ∈ subst_range δ'"
then obtain x where "x ∈ fv (s ⋅ θ)" "δ' x = t"
using ‹subst_domain δ' = fv (s ⋅ θ)› by auto
hence "t ⊑ s ⋅ θ ⋅ δ'" by (meson subst_mono_fv)
hence "fv t = {}" using ‹s' = s ⋅ θ ⋅ δ'› *(3) ground_subterm by blast
} thus ?thesis by force
qed
ultimately show thesis using that[of δ'] by fast
qed
have ξ: "transaction_decl_subst Var (Tatt t)"
and σ: "transaction_fresh_subst Var (Tatt t) (trms⇩l⇩s⇩s⇩t A)"
for t
unfolding transaction_decl_subst_def transaction_fresh_subst_def Tatt_def by simp_all
obtain α::"('fun,'atom,'sets,'lbl) prot_subst"
where α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)"
unfolding transaction_renaming_subst_def by blast
obtain J where J:
"welltyped_constraint_model J (A@[⟨⋆, send⟨[s ⋅ θ ⋅ δ]⟩⟩])" "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J ⊢ s ⋅ θ ⋅ α ⋅ J"
using welltyped_constraint_model_leakage_model_swap[OF A(3)[unfolded δ(5)] α δ(1-4)] by blast
define T where "T = dual⇩l⇩s⇩s⇩t (transaction_strand (Tatt (s ⋅ θ)) ⋅⇩l⇩s⇩s⇩t α)"
define B where "B ≡ A@T"
have "transaction_receive (Tatt t) = [⟨n, receive⟨[t]⟩⟩]"
"transaction_checks (Tatt t) = []"
"transaction_updates (Tatt t) = []"
"transaction_send (Tatt t) = [⟨n, send⟨[attack⟨ln n⟩]⟩⟩]"
for t
unfolding Tatt_def by simp_all
hence T_def': "T = [⟨n, send⟨[s ⋅ θ ⋅ α]⟩⟩, ⟨n, receive⟨[attack⟨ln n⟩]⟩⟩]"
using subst_lsst_append[of "transaction_receive (Tatt (s ⋅ θ))" _ α]
subst_lsst_singleton[of "ln n" "receive⟨[s ⋅ θ]⟩" α]
subst_lsst_singleton[of "ln n" "send⟨[attack⟨ln n⟩]⟩" α]
unfolding transaction_strand_def T_def by fastforce
have B0: "ik⇩l⇩s⇩s⇩t B ⋅⇩s⇩e⇩t J ⊢ attack⟨ln n⟩"
using in_ik⇩s⇩s⇩t_iff[of "attack⟨ln n⟩" "unlabel T"]
unfolding B_def T_def' by (force intro!: intruder_deduct.Axiom)
have B1: "B ∈ reachable_constraints P"
using reachable_constraints.step[OF A(1) θ(4) ξ σ α]
unfolding B_def T_def by simp
have "welltyped_constraint_model J B"
using J strand_sem_append_stateful[of "{}" "{}" "unlabel A" _ J]
unfolding defs' B_def T_def' by fastforce
hence B2: "welltyped_constraint_model J (B@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
using B0 strand_sem_append_stateful[of "{}" "{}" "unlabel B" "[send⟨[attack⟨ln n⟩]⟩]" J]
unfolding defs' B_def by auto
show False using P_wt_secure B1 B2 by blast
qed
lemma welltyped_leakage_free_value_constI:
assumes P:
"∀T ∈ set P. wellformed_transaction T"
"∀T ∈ set P. admissible_transaction_terms T"
"∀T ∈ set P. transaction_decl T () = []"
"∀T ∈ set P. bvars_transaction T = {}"
and P_fresh_declass:
"∀T ∈ set P. transaction_fresh T ≠ [] ⟶
(transaction_send T ≠ [] ∧ (let (l,a) = hd (transaction_send T)
in l = ⋆ ∧ is_Send a ∧ Var ` set (transaction_fresh T) ⊆ set (the_msgs a)))"
shows "welltyped_leakage_free_protocol [⟨m: value⟩⇩v] P"
proof (rule ccontr)
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define f where "f ≡ λM. {t ⋅ δ | t δ. Q M t δ}"
define Sec where "Sec ≡ f (set [⟨m: value⟩⇩v]) - {m. {} ⊢⇩c m}"
note defs = Sec_def f_def Q_def
note defs' = welltyped_constraint_model_def constraint_model_def
assume "¬welltyped_leakage_free_protocol [⟨m: value⟩⇩v] P"
then obtain A I s where A:
"A ∈ reachable_constraints P" "s ∈ Sec - declassified⇩l⇩s⇩s⇩t A I"
"welltyped_constraint_model I (A@[⟨⋆, send⟨[s]⟩⟩])"
unfolding welltyped_leakage_free_protocol_def defs by fastforce
have "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ s ⋅ I" using welltyped_constraint_model_deduct_split[OF A(3)] by simp
moreover have "s ⋅ I = s" using A(2) unfolding defs by fast
ultimately have s_deduct: "ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I) ⊢ s" by (metis ik⇩s⇩s⇩t_subst unlabel_subst)
note I0 = welltyped_constraint_model_prefix[OF A(3)]
have I1: "wt⇩s⇩u⇩b⇩s⇩t I" using A(3) unfolding defs' by blast
obtain f ts δ where f: "s = Fun f ts" "s = ⟨m: value⟩⇩v ⋅ δ" "¬{} ⊢⇩c s" "s ∉ declassified⇩l⇩s⇩s⇩t A I"
and δ: "wt⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" "fv s = {}"
using A(2) unfolding defs by (cases s) auto
have s1: "Γ s = TAtom Value"
by (metis Γ.simps(1) Γ⇩v_TAtom f(2) wt_subst_trm''[OF δ(1)])
have s2: "wf⇩t⇩r⇩m s"
using f(2) δ(2) by force
have s3: "ts = []"
using f(1) s1 s2 const_type_inv_wf by blast
obtain sn where sn: "s = Fun (Val sn) []"
using s1 f(3) Γ_Fu_simps(4) Γ_Set_simps(3) unfolding f(1) s3 by (cases f) auto
have "s ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I"
using private_fun_deduct_in_ik'[OF s_deduct[unfolded sn]]
by (metis sn public.simps(3) ik⇩s⇩s⇩t_subst unlabel_subst)
hence s4: "s ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t A"
using constraint_model_Val_const_in_constr_prefix[OF A(1) I0 P(1,2)]
unfolding sn by presburger
obtain B T ξ σ α where B:
"prefix (B@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) A"
"B ∈ reachable_constraints P" "T ∈ set P" "transaction_decl_subst ξ T"
"transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t B)"
"s ∈ subst_range σ"
using constraint_model_Value_in_constr_prefix_fresh_action'[OF A(1) P(2-) s4[unfolded sn]] sn
by blast
obtain Tts Tsnds sx
where T: "transaction_send T = ⟨⋆, send⟨Tts⟩⟩#Tsnds" "Var ` set (transaction_fresh T) ⊆ set Tts"
and sx: "Var sx ∈ set Tts" "σ sx = s"
using P_fresh_declass B(3,5,7)
unfolding transaction_fresh_subst_def is_Send_def
by (cases "transaction_send T") (fastforce,fastforce)
have ξ_elim: "ξ ∘⇩s σ ∘⇩s α = σ ∘⇩s α"
using admissible_transaction_decl_subst_empty'[OF bspec[OF P(3) B(3)] B(4)]
by simp
have s5: "s ∈ set (Tts ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α ∘⇩s I)"
using sx unfolding ξ_elim sn by force
have s6: "⟨⋆, receive⟨Tts ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α ∘⇩s I⟩⟩ ∈ set (A ⋅⇩l⇩s⇩s⇩t I)"
proof -
have "⟨⋆, send⟨Tts⟩⟩ ∈ set (transaction_send T)"
using T(1) by simp
hence "⟨⋆, send⟨Tts ⋅⇩l⇩i⇩s⇩t δ⟩⟩ ∈ set (transaction_send T ⋅⇩l⇩s⇩s⇩t δ)" for δ
unfolding subst_apply_labeled_stateful_strand_def by force
hence "⟨⋆, send⟨Tts ⋅⇩l⇩i⇩s⇩t δ⟩⟩ ∈ set (transaction_strand T ⋅⇩l⇩s⇩s⇩t δ)" for δ
using transaction_strand_subst_subsets(4)[of T δ] by fast
hence *: "⟨⋆, receive⟨Tts ⋅⇩l⇩i⇩s⇩t δ⟩⟩ ∈ set (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t δ))" for δ
using dual⇩l⇩s⇩s⇩t_steps_iff(1)[of ⋆ "Tts ⋅⇩l⇩i⇩s⇩t δ"] by blast
have "⟨⋆, receive⟨Tts ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α⟩⟩ ∈ set A"
using B(1) *[of "ξ ∘⇩s σ ∘⇩s α"] unfolding prefix_def by force
thus ?thesis
unfolding subst_apply_labeled_stateful_strand_def by force
qed
show False
using s6 f(4) ideduct_mono[OF Axiom[OF s5], of "⋃{set ts|ts. ⟨⋆,receive⟨ts⟩⟩ ∈ set (A ⋅⇩l⇩s⇩s⇩t I)}"]
unfolding declassified⇩l⇩s⇩s⇩t_def by blast
qed
lemma welltyped_leakage_free_priv_constI:
fixes c
defines "s ≡ Fun c []::('fun,'atom,'sets,'lbl) prot_term"
assumes c: "¬public c" "arity c = 0" "Γ s = TAtom ca" "ca ≠ Value"
and P: "∀T ∈ set P. ∀x ∈ vars_transaction T. is_Var (Γ⇩v x)"
"∀T ∈ set P. ∀x ∈ vars_transaction T ∪ set (transaction_fresh T). Γ s ≠ Γ⇩v x"
"∀T ∈ set P. ∀t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). s ∉ set (snd (Ana t))"
"∀T ∈ set P. s ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)"
"∀T ∈ set P. ∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
"∀T ∈ set P. wellformed_transaction T"
shows "∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[s]⟩⟩])"
(is "∀𝒜 ∈ ?R. ?P 𝒜")
and "welltyped_leakage_free_protocol [s] P"
proof -
show "∀𝒜 ∈ ?R. ?P 𝒜"
proof
fix A assume A: "A ∈ reachable_constraints P"
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define f where "f ≡ λM. {t ⋅ δ | t δ. Q M t δ}"
define Sec where "Sec ≡ f (set [s]) - {m. {} ⊢⇩c m}"
define f' where "f' ≡ λ(T,ξ,σ::('fun,'atom,'sets,'lbl) prot_subst,α).
dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define g' where "g' ≡ concat ∘ map f'"
let ?P_s_cases = "λM. s ∈ M ∨ (∃m ∈ subterms⇩s⇩e⇩t M. s ∈ set (snd (Ana m)))"
let ?P_s_cases' = "λM δ. s ∈ M ⋅⇩s⇩e⇩t δ ∨ (∃m ∈ subterms⇩s⇩e⇩t M ⋅⇩s⇩e⇩t δ. s ∈ set (snd (Ana m)))"
note defs = Sec_def f_def Q_def
note defs' = welltyped_constraint_model_def constraint_model_def
show "?P A"
proof (rule ccontr)
assume "¬?P A"
then obtain I where I: "welltyped_constraint_model I (A@[⟨⋆, send⟨[s]⟩⟩])" by blast
obtain Ts where Ts:
"A = g' Ts" "∀B. prefix B Ts ⟶ g' B ∈ reachable_constraints P"
"∀B T ξ σ α. prefix (B@[(T,ξ,σ,α)]) Ts ⟶
T ∈ set P ∧ transaction_decl_subst ξ T ∧
transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t (g' B)) ∧
transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t (g' B))"
using reachable_constraints_as_transaction_lists[OF A(1)] unfolding g'_def f'_def by blast
have "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ s ⋅ I" and I_s: "s ⋅ I = s"
using welltyped_constraint_model_deduct_split[OF I]
unfolding s_def by simp_all
hence s_deduct: "ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I) ⊢ s" "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ s"
by (metis ik⇩s⇩s⇩t_subst unlabel_subst, metis)
have I_wt: "wt⇩s⇩u⇩b⇩s⇩t I"
and I_wf: "wf⇩t⇩r⇩m⇩s (subst_range I)"
and I_grounds: "ground (subst_range I)"
and I_interp: "interpretation⇩s⇩u⇩b⇩s⇩t I"
using I unfolding defs' by (blast,blast,blast,blast)
have Sec_unfold: "Sec = {s}"
proof
have "¬{} ⊢⇩c s" using ideduct_synth_priv_const_in_ik[OF _ c(1)] unfolding s_def by blast
thus "{s} ⊆ Sec" unfolding defs s_def by fastforce
qed (auto simp add: defs s_def)
have s2: "wf⇩t⇩r⇩m s"
using c(1,2) unfolding s_def by fastforce
have A_ik_fv: "∃a. Γ⇩v x = TAtom a ∧ a ≠ ca" when x: "x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A)" for x
proof -
obtain T where T: "T ∈ set P" "Γ⇩v x ∈ Γ⇩v ` fv_transaction T"
using fv_ik⇩s⇩s⇩t_is_fv⇩s⇩s⇩t[OF x] reachable_constraints_var_types_in_transactions(1)[OF A P(5)]
by fast
then obtain y where y: "y ∈ vars_transaction T" "Γ⇩v y = Γ⇩v x"
using vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel (transaction_strand T)"] by fastforce
then obtain a where a: "Γ⇩v y = TAtom a" using P(1) T(1) by blast
hence "Γ⇩v x = TAtom a" "Γ s ≠ Γ⇩v x" "Γ s = TAtom ca" using y P(2) T(1) c(3) by auto
thus ?thesis by force
qed
have I_s_x: "¬s ⊑ I x" when x: "x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A)" for x
proof -
obtain a where a: "Γ⇩v x = TAtom a" "a ≠ ca" using A_ik_fv[OF x] by blast
hence a': "Γ (I x) = TAtom a" using wt_subst_trm''[OF I_wt, of "Var x"] by simp
obtain f ts where f: "I x = Fun f ts"
by (meson empty_fv_exists_fun interpretation_grounds_all[OF I_interp])
hence ts: "ts = []"
using I_wf const_type_inv_wf[OF a'[unfolded f]] by fastforce
have "c ≠ f" using f[unfolded ts] a a' c(3)[unfolded s_def] by force
thus ?thesis using f ts unfolding s_def by simp
qed
have A_ik_I_const: "∃f. arity f = 0 ∧ I x = Fun f []" when x: "x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A)" for x
using x A_ik_fv I_wt empty_fv_exists_fun[OF interpretation_grounds_all[OF I_interp, of x]]
wf_trm_subst_rangeD[OF I_wf, of x] const_type_inv const_type_inv_wf
by (metis (no_types, lifting) Γ.simps(1) wt⇩s⇩u⇩b⇩s⇩t_def)
hence A_ik_subst: "subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) = subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A) ⋅⇩s⇩e⇩t I"
using subterms_subst''[of "ik⇩l⇩s⇩s⇩t A" I] by blast
have sublmm1: "s ∈ set (snd (Ana m))"
when m: "m ⊑⇩s⇩e⇩t M" "s ∈ set (snd (Ana (m ⋅ δ)))"
and M: "⋀y. y ∈ fv⇩s⇩e⇩t M ⟹ ¬s ⊑ δ y"
for m M δ
proof -
have m_fun: "is_Fun m"
using m M Ana_subterm' vars_iff_subtermeq_set
unfolding s_def is_Var_def by (metis eval_term.simps(1))
obtain f K R ts i where f:
"m = ⟨f ts⟩⇩t" "arity⇩f f = length ts" "arity⇩f f > 0" "Ana⇩f f = (K, R)"
and i: "i < length R" "s = ts ! (R ! i) ⋅ δ"
and R_i: "∀i < length R. map ((!) ts) R ! i = ts ! (R ! i) ∧ R ! i < length ts"
proof -
obtain f ts K R where f:
"m ⋅ δ = ⟨f ts⟩⇩t" "arity⇩f f = length ts" "arity⇩f f > 0"
"Ana⇩f f = (K, R)" "Ana (m ⋅ δ) = (K ⋅⇩l⇩i⇩s⇩t (!) ts, map ((!) ts) R)"
using m(2) Ana_nonempty_inv[of "m ⋅ δ"] by force
obtain ts' where m': "m = ⟨f ts'⟩⇩t" "ts = ts' ⋅⇩l⇩i⇩s⇩t δ"
using f(1) m_fun by auto
have R_i: "map ((!) ts) R ! i = ts ! (R ! i)" "R ! i < length ts"
when i: "i < length R" for i
using i Ana⇩f_assm2_alt[OF f(4), of "R ! i"] f(2) by simp_all
then obtain i where i: "s = ts ! (R ! i)" "i < length R"
by (metis (no_types, lifting) m(2) f(5) in_set_conv_nth length_map snd_conv)
have ts': "arity⇩f f = length ts'" "length ts = length ts'" using m'(2) f(2) by simp_all
have s': "s = ts' ! (R ! i) ⋅ δ" using R_i(2)[OF i(2)] i(1) unfolding ts'(2) m'(2) by simp
show thesis using that f m' R_i ts' s' i by auto
qed
have "s = ts ! (R ! i)"
proof (cases "ts ! (R ! i)")
case (Var x)
hence "Var x ∈ set ts" using R_i i nth_mem by fastforce
hence "x ∈ fv⇩s⇩e⇩t M" using m(1) f(1) fv_subterms_set by fastforce
thus ?thesis using i M Var by fastforce
qed (use i s_def in fastforce)
thus "s ∈ set (snd (Ana m))" using f(1) Ana_Fu_intro[OF f(2-4)] i(1) by simp
qed
have "¬s ⊑ δ y"
when m: "m ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t (transaction_send T)" "s ∈ set (snd (Ana (m ⋅ δ)))"
and T: "T ∈ set P" and δ_wt: "wt⇩s⇩u⇩b⇩s⇩t δ"
and δ_ran: "⋀t. t ∈ subst_range δ ⟹ (∃c. t = Fun c [] ∧ arity c = 0) ∨ (∃x. t = Var x)"
and y: "y ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))"
for m T δ y
proof
assume "s ⊑ δ y"
hence "Γ⇩v y = Γ s" using wt_subst_trm''[OF δ_wt, of "Var y"] δ_ran[of "δ y"] by fastforce
moreover have "y ∈ vars_transaction T"
using y trms⇩s⇩s⇩t_fv_vars⇩s⇩s⇩t_subset unfolding vars_transaction_unfold[of T] by fastforce
ultimately show False using P(2) T by force
qed
hence sublmm2: "s ∈ set (snd (Ana m))"
when m: "m ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t (transaction_send T)" "s ∈ set (snd (Ana (m ⋅ δ)))"
and T: "T ∈ set P" and δ_wt: "wt⇩s⇩u⇩b⇩s⇩t δ"
and δ_ran: "⋀t. t ∈ subst_range δ ⟹ (∃c. t = Fun c [] ∧ arity c = 0) ∨ (∃x. t = Var x)"
for m T δ
using sublmm1[OF m] m T δ_wt δ_ran by blast
have "s ∈ ik⇩l⇩s⇩s⇩t A ∨ (∃m ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A) ⋅⇩s⇩e⇩t I. s ∈ set (snd (Ana m)))"
using private_const_deduct[OF c(1) s_deduct(2)[unfolded s_def]]
I_s_x const_mem_subst_cases[of c] A_ik_subst
unfolding s_def by blast
hence "?P_s_cases (ik⇩l⇩s⇩s⇩t A)" using sublmm1[of _ "ik⇩l⇩s⇩s⇩t A"] I_s_x by blast
then obtain T ξ σ α where T: "(T,ξ,σ,α) ∈ set Ts" "?P_s_cases (ik⇩l⇩s⇩s⇩t (f' (T,ξ,σ,α)))"
using ik⇩l⇩s⇩s⇩t_concat[of "map f' Ts"] Ts(1)[unfolded g'_def] by fastforce
obtain B where "prefix (B@[(T, ξ, σ, α)]) Ts" by (metis T(1) prefix_snoc_in_iff)
hence T_in_P: "T ∈ set P"
and T_wf: "wellformed_transaction T"
and ξ: "transaction_decl_subst ξ T"
and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t (concat (map f' B)))"
and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t (concat (map f' B)))"
using P(6) Ts(3)[unfolded g'_def] unfolding comp_def by (metis,metis,metis,metis,metis)
note ξσα_wt = transaction_decl_fresh_renaming_substs_wt[OF ξ σ α]
note ξσα_ran = transaction_decl_fresh_renaming_substs_range'(1)[OF ξ σ α]
have "subterms⇩s⇩e⇩t (M ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α) = subterms⇩s⇩e⇩t M ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" for M
using ξσα_ran subterms_subst''[of _ "ξ ∘⇩s σ ∘⇩s α"] by (meson subst_imgI)
hence s_cases: "?P_s_cases' (trms⇩l⇩s⇩s⇩t (transaction_send T)) (ξ ∘⇩s σ ∘⇩s α)"
using T(2) dual_transaction_ik_is_transaction_send'[OF T_wf, of "ξ ∘⇩s σ ∘⇩s α"]
unfolding f'_def by auto
from s_cases show False
proof
assume "s ∈ trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α"
then obtain t where t: "t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T)" "s = t ⋅ ξ ∘⇩s σ ∘⇩s α" by force
have "s ≠ t" using P(4) T_in_P t(1) by blast
then obtain x where x: "t = Var x" using t(2) unfolding s_def by (cases t) auto
have "Γ⇩v x = Γ s" using x t(2) wt_subst_trm''[OF ξσα_wt, of "Var x"] by simp
moreover have "x ∈ vars_transaction T"
using t(1) trms⇩s⇩s⇩t_fv_vars⇩s⇩s⇩t_subset unfolding x vars_transaction_unfold[of T] by fastforce
ultimately show False using P(2) T_in_P by force
qed (use sublmm2[OF _ _ T_in_P ξσα_wt ξσα_ran] P(3) T_in_P in blast)
qed
qed
thus "welltyped_leakage_free_protocol [s] P"
using welltyped_leakage_free_no_deduct_constI[of P c]
unfolding s_def by blast
qed
lemma welltyped_leakage_free_priv_constI':
assumes c: "¬public⇩f c" "arity⇩f c = 0" "Γ⇩f c = Some ca"
and P:
"∀T ∈ set P. wellformed_transaction T"
"∀T ∈ set P. ∀x ∈ vars_transaction T ∪ set (transaction_fresh T). Γ ⟨c⟩⇩c ≠ Γ⇩v x"
"∀T ∈ set P. ∀x ∈ vars_transaction T. is_Var (Γ⇩v x)"
"∀T ∈ set P. ∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
"∀T ∈ set P. ∀t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). ⟨c⟩⇩c ∉ set (snd (Ana t))"
"∀T ∈ set P. ⟨c⟩⇩c ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)"
shows "∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[⟨c⟩⇩c]⟩⟩])"
and "welltyped_leakage_free_protocol [⟨c⟩⇩c] P"
using c welltyped_leakage_free_priv_constI[OF _ _ _ _ P(3,2,5,6,4,1), of "Atom ca"]
by (force, force)
lemma welltyped_leakage_free_set_constI:
assumes P:
"∀T ∈ set P. wellformed_transaction T"
"∀T ∈ set P. ∀f ∈ ⋃(funs_term ` (trms⇩l⇩s⇩s⇩t (transaction_send T))). ¬is_Set f"
"∀T ∈ set P. ∀x ∈ vars_transaction T ∪ set (transaction_fresh T). Γ⇩v x ≠ TAtom SetType"
"∀T ∈ set P. ∀x ∈ vars_transaction T. is_Var (Γ⇩v x)"
"∀T ∈ set P. ∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
and c: "arity⇩s c = 0"
shows "∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[⟨c⟩⇩s]⟩⟩])"
and "welltyped_leakage_free_protocol [⟨c⟩⇩s] P"
proof -
have c'': "⟨c⟩⇩s ∉ subterms t"
when T: "T ∈ set P" and t: "t ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t (transaction_send T)" for T t
using t bspec[OF P(2) T] subtermeq_imp_funs_term_subset[of t]
funs_term_Fun_subterm'[of "Set c" "[]::('fun,'atom,'sets,'lbl) prot_term list"]
by fastforce
have P':
"∀T ∈ set P. ∀t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). ⟨c⟩⇩s ∉ set (snd (Ana t))"
"∀T ∈ set P. ⟨c⟩⇩s ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)"
subgoal using Ana_subterm' c'' by fast
subgoal using c'' by fast
done
have P'':
"∀T ∈ set P. ∀x ∈ vars_transaction T ∪ set (transaction_fresh T). Γ ⟨c⟩⇩s ≠ Γ⇩v x"
using P(3) Γ_consts_simps(4)[OF c] by fastforce
show "∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[⟨c⟩⇩s]⟩⟩])"
"welltyped_leakage_free_protocol [⟨c⟩⇩s] P"
using c welltyped_leakage_free_priv_constI[OF _ _ _ _ P(4) P'' P' P(5,1), of SetType]
by (force, force)
qed
lemma :
defines "s ≡ Fun OccursSec []"
assumes P:
"∀T ∈ set P. wellformed_transaction T"
"∀T ∈ set P. ∀x ∈ vars_transaction T ∪ set (transaction_fresh T). Γ⇩v x ≠ TAtom OccursSecType"
"∀T ∈ set P. ∀t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). Fun OccursSec [] ∉ set (snd (Ana t))"
"∀T ∈ set P. Fun OccursSec [] ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)"
"∀T ∈ set P. ∀x ∈ vars_transaction T. is_Var (Γ⇩v x)"
"∀T ∈ set P. ∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
shows "∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[s]⟩⟩])"
and "welltyped_leakage_free_protocol [s] P"
proof -
have P':
"∀T ∈ set P. ∀x ∈ vars_transaction T ∪ set (transaction_fresh T). Γ (Fun OccursSec []) ≠ Γ⇩v x"
using P(2) by auto
show "∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[s]⟩⟩])"
"welltyped_leakage_free_protocol [s] P"
using welltyped_leakage_free_priv_constI[OF _ _ _ _ P(5) P' P(3,4,6,1), of OccursSecType]
unfolding s_def by auto
qed
lemma welltyped_leakage_free_occurs_factI:
assumes P: "∀T ∈ set P. admissible_transaction' T"
and P_occ: "∀T ∈ set P. admissible_transaction_occurs_checks T"
and P_occ_star:
"∀T ∈ set P. ∀r ∈ set (transaction_send T).
OccursFact ∈ ⋃(funs_term ` (trms⇩s⇩s⇩t⇩p (snd r))) ⟶ fst r = ⋆"
shows "welltyped_leakage_free_protocol [occurs x] P"
proof -
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define f where "f ≡ λM. {t ⋅ δ | t δ. Q M t δ}"
define Sec where "Sec ≡ f (set [occurs x]) - {m. {} ⊢⇩c m}"
define f' where "f' ≡ λ(T,ξ,σ::('fun,'atom,'sets,'lbl) prot_subst,α).
dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define g' where "g' ≡ concat ∘ map f'"
note defs = Sec_def f_def Q_def
note defs' = welltyped_constraint_model_def constraint_model_def
show ?thesis
proof (rule ccontr)
assume "¬welltyped_leakage_free_protocol [occurs x] P"
then obtain A I k where A:
"A ∈ reachable_constraints P" "occurs k ∈ Sec - declassified⇩l⇩s⇩s⇩t A I"
"welltyped_constraint_model I (A@[⟨⋆, send⟨[occurs k]⟩⟩])"
unfolding welltyped_leakage_free_protocol_def defs by fastforce
note A' = welltyped_constraint_model_prefix[OF A(3)]
have occ_I: "occurs k ⋅ I = occurs k" using A(2) unfolding defs by auto
hence occ_in_ik: "occurs k ∈ ik⇩l⇩s⇩s⇩t A" "occurs k ∈ ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I"
using welltyped_constraint_model_deduct_split(2)[OF A(3)]
reachable_constraints_occurs_fact_deduct_in_ik[OF A(1) A' P P_occ, of k]
by (argo, argo)
then obtain l ts where ts: "(l,receive⟨ts⟩) ∈ set A" "occurs k ∈ set ts"
using in_ik⇩s⇩s⇩t_iff[of "occurs k" "unlabel A"] unfolding unlabel_def by force
obtain T a B α σ ξ
where B: "prefix (B@f' (T,ξ,σ,α)) A"
and T: "T ∈ set P" "transaction_decl_subst ξ T"
"transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)"
"transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t B)"
and a: "a ∈ set (transaction_strand T)" "fst (l,receive⟨ts⟩) = fst a"
"(l,receive⟨ts⟩) = dual⇩l⇩s⇩s⇩t⇩p a ⋅⇩l⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩s α"
using reachable_constraints_transaction_action_obtain[OF A(1) ts(1), of thesis]
unfolding f'_def by simp
obtain ts' where ts': "a = (l,send⟨ts'⟩)" "ts = ts' ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α"
using surj_pair[of a] a(2,3) by (cases "snd a") force+
obtain t where t: "t ∈ set ts'" "occurs k = t ⋅ ξ ∘⇩s σ ∘⇩s α"
using ts(2) unfolding ts'(2) by force
have occ_t: "OccursFact ∈ funs_term t"
proof (cases t)
case (Var y) thus ?thesis
using t(2) transaction_decl_fresh_renaming_substs_range'(1)[OF T(2-), of "occurs k"]
by fastforce
qed (use t(2) in simp)
have P_wf: "∀T ∈ set P. wellformed_transaction T"
using P admissible_transaction_is_wellformed_transaction(1) by blast
have l: "l = ⋆"
using wellformed_transaction_strand_memberD(8)[OF bspec[OF P_wf T(1)] a(1)[unfolded ts'(1)]]
t(1) T(1) P_occ_star occ_t
unfolding ts'(1) by fastforce
have "occurs k ∈ ⋃{set ts |ts. ⟨⋆, receive⟨ts⟩⟩ ∈ set (A ⋅⇩l⇩s⇩s⇩t I)}"
using subst_lsst_memI[OF ts(1), of I] subst_set_map[OF ts(2), of I]
unfolding occ_I l by auto
thus False using A(2) unfolding declassified⇩l⇩s⇩s⇩t_def by simp
qed
qed
lemma welltyped_leakage_free_setop_pairI:
assumes P:
"∀T ∈ set P. wellformed_transaction T"
"∀T ∈ set P. ∀x ∈ vars_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
"∀T ∈ set P. ∀f ∈ ⋃(funs_term ` (trms⇩l⇩s⇩s⇩t (transaction_send T))). ¬is_Set f"
"∀T ∈ set P. ∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value"
"∀T ∈ set P. transaction_decl T () = []"
"∀T ∈ set P. admissible_transaction_terms T"
and c: "arity⇩s c = 0"
shows "welltyped_leakage_free_protocol [pair (x, ⟨c⟩⇩s)] P"
proof -
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define f where "f ≡ λM. {t ⋅ δ | t δ. Q M t δ}"
define Sec where "Sec ≡ f (set [pair (x, ⟨c⟩⇩s)]) - {m. {} ⊢⇩c m}"
define f' where "f' ≡ λ(T,ξ,σ::('fun,'atom,'sets,'lbl) prot_subst,α).
dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define g' where "g' ≡ concat ∘ map f'"
note defs = Sec_def f_def Q_def
note defs' = welltyped_constraint_model_def constraint_model_def
have P':
"∀T ∈ set P. ∀x ∈ vars_transaction T ∪ set (transaction_fresh T). Γ⇩v x ≠ TAtom SetType"
"∀T ∈ set P. ∀x ∈ vars_transaction T. is_Var (Γ⇩v x)"
"∀T ∈ set P. ∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
subgoal using P(2,4) by fastforce
subgoal using P(2) by fastforce
subgoal using P(4) by fast
done
note 0 = welltyped_leakage_free_set_constI(1)[OF P(1,3) P' c]
show ?thesis
proof (rule ccontr)
assume "¬welltyped_leakage_free_protocol [pair (x, ⟨c⟩⇩s)] P"
then obtain A I k where A:
"A ∈ reachable_constraints P" "pair (k, ⟨c⟩⇩s) ∈ Sec - declassified⇩l⇩s⇩s⇩t A I"
"welltyped_constraint_model I (A@[⟨⋆, send⟨[pair (k, ⟨c⟩⇩s)]⟩⟩])"
unfolding welltyped_leakage_free_protocol_def defs pair_def by fastforce
note A' = welltyped_constraint_model_prefix[OF A(3)]
have "pair (k, ⟨c⟩⇩s) ⋅ I = pair (k, ⟨c⟩⇩s)" using A(2) unfolding defs by auto
hence "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ pair (k, ⟨c⟩⇩s)"
using welltyped_constraint_model_deduct_split(2)[OF A(3)] by argo
then obtain n where n: "intruder_deduct_num (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) n (pair (k, ⟨c⟩⇩s))"
using deduct_num_if_deduct by fast
have "wt⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩s (subst_range I)" "ik⇩l⇩s⇩s⇩t A ⊆ trms⇩l⇩s⇩s⇩t A"
using A(3) ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset unfolding defs' by simp_all
hence "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊆ SMP (trms⇩l⇩s⇩s⇩t A)" by blast
hence "Pair ∉ ⋃(funs_term ` (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I))"
using reachable_constraints_no_Pair_fun'[OF A(1) P(4-6)] P by blast
hence 1: "¬pair (k, ⟨c⟩⇩s) ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I"
using funs_term_Fun_subterm'[of Pair] unfolding pair_def by auto
have 2: "pair (k, ⟨c⟩⇩s) ∉ set (snd (Ana m))" when m: "m ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" for m
using m 1 term.dual_order.trans Ana_subterm'[of "pair (k, ⟨c⟩⇩s)" m] by auto
have "¬ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ ⟨c⟩⇩s"
using 0 A(1) A' welltyped_constraint_model_deduct_iff[of I A ⋆ "⟨c⟩⇩s"] by force
moreover have "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ ⟨c⟩⇩s"
using 1 2 deduct_inv[OF n] deduct_if_deduct_num[of "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" _ "⟨c⟩⇩s"]
unfolding pair_def by auto
ultimately show False by blast
qed
qed
lemma welltyped_leakage_free_short_term_secretI:
fixes c x y f n d l l'
defines "s ≡ ⟨f [⟨c⟩⇩c, ⟨x: value⟩⇩v]⟩⇩t"
and "Tatt ≡ Transaction (λ(). []) []
[⟨⋆, receive⟨[occurs ⟨y: value⟩⇩v]⟩⟩,
(l, receive⟨[⟨f [⟨c⟩⇩c, ⟨y: value⟩⇩v]⟩⇩t]⟩)]
[(l', ⟨⟨y: value⟩⇩v not in ⟨d⟩⇩s⟩)]
[]
[⟨n, send⟨[attack⟨ln n⟩]⟩⟩]"
assumes P:
"∀T ∈ set P. admissible_transaction' T"
"∀T ∈ set P. admissible_transaction_occurs_checks T"
"∀T ∈ set P. ∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
and subterms_sec:
"∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[⟨c⟩⇩c]⟩⟩])"
and P_sec:
"∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ.
welltyped_constraint_model ℐ⇩τ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
and P_Tatt: "Tatt ∈ set P"
and P_d:
"∀T ∈ set P. ∀a ∈ set (transaction_updates T).
is_Insert (snd a) ∧ the_set_term (snd a) = ⟨d⟩⇩s ⟶
transaction_send T ≠ [] ∧ (let (l,b) = hd (transaction_send T)
in l = ⋆ ∧ is_Send b ∧ ⟨f [⟨c⟩⇩c, the_elem_term (snd a)]⟩⇩t ∈ set (the_msgs b))"
shows "welltyped_leakage_free_protocol [s] P"
proof -
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define Sec where "Sec ≡ {t ⋅ δ | t δ. Q (set [s]) t δ} - {m. {} ⊢⇩c m}"
define f' where "f' ≡ λ(T,ξ,σ::('fun,'atom,'sets,'lbl) prot_subst,α).
dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define g' where "g' ≡ concat ∘ map f'"
note defs = Sec_def Q_def
note defs' = welltyped_constraint_model_def constraint_model_def
show ?thesis
proof (rule ccontr)
assume "¬welltyped_leakage_free_protocol [s] P"
then obtain A I k where A:
"A ∈ reachable_constraints P" "⟨f [⟨c⟩⇩c, k]⟩⇩t ∈ Sec - declassified⇩l⇩s⇩s⇩t A I"
"welltyped_constraint_model I (A@[⟨⋆, send⟨[⟨f [⟨c⟩⇩c, k]⟩⇩t]⟩⟩])"
unfolding welltyped_leakage_free_protocol_def defs s_def by fastforce
have I: "wt⇩s⇩u⇩b⇩s⇩t I" "interpretation⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩s (subst_range I)"
using A(3) unfolding defs' by (blast,blast,blast)
note A' = welltyped_constraint_model_prefix[OF A(3)]
have "strand_sem_stateful {} {} (unlabel A) I"
using A' unfolding defs' by simp
hence A'': "strand_sem_stateful {} {} (unlabel A) (I(z := k))"
when z: "z ∉ fv⇩l⇩s⇩s⇩t A" for z
using z strand_sem_model_swap[of "unlabel A" I "I(z := k)" "{}" "{}"] by auto
obtain δ where δ:
"δ (the_Var ⟨x: value⟩⇩v) = k" "wt⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)"
"fv (δ (the_Var ⟨x: value⟩⇩v)) = {}"
using A(2) unfolding defs s_def by auto
have k: "Γ k = TAtom Value" "fv k = {}" "wf⇩t⇩r⇩m k"
subgoal using δ(1) wt_subst_trm''[OF δ(2), of "⟨x: value⟩⇩v"] by simp
subgoal using δ(1,4) by blast
subgoal using δ(1,3) by force
done
then obtain fk where fk: "k = Fun fk []"
using const_type_inv_wf by (cases k) auto
have fk': "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ ⟨f [⟨c⟩⇩c, k]⟩⇩t"
using fk welltyped_constraint_model_deduct_split(2)[OF A(3)] by auto
have "¬welltyped_constraint_model I (A@[⟨⋆, send⟨[⟨c⟩⇩c]⟩⟩])"
using subterms_sec(1) A(1) by blast
hence "¬ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ ⟨c⟩⇩c"
using A' strand_sem_append_stateful[of "{}" "{}" "unlabel A" "unlabel [⟨⋆, send⟨[⟨c⟩⇩c]⟩⟩]" I]
unfolding defs' by auto
hence "⟨f [⟨c⟩⇩c, k]⟩⇩t ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I"
using fk' deduct_inv'[OF fk'] by force
moreover have "k ⊑ ⟨f [⟨c⟩⇩c, k]⟩⇩t" by simp
ultimately have k_in_ik: "k ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I"
using in_subterms_subset_Union by blast
hence "k ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t A ∨ (∃x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A). k ⊑ I x)"
using const_subterms_subst_cases[of fk I "ik⇩l⇩s⇩s⇩t A"]
unfolding fk by fast
hence "fk ∈ ⋃(funs_term ` ik⇩l⇩s⇩s⇩t A) ∨ (∃x ∈ fv⇩l⇩s⇩s⇩t A. k ⊑ I x)"
unfolding fk by (meson UN_iff funs_term_Fun_subterm' fv_ik⇩s⇩s⇩t_is_fv⇩s⇩s⇩t )
hence "fk ∈ ⋃(funs_term ` trms⇩l⇩s⇩s⇩t A) ∨ (∃x ∈ fv⇩l⇩s⇩s⇩t A. k ⊑ I x)"
using ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset by blast
moreover have "Γ⇩v x = TAtom Value" when x: "x ∈ fv⇩l⇩s⇩s⇩t A" for x
using x reachable_constraints_var_types_in_transactions(1)[OF A(1) P(3)]
P(1,2) admissible_transaction_Value_vars
by force
ultimately have "fk ∈ ⋃(funs_term ` trms⇩l⇩s⇩s⇩t A) ∨ (∃x ∈ fv⇩l⇩s⇩s⇩t A. Γ⇩v x = TAtom Value ∧ k ⊑ I x)"
by blast
hence "fk ∈ ⋃(funs_term ` trms⇩l⇩s⇩s⇩t A) ∨ (∃x ∈ fv⇩l⇩s⇩s⇩t A. Γ⇩v x = TAtom Value ∧ I x = k)"
using I(1,3) wf_trm_TComp_subterm wf_trm_subst_rangeD
unfolding fk wt⇩s⇩u⇩b⇩s⇩t_def by (metis Γ.simps(1) term.distinct(1))
then obtain kn where kn: "fk = Val kn"
using reachable_constraints_val_funs_private[OF A(1) P(1), of fk]
constraint_model_Value_term_is_Val[OF A(1) A' P(1,2)]
Fun_Value_type_inv[OF k(1)[unfolded fk]]
unfolding fk is_PubConstValue_def by (cases fk) force+
obtain α::"('fun,'atom,'sets,'lbl) prot_subst"
where α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)"
unfolding transaction_renaming_subst_def by blast
obtain y' yn' where y':
"α (the_Var ⟨y: value⟩⇩v) = Var y'" "y ≠ yn'" "Var y' = ⟨yn': value⟩⇩v"
using transaction_renaming_subst_is_renaming(1,2)[OF α] by force
define B where "B ≡ A@dual⇩l⇩s⇩s⇩t (transaction_strand Tatt ⋅⇩l⇩s⇩s⇩t α)"
define J where "J ≡ I(y' := k)"
have "y' ∈ range_vars α"
using y'(1) transaction_renaming_subst_is_renaming(3)[OF α]
by (metis (no_types, lifting) in_mono subst_fv_imgI term.set_intros(3))
hence y'': "y' ∉ vars⇩l⇩s⇩s⇩t A"
using transaction_renaming_subst_vars_disj(6)[OF α] by blast
have 0: "(k,⟨d⟩⇩s) ∉ set (db⇩l⇩s⇩s⇩t A I)"
proof
assume a: "(k,⟨d⟩⇩s) ∈ set (db⇩l⇩s⇩s⇩t A I)"
obtain l t t' where t: "(l,insert⟨t,t'⟩) ∈ set A" "t ⋅ I = k" "t' ⋅ I = ⟨d⟩⇩s"
using db⇩s⇩s⇩t_in_cases[OF a[unfolded db⇩s⇩s⇩t_def]] unfolding unlabel_def by auto
obtain T b B α σ ξ where T:
"prefix (B@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) A"
"T ∈ set P" "transaction_decl_subst ξ T"
"transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t B)"
"b ∈ set (transaction_strand T)" "(l, insert⟨t,t'⟩) = dual⇩l⇩s⇩s⇩t⇩p b ⋅⇩l⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩s α"
"fst (l, insert⟨t,t'⟩) = fst b"
using reachable_constraints_transaction_action_obtain[OF A(1) t(1)] by metis
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
obtain b' where "b = (l,b')"
using T(8) by (cases b) simp
then obtain tb tb'
where b': "b = (l,insert⟨tb,tb'⟩)"
and tb: "t = tb ⋅ θ"
and tb': "t' = tb' ⋅ θ"
using T(7) unfolding θ_def by (cases b') auto
note T_adm = bspec[OF P(1) T(2)]
note T_wf = admissible_transaction_is_wellformed_transaction(1,3)[OF T_adm]
have b: "b ∈ set (transaction_updates T)"
using transaction_strand_memberD[OF T(6)[unfolded b']]
wellformed_transaction_cases[OF T_wf(1)]
unfolding b' by blast
have "∃n. tb = ⟨n: value⟩⇩v" and *: "tb' = ⟨d⟩⇩s"
using tb tb' T(6) t(3) transaction_inserts_are_Value_vars[OF T_wf, of tb tb']
unfolding b' unlabel_def by (force,force)
have "is_Insert (snd b)" "the_set_term (snd b) = ⟨d⟩⇩s" "the_elem_term (snd b) = tb"
unfolding b' * by simp_all
hence "transaction_send T ≠ []"
"let (l, a) = hd (transaction_send T)
in l = ⋆ ∧ is_Send a ∧ ⟨f [⟨c⟩⇩c, tb]⟩⇩t ∈ set (the_msgs a)"
using P_d T(2) b by (fast,fast)
hence "∃ts. ⟨⋆,send⟨ts⟩⟩ ∈ set (transaction_send T) ∧ ⟨f [⟨c⟩⇩c, tb]⟩⇩t ∈ set ts"
unfolding is_Send_def by (cases "transaction_send T") auto
then obtain ts where ts: "⟨⋆,send⟨ts⟩⟩ ∈ set (transaction_strand T)" "⟨f [⟨c⟩⇩c, tb]⟩⇩t ∈ set ts"
unfolding transaction_strand_def by auto
have "⟨⋆,receive⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩⟩ ∈ set A" "⟨f [⟨c⟩⇩c, t]⟩⇩t ∈ set (ts ⋅⇩l⇩i⇩s⇩t θ)"
using subst_lsst_memI[OF ts(1), of θ] subst_set_map[OF ts(2), of θ]
dual⇩l⇩s⇩s⇩t_steps_iff(1)[of ⋆ "ts ⋅⇩l⇩i⇩s⇩t θ" "transaction_strand T ⋅⇩l⇩s⇩s⇩t θ"]
set_mono_prefix[OF T(1)[unfolded θ_def[symmetric]]]
unfolding tb by auto
hence "⟨f [⟨c⟩⇩c, k]⟩⇩t ∈ ⋃{set ts |ts. ⟨⋆, receive⟨ts⟩⟩ ∈ set (A ⋅⇩l⇩s⇩s⇩t I)}"
using t(2) subst_lsst_memI[of "⟨⋆,receive⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩⟩" A I] by force
thus False
using A(2) unfolding declassified⇩l⇩s⇩s⇩t_def by auto
qed
have "y' ∉ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A)"
using y'' fv_ik_subset_vars_sst'[of "unlabel A"] by blast
hence 1: "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I = ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J"
unfolding J_def by (metis (no_types, lifting) fv_subset image_cong in_mono repl_invariance)
have "(k,⟨d⟩⇩s) ∉ dbupd⇩s⇩s⇩t (unlabel A) I {}"
using 0 db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel A" I "[]"]
unfolding db⇩s⇩s⇩t_def by force
hence "(k,⟨d⟩⇩s) ∉ dbupd⇩s⇩s⇩t (unlabel A) J {}"
using y'' vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel A"]
db⇩s⇩s⇩t_subst_swap[of "unlabel A" I J "[]"]
db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel A" _ "[]"]
unfolding db⇩s⇩s⇩t_def J_def by (metis (no_types, lifting) Un_iff empty_set fun_upd_other)
hence "((Var y' ⋅ J, ⟨d⟩⇩s ⋅ J) ∉ dbupd⇩s⇩s⇩t (unlabel A) J {})"
unfolding J_def fk by simp
hence "strand_sem_stateful (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J) (dbupd⇩s⇩s⇩t (unlabel A) J {})
(unlabel [⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩]) J"
using stateful_strand_sem_NegChecks_no_bvars(1)[
of "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J" "dbupd⇩s⇩s⇩t (unlabel A) J {}" "Var y'" "⟨d⟩⇩s" J]
by simp
hence 2: "strand_sem_stateful {} {} (unlabel (A@[⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩])) J"
using A'' y' y'' vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel A"]
strand_sem_append_stateful[
of "{}" "{}" "unlabel A" "unlabel [⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩]" J]
unfolding J_def by simp
have B: "B ∈ reachable_constraints P"
using reachable_constraints.step[OF A(1) P_Tatt _ _ α, of Var Var]
unfolding B_def Tatt_def transaction_decl_subst_def transaction_fresh_subst_def by simp
have Tatt': "dual⇩l⇩s⇩s⇩t (transaction_strand Tatt ⋅⇩l⇩s⇩s⇩t α) =
[⟨⋆, send⟨[occurs (Var y')]⟩⟩,
(l, send⟨[⟨f [⟨c⟩⇩c, Var y']⟩⇩t]⟩),
(l', ⟨Var y' not in ⟨d⟩⇩s⟩),
⟨n, receive⟨[attack⟨ln n⟩]⟩⟩]"
using y'
unfolding Tatt_def transaction_strand_def dual⇩l⇩s⇩s⇩t_def subst_apply_labeled_stateful_strand_def
by auto
have J: "wt⇩s⇩u⇩b⇩s⇩t J" "interpretation⇩s⇩u⇩b⇩s⇩t J" "wf⇩t⇩r⇩m⇩s (subst_range J)"
unfolding J_def
subgoal using wt_subst_subst_upd[OF I(1)] k(1) y'(3) by simp
subgoal by (metis I(2) k(2) fun_upd_apply interpretation_grounds_all interpretation_substI)
subgoal using I(3) k(3) by fastforce
done
have 3: "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J ⊢ ⟨f [⟨c⟩⇩c, Var y']⟩⇩t ⋅ J"
using 1 fk fk' unfolding J_def by auto
have 4: "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J ⊢ occurs (Var y') ⋅ J"
using reachable_constraints_occurs_fact_ik_case'[
OF A(1) P(1,2) constraint_model_Val_const_in_constr_prefix'[
OF A(1) A' P(1) k_in_ik[unfolded fk kn]]]
intruder_deduct.Axiom[of "occurs k" "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J"]
unfolding J_def fk kn by fastforce
have "strand_sem_stateful {} {} (unlabel A) J"
using 2 strand_sem_append_stateful by force
hence "strand_sem_stateful {} {}
(unlabel (A@[⟨⋆, send⟨[occurs (Var y')]⟩⟩,
⟨n, send⟨[⟨f [⟨c⟩⇩c, Var y']⟩⇩t]⟩⟩,
⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩])) J"
using 2 3 4 strand_sem_append_stateful[of "{}" "{}" _ _ J]
unfolding unlabel_def ik⇩s⇩s⇩t_def by force
hence "strand_sem_stateful {} {} (unlabel (B@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])) J"
using strand_sem_receive_send_append[of "{}" "{}" _ J "attack⟨ln n⟩"]
strand_sem_append_stateful[of "{}" "{}" _ _ J]
unfolding B_def Tatt' by simp
hence "welltyped_constraint_model J (B@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
using B J unfolding defs' by blast
thus False using B(1) P_sec by blast
qed
qed
lemma welltyped_leakage_free_short_term_secretI':
fixes c x y f n d l l' τ
defines "s ≡ ⟨f [⟨c⟩⇩c, Var (TAtom τ,x)]⟩⇩t"
and "Tatt ≡ Transaction (λ(). []) []
[(l, receive⟨[⟨f [⟨c⟩⇩c, Var (TAtom τ,y)]⟩⇩t]⟩)]
[(l', ⟨Var (TAtom τ,y) not in ⟨d⟩⇩s⟩)]
[]
[⟨n, send⟨[attack⟨ln n⟩]⟩⟩]"
assumes P:
"∀T ∈ set P. wellformed_transaction T"
"∀T ∈ set P. ∀x ∈ set (unlabel (transaction_updates T)).
is_Update x ⟶ is_Fun (the_set_term x)"
and subterms_sec:
"∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[⟨⋆, send⟨[⟨c⟩⇩c]⟩⟩])"
and P_sec:
"∀𝒜 ∈ reachable_constraints P. ∄ℐ⇩τ.
welltyped_constraint_model ℐ⇩τ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
and P_Tatt: "Tatt ∈ set P"
and P_d:
"∀T ∈ set P. ∀a ∈ set (transaction_updates T).
is_Insert (snd a) ∧ the_set_term (snd a) = ⟨d⟩⇩s ⟶
transaction_send T ≠ [] ∧ (let (l,b) = hd (transaction_send T)
in l = ⋆ ∧ is_Send b ∧ ⟨f [⟨c⟩⇩c, the_elem_term (snd a)]⟩⇩t ∈ set (the_msgs b))"
shows "welltyped_leakage_free_protocol [s] P"
proof -
define Q where "Q ≡ λM t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}"
define Sec where "Sec ≡ {t ⋅ δ | t δ. Q (set [s]) t δ} - {m. {} ⊢⇩c m}"
define f' where "f' ≡ λ(T,ξ,σ::('fun,'atom,'sets,'lbl) prot_subst,α).
dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)"
define g' where "g' ≡ concat ∘ map f'"
note defs = Sec_def Q_def
note defs' = welltyped_constraint_model_def constraint_model_def
show ?thesis
proof (rule ccontr)
assume "¬welltyped_leakage_free_protocol [s] P"
then obtain A I k where A:
"A ∈ reachable_constraints P" "⟨f [⟨c⟩⇩c, k]⟩⇩t ∈ Sec - declassified⇩l⇩s⇩s⇩t A I"
"welltyped_constraint_model I (A@[⟨⋆, send⟨[⟨f [⟨c⟩⇩c, k]⟩⇩t]⟩⟩])"
unfolding welltyped_leakage_free_protocol_def defs s_def by fastforce
have I: "wt⇩s⇩u⇩b⇩s⇩t I" "interpretation⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩s (subst_range I)"
using A(3) unfolding defs' by (blast,blast,blast)
note A' = welltyped_constraint_model_prefix[OF A(3)]
have "strand_sem_stateful {} {} (unlabel A) I"
using A' unfolding defs' by simp
hence A'': "strand_sem_stateful {} {} (unlabel A) (I(z := k))"
when z: "z ∉ fv⇩l⇩s⇩s⇩t A" for z
using z strand_sem_model_swap[of "unlabel A" I "I(z := k)" "{}" "{}"] by auto
obtain δ where δ:
"δ (TAtom τ,x) = k" "wt⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" "fv (δ (TAtom τ,x)) = {}"
using A(2) unfolding defs s_def by auto
have k: "Γ k = TAtom τ" "fv k = {}" "wf⇩t⇩r⇩m k"
subgoal using δ(1) wt_subst_trm''[OF δ(2), of "Var (TAtom τ,x)"] by simp
subgoal using δ(1,4) by blast
subgoal using δ(1,3) by force
done
then obtain fk where fk: "k = Fun fk []"
using const_type_inv_wf by (cases k) auto
have fk': "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ ⟨f [⟨c⟩⇩c, k]⟩⇩t"
using fk welltyped_constraint_model_deduct_split(2)[OF A(3)] by auto
obtain α::"('fun,'atom,'sets,'lbl) prot_subst"
where α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)"
unfolding transaction_renaming_subst_def by blast
obtain y' yn' where y': "α (TAtom τ,y) = Var y'" "y ≠ yn'" "y' = (TAtom τ,yn')"
using transaction_renaming_subst_is_renaming(1,2)[OF α] by force
define B where "B ≡ A@dual⇩l⇩s⇩s⇩t (transaction_strand Tatt ⋅⇩l⇩s⇩s⇩t α)"
define J where "J ≡ I(y' := k)"
have "y' ∈ range_vars α"
using y'(1) transaction_renaming_subst_is_renaming(3)[OF α]
by (metis (no_types, lifting) in_mono subst_fv_imgI term.set_intros(3))
hence y'': "y' ∉ vars⇩l⇩s⇩s⇩t A"
using transaction_renaming_subst_vars_disj(6)[OF α] by blast
have 0: "(k,⟨d⟩⇩s) ∉ set (db⇩l⇩s⇩s⇩t A I)"
proof
assume a: "(k,⟨d⟩⇩s) ∈ set (db⇩l⇩s⇩s⇩t A I)"
obtain l t t' where t: "(l,insert⟨t,t'⟩) ∈ set A" "t ⋅ I = k" "t' ⋅ I = ⟨d⟩⇩s"
using db⇩s⇩s⇩t_in_cases[OF a[unfolded db⇩s⇩s⇩t_def]] unfolding unlabel_def by auto
obtain T b B α σ ξ where T:
"prefix (B@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) A"
"T ∈ set P" "transaction_decl_subst ξ T"
"transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t B)"
"b ∈ set (transaction_strand T)" "(l, insert⟨t,t'⟩) = dual⇩l⇩s⇩s⇩t⇩p b ⋅⇩l⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩s α"
"fst (l, insert⟨t,t'⟩) = fst b"
using reachable_constraints_transaction_action_obtain[OF A(1) t(1)] by metis
define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α"
obtain b' where "b = (l,b')"
using T(8) by (cases b) simp
then obtain tb tb'
where b': "b = (l,insert⟨tb,tb'⟩)"
and tb: "t = tb ⋅ θ"
and tb': "t' = tb' ⋅ θ"
using T(7) unfolding θ_def by (cases b') auto
note T_wf = bspec[OF P(1) T(2)] bspec[OF P(2) T(2)]
have b: "b ∈ set (transaction_updates T)"
using transaction_strand_memberD[OF T(6)[unfolded b']]
wellformed_transaction_cases[OF T_wf(1)]
unfolding b' by blast
have "is_Fun tb'"
using bspec[OF P(2) T(2)]
wellformed_transaction_strand_unlabel_memberD(6)[
OF T_wf(1) unlabel_in[OF T(6)[unfolded b']]]
by fastforce
hence *: "tb' = ⟨d⟩⇩s"
using t(3) unfolding b' tb' by force
have "is_Insert (snd b)" "the_set_term (snd b) = ⟨d⟩⇩s" "the_elem_term (snd b) = tb"
unfolding b' * by simp_all
hence "transaction_send T ≠ []"
"let (l, a) = hd (transaction_send T)
in l = ⋆ ∧ is_Send a ∧ ⟨f [⟨c⟩⇩c, tb]⟩⇩t ∈ set (the_msgs a)"
using P_d T(2) b by (fast,fast)
hence "∃ts. ⟨⋆,send⟨ts⟩⟩ ∈ set (transaction_send T) ∧ ⟨f [⟨c⟩⇩c, tb]⟩⇩t ∈ set ts"
unfolding is_Send_def by (cases "transaction_send T") auto
then obtain ts where ts: "⟨⋆,send⟨ts⟩⟩ ∈ set (transaction_strand T)" "⟨f [⟨c⟩⇩c, tb]⟩⇩t ∈ set ts"
unfolding transaction_strand_def by auto
have "⟨⋆,receive⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩⟩ ∈ set A" "⟨f [⟨c⟩⇩c, t]⟩⇩t ∈ set (ts ⋅⇩l⇩i⇩s⇩t θ)"
using subst_lsst_memI[OF ts(1), of θ] subst_set_map[OF ts(2), of θ]
dual⇩l⇩s⇩s⇩t_steps_iff(1)[of ⋆ "ts ⋅⇩l⇩i⇩s⇩t θ" "transaction_strand T ⋅⇩l⇩s⇩s⇩t θ"]
set_mono_prefix[OF T(1)[unfolded θ_def[symmetric]]]
unfolding tb by auto
hence "⟨f [⟨c⟩⇩c, k]⟩⇩t ∈ ⋃{set ts |ts. ⟨⋆, receive⟨ts⟩⟩ ∈ set (A ⋅⇩l⇩s⇩s⇩t I)}"
using t(2) subst_lsst_memI[of "⟨⋆,receive⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩⟩" A I] by force
thus False
using A(2) unfolding declassified⇩l⇩s⇩s⇩t_def by auto
qed
have "y' ∉ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A)"
using y'' fv_ik_subset_vars_sst'[of "unlabel A"] by blast
hence 1: "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I = ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J"
unfolding J_def by (metis (no_types, lifting) fv_subset image_cong in_mono repl_invariance)
have "(k,⟨d⟩⇩s) ∉ dbupd⇩s⇩s⇩t (unlabel A) I {}"
using 0 db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel A" I "[]"]
unfolding db⇩s⇩s⇩t_def by force
hence "(k,⟨d⟩⇩s) ∉ dbupd⇩s⇩s⇩t (unlabel A) J {}"
using y'' vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel A"]
db⇩s⇩s⇩t_subst_swap[of "unlabel A" I J "[]"]
db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel A" _ "[]"]
unfolding db⇩s⇩s⇩t_def J_def by (metis (no_types, lifting) Un_iff empty_set fun_upd_other)
hence "((Var y' ⋅ J, ⟨d⟩⇩s ⋅ J) ∉ dbupd⇩s⇩s⇩t (unlabel A) J {})"
unfolding J_def fk by simp
hence "strand_sem_stateful (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J) (dbupd⇩s⇩s⇩t (unlabel A) J {})
(unlabel [⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩]) J"
using stateful_strand_sem_NegChecks_no_bvars(1)[
of "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J" "dbupd⇩s⇩s⇩t (unlabel A) J {}" "Var y'" "⟨d⟩⇩s" J]
by simp
hence 2: "strand_sem_stateful {} {} (unlabel (A@[⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩])) J"
using A'' y' y'' vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel A"]
strand_sem_append_stateful[
of "{}" "{}" "unlabel A" "unlabel [⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩]" J]
unfolding J_def by simp
have B: "B ∈ reachable_constraints P"
using reachable_constraints.step[OF A(1) P_Tatt _ _ α, of Var Var]
unfolding B_def Tatt_def transaction_decl_subst_def transaction_fresh_subst_def by simp
have Tatt': "dual⇩l⇩s⇩s⇩t (transaction_strand Tatt ⋅⇩l⇩s⇩s⇩t α) =
[(l, send⟨[⟨f [⟨c⟩⇩c, Var y']⟩⇩t]⟩),
(l', ⟨Var y' not in ⟨d⟩⇩s⟩),
⟨n, receive⟨[attack⟨ln n⟩]⟩⟩]"
using y'
unfolding Tatt_def transaction_strand_def dual⇩l⇩s⇩s⇩t_def subst_apply_labeled_stateful_strand_def
by auto
have J: "wt⇩s⇩u⇩b⇩s⇩t J" "interpretation⇩s⇩u⇩b⇩s⇩t J" "wf⇩t⇩r⇩m⇩s (subst_range J)"
unfolding J_def
subgoal using wt_subst_subst_upd[OF I(1)] k(1) y'(3) by simp
subgoal by (metis I(2) k(2) fun_upd_apply interpretation_grounds_all interpretation_substI)
subgoal using I(3) k(3) by fastforce
done
have 3: "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t J ⊢ ⟨f [⟨c⟩⇩c, Var y']⟩⇩t ⋅ J"
using 1 fk fk' unfolding J_def by auto
have "strand_sem_stateful {} {} (unlabel A) J"
using 2 strand_sem_append_stateful by force
hence "strand_sem_stateful {} {}
(unlabel (A@[⟨n, send⟨[⟨f [⟨c⟩⇩c, Var y']⟩⇩t]⟩⟩,
⟨n, ⟨Var y' not in ⟨d⟩⇩s⟩⟩])) J"
using 2 3 strand_sem_append_stateful[of "{}" "{}" _ _ J]
unfolding unlabel_def ik⇩s⇩s⇩t_def by force
hence "strand_sem_stateful {} {} (unlabel (B@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])) J"
using strand_sem_receive_send_append[of "{}" "{}" _ J "attack⟨ln n⟩"]
strand_sem_append_stateful[of "{}" "{}" _ _ J]
unfolding B_def Tatt' by simp
hence "welltyped_constraint_model J (B@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
using B J unfolding defs' by blast
thus False using B(1) P_sec by blast
qed
qed
definition welltyped_leakage_free_invkey_conditions' where
"welltyped_leakage_free_invkey_conditions' invfun privfunsec declassifiedset S n P ≡
let f = λs. is_Var s ∧ fst (the_Var s) = TAtom Value;
g = λs. is_Fun s ∧ args s = [] ∧ is_Set (the_Fun s) ∧
arity⇩s (the_Set (the_Fun s)) = 0;
h = λs. is_Fun s ∧ args s = [] ∧ is_Fu (the_Fun s) ∧
public⇩f (the_Fu (the_Fun s)) ∧ arity⇩f (the_Fu (the_Fun s)) = 0
in (∀s∈set S.
f s ∨
(is_Fun s ∧ the_Fun s = Pair ∧ length (args s) = 2 ∧ g (args s ! 1)) ∨
g s ∨ h s ∨ s = ⟨privfunsec⟩⇩c ∨ s = Fun OccursSec [] ∨
(is_Fun s ∧ the_Fun s = OccursFact ∧ length (args s) = 2 ∧
args s ! 0 = Fun OccursSec []) ∨
(is_Fun s ∧ the_Fun s = Fu invfun ∧ length (args s) = 2 ∧
args s ! 0 = ⟨privfunsec⟩⇩c ∧ f (args s ! 1)) ∨
(is_Fun s ∧ is_Fu (the_Fun s) ∧ fv s = {} ∧
Transaction (λ(). []) [] [⟨n, receive⟨[s]⟩⟩] [] [] [⟨n, send⟨[attack⟨ln n⟩]⟩⟩]∈set P)) ∧
(¬public⇩f privfunsec ∧ arity⇩f privfunsec = 0 ∧ Γ⇩f privfunsec ≠ None) ∧
(∀T∈set P. transaction_fresh T ≠ [] ⟶
transaction_send T ≠ [] ∧
(let (l, a) = hd (transaction_send T)
in l = ⋆ ∧ is_Send a ∧ Var ` set (transaction_fresh T) ⊆ set (the_msgs a))) ∧
(∀T∈set P. ∀x∈vars_transaction T. is_Var (Γ⇩v x)) ∧
(∀T∈set P. ∀x∈set (transaction_fresh T). Γ⇩v x = Var Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)) ∧
(∀T∈set P. ∀f∈⋃(funs_term ` trms⇩l⇩s⇩s⇩t (transaction_send T)). ¬is_Set f) ∧
(∀T∈set P. ∀r∈set (transaction_send T).
OccursFact ∈ ⋃(funs_term ` trms⇩s⇩s⇩t⇩p (snd r)) ⟶ has_LabelS r) ∧
(∀T∈set P. ∀t∈subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)).
⟨privfunsec⟩⇩c ∉ set (snd (Ana t))) ∧
(∀T∈set P. ⟨privfunsec⟩⇩c ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)) ∧
(∀T∈set P. ∀a∈set (transaction_updates T).
is_Insert (snd a) ∧ the_set_term (snd a) = ⟨declassifiedset⟩⇩s ⟶
transaction_send T ≠ [] ∧
(let (l, b) = hd (transaction_send T)
in l = ⋆ ∧ is_Send b ∧
⟨invfun [⟨privfunsec⟩⇩c, the_elem_term (snd a)]⟩⇩t ∈ set (the_msgs b)))"
definition welltyped_leakage_free_invkey_conditions where
"welltyped_leakage_free_invkey_conditions invfun privfunsec declassifiedset S n P ≡
let Tatt = λR. Transaction (λ(). []) []
(R@[⟨n, receive⟨[⟨invfun [⟨privfunsec⟩⇩c, ⟨0: value⟩⇩v]⟩⇩t]⟩⟩])
[⟨⋆, ⟨⟨0: value⟩⇩v not in ⟨declassifiedset⟩⇩s⟩⟩]
[]
[⟨n, send⟨[attack⟨ln n⟩]⟩⟩]
in welltyped_leakage_free_invkey_conditions' invfun privfunsec declassifiedset S n P ∧
has_initial_value_producing_transaction P ∧
(if Tatt [⟨⋆, receive⟨[occurs ⟨0: value⟩⇩v]⟩⟩] ∈ set P
then ∀T∈set P. admissible_transaction' T ∧ admissible_transaction_occurs_checks T
else Tatt [] ∈ set P ∧
(∀T∈set P. wellformed_transaction T) ∧
(∀T∈set P. admissible_transaction_terms T) ∧
(∀T∈set P. bvars_transaction T = {}) ∧
(∀T∈set P. transaction_decl T () = []) ∧
(∀T∈set P. ∀x∈set (transaction_fresh T). let τ = fst x
in τ = TAtom Value ∧ τ ≠ Γ ⟨privfunsec⟩⇩c) ∧
(∀T∈set P. ∀x∈vars_transaction T. let τ = fst x
in is_Var τ ∧ (the_Var τ = Value ∨ is_Atom (the_Var τ)) ∧ τ ≠ Γ ⟨privfunsec⟩⇩c) ∧
(∀T∈set P. ∀t∈subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)).
Fun OccursSec [] ∉ set (snd (Ana t))) ∧
(∀T∈set P. Fun OccursSec [] ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)) ∧
(∀T∈set P. ∀x∈set (unlabel (transaction_updates T)).
is_Update x ⟶ is_Fun (the_set_term x)) ∧
(∀s∈set S. is_Fun s ⟶ the_Fun s ≠ OccursFact))"
lemma welltyped_leakage_free_invkeyI:
assumes P_wt_secure: "∀𝒜 ∈ reachable_constraints P.
∄ℐ. welltyped_constraint_model ℐ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
and a: "welltyped_leakage_free_invkey_conditions invfun privsec declassset S n P"
shows "welltyped_leakage_free_protocol S P"
proof -
let ?Tatt' = "λR C. Transaction (λ(). []) [] R C [] [⟨n, send⟨[attack⟨ln n⟩]⟩⟩]
::('fun,'atom,'sets,'lbl) prot_transaction"
let ?Tatt = "λR. ?Tatt' (R@[⟨n, receive⟨[⟨invfun [⟨privsec⟩⇩c, ⟨0: value⟩⇩v]⟩⇩t]⟩⟩])
[⟨⋆, ⟨⟨0: value⟩⇩v not in ⟨declassset⟩⇩s⟩⟩]"
define Tatt1 where "Tatt1 ≡ ?Tatt [⟨⋆, receive⟨[occurs ⟨0: value⟩⇩v]⟩⟩]"
define Tatt2 where "Tatt2 ≡ ?Tatt []"
define Tatt_lts where "Tatt_lts ≡ λs. ?Tatt' [⟨n, receive⟨[s]⟩⟩] []"
note defs = welltyped_leakage_free_invkey_conditions_def Let_def
note defs' = defs welltyped_leakage_free_invkey_conditions'_def
note Tatts = Tatt1_def Tatt2_def Tatt_lts_def
obtain at where 0: "¬public⇩f privsec" "arity⇩f privsec = 0" "Γ⇩f privsec = Some at"
using a unfolding defs' by fast
have *: "∀T∈set P. admissible_transaction' T" "∀T∈set P. admissible_transaction_occurs_checks T"
when "Tatt1 ∈ set P"
using a that unfolding defs Tatts by (meson,meson)
have **: "Tatt1 ∈ set P ∨ Tatt2 ∈ set P" using a unfolding defs Tatts by argo
have ***:
"∀T∈set P. ∀x∈set (transaction_fresh T). Γ⇩v x = TAtom Value ∧ Γ⇩v x ≠ Γ ⟨privsec⟩⇩c"
"∀T∈set P. ∀x∈vars_transaction T. ∃a. Γ⇩v x = TAtom a ∧
(a = Value ∨ (∃b. a = Atom b)) ∧ TAtom a ≠ Γ ⟨privsec⟩⇩c"
when "Tatt1 ∉ set P"
subgoal using a that Γ⇩v_TAtom''(2) unfolding defs Tatts by metis
subgoal
using a that Γ⇩v_TAtom''(1,2)
unfolding defs Tatts[symmetric] is_Atom_def is_Var_def by fastforce
done
have ****: "s ≠ occurs x"
when "Tatt1 ∉ set P" "s ∈ set S" for s x
using a that ** unfolding defs Tatts the_Fun_def by fastforce
have 1:
"∀T∈set P. transaction_fresh T ≠ [] ⟶
transaction_send T ≠ [] ∧
(let (l, a) = hd (transaction_send T)
in l = ⋆ ∧ is_Send a ∧ Var ` set (transaction_fresh T) ⊆ set (the_msgs a))"
"∀T∈set P. ∀x∈vars_transaction T. is_Var (Γ⇩v x)"
"∀T∈set P. ∀x∈set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
"∀T∈set P. ∀f∈⋃(funs_term ` trms⇩l⇩s⇩s⇩t (transaction_send T)). ¬is_Set f"
"∀T∈set P. ∀r∈set (transaction_send T).
OccursFact ∈ ⋃(funs_term ` trms⇩s⇩s⇩t⇩p (snd r)) ⟶ has_LabelS r"
"∀T∈set P. ∀t∈subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). ⟨privsec⟩⇩c ∉ set (snd (Ana t))"
"∀T∈set P. ⟨privsec⟩⇩c ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)"
"∀T∈set P. ∀a∈set (transaction_updates T).
is_Insert (snd a) ∧ the_set_term (snd a) = ⟨declassset⟩⇩s ⟶
transaction_send T ≠ [] ∧
(let (l, b) = hd (transaction_send T)
in l = ⋆ ∧ is_Send b ∧
⟨invfun [⟨privsec⟩⇩c, the_elem_term (snd a)]⟩⇩t ∈ set (the_msgs b))"
using a unfolding defs' by (meson,meson,meson,meson,meson,meson,meson,meson)
have 2:
"∀T∈set P. wellformed_transaction T"
"∀T∈set P. ∀x∈vars_transaction T ∪ set (transaction_fresh T). Γ ⟨privsec⟩⇩c ≠ Γ⇩v x"
"∀T∈set P. admissible_transaction_terms T"
"∀T∈set P. ∀x∈set (transaction_fresh T). Γ⇩v x = TAtom Value"
"∀T∈set P. transaction_decl T () = []"
"∀T∈set P. ∀x∈vars_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)"
"∀T∈set P. ∀x∈vars_transaction T ∪ set (transaction_fresh T). Γ⇩v x ≠ TAtom SetType"
"∀T∈set P. ∀x∈vars_transaction T ∪ set (transaction_fresh T). Γ⇩v x ≠ TAtom OccursSecType"
"∀T∈set P. ∀t∈subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). Fun OccursSec [] ∉ set (snd (Ana t))"
"∀T∈set P. Fun OccursSec [] ∉ trms⇩l⇩s⇩s⇩t (transaction_send T)"
"∀T∈set P. bvars_transaction T = {}"
"∀T∈set P. ∀x∈set (unlabel (transaction_updates T)). is_Update x ⟶ is_Fun (the_set_term x)"
subgoal using a * unfolding defs by (metis admissible_transaction_is_wellformed_transaction(1))
subgoal
apply (cases "Tatt1 ∈ set P")
subgoal using a * admissible_transactionE(2,3) Γ_Fu_simps(4) unfolding defs Tatts by force
subgoal using a Γ_Fu_simps(4) unfolding defs Tatts by fastforce
done
subgoal using a * admissible_transaction_is_wellformed_transaction(4) unfolding defs by metis
subgoal using a * admissible_transactionE(2) unfolding defs Tatts[symmetric] by fastforce
subgoal using a * admissible_transactionE(1) unfolding defs Tatts[symmetric] by metis
subgoal using * *** admissible_transactionE(3) by fast
subgoal using * *** admissible_transactionE(2,3) by (cases "Tatt1 ∈ set P") (force, fastforce)
subgoal using * *** admissible_transactionE(2,3) by (cases "Tatt1 ∈ set P") (force, fastforce)
subgoal using a * admissible_transaction_occurs_checksE6 unfolding defs by metis
subgoal using a * admissible_transaction_occurs_checksE5 unfolding defs by metis
subgoal
using a * admissible_transaction_no_bvars(2)
unfolding defs Tatts[symmetric] by fastforce
subgoal
using a * admissible_transaction_is_wellformed_transaction(3)
unfolding defs Tatts[symmetric] admissible_transaction_updates_def by fastforce
done
have Tatt_lts_case:
"∃θ. wt⇩s⇩u⇩b⇩s⇩t θ ∧ inj_on θ (fv s) ∧ θ ` fv s ⊆ range Var ∧
?Tatt' [⟨n, receive⟨[s ⋅ θ]⟩⟩] [] ∈ set P"
when s: "fv s = {}" "Tatt_lts s ∈ set P" for s
proof -
have "wt⇩s⇩u⇩b⇩s⇩t Var" "inj_on Var (fv s)" "Var ` fv s ⊆ range Var" "s ⋅ Var = s"
using s(1) by simp_all
thus ?thesis using s(2) unfolding Tatts by metis
qed
have Tatt1_case:
"?Tatt' [⟨⋆, receive⟨[occurs ⟨0: value⟩⇩v]⟩⟩, ⟨n, receive⟨[⟨invfun [⟨privsec⟩⇩c, ⟨0: value⟩⇩v]⟩⇩t]⟩⟩]
[⟨⋆, ⟨⟨0: value⟩⇩v not in ⟨declassset⟩⇩s⟩⟩] ∈ set P"
when "Tatt1 ∈ set P"
using that unfolding Tatts by auto
have Tatt2_case:
"?Tatt' [⟨n, receive⟨[⟨invfun [⟨privsec⟩⇩c, ⟨0: value⟩⇩v]⟩⇩t]⟩⟩]
[⟨⋆, ⟨⟨0: value⟩⇩v not in ⟨declassset⟩⇩s⟩⟩] ∈ set P"
when "Tatt2 ∈ set P"
using that unfolding Tatts by auto
note 3 = pair_def case_prod_conv
note 4 = welltyped_leakage_free_priv_constI'[OF 0(1-3) 2(1,2) 1(2,3,6,7)]
note 5 = welltyped_leakage_free_setop_pairI[OF 2(1,6) 1(4) 2(4,5,3), unfolded 3]
welltyped_leakage_free_set_constI(2)[OF 2(1) 1(4) 2(7) 1(2,3), unfolded 3]
welltyped_leakage_free_pub_constI
4(2)
welltyped_leakage_free_occurssec_constI(2)[OF 2(1,8-10) 1(2,3)]
welltyped_leakage_free_value_constI[OF 2(1,3,5,11) 1(1)]
welltyped_leakage_free_short_term_secretI'[
OF 2(1,12) 4(1) P_wt_secure Tatt2_case 1(8)]
welltyped_leakage_free_long_term_secretI[OF P_wt_secure Tatt_lts_case]
welltyped_leakage_free_short_term_secretI[
OF * 1(3) 4(1) P_wt_secure Tatt1_case 1(8)]
welltyped_leakage_free_occurs_factI[OF * 1(5)]
** ****
have 6: "is_Fun s ∧ length (args s) = 2 ⟷ (∃f t u. s = Fun f [t, u])"
for s::"('fun,'atom,'sets,'lbl) prot_term"
by auto
define pubconst_cond where
"pubconst_cond ≡ λs::('fun,'atom,'sets,'lbl) prot_term.
is_Fun s ∧ args s = [] ∧ is_Fu (the_Fun s) ∧
public⇩f (the_Fu (the_Fun s)) ∧ arity⇩f (the_Fu (the_Fun s)) = 0"
define valuevar_cond where
"valuevar_cond ≡ λs::('fun,'atom,'sets,'lbl) prot_term.
is_Var s ∧ fst (the_Var s) = TAtom Value"
define setconst_cond where
"setconst_cond ≡ λs::('fun,'atom,'sets,'lbl) prot_term.
is_Fun s ∧ args s = [] ∧ is_Set (the_Fun s) ∧ arity⇩s (the_Set (the_Fun s)) = 0"
define setop_pair_cond where
"setop_pair_cond ≡ λs.
is_Fun s ∧ the_Fun s = Pair ∧ length (args s) = 2 ∧ setconst_cond (args s ! 1)"
define occursfact_cond where
"occursfact_cond ≡ λs::('fun,'atom,'sets,'lbl) prot_term.
is_Fun s ∧ the_Fun s = OccursFact ∧ length (args s) = 2 ∧
args s ! 0 = Fun OccursSec []"
define invkey_cond where
"invkey_cond ≡ λs.
is_Fun s ∧ the_Fun s = Fu invfun ∧ length (args s) = 2 ∧
args s ! 0 = ⟨privsec⟩⇩c ∧ valuevar_cond (args s ! 1)"
define ground_lts_cond where
"ground_lts_cond ≡ λs. is_Fun s ∧ is_Fu (the_Fun s) ∧ fv s = {} ∧ Tatt_lts s ∈ set P"
note cond_defs =
pubconst_cond_def valuevar_cond_def setconst_cond_def setop_pair_cond_def
occursfact_cond_def invkey_cond_def ground_lts_cond_def
have "(∃m. s = ⟨m: value⟩⇩v) ⟷ valuevar_cond s"
"(∃x c. arity⇩s c = 0 ∧ s = Fun Pair [x, ⟨c⟩⇩s]) ⟷ setop_pair_cond s"
"(∃c. arity⇩s c = 0 ∧ s = ⟨c⟩⇩s) ⟷ setconst_cond s"
"(∃c. public⇩f c ∧ arity⇩f c = 0 ∧ s = ⟨c⟩⇩c) ⟷ pubconst_cond s"
"(∃x. s = occurs x) ⟷ occursfact_cond s"
"(∃x. s = ⟨invfun [⟨privsec⟩⇩c, ⟨x: value⟩⇩v]⟩⇩t) ⟷ invkey_cond s"
"(∃f ts. s = ⟨f ts⟩⇩t ∧ fv s = {} ∧ Tatt_lts s ∈ set P) ⟷ ground_lts_cond s"
for s::"('fun,'atom,'sets,'lbl) prot_term"
unfolding is_Set_def the_Set_def is_Fu_def cond_defs
by (fastforce,use 6[of s] in fastforce,fastforce,force,fastforce,fastforce,fastforce)
moreover have
"(∀s ∈ set S. valuevar_cond s ∨ setop_pair_cond s ∨ setconst_cond s ∨ pubconst_cond s ∨
s = ⟨privsec⟩⇩c ∨ s = Fun OccursSec [] ∨ occursfact_cond s ∨ invkey_cond s ∨
ground_lts_cond s) ∧
(¬public⇩f privsec ∧ arity⇩f privsec = 0 ∧ Γ⇩f privsec ≠ None)"
using a unfolding defs' cond_defs Tatts by meson
ultimately have 7:
"∀s ∈ set S.
(∃x c. arity⇩s c = 0 ∧ s = Fun Pair [x, ⟨c⟩⇩s]) ∨
(∃c. arity⇩s c = 0 ∧ s = ⟨c⟩⇩s) ∨
(∃c. public⇩f c ∧ arity⇩f c = 0 ∧ s = ⟨c⟩⇩c) ∨
s = ⟨privsec⟩⇩c ∨ s = Fun OccursSec [] ∨
(∃m. s = ⟨m: value⟩⇩v) ∨
(∃x. s = occurs x) ∨
(∃x. s = ⟨invfun [⟨privsec⟩⇩c, ⟨x: value⟩⇩v]⟩⇩t) ∨
(∃f ts. s = ⟨f ts⟩⇩t ∧ fv s = {} ∧ Tatt_lts s ∈ set P)"
unfolding Let_def by fastforce
show ?thesis
by (rule iffD2[OF welltyped_leakage_free_protocol_pointwise]; unfold list_all_iff; intro ballI)
(use bspec[OF 7] 5 in blast)
qed
end
end
locale composable_stateful_protocols =
pm: stateful_protocol_model arity⇩f arity⇩s public⇩f Ana⇩f Γ⇩f label_witness1 label_witness2
for arity⇩f::"'fun ⇒ nat"
and arity⇩s::"'sets ⇒ nat"
and public⇩f::"'fun ⇒ bool"
and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,nat) prot_fun, nat) term list × nat list)"
and Γ⇩f::"'fun ⇒ 'atom option"
and label_witness1::"nat"
and label_witness2::"nat"
+
fixes Pc::"('fun,'atom,'sets,nat) prot_transaction list"
and Ps Ps_with_star_projs::"('fun,'atom,'sets,nat) prot_transaction list list"
and Pc_SMP Sec_symbolic::"('fun,'atom,'sets,nat) prot_term list"
and Ps_GSMPs::"(nat × ('fun,'atom,'sets,nat) prot_term list) list"
assumes Pc_def: "Pc = concat Ps"
and Ps_with_star_projs_def: "let Pc' = Pc; L = [0..<length Ps] in
Ps_with_star_projs = map (λn. (map (transaction_proj n) Pc')) L ∧
set L = set (remdups (concat (
map (λT. map (the_LabelN ∘ fst)
(filter (Not ∘ has_LabelS) (transaction_strand T)))
Pc')))"
and Pc_wellformed_composable:
"list_all (list_all (Not ∘ has_LabelS) ∘ tl ∘ transaction_send) Pc"
"pm.wellformed_composable_protocols Ps Pc_SMP"
"pm.wellformed_composable_protocols' Ps"
"pm.composable_protocols Ps Ps_GSMPs Sec_symbolic"
begin
theorem composed_protocol_preserves_component_goals:
assumes components_leakage_free:
"list_all (pm.welltyped_leakage_free_protocol Sec_symbolic) Ps_with_star_projs"
and n_def: "n < length Ps_with_star_projs"
and P_def: "P = Ps_with_star_projs ! n"
and P_welltyped_secure:
"∀𝒜 ∈ pm.reachable_constraints P. ∄ℐ.
pm.welltyped_constraint_model ℐ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
shows "∀𝒜 ∈ pm.reachable_constraints Pc. ∄ℐ.
pm.constraint_model ℐ (𝒜@[⟨n, send⟨[attack⟨ln n⟩]⟩⟩])"
proof -
note 0 = Ps_with_star_projs_def[unfolded Let_def]
have 1:
"set Ps_with_star_projs =
(λn. map (transaction_proj n) Pc) `
set (remdups (concat (map (λT. map (the_LabelN ∘ fst)
(filter (Not ∘ has_LabelS) (transaction_strand T)))
Pc)))"
by (metis (mono_tags, lifting) 0 image_set)
have 2: "Ps_with_star_projs ! n = map (transaction_proj n) Pc"
using conjunct1[OF 0] n_def by fastforce
show ?thesis
by (rule pm.composable_protocols_par_comp_prot'[
OF Pc_def 1 Pc_wellformed_composable
components_leakage_free 2 P_welltyped_secure[unfolded P_def]])
qed
end
end