diff -r 7f0b4515588a -r daeb538c57bf src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Fri Sep 02 16:58:00 2011 -0700 +++ b/src/HOL/Nominal/Nominal.thy Sat Sep 03 17:32:34 2011 +0200 @@ -32,7 +32,7 @@ (* an auxiliary constant for the decision procedure involving *) (* permutations (to avoid loops when using perm-compositions) *) definition - "perm_aux pi x \ pi\x" + "perm_aux pi x = pi\x" (* overloaded permutation operations *) overloading @@ -51,61 +51,42 @@ begin definition - perm_fun_def: "perm_fun pi (f::'a\'b) \ (\x. pi\f((rev pi)\x))" - -fun - perm_bool :: "'x prm \ bool \ bool" -where + perm_fun_def: "perm_fun pi (f::'a\'b) = (\x. pi\f((rev pi)\x))" + +primrec perm_bool :: "'x prm \ bool \ bool" where true_eqvt: "perm_bool pi True = True" | false_eqvt: "perm_bool pi False = False" -fun - perm_unit :: "'x prm \ unit \ unit" -where +primrec perm_unit :: "'x prm \ unit \ unit" where "perm_unit pi () = ()" -fun - perm_prod :: "'x prm \ ('a\'b) \ ('a\'b)" -where +primrec perm_prod :: "'x prm \ ('a\'b) \ ('a\'b)" where "perm_prod pi (x,y) = (pi\x,pi\y)" -fun - perm_list :: "'x prm \ 'a list \ 'a list" -where +primrec perm_list :: "'x prm \ 'a list \ 'a list" where nil_eqvt: "perm_list pi [] = []" | cons_eqvt: "perm_list pi (x#xs) = (pi\x)#(pi\xs)" -fun - perm_option :: "'x prm \ 'a option \ 'a option" -where +primrec perm_option :: "'x prm \ 'a option \ 'a option" where some_eqvt: "perm_option pi (Some x) = Some (pi\x)" | none_eqvt: "perm_option pi None = None" -definition - perm_char :: "'x prm \ char \ char" -where - perm_char_def: "perm_char pi c \ c" - -definition - perm_nat :: "'x prm \ nat \ nat" -where - perm_nat_def: "perm_nat pi i \ i" - -definition - perm_int :: "'x prm \ int \ int" -where - perm_int_def: "perm_int pi i \ i" - -fun - perm_noption :: "'x prm \ 'a noption \ 'a noption" -where +definition perm_char :: "'x prm \ char \ char" where + perm_char_def: "perm_char pi c = c" + +definition perm_nat :: "'x prm \ nat \ nat" where + perm_nat_def: "perm_nat pi i = i" + +definition perm_int :: "'x prm \ int \ int" where + perm_int_def: "perm_int pi i = i" + +primrec perm_noption :: "'x prm \ 'a noption \ 'a noption" where nsome_eqvt: "perm_noption pi (nSome x) = nSome (pi\x)" | nnone_eqvt: "perm_noption pi nNone = nNone" -fun - perm_nprod :: "'x prm \ ('a, 'b) nprod \ ('a, 'b) nprod" -where +primrec perm_nprod :: "'x prm \ ('a, 'b) nprod \ ('a, 'b) nprod" where "perm_nprod pi (nPair x y) = nPair (pi\x) (pi\y)" + end @@ -188,18 +169,18 @@ (*==============================*) definition prm_eq :: "'x prm \ 'x prm \ bool" (" _ \ _ " [80,80] 80) where - "pi1 \ pi2 \ \a::'x. pi1\a = pi2\a" + "pi1 \ pi2 \ (\a::'x. pi1\a = pi2\a)" section {* Support, Freshness and Supports*} (*========================================*) definition supp :: "'a \ ('x set)" where - "supp x \ {a . (infinite {b . [(a,b)]\x \ x})}" + "supp x = {a . (infinite {b . [(a,b)]\x \ x})}" definition fresh :: "'x \ 'a \ bool" ("_ \ _" [80,80] 80) where - "a \ x \ a \ supp x" + "a \ x \ a \ supp x" definition supports :: "'x set \ 'a \ bool" (infixl "supports" 80) where - "S supports x \ \a b. (a\S \ b\S \ [(a,b)]\x=x)" + "S supports x \ (\a b. (a\S \ b\S \ [(a,b)]\x=x))" (* lemmas about supp *) lemma supp_fresh_iff: @@ -1731,7 +1712,7 @@ hence "infinite ({c. [(a,c)]\x = x \ [(b,c)]\x = x}-{a,b})" by (force dest: Diff_infinite_finite) hence "({c. [(a,c)]\x = x \ [(b,c)]\x = x}-{a,b}) \ {}" - by (metis Collect_def finite_set set_empty2) + by (metis finite_set set_empty2) hence "\c. c\({c. [(a,c)]\x = x \ [(b,c)]\x = x}-{a,b})" by (force) then obtain c where eq1: "[(a,c)]\x = x"