# HG changeset patch # User wenzelm # Date 1468226607 -7200 # Node ID aa03b0487bf5a6182dd9d72563c65776a9315211 # Parent ba7901e94e7bfa75aa752447490ec30fedd0b876 tuned; diff -r ba7901e94e7b -r aa03b0487bf5 src/HOL/Library/Groups_Big_Fun.thy --- a/src/HOL/Library/Groups_Big_Fun.thy Mon Jul 11 09:57:20 2016 +0200 +++ b/src/HOL/Library/Groups_Big_Fun.thy Mon Jul 11 10:43:27 2016 +0200 @@ -215,10 +215,8 @@ begin sublocale Sum_any: comm_monoid_fun plus 0 -defines - Sum_any = Sum_any.G -rewrites - "comm_monoid_set.F plus 0 = setsum" + defines Sum_any = Sum_any.G + rewrites "comm_monoid_set.F plus 0 = setsum" proof - show "comm_monoid_fun plus 0" .. then interpret Sum_any: comm_monoid_fun plus 0 . @@ -284,10 +282,8 @@ begin sublocale Prod_any: comm_monoid_fun times 1 -defines - Prod_any = Prod_any.G -rewrites - "comm_monoid_set.F times 1 = setprod" + defines Prod_any = Prod_any.G + rewrites "comm_monoid_set.F times 1 = setprod" proof - show "comm_monoid_fun times 1" .. then interpret Prod_any: comm_monoid_fun times 1 . diff -r ba7901e94e7b -r aa03b0487bf5 src/HOL/Library/Perm.thy --- a/src/HOL/Library/Perm.thy Mon Jul 11 09:57:20 2016 +0200 +++ b/src/HOL/Library/Perm.thy Mon Jul 11 10:43:27 2016 +0200 @@ -20,7 +20,9 @@ typedef 'a perm = "{f :: 'a \ 'a. bij f \ finite {a. f a \ a}}" morphisms "apply" Perm - by (rule exI [of _ id]) simp +proof + show "id \ ?perm" by simp +qed setup_lifting type_definition_perm diff -r ba7901e94e7b -r aa03b0487bf5 src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Mon Jul 11 09:57:20 2016 +0200 +++ b/src/HOL/Library/Polynomial.thy Mon Jul 11 10:43:27 2016 +0200 @@ -55,7 +55,8 @@ subsection \Definition of type \poly\\ typedef (overloaded) 'a poly = "{f :: nat \ 'a::zero. \\<^sub>\ n. f n = 0}" - morphisms coeff Abs_poly by (auto intro!: ALL_MOST) + morphisms coeff Abs_poly + by (auto intro!: ALL_MOST) setup_lifting type_definition_poly diff -r ba7901e94e7b -r aa03b0487bf5 src/HOL/Library/Saturated.thy --- a/src/HOL/Library/Saturated.thy Mon Jul 11 09:57:20 2016 +0200 +++ b/src/HOL/Library/Saturated.thy Mon Jul 11 10:43:27 2016 +0200 @@ -215,8 +215,7 @@ end interpretation Inf_sat: semilattice_neutr_set min "top :: 'a::len sat" -rewrites - "semilattice_neutr_set.F min (top :: 'a sat) = Inf" + rewrites "semilattice_neutr_set.F min (top :: 'a sat) = Inf" proof - show "semilattice_neutr_set min (top :: 'a sat)" by standard (simp add: min_def) @@ -225,8 +224,7 @@ qed interpretation Sup_sat: semilattice_neutr_set max "bot :: 'a::len sat" -rewrites - "semilattice_neutr_set.F max (bot :: 'a sat) = Sup" + rewrites "semilattice_neutr_set.F max (bot :: 'a sat) = Sup" proof - show "semilattice_neutr_set max (bot :: 'a sat)" by standard (simp add: max_def bot.extremum_unique)