# HG changeset patch # User haftmann # Date 1192454979 -7200 # Node ID 6394db28d7951392130763fa9254ff08bd62dd61 # Parent 4bfae4c030be2a800b0d1360cad52056a66317b7 explicit parameter for class finite diff -r 4bfae4c030be -r 6394db28d795 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Mon Oct 15 12:25:33 2007 +0200 +++ b/src/HOL/Finite_Set.thy Mon Oct 15 15:29:39 2007 +0200 @@ -2622,6 +2622,7 @@ setup {* Sign.add_path "finite" *} -- {*FIXME: name tweaking*} class finite (attach UNIV) = type + + fixes itself :: "'a itself" assumes finite_UNIV: "finite (UNIV \ 'a set)" setup {* Sign.parent_path *} hide const finite @@ -2633,6 +2634,7 @@ "UNIV = {()}" by auto instance unit :: finite + "Finite_Set.itself \ TYPE(unit)" proof have "finite {()}" by simp also note univ_unit [symmetric] @@ -2645,6 +2647,7 @@ "UNIV = {False, True}" by auto instance bool :: finite + "itself \ TYPE(bool)" proof have "finite {False, True}" by simp also note univ_bool [symmetric] @@ -2654,6 +2657,7 @@ lemmas [code func] = univ_bool instance * :: (finite, finite) finite + "itself \ TYPE('a\finite)" proof show "finite (UNIV :: ('a \ 'b) set)" proof (rule finite_Prod_UNIV) @@ -2667,6 +2671,7 @@ unfolding UNIV_Times_UNIV .. instance "+" :: (finite, finite) finite + "itself \ TYPE('a\finite + 'b\finite)" proof have a: "finite (UNIV :: 'a set)" by (rule finite) have b: "finite (UNIV :: 'b set)" by (rule finite) @@ -2680,6 +2685,7 @@ unfolding UNIV_Plus_UNIV .. instance set :: (finite) finite + "itself \ TYPE('a\finite set)" proof have "finite (UNIV :: 'a set)" by (rule finite) hence "finite (Pow (UNIV :: 'a set))" @@ -2694,6 +2700,7 @@ by (rule inj_onI, auto simp add: expand_set_eq expand_fun_eq) instance "fun" :: (finite, finite) finite + "itself \ TYPE('a\finite \ 'b\finite)" proof show "finite (UNIV :: ('a => 'b) set)" proof (rule finite_imageD) @@ -2703,6 +2710,7 @@ qed qed +hide (open) const itself subsection {* Equality and order on functions *}