diff -r fb201e383f1b -r da575186d492 Nominal/Ex/Pi.thy --- a/Nominal/Ex/Pi.thy Tue Feb 19 05:38:46 2013 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,575 +0,0 @@ -(* Theory be Kirstin Peters *) - -theory Pi - imports "../Nominal2" -begin - -atom_decl name - -subsection {* Capture-Avoiding Substitution of Names *} - -definition - subst_name :: "name \ name \ name \ name" ("_[_:::=_]" [110, 110, 110] 110) -where - "a[b:::=c] \ if (a = b) then c else a" - -declare subst_name_def[simp] - -lemma subst_name_mix_eqvt[eqvt]: - fixes p :: perm - and a :: name - and b :: name - and c :: name - - shows "p \ (a[b:::=c]) = (p \ a)[(p \ b):::=(p \ c)]" -proof - - show ?thesis - by(auto) -qed - -nominal_primrec - subst_name_list :: "name \ (name \ name) list \ name" -where - "subst_name_list a [] = a" -| "subst_name_list a ((b, c)#xs) = (if (a = b) then c else (subst_name_list a xs))" - apply(simp add: eqvt_def subst_name_list_graph_aux_def) - apply(simp) - apply(auto) - apply(rule_tac y="b" in list.exhaust) - by(auto) - -termination (eqvt) - by (lexicographic_order) - - -section {* The Synchronous Pi-Calculus *} - -subsection {* Syntax: Synchronous, Monadic Pi-Calculus with n-ary, Mixed Choice *} - -nominal_datatype - guardedTerm_mix = Output name name piMix ("_!<_>\._" [120, 120, 110] 110) - | Input name b::name P::piMix binds b in P ("_?<_>\._" [120, 120, 110] 110) - | Tau piMix ("<\\>._" [110] 110) - and sumList_mix = SumNil ("\\") - | AddSummand guardedTerm_mix sumList_mix (infixr "\\" 65) - and piMix = Res a::name P::piMix binds a in P ("<\_>\_" [100, 100] 100) - | Par piMix piMix (infixr "\\" 85) - | Match name name piMix ("[_\\_]_" [120, 120, 110] 110) - | Sum sumList_mix ("\\{_}" 90) - | Rep name b::name P::piMix binds b in P ("\_?<_>\._" [120, 120, 110] 110) - | Succ ("succ\") - -lemmas piMix_strong_induct = guardedTerm_mix_sumList_mix_piMix.strong_induct -lemmas piMix_fresh = guardedTerm_mix_sumList_mix_piMix.fresh -lemmas piMix_eq_iff = guardedTerm_mix_sumList_mix_piMix.eq_iff -lemmas piMix_distinct = guardedTerm_mix_sumList_mix_piMix.distinct -lemmas piMix_size = guardedTerm_mix_sumList_mix_piMix.size - -subsection {* Alpha-Conversion Lemmata *} - -lemma alphaRes_mix: - fixes a :: name - and P :: piMix - and z :: name - - assumes "atom z \ P" - - shows "<\a>\P = <\z>\((atom a \ atom z) \ P)" -proof(cases "a = z") - assume "a = z" - thus ?thesis - by(simp) -next - assume "a \ z" - thus ?thesis - using assms - by (simp add: flip_def piMix_eq_iff Abs1_eq_iff fresh_permute_left) -qed - -lemma alphaInput_mix: - fixes a :: name - and b :: name - and P :: piMix - and z :: name - - assumes "atom z \ P" - - shows "a?\.P = a?\.((atom b \ atom z) \ P)" -proof(cases "b = z") - assume "b = z" - thus ?thesis - by(simp) -next - assume "b \ z" - thus ?thesis - using assms - by(simp add: flip_def piMix_eq_iff Abs1_eq_iff fresh_permute_left) -qed - -lemma alphaRep_mix: - fixes a :: name - and b :: name - and P :: piMix - and z :: name - - assumes "atom z \ P" - - shows "\a?\.P = \a?\.((atom b \ atom z) \ P)" -proof(cases "b = z") - assume "b = z" - thus ?thesis - by(simp) -next - assume "b \ z" - thus ?thesis - using assms - by(simp add: flip_def piMix_eq_iff Abs1_eq_iff fresh_permute_left) -qed - -subsection {* Capture-Avoiding Substitution of Names *} - -lemma testl: - assumes a: "\y. f = Inl y" - shows "(p \ (Sum_Type.Projl f)) = Sum_Type.Projl (p \ f)" -using a by auto - -lemma testrr: - assumes a: "\y. f = Inr (Inr y)" - shows "(p \ (Sum_Type.Projr (Sum_Type.Projr f))) = Sum_Type.Projr (Sum_Type.Projr (p \ f))" -using a by auto - -lemma testlr: - assumes a: "\y. f = Inr (Inl y)" - shows "(p \ (Sum_Type.Projl (Sum_Type.Projr f))) = Sum_Type.Projl (Sum_Type.Projr (p \ f))" -using a by auto - -nominal_primrec (default "sum_case (\x. Inl undefined) (sum_case (\x. Inr (Inl undefined)) (\x. Inr (Inr undefined)))") - subsGuard_mix :: "guardedTerm_mix \ name \ name \ guardedTerm_mix" ("_[_::=\\_]" [100, 100, 100] 100) and - subsList_mix :: "sumList_mix \ name \ name \ sumList_mix" ("_[_::=\\_]" [100, 100, 100] 100) and - subs_mix :: "piMix \ name \ name \ piMix" ("_[_::=\_]" [100, 100, 100] 100) -where - "(a!\.P)[x::=\\y] = (a[x:::=y])!<(b[x:::=y])>\.(P[x::=\y])" -| "\atom b \ (x, y)\ \ (a?\.P)[x::=\\y] = (a[x:::=y])?\.(P[x::=\y])" -| "(<\\>.P)[x::=\\y] = <\\>.(P[x::=\y])" -| "(\\)[x::=\\y] = \\" -| "(g \\ xg)[x::=\\y] = (g[x::=\\y]) \\ (xg[x::=\\y])" -| "\atom a \ (x, y)\ \ (<\a>\P)[x::=\y] = <\a>\(P[x::=\y])" -| "(P \\ Q)[x::=\y] = (P[x::=\y]) \\ (Q[x::=\y])" -| "([a\\b]P)[x::=\y] = ([(a[x:::=y])\\(b[x:::=y])](P[x::=\y]))" -| "(\\{xg})[x::=\y] = \\{(xg[x::=\\y])}" -| "\atom b \ (x, y)\ \ (\a?\.P)[x::=\y] = \(a[x:::=y])?\.(P[x::=\y])" -| "(succ\)[x::=\y] = succ\" - using [[simproc del: alpha_lst]] - apply(auto simp add: piMix_distinct piMix_eq_iff) - apply(simp add: eqvt_def subsGuard_mix_subsList_mix_subs_mix_graph_aux_def) - --"Covered all cases" - apply(case_tac x) - apply(simp) - apply(case_tac a) - apply(simp) - apply (rule_tac y="aa" and c="(b, c)" in guardedTerm_mix_sumList_mix_piMix.strong_exhaust(1)) - apply(blast) - apply(auto simp add: fresh_star_def)[1] - apply(blast) - apply(simp) - apply(blast) - apply(simp) - apply(case_tac b) - apply(simp) - apply(case_tac a) - apply(simp) - apply (rule_tac ya="aa" in guardedTerm_mix_sumList_mix_piMix.strong_exhaust(2)) - apply(blast) - apply(blast) - apply(simp) - apply(case_tac ba) - apply(simp) - apply (rule_tac yb="a" and c="(bb,c)" in guardedTerm_mix_sumList_mix_piMix.strong_exhaust(3)) - using [[simproc del: alpha_lst]] - apply(auto simp add: fresh_star_def)[1] - apply(blast) - apply(blast) - apply(blast) - using [[simproc del: alpha_lst]] - apply(auto simp add: fresh_star_def)[1] - apply(blast) - apply(simp) - apply(blast) - --"compatibility" - apply (simp only: meta_eq_to_obj_eq[OF subs_mix_def, symmetric, unfolded fun_eq_iff]) - apply (subgoal_tac "eqvt_at (\(a, b, c). subs_mix a b c) (P, xa, ya)") - apply (thin_tac "eqvt_at subsGuard_mix_subsList_mix_subs_mix_sumC (Inr (Inr (P, xa, ya)))") - apply (thin_tac "eqvt_at subsGuard_mix_subsList_mix_subs_mix_sumC (Inr (Inr (Pa, xa, ya)))") - prefer 2 - apply (simp only: eqvt_at_def subs_mix_def) - apply rule - apply(simp (no_asm)) - apply (subst testrr) - apply (simp add: subsGuard_mix_subsList_mix_subs_mix_sumC_def) - apply (simp add: THE_default_def) -apply (case_tac "Ex1 (subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (P, xa, ya))))") -apply simp_all[2] -apply auto[1] -apply (erule_tac x="x" in allE) -apply simp -apply (thin_tac "\p\perm. - p \ The (subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (P, xa, ya)))) = - (if \!x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ P, p \ xa, p \ ya))) x - then THE x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ P, p \ xa, p \ ya))) x - else Inr (Inr undefined))") -apply (thin_tac "\p\perm. - p \ (if \!x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (Pa, xa, ya))) x - then THE x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (Pa, xa, ya))) x - else Inr (Inr undefined)) = - (if \!x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ Pa, p \ xa, p \ ya))) x - then THE x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ Pa, p \ xa, p \ ya))) x - else Inr (Inr undefined))") -apply (thin_tac "atom b \ (xa, ya)") -apply (thin_tac "atom ba \ (xa, ya)") -apply (thin_tac "[[atom b]]lst. P = [[atom ba]]lst. Pa") -apply(cases rule: subsGuard_mix_subsList_mix_subs_mix_graph.cases) -apply assumption -apply (metis Inr_not_Inl) -apply (metis Inr_not_Inl) -apply (metis Inr_not_Inl) -apply (metis Inr_inject Inr_not_Inl) -apply (metis Inr_inject Inr_not_Inl) -apply (rule_tac x="<\a>\Sum_Type.Projr - (Sum_Type.Projr - (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="Sum_Type.Projr - (Sum_Type.Projr (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y))))) \\ - Sum_Type.Projr - (Sum_Type.Projr (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Q, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="[(a[xb:::=y])\\(bb[xb:::=y])]Sum_Type.Projr - (Sum_Type.Projr -(subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="\\{Sum_Type.Projl - (Sum_Type.Projr - (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inl (xg, xb, y)))))}" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="\(a[xb:::=y])?\.Sum_Type.Projr - (Sum_Type.Projr - (subsGuard_mix_subsList_mix_subs_mix_sum - (Inr (Inr (Pb, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="succ\" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply simp -(* Here the only real goal compatibility is left *) - apply (erule Abs_lst1_fcb) - apply (simp_all add: Abs_fresh_iff fresh_fun_eqvt_app) - apply (subgoal_tac "atom ba \ (\(a, x, y). subs_mix a x y) (P, xa, ya)") - apply simp - apply (erule fresh_eqvt_at) - apply(simp add: finite_supp) - apply(simp) - apply(simp add: eqvt_at_def) - apply(drule_tac x="(b \ ba)" in spec) - apply(simp) - apply (simp only: meta_eq_to_obj_eq[OF subs_mix_def, symmetric, unfolded fun_eq_iff]) - apply(rename_tac b P ba xa ya Pa) - apply (subgoal_tac "eqvt_at (\(a, b, c). subs_mix a b c) (P, xa, ya)") - apply (thin_tac "eqvt_at subsGuard_mix_subsList_mix_subs_mix_sumC (Inr (Inr (P, xa, ya)))") - apply (thin_tac "eqvt_at subsGuard_mix_subsList_mix_subs_mix_sumC (Inr (Inr (Pa, xa, ya)))") - prefer 2 - apply (simp only: eqvt_at_def subs_mix_def) - apply rule - apply(simp (no_asm)) - apply (subst testrr) - apply (simp add: subsGuard_mix_subsList_mix_subs_mix_sumC_def) - apply (simp add: THE_default_def) -apply (case_tac "Ex1 (subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (P, xa, ya))))") -apply simp_all[2] -apply auto[1] -apply (erule_tac x="x" in allE) -apply simp -apply (thin_tac "\p\perm. - p \ The (subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (P, xa, ya)))) = - (if \!x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ P, p \ xa, p \ ya))) x - then THE x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ P, p \ xa, p \ ya))) x - else Inr (Inr undefined))") -apply (thin_tac "\p\perm. - p \ (if \!x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (Pa, xa, ya))) x - then THE x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (Pa, xa, ya))) x - else Inr (Inr undefined)) = - (if \!x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ Pa, p \ xa, p \ ya))) x - then THE x\guardedTerm_mix + sumList_mix + piMix. - subsGuard_mix_subsList_mix_subs_mix_graph (Inr (Inr (p \ Pa, p \ xa, p \ ya))) x - else Inr (Inr undefined))") -apply (thin_tac "atom b \ (xa, ya)") -apply (thin_tac "atom ba \ (xa, ya)") -apply (thin_tac "[[atom b]]lst. P = [[atom ba]]lst. Pa") -apply(cases rule: subsGuard_mix_subsList_mix_subs_mix_graph.cases) -apply assumption -apply (metis Inr_not_Inl) -apply (metis Inr_not_Inl) -apply (metis Inr_not_Inl) -apply (metis Inr_inject Inr_not_Inl) -apply (metis Inr_inject Inr_not_Inl) -apply (rule_tac x="<\a>\Sum_Type.Projr - (Sum_Type.Projr - (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="Sum_Type.Projr - (Sum_Type.Projr (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y))))) \\ - Sum_Type.Projr - (Sum_Type.Projr (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Q, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="[(a[xb:::=y])\\(bb[xb:::=y])]Sum_Type.Projr - (Sum_Type.Projr -(subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inr (Pb, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="\\{Sum_Type.Projl - (Sum_Type.Projr - (subsGuard_mix_subsList_mix_subs_mix_sum (Inr (Inl (xg, xb, y)))))}" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="\(a[xb:::=y])?\.Sum_Type.Projr - (Sum_Type.Projr - (subsGuard_mix_subsList_mix_subs_mix_sum - (Inr (Inr (Pb, xb, y)))))" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply (rule_tac x="succ\" in exI) -apply clarify -apply (rule the1_equality) -apply blast apply assumption -apply simp -(* Here the only real goal compatibility is left *) - apply (erule Abs_lst1_fcb) - apply (simp_all add: Abs_fresh_iff fresh_fun_eqvt_app) - apply (subgoal_tac "atom ba \ (\(a, x, y). subs_mix a x y) (P, xa, ya)") - apply simp - apply (erule fresh_eqvt_at) - apply(simp add: finite_supp) - apply(simp) - apply(simp add: eqvt_at_def) - apply(drule_tac x="(b \ ba)" in spec) - apply(simp) - done - -termination (eqvt) - by (lexicographic_order) - -lemma forget_mix: - fixes g :: guardedTerm_mix - and xg :: sumList_mix - and P :: piMix - and x :: name - and y :: name - - shows "atom x \ g \ g[x::=\\y] = g" - and "atom x \ xg \ xg[x::=\\y] = xg" - and "atom x \ P \ P[x::=\y] = P" -proof - - show "atom x \ g \ g[x::=\\y] = g" - and "atom x \ xg \ xg[x::=\\y] = xg" - and "atom x \ P \ P[x::=\y] = P" - using assms - apply(nominal_induct g and xg and P avoiding: x y rule: piMix_strong_induct) - by(auto simp add: piMix_eq_iff piMix_fresh fresh_at_base) -qed - -lemma fresh_fact_mix: - fixes g :: guardedTerm_mix - and xg :: sumList_mix - and P :: piMix - and x :: name - and y :: name - and z :: name - - assumes "atom z \ y" - - shows "(z = x \ atom z \ g) \ atom z \ g[x::=\\y]" - and "(z = x \ atom z \ xg) \ atom z \ xg[x::=\\y]" - and "(z = x \ atom z \ P) \ atom z \ P[x::=\y]" -proof - - show "(z = x \ atom z \ g) \ atom z \ g[x::=\\y]" - and "(z = x \ atom z \ xg) \ atom z \ xg[x::=\\y]" - and "(z = x \ atom z \ P) \ atom z \ P[x::=\y]" - using assms - apply(nominal_induct g and xg and P avoiding: x y z rule: piMix_strong_induct) - by(auto simp add: piMix_fresh fresh_at_base) -qed - -lemma substitution_lemma_mix: - fixes g :: guardedTerm_mix - and xg :: sumList_mix - and P :: piMix - and s :: name - and u :: name - and x :: name - and y :: name - - assumes "x \ y" - and "atom x \ u" - - shows "g[x::=\\s][y::=\\u] = g[y::=\\u][x::=\\s[y:::=u]]" - and "xg[x::=\\s][y::=\\u] = xg[y::=\\u][x::=\\s[y:::=u]]" - and "P[x::=\s][y::=\u] = P[y::=\u][x::=\s[y:::=u]]" -proof - - show "g[x::=\\s][y::=\\u] = g[y::=\\u][x::=\\s[y:::=u]]" - and "xg[x::=\\s][y::=\\u] = xg[y::=\\u][x::=\\s[y:::=u]]" - and "P[x::=\s][y::=\u] = P[y::=\u][x::=\s[y:::=u]]" - using assms - apply(nominal_induct g and xg and P avoiding: x y s u rule: piMix_strong_induct) - apply(simp_all add: fresh_fact_mix forget_mix) - by(auto simp add: fresh_at_base) -qed - -lemma perm_eq_subst_mix: - fixes g :: guardedTerm_mix - and xg :: sumList_mix - and P :: piMix - and x :: name - and y :: name - - shows "atom y \ g \ (atom x \ atom y) \ g = g[x::=\\y]" - and "atom y \ xg \ (atom x \ atom y) \ xg = xg[x::=\\y]" - and "atom y \ P \ (atom x \ atom y) \ P = P[x::=\y]" -proof - - show "atom y \ g \ (atom x \ atom y) \ g = g[x::=\\y]" - and "atom y \ xg \ (atom x \ atom y) \ xg = xg[x::=\\y]" - and "atom y \ P \ (atom x \ atom y) \ P = P[x::=\y]" - apply(nominal_induct g and xg and P avoiding: x y rule: piMix_strong_induct) - by(auto simp add: piMix_fresh fresh_at_base) -qed - -lemma subst_id_mix: - fixes g :: guardedTerm_mix - and xg :: sumList_mix - and P :: piMix - and x :: name - - shows "g[x::=\\x] = g" and "xg[x::=\\x] = xg" and "P[x::=\x] = P" -proof - - show "g[x::=\\x] = g" and "xg[x::=\\x] = xg" and "P[x::=\x] = P" - apply(nominal_induct g and xg and P avoiding: x rule: piMix_strong_induct) - by(auto) -qed - -lemma alphaRes_subst_mix: - fixes a :: name - and P :: piMix - and z :: name - - assumes "atom z \ P" - - shows "<\a>\P = <\z>\(P[a::=\z])" -proof(cases "a = z") - assume "a = z" - thus ?thesis - by(simp add: subst_id_mix) -next - assume "a \ z" - thus ?thesis - using assms - by(simp add: alphaRes_mix perm_eq_subst_mix) -qed - -lemma alphaInput_subst_mix: - fixes a :: name - and b :: name - and P :: piMix - and z :: name - - assumes "atom z \ P" - - shows "a?\.P = a?\.(P[b::=\z])" -proof(cases "b = z") - assume "b = z" - thus ?thesis - by(simp add: subst_id_mix) -next - assume "b \ z" - thus ?thesis - using assms - by(simp add: alphaInput_mix perm_eq_subst_mix) -qed - -lemma alphaRep_subst_mix: - fixes a :: name - and b :: name - and P :: piMix - and z :: name - - assumes "atom z \ P" - - shows "\a?\.P = \a?\.(P[b::=\z])" -proof(cases "b = z") - assume "b = z" - thus ?thesis - by(simp add: subst_id_mix) -next - assume "b \ z" - thus ?thesis - using assms - by(simp add: alphaRep_mix perm_eq_subst_mix) -qed - -inductive - fresh_list_guard_mix :: "name list \ guardedTerm_mix \ bool" -where - "fresh_list_guard_mix [] g" -| "\atom n \ g; fresh_list_guard_mix xn g\ \ fresh_list_guard_mix (n#xn) g" - -equivariance fresh_list_guard_mix -nominal_inductive fresh_list_guard_mix - done - -inductive - fresh_list_sumList_mix :: "name list \ sumList_mix \ bool" -where - "fresh_list_sumList_mix [] xg" -| "\atom n \ xg; fresh_list_sumList_mix xn xg\ \ fresh_list_sumList_mix (n#xn) xg" - -equivariance fresh_list_sumList_mix -nominal_inductive fresh_list_sumList_mix - done - -inductive - fresh_list_mix :: "name list \ piMix \ bool" -where - "fresh_list_mix [] P" -| "\atom n \ P; fresh_list_mix xn P\ \ fresh_list_mix (n#xn) P" - -equivariance fresh_list_mix -nominal_inductive fresh_list_mix - done - -end \ No newline at end of file