# HG changeset patch # User huffman # Date 1315087791 25200 # Node ID a9635943a3e969539d5fdfbaddbb16f18eaad54d # Parent 67b78d5dea5b231807652147327bd6a25d3bb6f3# Parent ccfc7c193d2b998079e8a0e643662ef6d7d3ca48 merged diff -r ccfc7c193d2b -r a9635943a3e9 src/HOL/Nominal/Examples/Pattern.thy --- a/src/HOL/Nominal/Examples/Pattern.thy Sat Sep 03 14:52:40 2011 -0700 +++ b/src/HOL/Nominal/Examples/Pattern.thy Sat Sep 03 15:09:51 2011 -0700 @@ -1,7 +1,7 @@ header {* Simply-typed lambda-calculus with let and tuple patterns *} theory Pattern -imports Nominal +imports "../Nominal" begin no_syntax @@ -622,7 +622,7 @@ (auto simp add: Cons_eq_append_conv fresh_star_def fresh_list_nil fresh_list_cons supp_list_nil supp_list_cons fresh_list_append supp_prod fresh_prod supp_atm fresh_atm - dest: notE [OF iffD1 [OF fresh_def [THEN meta_eq_to_obj_eq]]]) + dest: notE [OF iffD1 [OF fresh_def]]) lemma perm_mem_left: "(x::name) \ ((pi::name prm) \ A) \ (rev pi \ x) \ A" by (drule perm_boolI [of _ "rev pi"]) (simp add: eqvts perm_pi_simp) diff -r ccfc7c193d2b -r a9635943a3e9 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Sat Sep 03 14:52:40 2011 -0700 +++ b/src/HOL/Nominal/Nominal.thy Sat Sep 03 15:09:51 2011 -0700 @@ -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" diff -r ccfc7c193d2b -r a9635943a3e9 src/HOL/Nominal/nominal_atoms.ML --- a/src/HOL/Nominal/nominal_atoms.ML Sat Sep 03 14:52:40 2011 -0700 +++ b/src/HOL/Nominal/nominal_atoms.ML Sat Sep 03 15:09:51 2011 -0700 @@ -714,48 +714,48 @@ fold (fn ak_name => fn thy => let val qu_class = Sign.full_bname thy ("pt_"^ak_name); - val simp_s = HOL_basic_ss addsimps [defn]; + val simp_s = HOL_basic_ss addsimps [Simpdata.mk_eq defn]; val proof = EVERY [Class.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)]; in - AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy + AxClass.prove_arity (discrete_ty, [], [qu_class]) proof thy end) ak_names; fun discrete_fs_inst discrete_ty defn = fold (fn ak_name => fn thy => let val qu_class = Sign.full_bname thy ("fs_"^ak_name); - val supp_def = @{thm "Nominal.supp_def"}; - val simp_s = HOL_ss addsimps [supp_def,Collect_const,finite_emptyI,defn]; + val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"}; + val simp_s = HOL_ss addsimps [supp_def, Collect_const, finite_emptyI, Simpdata.mk_eq defn]; val proof = EVERY [Class.intro_classes_tac [], asm_simp_tac simp_s 1]; in - AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy + AxClass.prove_arity (discrete_ty, [], [qu_class]) proof thy end) ak_names; fun discrete_cp_inst discrete_ty defn = fold (fn ak_name' => (fold (fn ak_name => fn thy => let val qu_class = Sign.full_bname thy ("cp_"^ak_name^"_"^ak_name'); - val supp_def = @{thm "Nominal.supp_def"}; - val simp_s = HOL_ss addsimps [defn]; + val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"}; + val simp_s = HOL_ss addsimps [Simpdata.mk_eq defn]; val proof = EVERY [Class.intro_classes_tac [], asm_simp_tac simp_s 1]; in - AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy + AxClass.prove_arity (discrete_ty, [], [qu_class]) proof thy end) ak_names)) ak_names; in thy26 - |> discrete_pt_inst @{type_name nat} @{thm "perm_nat_def"} - |> discrete_fs_inst @{type_name nat} @{thm "perm_nat_def"} - |> discrete_cp_inst @{type_name nat} @{thm "perm_nat_def"} - |> discrete_pt_inst @{type_name bool} @{thm "perm_bool"} - |> discrete_fs_inst @{type_name bool} @{thm "perm_bool"} - |> discrete_cp_inst @{type_name bool} @{thm "perm_bool"} - |> discrete_pt_inst @{type_name int} @{thm "perm_int_def"} - |> discrete_fs_inst @{type_name int} @{thm "perm_int_def"} - |> discrete_cp_inst @{type_name int} @{thm "perm_int_def"} - |> discrete_pt_inst @{type_name char} @{thm "perm_char_def"} - |> discrete_fs_inst @{type_name char} @{thm "perm_char_def"} - |> discrete_cp_inst @{type_name char} @{thm "perm_char_def"} + |> discrete_pt_inst @{type_name nat} @{thm perm_nat_def} + |> discrete_fs_inst @{type_name nat} @{thm perm_nat_def} + |> discrete_cp_inst @{type_name nat} @{thm perm_nat_def} + |> discrete_pt_inst @{type_name bool} @{thm perm_bool} + |> discrete_fs_inst @{type_name bool} @{thm perm_bool} + |> discrete_cp_inst @{type_name bool} @{thm perm_bool} + |> discrete_pt_inst @{type_name int} @{thm perm_int_def} + |> discrete_fs_inst @{type_name int} @{thm perm_int_def} + |> discrete_cp_inst @{type_name int} @{thm perm_int_def} + |> discrete_pt_inst @{type_name char} @{thm perm_char_def} + |> discrete_fs_inst @{type_name char} @{thm perm_char_def} + |> discrete_cp_inst @{type_name char} @{thm perm_char_def} end; diff -r ccfc7c193d2b -r a9635943a3e9 src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Sat Sep 03 14:52:40 2011 -0700 +++ b/src/HOL/Nominal/nominal_datatype.ML Sat Sep 03 15:09:51 2011 -0700 @@ -155,9 +155,9 @@ val supp_prod = @{thm supp_prod}; val fresh_prod = @{thm fresh_prod}; val supports_fresh = @{thm supports_fresh}; -val supports_def = @{thm Nominal.supports_def}; -val fresh_def = @{thm fresh_def}; -val supp_def = @{thm supp_def}; +val supports_def = Simpdata.mk_eq @{thm Nominal.supports_def}; +val fresh_def = Simpdata.mk_eq @{thm fresh_def}; +val supp_def = Simpdata.mk_eq @{thm supp_def}; val rev_simps = @{thms rev.simps}; val app_simps = @{thms append.simps}; val at_fin_set_supp = @{thm at_fin_set_supp}; @@ -306,7 +306,7 @@ val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types); - val perm_fun_def = Global_Theory.get_thm thy2 "perm_fun_def"; + val perm_fun_def = Simpdata.mk_eq @{thm perm_fun_def}; val unfolded_perm_eq_thms = if length descr = length new_type_names then [] diff -r ccfc7c193d2b -r a9635943a3e9 src/HOL/Nominal/nominal_permeq.ML --- a/src/HOL/Nominal/nominal_permeq.ML Sat Sep 03 14:52:40 2011 -0700 +++ b/src/HOL/Nominal/nominal_permeq.ML Sat Sep 03 15:09:51 2011 -0700 @@ -56,10 +56,10 @@ val finite_Un = @{thm "finite_Un"}; val conj_absorb = @{thm "conj_absorb"}; val not_false = @{thm "not_False_eq_True"} -val perm_fun_def = @{thm "Nominal.perm_fun_def"}; +val perm_fun_def = Simpdata.mk_eq @{thm "Nominal.perm_fun_def"}; val perm_eq_app = @{thm "Nominal.pt_fun_app_eq"}; -val supports_def = @{thm "Nominal.supports_def"}; -val fresh_def = @{thm "Nominal.fresh_def"}; +val supports_def = Simpdata.mk_eq @{thm "Nominal.supports_def"}; +val fresh_def = Simpdata.mk_eq @{thm "Nominal.fresh_def"}; val fresh_prod = @{thm "Nominal.fresh_prod"}; val fresh_unit = @{thm "Nominal.fresh_unit"}; val supports_rule = @{thm "supports_finite"}; @@ -130,7 +130,7 @@ case redex of (* case pi o f == (%x. pi o (f ((rev pi)o x))) *) (Const("Nominal.perm",_) $ pi $ f) => - (if (applicable_fun f) then SOME (perm_fun_def) else NONE) + (if (applicable_fun f) then SOME perm_fun_def else NONE) | _ => NONE end