diff -r a44479bde681 -r 017e33849f4d Nominal/Nominal2_Base.thy --- a/Nominal/Nominal2_Base.thy Tue Mar 22 12:18:30 2016 +0000 +++ b/Nominal/Nominal2_Base.thy Thu Apr 19 13:57:17 2018 +0100 @@ -7,7 +7,8 @@ theory Nominal2_Base imports "~~/src/HOL/Library/Old_Datatype" "~~/src/HOL/Library/Infinite_Set" - "~~/src/HOL/Quotient_Examples/FSet" + "~~/src/HOL/Library/Multiset" + "~~/src/HOL/Library/FSet" "~~/src/HOL/Library/FinFun" keywords "atom_decl" "equivariance" :: thy_decl @@ -18,7 +19,7 @@ section {* Atoms and Sorts *} -text {* A simple implementation for atom_sorts is strings. *} +text {* A simple implementation for @{text atom_sorts} is strings. *} (* types atom_sort = string *) text {* To deal with Church-like binding we use trees of @@ -611,7 +612,7 @@ lemma permute_multiset [simp]: fixes M N::"('a::pt) multiset" shows "(p \ {#}) = ({#} ::('a::pt) multiset)" - and "(p \ {# x #}) = {# p \ x #}" + and "(p \ add_mset x M) = add_mset (p \ x) (p \ M)" and "(p \ (M + N)) = (p \ M) + (p \ N)" unfolding permute_multiset_def by (simp_all) @@ -619,38 +620,40 @@ subsection {* Permutations for @{typ "'a fset"} *} -lemma permute_fset_rsp[quot_respect]: - shows "(op = ===> list_eq ===> list_eq) permute permute" - unfolding rel_fun_def - by (simp add: set_eqvt[symmetric]) - instantiation fset :: (pt) pt begin -quotient_definition - "permute_fset :: perm \ 'a fset \ 'a fset" -is - "permute :: perm \ 'a list \ 'a list" - by (simp add: set_eqvt[symmetric]) - +context includes fset.lifting begin +lift_definition + "permute_fset" :: "perm \ 'a fset \ 'a fset" +is "permute :: perm \ 'a set \ 'a set" by simp +end + +context includes fset.lifting begin instance proof fix x :: "'a fset" and p q :: "perm" - show "0 \ x = x" by (descending) (simp) - show "(p + q) \ x = p \ q \ x" by (descending) (simp) + show "0 \ x = x" by transfer simp + show "(p + q) \ x = p \ q \ x" by transfer simp qed +end end +context includes fset.lifting +begin lemma permute_fset [simp]: fixes S::"('a::pt) fset" shows "(p \ {||}) = ({||} ::('a::pt) fset)" - and "(p \ insert_fset x S) = insert_fset (p \ x) (p \ S)" - by (lifting permute_list.simps) + and "(p \ finsert x S) = finsert (p \ x) (p \ S)" + apply (transfer, simp add: empty_eqvt) + apply (transfer, simp add: insert_eqvt) + done lemma fset_eqvt: shows "p \ (fset S) = fset (p \ S)" - by (lifting set_eqvt) + by transfer simp +end subsection {* Permutations for @{typ "('a, 'b) finfun"} *} @@ -762,7 +765,7 @@ proof qed (rule permute_int_def) -section {* Infrastructure for Equivariance and Perm_simp *} +section {* Infrastructure for Equivariance and @{text Perm_simp} *} subsection {* Basic functions about permutations *} @@ -792,7 +795,7 @@ (* multisets *) permute_multiset -subsection {* perm_simp infrastructure *} +subsection {* @{text perm_simp} infrastructure *} definition "unpermute p = permute (- p)" @@ -812,7 +815,7 @@ shows "p \ unpermute p x \ x" unfolding unpermute_def by simp -text {* provides perm_simp methods *} +text {* provides @{text perm_simp} methods *} ML_file "nominal_permeq.ML" @@ -974,10 +977,6 @@ unfolding permute_set_eq permute_fun_def by (auto simp: permute_bool_def) -lemma inter_eqvt [eqvt]: - shows "p \ (A \ B) = (p \ A) \ (p \ B)" - unfolding Int_def by simp - lemma Bex_eqvt [eqvt]: shows "p \ (\x \ S. P x) = (\x \ (p \ S). (p \ P) x)" unfolding Bex_def by simp @@ -999,14 +998,22 @@ unfolding UNIV_def by (perm_simp) (rule refl) +lemma inter_eqvt [eqvt]: + shows "p \ (A \ B) = (p \ A) \ (p \ B)" + unfolding Int_def by simp + +lemma Inter_eqvt [eqvt]: + shows "p \ \S = \(p \ S)" + unfolding Inter_eq by simp + lemma union_eqvt [eqvt]: shows "p \ (A \ B) = (p \ A) \ (p \ B)" unfolding Un_def by simp -lemma UNION_eqvt [eqvt]: - shows "p \ (UNION A f) = (UNION (p \ A) (p \ f))" -unfolding UNION_eq -by (perm_simp) (simp) +lemma Union_eqvt [eqvt]: + shows "p \ \A = \(p \ A)" + unfolding Union_eq + by perm_simp rule lemma Diff_eqvt [eqvt]: fixes A B :: "'a::pt set" @@ -1030,16 +1037,6 @@ shows "p \ (f -` A) = (p \ f) -` (p \ A)" unfolding vimage_def by simp -lemma Union_eqvt [eqvt]: - shows "p \ (\ S) = \ (p \ S)" - unfolding Union_eq by simp - -lemma Inter_eqvt [eqvt]: - shows "p \ (\ S) = \ (p \ S)" - unfolding Inter_eq by simp - -thm foldr.simps - lemma foldr_eqvt[eqvt]: "p \ foldr f xs = foldr (p \ f) (p \ xs)" apply(induct xs) @@ -1052,7 +1049,6 @@ lemma Sigma_eqvt: shows "(p \ (X \ Y)) = (p \ X) \ (p \ Y)" unfolding Sigma_def -unfolding SUP_def by (perm_simp) (rule refl) text {* @@ -1109,7 +1105,7 @@ instance apply standard -unfolding Inf_fun_def INF_def +unfolding Inf_fun_def apply(perm_simp) apply(rule refl) done @@ -1200,9 +1196,10 @@ subsubsection {* Equivariance for @{typ "'a fset"} *} +context includes fset.lifting begin lemma in_fset_eqvt [eqvt]: shows "(p \ (x |\| S)) = ((p \ x) |\| (p \ S))" -unfolding in_fset by simp + by transfer simp lemma union_fset_eqvt [eqvt]: shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" @@ -1210,21 +1207,16 @@ lemma inter_fset_eqvt [eqvt]: shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" - apply(descending) - unfolding list_eq_def inter_list_def - apply(simp) - done + by transfer simp lemma subset_fset_eqvt [eqvt]: shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" - apply(descending) - unfolding sub_list_def - apply(simp) - done + by transfer simp lemma map_fset_eqvt [eqvt]: - shows "p \ (map_fset f S) = map_fset (p \ f) (p \ S)" - by (lifting map_eqvt) + shows "p \ (f |`| S) = (p \ f) |`| (p \ S)" + by transfer simp +end subsubsection {* Equivariance for @{typ "('a, 'b) finfun"} *} @@ -1906,7 +1898,7 @@ end *} -subsection {* helper functions for nominal_functions *} +subsection {* helper functions for @{text nominal_functions} *} lemma THE_defaultI2: assumes "\!x. P x" "\x. P x \ Q x" @@ -2146,30 +2138,22 @@ qed lemma Union_supports_multiset: - shows "\{supp x | x. x :# M} supports M" + shows "\{supp x | x. x \# M} supports M" proof - - have sw: "\a b. ((\x. x :# M \ (a \ b) \ x = x) \ (a \ b) \ M = M)" - unfolding permute_multiset_def - apply(induct M) - apply(simp_all) - done - show "(\{supp x | x. x :# M}) supports M" - unfolding supports_def - apply(clarify) - apply(rule sw) - apply(rule swap_fresh_fresh) - apply(simp_all only: fresh_def) - apply(auto) - apply(metis neq0_conv)+ - done + have sw: "\a b. ((\x. x \# M \ (a \ b) \ x = x) \ (a \ b) \ M = M)" + unfolding permute_multiset_def by (induct M) simp_all + have "(\x\set_mset M. supp x) supports M" + by (auto intro!: sw swap_fresh_fresh simp add: fresh_def supports_def) + also have "(\x\set_mset M. supp x) = (\{supp x | x. x \# M})" + by auto + finally show "(\{supp x | x. x \# M}) supports M" . qed lemma Union_included_multiset: fixes M::"('a::fs multiset)" shows "(\{supp x | x. x \# M}) \ supp M" proof - - have "(\{supp x | x. x \# M}) = (\{supp x | x. x \ set_mset M})" by simp - also have "... \ (\x \ set_mset M. supp x)" by auto + have "(\{supp x | x. x \# M}) = (\x \ set_mset M. supp x)" by auto also have "... = supp (set_mset M)" by (simp add: supp_of_finite_sets) also have " ... \ supp M" by (rule supp_set_mset) @@ -2178,7 +2162,7 @@ lemma supp_of_multisets: fixes M::"('a::fs multiset)" - shows "(supp M) = (\{supp x | x. x :# M})" + shows "(supp M) = (\{supp x | x. x \# M})" apply(rule subset_antisym) apply(rule supp_is_subset) apply(rule Union_supports_multiset) @@ -2221,20 +2205,20 @@ unfolding fresh_def by (simp) -lemma supp_insert_fset [simp]: +lemma supp_finsert [simp]: fixes x::"'a::fs" and S::"'a fset" - shows "supp (insert_fset x S) = supp x \ supp S" + shows "supp (finsert x S) = supp x \ supp S" apply(subst supp_fset[symmetric]) apply(simp add: supp_of_finite_insert) done -lemma fresh_insert_fset: +lemma fresh_finsert: fixes x::"'a::fs" and S::"'a fset" - shows "a \ insert_fset x S \ a \ x \ a \ S" + shows "a \ finsert x S \ a \ x \ a \ S" unfolding fresh_def - by (simp) + by simp lemma fset_finite_supp: fixes S::"('a::fs) fset" @@ -3121,7 +3105,7 @@ unfolding atom_eqvt [symmetric] by simp_all -text {* the following two lemmas do not hold for at_base, +text {* the following two lemmas do not hold for @{text at_base}, only for single sort atoms from at *} lemma flip_triple: