# HG changeset patch # User Christian Urban # Date 1282990523 -28800 # Node ID 76be909eaf048f2bed22ea7ca4651aab8cf2319e # Parent 63c936b09764bbb01e56fafb7e25673c3bc9b2c3 slight cleaning diff -r 63c936b09764 -r 76be909eaf04 Nominal/Abs.thy --- a/Nominal/Abs.thy Sat Aug 28 13:41:31 2010 +0800 +++ b/Nominal/Abs.thy Sat Aug 28 18:15:23 2010 +0800 @@ -568,16 +568,6 @@ and "Abs_res bs x = Abs_res cs y \ (\p. (bs, x) \res (op =) supp p (cs, y))" and "Abs_lst bsl x = Abs_lst csl y \ (\p. (bsl, x) \lst (op =) supp p (csl, y))" by (lifting alphas_abs) - -lemma split_rsp2[quot_respect]: "((R1 ===> R2 ===> prod_rel R1 R2 ===> op =) ===> - prod_rel R1 R2 ===> prod_rel R1 R2 ===> op =) split split" - by auto - -lemma split_prs2[quot_preserve]: - assumes q1: "Quotient R1 Abs1 Rep1" - and q2: "Quotient R2 Abs2 Rep2" - shows "((Abs1 ---> Abs2 ---> prod_fun Abs1 Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> prod_fun Rep1 Rep2 ---> id) split = split" - by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) *) diff -r 63c936b09764 -r 76be909eaf04 Nominal/Nominal2_FSet.thy --- a/Nominal/Nominal2_FSet.thy Sat Aug 28 13:41:31 2010 +0800 +++ b/Nominal/Nominal2_FSet.thy Sat Aug 28 18:15:23 2010 +0800 @@ -5,20 +5,16 @@ FSet begin -lemma "p \ {} = {}" -apply(perm_simp) -by simp - lemma permute_rsp_fset[quot_respect]: - "(op = ===> list_eq ===> list_eq) permute permute" - apply (simp add: eqvts[symmetric]) - apply clarify - apply (subst permute_minus_cancel(1)[symmetric, of "xb"]) - apply (subst mem_eqvt[symmetric]) - apply (subst (2) permute_minus_cancel(1)[symmetric, of "xb"]) - apply (subst mem_eqvt[symmetric]) - apply (erule_tac x="- x \ xb" in allE) - apply simp + shows "(op = ===> list_eq ===> list_eq) permute permute" + apply(simp) + apply(clarify) + apply(simp add: eqvts[symmetric]) + apply(subst permute_minus_cancel(1)[symmetric, of "xb"]) + apply(subst mem_eqvt[symmetric]) + apply(subst (2) permute_minus_cancel(1)[symmetric, of "xb"]) + apply(subst mem_eqvt[symmetric]) + apply(simp) done instantiation fset :: (pt) pt @@ -40,29 +36,14 @@ end -lemma "p \ {} = {}" -apply(perm_simp) -by simp - lemma permute_fset[simp]: fixes S::"('a::pt) fset" shows "(p \ {||}) = ({||} ::('a::pt) fset)" and "p \ finsert x S = finsert (p \ x) (p \ S)" by (lifting permute_list.simps) -lemma "p \ {} = {}" -apply(perm_simp) -by simp - -ML {* @{term "{}"} ; @{term "{||}"} *} - declare permute_fset[eqvt] -lemma "p \ {} = {}" -apply(perm_simp) -by simp - - lemma fmap_eqvt[eqvt]: shows "p \ (fmap f S) = fmap (p \ f) (p \ S)" by (lifting map_eqvt) @@ -80,6 +61,11 @@ unfolding supp_def by (perm_simp) (simp add: fset_cong) +lemma supp_fempty: + shows "supp {||} = {}" + unfolding supp_def + by simp + lemma supp_finsert: fixes x::"'a::fs" shows "supp (finsert x S) = supp x \ supp S" @@ -89,10 +75,6 @@ apply(simp add: supp_fset_to_set) done -lemma supp_fempty: - shows "supp {||} = {}" - unfolding supp_def - by simp instance fset :: (fs) fs apply (default) @@ -125,8 +107,9 @@ done lemma fresh_star_atom: - "fset_to_set s \* (a :: _ :: at_base) \ atom a \ fset_to_set s" - apply (induct s) + fixes a::"'a::at_base" + shows "fset_to_set S \* a \ atom a \ fset_to_set S" + apply (induct S) apply (simp add: fresh_set_empty) apply simp apply (unfold fresh_def) @@ -140,8 +123,4 @@ apply auto done -lemma "p \ {} = {}" -apply(perm_simp) -by simp - end