# HG changeset patch # User huffman # Date 1243868257 25200 # Node ID 14a58e2ca3741863693e6740bc87f4dbc6b245d5 # Parent b3b534f06c2d472a9d68a18f364ef1618d0c19f7 add [code del] declarations diff -r b3b534f06c2d -r 14a58e2ca374 src/HOL/Lim.thy --- a/src/HOL/Lim.thy Mon Jun 01 07:45:49 2009 -0700 +++ b/src/HOL/Lim.thy Mon Jun 01 07:57:37 2009 -0700 @@ -14,7 +14,7 @@ definition at :: "'a::metric_space \ 'a filter" where - "at a = Abs_filter (\P. \r>0. \x. x \ a \ dist x a < r \ P x)" + [code del]: "at a = Abs_filter (\P. \r>0. \x. x \ a \ dist x a < r \ P x)" definition LIM :: "['a::metric_space \ 'b::metric_space, 'a, 'b] \ bool" diff -r b3b534f06c2d -r 14a58e2ca374 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Mon Jun 01 07:45:49 2009 -0700 +++ b/src/HOL/Limits.thy Mon Jun 01 07:57:37 2009 -0700 @@ -20,7 +20,7 @@ definition eventually :: "('a \ bool) \ 'a filter \ bool" where - "eventually P F \ Rep_filter F P" + [simp del]: "eventually P F \ Rep_filter F P" lemma eventually_True [simp]: "eventually (\x. True) F" unfolding eventually_def using Rep_filter [of F] by blast @@ -85,7 +85,7 @@ definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" where - "Zfun S F = (\r>0. eventually (\i. norm (S i) < r) F)" + [code del]: "Zfun S F = (\r>0. eventually (\i. norm (S i) < r) F)" lemma ZfunI: "(\r. 0 < r \ eventually (\i. norm (S i) < r) F) \ Zfun S F" @@ -228,7 +228,7 @@ definition tendsto :: "('a \ 'b::metric_space) \ 'b \ 'a filter \ bool" where - "tendsto f l net \ (\e>0. eventually (\x. dist (f x) l < e) net)" + [code del]: "tendsto f l net \ (\e>0. eventually (\x. dist (f x) l < e) net)" lemma tendstoI: "(\e. 0 < e \ eventually (\x. dist (f x) l < e) net) diff -r b3b534f06c2d -r 14a58e2ca374 src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Mon Jun 01 07:45:49 2009 -0700 +++ b/src/HOL/SEQ.thy Mon Jun 01 07:57:37 2009 -0700 @@ -14,7 +14,7 @@ definition sequentially :: "nat filter" where - "sequentially = Abs_filter (\P. \N. \n\N. P n)" + [code del]: "sequentially = Abs_filter (\P. \N. \n\N. P n)" definition Zseq :: "[nat \ 'a::real_normed_vector] \ bool" where