# HG changeset patch # User berghofe # Date 1326235793 -3600 # Node ID 47bcf3d5d1f0e095c26396877cdd3473c4d25558 # Parent adac34829e108fdab6525790edb0e3e92a6e7d0c Reverted several lemmas involving sets to the state before the removal of the set type. diff -r adac34829e10 -r 47bcf3d5d1f0 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Tue Jan 10 18:12:55 2012 +0100 +++ b/src/HOL/Nominal/Nominal.thy Tue Jan 10 23:49:53 2012 +0100 @@ -58,7 +58,7 @@ "perm_bool pi b = b" definition perm_set :: "'x prm \ 'a set \ 'a set" where - "perm_set pi A = {x. rev pi \ x \ A}" + "perm_set pi X = {pi \ x | x. x \ X}" primrec perm_unit :: "'x prm \ unit \ unit" where "perm_unit pi () = ()" @@ -137,11 +137,15 @@ (* permutation on sets *) lemma empty_eqvt: shows "pi\{} = {}" - by (simp add: perm_set_def fun_eq_iff) + by (simp add: perm_set_def) lemma union_eqvt: shows "(pi\(X\Y)) = (pi\X) \ (pi\Y)" - by (simp add: perm_set_def fun_eq_iff Un_def) + by (auto simp add: perm_set_def) + +lemma insert_eqvt: + shows "pi\(insert x X) = insert (pi\x) (pi\X)" + by (auto simp add: perm_set_def) (* permutations on products *) lemma fst_eqvt: @@ -166,6 +170,12 @@ shows "pi\(rev l) = rev (pi\l)" by (induct l) (simp_all add: append_eqvt) +lemma set_eqvt: + fixes pi :: "'x prm" + and xs :: "'a list" + shows "pi\(set xs) = set (pi\xs)" +by (induct xs) (auto simp add: empty_eqvt insert_eqvt) + (* permutation on characters and strings *) lemma perm_string: fixes s::"string" @@ -1018,16 +1028,13 @@ by (simp add: pt_def perm_bool_def) lemma pt_set_inst: - assumes pta: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "pt TYPE('a set) TYPE('x)" -proof - - have *: "\pi A. pi \ A = {x. (pi \ (\x. x \ A)) x}" - by (simp add: perm_fun_def perm_bool_def perm_set_def) - from pta pt_bool_inst at - have "pt TYPE('a \ bool) TYPE('x)" by (rule pt_fun_inst) - then show ?thesis by (simp add: pt_def *) -qed + assumes pt: "pt TYPE('a) TYPE('x)" + shows "pt TYPE('a set) TYPE('x)" +apply(simp add: pt_def) +apply(simp_all add: perm_set_def) +apply(simp add: pt1[OF pt]) +apply(force simp add: pt2[OF pt] pt3[OF pt]) +done lemma pt_unit_inst: shows "pt TYPE(unit) TYPE('x)" @@ -1245,43 +1252,13 @@ apply(rule pt1[OF pt]) done -lemma perm_set_eq: - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "(pi::'x prm)\(X::'a set) = {pi\x | x. x\X}" - apply (auto simp add: perm_fun_def perm_bool perm_set_def) - apply (rule_tac x="rev pi \ x" in exI) - apply (simp add: pt_pi_rev [OF pt at]) - apply (simp add: pt_rev_pi [OF pt at]) - done - -lemma pt_insert_eqvt: - fixes pi::"'x prm" - and x::"'a" - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "(pi\(insert x X)) = insert (pi\x) (pi\X)" - by (auto simp add: perm_set_eq [OF pt at]) - -lemma pt_set_eqvt: - fixes pi :: "'x prm" - and xs :: "'a list" - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "pi\(set xs) = set (pi\xs)" -by (induct xs) (auto simp add: empty_eqvt pt_insert_eqvt [OF pt at]) - lemma supp_singleton: - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "(supp {x::'a} :: 'x set) = supp x" - by (force simp add: supp_def perm_set_eq [OF pt at]) + shows "supp {x} = supp x" + by (force simp add: supp_def perm_set_def) lemma fresh_singleton: - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "(a::'x)\{x::'a} = a\x" - by (simp add: fresh_def supp_singleton [OF pt at]) + shows "a\{x} = a\x" + by (simp add: fresh_def supp_singleton) lemma pt_set_bij1: fixes pi :: "'x prm" @@ -1290,7 +1267,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "((pi\x)\X) = (x\((rev pi)\X))" - by (force simp add: perm_set_eq [OF pt at] pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at]) + by (force simp add: perm_set_def pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at]) lemma pt_set_bij1a: fixes pi :: "'x prm" @@ -1299,7 +1276,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "(x\(pi\X)) = (((rev pi)\x)\X)" - by (force simp add: perm_set_eq [OF pt at] pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at]) + by (force simp add: perm_set_def pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at]) lemma pt_set_bij: fixes pi :: "'x prm" @@ -1308,7 +1285,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "((pi\x)\(pi\X)) = (x\X)" - by (simp add: perm_set_eq [OF pt at] pt_bij[OF pt, OF at]) + by (simp add: perm_set_def pt_bij[OF pt, OF at]) lemma pt_in_eqvt: fixes pi :: "'x prm" @@ -1355,7 +1332,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "(pi\(X\Y)) = ((pi\X)\(pi\Y))" -by (auto simp add: perm_set_eq [OF pt at] perm_bool pt_bij[OF pt, OF at]) +by (auto simp add: perm_set_def perm_bool pt_bij[OF pt, OF at]) lemma pt_set_diff_eqvt: fixes X::"'a set" @@ -1364,14 +1341,14 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "pi\(X - Y) = (pi\X) - (pi\Y)" - by (auto simp add: perm_set_eq [OF pt at] pt_bij[OF pt, OF at]) + by (auto simp add: perm_set_def pt_bij[OF pt, OF at]) lemma pt_Collect_eqvt: fixes pi::"'x prm" assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "pi\{x::'a. P x} = {x. P ((rev pi)\x)}" -apply(auto simp add: perm_set_eq [OF pt at] pt_rev_pi[OF pt, OF at]) +apply(auto simp add: perm_set_def pt_rev_pi[OF pt, OF at]) apply(rule_tac x="(rev pi)\x" in exI) apply(simp add: pt_pi_rev[OF pt, OF at]) done @@ -1415,7 +1392,7 @@ and at: "at TYPE('y)" shows "finite (pi\X) = finite X" proof - - have image: "(pi\X) = (perm pi ` X)" by (force simp only: perm_set_eq [OF pt at]) + have image: "(pi\X) = (perm pi ` X)" by (force simp only: perm_set_def) show ?thesis proof (rule iffI) assume "finite (pi\X)" @@ -1445,17 +1422,17 @@ and cp: "cp TYPE('a) TYPE('x) TYPE('y)" shows "(pi\((supp x)::'y set)) = supp (pi\x)" (is "?LHS = ?RHS") proof - - have "?LHS = {pi\a | a. infinite {b. [(a,b)]\x \ x}}" by (simp add: supp_def perm_set_eq [OF ptb at]) + have "?LHS = {pi\a | a. infinite {b. [(a,b)]\x \ x}}" by (simp add: supp_def perm_set_def) also have "\ = {pi\a | a. infinite {pi\b | b. [(a,b)]\x \ x}}" proof (rule Collect_permI, rule allI, rule iffI) fix a assume "infinite {b::'y. [(a,b)]\x \ x}" hence "infinite (pi\{b::'y. [(a,b)]\x \ x})" by (simp add: pt_set_infinite_ineq[OF ptb, OF at]) - thus "infinite {pi\b |b::'y. [(a,b)]\x \ x}" by (simp add: perm_set_eq [OF ptb at]) + thus "infinite {pi\b |b::'y. [(a,b)]\x \ x}" by (simp add: perm_set_def) next fix a assume "infinite {pi\b |b::'y. [(a,b)]\x \ x}" - hence "infinite (pi\{b::'y. [(a,b)]\x \ x})" by (simp add: perm_set_eq [OF ptb at]) + hence "infinite (pi\{b::'y. [(a,b)]\x \ x})" by (simp add: perm_set_def) thus "infinite {b::'y. [(a,b)]\x \ x}" by (simp add: pt_set_infinite_ineq[OF ptb, OF at]) qed @@ -1981,7 +1958,7 @@ shows "X supports X" proof - have "\a b. a\X \ b\X \ [(a,b)]\X = X" - by (auto simp add: perm_set_eq [OF at_pt_inst [OF at] at] at_calc[OF at]) + by (auto simp add: perm_set_def at_calc[OF at]) then show ?thesis by (simp add: supports_def) qed @@ -2008,7 +1985,7 @@ { fix a::"'x" assume asm: "a\X" hence "\b\(UNIV-X). [(a,b)]\X\X" - by (auto simp add: perm_set_eq [OF at_pt_inst [OF at] at] at_calc[OF at]) + by (auto simp add: perm_set_def at_calc[OF at]) with inf have "infinite {b\(UNIV-X). [(a,b)]\X\X}" by (rule infinite_Collection) hence "infinite {b. [(a,b)]\X\X}" by (rule_tac infinite_super, auto) hence "a\(supp X)" by (simp add: supp_def) @@ -2229,7 +2206,7 @@ proof (rule equalityI) case goal1 show "pi\(\x\X. f x) \ (\x\(pi\X). (pi\f) x)" - apply(auto simp add: perm_set_eq [OF pt at] perm_set_eq [OF at_pt_inst [OF at] at]) + apply(auto simp add: perm_set_def) apply(rule_tac x="pi\xb" in exI) apply(rule conjI) apply(rule_tac x="xb" in exI) @@ -2245,7 +2222,7 @@ next case goal2 show "(\x\(pi\X). (pi\f) x) \ pi\(\x\X. f x)" - apply(auto simp add: perm_set_eq [OF pt at] perm_set_eq [OF at_pt_inst [OF at] at]) + apply(auto simp add: perm_set_def) apply(rule_tac x="(rev pi)\x" in exI) apply(rule conjI) apply(simp add: pt_pi_rev[OF pt_x, OF at]) @@ -2278,7 +2255,7 @@ apply(rule allI)+ apply(rule impI) apply(erule conjE) - apply(simp add: perm_set_eq [OF pt at]) + apply(simp add: perm_set_def) apply(auto) apply(subgoal_tac "[(a,b)]\xa = xa")(*A*) apply(simp) @@ -2362,7 +2339,7 @@ also have "\ = (supp {x})\(supp X)" by (rule supp_fin_union[OF pt, OF at, OF fs], simp_all add: f) finally show "(supp (insert x X)) = (supp x)\((supp X)::'x set)" - by (simp add: supp_singleton [OF pt at]) + by (simp add: supp_singleton) qed lemma fresh_fin_union: @@ -2513,10 +2490,10 @@ apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) apply(drule_tac x="pi\xa" in bspec) apply(simp add: pt_set_bij1[OF ptb, OF at]) -apply(simp add: pt_set_eqvt [OF ptb at] pt_rev_pi[OF pt_list_inst[OF ptb], OF at]) +apply(simp add: set_eqvt pt_rev_pi[OF pt_list_inst[OF ptb], OF at]) apply(simp add: pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp]) apply(drule_tac x="(rev pi)\xa" in bspec) -apply(simp add: pt_set_bij1[OF ptb, OF at] pt_set_eqvt [OF ptb at]) +apply(simp add: pt_set_bij1[OF ptb, OF at] set_eqvt) apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) done @@ -2638,10 +2615,7 @@ have "({}::'a set) \ As = {}" by simp moreover have "set ([]::'a prm) \ {} \ {}" by simp - moreover - have "([]::'a prm)\{} = ({}::'a set)" - by (rule pt1 [OF pt_set_inst, OF at_pt_inst [OF at] at]) - ultimately show ?case by simp + ultimately show ?case by (simp add: empty_eqvt) next case (insert x Xs) then have ih: "\pi. (pi\Xs)\*c \ (pi\Xs) \ As = {} \ set pi \ Xs \ (pi\Xs)" by simp @@ -2655,7 +2629,7 @@ obtain y::"'a" where fr: "y\(c,pi\Xs,As)" apply(rule_tac at_exists_fresh[OF at, where x="(c,pi\Xs,As)"]) apply(auto simp add: supp_prod at_supp[OF at] at_fin_set_supp[OF at] - pt_supp_finite_pi[OF pt_set_inst[OF at_pt_inst[OF at] at] at]) + pt_supp_finite_pi[OF pt_set_inst[OF at_pt_inst[OF at]] at]) done have "({y}\(pi\Xs))\*c" using a1 fr by (simp add: fresh_star_def) moreover @@ -2671,20 +2645,18 @@ have eq: "[(pi\x,y)]\(pi\Xs) = (pi\Xs)" proof - have "(pi\x)\(pi\Xs)" using b d2 - by (simp add: pt_fresh_bij [OF pt_set_inst, OF at_pt_inst [OF at], - OF at, OF at] + by (simp add: pt_fresh_bij [OF pt_set_inst [OF at_pt_inst [OF at]], OF at] at_fin_set_fresh [OF at]) moreover have "y\(pi\Xs)" using fr by simp ultimately show "[(pi\x,y)]\(pi\Xs) = (pi\Xs)" - by (simp add: pt_fresh_fresh[OF pt_set_inst, - OF at_pt_inst[OF at], OF at, OF at]) + by (simp add: pt_fresh_fresh[OF pt_set_inst + [OF at_pt_inst[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]) + by (simp add: pt2[symmetric, OF pt_set_inst [OF at_pt_inst[OF at]]]) also have "\ = {y}\([(pi\x,y)]\(pi\Xs))" - by (simp only: union_eqvt perm_set_eq[OF at_pt_inst[OF at], OF at] at_calc[OF at])(auto) + by (simp only: union_eqvt perm_set_def at_calc[OF at])(auto) finally show "(((pi\x,y)#pi)\({x} \ Xs)) = {y}\(pi\Xs)" using eq by simp qed ultimately @@ -2714,6 +2686,17 @@ apply(auto) done +lemma cp_set_inst: + assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" + shows "cp TYPE ('a set) TYPE('x) TYPE('y)" +using c1 +apply(simp add: cp_def) +apply(auto) +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi2\xc" in exI) +apply(auto) +done + lemma cp_option_inst: assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" shows "cp TYPE ('a option) TYPE('x) TYPE('y)" @@ -3596,7 +3579,7 @@ plus_int_eqvt minus_int_eqvt mult_int_eqvt div_int_eqvt (* sets *) - union_eqvt empty_eqvt + union_eqvt empty_eqvt insert_eqvt set_eqvt (* the lemmas numeral_nat_eqvt numeral_int_eqvt do not conform with the *)