# HG changeset patch # User Andreas Lochbihler # Date 1338360494 -7200 # Node ID 1c5171abe5cc16af81aedff8299d9356526a8fdf # Parent 65eb8910a7797a33f51e8ae24570536f9cba68f0 removed subscripts from FinFun type syntax diff -r 65eb8910a779 -r 1c5171abe5cc src/HOL/Library/FinFun.thy --- a/src/HOL/Library/FinFun.thy Wed May 30 08:34:14 2012 +0200 +++ b/src/HOL/Library/FinFun.thy Wed May 30 08:48:14 2012 +0200 @@ -83,7 +83,7 @@ definition "finfun = {f::'a\'b. \b. finite {a. f a \ b}}" -typedef (open) ('a,'b) finfun ("(_ \\<^isub>f /_)" [22, 21] 21) = "finfun :: ('a => 'b) set" +typedef (open) ('a,'b) finfun ("(_ =>f /_)" [22, 21] 21) = "finfun :: ('a => 'b) set" morphisms finfun_apply Abs_finfun proof - have "\f. finite {x. f x \ undefined}" @@ -93,6 +93,8 @@ then show ?thesis unfolding finfun_def by auto qed +type_notation finfun ("(_ \f /_)" [22, 21] 21) + setup_lifting type_definition_finfun lemma fun_upd_finfun: "y(a := b) \ finfun \ y \ finfun" @@ -220,7 +222,7 @@ lemma Abs_finfun_inj_finite: assumes fin: "finite (UNIV :: 'a set)" - shows "inj (Abs_finfun :: ('a \ 'b) \ 'a \\<^isub>f 'b)" + shows "inj (Abs_finfun :: ('a \ 'b) \ 'a \f 'b)" proof(rule inj_onI) fix x y :: "'a \ 'b" assume "Abs_finfun x = Abs_finfun y" @@ -274,12 +276,13 @@ qed -subsection {* Kernel functions for type @{typ "'a \\<^isub>f 'b"} *} +subsection {* Kernel functions for type @{typ "'a \f 'b"} *} -lift_definition finfun_const :: "'b \ 'a \\<^isub>f 'b" ("\\<^isup>f/ _" [0] 1) +lift_definition finfun_const :: "'b \ 'a \f 'b" ("\\<^isup>f/ _" [0] 1) is "\ b x. b" by (rule const_finfun) -lift_definition finfun_update :: "'a \\<^isub>f 'b \ 'a \ 'b \ 'a \\<^isub>f 'b" ("_'(\<^sup>f/ _ := _')" [1000,0,0] 1000) is "fun_upd" by (simp add: fun_upd_finfun) +lift_definition finfun_update :: "'a \f 'b \ 'a \ 'b \ 'a \f 'b" ("_'(\<^sup>f/ _ := _')" [1000,0,0] 1000) is "fun_upd" +by (simp add: fun_upd_finfun) lemma finfun_update_twist: "a \ a' \ f(\<^sup>f a := b)(\<^sup>f a' := b') = f(\<^sup>f a' := b')(\<^sup>f a := b)" by transfer (simp add: fun_upd_twist) @@ -293,7 +296,7 @@ subsection {* Code generator setup *} -definition finfun_update_code :: "'a \\<^isub>f 'b \ 'a \ 'b \ 'a \\<^isub>f 'b" ("_'(\<^sup>fc/ _ := _')" [1000,0,0] 1000) +definition finfun_update_code :: "'a \f 'b \ 'a \ 'b \ 'a \f 'b" ("_'(\<^sup>fc/ _ := _')" [1000,0,0] 1000) where [simp, code del]: "finfun_update_code = finfun_update" code_datatype finfun_const finfun_update_code @@ -309,7 +312,7 @@ subsection {* Setup for quickcheck *} -quickcheck_generator finfun constructors: finfun_update_code, "finfun_const :: 'b => 'a \\<^isub>f 'b" +quickcheck_generator finfun constructors: finfun_update_code, "finfun_const :: 'b \ 'a \f 'b" subsection {* @{text "finfun_update"} as instance of @{text "comp_fun_commute"} *} @@ -414,7 +417,7 @@ case True thus ?thesis by(simp add: finfun_default_aux_def) qed -lift_definition finfun_default :: "'a \\<^isub>f 'b \ 'b" +lift_definition finfun_default :: "'a \f 'b \ 'b" is "finfun_default_aux" .. lemma finfun_apply_transfer [transfer_rule]: @@ -424,7 +427,7 @@ lemma finite_finfun_default: "finite {a. finfun_apply f a \ finfun_default f}" by transfer (erule finite_finfun_default_aux) -lemma finfun_default_const: "finfun_default ((\\<^isup>f b) :: 'a \\<^isub>f 'b) = (if finite (UNIV :: 'a set) then undefined else b)" +lemma finfun_default_const: "finfun_default ((\\<^isup>f b) :: 'a \f 'b) = (if finite (UNIV :: 'a set) then undefined else b)" by(transfer)(auto simp add: finfun_default_aux_infinite finfun_default_aux_def) lemma finfun_default_update_const: @@ -432,7 +435,7 @@ by transfer (simp add: finfun_default_aux_update_const) lemma finfun_default_const_code [code]: - "finfun_default ((\\<^isup>f c) :: ('a :: card_UNIV) \\<^isub>f 'b) = (if card_UNIV (TYPE('a)) = 0 then c else undefined)" + "finfun_default ((\\<^isup>f c) :: ('a :: card_UNIV) \f 'b) = (if card_UNIV (TYPE('a)) = 0 then c else undefined)" by(simp add: finfun_default_const card_UNIV_eq_0_infinite_UNIV) lemma finfun_default_update_code [code]: @@ -441,7 +444,7 @@ subsection {* Recursion combinator and well-formedness conditions *} -definition finfun_rec :: "('b \ 'c) \ ('a \ 'b \ 'c \ 'c) \ ('a \\<^isub>f 'b) \ 'c" +definition finfun_rec :: "('b \ 'c) \ ('a \ 'b \ 'c \ 'c) \ ('a \f 'b) \ 'c" where [code del]: "finfun_rec cnst upd f \ let b = finfun_default f; @@ -740,7 +743,7 @@ "finfun_rec cnst upd (\\<^isup>f c) = cnst c" proof(cases "finite (UNIV :: 'a set)") case False - hence "finfun_default ((\\<^isup>f c) :: 'a \\<^isub>f 'b) = c" by(simp add: finfun_default_const) + hence "finfun_default ((\\<^isup>f c) :: 'a \f 'b) = c" by(simp add: finfun_default_const) moreover have "(THE g :: 'a \ 'b. (\\<^isup>f c) = Abs_finfun (map_default c g) \ finite (dom g) \ c \ ran g) = empty" proof (rule the_equality) show "(\\<^isup>f c) = Abs_finfun (map_default c empty) \ finite (dom empty) \ c \ ran empty" @@ -757,7 +760,7 @@ ultimately show ?thesis by(simp add: finfun_rec_def) next case True - hence default: "finfun_default ((\\<^isup>f c) :: 'a \\<^isub>f 'b) = undefined" by(simp add: finfun_default_const) + hence default: "finfun_default ((\\<^isup>f c) :: 'a \f 'b) = undefined" by(simp add: finfun_default_const) let ?the = "\g :: 'a \ 'b. (\\<^isup>f c) = Abs_finfun (map_default undefined g) \ finite (dom g) \ undefined \ ran g" show ?thesis proof(cases "c = undefined") @@ -840,14 +843,14 @@ by(atomize_elim)(rule finfun_exhaust_disj) lemma finfun_rec_unique: - fixes f :: "'a \\<^isub>f 'b \ 'c" + fixes f :: "'a \f 'b \ 'c" assumes c: "\c. f (\\<^isup>f c) = cnst c" and u: "\g a b. f (g(\<^sup>f a := b)) = upd g a b (f g)" and c': "\c. f' (\\<^isup>f c) = cnst c" and u': "\g a b. f' (g(\<^sup>f a := b)) = upd g a b (f' g)" shows "f = f'" proof - fix g :: "'a \\<^isub>f 'b" + fix g :: "'a \f 'b" show "f g = f' g" by(induct g rule: finfun_weak_induct)(auto simp add: c u c' u') qed @@ -910,7 +913,7 @@ subsection {* Function composition *} -definition finfun_comp :: "('a \ 'b) \ 'c \\<^isub>f 'a \ 'c \\<^isub>f 'b" (infixr "\\<^isub>f" 55) +definition finfun_comp :: "('a \ 'b) \ 'c \f 'a \ 'c \f 'b" (infixr "\\<^isub>f" 55) where [code del]: "g \\<^isub>f f = finfun_rec (\b. (\\<^isup>f g b)) (\a b c. c(\<^sup>f a := g b)) f" interpretation finfun_comp_aux: finfun_rec_wf_aux "(\b. (\\<^isup>f g b))" "(\a b c. c(\<^sup>f a := g b))" @@ -968,7 +971,7 @@ thus ?thesis by(auto simp add: fun_eq_iff) qed -definition finfun_comp2 :: "'b \\<^isub>f 'c \ ('a \ 'b) \ 'a \\<^isub>f 'c" (infixr "\<^sub>f\" 55) +definition finfun_comp2 :: "'b \f 'c \ ('a \ 'b) \ 'a \f 'c" (infixr "\<^sub>f\" 55) where [code del]: "finfun_comp2 g f = Abs_finfun (finfun_apply g \ f)" lemma finfun_comp2_const [code, simp]: "finfun_comp2 (\\<^isup>f c) f = (\\<^isup>f c)" @@ -991,7 +994,7 @@ subsection {* Universal quantification *} -definition finfun_All_except :: "'a list \ 'a \\<^isub>f bool \ bool" +definition finfun_All_except :: "'a list \ 'a \f bool \ bool" where [code del]: "finfun_All_except A P \ \a. a \ set A \ P\<^sub>f a" lemma finfun_All_except_const: "finfun_All_except A (\\<^isup>f b) \ b \ set A = UNIV" @@ -1010,7 +1013,7 @@ shows "finfun_All_except A (finfun_update_code f a b) = ((a \ set A \ b) \ finfun_All_except (a # A) f)" by(simp add: finfun_All_except_update) -definition finfun_All :: "'a \\<^isub>f bool \ bool" +definition finfun_All :: "'a \f bool \ bool" where "finfun_All = finfun_All_except []" lemma finfun_All_const [simp]: "finfun_All (\\<^isup>f b) = b" @@ -1023,7 +1026,7 @@ by(simp add: finfun_All_def finfun_All_except_def) -definition finfun_Ex :: "'a \\<^isub>f bool \ bool" +definition finfun_Ex :: "'a \f bool \ bool" where "finfun_Ex P = Not (finfun_All (Not \\<^isub>f P))" lemma finfun_Ex_Ex: "finfun_Ex P = Ex P\<^sub>f" @@ -1035,7 +1038,7 @@ subsection {* A diagonal operator for FinFuns *} -definition finfun_Diag :: "'a \\<^isub>f 'b \ 'a \\<^isub>f 'c \ 'a \\<^isub>f ('b \ 'c)" ("(1'(_,/ _')\<^sup>f)" [0, 0] 1000) +definition finfun_Diag :: "'a \f 'b \ 'a \f 'c \ 'a \f ('b \ 'c)" ("(1'(_,/ _')\<^sup>f)" [0, 0] 1000) where [code del]: "finfun_Diag f g = finfun_rec (\b. Pair b \\<^isub>f g) (\a b c. c(\<^sup>f a := (b, g\<^sub>f a))) f" interpretation finfun_Diag_aux: finfun_rec_wf_aux "\b. Pair b \\<^isub>f g" "\a b c. c(\<^sup>f a := (b, g\<^sub>f a))" @@ -1100,7 +1103,7 @@ "(f, g)\<^sup>f = Abs_finfun ((\x. (finfun_apply f x, finfun_apply g x)))" including finfun proof - - have "(\f :: 'a \\<^isub>f 'b. (f, g)\<^sup>f) = (\f. Abs_finfun ((\x. (finfun_apply f x, finfun_apply g x))))" + have "(\f :: 'a \f 'b. (f, g)\<^sup>f) = (\f. Abs_finfun ((\x. (finfun_apply f x, finfun_apply g x))))" proof(rule finfun_rec_unique) { fix c show "Abs_finfun (\x. (finfun_apply (\\<^isup>f c) x, finfun_apply g x)) = Pair c \\<^isub>f g" by(simp add: finfun_comp_conv_comp o_def finfun_const_def) } @@ -1115,7 +1118,7 @@ lemma finfun_Diag_eq: "(f, g)\<^sup>f = (f', g')\<^sup>f \ f = f' \ g = g'" by(auto simp add: expand_finfun_eq fun_eq_iff) -definition finfun_fst :: "'a \\<^isub>f ('b \ 'c) \ 'a \\<^isub>f 'b" +definition finfun_fst :: "'a \f ('b \ 'c) \ 'a \f 'b" where [code]: "finfun_fst f = fst \\<^isub>f f" lemma finfun_fst_const: "finfun_fst (\\<^isup>f bc) = (\\<^isup>f fst bc)" @@ -1135,7 +1138,7 @@ by(simp add: finfun_fst_def [abs_def] finfun_comp_conv_comp) -definition finfun_snd :: "'a \\<^isub>f ('b \ 'c) \ 'a \\<^isub>f 'c" +definition finfun_snd :: "'a \f ('b \ 'c) \ 'a \f 'c" where [code]: "finfun_snd f = snd \\<^isub>f f" lemma finfun_snd_const: "finfun_snd (\\<^isup>f bc) = (\\<^isup>f snd bc)" @@ -1161,7 +1164,7 @@ subsection {* Currying for FinFuns *} -definition finfun_curry :: "('a \ 'b) \\<^isub>f 'c \ 'a \\<^isub>f 'b \\<^isub>f 'c" +definition finfun_curry :: "('a \ 'b) \f 'c \ 'a \f 'b \f 'c" where [code del]: "finfun_curry = finfun_rec (finfun_const \ finfun_const) (\(a, b) c f. f(\<^sup>f a := (f\<^sub>f a)(\<^sup>f b := c)))" interpretation finfun_curry_aux: finfun_rec_wf_aux "finfun_const \ finfun_const" "\(a, b) c f. f(\<^sup>f a := (f\<^sub>f a)(\<^sup>f b := c))" @@ -1211,22 +1214,23 @@ qed lemma finfun_curry_conv_curry: - fixes f :: "('a \ 'b) \\<^isub>f 'c" + fixes f :: "('a \ 'b) \f 'c" shows "finfun_curry f = Abs_finfun (\a. Abs_finfun (curry (finfun_apply f) a))" including finfun proof - - have "finfun_curry = (\f :: ('a \ 'b) \\<^isub>f 'c. Abs_finfun (\a. Abs_finfun (curry (finfun_apply f) a)))" + have "finfun_curry = (\f :: ('a \ 'b) \f 'c. Abs_finfun (\a. Abs_finfun (curry (finfun_apply f) a)))" proof(rule finfun_rec_unique) - { fix c show "finfun_curry (\\<^isup>f c) = (\\<^isup>f \\<^isup>f c)" by simp } - { fix f a c show "finfun_curry (f(\<^sup>f a := c)) = (finfun_curry f)(\<^sup>f fst a := ((finfun_curry f)\<^sub>f (fst a))(\<^sup>f snd a := c))" - by(cases a) simp } - { fix c show "Abs_finfun (\a. Abs_finfun (curry (finfun_apply (\\<^isup>f c)) a)) = (\\<^isup>f \\<^isup>f c)" - by(simp add: finfun_curry_def finfun_const_def curry_def) } - { fix g a b - show "Abs_finfun (\aa. Abs_finfun (curry (finfun_apply g(\<^sup>f a := b)) aa)) = - (Abs_finfun (\a. Abs_finfun (curry (finfun_apply g) a)))(\<^sup>f - fst a := ((Abs_finfun (\a. Abs_finfun (curry (finfun_apply g) a)))\<^sub>f (fst a))(\<^sup>f snd a := b))" - by(cases a)(auto intro!: ext arg_cong[where f=Abs_finfun] simp add: finfun_curry_def finfun_update_def finfun_curry finfun_Abs_finfun_curry) } + fix c show "finfun_curry (\\<^isup>f c) = (\\<^isup>f \\<^isup>f c)" by simp + fix f a + show "finfun_curry (f(\<^sup>f a := c)) = (finfun_curry f)(\<^sup>f fst a := ((finfun_curry f)\<^sub>f (fst a))(\<^sup>f snd a := c))" + by(cases a) simp + show "Abs_finfun (\a. Abs_finfun (curry (finfun_apply (\\<^isup>f c)) a)) = (\\<^isup>f \\<^isup>f c)" + by(simp add: finfun_curry_def finfun_const_def curry_def) + fix g b + show "Abs_finfun (\aa. Abs_finfun (curry (finfun_apply g(\<^sup>f a := b)) aa)) = + (Abs_finfun (\a. Abs_finfun (curry (finfun_apply g) a)))(\<^sup>f + fst a := ((Abs_finfun (\a. Abs_finfun (curry (finfun_apply g) a)))\<^sub>f (fst a))(\<^sup>f snd a := b))" + by(cases a)(auto intro!: ext arg_cong[where f=Abs_finfun] simp add: finfun_curry_def finfun_update_def finfun_curry finfun_Abs_finfun_curry) qed thus ?thesis by(auto simp add: fun_eq_iff) qed @@ -1242,12 +1246,12 @@ end lemma [code nbe]: - "HOL.equal (f :: _ \\<^isub>f _) f \ True" + "HOL.equal (f :: _ \f _) f \ True" by (fact equal_refl) subsection {* An operator that explicitly removes all redundant updates in the generated representations *} -definition finfun_clearjunk :: "'a \\<^isub>f 'b \ 'a \\<^isub>f 'b" +definition finfun_clearjunk :: "'a \f 'b \ 'a \f 'b" where [simp, code del]: "finfun_clearjunk = id" lemma finfun_clearjunk_const [code]: "finfun_clearjunk (\\<^isup>f b) = (\\<^isup>f b)" @@ -1258,11 +1262,11 @@ subsection {* The domain of a FinFun as a FinFun *} -definition finfun_dom :: "('a \\<^isub>f 'b) \ ('a \\<^isub>f bool)" +definition finfun_dom :: "('a \f 'b) \ ('a \f bool)" where [code del]: "finfun_dom f = Abs_finfun (\a. f\<^sub>f a \ finfun_default f)" lemma finfun_dom_const: - "finfun_dom ((\\<^isup>f c) :: 'a \\<^isub>f 'b) = (\\<^isup>f finite (UNIV :: 'a set) \ c \ undefined)" + "finfun_dom ((\\<^isup>f c) :: 'a \f 'b) = (\\<^isup>f finite (UNIV :: 'a set) \ c \ undefined)" unfolding finfun_dom_def finfun_default_const by(auto)(simp_all add: finfun_const_def) @@ -1272,7 +1276,7 @@ *} lemma finfun_dom_const_code [code]: - "finfun_dom ((\\<^isup>f c) :: ('a :: card_UNIV) \\<^isub>f 'b) = + "finfun_dom ((\\<^isup>f c) :: ('a :: card_UNIV) \f 'b) = (if card_UNIV (TYPE('a)) = 0 then (\\<^isup>f False) else FinFun.code_abort (\_. finfun_dom (\\<^isup>f c)))" unfolding card_UNIV_eq_0_infinite_UNIV by(simp add: finfun_dom_const) @@ -1310,7 +1314,7 @@ subsection {* The domain of a FinFun as a sorted list *} -definition finfun_to_list :: "('a :: linorder) \\<^isub>f 'b \ 'a list" +definition finfun_to_list :: "('a :: linorder) \f 'b \ 'a list" where "finfun_to_list f = (THE xs. set xs = {x. (finfun_dom f)\<^sub>f x} \ sorted xs \ distinct xs)" @@ -1331,13 +1335,13 @@ by auto lemma finfun_to_list_const: - "finfun_to_list ((\\<^isup>f c) :: ('a :: {linorder} \\<^isub>f 'b)) = + "finfun_to_list ((\\<^isup>f c) :: ('a :: {linorder} \f 'b)) = (if \ finite (UNIV :: 'a set) \ c = undefined then [] else THE xs. set xs = UNIV \ sorted xs \ distinct xs)" by(auto simp add: finfun_to_list_def finfun_const_False_conv_bot finfun_const_True_conv_top finfun_dom_const) lemma finfun_to_list_const_code [code]: - "finfun_to_list ((\\<^isup>f c) :: ('a :: {linorder, card_UNIV} \\<^isub>f 'b)) = - (if card_UNIV (TYPE('a)) = 0 then [] else FinFun.code_abort (\_. finfun_to_list ((\\<^isup>f c) :: ('a \\<^isub>f 'b))))" + "finfun_to_list ((\\<^isup>f c) :: ('a :: {linorder, card_UNIV} \f 'b)) = + (if card_UNIV (TYPE('a)) = 0 then [] else FinFun.code_abort (\_. finfun_to_list ((\\<^isup>f c) :: ('a \f 'b))))" unfolding card_UNIV_eq_0_infinite_UNIV by(auto simp add: finfun_to_list_const) diff -r 65eb8910a779 -r 1c5171abe5cc src/HOL/ex/FinFunPred.thy --- a/src/HOL/ex/FinFunPred.thy Wed May 30 08:34:14 2012 +0200 +++ b/src/HOL/ex/FinFunPred.thy Wed May 30 08:48:14 2012 +0200 @@ -8,14 +8,14 @@ text {* Instantiate FinFun predicates just like predicates *} -type_synonym 'a pred\<^isub>f = "'a \\<^isub>f bool" +type_synonym 'a pred\<^isub>f = "'a \f bool" instantiation "finfun" :: (type, ord) ord begin definition le_finfun_def [code del]: "f \ g \ (\x. f\<^sub>f x \ g\<^sub>f x)" -definition [code del]: "(f\'a \\<^isub>f 'b) < g \ f \ g \ \ f \ g" +definition [code del]: "(f\'a \f 'b) < g \ f \ g \ \ f \ g" instance ..