Theory Chain
theory Chain
imports "HOL-Nominal.Nominal"
begin
lemma pt_set_nil:
fixes Xs :: "'a set"
assumes pt: "pt TYPE('a) TYPE ('x)"
and at: "at TYPE('x)"
shows "([]::'x prm)∙Xs = Xs"
by(auto simp add: perm_set_def pt1[OF pt])
lemma pt_set_append:
fixes pi1 :: "'x prm"
and pi2 :: "'x prm"
and Xs :: "'a set"
assumes pt: "pt TYPE('a) TYPE ('x)"
and at: "at TYPE('x)"
shows "(pi1@pi2)∙Xs = pi1∙(pi2∙Xs)"
by(auto simp add: perm_set_def pt2[OF pt])
lemma pt_set_prm_eq:
fixes pi1 :: "'x prm"
and pi2 :: "'x prm"
and Xs :: "'a set"
assumes pt: "pt TYPE('a) TYPE ('x)"
and at: "at TYPE('x)"
shows "pi1 ≜ pi2 ⟹ pi1∙Xs = pi2∙Xs"
by(auto simp add: perm_set_def pt3[OF pt])
lemma pt_set_inst:
assumes pt: "pt TYPE('a) TYPE ('x)"
and at: "at TYPE('x)"
shows "pt TYPE('a set) TYPE('x)"
apply(simp add: pt_def pt_set_nil[OF pt, OF at] pt_set_append[OF pt, OF at])
apply(clarify)
by(rule pt_set_prm_eq[OF pt, OF at])
lemma pt_ball_eqvt:
fixes pi :: "'a prm"
and Xs :: "'b set"
and P :: "'b ⇒ bool"
assumes pt: "pt TYPE('b) TYPE('a)"
and at: "at TYPE('a)"
shows "(pi ∙ (∀x ∈ Xs. P x)) = (∀x ∈ (pi ∙ Xs). pi ∙ P (rev pi ∙ x))"
apply(auto simp add: perm_bool)
apply(drule_tac pi="rev pi" in pt_set_bij2[OF pt, OF at])
apply(simp add: pt_rev_pi[OF pt_set_inst[OF pt, OF at], OF at])
apply(drule_tac pi="pi" in pt_set_bij2[OF pt, OF at])
apply(erule_tac x="pi ∙ x" in ballE)
apply(simp add: pt_rev_pi[OF pt, OF at])
by simp
lemma perm_cart_prod:
fixes Xs :: "'b set"
and Ys :: "'c set"
and p :: "'a prm"
assumes pt1: "pt TYPE('b) TYPE('a)"
and pt2: "pt TYPE('c) TYPE('a)"
and at: "at TYPE('a)"
shows "(p ∙ (Xs × Ys)) = (((p ∙ Xs) × (p ∙ Ys))::(('b × 'c) set))"
by(auto simp add: perm_set_def)
lemma supp_member:
fixes Xs :: "'b set"
and x :: 'a
assumes pt: "pt TYPE('b) TYPE('a)"
and at: "at TYPE('a)"
and fs: "fs TYPE('b) TYPE('a)"
and "finite Xs"
and "x ∈ ((supp Xs)::'a set)"
obtains X where "(X::'b) ∈ Xs" and "x ∈ supp X"
proof -
from ‹finite Xs› ‹x ∈ supp Xs› have "∃X::'b. (X ∈ Xs) ∧ (x ∈ (supp X))"
proof(induct rule: finite_induct)
case empty
from ‹x ∈ ((supp {})::'a set)› have False
by(simp add: supp_set_empty)
thus ?case by simp
next
case(insert y Xs)
show ?case
proof(case_tac "x ∈ supp y")
assume "x ∈ supp y"
thus ?case by force
next
assume "x ∉ supp y"
with ‹x ∈ supp(insert y Xs)› have "x ∈ supp Xs"
by(simp add: supp_fin_insert[OF pt, OF at, OF fs, OF ‹finite Xs›])
with ‹x ∈ supp Xs ⟹ ∃X. X ∈ Xs ∧ x ∈ supp X›
show ?case by force
qed
qed
with that show ?thesis by blast
qed
lemma supp_cart_prod_empty[simp]:
fixes Xs :: "'b set"
shows "supp (Xs × {}) = ({}::'a set)"
and "supp ({} × Xs) = ({}::'a set)"
by(auto simp add: supp_set_empty)
lemma supp_cart_prod:
fixes Xs :: "'b set"
and Ys :: "'c set"
assumes pt1: "pt TYPE('b) TYPE('a)"
and pt2: "pt TYPE('c) TYPE('a)"
and fs1: "fs TYPE('b) TYPE('a)"
and fs2: "fs TYPE('c) TYPE('a)"
and at: "at TYPE('a)"
and f1: "finite Xs"
and f2: "finite Ys"
and a: "Xs ≠ {}"
and b: "Ys ≠ {}"
shows "((supp (Xs × Ys))::'a set) = ((supp Xs) ∪ (supp Ys))"
proof -
from f1 f2 have f3: "finite(Xs × Ys)" by simp
show ?thesis
apply(simp add: supp_of_fin_sets[OF pt_prod_inst[OF pt1, OF pt2], OF at, OF fs_prod_inst[OF fs1, OF fs2], OF f3] supp_prod)
apply(rule equalityI)
using Union_included_in_supp[OF pt1, OF at, OF fs1, OF f1] Union_included_in_supp[OF pt2, OF at, OF fs2, OF f2]
apply(force simp add: supp_prod)
using a b
apply(auto simp add: supp_prod)
using supp_member[OF pt1, OF at, OF fs1, OF f1]
apply blast
using supp_member[OF pt2, OF at, OF fs2, OF f2]
by blast
qed
lemma fresh_cart_prod:
fixes x :: 'a
and Xs :: "'b set"
and Ys :: "'c set"
assumes pt1: "pt TYPE('b) TYPE('a)"
and pt2: "pt TYPE('c) TYPE('a)"
and fs1: "fs TYPE('b) TYPE('a)"
and fs2: "fs TYPE('c) TYPE('a)"
and at: "at TYPE('a)"
and f1: "finite Xs"
and f2: "finite Ys"
and a: "Xs ≠ {}"
and b: "Ys ≠ {}"
shows "(x ♯ (Xs × Ys)) = (x ♯ Xs ∧ x ♯ Ys)"
using assms
by(simp add: supp_cart_prod fresh_def)
lemma fresh_star_cart_prod:
fixes Zs :: "'a set"
and xvec :: "'a list"
and Xs :: "'b set"
and Ys :: "'c set"
assumes pt1: "pt TYPE('b) TYPE('a)"
and pt2: "pt TYPE('c) TYPE('a)"
and fs1: "fs TYPE('b) TYPE('a)"
and fs2: "fs TYPE('c) TYPE('a)"
and at: "at TYPE('a)"
and f1: "finite Xs"
and f2: "finite Ys"
and a: "Xs ≠ {}"
and b: "Ys ≠ {}"
shows "(Zs ♯* (Xs × Ys)) = (Zs ♯* Xs ∧ Zs ♯* Ys)"
and "(xvec ♯* (Xs × Ys)) = (xvec ♯* Xs ∧ xvec ♯* Ys)"
using assms
by(force simp add: fresh_cart_prod fresh_star_def)+
lemma permCommute:
fixes p :: "'a prm"
and q :: "'a prm"
and P :: 'x
and Xs :: "'a set"
and Ys :: "'a set"
assumes pt: "pt TYPE('x) TYPE('a)"
and at: "at TYPE('a)"
and a: "(set p) ⊆ Xs × Ys"
and b: "Xs ♯* q"
and c: "Ys ♯* q"
shows "p ∙ q ∙ P = q ∙ p ∙ P"
proof -
have "p ∙ q ∙ P = (p ∙ q) ∙ p ∙ P"
by(rule pt_perm_compose[OF pt, OF at])
moreover from at have "pt TYPE('a) TYPE('a)"
by(rule at_pt_inst)
hence "pt TYPE(('a × 'a) list) TYPE('a)"
by(force intro: pt_prod_inst pt_list_inst)
hence "p ∙ q = q" using at a b c
by(rule pt_freshs_freshs)
ultimately show ?thesis by simp
qed
definition
distinctPerm :: "'a prm ⇒ bool" where
"distinctPerm p ≡ distinct((map fst p)@(map snd p))"
lemma at_set_avoiding_aux':
fixes Xs::"'a set"
and As::"'a set"
assumes at: "at TYPE('a)"
and a: "finite Xs"
and b: "Xs ⊆ As"
and c: "finite As"
and d: "finite ((supp c)::'a set)"
shows "∃(Ys::'a set) (pi::'a prm). Ys♯*c ∧ Ys ∩ As = {} ∧ (pi∙Xs=Ys) ∧
set pi ⊆ Xs × Ys ∧ finite Ys ∧ (distinctPerm pi)"
using a b c
proof (induct)
case empty
have "({}::'a set)♯*c" by (simp add: fresh_star_def)
moreover
have "({}::'a set) ∩ As = {}" by simp
moreover
have "([]::'a prm)∙{} = ({}::'a set)"
by(rule pt1) (metis Nominal.pt_set_inst at at_pt_inst)
moreover
have "set ([]::'a prm) ⊆ {} × {}" by simp
moreover
have "finite ({}::'a set)" by simp
moreover have "distinctPerm([]::'a prm)"
by(simp add: distinctPerm_def)
ultimately show ?case by blast
next
case (insert x Xs)
then have ih: "∃Ys pi. Ys♯*c ∧ Ys ∩ As = {} ∧ pi∙Xs = Ys ∧ set pi ⊆ Xs × Ys ∧ finite Ys ∧ distinctPerm pi" by simp
then obtain Ys pi where a1: "Ys♯*c" and a2: "Ys ∩ As = {}" and a3: "(pi::'a prm)∙Xs = Ys" and
a4: "set pi ⊆ Xs × Ys" and a5: "finite Ys"
and a6: "distinctPerm pi" by blast
have b: "x∉Xs" by fact
have d1: "finite As" by fact
have d2: "finite Xs" by fact
have d3: "insert x Xs ⊆ As" by fact
have d4: "finite((supp pi)::'a set)"
by(induct pi)
(auto simp add: supp_list_nil supp_prod at_fin_set_supp[OF at]
supp_list_cons at_supp[OF at])
have "∃y::'a. y♯(c,x,Ys,As,pi)" using d d1 a5 d4
by (rule_tac at_exists_fresh'[OF at])
(simp add: supp_prod at_supp[OF at] at_fin_set_supp[OF at])
then obtain y::"'a" where e: "y♯(c,x,Ys,As,pi)" by blast
have "({y}∪Ys)♯*c" using a1 e by (simp add: fresh_star_def)
moreover
have "({y}∪Ys)∩As = {}" using a2 d1 e by (simp add: fresh_prod at_fin_set_fresh[OF at])
moreover
have "(((pi∙x,y)#pi)∙(insert x Xs)) = {y}∪Ys"
proof -
have eq: "[(pi∙x,y)]∙Ys = Ys"
proof -
have "(pi∙x)♯Ys" using a3[symmetric] b d2
by(simp add: pt_fresh_bij[OF pt_set_inst, OF at_pt_inst[OF at], OF at, OF at]
at_fin_set_fresh[OF at d2])
moreover
have "y♯Ys" using e by simp
ultimately show "[(pi∙x,y)]∙Ys = Ys"
by (simp add: pt_fresh_fresh[OF pt_set_inst, OF at_pt_inst[OF at], OF at, OF at])
qed
have "(((pi∙x,y)#pi)∙({x}∪Xs)) = ([(pi∙x,y)]∙(pi∙({x}∪Xs)))"
by (simp add: pt2[symmetric, OF pt_set_inst, OF at_pt_inst[OF at], OF at])
also have "… = {y}∪([(pi∙x,y)]∙(pi∙Xs))"
by (simp only: union_eqvt perm_set_def at_calc[OF at])(auto)
also have "… = {y}∪([(pi∙x,y)]∙Ys)" using a3 by simp
also have "… = {y}∪Ys" using eq by simp
finally show "(((pi∙x,y)#pi)∙(insert x Xs)) = {y}∪Ys" by auto
qed
moreover
have "pi∙x=x" using a4 b a2 a3 d3 by (rule_tac at_prm_fresh2[OF at]) (auto)
then have "set ((pi∙x,y)#pi) ⊆ (insert x Xs) × ({y}∪Ys)" using a4 by auto
moreover
have "finite ({y}∪Ys)" using a5 by simp
moreover from ‹Ys ∩ As = {}› ‹(insert x Xs) ⊆ As› ‹finite Ys› have "x ∉ Ys"
by(auto simp add: fresh_def at_fin_set_supp[OF at])
with a6 ‹pi ∙ x = x› ‹x ∉ Xs› ‹(set pi) ⊆ Xs × Ys› e have "distinctPerm((pi ∙ x, y)#pi)"
apply(auto simp add: distinctPerm_def fresh_prod at_fresh[OF at])
proof -
fix a b
assume "b ♯ pi" and "(a, b) ∈ set pi"
thus False
by(induct pi)
(auto simp add: supp_list_cons supp_prod at_supp[OF at] fresh_list_cons fresh_prod at_fresh[OF at])
next
fix a b
assume "a ♯ pi" and "(a, b) ∈ set pi"
thus False
by(induct pi)
(auto simp add: supp_list_cons supp_prod at_supp[OF at] fresh_list_cons fresh_prod at_fresh[OF at])
qed
ultimately
show ?case by blast
qed
lemma at_set_avoiding:
fixes Xs::"'a set"
assumes at: "at TYPE('a)"
and a: "finite Xs"
and b: "finite ((supp c)::'a set)"
obtains pi::"'a prm" where "(pi ∙ Xs) ♯* c" and "set pi ⊆ Xs × (pi ∙ Xs)" and "distinctPerm pi"
using a b
by (frule_tac As="Xs" in at_set_avoiding_aux'[OF at]) auto
lemma pt_swap:
fixes x :: 'a
and a :: 'x
and b :: 'x
assumes pt: "pt TYPE('a) TYPE('x)"
and at: "at TYPE('x)"
shows "[(a, b)] ∙ x = [(b, a)] ∙ x"
proof -
show ?thesis by(simp add: pt3[OF pt] at_ds5[OF at])
qed