merged;
authorwenzelm
Wed, 13 Feb 2013 11:46:48 +0100
changeset 522250a55ac5bdd92
parent 52224 175b43e0b9ce
parent 52222 faf7f0d4f9eb
child 52226 ced7163f1fe4
merged;
NEWS
     1.1 --- a/CONTRIBUTORS	Tue Feb 12 17:39:45 2013 +0100
     1.2 +++ b/CONTRIBUTORS	Wed Feb 13 11:46:48 2013 +0100
     1.3 @@ -3,6 +3,9 @@
     1.4  who is listed as an author in one of the source files of this Isabelle
     1.5  distribution.
     1.6  
     1.7 +Contributions to this Isabelle version
     1.8 +--------------------------------------
     1.9 +
    1.10  Contributions to Isabelle2013
    1.11  -----------------------------
    1.12  
     2.1 --- a/NEWS	Tue Feb 12 17:39:45 2013 +0100
     2.2 +++ b/NEWS	Wed Feb 13 11:46:48 2013 +0100
     2.3 @@ -1,6 +1,16 @@
     2.4  Isabelle NEWS -- history user-relevant changes
     2.5  ==============================================
     2.6  
     2.7 +New in this Isabelle version
     2.8 +----------------------------
     2.9 +
    2.10 +*** HOL ***
    2.11 +
    2.12 +* Theory "RealVector" and "Limits": Introduce type class
    2.13 +(lin)order_topology. Allows to generalize theorems about limits and
    2.14 +order. Instances are reals and extended reals.
    2.15 +
    2.16 +
    2.17  New in Isabelle2013 (February 2013)
    2.18  -----------------------------------
    2.19  
     3.1 --- a/src/Doc/ProgProve/Logic.thy	Tue Feb 12 17:39:45 2013 +0100
     3.2 +++ b/src/Doc/ProgProve/Logic.thy	Wed Feb 13 11:46:48 2013 +0100
     3.3 @@ -72,6 +72,7 @@
     3.4  \end{warn}
     3.5  
     3.6  \subsection{Sets}
     3.7 +\label{sec:Sets}
     3.8  
     3.9  Sets of elements of type @{typ 'a} have type @{typ"'a set"}.
    3.10  They can be finite or infinite. Sets come with the usual notation:
     4.1 --- a/src/Doc/Sledgehammer/document/root.tex	Tue Feb 12 17:39:45 2013 +0100
     4.2 +++ b/src/Doc/Sledgehammer/document/root.tex	Wed Feb 13 11:46:48 2013 +0100
     4.3 @@ -1063,10 +1063,11 @@
     4.4  variable \texttt{MASH} to \texttt{yes}. Persistent data is stored in the
     4.5  directory \texttt{\$ISABELLE\_HOME\_USER/mash}.
     4.6  
     4.7 -\item[\labelitemi] \textbf{\textit{mesh}:} A combination of MePo and MaSh.
     4.8 +\item[\labelitemi] \textbf{\textit{mesh}:} The MeSh filter, which combines the
     4.9 +rankings from MePo and MaSh.
    4.10  
    4.11 -\item[\labelitemi] \textbf{\textit{smart}:} Use MeSh if MaSh is
    4.12 -enabled and the target prover is an ATP; otherwise, use MePo.
    4.13 +\item[\labelitemi] \textbf{\textit{smart}:} A mixture of MePo, MaSh, and MeSh if
    4.14 +MaSh is enabled; otherwise, MePo.
    4.15  \end{enum}
    4.16  
    4.17  \opdefault{max\_facts}{smart\_int}{smart}
     5.1 --- a/src/HOL/BNF/Examples/Stream.thy	Tue Feb 12 17:39:45 2013 +0100
     5.2 +++ b/src/HOL/BNF/Examples/Stream.thy	Wed Feb 13 11:46:48 2013 +0100
     5.3 @@ -12,7 +12,7 @@
     5.4  imports "../BNF"
     5.5  begin
     5.6  
     5.7 -codata 'a stream = Stream (shd: 'a) (stl: "'a stream")
     5.8 +codata 'a stream = Stream (shd: 'a) (stl: "'a stream") (infixr "##" 65)
     5.9  
    5.10  (* TODO: Provide by the package*)
    5.11  theorem stream_set_induct:
    5.12 @@ -42,7 +42,7 @@
    5.13  
    5.14  primrec shift :: "'a list \<Rightarrow> 'a stream \<Rightarrow> 'a stream" (infixr "@-" 65) where
    5.15    "shift [] s = s"
    5.16 -| "shift (x # xs) s = Stream x (shift xs s)"
    5.17 +| "shift (x # xs) s = x ## shift xs s"
    5.18  
    5.19  lemma shift_append[simp]: "(xs @ ys) @- s = xs @- ys @- s"
    5.20  by (induct xs) auto
    5.21 @@ -71,10 +71,10 @@
    5.22    thus ?case using stream.unfold[of hd "\<lambda>xs. tl xs @ [hd xs]" u] by (auto simp: cycle_def)
    5.23  qed auto
    5.24  
    5.25 -lemma cycle_Cons: "cycle (x # xs) = Stream x (cycle (xs @ [x]))"
    5.26 -proof (coinduct rule: stream.coinduct[of "\<lambda>s1 s2. \<exists>x xs. s1 = cycle (x # xs) \<and> s2 = Stream x (cycle (xs @ [x]))"])
    5.27 +lemma cycle_Cons: "cycle (x # xs) = x ## cycle (xs @ [x])"
    5.28 +proof (coinduct rule: stream.coinduct[of "\<lambda>s1 s2. \<exists>x xs. s1 = cycle (x # xs) \<and> s2 = x ## cycle (xs @ [x])"])
    5.29    case (2 s1 s2)
    5.30 -  then obtain x xs where "s1 = cycle (x # xs) \<and> s2 = Stream x (cycle (xs @ [x]))" by blast
    5.31 +  then obtain x xs where "s1 = cycle (x # xs) \<and> s2 = x ## cycle (xs @ [x])" by blast
    5.32    thus ?case
    5.33      by (auto simp: cycle_def intro!: exI[of _ "hd (xs @ [x])"] exI[of _ "tl (xs @ [x])"] stream.unfold)
    5.34  qed auto
    5.35 @@ -83,7 +83,7 @@
    5.36    streams :: "'a set => 'a stream set"
    5.37    for A :: "'a set"
    5.38  where
    5.39 -  Stream[intro!, simp, no_atp]: "\<lbrakk>a \<in> A; s \<in> streams A\<rbrakk> \<Longrightarrow> Stream a s \<in> streams A"
    5.40 +  Stream[intro!, simp, no_atp]: "\<lbrakk>a \<in> A; s \<in> streams A\<rbrakk> \<Longrightarrow> a ## s \<in> streams A"
    5.41  
    5.42  lemma shift_streams: "\<lbrakk>w \<in> lists A; s \<in> streams A\<rbrakk> \<Longrightarrow> w @- s \<in> streams A"
    5.43  by (induct w) auto
    5.44 @@ -91,13 +91,13 @@
    5.45  lemma stream_set_streams:
    5.46    assumes "stream_set s \<subseteq> A"
    5.47    shows "s \<in> streams A"
    5.48 -proof (coinduct rule: streams.coinduct[of "\<lambda>s'. \<exists>a s. s' = Stream a s \<and> a \<in> A \<and> stream_set s \<subseteq> A"])
    5.49 +proof (coinduct rule: streams.coinduct[of "\<lambda>s'. \<exists>a s. s' = a ## s \<and> a \<in> A \<and> stream_set s \<subseteq> A"])
    5.50    case streams from assms show ?case by (cases s) auto
    5.51  next
    5.52 -  fix s' assume "\<exists>a s. s' = Stream a s \<and> a \<in> A \<and> stream_set s \<subseteq> A"
    5.53 +  fix s' assume "\<exists>a s. s' = a ## s \<and> a \<in> A \<and> stream_set s \<subseteq> A"
    5.54    then guess a s by (elim exE)
    5.55 -  with assms show "\<exists>a l. s' = Stream a l \<and>
    5.56 -    a \<in> A \<and> ((\<exists>a s. l = Stream a s \<and> a \<in> A \<and> stream_set s \<subseteq> A) \<or> l \<in> streams A)"
    5.57 +  with assms show "\<exists>a l. s' = a ## l \<and>
    5.58 +    a \<in> A \<and> ((\<exists>a s. l = a ## s \<and> a \<in> A \<and> stream_set s \<subseteq> A) \<or> l \<in> streams A)"
    5.59      by (cases s) auto
    5.60  qed
    5.61  
    5.62 @@ -105,17 +105,17 @@
    5.63  subsection {* flatten a stream of lists *}
    5.64  
    5.65  definition flat where
    5.66 -  "flat \<equiv> stream_unfold (hd o shd) (\<lambda>s. if tl (shd s) = [] then stl s else Stream (tl (shd s)) (stl s))"
    5.67 +  "flat \<equiv> stream_unfold (hd o shd) (\<lambda>s. if tl (shd s) = [] then stl s else tl (shd s) ## stl s)"
    5.68  
    5.69  lemma flat_simps[simp]:
    5.70    "shd (flat ws) = hd (shd ws)"
    5.71 -  "stl (flat ws) = flat (if tl (shd ws) = [] then stl ws else Stream (tl (shd ws)) (stl ws))"
    5.72 +  "stl (flat ws) = flat (if tl (shd ws) = [] then stl ws else tl (shd ws) ## stl ws)"
    5.73  unfolding flat_def by auto
    5.74  
    5.75 -lemma flat_Cons[simp]: "flat (Stream (x#xs) w) = Stream x (flat (if xs = [] then w else Stream xs w))"
    5.76 -unfolding flat_def using stream.unfold[of "hd o shd" _ "Stream (x#xs) w"] by auto
    5.77 +lemma flat_Cons[simp]: "flat ((x # xs) ## ws) = x ## flat (if xs = [] then ws else xs ## ws)"
    5.78 +unfolding flat_def using stream.unfold[of "hd o shd" _ "(x # xs) ## ws"] by auto
    5.79  
    5.80 -lemma flat_Stream[simp]: "xs \<noteq> [] \<Longrightarrow> flat (Stream xs ws) = xs @- flat ws"
    5.81 +lemma flat_Stream[simp]: "xs \<noteq> [] \<Longrightarrow> flat (xs ## ws) = xs @- flat ws"
    5.82  by (induct xs) auto
    5.83  
    5.84  lemma flat_unfold: "shd ws \<noteq> [] \<Longrightarrow> flat ws = shd ws @- flat (stl ws)"
    5.85 @@ -132,9 +132,9 @@
    5.86    "sdrop 0 s = s"
    5.87  | "sdrop (Suc n) s = sdrop n (stl s)"
    5.88  
    5.89 -primrec snth :: "nat \<Rightarrow> 'a stream \<Rightarrow> 'a" where
    5.90 -  "snth 0 s = shd s"
    5.91 -| "snth (Suc n) s = snth n (stl s)"
    5.92 +primrec snth :: "'a stream \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 100) where
    5.93 +  "s !! 0 = shd s"
    5.94 +| "s !! Suc n = stl s !! n"
    5.95  
    5.96  lemma stake_sdrop: "stake n s @- sdrop n s = s"
    5.97  by (induct n arbitrary: s) auto
     6.1 --- a/src/HOL/Codegenerator_Test/RBT_Set_Test.thy	Tue Feb 12 17:39:45 2013 +0100
     6.2 +++ b/src/HOL/Codegenerator_Test/RBT_Set_Test.thy	Wed Feb 13 11:46:48 2013 +0100
     6.3 @@ -46,6 +46,27 @@
     6.4    or implemented for RBT trees.
     6.5  *)
     6.6  
     6.7 -export_code _ checking SML OCaml? Haskell? Scala?
     6.8 +export_code _ checking SML OCaml? Haskell?
     6.9 +
    6.10 +(* Extra setup to make Scala happy *)
    6.11 +(* If the compilation fails with an error "ambiguous implicit values",
    6.12 +   read \<section>7.1 in the Code Generation Manual *)
    6.13 +
    6.14 +class ccpo_linorder = ccpo + linorder
    6.15 +
    6.16 +lemma [code]:
    6.17 +  "(Complete_Partial_Order.admissible :: ('a :: ccpo_linorder \<Rightarrow> bool) \<Rightarrow> bool) P = 
    6.18 +    (\<forall>A. Complete_Partial_Order.chain (op \<le>) A \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (Sup A))"
    6.19 +unfolding admissible_def by blast
    6.20 +
    6.21 +lemma [code]:
    6.22 +  "(gfp :: ('a :: complete_linorder \<Rightarrow> 'a) \<Rightarrow> 'a) f = Sup {u. u \<le> f u}"
    6.23 +unfolding gfp_def by blast
    6.24 +
    6.25 +lemma [code]:
    6.26 +  "(lfp :: ('a :: complete_linorder \<Rightarrow> 'a) \<Rightarrow> 'a) f = Inf {u. f u \<le> u}"
    6.27 +unfolding lfp_def by blast
    6.28 +
    6.29 +export_code _ checking Scala?
    6.30  
    6.31  end
     7.1 --- a/src/HOL/Complex.thy	Tue Feb 12 17:39:45 2013 +0100
     7.2 +++ b/src/HOL/Complex.thy	Wed Feb 13 11:46:48 2013 +0100
     7.3 @@ -288,8 +288,6 @@
     7.4  
     7.5  instance proof
     7.6    fix r :: real and x y :: complex and S :: "complex set"
     7.7 -  show "0 \<le> norm x"
     7.8 -    by (induct x) simp
     7.9    show "(norm x = 0) = (x = 0)"
    7.10      by (induct x) simp
    7.11    show "norm (x + y) \<le> norm x + norm y"
     8.1 --- a/src/HOL/HOL.thy	Tue Feb 12 17:39:45 2013 +0100
     8.2 +++ b/src/HOL/HOL.thy	Wed Feb 13 11:46:48 2013 +0100
     8.3 @@ -1265,15 +1265,15 @@
     8.4            val fx_g = Simplifier.rewrite ss (Thm.apply cf cx);
     8.5            val (_ $ _ $ g) = prop_of fx_g;
     8.6            val g' = abstract_over (x,g);
     8.7 +          val abs_g'= Abs (n,xT,g');
     8.8          in (if (g aconv g')
     8.9               then
    8.10                  let
    8.11                    val rl =
    8.12                      cterm_instantiate [(f_Let_unfold, cf), (x_Let_unfold, cx)] @{thm Let_unfold};
    8.13                  in SOME (rl OF [fx_g]) end
    8.14 -             else if Term.betapply (f, x) aconv g then NONE (*avoid identity conversion*)
    8.15 +             else if (Envir.beta_eta_contract f) aconv (Envir.beta_eta_contract abs_g') then NONE (*avoid identity conversion*)
    8.16               else let
    8.17 -                   val abs_g'= Abs (n,xT,g');
    8.18                     val g'x = abs_g'$x;
    8.19                     val g_g'x = Thm.symmetric (Thm.beta_conversion false (cterm_of thy g'x));
    8.20                     val rl = cterm_instantiate
     9.1 --- a/src/HOL/IMP/AExp.thy	Tue Feb 12 17:39:45 2013 +0100
     9.2 +++ b/src/HOL/IMP/AExp.thy	Wed Feb 13 11:46:48 2013 +0100
     9.3 @@ -32,7 +32,7 @@
     9.4  syntax 
     9.5    "_State" :: "updbinds => 'a" ("<_>")
     9.6  translations
     9.7 -  "_State ms" => "_Update <> ms"
     9.8 +  "_State ms" == "_Update <> ms"
     9.9  
    9.10  text {* \noindent
    9.11    We can now write a series of updates to the function @{text "\<lambda>x. 0"} compactly:
    10.1 --- a/src/HOL/IMP/Abs_Int0.thy	Tue Feb 12 17:39:45 2013 +0100
    10.2 +++ b/src/HOL/IMP/Abs_Int0.thy	Wed Feb 13 11:46:48 2013 +0100
    10.3 @@ -242,9 +242,8 @@
    10.4    and gamma_Top[simp]: "\<gamma> \<top> = UNIV"
    10.5  fixes num' :: "val \<Rightarrow> 'av"
    10.6  and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av"
    10.7 -  assumes gamma_num': "i : \<gamma>(num' i)"
    10.8 -  and gamma_plus':
    10.9 - "i1 : \<gamma> a1 \<Longrightarrow> i2 : \<gamma> a2 \<Longrightarrow> i1+i2 : \<gamma>(plus' a1 a2)"
   10.10 +  assumes gamma_num': "i \<in> \<gamma>(num' i)"
   10.11 +  and gamma_plus': "i1 \<in> \<gamma> a1 \<Longrightarrow> i2 \<in> \<gamma> a2 \<Longrightarrow> i1+i2 \<in> \<gamma>(plus' a1 a2)"
   10.12  
   10.13  type_synonym 'av st = "(vname \<Rightarrow> 'av)"
   10.14  
   10.15 @@ -263,7 +262,8 @@
   10.16    x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S))}" |
   10.17  "step' S (C1; C2) = step' S C1; step' (post C1) C2" |
   10.18  "step' S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) =
   10.19 -   IF b THEN {S} step' P1 C1 ELSE {S} step' P2 C2 {post C1 \<squnion> post C2}" |
   10.20 +   IF b THEN {S} step' P1 C1 ELSE {S} step' P2 C2
   10.21 +   {post C1 \<squnion> post C2}" |
   10.22  "step' S ({I} WHILE b DO {P} C {Q}) =
   10.23    {S \<squnion> post C} WHILE b DO {I} step' P C {I}"
   10.24  
   10.25 @@ -290,8 +290,6 @@
   10.26  lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
   10.27  by (simp add: Top_option_def)
   10.28  
   10.29 -(* FIXME (maybe also le \<rightarrow> sqle?) *)
   10.30 -
   10.31  lemma mono_gamma_s: "f1 \<sqsubseteq> f2 \<Longrightarrow> \<gamma>\<^isub>s f1 \<subseteq> \<gamma>\<^isub>s f2"
   10.32  by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)
   10.33  
    11.1 --- a/src/HOL/IMP/Abs_Int1.thy	Tue Feb 12 17:39:45 2013 +0100
    11.2 +++ b/src/HOL/IMP/Abs_Int1.thy	Wed Feb 13 11:46:48 2013 +0100
    11.3 @@ -71,7 +71,7 @@
    11.4     {S \<squnion> post C} WHILE b DO {I} step' P C {I}"
    11.5  
    11.6  definition AI :: "com \<Rightarrow> 'av st option acom option" where
    11.7 -"AI c = pfp (step' (top c)) (bot c)"
    11.8 +"AI c = pfp (step' (top(vars c))) (bot c)"
    11.9  
   11.10  
   11.11  lemma strip_step'[simp]: "strip(step' S C) = strip C"
   11.12 @@ -110,21 +110,21 @@
   11.13  
   11.14  lemma AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
   11.15  proof(simp add: CS_def AI_def)
   11.16 -  assume 1: "pfp (step' (top c)) (bot c) = Some C"
   11.17 +  assume 1: "pfp (step' (top(vars c))) (bot c) = Some C"
   11.18    have "C \<in> L(vars c)"
   11.19      by(rule pfp_inv[where P = "%C. C \<in> L(vars c)", OF 1 _ bot_in_L])
   11.20        (erule step'_in_L[OF _ top_in_L])
   11.21 -  have pfp': "step' (top c) C \<sqsubseteq> C" by(rule pfp_pfp[OF 1])
   11.22 -  have 2: "step (\<gamma>\<^isub>o(top c)) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
   11.23 +  have pfp': "step' (top(vars c)) C \<sqsubseteq> C" by(rule pfp_pfp[OF 1])
   11.24 +  have 2: "step (\<gamma>\<^isub>o(top(vars c))) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
   11.25    proof(rule order_trans)
   11.26 -    show "step (\<gamma>\<^isub>o (top c)) (\<gamma>\<^isub>c C) \<le>  \<gamma>\<^isub>c (step' (top c) C)"
   11.27 +    show "step (\<gamma>\<^isub>o (top(vars c))) (\<gamma>\<^isub>c C) \<le>  \<gamma>\<^isub>c (step' (top(vars c)) C)"
   11.28        by(rule step_step'[OF `C \<in> L(vars c)` top_in_L])
   11.29 -    show "\<gamma>\<^isub>c (step' (top c) C) \<le> \<gamma>\<^isub>c C"
   11.30 +    show "\<gamma>\<^isub>c (step' (top(vars c)) C) \<le> \<gamma>\<^isub>c C"
   11.31        by(rule mono_gamma_c[OF pfp'])
   11.32    qed
   11.33    have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1])
   11.34 -  have "lfp c (step (\<gamma>\<^isub>o(top c))) \<le> \<gamma>\<^isub>c C"
   11.35 -    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o(top c))", OF 3 2])
   11.36 +  have "lfp c (step (\<gamma>\<^isub>o(top(vars c)))) \<le> \<gamma>\<^isub>c C"
   11.37 +    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o(top(vars c)))", OF 3 2])
   11.38    thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
   11.39  qed
   11.40  
   11.41 @@ -153,8 +153,8 @@
   11.42              split: option.split)
   11.43  done
   11.44  
   11.45 -lemma mono_step'_top: "C \<in> L(vars c) \<Longrightarrow> C' \<in> L(vars c) \<Longrightarrow>
   11.46 -  C \<sqsubseteq> C' \<Longrightarrow> step' (top c) C \<sqsubseteq> step' (top c) C'"
   11.47 +lemma mono_step'_top: "C \<in> L X \<Longrightarrow> C' \<in> L X \<Longrightarrow>
   11.48 +  C \<sqsubseteq> C' \<Longrightarrow> step' (top X) C \<sqsubseteq> step' (top X) C'"
   11.49  by (metis top_in_L mono_step' preord_class.le_refl)
   11.50  
   11.51  lemma pfp_bot_least:
   11.52 @@ -167,7 +167,7 @@
   11.53  by (simp_all add: assms(4,5) bot_least)
   11.54  
   11.55  lemma AI_least_pfp: assumes "AI c = Some C"
   11.56 -and "step' (top c) C' \<sqsubseteq> C'" "strip C' = c" "C' \<in> L(vars c)"
   11.57 +and "step' (top(vars c)) C' \<sqsubseteq> C'" "strip C' = c" "C' \<in> L(vars c)"
   11.58  shows "C \<sqsubseteq> C'"
   11.59  apply(rule pfp_bot_least[OF _ _ assms(2-4) assms(1)[unfolded AI_def]])
   11.60  by (simp_all add: mono_step'_top)
    12.1 --- a/src/HOL/IMP/Abs_Int1_const.thy	Tue Feb 12 17:39:45 2013 +0100
    12.2 +++ b/src/HOL/IMP/Abs_Int1_const.thy	Wed Feb 13 11:46:48 2013 +0100
    12.3 @@ -74,23 +74,23 @@
    12.4  
    12.5  subsubsection "Tests"
    12.6  
    12.7 -definition "steps c i = (step_const(top c) ^^ i) (bot c)"
    12.8 +definition "steps c i = (step_const(top(vars c)) ^^ i) (bot c)"
    12.9  
   12.10  value "show_acom (steps test1_const 0)"
   12.11  value "show_acom (steps test1_const 1)"
   12.12  value "show_acom (steps test1_const 2)"
   12.13  value "show_acom (steps test1_const 3)"
   12.14 -value "show_acom_opt (AI_const test1_const)"
   12.15 +value "show_acom (the(AI_const test1_const))"
   12.16  
   12.17 -value "show_acom_opt (AI_const test2_const)"
   12.18 -value "show_acom_opt (AI_const test3_const)"
   12.19 +value "show_acom (the(AI_const test2_const))"
   12.20 +value "show_acom (the(AI_const test3_const))"
   12.21  
   12.22  value "show_acom (steps test4_const 0)"
   12.23  value "show_acom (steps test4_const 1)"
   12.24  value "show_acom (steps test4_const 2)"
   12.25  value "show_acom (steps test4_const 3)"
   12.26  value "show_acom (steps test4_const 4)"
   12.27 -value "show_acom_opt (AI_const test4_const)"
   12.28 +value "show_acom (the(AI_const test4_const))"
   12.29  
   12.30  value "show_acom (steps test5_const 0)"
   12.31  value "show_acom (steps test5_const 1)"
   12.32 @@ -99,7 +99,7 @@
   12.33  value "show_acom (steps test5_const 4)"
   12.34  value "show_acom (steps test5_const 5)"
   12.35  value "show_acom (steps test5_const 6)"
   12.36 -value "show_acom_opt (AI_const test5_const)"
   12.37 +value "show_acom (the(AI_const test5_const))"
   12.38  
   12.39  value "show_acom (steps test6_const 0)"
   12.40  value "show_acom (steps test6_const 1)"
   12.41 @@ -115,7 +115,7 @@
   12.42  value "show_acom (steps test6_const 11)"
   12.43  value "show_acom (steps test6_const 12)"
   12.44  value "show_acom (steps test6_const 13)"
   12.45 -value "show_acom_opt (AI_const test6_const)"
   12.46 +value "show_acom (the(AI_const test6_const))"
   12.47  
   12.48  
   12.49  text{* Monotonicity: *}
    13.1 --- a/src/HOL/IMP/Abs_Int1_parity.thy	Tue Feb 12 17:39:45 2013 +0100
    13.2 +++ b/src/HOL/IMP/Abs_Int1_parity.thy	Wed Feb 13 11:46:48 2013 +0100
    13.3 @@ -121,13 +121,13 @@
    13.4  definition "test1_parity =
    13.5    ''x'' ::= N 1;
    13.6    WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)"
    13.7 -value [code] "show_acom_opt (AI_parity test1_parity)"
    13.8 +value [code] "show_acom (the(AI_parity test1_parity))"
    13.9  
   13.10  definition "test2_parity =
   13.11    ''x'' ::= N 1;
   13.12    WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)"
   13.13  
   13.14 -definition "steps c i = (step_parity(top c) ^^ i) (bot c)"
   13.15 +definition "steps c i = (step_parity(top(vars c)) ^^ i) (bot c)"
   13.16  
   13.17  value "show_acom (steps test2_parity 0)"
   13.18  value "show_acom (steps test2_parity 1)"
   13.19 @@ -136,7 +136,7 @@
   13.20  value "show_acom (steps test2_parity 4)"
   13.21  value "show_acom (steps test2_parity 5)"
   13.22  value "show_acom (steps test2_parity 6)"
   13.23 -value "show_acom_opt (AI_parity test2_parity)"
   13.24 +value "show_acom (the(AI_parity test2_parity))"
   13.25  
   13.26  
   13.27  subsubsection "Termination"
    14.1 --- a/src/HOL/IMP/Abs_Int2.thy	Tue Feb 12 17:39:45 2013 +0100
    14.2 +++ b/src/HOL/IMP/Abs_Int2.thy	Wed Feb 13 11:46:48 2013 +0100
    14.3 @@ -105,8 +105,8 @@
    14.4    (if res then bfilter b1 True (bfilter b2 True S)
    14.5     else bfilter b1 False S \<squnion> bfilter b2 False S)" |
    14.6  "bfilter (Less e1 e2) res S =
    14.7 -  (let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S)
    14.8 -   in afilter e1 res1 (afilter e2 res2 S))"
    14.9 +  (let (a1,a2) = filter_less' res (aval'' e1 S) (aval'' e2 S)
   14.10 +   in afilter e1 a1 (afilter e2 a2 S))"
   14.11  
   14.12  lemma afilter_in_L: "S \<in> L X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow> afilter e a S \<in> L X"
   14.13  by(induction e arbitrary: a S)
   14.14 @@ -167,7 +167,7 @@
   14.15     {bfilter b False I}"
   14.16  
   14.17  definition AI :: "com \<Rightarrow> 'av st option acom option" where
   14.18 -"AI c = pfp (step' \<top>\<^bsub>c\<^esub>) (bot c)"
   14.19 +"AI c = pfp (step' \<top>\<^bsub>vars c\<^esub>) (bot c)"
   14.20  
   14.21  lemma strip_step'[simp]: "strip(step' S c) = strip c"
   14.22  by(induct c arbitrary: S) (simp_all add: Let_def)
   14.23 @@ -209,21 +209,21 @@
   14.24  
   14.25  lemma AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
   14.26  proof(simp add: CS_def AI_def)
   14.27 -  assume 1: "pfp (step' (top c)) (bot c) = Some C"
   14.28 +  assume 1: "pfp (step' (top(vars c))) (bot c) = Some C"
   14.29    have "C \<in> L(vars c)"
   14.30      by(rule pfp_inv[where P = "%C. C \<in> L(vars c)", OF 1 _ bot_in_L])
   14.31        (erule step'_in_L[OF _ top_in_L])
   14.32 -  have pfp': "step' (top c) C \<sqsubseteq> C" by(rule pfp_pfp[OF 1])
   14.33 -  have 2: "step (\<gamma>\<^isub>o(top c)) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
   14.34 +  have pfp': "step' (top(vars c)) C \<sqsubseteq> C" by(rule pfp_pfp[OF 1])
   14.35 +  have 2: "step (\<gamma>\<^isub>o(top(vars c))) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
   14.36    proof(rule order_trans)
   14.37 -    show "step (\<gamma>\<^isub>o (top c)) (\<gamma>\<^isub>c C) \<le>  \<gamma>\<^isub>c (step' (top c) C)"
   14.38 +    show "step (\<gamma>\<^isub>o (top(vars c))) (\<gamma>\<^isub>c C) \<le>  \<gamma>\<^isub>c (step' (top(vars c)) C)"
   14.39        by(rule step_step'[OF `C \<in> L(vars c)` top_in_L])
   14.40 -    show "\<gamma>\<^isub>c (step' (top c) C) \<le> \<gamma>\<^isub>c C"
   14.41 +    show "\<gamma>\<^isub>c (step' (top(vars c)) C) \<le> \<gamma>\<^isub>c C"
   14.42        by(rule mono_gamma_c[OF pfp'])
   14.43    qed
   14.44    have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1])
   14.45 -  have "lfp c (step (\<gamma>\<^isub>o(top c))) \<le> \<gamma>\<^isub>c C"
   14.46 -    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o(top c))", OF 3 2])
   14.47 +  have "lfp c (step (\<gamma>\<^isub>o(top(vars c)))) \<le> \<gamma>\<^isub>c C"
   14.48 +    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o(top(vars c)))", OF 3 2])
   14.49    thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
   14.50  qed
   14.51  
   14.52 @@ -281,8 +281,8 @@
   14.53              split: option.split)
   14.54  done
   14.55  
   14.56 -lemma mono_step'_top: "C1 \<in> L(vars c) \<Longrightarrow> C2 \<in> L(vars c) \<Longrightarrow>
   14.57 -  C1 \<sqsubseteq> C2 \<Longrightarrow> step' (top c) C1 \<sqsubseteq> step' (top c) C2"
   14.58 +lemma mono_step'_top: "C1 \<in> L X \<Longrightarrow> C2 \<in> L X \<Longrightarrow>
   14.59 +  C1 \<sqsubseteq> C2 \<Longrightarrow> step' (top X) C1 \<sqsubseteq> step' (top X) C2"
   14.60  by (metis top_in_L mono_step' preord_class.le_refl)
   14.61  
   14.62  end
    15.1 --- a/src/HOL/IMP/Abs_Int2_ivl.thy	Tue Feb 12 17:39:45 2013 +0100
    15.2 +++ b/src/HOL/IMP/Abs_Int2_ivl.thy	Wed Feb 13 11:46:48 2013 +0100
    15.3 @@ -255,7 +255,7 @@
    15.4  value "show_acom_opt (AI_ivl test4_const)"
    15.5  value "show_acom_opt (AI_ivl test6_const)"
    15.6  
    15.7 -definition "steps c i = (step_ivl(top c) ^^ i) (bot c)"
    15.8 +definition "steps c i = (step_ivl(top(vars c)) ^^ i) (bot c)"
    15.9  
   15.10  value "show_acom_opt (AI_ivl test2_ivl)"
   15.11  value "show_acom (steps test2_ivl 0)"
    16.1 --- a/src/HOL/IMP/Abs_Int3.thy	Tue Feb 12 17:39:45 2013 +0100
    16.2 +++ b/src/HOL/IMP/Abs_Int3.thy	Wed Feb 13 11:46:48 2013 +0100
    16.3 @@ -358,24 +358,24 @@
    16.4  begin
    16.5  
    16.6  definition AI_wn :: "com \<Rightarrow> 'av st option acom option" where
    16.7 -"AI_wn c = pfp_wn (step' (top c)) (bot c)"
    16.8 +"AI_wn c = pfp_wn (step' (top(vars c))) (bot c)"
    16.9  
   16.10  lemma AI_wn_sound: "AI_wn c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
   16.11  proof(simp add: CS_def AI_wn_def)
   16.12 -  assume 1: "pfp_wn (step' (top c)) (bot c) = Some C"
   16.13 -  have 2: "(strip C = c & C \<in> L(vars c)) \<and> step' \<top>\<^bsub>c\<^esub> C \<sqsubseteq> C"
   16.14 +  assume 1: "pfp_wn (step' (top(vars c))) (bot c) = Some C"
   16.15 +  have 2: "(strip C = c & C \<in> L(vars c)) \<and> step' \<top>\<^bsub>vars c\<^esub> C \<sqsubseteq> C"
   16.16      by(rule pfp_wn_pfp[where x="bot c"])
   16.17        (simp_all add: 1 mono_step'_top widen_c_in_L narrow_c_in_L)
   16.18 -  have pfp: "step (\<gamma>\<^isub>o(top c)) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
   16.19 +  have pfp: "step (\<gamma>\<^isub>o(top(vars c))) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
   16.20    proof(rule order_trans)
   16.21 -    show "step (\<gamma>\<^isub>o (top c)) (\<gamma>\<^isub>c C) \<le>  \<gamma>\<^isub>c (step' (top c) C)"
   16.22 +    show "step (\<gamma>\<^isub>o (top(vars c))) (\<gamma>\<^isub>c C) \<le>  \<gamma>\<^isub>c (step' (top(vars c)) C)"
   16.23        by(rule step_step'[OF conjunct2[OF conjunct1[OF 2]] top_in_L])
   16.24      show "... \<le> \<gamma>\<^isub>c C"
   16.25        by(rule mono_gamma_c[OF conjunct2[OF 2]])
   16.26    qed
   16.27    have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp_wn[OF _ 1])
   16.28 -  have "lfp c (step (\<gamma>\<^isub>o (top c))) \<le> \<gamma>\<^isub>c C"
   16.29 -    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o(top c))", OF 3 pfp])
   16.30 +  have "lfp c (step (\<gamma>\<^isub>o (top(vars c)))) \<le> \<gamma>\<^isub>c C"
   16.31 +    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o(top(vars c)))", OF 3 pfp])
   16.32    thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
   16.33  qed
   16.34  
   16.35 @@ -396,9 +396,9 @@
   16.36  by(rule refl)
   16.37  
   16.38  definition "step_up_ivl n =
   16.39 -  ((\<lambda>C. C \<nabla> step_ivl (top(strip C)) C)^^n)"
   16.40 +  ((\<lambda>C. C \<nabla> step_ivl (top(vars(strip C))) C)^^n)"
   16.41  definition "step_down_ivl n =
   16.42 -  ((\<lambda>C. C \<triangle> step_ivl (top (strip C)) C)^^n)"
   16.43 +  ((\<lambda>C. C \<triangle> step_ivl (top(vars(strip C))) C)^^n)"
   16.44  
   16.45  text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as
   16.46  the loop took to execute. In contrast, @{const AI_ivl'} converges in a
   16.47 @@ -629,14 +629,14 @@
   16.48  
   16.49  
   16.50  lemma iter_winden_step_ivl_termination:
   16.51 -  "\<exists>C. iter_widen (step_ivl (top c)) (bot c) = Some C"
   16.52 +  "\<exists>C. iter_widen (step_ivl (top(vars c))) (bot c) = Some C"
   16.53  apply(rule iter_widen_termination[where m = "m_c" and P = "%C. C \<in> Lc c"])
   16.54  apply (simp_all add: step'_in_Lc m_c_widen)
   16.55  done
   16.56  
   16.57  lemma iter_narrow_step_ivl_termination:
   16.58 -  "C0 \<in> Lc c \<Longrightarrow> step_ivl (top c) C0 \<sqsubseteq> C0 \<Longrightarrow>
   16.59 -  \<exists>C. iter_narrow (step_ivl (top c)) C0 = Some C"
   16.60 +  "C0 \<in> Lc c \<Longrightarrow> step_ivl (top(vars c)) C0 \<sqsubseteq> C0 \<Longrightarrow>
   16.61 +  \<exists>C. iter_narrow (step_ivl (top(vars c))) C0 = Some C"
   16.62  apply(rule iter_narrow_termination[where n = "n_c" and P = "%C. C \<in> Lc c"])
   16.63  apply (simp add: step'_in_Lc)
   16.64  apply (simp)
   16.65 @@ -654,7 +654,7 @@
   16.66  apply(auto simp: AI_wn_def pfp_wn_def iter_winden_step_ivl_termination
   16.67             split: option.split)
   16.68  apply(rule iter_narrow_step_ivl_termination)
   16.69 -apply(blast intro: iter_widen_inv[where f = "step' \<top>\<^bsub>c\<^esub>" and P = "%C. C \<in> Lc c"] bot_in_Lc Lc_widen step'_in_Lc[where S = "\<top>\<^bsub>c\<^esub>" and c=c, simplified])
   16.70 +apply(blast intro: iter_widen_inv[where f = "step' \<top>\<^bsub>vars c\<^esub>" and P = "%C. C \<in> Lc c"] bot_in_Lc Lc_widen step'_in_Lc[where S = "\<top>\<^bsub>vars c\<^esub>" and c=c, simplified])
   16.71  apply(erule iter_widen_pfp)
   16.72  done
   16.73  
    17.1 --- a/src/HOL/IMP/Abs_State.thy	Tue Feb 12 17:39:45 2013 +0100
    17.2 +++ b/src/HOL/IMP/Abs_State.thy	Wed Feb 13 11:46:48 2013 +0100
    17.3 @@ -60,12 +60,12 @@
    17.4  end
    17.5  
    17.6  class semilatticeL = join + L +
    17.7 -fixes top :: "com \<Rightarrow> 'a"
    17.8 +fixes top :: "vname set \<Rightarrow> 'a"
    17.9  assumes join_ge1 [simp]: "x \<in> L X \<Longrightarrow> y \<in> L X \<Longrightarrow> x \<sqsubseteq> x \<squnion> y"
   17.10  and join_ge2 [simp]: "x \<in> L X \<Longrightarrow> y \<in> L X \<Longrightarrow> y \<sqsubseteq> x \<squnion> y"
   17.11  and join_least[simp]: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z"
   17.12 -and top[simp]: "x \<in> L(vars c) \<Longrightarrow> x \<sqsubseteq> top c"
   17.13 -and top_in_L[simp]: "top c \<in> L(vars c)"
   17.14 +and top[simp]: "x \<in> L X \<Longrightarrow> x \<sqsubseteq> top X"
   17.15 +and top_in_L[simp]: "top X \<in> L X"
   17.16  and join_in_L[simp]: "x \<in> L X \<Longrightarrow> y \<in> L X \<Longrightarrow> x \<squnion> y \<in> L X"
   17.17  
   17.18  notation (input) top ("\<top>\<^bsub>_\<^esub>")
   17.19 @@ -158,7 +158,7 @@
   17.20  instantiation st :: (semilattice) semilatticeL
   17.21  begin
   17.22  
   17.23 -definition top_st where "top c = FunDom (\<lambda>x. \<top>) (vars c)"
   17.24 +definition top_st where "top X = FunDom (\<lambda>x. \<top>) X"
   17.25  
   17.26  instance
   17.27  proof
   17.28 @@ -201,8 +201,6 @@
   17.29  lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o (top c) = UNIV"
   17.30  by (simp add: top_option_def)
   17.31  
   17.32 -(* FIXME (maybe also le \<rightarrow> sqle?) *)
   17.33 -
   17.34  lemma mono_gamma_s: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>s f \<subseteq> \<gamma>\<^isub>s g"
   17.35  apply(simp add:\<gamma>_st_def subset_iff le_st_def split: if_splits)
   17.36  by (metis mono_gamma subsetD)
    18.1 --- a/src/HOL/IMP/Collecting.thy	Tue Feb 12 17:39:45 2013 +0100
    18.2 +++ b/src/HOL/IMP/Collecting.thy	Wed Feb 13 11:46:48 2013 +0100
    18.3 @@ -142,14 +142,14 @@
    18.4  subsubsection "Collecting semantics"
    18.5  
    18.6  fun step :: "state set \<Rightarrow> state set acom \<Rightarrow> state set acom" where
    18.7 -"step S (SKIP {P}) = (SKIP {S})" |
    18.8 -"step S (x ::= e {P}) =
    18.9 +"step S (SKIP {Q}) = (SKIP {S})" |
   18.10 +"step S (x ::= e {Q}) =
   18.11    x ::= e {{s(x := aval e s) |s. s : S}}" |
   18.12  "step S (C1; C2) = step S C1; step (post C1) C2" |
   18.13 -"step S (IF b THEN {P1} C1 ELSE {P2} C2 {P}) =
   18.14 +"step S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) =
   18.15    IF b THEN {{s:S. bval b s}} step P1 C1 ELSE {{s:S. \<not> bval b s}} step P2 C2
   18.16    {post C1 \<union> post C2}" |
   18.17 -"step S ({I} WHILE b DO {P} C {P'}) =
   18.18 +"step S ({I} WHILE b DO {P} C {Q}) =
   18.19    {S \<union> post C} WHILE b DO {{s:I. bval b s}} step P C {{s:I. \<not> bval b s}}"
   18.20  
   18.21  definition CS :: "com \<Rightarrow> state set acom" where
    19.1 --- a/src/HOL/IMP/Collecting_Examples.thy	Tue Feb 12 17:39:45 2013 +0100
    19.2 +++ b/src/HOL/IMP/Collecting_Examples.thy	Wed Feb 13 11:46:48 2013 +0100
    19.3 @@ -7,13 +7,20 @@
    19.4  lemma insert_code [code]:  "insert x (set xs) = set (x#xs)"
    19.5  by simp
    19.6  
    19.7 +text{* In order to display commands annotated with state sets,
    19.8 +states must be translated into a printable format as lists of pairs,
    19.9 +for a given set of variable names. This is what @{text show_acom} does: *}
   19.10 +
   19.11 +definition show_acom ::
   19.12 +  "vname list \<Rightarrow> state set acom \<Rightarrow> (vname*val)list set acom" where
   19.13 +"show_acom xs = map_acom (\<lambda>S. (\<lambda>s. map (\<lambda>x. (x, s x)) xs) ` S)"
   19.14 +
   19.15 +
   19.16  text{* The example: *}
   19.17  definition "c = WHILE Less (V ''x'') (N 3)
   19.18                  DO ''x'' ::= Plus (V ''x'') (N 2)"
   19.19  definition C0 :: "state set acom" where "C0 = anno {} c"
   19.20  
   19.21 -definition "show_acom xs = map_acom (\<lambda>S. (\<lambda>s. [(x,s x). x \<leftarrow> xs]) ` S)"
   19.22 -
   19.23  text{* Collecting semantics: *}
   19.24  
   19.25  value "show_acom [''x''] (((step {<>}) ^^ 1) C0)"
    20.1 --- a/src/HOL/IMP/Procs.thy	Tue Feb 12 17:39:45 2013 +0100
    20.2 +++ b/src/HOL/IMP/Procs.thy	Wed Feb 13 11:46:48 2013 +0100
    20.3 @@ -18,12 +18,11 @@
    20.4  
    20.5  definition "test_com =
    20.6  {VAR ''x'';;
    20.7 - ''x'' ::= N 0;
    20.8 - {PROC ''p'' = ''x'' ::= Plus (V ''x'') (V ''x'');;
    20.9 + {PROC ''p'' = ''x'' ::= N 1;;
   20.10    {PROC ''q'' = CALL ''p'';;
   20.11     {VAR ''x'';;
   20.12 -    ''x'' ::= N 5;
   20.13 -    {PROC ''p'' = ''x'' ::= Plus (V ''x'') (N 1);;
   20.14 +    ''x'' ::= N 2;
   20.15 +    {PROC ''p'' = ''x'' ::= N 3;;
   20.16       CALL ''q''; ''y'' ::= V ''x''}}}}}"
   20.17  
   20.18  end
    21.1 --- a/src/HOL/IMP/Procs_Dyn_Vars_Dyn.thy	Tue Feb 12 17:39:45 2013 +0100
    21.2 +++ b/src/HOL/IMP/Procs_Dyn_Vars_Dyn.thy	Wed Feb 13 11:46:48 2013 +0100
    21.3 @@ -31,9 +31,6 @@
    21.4  
    21.5  code_pred big_step .
    21.6  
    21.7 -values "{map t [''x'',''y'',''z''] |t. 
    21.8 -          (\<lambda>p. SKIP) \<turnstile> (CALL ''p'', <''x'' := 42, ''y'' := 43>) \<Rightarrow> t}"
    21.9 -
   21.10 -values "{map t [''x'',''y'',''z''] |t. (\<lambda>p. SKIP) \<turnstile> (test_com, <>) \<Rightarrow> t}"
   21.11 +values "{map t [''x'',''y''] |t. (\<lambda>p. SKIP) \<turnstile> (test_com, <>) \<Rightarrow> t}"
   21.12  
   21.13  end
    22.1 --- a/src/HOL/IMP/Procs_Stat_Vars_Dyn.thy	Tue Feb 12 17:39:45 2013 +0100
    22.2 +++ b/src/HOL/IMP/Procs_Stat_Vars_Dyn.thy	Wed Feb 13 11:46:48 2013 +0100
    22.3 @@ -33,9 +33,6 @@
    22.4  
    22.5  code_pred big_step .
    22.6  
    22.7 -values "{map t [''x'', ''y'', ''z''] |t. 
    22.8 -            [] \<turnstile> (CALL ''p'', <''x'' := 42, ''y'' := 43>) \<Rightarrow> t}"
    22.9 -
   22.10 -values "{map t [''x'', ''y'', ''z''] |t. [] \<turnstile> (test_com, <>) \<Rightarrow> t}"
   22.11 +values "{map t [''x'', ''y''] |t. [] \<turnstile> (test_com, <>) \<Rightarrow> t}"
   22.12  
   22.13  end
    23.1 --- a/src/HOL/IMP/Procs_Stat_Vars_Stat.thy	Tue Feb 12 17:39:45 2013 +0100
    23.2 +++ b/src/HOL/IMP/Procs_Stat_Vars_Stat.thy	Wed Feb 13 11:46:48 2013 +0100
    23.3 @@ -32,7 +32,7 @@
    23.4     e \<turnstile> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
    23.5  
    23.6  Var: "(pe,ve(x:=f),f+1) \<turnstile> (c,s) \<Rightarrow> t  \<Longrightarrow>
    23.7 -      (pe,ve,f) \<turnstile> ({VAR x;; c}, s) \<Rightarrow> t(f := s f)" |
    23.8 +      (pe,ve,f) \<turnstile> ({VAR x;; c}, s) \<Rightarrow> t" |
    23.9  
   23.10  Call1: "((p,c,ve)#pe,ve,f) \<turnstile> (c, s) \<Rightarrow> t  \<Longrightarrow>
   23.11          ((p,c,ve)#pe,ve',f) \<turnstile> (CALL p, s) \<Rightarrow> t" |
   23.12 @@ -45,10 +45,8 @@
   23.13  code_pred big_step .
   23.14  
   23.15  
   23.16 -values "{map t [0,1] |t. ([], <>, 0) \<turnstile> (CALL ''p'', nth [42, 43]) \<Rightarrow> t}"
   23.17 -
   23.18 -values "{map t [0, 1, 2] |t.
   23.19 -  ([], <''x'' := 0, ''y'' := 1,''z'' := 2>, 0)
   23.20 +values "{map t [10,11] |t.
   23.21 +  ([], <''x'' := 10, ''y'' := 11>, 12)
   23.22    \<turnstile> (test_com, <>) \<Rightarrow> t}"
   23.23  
   23.24  end
    24.1 --- a/src/HOL/Library/Extended_Nat.thy	Tue Feb 12 17:39:45 2013 +0100
    24.2 +++ b/src/HOL/Library/Extended_Nat.thy	Wed Feb 13 11:46:48 2013 +0100
    24.3 @@ -6,7 +6,7 @@
    24.4  header {* Extended natural numbers (i.e. with infinity) *}
    24.5  
    24.6  theory Extended_Nat
    24.7 -imports Main
    24.8 +imports "~~/src/HOL/Main"
    24.9  begin
   24.10  
   24.11  class infinity =
    25.1 --- a/src/HOL/Library/Extended_Real.thy	Tue Feb 12 17:39:45 2013 +0100
    25.2 +++ b/src/HOL/Library/Extended_Real.thy	Wed Feb 13 11:46:48 2013 +0100
    25.3 @@ -8,7 +8,7 @@
    25.4  header {* Extended real number line *}
    25.5  
    25.6  theory Extended_Real
    25.7 -imports Complex_Main Extended_Nat
    25.8 +imports "~~/src/HOL/Complex_Main" Extended_Nat
    25.9  begin
   25.10  
   25.11  text {*
   25.12 @@ -18,6 +18,69 @@
   25.13  
   25.14  *}
   25.15  
   25.16 +lemma LIMSEQ_SUP:
   25.17 +  fixes X :: "nat \<Rightarrow> 'a :: {complete_linorder, linorder_topology}"
   25.18 +  assumes "incseq X"
   25.19 +  shows "X ----> (SUP i. X i)"
   25.20 +  using `incseq X`
   25.21 +  by (intro increasing_tendsto)
   25.22 +     (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans)
   25.23 +
   25.24 +lemma eventually_const: "\<not> trivial_limit net \<Longrightarrow> eventually (\<lambda>x. P) net \<longleftrightarrow> P"
   25.25 +  by (cases P) (simp_all add: eventually_False)
   25.26 +
   25.27 +lemma (in complete_lattice) Inf_le_Sup: "A \<noteq> {} \<Longrightarrow> Inf A \<le> Sup A"
   25.28 +  by (metis Sup_upper2 Inf_lower ex_in_conv)
   25.29 +
   25.30 +lemma (in complete_lattice) INF_le_SUP: "A \<noteq> {} \<Longrightarrow> INFI A f \<le> SUPR A f"
   25.31 +  unfolding INF_def SUP_def by (rule Inf_le_Sup) auto
   25.32 +
   25.33 +lemma (in complete_linorder) le_Sup_iff:
   25.34 +  "x \<le> Sup A \<longleftrightarrow> (\<forall>y<x. \<exists>a\<in>A. y < a)"
   25.35 +proof safe
   25.36 +  fix y assume "x \<le> Sup A" "y < x"
   25.37 +  then have "y < Sup A" by auto
   25.38 +  then show "\<exists>a\<in>A. y < a"
   25.39 +    unfolding less_Sup_iff .
   25.40 +qed (auto elim!: allE[of _ "Sup A"] simp add: not_le[symmetric] Sup_upper)
   25.41 +
   25.42 +lemma (in complete_linorder) le_SUP_iff:
   25.43 +  "x \<le> SUPR A f \<longleftrightarrow> (\<forall>y<x. \<exists>i\<in>A. y < f i)"
   25.44 +  unfolding le_Sup_iff SUP_def by simp
   25.45 +
   25.46 +lemma (in complete_linorder) Inf_le_iff:
   25.47 +  "Inf A \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>a\<in>A. y > a)"
   25.48 +proof safe
   25.49 +  fix y assume "x \<ge> Inf A" "y > x"
   25.50 +  then have "y > Inf A" by auto
   25.51 +  then show "\<exists>a\<in>A. y > a"
   25.52 +    unfolding Inf_less_iff .
   25.53 +qed (auto elim!: allE[of _ "Inf A"] simp add: not_le[symmetric] Inf_lower)
   25.54 +
   25.55 +lemma (in complete_linorder) le_INF_iff:
   25.56 +  "INFI A f \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>i\<in>A. y > f i)"
   25.57 +  unfolding Inf_le_iff INF_def by simp
   25.58 +
   25.59 +lemma (in complete_lattice) Sup_eqI:
   25.60 +  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
   25.61 +  assumes "\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y"
   25.62 +  shows "Sup A = x"
   25.63 +  by (metis antisym Sup_least Sup_upper assms)
   25.64 +
   25.65 +lemma (in complete_lattice) Inf_eqI:
   25.66 +  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> i"
   25.67 +  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x"
   25.68 +  shows "Inf A = x"
   25.69 +  by (metis antisym Inf_greatest Inf_lower assms)
   25.70 +
   25.71 +lemma (in complete_lattice) SUP_eqI:
   25.72 +  "(\<And>i. i \<in> A \<Longrightarrow> f i \<le> x) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> (SUP i:A. f i) = x"
   25.73 +  unfolding SUP_def by (rule Sup_eqI) auto
   25.74 +
   25.75 +lemma (in complete_lattice) INF_eqI:
   25.76 +  "(\<And>i. i \<in> A \<Longrightarrow> x \<le> f i) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<ge> y) \<Longrightarrow> x \<ge> y) \<Longrightarrow> (INF i:A. f i) = x"
   25.77 +  unfolding INF_def by (rule Inf_eqI) auto
   25.78 +
   25.79  lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \<longleftrightarrow> x = bot"
   25.80  proof
   25.81    assume "{x..} = UNIV"
   25.82 @@ -1353,22 +1416,6 @@
   25.83    unfolding Sup_ereal_def
   25.84    by (auto intro!: Least_equality)
   25.85  
   25.86 -lemma ereal_SUPI:
   25.87 -  fixes x :: ereal
   25.88 -  assumes "!!i. i : A ==> f i <= x"
   25.89 -  assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y"
   25.90 -  shows "(SUP i:A. f i) = x"
   25.91 -  unfolding SUP_def Sup_ereal_def
   25.92 -  using assms by (auto intro!: Least_equality)
   25.93 -
   25.94 -lemma ereal_INFI:
   25.95 -  fixes x :: ereal
   25.96 -  assumes "!!i. i : A ==> f i >= x"
   25.97 -  assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y"
   25.98 -  shows "(INF i:A. f i) = x"
   25.99 -  unfolding INF_def Inf_ereal_def
  25.100 -  using assms by (auto intro!: Greatest_equality)
  25.101 -
  25.102  lemma Sup_ereal_close:
  25.103    fixes e :: ereal
  25.104    assumes "0 < e" and S: "\<bar>Sup S\<bar> \<noteq> \<infinity>" "S \<noteq> {}"
  25.105 @@ -1402,8 +1449,7 @@
  25.106  
  25.107  lemma ereal_le_Sup:
  25.108    fixes x :: ereal
  25.109 -  shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))"
  25.110 -(is "?lhs <-> ?rhs")
  25.111 +  shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))" (is "?lhs = ?rhs")
  25.112  proof-
  25.113  { assume "?rhs"
  25.114    { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i)<x" by (simp add: not_le)
  25.115 @@ -1476,6 +1522,17 @@
  25.116      using assms by (metis SUP_least SUP_upper2)
  25.117  qed
  25.118  
  25.119 +lemma INFI_eq:
  25.120 +  assumes "\<forall>i\<in>A. \<exists>j\<in>B. f i \<ge> g j"
  25.121 +  assumes "\<forall>j\<in>B. \<exists>i\<in>A. g j \<ge> f i"
  25.122 +  shows "(INF i:A. f i) = (INF j:B. g j)"
  25.123 +proof (intro antisym)
  25.124 +  show "(INF i:A. f i) \<le> (INF j:B. g j)"
  25.125 +    using assms by (metis INF_greatest INF_lower2)
  25.126 +  show "(INF i:B. g i) \<le> (INF j:A. f j)"
  25.127 +    using assms by (metis INF_greatest INF_lower2)
  25.128 +qed
  25.129 +
  25.130  lemma SUP_ereal_le_addI:
  25.131    fixes f :: "'i \<Rightarrow> ereal"
  25.132    assumes "\<And>i. f i + y \<le> z" and "y \<noteq> -\<infinity>"
  25.133 @@ -1491,7 +1548,7 @@
  25.134    fixes f g :: "nat \<Rightarrow> ereal"
  25.135    assumes "incseq f" "incseq g" and pos: "\<And>i. f i \<noteq> -\<infinity>" "\<And>i. g i \<noteq> -\<infinity>"
  25.136    shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
  25.137 -proof (rule ereal_SUPI)
  25.138 +proof (rule SUP_eqI)
  25.139    fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> f i + g i \<le> y"
  25.140    have f: "SUPR UNIV f \<noteq> -\<infinity>" using pos
  25.141      unfolding SUP_def Sup_eq_MInfty by (auto dest: image_eqD)
  25.142 @@ -1531,7 +1588,7 @@
  25.143  lemma SUPR_ereal_cmult:
  25.144    fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i" "0 \<le> c"
  25.145    shows "(SUP i. c * f i) = c * SUPR UNIV f"
  25.146 -proof (rule ereal_SUPI)
  25.147 +proof (rule SUP_eqI)
  25.148    fix i have "f i \<le> SUPR UNIV f" by (rule SUP_upper) auto
  25.149    then show "c * f i \<le> c * SUPR UNIV f"
  25.150      using `0 \<le> c` by (rule ereal_mult_left_mono)
  25.151 @@ -1598,7 +1655,7 @@
  25.152    qed
  25.153    from choice[OF this] guess f .. note f = this
  25.154    have "SUPR UNIV f = Sup A"
  25.155 -  proof (rule ereal_SUPI)
  25.156 +  proof (rule SUP_eqI)
  25.157      fix i show "f i \<le> Sup A" using f
  25.158        by (auto intro!: complete_lattice_class.Sup_upper)
  25.159    next
  25.160 @@ -1801,18 +1858,84 @@
  25.161  
  25.162  subsubsection "Topological space"
  25.163  
  25.164 -instantiation ereal :: topological_space
  25.165 +instantiation ereal :: linorder_topology
  25.166  begin
  25.167  
  25.168 -definition "open A \<longleftrightarrow> open (ereal -` A)
  25.169 -       \<and> (\<infinity> \<in> A \<longrightarrow> (\<exists>x. {ereal x <..} \<subseteq> A))
  25.170 -       \<and> (-\<infinity> \<in> A \<longrightarrow> (\<exists>x. {..<ereal x} \<subseteq> A))"
  25.171 +definition "open_ereal" :: "ereal set \<Rightarrow> bool" where
  25.172 +  open_ereal_generated: "open_ereal = generate_topology (range lessThan \<union> range greaterThan)"
  25.173 +
  25.174 +instance
  25.175 +  by default (simp add: open_ereal_generated)
  25.176 +end
  25.177  
  25.178  lemma open_PInfty: "open A \<Longrightarrow> \<infinity> \<in> A \<Longrightarrow> (\<exists>x. {ereal x<..} \<subseteq> A)"
  25.179 -  unfolding open_ereal_def by auto
  25.180 +  unfolding open_ereal_generated
  25.181 +proof (induct rule: generate_topology.induct)
  25.182 +  case (Int A B)
  25.183 +  moreover then obtain x z where "\<infinity> \<in> A \<Longrightarrow> {ereal x <..} \<subseteq> A" "\<infinity> \<in> B \<Longrightarrow> {ereal z <..} \<subseteq> B"
  25.184 +      by auto
  25.185 +  ultimately show ?case
  25.186 +    by (intro exI[of _ "max x z"]) fastforce
  25.187 +next
  25.188 +  { fix x have "x \<noteq> \<infinity> \<Longrightarrow> \<exists>t. x \<le> ereal t" by (cases x) auto }
  25.189 +  moreover case (Basis S)
  25.190 +  ultimately show ?case
  25.191 +    by (auto split: ereal.split)
  25.192 +qed (fastforce simp add: vimage_Union)+
  25.193  
  25.194  lemma open_MInfty: "open A \<Longrightarrow> -\<infinity> \<in> A \<Longrightarrow> (\<exists>x. {..<ereal x} \<subseteq> A)"
  25.195 -  unfolding open_ereal_def by auto
  25.196 +  unfolding open_ereal_generated
  25.197 +proof (induct rule: generate_topology.induct)
  25.198 +  case (Int A B)
  25.199 +  moreover then obtain x z where "-\<infinity> \<in> A \<Longrightarrow> {..< ereal x} \<subseteq> A" "-\<infinity> \<in> B \<Longrightarrow> {..< ereal z} \<subseteq> B"
  25.200 +      by auto
  25.201 +  ultimately show ?case
  25.202 +    by (intro exI[of _ "min x z"]) fastforce
  25.203 +next
  25.204 +  { fix x have "x \<noteq> - \<infinity> \<Longrightarrow> \<exists>t. ereal t \<le> x" by (cases x) auto }
  25.205 +  moreover case (Basis S)
  25.206 +  ultimately show ?case
  25.207 +    by (auto split: ereal.split)
  25.208 +qed (fastforce simp add: vimage_Union)+
  25.209 +
  25.210 +lemma open_ereal_vimage: "open S \<Longrightarrow> open (ereal -` S)"
  25.211 +  unfolding open_ereal_generated
  25.212 +proof (induct rule: generate_topology.induct)
  25.213 +  case (Int A B) then show ?case by auto
  25.214 +next
  25.215 +  { fix x have
  25.216 +      "ereal -` {..<x} = (case x of PInfty \<Rightarrow> UNIV | MInfty \<Rightarrow> {} | ereal r \<Rightarrow> {..<r})"
  25.217 +      "ereal -` {x<..} = (case x of PInfty \<Rightarrow> {} | MInfty \<Rightarrow> UNIV | ereal r \<Rightarrow> {r<..})"
  25.218 +      by (induct x) auto }
  25.219 +  moreover case (Basis S)
  25.220 +  ultimately show ?case
  25.221 +    by (auto split: ereal.split)
  25.222 +qed (fastforce simp add: vimage_Union)+
  25.223 +
  25.224 +lemma open_ereal: "open S \<Longrightarrow> open (ereal ` S)"
  25.225 +  unfolding open_generated_order[where 'a=real]
  25.226 +proof (induct rule: generate_topology.induct)
  25.227 +  case (Basis S)
  25.228 +  moreover { fix x have "ereal ` {..< x} = { -\<infinity> <..< ereal x }" by auto (case_tac xa, auto) }
  25.229 +  moreover { fix x have "ereal ` {x <..} = { ereal x <..< \<infinity> }" by auto (case_tac xa, auto) }
  25.230 +  ultimately show ?case
  25.231 +     by auto
  25.232 +qed (auto simp add: image_Union image_Int)
  25.233 +
  25.234 +lemma open_ereal_def: "open A \<longleftrightarrow> open (ereal -` A) \<and> (\<infinity> \<in> A \<longrightarrow> (\<exists>x. {ereal x <..} \<subseteq> A)) \<and> (-\<infinity> \<in> A \<longrightarrow> (\<exists>x. {..<ereal x} \<subseteq> A))"
  25.235 +  (is "open A \<longleftrightarrow> ?rhs")
  25.236 +proof
  25.237 +  assume "open A" then show ?rhs
  25.238 +    using open_PInfty open_MInfty open_ereal_vimage by auto
  25.239 +next
  25.240 +  assume "?rhs"
  25.241 +  then obtain x y where A: "open (ereal -` A)" "\<infinity> \<in> A \<Longrightarrow> {ereal x<..} \<subseteq> A" "-\<infinity> \<in> A \<Longrightarrow> {..< ereal y} \<subseteq> A"
  25.242 +    by auto
  25.243 +  have *: "A = ereal ` (ereal -` A) \<union> (if \<infinity> \<in> A then {ereal x<..} else {}) \<union> (if -\<infinity> \<in> A then {..< ereal y} else {})"
  25.244 +    using A(2,3) by auto
  25.245 +  from open_ereal[OF A(1)] show "open A"
  25.246 +    by (subst *) (auto simp: open_Un)
  25.247 +qed
  25.248  
  25.249  lemma open_PInfty2: assumes "open A" "\<infinity> \<in> A" obtains x where "{ereal x<..} \<subseteq> A"
  25.250    using open_PInfty[OF assms] by auto
  25.251 @@ -1821,85 +1944,17 @@
  25.252    using open_MInfty[OF assms] by auto
  25.253  
  25.254  lemma ereal_openE: assumes "open A" obtains x y where
  25.255 -  "open (ereal -` A)"
  25.256 -  "\<infinity> \<in> A \<Longrightarrow> {ereal x<..} \<subseteq> A"
  25.257 -  "-\<infinity> \<in> A \<Longrightarrow> {..<ereal y} \<subseteq> A"
  25.258 +  "open (ereal -` A)" "\<infinity> \<in> A \<Longrightarrow> {ereal x<..} \<subseteq> A" "-\<infinity> \<in> A \<Longrightarrow> {..<ereal y} \<subseteq> A"
  25.259    using assms open_ereal_def by auto
  25.260  
  25.261 -instance
  25.262 -proof
  25.263 -  let ?U = "UNIV::ereal set"
  25.264 -  show "open ?U" unfolding open_ereal_def
  25.265 -    by (auto intro!: exI[of _ 0])
  25.266 -next
  25.267 -  fix S T::"ereal set" assume "open S" and "open T"
  25.268 -  from `open S`[THEN ereal_openE] guess xS yS .
  25.269 -  moreover from `open T`[THEN ereal_openE] guess xT yT .
  25.270 -  ultimately have
  25.271 -    "open (ereal -` (S \<inter> T))"
  25.272 -    "\<infinity> \<in> S \<inter> T \<Longrightarrow> {ereal (max xS xT) <..} \<subseteq> S \<inter> T"
  25.273 -    "-\<infinity> \<in> S \<inter> T \<Longrightarrow> {..< ereal (min yS yT)} \<subseteq> S \<inter> T"
  25.274 -    by auto
  25.275 -  then show "open (S Int T)" unfolding open_ereal_def by blast
  25.276 -next
  25.277 -  fix K :: "ereal set set" assume "\<forall>S\<in>K. open S"
  25.278 -  then have *: "\<forall>S. \<exists>x y. S \<in> K \<longrightarrow> open (ereal -` S) \<and>
  25.279 -    (\<infinity> \<in> S \<longrightarrow> {ereal x <..} \<subseteq> S) \<and> (-\<infinity> \<in> S \<longrightarrow> {..< ereal y} \<subseteq> S)"
  25.280 -    by (auto simp: open_ereal_def)
  25.281 -  then show "open (Union K)" unfolding open_ereal_def
  25.282 -  proof (intro conjI impI)
  25.283 -    show "open (ereal -` \<Union>K)"
  25.284 -      using *[THEN choice] by (auto simp: vimage_Union)
  25.285 -  qed ((metis UnionE Union_upper subset_trans *)+)
  25.286 -qed
  25.287 -end
  25.288 -
  25.289 -lemma open_ereal: "open S \<Longrightarrow> open (ereal ` S)"
  25.290 -  by (auto simp: inj_vimage_image_eq open_ereal_def)
  25.291 -
  25.292 -lemma open_ereal_vimage: "open S \<Longrightarrow> open (ereal -` S)"
  25.293 -  unfolding open_ereal_def by auto
  25.294 -
  25.295 -lemma open_ereal_lessThan[intro, simp]: "open {..< a :: ereal}"
  25.296 -proof -
  25.297 -  have "\<And>x. ereal -` {..<ereal x} = {..< x}"
  25.298 -    "ereal -` {..< \<infinity>} = UNIV" "ereal -` {..< -\<infinity>} = {}" by auto
  25.299 -  then show ?thesis by (cases a) (auto simp: open_ereal_def)
  25.300 -qed
  25.301 -
  25.302 -lemma open_ereal_greaterThan[intro, simp]:
  25.303 -  "open {a :: ereal <..}"
  25.304 -proof -
  25.305 -  have "\<And>x. ereal -` {ereal x<..} = {x<..}"
  25.306 -    "ereal -` {\<infinity><..} = {}" "ereal -` {-\<infinity><..} = UNIV" by auto
  25.307 -  then show ?thesis by (cases a) (auto simp: open_ereal_def)
  25.308 -qed
  25.309 -
  25.310 -lemma ereal_open_greaterThanLessThan[intro, simp]: "open {a::ereal <..< b}"
  25.311 -  unfolding greaterThanLessThan_def by auto
  25.312 -
  25.313 -lemma closed_ereal_atLeast[simp, intro]: "closed {a :: ereal ..}"
  25.314 -proof -
  25.315 -  have "- {a ..} = {..< a}" by auto
  25.316 -  then show "closed {a ..}"
  25.317 -    unfolding closed_def using open_ereal_lessThan by auto
  25.318 -qed
  25.319 -
  25.320 -lemma closed_ereal_atMost[simp, intro]: "closed {.. b :: ereal}"
  25.321 -proof -
  25.322 -  have "- {.. b} = {b <..}" by auto
  25.323 -  then show "closed {.. b}"
  25.324 -    unfolding closed_def using open_ereal_greaterThan by auto
  25.325 -qed
  25.326 -
  25.327 -lemma closed_ereal_atLeastAtMost[simp, intro]:
  25.328 -  shows "closed {a :: ereal .. b}"
  25.329 -  unfolding atLeastAtMost_def by auto
  25.330 -
  25.331 -lemma closed_ereal_singleton:
  25.332 -  "closed {a :: ereal}"
  25.333 -by (metis atLeastAtMost_singleton closed_ereal_atLeastAtMost)
  25.334 -
  25.335 +lemmas open_ereal_lessThan = open_lessThan[where 'a=ereal]
  25.336 +lemmas open_ereal_greaterThan = open_greaterThan[where 'a=ereal]
  25.337 +lemmas ereal_open_greaterThanLessThan = open_greaterThanLessThan[where 'a=ereal]
  25.338 +lemmas closed_ereal_atLeast = closed_atLeast[where 'a=ereal]
  25.339 +lemmas closed_ereal_atMost = closed_atMost[where 'a=ereal]
  25.340 +lemmas closed_ereal_atLeastAtMost = closed_atLeastAtMost[where 'a=ereal]
  25.341 +lemmas closed_ereal_singleton = closed_singleton[where 'a=ereal]
  25.342 +  
  25.343  lemma ereal_open_cont_interval:
  25.344    fixes S :: "ereal set"
  25.345    assumes "open S" "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
  25.346 @@ -1928,28 +1983,6 @@
  25.347    show thesis by auto
  25.348  qed
  25.349  
  25.350 -instance ereal :: t2_space
  25.351 -proof
  25.352 -  fix x y :: ereal assume "x ~= y"
  25.353 -  let "?P x (y::ereal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}"
  25.354 -
  25.355 -  { fix x y :: ereal assume "x < y"
  25.356 -    from ereal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
  25.357 -    have "?P x y"
  25.358 -      apply (rule exI[of _ "{..<z}"])
  25.359 -      apply (rule exI[of _ "{z<..}"])
  25.360 -      using z by auto }
  25.361 -  note * = this
  25.362 -
  25.363 -  from `x ~= y`
  25.364 -  show "EX U V. open U & open V & x : U & y : V & U Int V = {}"
  25.365 -  proof (cases rule: linorder_cases)
  25.366 -    assume "x = y" with `x ~= y` show ?thesis by simp
  25.367 -  next assume "x < y" from *[OF this] show ?thesis by auto
  25.368 -  next assume "y < x" from *[OF this] show ?thesis by auto
  25.369 -  qed
  25.370 -qed
  25.371 -
  25.372  subsubsection {* Convergent sequences *}
  25.373  
  25.374  lemma lim_ereal[simp]:
  25.375 @@ -1979,152 +2012,67 @@
  25.376      by (rule eventually_mono)
  25.377  qed
  25.378  
  25.379 -lemma Lim_PInfty: "f ----> \<infinity> <-> (ALL B. EX N. ALL n>=N. f n >= ereal B)" (is "?l = ?r")
  25.380 -proof
  25.381 -  assume ?r
  25.382 -  show ?l
  25.383 -    apply(rule topological_tendstoI)
  25.384 -    unfolding eventually_sequentially
  25.385 -  proof-
  25.386 -    fix S :: "ereal set" assume "open S" "\<infinity> : S"
  25.387 -    from open_PInfty[OF this] guess B .. note B=this
  25.388 -    from `?r`[rule_format,of "B+1"] guess N .. note N=this
  25.389 -    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
  25.390 -    proof safe case goal1
  25.391 -      have "ereal B < ereal (B + 1)" by auto
  25.392 -      also have "... <= f n" using goal1 N by auto
  25.393 -      finally show ?case using B by fastforce
  25.394 -    qed
  25.395 -  qed
  25.396 +lemma tendsto_PInfty: "(f ---> \<infinity>) F \<longleftrightarrow> (\<forall>r. eventually (\<lambda>x. ereal r < f x) F)"
  25.397 +proof -
  25.398 +  { fix l :: ereal assume "\<forall>r. eventually (\<lambda>x. ereal r < f x) F"
  25.399 +    from this[THEN spec, of "real l"]
  25.400 +    have "l \<noteq> \<infinity> \<Longrightarrow> eventually (\<lambda>x. l < f x) F"
  25.401 +      by (cases l) (auto elim: eventually_elim1) }
  25.402 +  then show ?thesis
  25.403 +    by (auto simp: order_tendsto_iff)
  25.404 +qed
  25.405 +
  25.406 +lemma tendsto_MInfty: "(f ---> -\<infinity>) F \<longleftrightarrow> (\<forall>r. eventually (\<lambda>x. f x < ereal r) F)"
  25.407 +  unfolding tendsto_def
  25.408 +proof safe
  25.409 +  fix S :: "ereal set" assume "open S" "-\<infinity> \<in> S"
  25.410 +  from open_MInfty[OF this] guess B .. note B = this
  25.411 +  moreover
  25.412 +  assume "\<forall>r::real. eventually (\<lambda>z. f z < r) F"
  25.413 +  then have "eventually (\<lambda>z. f z \<in> {..< B}) F" by auto
  25.414 +  ultimately show "eventually (\<lambda>z. f z \<in> S) F" by (auto elim!: eventually_elim1)
  25.415  next
  25.416 -  assume ?l
  25.417 -  show ?r
  25.418 -  proof fix B::real have "open {ereal B<..}" "\<infinity> : {ereal B<..}" by auto
  25.419 -    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
  25.420 -    guess N .. note N=this
  25.421 -    show "EX N. ALL n>=N. ereal B <= f n" apply(rule_tac x=N in exI) using N by auto
  25.422 -  qed
  25.423 +  fix x assume "\<forall>S. open S \<longrightarrow> -\<infinity> \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) F"
  25.424 +  from this[rule_format, of "{..< ereal x}"]
  25.425 +  show "eventually (\<lambda>y. f y < ereal x) F" by auto
  25.426  qed
  25.427  
  25.428 -
  25.429 -lemma Lim_MInfty: "f ----> (-\<infinity>) <-> (ALL B. EX N. ALL n>=N. f n <= ereal B)" (is "?l = ?r")
  25.430 -proof
  25.431 -  assume ?r
  25.432 -  show ?l
  25.433 -    apply(rule topological_tendstoI)
  25.434 -    unfolding eventually_sequentially
  25.435 -  proof-
  25.436 -    fix S :: "ereal set"
  25.437 -    assume "open S" "(-\<infinity>) : S"
  25.438 -    from open_MInfty[OF this] guess B .. note B=this
  25.439 -    from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this
  25.440 -    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
  25.441 -    proof safe case goal1
  25.442 -      have "ereal (B - 1) >= f n" using goal1 N by auto
  25.443 -      also have "... < ereal B" by auto
  25.444 -      finally show ?case using B by fastforce
  25.445 -    qed
  25.446 -  qed
  25.447 -next assume ?l show ?r
  25.448 -  proof fix B::real have "open {..<ereal B}" "(-\<infinity>) : {..<ereal B}" by auto
  25.449 -    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
  25.450 -    guess N .. note N=this
  25.451 -    show "EX N. ALL n>=N. ereal B >= f n" apply(rule_tac x=N in exI) using N by auto
  25.452 -  qed
  25.453 -qed
  25.454 -
  25.455 -
  25.456 -lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= ereal B" shows "l ~= \<infinity>"
  25.457 -proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\<infinity>"
  25.458 -  from lim[unfolded this Lim_PInfty,rule_format,of "?B"]
  25.459 -  guess N .. note N=this[rule_format,OF le_refl]
  25.460 -  hence "ereal ?B <= ereal B" using assms(2)[of N] by(rule order_trans)
  25.461 -  hence "ereal ?B < ereal ?B" apply (rule le_less_trans) by auto
  25.462 -  thus False by auto
  25.463 -qed
  25.464 -
  25.465 -
  25.466 -lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= ereal B" shows "l ~= (-\<infinity>)"
  25.467 -proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\<infinity>)"
  25.468 -  from lim[unfolded this Lim_MInfty,rule_format,of "?B"]
  25.469 -  guess N .. note N=this[rule_format,OF le_refl]
  25.470 -  hence "ereal B <= ereal ?B" using assms(2)[of N] order_trans[of "ereal B" "f N" "ereal(B - 1)"] by blast
  25.471 -  thus False by auto
  25.472 -qed
  25.473 -
  25.474 +lemma Lim_PInfty: "f ----> \<infinity> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n \<ge> ereal B)"
  25.475 +  unfolding tendsto_PInfty eventually_sequentially
  25.476 +proof safe
  25.477 +  fix r assume "\<forall>r. \<exists>N. \<forall>n\<ge>N. ereal r \<le> f n"
  25.478 +  from this[rule_format, of "r+1"] guess N ..
  25.479 +  moreover have "ereal r < ereal (r + 1)" by auto
  25.480 +  ultimately show "\<exists>N. \<forall>n\<ge>N. ereal r < f n"
  25.481 +    by (blast intro: less_le_trans)
  25.482 +qed (blast intro: less_imp_le)
  25.483 +
  25.484 +lemma Lim_MInfty: "f ----> -\<infinity> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. ereal B \<ge> f n)"
  25.485 +  unfolding tendsto_MInfty eventually_sequentially
  25.486 +proof safe
  25.487 +  fix r assume "\<forall>r. \<exists>N. \<forall>n\<ge>N. f n \<le> ereal r"
  25.488 +  from this[rule_format, of "r - 1"] guess N ..
  25.489 +  moreover have "ereal (r - 1) < ereal r" by auto
  25.490 +  ultimately show "\<exists>N. \<forall>n\<ge>N. f n < ereal r"
  25.491 +    by (blast intro: le_less_trans)
  25.492 +qed (blast intro: less_imp_le)
  25.493 +
  25.494 +lemma Lim_bounded_PInfty: "f ----> l \<Longrightarrow> (\<And>n. f n \<le> ereal B) \<Longrightarrow> l \<noteq> \<infinity>"
  25.495 +  using LIMSEQ_le_const2[of f l "ereal B"] by auto
  25.496 +
  25.497 +lemma Lim_bounded_MInfty: "f ----> l \<Longrightarrow> (\<And>n. ereal B \<le> f n) \<Longrightarrow> l \<noteq> -\<infinity>"
  25.498 +  using LIMSEQ_le_const[of f l "ereal B"] by auto
  25.499  
  25.500  lemma tendsto_explicit:
  25.501    "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))"
  25.502    unfolding tendsto_def eventually_sequentially by auto
  25.503  
  25.504 -
  25.505 -lemma tendsto_obtains_N:
  25.506 -  assumes "f ----> f0"
  25.507 -  assumes "open S" "f0 : S"
  25.508 -  obtains N where "ALL n>=N. f n : S"
  25.509 -  using tendsto_explicit[of f f0] assms by auto
  25.510 -
  25.511 -
  25.512 -lemma tail_same_limit:
  25.513 -  fixes X Y N
  25.514 -  assumes "X ----> L" "ALL n>=N. X n = Y n"
  25.515 -  shows "Y ----> L"
  25.516 -proof-
  25.517 -{ fix S assume "open S" and "L:S"
  25.518 -  then obtain N1 where "ALL n>=N1. X n : S"
  25.519 -     using assms unfolding tendsto_def eventually_sequentially by auto
  25.520 -  hence "ALL n>=max N N1. Y n : S" using assms by auto
  25.521 -  hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto
  25.522 -}
  25.523 -thus ?thesis using tendsto_explicit by auto
  25.524 -qed
  25.525 -
  25.526 -
  25.527  lemma Lim_bounded_PInfty2:
  25.528 -assumes lim:"f ----> l" and "ALL n>=N. f n <= ereal B"
  25.529 -shows "l ~= \<infinity>"
  25.530 -proof-
  25.531 -  def g == "(%n. if n>=N then f n else ereal B)"
  25.532 -  hence "g ----> l" using tail_same_limit[of f l N g] lim by auto
  25.533 -  moreover have "!!n. g n <= ereal B" using g_def assms by auto
  25.534 -  ultimately show ?thesis using  Lim_bounded_PInfty by auto
  25.535 -qed
  25.536 -
  25.537 -lemma Lim_bounded_ereal:
  25.538 -  assumes lim:"f ----> (l :: ereal)"
  25.539 -  and "ALL n>=M. f n <= C"
  25.540 -  shows "l<=C"
  25.541 -proof-
  25.542 -{ assume "l=(-\<infinity>)" hence ?thesis by auto }
  25.543 -moreover
  25.544 -{ assume "~(l=(-\<infinity>))"
  25.545 -  { assume "C=\<infinity>" hence ?thesis by auto }
  25.546 -  moreover
  25.547 -  { assume "C=(-\<infinity>)" hence "ALL n>=M. f n = (-\<infinity>)" using assms by auto
  25.548 -    hence "l=(-\<infinity>)" using assms
  25.549 -       tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\<lambda>n. -\<infinity>" "-\<infinity>" M f, OF tendsto_const] by auto
  25.550 -    hence ?thesis by auto }
  25.551 -  moreover
  25.552 -  { assume "EX B. C = ereal B"
  25.553 -    then obtain B where B_def: "C=ereal B" by auto
  25.554 -    hence "~(l=\<infinity>)" using Lim_bounded_PInfty2 assms by auto
  25.555 -    then obtain m where m_def: "ereal m=l" using `~(l=(-\<infinity>))` by (cases l) auto
  25.556 -    then obtain N where N_def: "ALL n>=N. f n : {ereal(m - 1) <..< ereal(m+1)}"
  25.557 -       apply (subst tendsto_obtains_N[of f l "{ereal(m - 1) <..< ereal(m+1)}"]) using assms by auto
  25.558 -    { fix n assume "n>=N"
  25.559 -      hence "EX r. ereal r = f n" using N_def by (cases "f n") auto
  25.560 -    } then obtain g where g_def: "ALL n>=N. ereal (g n) = f n" by metis
  25.561 -    hence "(%n. ereal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto
  25.562 -    hence *: "(%n. g n) ----> m" using m_def by auto
  25.563 -    { fix n assume "n>=max N M"
  25.564 -      hence "ereal (g n) <= ereal B" using assms g_def B_def by auto
  25.565 -      hence "g n <= B" by auto
  25.566 -    } hence "EX N. ALL n>=N. g n <= B" by blast
  25.567 -    hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto
  25.568 -    hence ?thesis using m_def B_def by auto
  25.569 -  } ultimately have ?thesis by (cases C) auto
  25.570 -} ultimately show ?thesis by blast
  25.571 -qed
  25.572 +  "f ----> l \<Longrightarrow> ALL n>=N. f n <= ereal B \<Longrightarrow> l ~= \<infinity>"
  25.573 +  using LIMSEQ_le_const2[of f l "ereal B"] by fastforce
  25.574 +
  25.575 +lemma Lim_bounded_ereal: "f ----> (l :: ereal) \<Longrightarrow> ALL n>=M. f n <= C \<Longrightarrow> l<=C"
  25.576 +  by (intro LIMSEQ_le_const2) auto
  25.577  
  25.578  lemma real_of_ereal_mult[simp]:
  25.579    fixes a b :: ereal shows "real (a * b) = real a * real b"
  25.580 @@ -2204,6 +2152,77 @@
  25.581  lemma ereal_mult_m1[simp]: "x * ereal (-1) = -x"
  25.582    by (cases x) auto
  25.583  
  25.584 +lemma ereal_real': assumes "\<bar>x\<bar> \<noteq> \<infinity>" shows "ereal (real x) = x"
  25.585 +  using assms by auto
  25.586 +
  25.587 +lemma ereal_le_ereal_bounded:
  25.588 +  fixes x y z :: ereal
  25.589 +  assumes "z \<le> y"
  25.590 +  assumes *: "\<And>B. z < B \<Longrightarrow> B < x \<Longrightarrow> B \<le> y"
  25.591 +  shows "x \<le> y"
  25.592 +proof (rule ereal_le_ereal)
  25.593 +  fix B assume "B < x"
  25.594 +  show "B \<le> y"
  25.595 +  proof cases
  25.596 +    assume "z < B" from *[OF this `B < x`] show "B \<le> y" .
  25.597 +  next
  25.598 +    assume "\<not> z < B" with `z \<le> y` show "B \<le> y" by auto
  25.599 +  qed
  25.600 +qed
  25.601 +
  25.602 +lemma fixes x y :: ereal
  25.603 +  shows Sup_atMost[simp]: "Sup {.. y} = y"
  25.604 +    and Sup_lessThan[simp]: "Sup {..< y} = y"
  25.605 +    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
  25.606 +    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
  25.607 +    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
  25.608 +  by (auto simp: Sup_ereal_def intro!: Least_equality
  25.609 +           intro: ereal_le_ereal ereal_le_ereal_bounded[of x])
  25.610 +
  25.611 +lemma Sup_greaterThanlessThan[simp]:
  25.612 +  fixes x y :: ereal assumes "x < y" shows "Sup { x <..< y} = y"
  25.613 +  unfolding Sup_ereal_def
  25.614 +proof (intro Least_equality ereal_le_ereal_bounded[of _ _ y])
  25.615 +  fix z assume z: "\<forall>u\<in>{x<..<y}. u \<le> z"
  25.616 +  from ereal_dense[OF `x < y`] guess w .. note w = this
  25.617 +  with z[THEN bspec, of w] show "x \<le> z" by auto
  25.618 +qed auto
  25.619 +
  25.620 +lemma real_ereal_id: "real o ereal = id"
  25.621 +proof-
  25.622 +  { fix x have "(real o ereal) x = id x" by auto }
  25.623 +  then show ?thesis using ext by blast
  25.624 +qed
  25.625 +
  25.626 +lemma open_image_ereal: "open(UNIV-{ \<infinity> , (-\<infinity> :: ereal)})"
  25.627 +by (metis range_ereal open_ereal open_UNIV)
  25.628 +
  25.629 +lemma ereal_le_distrib:
  25.630 +  fixes a b c :: ereal shows "c * (a + b) \<le> c * a + c * b"
  25.631 +  by (cases rule: ereal3_cases[of a b c])
  25.632 +     (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
  25.633 +
  25.634 +lemma ereal_pos_distrib:
  25.635 +  fixes a b c :: ereal assumes "0 \<le> c" "c \<noteq> \<infinity>" shows "c * (a + b) = c * a + c * b"
  25.636 +  using assms by (cases rule: ereal3_cases[of a b c])
  25.637 +                 (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
  25.638 +
  25.639 +lemma ereal_pos_le_distrib:
  25.640 +fixes a b c :: ereal
  25.641 +assumes "c>=0"
  25.642 +shows "c * (a + b) <= c * a + c * b"
  25.643 +  using assms by (cases rule: ereal3_cases[of a b c])
  25.644 +                 (auto simp add: field_simps)
  25.645 +
  25.646 +lemma ereal_max_mono:
  25.647 +  "[| (a::ereal) <= b; c <= d |] ==> max a c <= max b d"
  25.648 +  by (metis sup_ereal_def sup_mono)
  25.649 +
  25.650 +
  25.651 +lemma ereal_max_least:
  25.652 +  "[| (a::ereal) <= x; c <= x |] ==> max a c <= x"
  25.653 +  by (metis sup_ereal_def sup_least)
  25.654 +
  25.655  lemma ereal_LimI_finite:
  25.656    fixes x :: ereal
  25.657    assumes "\<bar>x\<bar> \<noteq> \<infinity>"
  25.658 @@ -2234,6 +2253,12 @@
  25.659    qed
  25.660  qed
  25.661  
  25.662 +lemma tendsto_obtains_N:
  25.663 +  assumes "f ----> f0"
  25.664 +  assumes "open S" "f0 : S"
  25.665 +  obtains N where "ALL n>=N. f n : S"
  25.666 +  using tendsto_explicit[of f f0] assms by auto
  25.667 +
  25.668  lemma ereal_LimI_finite_iff:
  25.669    fixes x :: ereal
  25.670    assumes "\<bar>x\<bar> \<noteq> \<infinity>"
  25.671 @@ -2257,64 +2282,302 @@
  25.672  subsubsection {* @{text Liminf} and @{text Limsup} *}
  25.673  
  25.674  definition
  25.675 -  "Liminf net f = (GREATEST l. \<forall>y<l. eventually (\<lambda>x. y < f x) net)"
  25.676 +  "Liminf F f = (SUP P:{P. eventually P F}. INF x:{x. P x}. f x)"
  25.677  
  25.678  definition
  25.679 -  "Limsup net f = (LEAST l. \<forall>y>l. eventually (\<lambda>x. f x < y) net)"
  25.680 -
  25.681 -lemma Liminf_Sup:
  25.682 -  fixes f :: "'a => 'b::complete_linorder"
  25.683 -  shows "Liminf net f = Sup {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net}"
  25.684 -  by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def)
  25.685 -
  25.686 -lemma Limsup_Inf:
  25.687 -  fixes f :: "'a => 'b::complete_linorder"
  25.688 -  shows "Limsup net f = Inf {l. \<forall>y>l. eventually (\<lambda>x. f x < y) net}"
  25.689 -  by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def)
  25.690 -
  25.691 -lemma ereal_SupI:
  25.692 -  fixes x :: ereal
  25.693 -  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
  25.694 -  assumes "\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y"
  25.695 -  shows "Sup A = x"
  25.696 -  unfolding Sup_ereal_def
  25.697 -  using assms by (auto intro!: Least_equality)
  25.698 -
  25.699 -lemma ereal_InfI:
  25.700 -  fixes x :: ereal
  25.701 -  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> i"
  25.702 -  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x"
  25.703 -  shows "Inf A = x"
  25.704 -  unfolding Inf_ereal_def
  25.705 -  using assms by (auto intro!: Greatest_equality)
  25.706 -
  25.707 -lemma Limsup_const:
  25.708 -  fixes c :: "'a::complete_linorder"
  25.709 -  assumes ntriv: "\<not> trivial_limit net"
  25.710 -  shows "Limsup net (\<lambda>x. c) = c"
  25.711 -  unfolding Limsup_Inf
  25.712 -proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower)
  25.713 -  fix x assume *: "\<forall>y>x. eventually (\<lambda>_. c < y) net"
  25.714 -  show "c \<le> x"
  25.715 -  proof (rule ccontr)
  25.716 -    assume "\<not> c \<le> x" then have "x < c" by auto
  25.717 -    then show False using ntriv * by (auto simp: trivial_limit_def)
  25.718 +  "Limsup F f = (INF P:{P. eventually P F}. SUP x:{x. P x}. f x)"
  25.719 +
  25.720 +abbreviation "liminf \<equiv> Liminf sequentially"
  25.721 +
  25.722 +abbreviation "limsup \<equiv> Limsup sequentially"
  25.723 +
  25.724 +lemma Liminf_eqI:
  25.725 +  "(\<And>P. eventually P F \<Longrightarrow> INFI (Collect P) f \<le> x) \<Longrightarrow>  
  25.726 +    (\<And>y. (\<And>P. eventually P F \<Longrightarrow> INFI (Collect P) f \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> Liminf F f = x"
  25.727 +  unfolding Liminf_def by (auto intro!: SUP_eqI)
  25.728 +
  25.729 +lemma Limsup_eqI:
  25.730 +  "(\<And>P. eventually P F \<Longrightarrow> x \<le> SUPR (Collect P) f) \<Longrightarrow>  
  25.731 +    (\<And>y. (\<And>P. eventually P F \<Longrightarrow> y \<le> SUPR (Collect P) f) \<Longrightarrow> y \<le> x) \<Longrightarrow> Limsup F f = x"
  25.732 +  unfolding Limsup_def by (auto intro!: INF_eqI)
  25.733 +
  25.734 +lemma liminf_SUPR_INFI:
  25.735 +  fixes f :: "nat \<Rightarrow> 'a :: complete_lattice"
  25.736 +  shows "liminf f = (SUP n. INF m:{n..}. f m)"
  25.737 +  unfolding Liminf_def eventually_sequentially
  25.738 +  by (rule SUPR_eq) (auto simp: atLeast_def intro!: INF_mono)
  25.739 +
  25.740 +lemma limsup_INFI_SUPR:
  25.741 +  fixes f :: "nat \<Rightarrow> 'a :: complete_lattice"
  25.742 +  shows "limsup f = (INF n. SUP m:{n..}. f m)"
  25.743 +  unfolding Limsup_def eventually_sequentially
  25.744 +  by (rule INFI_eq) (auto simp: atLeast_def intro!: SUP_mono)
  25.745 +
  25.746 +lemma Limsup_const: 
  25.747 +  assumes ntriv: "\<not> trivial_limit F"
  25.748 +  shows "Limsup F (\<lambda>x. c) = (c::'a::complete_lattice)"
  25.749 +proof -
  25.750 +  have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto
  25.751 +  have "\<And>P. eventually P F \<Longrightarrow> (SUP x : {x. P x}. c) = c"
  25.752 +    using ntriv by (intro SUP_const) (auto simp: eventually_False *)
  25.753 +  then show ?thesis
  25.754 +    unfolding Limsup_def using eventually_True
  25.755 +    by (subst INF_cong[where D="\<lambda>x. c"])
  25.756 +       (auto intro!: INF_const simp del: eventually_True)
  25.757 +qed
  25.758 +
  25.759 +lemma Liminf_const:
  25.760 +  assumes ntriv: "\<not> trivial_limit F"
  25.761 +  shows "Liminf F (\<lambda>x. c) = (c::'a::complete_lattice)"
  25.762 +proof -
  25.763 +  have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto
  25.764 +  have "\<And>P. eventually P F \<Longrightarrow> (INF x : {x. P x}. c) = c"
  25.765 +    using ntriv by (intro INF_const) (auto simp: eventually_False *)
  25.766 +  then show ?thesis
  25.767 +    unfolding Liminf_def using eventually_True
  25.768 +    by (subst SUP_cong[where D="\<lambda>x. c"])
  25.769 +       (auto intro!: SUP_const simp del: eventually_True)
  25.770 +qed
  25.771 +
  25.772 +lemma Liminf_mono:
  25.773 +  fixes f g :: "'a => 'b :: complete_lattice"
  25.774 +  assumes ev: "eventually (\<lambda>x. f x \<le> g x) F"
  25.775 +  shows "Liminf F f \<le> Liminf F g"
  25.776 +  unfolding Liminf_def
  25.777 +proof (safe intro!: SUP_mono)
  25.778 +  fix P assume "eventually P F"
  25.779 +  with ev have "eventually (\<lambda>x. f x \<le> g x \<and> P x) F" (is "eventually ?Q F") by (rule eventually_conj)
  25.780 +  then show "\<exists>Q\<in>{P. eventually P F}. INFI (Collect P) f \<le> INFI (Collect Q) g"
  25.781 +    by (intro bexI[of _ ?Q]) (auto intro!: INF_mono)
  25.782 +qed
  25.783 +
  25.784 +lemma Liminf_eq:
  25.785 +  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
  25.786 +  assumes "eventually (\<lambda>x. f x = g x) F"
  25.787 +  shows "Liminf F f = Liminf F g"
  25.788 +  by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
  25.789 +
  25.790 +lemma Limsup_mono:
  25.791 +  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
  25.792 +  assumes ev: "eventually (\<lambda>x. f x \<le> g x) F"
  25.793 +  shows "Limsup F f \<le> Limsup F g"
  25.794 +  unfolding Limsup_def
  25.795 +proof (safe intro!: INF_mono)
  25.796 +  fix P assume "eventually P F"
  25.797 +  with ev have "eventually (\<lambda>x. f x \<le> g x \<and> P x) F" (is "eventually ?Q F") by (rule eventually_conj)
  25.798 +  then show "\<exists>Q\<in>{P. eventually P F}. SUPR (Collect Q) f \<le> SUPR (Collect P) g"
  25.799 +    by (intro bexI[of _ ?Q]) (auto intro!: SUP_mono)
  25.800 +qed
  25.801 +
  25.802 +lemma Limsup_eq:
  25.803 +  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
  25.804 +  assumes "eventually (\<lambda>x. f x = g x) net"
  25.805 +  shows "Limsup net f = Limsup net g"
  25.806 +  by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
  25.807 +
  25.808 +lemma Liminf_le_Limsup:
  25.809 +  fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
  25.810 +  assumes ntriv: "\<not> trivial_limit F"
  25.811 +  shows "Liminf F f \<le> Limsup F f"
  25.812 +  unfolding Limsup_def Liminf_def
  25.813 +  apply (rule complete_lattice_class.SUP_least)
  25.814 +  apply (rule complete_lattice_class.INF_greatest)
  25.815 +proof safe
  25.816 +  fix P Q assume "eventually P F" "eventually Q F"
  25.817 +  then have "eventually (\<lambda>x. P x \<and> Q x) F" (is "eventually ?C F") by (rule eventually_conj)
  25.818 +  then have not_False: "(\<lambda>x. P x \<and> Q x) \<noteq> (\<lambda>x. False)"
  25.819 +    using ntriv by (auto simp add: eventually_False)
  25.820 +  have "INFI (Collect P) f \<le> INFI (Collect ?C) f"
  25.821 +    by (rule INF_mono) auto
  25.822 +  also have "\<dots> \<le> SUPR (Collect ?C) f"
  25.823 +    using not_False by (intro INF_le_SUP) auto
  25.824 +  also have "\<dots> \<le> SUPR (Collect Q) f"
  25.825 +    by (rule SUP_mono) auto
  25.826 +  finally show "INFI (Collect P) f \<le> SUPR (Collect Q) f" .
  25.827 +qed
  25.828 +
  25.829 +lemma
  25.830 +  fixes X :: "nat \<Rightarrow> ereal"
  25.831 +  shows ereal_incseq_uminus[simp]: "incseq (\<lambda>i. - X i) = decseq X"
  25.832 +    and ereal_decseq_uminus[simp]: "decseq (\<lambda>i. - X i) = incseq X"
  25.833 +  unfolding incseq_def decseq_def by auto
  25.834 +
  25.835 +lemma Liminf_bounded:
  25.836 +  fixes X Y :: "'a \<Rightarrow> 'b::complete_lattice"
  25.837 +  assumes ntriv: "\<not> trivial_limit F"
  25.838 +  assumes le: "eventually (\<lambda>n. C \<le> X n) F"
  25.839 +  shows "C \<le> Liminf F X"
  25.840 +  using Liminf_mono[OF le] Liminf_const[OF ntriv, of C] by simp
  25.841 +
  25.842 +lemma Limsup_bounded:
  25.843 +  fixes X Y :: "'a \<Rightarrow> 'b::complete_lattice"
  25.844 +  assumes ntriv: "\<not> trivial_limit F"
  25.845 +  assumes le: "eventually (\<lambda>n. X n \<le> C) F"
  25.846 +  shows "Limsup F X \<le> C"
  25.847 +  using Limsup_mono[OF le] Limsup_const[OF ntriv, of C] by simp
  25.848 +
  25.849 +lemma le_Liminf_iff:
  25.850 +  fixes X :: "_ \<Rightarrow> _ :: complete_linorder"
  25.851 +  shows "C \<le> Liminf F X \<longleftrightarrow> (\<forall>y<C. eventually (\<lambda>x. y < X x) F)"
  25.852 +proof -
  25.853 +  { fix y P assume "eventually P F" "y < INFI (Collect P) X"
  25.854 +    then have "eventually (\<lambda>x. y < X x) F"
  25.855 +      by (auto elim!: eventually_elim1 dest: less_INF_D) }
  25.856 +  moreover
  25.857 +  { fix y P assume "y < C" and y: "\<forall>y<C. eventually (\<lambda>x. y < X x) F"
  25.858 +    have "\<exists>P. eventually P F \<and> y < INFI (Collect P) X"
  25.859 +    proof cases
  25.860 +      assume "\<exists>z. y < z \<and> z < C"
  25.861 +      then guess z ..
  25.862 +      moreover then have "z \<le> INFI {x. z < X x} X"
  25.863 +        by (auto intro!: INF_greatest)
  25.864 +      ultimately show ?thesis
  25.865 +        using y by (intro exI[of _ "\<lambda>x. z < X x"]) auto
  25.866 +    next
  25.867 +      assume "\<not> (\<exists>z. y < z \<and> z < C)"
  25.868 +      then have "C \<le> INFI {x. y < X x} X"
  25.869 +        by (intro INF_greatest) auto
  25.870 +      with `y < C` show ?thesis
  25.871 +        using y by (intro exI[of _ "\<lambda>x. y < X x"]) auto
  25.872 +    qed }
  25.873 +  ultimately show ?thesis
  25.874 +    unfolding Liminf_def le_SUP_iff by auto
  25.875 +qed
  25.876 +
  25.877 +lemma lim_imp_Liminf:
  25.878 +  fixes f :: "'a \<Rightarrow> _ :: {complete_linorder, linorder_topology}"
  25.879 +  assumes ntriv: "\<not> trivial_limit F"
  25.880 +  assumes lim: "(f ---> f0) F"
  25.881 +  shows "Liminf F f = f0"
  25.882 +proof (intro Liminf_eqI)
  25.883 +  fix P assume P: "eventually P F"
  25.884 +  then have "eventually (\<lambda>x. INFI (Collect P) f \<le> f x) F"
  25.885 +    by eventually_elim (auto intro!: INF_lower)
  25.886 +  then show "INFI (Collect P) f \<le> f0"
  25.887 +    by (rule tendsto_le[OF ntriv lim tendsto_const])
  25.888 +next
  25.889 +  fix y assume upper: "\<And>P. eventually P F \<Longrightarrow> INFI (Collect P) f \<le> y"
  25.890 +  show "f0 \<le> y"
  25.891 +  proof cases
  25.892 +    assume "\<exists>z. y < z \<and> z < f0"
  25.893 +    then guess z ..
  25.894 +    moreover have "z \<le> INFI {x. z < f x} f"
  25.895 +      by (rule INF_greatest) simp
  25.896 +    ultimately show ?thesis
  25.897 +      using lim[THEN topological_tendstoD, THEN upper, of "{z <..}"] by auto
  25.898 +  next
  25.899 +    assume discrete: "\<not> (\<exists>z. y < z \<and> z < f0)"
  25.900 +    show ?thesis
  25.901 +    proof (rule classical)
  25.902 +      assume "\<not> f0 \<le> y"
  25.903 +      then have "eventually (\<lambda>x. y < f x) F"
  25.904 +        using lim[THEN topological_tendstoD, of "{y <..}"] by auto
  25.905 +      then have "eventually (\<lambda>x. f0 \<le> f x) F"
  25.906 +        using discrete by (auto elim!: eventually_elim1)
  25.907 +      then have "INFI {x. f0 \<le> f x} f \<le> y"
  25.908 +        by (rule upper)
  25.909 +      moreover have "f0 \<le> INFI {x. f0 \<le> f x} f"
  25.910 +        by (intro INF_greatest) simp
  25.911 +      ultimately show "f0 \<le> y" by simp
  25.912 +    qed
  25.913    qed
  25.914 -qed auto
  25.915 -
  25.916 -lemma Liminf_const:
  25.917 -  fixes c :: "'a::complete_linorder"
  25.918 -  assumes ntriv: "\<not> trivial_limit net"
  25.919 -  shows "Liminf net (\<lambda>x. c) = c"
  25.920 -  unfolding Liminf_Sup
  25.921 -proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
  25.922 -  fix x assume *: "\<forall>y<x. eventually (\<lambda>_. y < c) net"
  25.923 -  show "x \<le> c"
  25.924 -  proof (rule ccontr)
  25.925 -    assume "\<not> x \<le> c" then have "c < x" by auto
  25.926 -    then show False using ntriv * by (auto simp: trivial_limit_def)
  25.927 +qed
  25.928 +
  25.929 +lemma lim_imp_Limsup:
  25.930 +  fixes f :: "'a \<Rightarrow> _ :: {complete_linorder, linorder_topology}"
  25.931 +  assumes ntriv: "\<not> trivial_limit F"
  25.932 +  assumes lim: "(f ---> f0) F"
  25.933 +  shows "Limsup F f = f0"
  25.934 +proof (intro Limsup_eqI)
  25.935 +  fix P assume P: "eventually P F"
  25.936 +  then have "eventually (\<lambda>x. f x \<le> SUPR (Collect P) f) F"
  25.937 +    by eventually_elim (auto intro!: SUP_upper)
  25.938 +  then show "f0 \<le> SUPR (Collect P) f"
  25.939 +    by (rule tendsto_le[OF ntriv tendsto_const lim])
  25.940 +next
  25.941 +  fix y assume lower: "\<And>P. eventually P F \<Longrightarrow> y \<le> SUPR (Collect P) f"
  25.942 +  show "y \<le> f0"
  25.943 +  proof cases
  25.944 +    assume "\<exists>z. f0 < z \<and> z < y"
  25.945 +    then guess z ..
  25.946 +    moreover have "SUPR {x. f x < z} f \<le> z"
  25.947 +      by (rule SUP_least) simp
  25.948 +    ultimately show ?thesis
  25.949 +      using lim[THEN topological_tendstoD, THEN lower, of "{..< z}"] by auto
  25.950 +  next
  25.951 +    assume discrete: "\<not> (\<exists>z. f0 < z \<and> z < y)"
  25.952 +    show ?thesis
  25.953 +    proof (rule classical)
  25.954 +      assume "\<not> y \<le> f0"
  25.955 +      then have "eventually (\<lambda>x. f x < y) F"
  25.956 +        using lim[THEN topological_tendstoD, of "{..< y}"] by auto
  25.957 +      then have "eventually (\<lambda>x. f x \<le> f0) F"
  25.958 +        using discrete by (auto elim!: eventually_elim1 simp: not_less)
  25.959 +      then have "y \<le> SUPR {x. f x \<le> f0} f"
  25.960 +        by (rule lower)
  25.961 +      moreover have "SUPR {x. f x \<le> f0} f \<le> f0"
  25.962 +        by (intro SUP_least) simp
  25.963 +      ultimately show "y \<le> f0" by simp
  25.964 +    qed
  25.965    qed
  25.966 -qed auto
  25.967 +qed
  25.968 +
  25.969 +lemma Liminf_eq_Limsup:
  25.970 +  fixes f0 :: "'a :: {complete_linorder, linorder_topology}"
  25.971 +  assumes ntriv: "\<not> trivial_limit F"
  25.972 +    and lim: "Liminf F f = f0" "Limsup F f = f0"
  25.973 +  shows "(f ---> f0) F"
  25.974 +proof (rule order_tendstoI)
  25.975 +  fix a assume "f0 < a"
  25.976 +  with assms have "Limsup F f < a" by simp
  25.977 +  then obtain P where "eventually P F" "SUPR (Collect P) f < a"
  25.978 +    unfolding Limsup_def INF_less_iff by auto
  25.979 +  then show "eventually (\<lambda>x. f x < a) F"
  25.980 +    by (auto elim!: eventually_elim1 dest: SUP_lessD)
  25.981 +next
  25.982 +  fix a assume "a < f0"
  25.983 +  with assms have "a < Liminf F f" by simp
  25.984 +  then obtain P where "eventually P F" "a < INFI (Collect P) f"
  25.985 +    unfolding Liminf_def less_SUP_iff by auto
  25.986 +  then show "eventually (\<lambda>x. a < f x) F"
  25.987 +    by (auto elim!: eventually_elim1 dest: less_INF_D)
  25.988 +qed
  25.989 +
  25.990 +lemma tendsto_iff_Liminf_eq_Limsup:
  25.991 +  fixes f0 :: "'a :: {complete_linorder, linorder_topology}"
  25.992 +  shows "\<not> trivial_limit F \<Longrightarrow> (f ---> f0) F \<longleftrightarrow> (Liminf F f = f0 \<and> Limsup F f = f0)"
  25.993 +  by (metis Liminf_eq_Limsup lim_imp_Limsup lim_imp_Liminf)
  25.994 +
  25.995 +lemma liminf_bounded_iff:
  25.996 +  fixes x :: "nat \<Rightarrow> ereal"
  25.997 +  shows "C \<le> liminf x \<longleftrightarrow> (\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n)" (is "?lhs <-> ?rhs")
  25.998 +  unfolding le_Liminf_iff eventually_sequentially ..
  25.999 +
 25.1000 +lemma liminf_subseq_mono:
 25.1001 +  fixes X :: "nat \<Rightarrow> 'a :: complete_linorder"
 25.1002 +  assumes "subseq r"
 25.1003 +  shows "liminf X \<le> liminf (X \<circ> r) "
 25.1004 +proof-
 25.1005 +  have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
 25.1006 +  proof (safe intro!: INF_mono)
 25.1007 +    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
 25.1008 +      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
 25.1009 +  qed
 25.1010 +  then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def)
 25.1011 +qed
 25.1012 +
 25.1013 +lemma limsup_subseq_mono:
 25.1014 +  fixes X :: "nat \<Rightarrow> 'a :: complete_linorder"
 25.1015 +  assumes "subseq r"
 25.1016 +  shows "limsup (X \<circ> r) \<le> limsup X"
 25.1017 +proof-
 25.1018 +  have "\<And>n. (SUP m:{n..}. (X \<circ> r) m) \<le> (SUP m:{n..}. X m)"
 25.1019 +  proof (safe intro!: SUP_mono)
 25.1020 +    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. (X \<circ> r) m \<le> X ma"
 25.1021 +      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
 25.1022 +  qed
 25.1023 +  then show ?thesis by (auto intro!: INF_mono simp: limsup_INFI_SUPR comp_def)
 25.1024 +qed
 25.1025  
 25.1026  definition (in order) mono_set:
 25.1027    "mono_set S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
 25.1028 @@ -2353,271 +2616,6 @@
 25.1029    qed
 25.1030  qed auto
 25.1031  
 25.1032 -lemma lim_imp_Liminf:
 25.1033 -  fixes f :: "'a \<Rightarrow> ereal"
 25.1034 -  assumes ntriv: "\<not> trivial_limit net"
 25.1035 -  assumes lim: "(f ---> f0) net"
 25.1036 -  shows "Liminf net f = f0"
 25.1037 -  unfolding Liminf_Sup
 25.1038 -proof (safe intro!: ereal_SupI)
 25.1039 -  fix y assume *: "\<forall>y'<y. eventually (\<lambda>x. y' < f x) net"
 25.1040 -  show "y \<le> f0"
 25.1041 -  proof (rule ereal_le_ereal)
 25.1042 -    fix B assume "B < y"
 25.1043 -    { assume "f0 < B"
 25.1044 -      then have "eventually (\<lambda>x. f x < B \<and> B < f x) net"
 25.1045 -         using topological_tendstoD[OF lim, of "{..<B}"] *[rule_format, OF `B < y`]
 25.1046 -         by (auto intro: eventually_conj)
 25.1047 -      also have "(\<lambda>x. f x < B \<and> B < f x) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
 25.1048 -      finally have False using ntriv[unfolded trivial_limit_def] by auto
 25.1049 -    } then show "B \<le> f0" by (metis linorder_le_less_linear)
 25.1050 -  qed
 25.1051 -next
 25.1052 -  fix y assume *: "\<forall>z. z \<in> {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net} \<longrightarrow> z \<le> y"
 25.1053 -  show "f0 \<le> y"
 25.1054 -  proof (safe intro!: *[rule_format])
 25.1055 -    fix y assume "y < f0" then show "eventually (\<lambda>x. y < f x) net"
 25.1056 -      using lim[THEN topological_tendstoD, of "{y <..}"] by auto
 25.1057 -  qed
 25.1058 -qed
 25.1059 -
 25.1060 -lemma ereal_Liminf_le_Limsup:
 25.1061 -  fixes f :: "'a \<Rightarrow> ereal"
 25.1062 -  assumes ntriv: "\<not> trivial_limit net"
 25.1063 -  shows "Liminf net f \<le> Limsup net f"
 25.1064 -  unfolding Limsup_Inf Liminf_Sup
 25.1065 -proof (safe intro!: complete_lattice_class.Inf_greatest  complete_lattice_class.Sup_least)
 25.1066 -  fix u v assume *: "\<forall>y<u. eventually (\<lambda>x. y < f x) net" "\<forall>y>v. eventually (\<lambda>x. f x < y) net"
 25.1067 -  show "u \<le> v"
 25.1068 -  proof (rule ccontr)
 25.1069 -    assume "\<not> u \<le> v"
 25.1070 -    then obtain t where "t < u" "v < t"
 25.1071 -      using ereal_dense[of v u] by (auto simp: not_le)
 25.1072 -    then have "eventually (\<lambda>x. t < f x \<and> f x < t) net"
 25.1073 -      using * by (auto intro: eventually_conj)
 25.1074 -    also have "(\<lambda>x. t < f x \<and> f x < t) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
 25.1075 -    finally show False using ntriv by (auto simp: trivial_limit_def)
 25.1076 -  qed
 25.1077 -qed
 25.1078 -
 25.1079 -lemma Liminf_mono:
 25.1080 -  fixes f g :: "'a => ereal"
 25.1081 -  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
 25.1082 -  shows "Liminf net f \<le> Liminf net g"
 25.1083 -  unfolding Liminf_Sup
 25.1084 -proof (safe intro!: Sup_mono bexI)
 25.1085 -  fix a y assume "\<forall>y<a. eventually (\<lambda>x. y < f x) net" and "y < a"
 25.1086 -  then have "eventually (\<lambda>x. y < f x) net" by auto
 25.1087 -  then show "eventually (\<lambda>x. y < g x) net"
 25.1088 -    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
 25.1089 -qed simp
 25.1090 -
 25.1091 -lemma Liminf_eq:
 25.1092 -  fixes f g :: "'a \<Rightarrow> ereal"
 25.1093 -  assumes "eventually (\<lambda>x. f x = g x) net"
 25.1094 -  shows "Liminf net f = Liminf net g"
 25.1095 -  by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
 25.1096 -
 25.1097 -lemma Liminf_mono_all:
 25.1098 -  fixes f g :: "'a \<Rightarrow> ereal"
 25.1099 -  assumes "\<And>x. f x \<le> g x"
 25.1100 -  shows "Liminf net f \<le> Liminf net g"
 25.1101 -  using assms by (intro Liminf_mono always_eventually) auto
 25.1102 -
 25.1103 -lemma Limsup_mono:
 25.1104 -  fixes f g :: "'a \<Rightarrow> ereal"
 25.1105 -  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
 25.1106 -  shows "Limsup net f \<le> Limsup net g"
 25.1107 -  unfolding Limsup_Inf
 25.1108 -proof (safe intro!: Inf_mono bexI)
 25.1109 -  fix a y assume "\<forall>y>a. eventually (\<lambda>x. g x < y) net" and "a < y"
 25.1110 -  then have "eventually (\<lambda>x. g x < y) net" by auto
 25.1111 -  then show "eventually (\<lambda>x. f x < y) net"
 25.1112 -    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
 25.1113 -qed simp
 25.1114 -
 25.1115 -lemma Limsup_mono_all:
 25.1116 -  fixes f g :: "'a \<Rightarrow> ereal"
 25.1117 -  assumes "\<And>x. f x \<le> g x"
 25.1118 -  shows "Limsup net f \<le> Limsup net g"
 25.1119 -  using assms by (intro Limsup_mono always_eventually) auto
 25.1120 -
 25.1121 -lemma Limsup_eq:
 25.1122 -  fixes f g :: "'a \<Rightarrow> ereal"
 25.1123 -  assumes "eventually (\<lambda>x. f x = g x) net"
 25.1124 -  shows "Limsup net f = Limsup net g"
 25.1125 -  by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
 25.1126 -
 25.1127 -abbreviation "liminf \<equiv> Liminf sequentially"
 25.1128 -
 25.1129 -abbreviation "limsup \<equiv> Limsup sequentially"
 25.1130 -
 25.1131 -lemma liminf_SUPR_INFI:
 25.1132 -  fixes f :: "nat \<Rightarrow> ereal"
 25.1133 -  shows "liminf f = (SUP n. INF m:{n..}. f m)"
 25.1134 -  unfolding Liminf_Sup eventually_sequentially
 25.1135 -proof (safe intro!: antisym complete_lattice_class.Sup_least)
 25.1136 -  fix x assume *: "\<forall>y<x. \<exists>N. \<forall>n\<ge>N. y < f n" show "x \<le> (SUP n. INF m:{n..}. f m)"
 25.1137 -  proof (rule ereal_le_ereal)
 25.1138 -    fix y assume "y < x"
 25.1139 -    with * obtain N where "\<And>n. N \<le> n \<Longrightarrow> y < f n" by auto
 25.1140 -    then have "y \<le> (INF m:{N..}. f m)" by (force simp: le_INF_iff)
 25.1141 -    also have "\<dots> \<le> (SUP n. INF m:{n..}. f m)" by (intro SUP_upper) auto
 25.1142 -    finally show "y \<le> (SUP n. INF m:{n..}. f m)" .
 25.1143 -  qed
 25.1144 -next
 25.1145 -  show "(SUP n. INF m:{n..}. f m) \<le> Sup {l. \<forall>y<l. \<exists>N. \<forall>n\<ge>N. y < f n}"
 25.1146 -  proof (unfold SUP_def, safe intro!: Sup_mono bexI)
 25.1147 -    fix y n assume "y < INFI {n..} f"
 25.1148 -    from less_INF_D[OF this] show "\<exists>N. \<forall>n\<ge>N. y < f n" by (intro exI[of _ n]) auto
 25.1149 -  qed (rule order_refl)
 25.1150 -qed
 25.1151 -
 25.1152 -lemma tail_same_limsup:
 25.1153 -  fixes X Y :: "nat => ereal"
 25.1154 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
 25.1155 -  shows "limsup X = limsup Y"
 25.1156 -  using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto
 25.1157 -
 25.1158 -lemma tail_same_liminf:
 25.1159 -  fixes X Y :: "nat => ereal"
 25.1160 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
 25.1161 -  shows "liminf X = liminf Y"
 25.1162 -  using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto
 25.1163 -
 25.1164 -lemma liminf_mono:
 25.1165 -  fixes X Y :: "nat \<Rightarrow> ereal"
 25.1166 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
 25.1167 -  shows "liminf X \<le> liminf Y"
 25.1168 -  using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto
 25.1169 -
 25.1170 -lemma limsup_mono:
 25.1171 -  fixes X Y :: "nat => ereal"
 25.1172 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
 25.1173 -  shows "limsup X \<le> limsup Y"
 25.1174 -  using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto
 25.1175 -
 25.1176 -lemma
 25.1177 -  fixes X :: "nat \<Rightarrow> ereal"
 25.1178 -  shows ereal_incseq_uminus[simp]: "incseq (\<lambda>i. - X i) = decseq X"
 25.1179 -    and ereal_decseq_uminus[simp]: "decseq (\<lambda>i. - X i) = incseq X"
 25.1180 -  unfolding incseq_def decseq_def by auto
 25.1181 -
 25.1182 -lemma liminf_bounded:
 25.1183 -  fixes X Y :: "nat \<Rightarrow> ereal"
 25.1184 -  assumes "\<And>n. N \<le> n \<Longrightarrow> C \<le> X n"
 25.1185 -  shows "C \<le> liminf X"
 25.1186 -  using liminf_mono[of N "\<lambda>n. C" X] assms Liminf_const[of sequentially C] by simp
 25.1187 -
 25.1188 -lemma limsup_bounded:
 25.1189 -  fixes X Y :: "nat => ereal"
 25.1190 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= C"
 25.1191 -  shows "limsup X \<le> C"
 25.1192 -  using limsup_mono[of N X "\<lambda>n. C"] assms Limsup_const[of sequentially C] by simp
 25.1193 -
 25.1194 -lemma liminf_bounded_iff:
 25.1195 -  fixes x :: "nat \<Rightarrow> ereal"
 25.1196 -  shows "C \<le> liminf x \<longleftrightarrow> (\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n)" (is "?lhs <-> ?rhs")
 25.1197 -proof safe
 25.1198 -  fix B assume "B < C" "C \<le> liminf x"
 25.1199 -  then have "B < liminf x" by auto
 25.1200 -  then obtain N where "B < (INF m:{N..}. x m)"
 25.1201 -    unfolding liminf_SUPR_INFI SUP_def less_Sup_iff by auto
 25.1202 -  from less_INF_D[OF this] show "\<exists>N. \<forall>n\<ge>N. B < x n" by auto
 25.1203 -next
 25.1204 -  assume *: "\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n"
 25.1205 -  { fix B assume "B<C"
 25.1206 -    then obtain N where "\<forall>n\<ge>N. B < x n" using `?rhs` by auto
 25.1207 -    hence "B \<le> (INF m:{N..}. x m)" by (intro INF_greatest) auto
 25.1208 -    also have "... \<le> liminf x" unfolding liminf_SUPR_INFI by (intro SUP_upper) simp
 25.1209 -    finally have "B \<le> liminf x" .
 25.1210 -  } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear)
 25.1211 -qed
 25.1212 -
 25.1213 -lemma liminf_subseq_mono:
 25.1214 -  fixes X :: "nat \<Rightarrow> ereal"
 25.1215 -  assumes "subseq r"
 25.1216 -  shows "liminf X \<le> liminf (X \<circ> r) "
 25.1217 -proof-
 25.1218 -  have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
 25.1219 -  proof (safe intro!: INF_mono)
 25.1220 -    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
 25.1221 -      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
 25.1222 -  qed
 25.1223 -  then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def)
 25.1224 -qed
 25.1225 -
 25.1226 -lemma ereal_real': assumes "\<bar>x\<bar> \<noteq> \<infinity>" shows "ereal (real x) = x"
 25.1227 -  using assms by auto
 25.1228 -
 25.1229 -lemma ereal_le_ereal_bounded:
 25.1230 -  fixes x y z :: ereal
 25.1231 -  assumes "z \<le> y"
 25.1232 -  assumes *: "\<And>B. z < B \<Longrightarrow> B < x \<Longrightarrow> B \<le> y"
 25.1233 -  shows "x \<le> y"
 25.1234 -proof (rule ereal_le_ereal)
 25.1235 -  fix B assume "B < x"
 25.1236 -  show "B \<le> y"
 25.1237 -  proof cases
 25.1238 -    assume "z < B" from *[OF this `B < x`] show "B \<le> y" .
 25.1239 -  next
 25.1240 -    assume "\<not> z < B" with `z \<le> y` show "B \<le> y" by auto
 25.1241 -  qed
 25.1242 -qed
 25.1243 -
 25.1244 -lemma fixes x y :: ereal
 25.1245 -  shows Sup_atMost[simp]: "Sup {.. y} = y"
 25.1246 -    and Sup_lessThan[simp]: "Sup {..< y} = y"
 25.1247 -    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
 25.1248 -    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
 25.1249 -    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
 25.1250 -  by (auto simp: Sup_ereal_def intro!: Least_equality
 25.1251 -           intro: ereal_le_ereal ereal_le_ereal_bounded[of x])
 25.1252 -
 25.1253 -lemma Sup_greaterThanlessThan[simp]:
 25.1254 -  fixes x y :: ereal assumes "x < y" shows "Sup { x <..< y} = y"
 25.1255 -  unfolding Sup_ereal_def
 25.1256 -proof (intro Least_equality ereal_le_ereal_bounded[of _ _ y])
 25.1257 -  fix z assume z: "\<forall>u\<in>{x<..<y}. u \<le> z"
 25.1258 -  from ereal_dense[OF `x < y`] guess w .. note w = this
 25.1259 -  with z[THEN bspec, of w] show "x \<le> z" by auto
 25.1260 -qed auto
 25.1261 -
 25.1262 -lemma real_ereal_id: "real o ereal = id"
 25.1263 -proof-
 25.1264 -  { fix x have "(real o ereal) x = id x" by auto }
 25.1265 -  then show ?thesis using ext by blast
 25.1266 -qed
 25.1267 -
 25.1268 -lemma open_image_ereal: "open(UNIV-{ \<infinity> , (-\<infinity> :: ereal)})"
 25.1269 -by (metis range_ereal open_ereal open_UNIV)
 25.1270 -
 25.1271 -lemma ereal_le_distrib:
 25.1272 -  fixes a b c :: ereal shows "c * (a + b) \<le> c * a + c * b"
 25.1273 -  by (cases rule: ereal3_cases[of a b c])
 25.1274 -     (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
 25.1275 -
 25.1276 -lemma ereal_pos_distrib:
 25.1277 -  fixes a b c :: ereal assumes "0 \<le> c" "c \<noteq> \<infinity>" shows "c * (a + b) = c * a + c * b"
 25.1278 -  using assms by (cases rule: ereal3_cases[of a b c])
 25.1279 -                 (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
 25.1280 -
 25.1281 -lemma ereal_pos_le_distrib:
 25.1282 -fixes a b c :: ereal
 25.1283 -assumes "c>=0"
 25.1284 -shows "c * (a + b) <= c * a + c * b"
 25.1285 -  using assms by (cases rule: ereal3_cases[of a b c])
 25.1286 -                 (auto simp add: field_simps)
 25.1287 -
 25.1288 -lemma ereal_max_mono:
 25.1289 -  "[| (a::ereal) <= b; c <= d |] ==> max a c <= max b d"
 25.1290 -  by (metis sup_ereal_def sup_mono)
 25.1291 -
 25.1292 -
 25.1293 -lemma ereal_max_least:
 25.1294 -  "[| (a::ereal) <= x; c <= x |] ==> max a c <= x"
 25.1295 -  by (metis sup_ereal_def sup_least)
 25.1296 -
 25.1297  subsubsection {* Tests for code generator *}
 25.1298  
 25.1299  (* A small list of simple arithmetic expressions *)
    26.1 --- a/src/HOL/Library/FrechetDeriv.thy	Tue Feb 12 17:39:45 2013 +0100
    26.2 +++ b/src/HOL/Library/FrechetDeriv.thy	Wed Feb 13 11:46:48 2013 +0100
    26.3 @@ -5,7 +5,7 @@
    26.4  header {* Frechet Derivative *}
    26.5  
    26.6  theory FrechetDeriv
    26.7 -imports Complex_Main
    26.8 +imports "~~/src/HOL/Complex_Main"
    26.9  begin
   26.10  
   26.11  definition
    27.1 --- a/src/HOL/Library/Inner_Product.thy	Tue Feb 12 17:39:45 2013 +0100
    27.2 +++ b/src/HOL/Library/Inner_Product.thy	Wed Feb 13 11:46:48 2013 +0100
    27.3 @@ -117,8 +117,6 @@
    27.4  subclass real_normed_vector
    27.5  proof
    27.6    fix a :: real and x y :: 'a
    27.7 -  show "0 \<le> norm x"
    27.8 -    unfolding norm_eq_sqrt_inner by simp
    27.9    show "norm x = 0 \<longleftrightarrow> x = 0"
   27.10      unfolding norm_eq_sqrt_inner by simp
   27.11    show "norm (x + y) \<le> norm x + norm y"
    28.1 --- a/src/HOL/Library/Product_Vector.thy	Tue Feb 12 17:39:45 2013 +0100
    28.2 +++ b/src/HOL/Library/Product_Vector.thy	Wed Feb 13 11:46:48 2013 +0100
    28.3 @@ -408,8 +408,6 @@
    28.4  
    28.5  instance proof
    28.6    fix r :: real and x y :: "'a \<times> 'b"
    28.7 -  show "0 \<le> norm x"
    28.8 -    unfolding norm_prod_def by simp
    28.9    show "norm x = 0 \<longleftrightarrow> x = 0"
   28.10      unfolding norm_prod_def
   28.11      by (simp add: prod_eq_iff)
    29.1 --- a/src/HOL/Library/Product_plus.thy	Tue Feb 12 17:39:45 2013 +0100
    29.2 +++ b/src/HOL/Library/Product_plus.thy	Wed Feb 13 11:46:48 2013 +0100
    29.3 @@ -5,7 +5,7 @@
    29.4  header {* Additive group operations on product types *}
    29.5  
    29.6  theory Product_plus
    29.7 -imports Main
    29.8 +imports "~~/src/HOL/Main"
    29.9  begin
   29.10  
   29.11  subsection {* Operations *}
    30.1 --- a/src/HOL/Library/RBT_Set.thy	Tue Feb 12 17:39:45 2013 +0100
    30.2 +++ b/src/HOL/Library/RBT_Set.thy	Wed Feb 13 11:46:48 2013 +0100
    30.3 @@ -63,8 +63,6 @@
    30.4  lemma [code, code del]:
    30.5    "Bex = Bex" ..
    30.6  
    30.7 -term can_select
    30.8 -
    30.9  lemma [code, code del]:
   30.10    "can_select = can_select" ..
   30.11  
   30.12 @@ -526,6 +524,8 @@
   30.13  
   30.14  code_datatype Set Coset
   30.15  
   30.16 +declare set.simps[code]
   30.17 +
   30.18  lemma empty_Set [code]:
   30.19    "Set.empty = Set RBT.empty"
   30.20  by (auto simp: Set_def)
    31.1 --- a/src/HOL/Limits.thy	Tue Feb 12 17:39:45 2013 +0100
    31.2 +++ b/src/HOL/Limits.thy	Wed Feb 13 11:46:48 2013 +0100
    31.3 @@ -832,6 +832,35 @@
    31.4    "(f ---> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)"
    31.5    using tendstoI tendstoD by fast
    31.6  
    31.7 +lemma order_tendstoI:
    31.8 +  fixes y :: "_ :: order_topology"
    31.9 +  assumes "\<And>a. a < y \<Longrightarrow> eventually (\<lambda>x. a < f x) F"
   31.10 +  assumes "\<And>a. y < a \<Longrightarrow> eventually (\<lambda>x. f x < a) F"
   31.11 +  shows "(f ---> y) F"
   31.12 +proof (rule topological_tendstoI)
   31.13 +  fix S assume "open S" "y \<in> S"
   31.14 +  then show "eventually (\<lambda>x. f x \<in> S) F"
   31.15 +    unfolding open_generated_order
   31.16 +  proof induct
   31.17 +    case (UN K)
   31.18 +    then obtain k where "y \<in> k" "k \<in> K" by auto
   31.19 +    with UN(2)[of k] show ?case
   31.20 +      by (auto elim: eventually_elim1)
   31.21 +  qed (insert assms, auto elim: eventually_elim2)
   31.22 +qed
   31.23 +
   31.24 +lemma order_tendstoD:
   31.25 +  fixes y :: "_ :: order_topology"
   31.26 +  assumes y: "(f ---> y) F"
   31.27 +  shows "a < y \<Longrightarrow> eventually (\<lambda>x. a < f x) F"
   31.28 +    and "y < a \<Longrightarrow> eventually (\<lambda>x. f x < a) F"
   31.29 +  using topological_tendstoD[OF y, of "{..< a}"] topological_tendstoD[OF y, of "{a <..}"] by auto
   31.30 +
   31.31 +lemma order_tendsto_iff: 
   31.32 +  fixes f :: "_ \<Rightarrow> 'a :: order_topology"
   31.33 +  shows "(f ---> x) F \<longleftrightarrow>(\<forall>l<x. eventually (\<lambda>x. l < f x) F) \<and> (\<forall>u>x. eventually (\<lambda>x. f x < u) F)"
   31.34 +  by (metis order_tendstoI order_tendstoD)
   31.35 +
   31.36  lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\<lambda>x. f x - a) F"
   31.37    by (simp only: tendsto_iff Zfun_def dist_norm)
   31.38  
   31.39 @@ -901,6 +930,20 @@
   31.40      using le_less_trans by (rule eventually_elim2)
   31.41  qed
   31.42  
   31.43 +lemma increasing_tendsto:
   31.44 +  fixes f :: "_ \<Rightarrow> 'a::order_topology"
   31.45 +  assumes bdd: "eventually (\<lambda>n. f n \<le> l) F"
   31.46 +      and en: "\<And>x. x < l \<Longrightarrow> eventually (\<lambda>n. x < f n) F"
   31.47 +  shows "(f ---> l) F"
   31.48 +  using assms by (intro order_tendstoI) (auto elim!: eventually_elim1)
   31.49 +
   31.50 +lemma decreasing_tendsto:
   31.51 +  fixes f :: "_ \<Rightarrow> 'a::order_topology"
   31.52 +  assumes bdd: "eventually (\<lambda>n. l \<le> f n) F"
   31.53 +      and en: "\<And>x. l < x \<Longrightarrow> eventually (\<lambda>n. f n < x) F"
   31.54 +  shows "(f ---> l) F"
   31.55 +  using assms by (intro order_tendstoI) (auto elim!: eventually_elim1)
   31.56 +
   31.57  subsubsection {* Distance and norms *}
   31.58  
   31.59  lemma tendsto_dist [tendsto_intros]:
   31.60 @@ -1002,22 +1045,21 @@
   31.61      by (simp add: tendsto_const)
   31.62  qed
   31.63  
   31.64 -lemma real_tendsto_sandwich:
   31.65 -  fixes f g h :: "'a \<Rightarrow> real"
   31.66 +lemma tendsto_sandwich:
   31.67 +  fixes f g h :: "'a \<Rightarrow> 'b::order_topology"
   31.68    assumes ev: "eventually (\<lambda>n. f n \<le> g n) net" "eventually (\<lambda>n. g n \<le> h n) net"
   31.69    assumes lim: "(f ---> c) net" "(h ---> c) net"
   31.70    shows "(g ---> c) net"
   31.71 -proof -
   31.72 -  have "((\<lambda>n. g n - f n) ---> 0) net"
   31.73 -  proof (rule metric_tendsto_imp_tendsto)
   31.74 -    show "eventually (\<lambda>n. dist (g n - f n) 0 \<le> dist (h n - f n) 0) net"
   31.75 -      using ev by (rule eventually_elim2) (simp add: dist_real_def)
   31.76 -    show "((\<lambda>n. h n - f n) ---> 0) net"
   31.77 -      using tendsto_diff[OF lim(2,1)] by simp
   31.78 -  qed
   31.79 -  from tendsto_add[OF this lim(1)] show ?thesis by simp
   31.80 +proof (rule order_tendstoI)
   31.81 +  fix a show "a < c \<Longrightarrow> eventually (\<lambda>x. a < g x) net"
   31.82 +    using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2)
   31.83 +next
   31.84 +  fix a show "c < a \<Longrightarrow> eventually (\<lambda>x. g x < a) net"
   31.85 +    using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2)
   31.86  qed
   31.87  
   31.88 +lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real]
   31.89 +
   31.90  subsubsection {* Linear operators and multiplication *}
   31.91  
   31.92  lemma (in bounded_linear) tendsto:
   31.93 @@ -1082,29 +1124,30 @@
   31.94      by (simp add: tendsto_const)
   31.95  qed
   31.96  
   31.97 -lemma tendsto_le_const:
   31.98 -  fixes f :: "_ \<Rightarrow> real" 
   31.99 +lemma tendsto_le:
  31.100 +  fixes f g :: "'a \<Rightarrow> 'b::linorder_topology"
  31.101    assumes F: "\<not> trivial_limit F"
  31.102 -  assumes x: "(f ---> x) F" and a: "eventually (\<lambda>x. a \<le> f x) F"
  31.103 -  shows "a \<le> x"
  31.104 +  assumes x: "(f ---> x) F" and y: "(g ---> y) F"
  31.105 +  assumes ev: "eventually (\<lambda>x. g x \<le> f x) F"
  31.106 +  shows "y \<le> x"
  31.107  proof (rule ccontr)
  31.108 -  assume "\<not> a \<le> x"
  31.109 -  with x have "eventually (\<lambda>x. f x < a) F"
  31.110 -    by (auto simp add: tendsto_def elim!: allE[of _ "{..< a}"])
  31.111 -  with a have "eventually (\<lambda>x. False) F"
  31.112 -    by eventually_elim auto
  31.113 +  assume "\<not> y \<le> x"
  31.114 +  with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{..<a} \<inter> {b<..} = {}"
  31.115 +    by (auto simp: not_le)
  31.116 +  then have "eventually (\<lambda>x. f x < a) F" "eventually (\<lambda>x. b < g x) F"
  31.117 +    using x y by (auto intro: order_tendstoD)
  31.118 +  with ev have "eventually (\<lambda>x. False) F"
  31.119 +    by eventually_elim (insert xy, fastforce)
  31.120    with F show False
  31.121      by (simp add: eventually_False)
  31.122  qed
  31.123  
  31.124 -lemma tendsto_le:
  31.125 -  fixes f g :: "_ \<Rightarrow> real" 
  31.126 +lemma tendsto_le_const:
  31.127 +  fixes f :: "'a \<Rightarrow> 'b::linorder_topology"
  31.128    assumes F: "\<not> trivial_limit F"
  31.129 -  assumes x: "(f ---> x) F" and y: "(g ---> y) F"
  31.130 -  assumes ev: "eventually (\<lambda>x. g x \<le> f x) F"
  31.131 -  shows "y \<le> x"
  31.132 -  using tendsto_le_const[OF F tendsto_diff[OF x y], of 0] ev
  31.133 -  by (simp add: sign_simps)
  31.134 +  assumes x: "(f ---> x) F" and a: "eventually (\<lambda>x. a \<le> f x) F"
  31.135 +  shows "a \<le> x"
  31.136 +  using F x tendsto_const a by (rule tendsto_le)
  31.137  
  31.138  subsubsection {* Inverse and division *}
  31.139  
    32.1 --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Tue Feb 12 17:39:45 2013 +0100
    32.2 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Feb 13 11:46:48 2013 +0100
    32.3 @@ -456,13 +456,14 @@
    32.4          NONE => I
    32.5        | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
    32.6      fun failed failure =
    32.7 -      ({outcome = SOME failure, used_facts = [], run_time = Time.zeroTime,
    32.8 +      ({outcome = SOME failure, used_facts = [], used_from = [],
    32.9 +        run_time = Time.zeroTime,
   32.10          preplay = Lazy.value (Sledgehammer_Reconstruct.Failed_to_Play
   32.11            Sledgehammer_Provers.plain_metis),
   32.12          message = K "", message_tail = ""}, ~1)
   32.13 -    val ({outcome, used_facts, run_time, preplay, message, message_tail}
   32.14 +    val ({outcome, used_facts, run_time, preplay, message, message_tail, ...}
   32.15           : Sledgehammer_Provers.prover_result,
   32.16 -        time_isa) = time_limit (Mirabelle.cpu_time (fn () =>
   32.17 +         time_isa) = time_limit (Mirabelle.cpu_time (fn () =>
   32.18        let
   32.19          val _ = if is_appropriate_prop concl_t then ()
   32.20                  else raise Fail "inappropriate"
   32.21 @@ -473,25 +474,20 @@
   32.22            Sledgehammer_Fact.nearly_all_facts ctxt ho_atp
   32.23                Sledgehammer_Fact.no_fact_override reserved css_table chained_ths
   32.24                hyp_ts concl_t
   32.25 +        val factss =
   32.26 +          facts
   32.27            |> filter (is_appropriate_prop o prop_of o snd)
   32.28            |> Sledgehammer_MaSh.relevant_facts ctxt params prover_name
   32.29                   (the_default default_max_facts max_facts)
   32.30                   Sledgehammer_Fact.no_fact_override hyp_ts concl_t
   32.31 -          |> map (apfst (apfst (fn name => name ())))
   32.32 -          |> tap (fn facts =>
   32.33 +          |> tap (fn factss =>
   32.34                       "Line " ^ str0 (Position.line_of pos) ^ ": " ^
   32.35 -                     (if null facts then
   32.36 -                        "Found no relevant facts."
   32.37 -                      else
   32.38 -                        "Including " ^ string_of_int (length facts) ^
   32.39 -                        " relevant fact(s):\n" ^
   32.40 -                        (facts |> map (fst o fst) |> space_implode " ") ^ ".")
   32.41 +                     Sledgehammer_Run.string_of_factss factss
   32.42                       |> Output.urgent_message)
   32.43          val prover = get_prover ctxt prover_name params goal facts
   32.44          val problem =
   32.45            {state = st', goal = goal, subgoal = i,
   32.46 -           subgoal_count = Sledgehammer_Util.subgoal_count st,
   32.47 -           facts = facts |> map Sledgehammer_Provers.Untranslated_Fact}
   32.48 +           subgoal_count = Sledgehammer_Util.subgoal_count st, factss = factss}
   32.49        in prover params (K (K (K ""))) problem end)) ()
   32.50        handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut
   32.51             | Fail "inappropriate" => failed ATP_Proof.Inappropriate
    33.1 --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Tue Feb 12 17:39:45 2013 +0100
    33.2 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Wed Feb 13 11:46:48 2013 +0100
    33.3 @@ -136,7 +136,7 @@
    33.4             |> Sledgehammer_MePo.mepo_suggested_facts ctxt params
    33.5                    default_prover (the_default default_max_facts max_facts)
    33.6                    (SOME relevance_fudge) hyp_ts concl_t
    33.7 -            |> map ((fn name => name ()) o fst o fst)
    33.8 +            |> map (fst o fst)
    33.9           val (found_facts, lost_facts) =
   33.10             flat proofs |> sort_distinct string_ord
   33.11             |> map (fn fact => (find_index (curry (op =) fact) facts, fact))
    34.1 --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Tue Feb 12 17:39:45 2013 +0100
    34.2 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Wed Feb 13 11:46:48 2013 +0100
    34.3 @@ -1,4 +1,3 @@
    34.4 -
    34.5  header {*Instanciates the finite cartesian product of euclidean spaces as a euclidean space.*}
    34.6  
    34.7  theory Cartesian_Euclidean_Space
    34.8 @@ -828,7 +827,7 @@
    34.9  
   34.10  lemma compact_lemma_cart:
   34.11    fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n"
   34.12 -  assumes "bounded s" and "\<forall>n. f n \<in> s"
   34.13 +  assumes f: "bounded (range f)"
   34.14    shows "\<forall>d.
   34.15          \<exists>l r. subseq r \<and>
   34.16          (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
   34.17 @@ -842,16 +841,17 @@
   34.18      thus ?case unfolding subseq_def by auto
   34.19    next
   34.20      case (insert k d)
   34.21 -    have s': "bounded ((\<lambda>x. x $ k) ` s)"
   34.22 -      using `bounded s` by (rule bounded_component_cart)
   34.23      obtain l1::"'a^'n" and r1 where r1:"subseq r1"
   34.24        and lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially"
   34.25        using insert(3) by auto
   34.26 -    have f': "\<forall>n. f (r1 n) $ k \<in> (\<lambda>x. x $ k) ` s"
   34.27 -      using `\<forall>n. f n \<in> s` by simp
   34.28 -    obtain l2 r2 where r2: "subseq r2"
   34.29 +    have s': "bounded ((\<lambda>x. x $ k) ` range f)" using `bounded (range f)`
   34.30 +      by (auto intro!: bounded_component_cart)
   34.31 +    have f': "\<forall>n. f (r1 n) $ k \<in> (\<lambda>x. x $ k) ` range f" by simp
   34.32 +    have "bounded (range (\<lambda>i. f (r1 i) $ k))"
   34.33 +      by (metis (lifting) bounded_subset image_subsetI f' s')
   34.34 +    then obtain l2 r2 where r2: "subseq r2"
   34.35        and lr2: "((\<lambda>i. f (r1 (r2 i)) $ k) ---> l2) sequentially"
   34.36 -      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
   34.37 +      using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) $ k"] by (auto simp: o_def)
   34.38      def r \<equiv> "r1 \<circ> r2"
   34.39      have r: "subseq r"
   34.40        using r1 and r2 unfolding r_def o_def subseq_def by auto
   34.41 @@ -873,11 +873,11 @@
   34.42  
   34.43  instance vec :: (heine_borel, finite) heine_borel
   34.44  proof
   34.45 -  fix s :: "('a ^ 'b) set" and f :: "nat \<Rightarrow> 'a ^ 'b"
   34.46 -  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
   34.47 +  fix f :: "nat \<Rightarrow> 'a ^ 'b"
   34.48 +  assume f: "bounded (range f)"
   34.49    then obtain l r where r: "subseq r"
   34.50        and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially"
   34.51 -    using compact_lemma_cart [OF s f] by blast
   34.52 +    using compact_lemma_cart [OF f] by blast
   34.53    let ?d = "UNIV::'b set"
   34.54    { fix e::real assume "e>0"
   34.55      hence "0 < e / (real_of_nat (card ?d))"
    35.1 --- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Tue Feb 12 17:39:45 2013 +0100
    35.2 +++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Wed Feb 13 11:46:48 2013 +0100
    35.3 @@ -288,16 +288,9 @@
    35.4    assumes lim:"f ----> (l :: ereal)"
    35.5      and ge: "ALL n>=N. f n >= C"
    35.6    shows "l>=C"
    35.7 -proof -
    35.8 -  def g == "(%i. -(f i))"
    35.9 -  { fix n
   35.10 -    assume "n>=N"
   35.11 -    then have "g n <= -C" using assms ereal_minus_le_minus g_def by auto }
   35.12 -  then have "ALL n>=N. g n <= -C" by auto
   35.13 -  moreover have limg: "g ----> (-l)" using g_def ereal_lim_uminus lim by auto
   35.14 -  ultimately have "-l <= -C" using Lim_bounded_ereal[of g "-l" _ "-C"] by auto
   35.15 -  then show ?thesis using ereal_minus_le_minus by auto
   35.16 -qed
   35.17 +  using ge
   35.18 +  by (intro tendsto_le[OF trivial_limit_sequentially lim tendsto_const])
   35.19 +     (auto simp: eventually_sequentially)
   35.20  
   35.21  lemma ereal_open_atLeast: fixes x :: ereal shows "open {x..} \<longleftrightarrow> x = -\<infinity>"
   35.22  proof
   35.23 @@ -326,54 +319,69 @@
   35.24    fixes f :: "'a => ereal"
   35.25    shows "Liminf net f =
   35.26      Sup {l. \<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
   35.27 -  unfolding Liminf_Sup
   35.28 -proof (intro arg_cong[where f="\<lambda>P. Sup (Collect P)"] ext iffI allI impI)
   35.29 -  fix l S
   35.30 -  assume ev: "\<forall>y<l. eventually (\<lambda>x. y < f x) net" and "open S" "mono_set S" "l \<in> S"
   35.31 -  then have "S = UNIV \<or> S = {Inf S <..}"
   35.32 -    using ereal_open_mono_set[of S] by auto
   35.33 -  then show "eventually (\<lambda>x. f x \<in> S) net"
   35.34 -  proof
   35.35 -    assume S: "S = {Inf S<..}"
   35.36 -    then have "Inf S < l" using `l \<in> S` by auto
   35.37 -    then have "eventually (\<lambda>x. Inf S < f x) net" using ev by auto
   35.38 -    then show "eventually (\<lambda>x. f x \<in> S) net" by (subst S) auto
   35.39 -  qed auto
   35.40 +    (is "_ = Sup ?A")
   35.41 +proof (safe intro!: Liminf_eqI complete_lattice_class.Sup_upper complete_lattice_class.Sup_least)
   35.42 +  fix P assume P: "eventually P net"
   35.43 +  fix S assume S: "mono_set S" "INFI (Collect P) f \<in> S"
   35.44 +  { fix x assume "P x"
   35.45 +    then have "INFI (Collect P) f \<le> f x"
   35.46 +      by (intro complete_lattice_class.INF_lower) simp
   35.47 +    with S have "f x \<in> S"
   35.48 +      by (simp add: mono_set) }
   35.49 +  with P show "eventually (\<lambda>x. f x \<in> S) net"
   35.50 +    by (auto elim: eventually_elim1)
   35.51  next
   35.52 -  fix l y
   35.53 -  assume S: "\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually  (\<lambda>x. f x \<in> S) net" "y < l"
   35.54 -  have "eventually  (\<lambda>x. f x \<in> {y <..}) net"
   35.55 -    using `y < l` by (intro S[rule_format]) auto
   35.56 -  then show "eventually (\<lambda>x. y < f x) net" by auto
   35.57 +  fix y l
   35.58 +  assume S: "\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually  (\<lambda>x. f x \<in> S) net"
   35.59 +  assume P: "\<forall>P. eventually P net \<longrightarrow> INFI (Collect P) f \<le> y"
   35.60 +  show "l \<le> y"
   35.61 +  proof (rule ereal_le_ereal)
   35.62 +    fix B assume "B < l"
   35.63 +    then have "eventually (\<lambda>x. f x \<in> {B <..}) net"
   35.64 +      by (intro S[rule_format]) auto
   35.65 +    then have "INFI {x. B < f x} f \<le> y"
   35.66 +      using P by auto
   35.67 +    moreover have "B \<le> INFI {x. B < f x} f"
   35.68 +      by (intro INF_greatest) auto
   35.69 +    ultimately show "B \<le> y"
   35.70 +      by simp
   35.71 +  qed
   35.72  qed
   35.73  
   35.74  lemma ereal_Limsup_Inf_monoset:
   35.75    fixes f :: "'a => ereal"
   35.76    shows "Limsup net f =
   35.77      Inf {l. \<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
   35.78 -  unfolding Limsup_Inf
   35.79 -proof (intro arg_cong[where f="\<lambda>P. Inf (Collect P)"] ext iffI allI impI)
   35.80 -  fix l S
   35.81 -  assume ev: "\<forall>y>l. eventually (\<lambda>x. f x < y) net" and "open S" "mono_set (uminus`S)" "l \<in> S"
   35.82 -  then have "open (uminus`S) \<and> mono_set (uminus`S)" by (simp add: ereal_open_uminus)
   35.83 -  then have "S = UNIV \<or> S = {..< Sup S}"
   35.84 -    unfolding ereal_open_mono_set ereal_Inf_uminus_image_eq ereal_image_uminus_shift by simp
   35.85 -  then show "eventually (\<lambda>x. f x \<in> S) net"
   35.86 -  proof
   35.87 -    assume S: "S = {..< Sup S}"
   35.88 -    then have "l < Sup S" using `l \<in> S` by auto
   35.89 -    then have "eventually (\<lambda>x. f x < Sup S) net" using ev by auto
   35.90 -    then show "eventually (\<lambda>x. f x \<in> S) net"  by (subst S) auto
   35.91 -  qed auto
   35.92 +    (is "_ = Inf ?A")
   35.93 +proof (safe intro!: Limsup_eqI complete_lattice_class.Inf_lower complete_lattice_class.Inf_greatest)
   35.94 +  fix P assume P: "eventually P net"
   35.95 +  fix S assume S: "mono_set (uminus`S)" "SUPR (Collect P) f \<in> S"
   35.96 +  { fix x assume "P x"
   35.97 +    then have "f x \<le> SUPR (Collect P) f"
   35.98 +      by (intro complete_lattice_class.SUP_upper) simp
   35.99 +    with S(1)[unfolded mono_set, rule_format, of "- SUPR (Collect P) f" "- f x"] S(2)
  35.100 +    have "f x \<in> S"
  35.101 +      by (simp add: inj_image_mem_iff) }
  35.102 +  with P show "eventually (\<lambda>x. f x \<in> S) net"
  35.103 +    by (auto elim: eventually_elim1)
  35.104  next
  35.105 -  fix l y
  35.106 -  assume S: "\<forall>S. open S \<longrightarrow> mono_set (uminus`S) \<longrightarrow> l \<in> S \<longrightarrow> eventually  (\<lambda>x. f x \<in> S) net" "l < y"
  35.107 -  have "eventually  (\<lambda>x. f x \<in> {..< y}) net"
  35.108 -    using `l < y` by (intro S[rule_format]) auto
  35.109 -  then show "eventually (\<lambda>x. f x < y) net" by auto
  35.110 +  fix y l
  35.111 +  assume S: "\<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually  (\<lambda>x. f x \<in> S) net"
  35.112 +  assume P: "\<forall>P. eventually P net \<longrightarrow> y \<le> SUPR (Collect P) f"
  35.113 +  show "y \<le> l"
  35.114 +  proof (rule ereal_ge_ereal, safe)
  35.115 +    fix B assume "l < B"
  35.116 +    then have "eventually (\<lambda>x. f x \<in> {..< B}) net"
  35.117 +      by (intro S[rule_format]) auto
  35.118 +    then have "y \<le> SUPR {x. f x < B} f"
  35.119 +      using P by auto
  35.120 +    moreover have "SUPR {x. f x < B} f \<le> B"
  35.121 +      by (intro SUP_least) auto
  35.122 +    ultimately show "y \<le> B"
  35.123 +      by simp
  35.124 +  qed
  35.125  qed
  35.126  
  35.127 -
  35.128  lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::ereal set)"
  35.129    using ereal_open_uminus[of S] ereal_open_uminus[of "uminus`S"] by auto
  35.130  
  35.131 @@ -434,21 +442,16 @@
  35.132    shows "(f ---> \<infinity>) net \<longleftrightarrow> Liminf net f = \<infinity>"
  35.133  proof (intro lim_imp_Liminf iffI assms)
  35.134    assume rhs: "Liminf net f = \<infinity>"
  35.135 -  { fix S :: "ereal set"
  35.136 -    assume "open S & \<infinity> : S"
  35.137 -    then obtain m where "{ereal m<..} <= S" using open_PInfty2 by auto
  35.138 -    moreover
  35.139 -    have "eventually (\<lambda>x. f x \<in> {ereal m<..}) net"
  35.140 -      using rhs
  35.141 -      unfolding Liminf_Sup top_ereal_def[symmetric] Sup_eq_top_iff
  35.142 -      by (auto elim!: allE[where x="ereal m"] simp: top_ereal_def)
  35.143 -    ultimately
  35.144 -    have "eventually (%x. f x : S) net"
  35.145 -      apply (subst eventually_mono)
  35.146 -      apply auto
  35.147 -      done
  35.148 -  }
  35.149 -  then show "(f ---> \<infinity>) net" unfolding tendsto_def by auto
  35.150 +  show "(f ---> \<infinity>) net"
  35.151 +    unfolding tendsto_PInfty
  35.152 +  proof
  35.153 +    fix r :: real
  35.154 +    have "ereal r < top" unfolding top_ereal_def by simp
  35.155 +    with rhs obtain P where "eventually P net" "r < INFI (Collect P) f"
  35.156 +      unfolding Liminf_def SUP_eq_top_iff top_ereal_def[symmetric] by auto
  35.157 +    then show "eventually (\<lambda>x. ereal r < f x) net"
  35.158 +      by (auto elim!: eventually_elim1 dest: less_INF_D)
  35.159 +  qed
  35.160  qed
  35.161  
  35.162  lemma Limsup_MInfty:
  35.163 @@ -474,12 +477,12 @@
  35.164    show "(f ---> f0) net"
  35.165    proof (rule topological_tendstoI)
  35.166      fix S
  35.167 -    assume "open S""f0 \<in> S"
  35.168 +    assume "open S" "f0 \<in> S"
  35.169      then obtain a b where "a < Liminf net f" "Limsup net f < b" "{a<..<b} \<subseteq> S"
  35.170        using ereal_open_cont_interval2[of S f0] real lim by auto
  35.171      then have "eventually (\<lambda>x. f x \<in> {a<..<b}) net"
  35.172 -      unfolding Liminf_Sup Limsup_Inf less_Sup_iff Inf_less_iff
  35.173 -      by (auto intro!: eventually_conj)
  35.174 +      unfolding Liminf_def Limsup_def less_SUP_iff INF_less_iff
  35.175 +      by (auto intro!: eventually_conj elim: eventually_elim1 dest: less_INF_D SUP_lessD)
  35.176      with `{a<..<b} \<subseteq> S` show "eventually (%x. f x : S) net"
  35.177        by (rule_tac eventually_mono) auto
  35.178    qed
  35.179 @@ -518,7 +521,7 @@
  35.180    assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
  35.181      and "X ----> x" "Y ----> y"
  35.182    shows "x <= y"
  35.183 -  by (metis ereal_Liminf_eq_Limsup_iff[OF trivial_limit_sequentially] assms liminf_mono)
  35.184 +  using assms(1) by (intro LIMSEQ_le[OF assms(2,3)]) auto
  35.185  
  35.186  lemma incseq_le_ereal:
  35.187    fixes X :: "nat \<Rightarrow> ereal"
  35.188 @@ -577,114 +580,24 @@
  35.189    by (metis abs_less_iff assms leI le_max_iff_disj
  35.190      less_eq_real_def less_le_not_le less_minus_iff minus_minus)
  35.191  
  35.192 -lemma bounded_increasing_convergent2:
  35.193 -  fixes f::"nat => real"
  35.194 -  assumes "ALL n. f n <= B" "ALL n m. n>=m --> f n >= f m"
  35.195 -  shows "EX l. (f ---> l) sequentially"
  35.196 -proof -
  35.197 -  def N == "max (abs (f 0)) (abs B)"
  35.198 -  { fix n
  35.199 -    have "abs (f n) <= N"
  35.200 -      unfolding N_def
  35.201 -      apply (subst bounded_abs)
  35.202 -      using assms apply auto
  35.203 -      done }
  35.204 -  then have "bounded {f n| n::nat. True}"
  35.205 -    unfolding bounded_real by auto
  35.206 -  then show ?thesis
  35.207 -    apply (rule Topology_Euclidean_Space.bounded_increasing_convergent)
  35.208 -    using assms apply auto
  35.209 -    done
  35.210 -qed
  35.211 -
  35.212  lemma lim_ereal_increasing:
  35.213 -  assumes "\<And>n m. n >= m \<Longrightarrow> f n >= f m"
  35.214 -  obtains l where "f ----> (l::ereal)"
  35.215 -proof (cases "f = (\<lambda>x. - \<infinity>)")
  35.216 -  case True
  35.217 -  then show thesis
  35.218 -    using tendsto_const[of "- \<infinity>" sequentially] by (intro that[of "-\<infinity>"]) auto
  35.219 -next
  35.220 -  case False
  35.221 -  then obtain N where N_def: "f N > (-\<infinity>)" by (auto simp: fun_eq_iff)
  35.222 -  have "ALL n>=N. f n >= f N" using assms by auto
  35.223 -  then have minf: "ALL n>=N. f n > (-\<infinity>)" using N_def by auto
  35.224 -  def Y == "(%n. (if n>=N then f n else f N))"
  35.225 -  then have incy: "!!n m. n>=m ==> Y n >= Y m" using assms by auto
  35.226 -  from minf have minfy: "ALL n. Y n ~= (-\<infinity>)" using Y_def by auto
  35.227 -  show thesis
  35.228 -  proof (cases "EX B. ALL n. f n < ereal B")
  35.229 -    case False
  35.230 -    then show thesis
  35.231 -      apply -
  35.232 -      apply (rule that[of \<infinity>])
  35.233 -      unfolding Lim_PInfty not_ex not_all
  35.234 -      apply safe
  35.235 -      apply (erule_tac x=B in allE, safe)
  35.236 -      apply (rule_tac x=x in exI, safe)
  35.237 -      apply (rule order_trans[OF _ assms[rule_format]])
  35.238 -      apply auto
  35.239 -      done
  35.240 -  next
  35.241 -    case True
  35.242 -    then guess B ..
  35.243 -    then have "ALL n. Y n < ereal B" using Y_def by auto
  35.244 -    note B = this[rule_format]
  35.245 -    { fix n
  35.246 -      have "Y n < \<infinity>"
  35.247 -        using B[of n]
  35.248 -        apply (subst less_le_trans)
  35.249 -        apply auto
  35.250 -        done
  35.251 -      then have "Y n ~= \<infinity> & Y n ~= (-\<infinity>)" using minfy by auto
  35.252 -    }
  35.253 -    then have *: "ALL n. \<bar>Y n\<bar> \<noteq> \<infinity>" by auto
  35.254 -    { fix n
  35.255 -      have "real (Y n) < B"
  35.256 -      proof -
  35.257 -        case goal1
  35.258 -        then show ?case
  35.259 -          using B[of n] apply-apply(subst(asm) ereal_real'[THEN sym]) defer defer
  35.260 -          unfolding ereal_less using * by auto
  35.261 -      qed
  35.262 -    }
  35.263 -    then have B': "ALL n. (real (Y n) <= B)" using less_imp_le by auto
  35.264 -    have "EX l. (%n. real (Y n)) ----> l"
  35.265 -      apply (rule bounded_increasing_convergent2)
  35.266 -    proof safe
  35.267 -      show "\<And>n. real (Y n) <= B" using B' by auto
  35.268 -      fix n m :: nat
  35.269 -      assume "n<=m"
  35.270 -      then have "ereal (real (Y n)) <= ereal (real (Y m))"
  35.271 -        using incy[rule_format,of n m] apply(subst ereal_real)+
  35.272 -        using *[rule_format, of n] *[rule_format, of m] by auto
  35.273 -      then show "real (Y n) <= real (Y m)" by auto
  35.274 -    qed
  35.275 -    then guess l .. note l=this
  35.276 -    have "Y ----> ereal l"
  35.277 -      using l
  35.278 -      apply -
  35.279 -      apply (subst(asm) lim_ereal[THEN sym])
  35.280 -      unfolding ereal_real
  35.281 -      using * apply auto
  35.282 -      done
  35.283 -    then show thesis
  35.284 -      apply -
  35.285 -      apply (rule that[of "ereal l"])
  35.286 -      apply (subst tail_same_limit[of Y _ N])
  35.287 -      using Y_def apply auto
  35.288 -      done
  35.289 -  qed
  35.290 +  assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<ge> f m"
  35.291 +  obtains l where "f ----> (l::'a::{complete_linorder, linorder_topology})"
  35.292 +proof
  35.293 +  show "f ----> (SUP n. f n)"
  35.294 +    using assms
  35.295 +    by (intro increasing_tendsto)
  35.296 +       (auto simp: SUP_upper eventually_sequentially less_SUP_iff intro: less_le_trans)
  35.297  qed
  35.298  
  35.299  lemma lim_ereal_decreasing:
  35.300 -  assumes "\<And>n m. n >= m \<Longrightarrow> f n <= f m"
  35.301 -  obtains l where "f ----> (l::ereal)"
  35.302 -proof -
  35.303 -  from lim_ereal_increasing[of "\<lambda>x. - f x"] assms
  35.304 -  obtain l where "(\<lambda>x. - f x) ----> l" by auto
  35.305 -  from ereal_lim_mult[OF this, of "- 1"] show thesis
  35.306 -    by (intro that[of "-l"]) (simp add: ereal_uminus_eq_reorder)
  35.307 +  assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<le> f m"
  35.308 +  obtains l where "f ----> (l::'a::{complete_linorder, linorder_topology})"
  35.309 +proof
  35.310 +  show "f ----> (INF n. f n)"
  35.311 +    using assms
  35.312 +    by (intro decreasing_tendsto)
  35.313 +       (auto simp: INF_lower eventually_sequentially INF_less_iff intro: le_less_trans)
  35.314  qed
  35.315  
  35.316  lemma compact_ereal:
  35.317 @@ -711,37 +624,17 @@
  35.318    by (metis Lim_bounded2_ereal assms complete_lattice_class.Inf_lower)
  35.319  
  35.320  lemma SUP_Lim_ereal:
  35.321 -  fixes X :: "nat \<Rightarrow> ereal"
  35.322 -  assumes "incseq X" "X ----> l"
  35.323 +  fixes X :: "nat \<Rightarrow> 'a::{complete_linorder, linorder_topology}"
  35.324 +  assumes inc: "incseq X" and l: "X ----> l"
  35.325    shows "(SUP n. X n) = l"
  35.326 -proof (rule ereal_SUPI)
  35.327 -  fix n from assms show "X n \<le> l"
  35.328 -    by (intro incseq_le_ereal) (simp add: incseq_def)
  35.329 -next
  35.330 -  fix y assume "\<And>n. n \<in> UNIV \<Longrightarrow> X n \<le> y"
  35.331 -  with ereal_Sup_lim[OF _ `X ----> l`, of "{..y}"] show "l \<le> y" by auto
  35.332 -qed
  35.333 -
  35.334 -lemma LIMSEQ_ereal_SUPR:
  35.335 -  fixes X :: "nat \<Rightarrow> ereal"
  35.336 -  assumes "incseq X"
  35.337 -  shows "X ----> (SUP n. X n)"
  35.338 -proof (rule lim_ereal_increasing)
  35.339 -  fix n m :: nat
  35.340 -  assume "m \<le> n"
  35.341 -  then show "X m \<le> X n" using `incseq X` by (simp add: incseq_def)
  35.342 -next
  35.343 -  fix l
  35.344 -  assume "X ----> l"
  35.345 -  with SUP_Lim_ereal[of X, OF assms this] show ?thesis by simp
  35.346 -qed
  35.347 +  using LIMSEQ_SUP[OF inc] tendsto_unique[OF trivial_limit_sequentially l] by simp
  35.348  
  35.349  lemma INF_Lim_ereal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::ereal)"
  35.350    using SUP_Lim_ereal[of "\<lambda>i. - X i" "- l"]
  35.351    by (simp add: ereal_SUPR_uminus ereal_lim_uminus)
  35.352  
  35.353  lemma LIMSEQ_ereal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: ereal)"
  35.354 -  using LIMSEQ_ereal_SUPR[of "\<lambda>i. - X i"]
  35.355 +  using LIMSEQ_SUP[of "\<lambda>i. - X i"]
  35.356    by (simp add: ereal_SUPR_uminus ereal_lim_uminus)
  35.357  
  35.358  lemma SUP_eq_LIMSEQ:
  35.359 @@ -755,165 +648,66 @@
  35.360      from SUP_Lim_ereal[OF inc this]
  35.361      show "(SUP n. ereal (f n)) = ereal x" . }
  35.362    { assume "(SUP n. ereal (f n)) = ereal x"
  35.363 -    with LIMSEQ_ereal_SUPR[OF inc]
  35.364 +    with LIMSEQ_SUP[OF inc]
  35.365      show "f ----> x" by auto }
  35.366  qed
  35.367  
  35.368  lemma Liminf_within:
  35.369 -  fixes f :: "'a::metric_space => ereal"
  35.370 -  shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
  35.371 -proof -
  35.372 -  let ?l="(SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
  35.373 -  { fix T
  35.374 -    assume T_def: "open T & mono_set T & ?l:T"
  35.375 -    have "EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T"
  35.376 -    proof -
  35.377 -      { assume "T=UNIV" then have ?thesis by (simp add: gt_ex) }
  35.378 -      moreover
  35.379 -      { assume "~(T=UNIV)"
  35.380 -        then obtain B where "T={B<..}" using T_def ereal_open_mono_set[of T] by auto
  35.381 -        then have "B<?l" using T_def by auto
  35.382 -        then obtain d where d_def: "0<d & B<(INF y:(S Int ball x d - {x}). f y)"
  35.383 -          unfolding less_SUP_iff by auto
  35.384 -        { fix y assume "y:S & 0 < dist y x & dist y x < d"
  35.385 -          then have "y:(S Int ball x d - {x})" unfolding ball_def by (auto simp add: dist_commute)
  35.386 -          then have "f y:T" using d_def INF_lower[of y "S Int ball x d - {x}" f] `T={B<..}` by auto
  35.387 -        } then have ?thesis apply(rule_tac x="d" in exI) using d_def by auto
  35.388 -      }
  35.389 -      ultimately show ?thesis by auto
  35.390 -    qed
  35.391 -  }
  35.392 -  moreover
  35.393 -  { fix z
  35.394 -    assume a: "ALL T. open T --> mono_set T --> z : T -->
  35.395 -       (EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T)"
  35.396 -    { fix B
  35.397 -      assume "B<z"
  35.398 -      then obtain d where d_def: "d>0 & (ALL y:S. 0 < dist y x & dist y x < d --> B < f y)"
  35.399 -         using a[rule_format, of "{B<..}"] mono_greaterThan by auto
  35.400 -      { fix y
  35.401 -        assume "y:(S Int ball x d - {x})"
  35.402 -        then have "y:S & 0 < dist y x & dist y x < d"
  35.403 -          unfolding ball_def
  35.404 -          apply (simp add: dist_commute)
  35.405 -          apply (metis dist_eq_0_iff less_le zero_le_dist)
  35.406 -          done
  35.407 -        then have "B <= f y" using d_def by auto
  35.408 -      }
  35.409 -      then have "B <= INFI (S Int ball x d - {x}) f"
  35.410 -        apply (subst INF_greatest)
  35.411 -        apply auto
  35.412 -        done
  35.413 -      also have "...<=?l"
  35.414 -        apply (subst SUP_upper)
  35.415 -        using d_def apply auto
  35.416 -        done
  35.417 -      finally have "B<=?l" by auto
  35.418 -    }
  35.419 -    then have "z <= ?l" using ereal_le_ereal[of z "?l"] by auto
  35.420 -  }
  35.421 -  ultimately show ?thesis
  35.422 -    unfolding ereal_Liminf_Sup_monoset eventually_within
  35.423 -    apply (subst ereal_SupI[of _ "(SUP e:{0<..}. INFI (S Int ball x e - {x}) f)"])
  35.424 -    apply auto
  35.425 -    done
  35.426 +  fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
  35.427 +  shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S \<inter> ball x e - {x}). f y)"
  35.428 +  unfolding Liminf_def eventually_within
  35.429 +proof (rule SUPR_eq, simp_all add: Ball_def Bex_def, safe)
  35.430 +  fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y"
  35.431 +  then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
  35.432 +    by (auto simp: zero_less_dist_iff dist_commute)
  35.433 +  then show "\<exists>r>0. INFI (Collect P) f \<le> INFI (S \<inter> ball x r - {x}) f"
  35.434 +    by (intro exI[of _ d] INF_mono conjI `0 < d`) auto
  35.435 +next
  35.436 +  fix d :: real assume "0 < d"
  35.437 +  then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> 0 < dist xa x \<and> dist xa x < d \<longrightarrow> P xa) \<and>
  35.438 +    INFI (S \<inter> ball x d - {x}) f \<le> INFI (Collect P) f"
  35.439 +    by (intro exI[of _ "\<lambda>y. y \<in> S \<inter> ball x d - {x}"])
  35.440 +       (auto intro!: INF_mono exI[of _ d] simp: dist_commute)
  35.441  qed
  35.442  
  35.443  lemma Limsup_within:
  35.444 -  fixes f :: "'a::metric_space => ereal"
  35.445 -  shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
  35.446 -proof -
  35.447 -  let ?l = "(INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
  35.448 -  { fix T
  35.449 -    assume T_def: "open T & mono_set (uminus ` T) & ?l:T"
  35.450 -    have "EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T"
  35.451 -    proof -
  35.452 -      { assume "T = UNIV"
  35.453 -        then have ?thesis by (simp add: gt_ex) }
  35.454 -      moreover
  35.455 -      { assume "T \<noteq> UNIV"
  35.456 -        then have "~(uminus ` T = UNIV)"
  35.457 -          by (metis Int_UNIV_right Int_absorb1 image_mono ereal_minus_minus_image subset_UNIV)
  35.458 -        then have "uminus ` T = {Inf (uminus ` T)<..}"
  35.459 -          using T_def ereal_open_mono_set[of "uminus ` T"] ereal_open_uminus[of T] by auto
  35.460 -        then obtain B where "T={..<B}"
  35.461 -          unfolding ereal_Inf_uminus_image_eq ereal_uminus_lessThan[symmetric]
  35.462 -          unfolding inj_image_eq_iff[OF ereal_inj_on_uminus] by simp
  35.463 -        then have "?l<B" using T_def by auto
  35.464 -        then obtain d where d_def: "0<d & (SUP y:(S Int ball x d - {x}). f y)<B"
  35.465 -          unfolding INF_less_iff by auto
  35.466 -        { fix y
  35.467 -          assume "y:S & 0 < dist y x & dist y x < d"
  35.468 -          then have "y:(S Int ball x d - {x})"
  35.469 -            unfolding ball_def by (auto simp add: dist_commute)
  35.470 -          then have "f y:T"
  35.471 -            using d_def SUP_upper[of y "S Int ball x d - {x}" f] `T={..<B}` by auto
  35.472 -        }
  35.473 -        then have ?thesis
  35.474 -          apply (rule_tac x="d" in exI)
  35.475 -          using d_def apply auto
  35.476 -          done
  35.477 -      }
  35.478 -      ultimately show ?thesis by auto
  35.479 -    qed
  35.480 -  }
  35.481 -  moreover
  35.482 -  { fix z
  35.483 -    assume a: "ALL T. open T --> mono_set (uminus ` T) --> z : T -->
  35.484 -       (EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T)"
  35.485 -    { fix B
  35.486 -      assume "z<B"
  35.487 -      then obtain d where d_def: "d>0 & (ALL y:S. 0 < dist y x & dist y x < d --> f y<B)"
  35.488 -         using a[rule_format, of "{..<B}"] by auto
  35.489 -      { fix y
  35.490 -        assume "y:(S Int ball x d - {x})"
  35.491 -        then have "y:S & 0 < dist y x & dist y x < d"
  35.492 -          unfolding ball_def
  35.493 -          apply (simp add: dist_commute)
  35.494 -          apply (metis dist_eq_0_iff less_le zero_le_dist)
  35.495 -          done
  35.496 -        then have "f y <= B" using d_def by auto
  35.497 -      }
  35.498 -      then have "SUPR (S Int ball x d - {x}) f <= B"
  35.499 -        apply (subst SUP_least)
  35.500 -        apply auto
  35.501 -        done
  35.502 -      moreover
  35.503 -      have "?l<=SUPR (S Int ball x d - {x}) f"
  35.504 -        apply (subst INF_lower)
  35.505 -        using d_def apply auto
  35.506 -        done
  35.507 -      ultimately have "?l<=B" by auto
  35.508 -    } then have "?l <= z" using ereal_ge_ereal[of z "?l"] by auto
  35.509 -  }
  35.510 -  ultimately show ?thesis
  35.511 -    unfolding ereal_Limsup_Inf_monoset eventually_within
  35.512 -    apply (subst ereal_InfI)
  35.513 -    apply auto
  35.514 -    done
  35.515 +  fixes f :: "'a::metric_space => 'b::complete_lattice"
  35.516 +  shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S \<inter> ball x e - {x}). f y)"
  35.517 +  unfolding Limsup_def eventually_within
  35.518 +proof (rule INFI_eq, simp_all add: Ball_def Bex_def, safe)
  35.519 +  fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y"
  35.520 +  then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
  35.521 +    by (auto simp: zero_less_dist_iff dist_commute)
  35.522 +  then show "\<exists>r>0. SUPR (S \<inter> ball x r - {x}) f \<le> SUPR (Collect P) f"
  35.523 +    by (intro exI[of _ d] SUP_mono conjI `0 < d`) auto
  35.524 +next
  35.525 +  fix d :: real assume "0 < d"
  35.526 +  then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> 0 < dist xa x \<and> dist xa x < d \<longrightarrow> P xa) \<and>
  35.527 +    SUPR (Collect P) f \<le> SUPR (S \<inter> ball x d - {x}) f"
  35.528 +    by (intro exI[of _ "\<lambda>y. y \<in> S \<inter> ball x d - {x}"])
  35.529 +       (auto intro!: SUP_mono exI[of _ d] simp: dist_commute)
  35.530  qed
  35.531  
  35.532 -
  35.533  lemma Liminf_within_UNIV:
  35.534 -  fixes f :: "'a::metric_space => ereal"
  35.535 +  fixes f :: "'a::metric_space => _"
  35.536    shows "Liminf (at x) f = Liminf (at x within UNIV) f"
  35.537    by simp (* TODO: delete *)
  35.538  
  35.539  
  35.540  lemma Liminf_at:
  35.541 -  fixes f :: "'a::metric_space => ereal"
  35.542 +  fixes f :: "'a::metric_space => _"
  35.543    shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
  35.544    using Liminf_within[of x UNIV f] by simp
  35.545  
  35.546  
  35.547  lemma Limsup_within_UNIV:
  35.548 -  fixes f :: "'a::metric_space => ereal"
  35.549 +  fixes f :: "'a::metric_space => _"
  35.550    shows "Limsup (at x) f = Limsup (at x within UNIV) f"
  35.551    by simp (* TODO: delete *)
  35.552  
  35.553  
  35.554  lemma Limsup_at:
  35.555 -  fixes f :: "'a::metric_space => ereal"
  35.556 +  fixes f :: "'a::metric_space => _"
  35.557    shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
  35.558    using Limsup_within[of x UNIV f] by simp
  35.559  
  35.560 @@ -1295,7 +1089,7 @@
  35.561  proof -
  35.562    have "incseq (\<lambda>i. \<Sum>j=0..<i. f j)"
  35.563      using ereal_add_mono[OF _ assms] by (auto intro!: incseq_SucI)
  35.564 -  from LIMSEQ_ereal_SUPR[OF this]
  35.565 +  from LIMSEQ_SUP[OF this]
  35.566    show ?thesis unfolding sums_def by (simp add: atLeast0LessThan)
  35.567  qed
  35.568  
    36.1 --- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Tue Feb 12 17:39:45 2013 +0100
    36.2 +++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Wed Feb 13 11:46:48 2013 +0100
    36.3 @@ -387,9 +387,6 @@
    36.4  
    36.5  instance proof
    36.6    fix a :: real and x y :: "'a ^ 'b"
    36.7 -  show "0 \<le> norm x"
    36.8 -    unfolding norm_vec_def
    36.9 -    by (rule setL2_nonneg)
   36.10    show "norm x = 0 \<longleftrightarrow> x = 0"
   36.11      unfolding norm_vec_def
   36.12      by (simp add: setL2_eq_0_iff vec_eq_iff)
    37.1 --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Tue Feb 12 17:39:45 2013 +0100
    37.2 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Wed Feb 13 11:46:48 2013 +0100
    37.3 @@ -47,6 +47,17 @@
    37.4  definition "topological_basis B =
    37.5    ((\<forall>b\<in>B. open b) \<and> (\<forall>x. open x \<longrightarrow> (\<exists>B'. B' \<subseteq> B \<and> Union B' = x)))"
    37.6  
    37.7 +lemma "topological_basis B = (\<forall>x. open x \<longleftrightarrow> (\<exists>B'. B' \<subseteq> B \<and> Union B' = x))"
    37.8 +  unfolding topological_basis_def
    37.9 +  apply safe
   37.10 +     apply fastforce
   37.11 +    apply fastforce
   37.12 +   apply (erule_tac x="x" in allE)
   37.13 +   apply simp
   37.14 +   apply (rule_tac x="{x}" in exI)
   37.15 +  apply auto
   37.16 +  done
   37.17 +
   37.18  lemma topological_basis_iff:
   37.19    assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
   37.20    shows "topological_basis B \<longleftrightarrow> (\<forall>O'. open O' \<longrightarrow> (\<forall>x\<in>O'. \<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'))"
   37.21 @@ -2143,6 +2154,9 @@
   37.22    bounded :: "'a set \<Rightarrow> bool" where
   37.23    "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"
   37.24  
   37.25 +lemma bounded_subset_cball: "bounded S \<longleftrightarrow> (\<exists>e x. S \<subseteq> cball x e)"
   37.26 +  unfolding bounded_def subset_eq by auto
   37.27 +
   37.28  lemma bounded_any_center: "bounded S \<longleftrightarrow> (\<exists>e. \<forall>y\<in>S. dist a y \<le> e)"
   37.29  unfolding bounded_def
   37.30  apply safe
   37.31 @@ -3202,16 +3216,16 @@
   37.32  
   37.33  class heine_borel = metric_space +
   37.34    assumes bounded_imp_convergent_subsequence:
   37.35 -    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
   37.36 -      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
   37.37 +    "bounded (range f) \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
   37.38  
   37.39  lemma bounded_closed_imp_seq_compact:
   37.40    fixes s::"'a::heine_borel set"
   37.41    assumes "bounded s" and "closed s" shows "seq_compact s"
   37.42  proof (unfold seq_compact_def, clarify)
   37.43    fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
   37.44 +  with `bounded s` have "bounded (range f)" by (auto intro: bounded_subset)
   37.45    obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
   37.46 -    using bounded_imp_convergent_subsequence [OF `bounded s` `\<forall>n. f n \<in> s`] by auto
   37.47 +    using bounded_imp_convergent_subsequence [OF `bounded (range f)`] by auto
   37.48    from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
   37.49    have "l \<in> s" using `closed s` fr l
   37.50      unfolding closed_sequential_limits by blast
   37.51 @@ -3239,20 +3253,20 @@
   37.52  
   37.53  instance real :: heine_borel
   37.54  proof
   37.55 -  fix s :: "real set" and f :: "nat \<Rightarrow> real"
   37.56 -  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
   37.57 +  fix f :: "nat \<Rightarrow> real"
   37.58 +  assume f: "bounded (range f)"
   37.59    obtain r where r: "subseq r" "monoseq (f \<circ> r)"
   37.60      unfolding comp_def by (metis seq_monosub)
   37.61    moreover
   37.62    then have "Bseq (f \<circ> r)"
   37.63 -    unfolding Bseq_eq_bounded using s f by (auto intro: bounded_subset)
   37.64 +    unfolding Bseq_eq_bounded using f by (auto intro: bounded_subset)
   37.65    ultimately show "\<exists>l r. subseq r \<and> (f \<circ> r) ----> l"
   37.66      using Bseq_monoseq_convergent[of "f \<circ> r"] by (auto simp: convergent_def)
   37.67  qed
   37.68  
   37.69  lemma compact_lemma:
   37.70    fixes f :: "nat \<Rightarrow> 'a::euclidean_space"
   37.71 -  assumes "bounded s" and "\<forall>n. f n \<in> s"
   37.72 +  assumes "bounded (range f)"
   37.73    shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r. subseq r \<and>
   37.74          (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)"
   37.75  proof safe
   37.76 @@ -3262,14 +3276,16 @@
   37.77        (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)"
   37.78    proof(induct d) case empty thus ?case unfolding subseq_def by auto
   37.79    next case (insert k d) have k[intro]:"k\<in>Basis" using insert by auto
   37.80 -    have s': "bounded ((\<lambda>x. x \<bullet> k) ` s)" using `bounded s`
   37.81 +    have s': "bounded ((\<lambda>x. x \<bullet> k) ` range f)" using `bounded (range f)`
   37.82        by (auto intro!: bounded_linear_image bounded_linear_inner_left)
   37.83      obtain l1::"'a" and r1 where r1:"subseq r1" and
   37.84        lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
   37.85        using insert(3) using insert(4) by auto
   37.86 -    have f': "\<forall>n. f (r1 n) \<bullet> k \<in> (\<lambda>x. x \<bullet> k) ` s" using `\<forall>n. f n \<in> s` by simp
   37.87 -    obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) \<bullet> k) ---> l2) sequentially"
   37.88 -      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
   37.89 +    have f': "\<forall>n. f (r1 n) \<bullet> k \<in> (\<lambda>x. x \<bullet> k) ` range f" by simp
   37.90 +    have "bounded (range (\<lambda>i. f (r1 i) \<bullet> k))"
   37.91 +      by (metis (lifting) bounded_subset f' image_subsetI s')
   37.92 +    then obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) \<bullet> k) ---> l2) sequentially"
   37.93 +      using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) \<bullet> k"] by (auto simp: o_def)
   37.94      def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
   37.95        using r1 and r2 unfolding r_def o_def subseq_def by auto
   37.96      moreover
   37.97 @@ -3289,11 +3305,11 @@
   37.98  
   37.99  instance euclidean_space \<subseteq> heine_borel
  37.100  proof
  37.101 -  fix s :: "'a set" and f :: "nat \<Rightarrow> 'a"
  37.102 -  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
  37.103 +  fix f :: "nat \<Rightarrow> 'a"
  37.104 +  assume f: "bounded (range f)"
  37.105    then obtain l::'a and r where r: "subseq r"
  37.106      and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially"
  37.107 -    using compact_lemma [OF s f] by blast
  37.108 +    using compact_lemma [OF f] by blast
  37.109    { fix e::real assume "e>0"
  37.110      hence "0 < e / real_of_nat DIM('a)" by (auto intro!: divide_pos_pos DIM_positive)
  37.111      with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))) sequentially"
  37.112 @@ -3338,19 +3354,16 @@
  37.113  
  37.114  instance prod :: (heine_borel, heine_borel) heine_borel
  37.115  proof
  37.116 -  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
  37.117 -  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
  37.118 -  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
  37.119 -  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
  37.120 -  obtain l1 r1 where r1: "subseq r1"
  37.121 -    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
  37.122 -    using bounded_imp_convergent_subsequence [OF s1 f1]
  37.123 -    unfolding o_def by fast
  37.124 -  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
  37.125 -  from f have f2: "\<forall>n. snd (f (r1 n)) \<in> snd ` s" by simp
  37.126 +  fix f :: "nat \<Rightarrow> 'a \<times> 'b"
  37.127 +  assume f: "bounded (range f)"
  37.128 +  from f have s1: "bounded (range (fst \<circ> f))" unfolding image_comp by (rule bounded_fst)
  37.129 +  obtain l1 r1 where r1: "subseq r1" and l1: "(\<lambda>n. fst (f (r1 n))) ----> l1"
  37.130 +    using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast
  37.131 +  from f have s2: "bounded (range (snd \<circ> f \<circ> r1))"
  37.132 +    by (auto simp add: image_comp intro: bounded_snd bounded_subset)
  37.133    obtain l2 r2 where r2: "subseq r2"
  37.134      and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
  37.135 -    using bounded_imp_convergent_subsequence [OF s2 f2]
  37.136 +    using bounded_imp_convergent_subsequence [OF s2]
  37.137      unfolding o_def by fast
  37.138    have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
  37.139      using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def .
    38.1 --- a/src/HOL/Probability/Discrete_Topology.thy	Tue Feb 12 17:39:45 2013 +0100
    38.2 +++ b/src/HOL/Probability/Discrete_Topology.thy	Wed Feb 13 11:46:48 2013 +0100
    38.3 @@ -3,7 +3,7 @@
    38.4  *)
    38.5  
    38.6  theory Discrete_Topology
    38.7 -imports Multivariate_Analysis
    38.8 +imports "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
    38.9  begin
   38.10  
   38.11  text {* Copy of discrete types with discrete topology. This space is polish. *}
    39.1 --- a/src/HOL/Probability/Lebesgue_Integration.thy	Tue Feb 12 17:39:45 2013 +0100
    39.2 +++ b/src/HOL/Probability/Lebesgue_Integration.thy	Wed Feb 13 11:46:48 2013 +0100
    39.3 @@ -297,7 +297,7 @@
    39.4      qed
    39.5    next
    39.6      fix x show "(SUP i. ?g i x) = max 0 (u x)"
    39.7 -    proof (rule ereal_SUPI)
    39.8 +    proof (rule SUP_eqI)
    39.9        fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def
   39.10          by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg
   39.11                                       mult_nonpos_nonneg mult_nonneg_nonneg)
   39.12 @@ -2147,7 +2147,7 @@
   39.13      assume neq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
   39.14      have "0 = limsup (\<lambda>n. 0 :: ereal)" by (simp add: Limsup_const)
   39.15      also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
   39.16 -      by (intro limsup_mono positive_integral_positive)
   39.17 +      by (simp add: Limsup_mono  positive_integral_positive)
   39.18      finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)" .
   39.19      have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (ereal (?diff n x))) \<partial>M)"
   39.20        using u'
   39.21 @@ -2178,11 +2178,11 @@
   39.22    qed
   39.23  
   39.24    have "liminf ?f \<le> limsup ?f"
   39.25 -    by (intro ereal_Liminf_le_Limsup trivial_limit_sequentially)
   39.26 +    by (intro Liminf_le_Limsup trivial_limit_sequentially)
   39.27    moreover
   39.28    { have "0 = liminf (\<lambda>n. 0 :: ereal)" by (simp add: Liminf_const)
   39.29      also have "\<dots> \<le> liminf ?f"
   39.30 -      by (intro liminf_mono positive_integral_positive)
   39.31 +      by (simp add: Liminf_mono positive_integral_positive)
   39.32      finally have "0 \<le> liminf ?f" . }
   39.33    ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0"
   39.34      using `limsup ?f = 0` by auto
    40.1 --- a/src/HOL/Probability/Lebesgue_Measure.thy	Tue Feb 12 17:39:45 2013 +0100
    40.2 +++ b/src/HOL/Probability/Lebesgue_Measure.thy	Wed Feb 13 11:46:48 2013 +0100
    40.3 @@ -698,7 +698,7 @@
    40.4    have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')"
    40.5    proof (rule tendsto_unique[OF trivial_limit_sequentially])
    40.6      have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
    40.7 -      unfolding u_eq by (intro LIMSEQ_ereal_SUPR incseq_positive_integral u)
    40.8 +      unfolding u_eq by (intro LIMSEQ_SUP incseq_positive_integral u)
    40.9      also note positive_integral_monotone_convergence_SUP
   40.10        [OF u(2)  borel_measurable_simple_function[OF u(1)] u(5), symmetric]
   40.11      finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
    41.1 --- a/src/HOL/Probability/Measure_Space.thy	Tue Feb 12 17:39:45 2013 +0100
    41.2 +++ b/src/HOL/Probability/Measure_Space.thy	Wed Feb 13 11:46:48 2013 +0100
    41.3 @@ -50,7 +50,7 @@
    41.4    then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
    41.5      by (auto simp: setsum_cases)
    41.6    moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
    41.7 -  proof (rule ereal_SUPI)
    41.8 +  proof (rule SUP_eqI)
    41.9      fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
   41.10      from this[of "Suc i"] show "f i \<le> y" by auto
   41.11    qed (insert assms, simp)
   41.12 @@ -523,7 +523,7 @@
   41.13  lemma SUP_emeasure_incseq:
   41.14    assumes A: "range A \<subseteq> sets M" "incseq A"
   41.15    shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
   41.16 -  using LIMSEQ_ereal_SUPR[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A]
   41.17 +  using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A]
   41.18    by (simp add: LIMSEQ_unique)
   41.19  
   41.20  lemma decseq_emeasure:
   41.21 @@ -1570,7 +1570,7 @@
   41.22        have "incseq (\<lambda>i. ?M (F i))"
   41.23          using `incseq F` unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
   41.24        then have "(\<lambda>i. ?M (F i)) ----> (SUP n. ?M (F n))"
   41.25 -        by (rule LIMSEQ_ereal_SUPR)
   41.26 +        by (rule LIMSEQ_SUP)
   41.27  
   41.28        moreover have "(SUP n. ?M (F n)) = \<infinity>"
   41.29        proof (rule SUP_PInfty)
    42.1 --- a/src/HOL/Probability/Regularity.thy	Tue Feb 12 17:39:45 2013 +0100
    42.2 +++ b/src/HOL/Probability/Regularity.thy	Wed Feb 13 11:46:48 2013 +0100
    42.3 @@ -16,7 +16,7 @@
    42.4    assumes f_nonneg: "\<And>i. 0 \<le> f i"
    42.5    assumes approx: "\<And>e. (e::real) > 0 \<Longrightarrow> \<exists>i \<in> A. x \<le> f i + e"
    42.6    shows "x = (SUP i : A. f i)"
    42.7 -proof (subst eq_commute, rule ereal_SUPI)
    42.8 +proof (subst eq_commute, rule SUP_eqI)
    42.9    show "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x" using f_bound by simp
   42.10  next
   42.11    fix y :: ereal assume f_le_y: "(\<And>i::'a. i \<in> A \<Longrightarrow> f i \<le> y)"
   42.12 @@ -45,7 +45,7 @@
   42.13    assumes f_nonneg: "\<And>i. 0 \<le> f i"
   42.14    assumes approx: "\<And>e. (e::real) > 0 \<Longrightarrow> \<exists>i \<in> A. f i \<le> x + e"
   42.15    shows "x = (INF i : A. f i)"
   42.16 -proof (subst eq_commute, rule ereal_INFI)
   42.17 +proof (subst eq_commute, rule INF_eqI)
   42.18    show "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i" using f_bound by simp
   42.19  next
   42.20    fix y :: ereal assume f_le_y: "(\<And>i::'a. i \<in> A \<Longrightarrow> y \<le> f i)"
   42.21 @@ -79,7 +79,7 @@
   42.22    from INF have "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> f i) \<Longrightarrow> y \<le> x" by (auto intro: INF_greatest)
   42.23    ultimately
   42.24    have "(INF i : A. f i) = x + e" using `e > 0`
   42.25 -    by (intro ereal_INFI)
   42.26 +    by (intro INF_eqI)
   42.27        (force, metis add.comm_neutral add_left_mono ereal_less(1)
   42.28          linorder_not_le not_less_iff_gr_or_eq)
   42.29    thus False using assms by auto
   42.30 @@ -97,7 +97,7 @@
   42.31    from SUP have "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> y \<ge> x" by (auto intro: SUP_least)
   42.32    ultimately
   42.33    have "(SUP i : A. f i) = x - e" using `e > 0` `\<bar>x\<bar> \<noteq> \<infinity>`
   42.34 -    by (intro ereal_SUPI)
   42.35 +    by (intro SUP_eqI)
   42.36         (metis PInfty_neq_ereal(2) abs_ereal.simps(1) ereal_minus_le linorder_linear,
   42.37          metis ereal_between(1) ereal_less(2) less_eq_ereal_def order_trans)
   42.38    thus False using assms by auto
   42.39 @@ -290,7 +290,7 @@
   42.40        have "emeasure M A \<le> (INF U:{U. A \<subseteq> U \<and> open U}. emeasure M U)"
   42.41          by (rule INF_greatest) (auto intro!: emeasure_mono simp: sb)
   42.42        ultimately show ?thesis by simp
   42.43 -    qed (auto intro!: ereal_INFI)
   42.44 +    qed (auto intro!: INF_eqI)
   42.45      note `?inner A` `?outer A` }
   42.46    note closed_in_D = this
   42.47    from `B \<in> sets borel`
   42.48 @@ -299,8 +299,8 @@
   42.49    then show "?inner B" "?outer B"
   42.50    proof (induct B rule: sigma_sets_induct_disjoint)
   42.51      case empty
   42.52 -    { case 1 show ?case by (intro ereal_SUPI[symmetric]) auto }
   42.53 -    { case 2 show ?case by (intro ereal_INFI[symmetric]) (auto elim!: meta_allE[of _ "{}"]) }
   42.54 +    { case 1 show ?case by (intro SUP_eqI[symmetric]) auto }
   42.55 +    { case 2 show ?case by (intro INF_eqI[symmetric]) (auto elim!: meta_allE[of _ "{}"]) }
   42.56    next
   42.57      case (basic B)
   42.58      { case 1 from basic closed_in_D show ?case by auto }
    43.1 --- a/src/HOL/RealVector.thy	Tue Feb 12 17:39:45 2013 +0100
    43.2 +++ b/src/HOL/RealVector.thy	Wed Feb 13 11:46:48 2013 +0100
    43.3 @@ -508,6 +508,56 @@
    43.4  
    43.5  end
    43.6  
    43.7 +inductive generate_topology for S where
    43.8 +  UNIV: "generate_topology S UNIV"
    43.9 +| Int: "generate_topology S a \<Longrightarrow> generate_topology S b \<Longrightarrow> generate_topology S (a \<inter> b)"
   43.10 +| UN: "(\<And>k. k \<in> K \<Longrightarrow> generate_topology S k) \<Longrightarrow> generate_topology S (\<Union>K)"
   43.11 +| Basis: "s \<in> S \<Longrightarrow> generate_topology S s"
   43.12 +
   43.13 +hide_fact (open) UNIV Int UN Basis 
   43.14 +
   43.15 +lemma generate_topology_Union: 
   43.16 +  "(\<And>k. k \<in> I \<Longrightarrow> generate_topology S (K k)) \<Longrightarrow> generate_topology S (\<Union>k\<in>I. K k)"
   43.17 +  unfolding SUP_def by (intro generate_topology.UN) auto
   43.18 +
   43.19 +lemma topological_space_generate_topology:
   43.20 +  "class.topological_space (generate_topology S)"
   43.21 +  by default (auto intro: generate_topology.intros)
   43.22 +
   43.23 +class order_topology = order + "open" +
   43.24 +  assumes open_generated_order: "open = generate_topology (range (\<lambda>a. {..< a}) \<union> range (\<lambda>a. {a <..}))"
   43.25 +begin
   43.26 +
   43.27 +subclass topological_space
   43.28 +  unfolding open_generated_order
   43.29 +  by (rule topological_space_generate_topology)
   43.30 +
   43.31 +lemma open_greaterThan [simp]: "open {a <..}"
   43.32 +  unfolding open_generated_order by (auto intro: generate_topology.Basis)
   43.33 +
   43.34 +lemma open_lessThan [simp]: "open {..< a}"
   43.35 +  unfolding open_generated_order by (auto intro: generate_topology.Basis)
   43.36 +
   43.37 +lemma open_greaterThanLessThan [simp]: "open {a <..< b}"
   43.38 +   unfolding greaterThanLessThan_eq by (simp add: open_Int)
   43.39 +
   43.40 +end
   43.41 +
   43.42 +class linorder_topology = linorder + order_topology
   43.43 +
   43.44 +lemma closed_atMost [simp]: "closed {.. a::'a::linorder_topology}"
   43.45 +  by (simp add: closed_open)
   43.46 +
   43.47 +lemma closed_atLeast [simp]: "closed {a::'a::linorder_topology ..}"
   43.48 +  by (simp add: closed_open)
   43.49 +
   43.50 +lemma closed_atLeastAtMost [simp]: "closed {a::'a::linorder_topology .. b}"
   43.51 +proof -
   43.52 +  have "{a .. b} = {a ..} \<inter> {.. b}"
   43.53 +    by auto
   43.54 +  then show ?thesis
   43.55 +    by (simp add: closed_Int)
   43.56 +qed
   43.57  
   43.58  subsection {* Metric spaces *}
   43.59  
   43.60 @@ -621,10 +671,20 @@
   43.61    assumes dist_norm: "dist x y = norm (x - y)"
   43.62  
   43.63  class real_normed_vector = real_vector + sgn_div_norm + dist_norm + open_dist +
   43.64 -  assumes norm_ge_zero [simp]: "0 \<le> norm x"
   43.65 -  and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
   43.66 +  assumes norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
   43.67    and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
   43.68    and norm_scaleR [simp]: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
   43.69 +begin
   43.70 +
   43.71 +lemma norm_ge_zero [simp]: "0 \<le> norm x"
   43.72 +proof -
   43.73 +  have "0 = norm (x + -1 *\<^sub>R x)" 
   43.74 +    using scaleR_add_left[of 1 "-1" x] norm_scaleR[of 0 x] by (simp add: scaleR_one)
   43.75 +  also have "\<dots> \<le> norm x + norm (-1 *\<^sub>R x)" by (rule norm_triangle_ineq)
   43.76 +  finally show ?thesis by simp
   43.77 +qed
   43.78 +
   43.79 +end
   43.80  
   43.81  class real_normed_algebra = real_algebra + real_normed_vector +
   43.82    assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
   43.83 @@ -850,7 +910,6 @@
   43.84  apply (rule dist_real_def)
   43.85  apply (rule open_real_def)
   43.86  apply (simp add: sgn_real_def)
   43.87 -apply (rule abs_ge_zero)
   43.88  apply (rule abs_eq_0)
   43.89  apply (rule abs_triangle_ineq)
   43.90  apply (rule abs_mult)
   43.91 @@ -859,46 +918,46 @@
   43.92  
   43.93  end
   43.94  
   43.95 -lemma open_real_lessThan [simp]:
   43.96 -  fixes a :: real shows "open {..<a}"
   43.97 -unfolding open_real_def dist_real_def
   43.98 -proof (clarify)
   43.99 -  fix x assume "x < a"
  43.100 -  hence "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
  43.101 -  thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
  43.102 +instance real :: linorder_topology
  43.103 +proof
  43.104 +  show "(open :: real set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)"
  43.105 +  proof (rule ext, safe)
  43.106 +    fix S :: "real set" assume "open S"
  43.107 +    then guess f unfolding open_real_def bchoice_iff ..
  43.108 +    then have *: "S = (\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x})"
  43.109 +      by (fastforce simp: dist_real_def)
  43.110 +    show "generate_topology (range lessThan \<union> range greaterThan) S"
  43.111 +      apply (subst *)
  43.112 +      apply (intro generate_topology_Union generate_topology.Int)
  43.113 +      apply (auto intro: generate_topology.Basis)
  43.114 +      done
  43.115 +  next
  43.116 +    fix S :: "real set" assume "generate_topology (range lessThan \<union> range greaterThan) S"
  43.117 +    moreover have "\<And>a::real. open {..<a}"
  43.118 +      unfolding open_real_def dist_real_def
  43.119 +    proof clarify
  43.120 +      fix x a :: real assume "x < a"
  43.121 +      hence "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
  43.122 +      thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
  43.123 +    qed
  43.124 +    moreover have "\<And>a::real. open {a <..}"
  43.125 +      unfolding open_real_def dist_real_def
  43.126 +    proof clarify
  43.127 +      fix x a :: real assume "a < x"
  43.128 +      hence "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
  43.129 +      thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
  43.130 +    qed
  43.131 +    ultimately show "open S"
  43.132 +      by induct auto
  43.133 +  qed
  43.134  qed
  43.135  
  43.136 -lemma open_real_greaterThan [simp]:
  43.137 -  fixes a :: real shows "open {a<..}"
  43.138 -unfolding open_real_def dist_real_def
  43.139 -proof (clarify)
  43.140 -  fix x assume "a < x"
  43.141 -  hence "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
  43.142 -  thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
  43.143 -qed
  43.144 -
  43.145 -lemma open_real_greaterThanLessThan [simp]:
  43.146 -  fixes a b :: real shows "open {a<..<b}"
  43.147 -proof -
  43.148 -  have "{a<..<b} = {a<..} \<inter> {..<b}" by auto
  43.149 -  thus "open {a<..<b}" by (simp add: open_Int)
  43.150 -qed
  43.151 -
  43.152 -lemma closed_real_atMost [simp]: 
  43.153 -  fixes a :: real shows "closed {..a}"
  43.154 -unfolding closed_open by simp
  43.155 -
  43.156 -lemma closed_real_atLeast [simp]:
  43.157 -  fixes a :: real shows "closed {a..}"
  43.158 -unfolding closed_open by simp
  43.159 -
  43.160 -lemma closed_real_atLeastAtMost [simp]:
  43.161 -  fixes a b :: real shows "closed {a..b}"
  43.162 -proof -
  43.163 -  have "{a..b} = {a..} \<inter> {..b}" by auto
  43.164 -  thus "closed {a..b}" by (simp add: closed_Int)
  43.165 -qed
  43.166 -
  43.167 +lemmas open_real_greaterThan = open_greaterThan[where 'a=real]
  43.168 +lemmas open_real_lessThan = open_lessThan[where 'a=real]
  43.169 +lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real]
  43.170 +lemmas closed_real_atMost = closed_atMost[where 'a=real]
  43.171 +lemmas closed_real_atLeast = closed_atLeast[where 'a=real]
  43.172 +lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real]
  43.173  
  43.174  subsection {* Extra type constraints *}
  43.175  
  43.176 @@ -1172,6 +1231,30 @@
  43.177  instance t2_space \<subseteq> t1_space
  43.178  proof qed (fast dest: hausdorff)
  43.179  
  43.180 +lemma (in linorder) less_separate:
  43.181 +  assumes "x < y"
  43.182 +  shows "\<exists>a b. x \<in> {..< a} \<and> y \<in> {b <..} \<and> {..< a} \<inter> {b <..} = {}"
  43.183 +proof cases
  43.184 +  assume "\<exists>z. x < z \<and> z < y"
  43.185 +  then guess z ..
  43.186 +  then have "x \<in> {..< z} \<and> y \<in> {z <..} \<and> {z <..} \<inter> {..< z} = {}"
  43.187 +    by auto
  43.188 +  then show ?thesis by blast
  43.189 +next
  43.190 +  assume "\<not> (\<exists>z. x < z \<and> z < y)"
  43.191 +  with `x < y` have "x \<in> {..< y} \<and> y \<in> {x <..} \<and> {x <..} \<inter> {..< y} = {}"
  43.192 +    by auto
  43.193 +  then show ?thesis by blast
  43.194 +qed
  43.195 +
  43.196 +instance linorder_topology \<subseteq> t2_space
  43.197 +proof
  43.198 +  fix x y :: 'a
  43.199 +  from less_separate[of x y] less_separate[of y x]
  43.200 +  show "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
  43.201 +    by (elim neqE) (metis open_lessThan open_greaterThan Int_commute)+
  43.202 +qed
  43.203 +
  43.204  instance metric_space \<subseteq> t2_space
  43.205  proof
  43.206    fix x y :: "'a::metric_space"
    44.1 --- a/src/HOL/SEQ.thy	Tue Feb 12 17:39:45 2013 +0100
    44.2 +++ b/src/HOL/SEQ.thy	Wed Feb 13 11:46:48 2013 +0100
    44.3 @@ -368,19 +368,16 @@
    44.4        and bdd: "\<And>n. f n \<le> l"
    44.5        and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
    44.6    shows "f ----> l"
    44.7 -  unfolding LIMSEQ_def
    44.8 -proof safe
    44.9 -  fix r :: real assume "0 < r"
   44.10 -  with bdd en[of "r / 2"] obtain n where n: "dist (f n) l \<le> r / 2"
   44.11 -    by (auto simp add: field_simps dist_real_def)
   44.12 -  { fix N assume "n \<le> N"
   44.13 -    then have "dist (f N) l \<le> dist (f n) l"
   44.14 -      using incseq_SucI[of f] inc bdd by (auto dest!: incseqD simp: dist_real_def)
   44.15 -    then have "dist (f N) l < r"
   44.16 -      using `0 < r` n by simp }
   44.17 -  with `0 < r` show "\<exists>no. \<forall>n\<ge>no. dist (f n) l < r"
   44.18 -    by (auto simp add: LIMSEQ_def field_simps intro!: exI[of _ n])
   44.19 -qed
   44.20 +proof (rule increasing_tendsto)
   44.21 +  fix x assume "x < l"
   44.22 +  with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x"
   44.23 +    by auto
   44.24 +  from en[OF `0 < e`] obtain n where "l - e \<le> f n"
   44.25 +    by (auto simp: field_simps)
   44.26 +  with `e < l - x` `0 < e` have "x < f n" by simp
   44.27 +  with incseq_SucI[of f, OF inc] show "eventually (\<lambda>n. x < f n) sequentially"
   44.28 +    by (auto simp: eventually_sequentially incseq_def intro: less_le_trans)
   44.29 +qed (insert bdd, auto)
   44.30  
   44.31  lemma Bseq_inverse_lemma:
   44.32    fixes x :: "'a::real_normed_div_algebra"
   44.33 @@ -437,15 +434,15 @@
   44.34    by auto
   44.35  
   44.36  lemma LIMSEQ_le_const:
   44.37 -  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
   44.38 +  "\<lbrakk>X ----> (x::'a::linorder_topology); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
   44.39    using tendsto_le_const[of sequentially X x a] by (simp add: eventually_sequentially)
   44.40  
   44.41  lemma LIMSEQ_le:
   44.42 -  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
   44.43 +  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::'a::linorder_topology)"
   44.44    using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially)
   44.45  
   44.46  lemma LIMSEQ_le_const2:
   44.47 -  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
   44.48 +  "\<lbrakk>X ----> (x::'a::linorder_topology); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
   44.49    by (rule LIMSEQ_le[of X x "\<lambda>n. a"]) (auto simp: tendsto_const)
   44.50  
   44.51  subsection {* Convergence *}
    45.1 --- a/src/HOL/Series.thy	Tue Feb 12 17:39:45 2013 +0100
    45.2 +++ b/src/HOL/Series.thy	Wed Feb 13 11:46:48 2013 +0100
    45.3 @@ -367,7 +367,7 @@
    45.4  
    45.5  lemma pos_summable:
    45.6    fixes f:: "nat \<Rightarrow> real"
    45.7 -  assumes pos: "!!n. 0 \<le> f n" and le: "!!n. setsum f {0..<n} \<le> x"
    45.8 +  assumes pos: "\<And>n. 0 \<le> f n" and le: "\<And>n. setsum f {0..<n} \<le> x"
    45.9    shows "summable f"
   45.10  proof -
   45.11    have "convergent (\<lambda>n. setsum f {0..<n})"
   45.12 @@ -386,35 +386,65 @@
   45.13  qed
   45.14  
   45.15  lemma series_pos_le:
   45.16 -  fixes f :: "nat \<Rightarrow> real"
   45.17 +  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
   45.18    shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 \<le> f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} \<le> suminf f"
   45.19 -apply (drule summable_sums)
   45.20 -apply (simp add: sums_def)
   45.21 -apply (cut_tac k = "setsum f {0..<n}" in tendsto_const)
   45.22 -apply (erule LIMSEQ_le, blast)
   45.23 -apply (rule_tac x="n" in exI, clarify)
   45.24 -apply (rule setsum_mono2)
   45.25 -apply auto
   45.26 -done
   45.27 +  apply (drule summable_sums)
   45.28 +  apply (simp add: sums_def)
   45.29 +  apply (rule LIMSEQ_le_const)
   45.30 +  apply assumption
   45.31 +  apply (intro exI[of _ n])
   45.32 +  apply (auto intro!: setsum_mono2)
   45.33 +  done
   45.34  
   45.35  lemma series_pos_less:
   45.36 -  fixes f :: "nat \<Rightarrow> real"
   45.37 +  fixes f :: "nat \<Rightarrow> 'a::{ordered_ab_semigroup_add_imp_le, ordered_comm_monoid_add, linorder_topology}"
   45.38    shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 < f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} < suminf f"
   45.39 -apply (rule_tac y="setsum f {0..<Suc n}" in order_less_le_trans)
   45.40 -apply simp
   45.41 -apply (erule series_pos_le)
   45.42 -apply (simp add: order_less_imp_le)
   45.43 -done
   45.44 +  apply (rule_tac y="setsum f {0..<Suc n}" in order_less_le_trans)
   45.45 +  using add_less_cancel_left [of "setsum f {0..<n}" 0 "f n"]
   45.46 +  apply simp
   45.47 +  apply (erule series_pos_le)
   45.48 +  apply (simp add: order_less_imp_le)
   45.49 +  done
   45.50 +
   45.51 +lemma suminf_eq_zero_iff:
   45.52 +  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
   45.53 +  shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> suminf f = 0 \<longleftrightarrow> (\<forall>n. f n = 0)"
   45.54 +proof
   45.55 +  assume "summable f" "suminf f = 0" and pos: "\<forall>n. 0 \<le> f n"
   45.56 +  then have "f sums 0"
   45.57 +    by (simp add: sums_iff)
   45.58 +  then have f: "(\<lambda>n. \<Sum>i<n. f i) ----> 0"
   45.59 +    by (simp add: sums_def atLeast0LessThan)
   45.60 +  have "\<And>i. (\<Sum>n\<in>{i}. f n) \<le> 0"
   45.61 +  proof (rule LIMSEQ_le_const[OF f])
   45.62 +    fix i show "\<exists>N. \<forall>n\<ge>N. (\<Sum>n\<in>{i}. f n) \<le> setsum f {..<n}"
   45.63 +      using pos by (intro exI[of _ "Suc i"] allI impI setsum_mono2) auto
   45.64 +  qed
   45.65 +  with pos show "\<forall>n. f n = 0"
   45.66 +    by (auto intro!: antisym)
   45.67 +next
   45.68 +  assume "\<forall>n. f n = 0"
   45.69 +  then have "f = (\<lambda>n. 0)"
   45.70 +    by auto
   45.71 +  then show "suminf f = 0"
   45.72 +    by simp
   45.73 +qed
   45.74 +
   45.75 +lemma suminf_gt_zero_iff:
   45.76 +  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
   45.77 +  shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 < suminf f \<longleftrightarrow> (\<exists>i. 0 < f i)"
   45.78 +  using series_pos_le[of f 0] suminf_eq_zero_iff[of f]
   45.79 +  by (simp add: less_le)
   45.80  
   45.81  lemma suminf_gt_zero:
   45.82 -  fixes f :: "nat \<Rightarrow> real"
   45.83 +  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
   45.84    shows "\<lbrakk>summable f; \<forall>n. 0 < f n\<rbrakk> \<Longrightarrow> 0 < suminf f"
   45.85 -by (drule_tac n="0" in series_pos_less, simp_all)
   45.86 +  using suminf_gt_zero_iff[of f] by (simp add: less_imp_le)
   45.87  
   45.88  lemma suminf_ge_zero:
   45.89 -  fixes f :: "nat \<Rightarrow> real"
   45.90 +  fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
   45.91    shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 \<le> suminf f"
   45.92 -by (drule_tac n="0" in series_pos_le, simp_all)
   45.93 +  by (drule_tac n="0" in series_pos_le, simp_all)
   45.94  
   45.95  lemma sumr_pos_lt_pair:
   45.96    fixes f :: "nat \<Rightarrow> real"
   45.97 @@ -493,9 +523,14 @@
   45.98  done
   45.99  
  45.100  lemma suminf_le:
  45.101 -  fixes x :: real
  45.102 +  fixes x :: "'a :: {ordered_comm_monoid_add, linorder_topology}"
  45.103    shows "summable f \<Longrightarrow> (!!n. setsum f {0..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
  45.104 -  by (simp add: summable_convergent_sumr_iff suminf_eq_lim lim_le)
  45.105 +  apply (drule summable_sums)
  45.106 +  apply (simp add: sums_def)
  45.107 +  apply (rule LIMSEQ_le_const2)
  45.108 +  apply assumption
  45.109 +  apply auto
  45.110 +  done
  45.111  
  45.112  lemma summable_Cauchy:
  45.113       "summable (f::nat \<Rightarrow> 'a::banach) =
  45.114 @@ -575,7 +610,7 @@
  45.115  text{*Limit comparison property for series (c.f. jrh)*}
  45.116  
  45.117  lemma summable_le:
  45.118 -  fixes f g :: "nat \<Rightarrow> real"
  45.119 +  fixes f g :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
  45.120    shows "\<lbrakk>\<forall>n. f n \<le> g n; summable f; summable g\<rbrakk> \<Longrightarrow> suminf f \<le> suminf g"
  45.121  apply (drule summable_sums)+
  45.122  apply (simp only: sums_def, erule (1) LIMSEQ_le)
  45.123 @@ -597,15 +632,7 @@
  45.124    fixes f::"nat\<Rightarrow>real"
  45.125    assumes gt0: "\<forall>n. 0 \<le> f n" and sm: "summable f"
  45.126    shows "0 \<le> suminf f"
  45.127 -proof -
  45.128 -  let ?g = "(\<lambda>n. (0::real))"
  45.129 -  from gt0 have "\<forall>n. ?g n \<le> f n" by simp
  45.130 -  moreover have "summable ?g" by (rule summable_zero)
  45.131 -  moreover from sm have "summable f" .
  45.132 -  ultimately have "suminf ?g \<le> suminf f" by (rule summable_le)
  45.133 -  then show "0 \<le> suminf f" by simp
  45.134 -qed
  45.135 -
  45.136 +  using suminf_ge_zero[OF sm gt0] by simp
  45.137  
  45.138  text{*Absolute convergence imples normal convergence*}
  45.139  lemma summable_norm_cancel:
    46.1 --- a/src/HOL/Set_Interval.thy	Tue Feb 12 17:39:45 2013 +0100
    46.2 +++ b/src/HOL/Set_Interval.thy	Wed Feb 13 11:46:48 2013 +0100
    46.3 @@ -117,6 +117,11 @@
    46.4  lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}"
    46.5  by (blast intro: order_antisym)
    46.6  
    46.7 +lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \<inter> { b <..} = { max a b <..}"
    46.8 +  by auto
    46.9 +
   46.10 +lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \<inter> {..< b} = {..< min a b}"
   46.11 +  by auto
   46.12  
   46.13  subsection {* Logical Equivalences for Set Inclusion and Equality *}
   46.14  
   46.15 @@ -190,6 +195,9 @@
   46.16  breaks many proofs. Since it only helps blast, it is better to leave well
   46.17  alone *}
   46.18  
   46.19 +lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \<inter> {..< b }"
   46.20 +  by auto
   46.21 +
   46.22  end
   46.23  
   46.24  subsubsection{* Emptyness, singletons, subset *}
    47.1 --- a/src/HOL/TPTP/MaSh_Export.thy	Tue Feb 12 17:39:45 2013 +0100
    47.2 +++ b/src/HOL/TPTP/MaSh_Export.thy	Wed Feb 13 11:46:48 2013 +0100
    47.3 @@ -16,6 +16,8 @@
    47.4  
    47.5  declare [[sledgehammer_instantiate_inducts = false]]
    47.6  
    47.7 +hide_fact (open) HOL.ext
    47.8 +
    47.9  ML {*
   47.10  !Multithreading.max_threads
   47.11  *}
    48.1 --- a/src/HOL/TPTP/mash_eval.ML	Tue Feb 12 17:39:45 2013 +0100
    48.2 +++ b/src/HOL/TPTP/mash_eval.ML	Wed Feb 13 11:46:48 2013 +0100
    48.3 @@ -30,11 +30,10 @@
    48.4  open Sledgehammer_Provers
    48.5  open Sledgehammer_Isar
    48.6  
    48.7 -val MePoN = "MePo"
    48.8 -val MaSh_IsarN = "MaSh-Isar"
    48.9 -val MaSh_ProverN = "MaSh-Prover"
   48.10 -val MeSh_IsarN = "MeSh-Isar"
   48.11 -val MeSh_ProverN = "MeSh-Prover"
   48.12 +val MaSh_IsarN = MaShN ^ "-Isar"
   48.13 +val MaSh_ProverN = MaShN ^ "-Prover"
   48.14 +val MeSh_IsarN = MeShN ^ "-Isar"
   48.15 +val MeSh_ProverN = MeShN ^ "-Prover"
   48.16  val IsarN = "Isar"
   48.17  
   48.18  fun in_range (from, to) j =
   48.19 @@ -44,6 +43,7 @@
   48.20          mepo_file_name mash_isar_file_name mash_prover_file_name
   48.21          mesh_isar_file_name mesh_prover_file_name report_file_name =
   48.22    let
   48.23 +    val zeros = [0, 0, 0, 0, 0, 0]
   48.24      val report_path = report_file_name |> Path.explode
   48.25      val _ = File.write report_path ""
   48.26      fun print s = File.append report_path (s ^ "\n")
   48.27 @@ -71,7 +71,7 @@
   48.28      val str_of_method = enclose "  " ": "
   48.29      fun str_of_result method facts ({outcome, run_time, used_facts, ...}
   48.30                                       : prover_result) =
   48.31 -      let val facts = facts |> map (fn ((name, _), _) => name ()) in
   48.32 +      let val facts = facts |> map (fst o fst) in
   48.33          str_of_method method ^
   48.34          (if is_none outcome then
   48.35             "Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^
   48.36 @@ -111,7 +111,7 @@
   48.37            val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
   48.38            val isar_deps = isar_dependencies_of name_tabs th
   48.39            val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
   48.40 -          val find_suggs = find_suggested_facts facts
   48.41 +          val find_suggs = find_suggested_facts facts #> map fact_of_raw_fact
   48.42            fun get_facts [] compute = compute facts
   48.43              | get_facts suggs _ = find_suggs suggs
   48.44            val mepo_facts =
   48.45 @@ -121,7 +121,8 @@
   48.46              |> weight_mepo_facts
   48.47            fun mash_of suggs =
   48.48              get_facts suggs (fn _ =>
   48.49 -                find_mash_suggestions slack_max_facts suggs facts [] [] |> fst)
   48.50 +                find_mash_suggestions slack_max_facts suggs facts [] []
   48.51 +                |> fst |> map fact_of_raw_fact)
   48.52              |> weight_mash_facts
   48.53            val mash_isar_facts = mash_of mash_isar_suggs
   48.54            val mash_prover_facts = mash_of mash_prover_suggs
   48.55 @@ -160,6 +161,7 @@
   48.56                    |> map (get #> nickify)
   48.57                    |> maybe_instantiate_inducts ctxt hyp_ts concl_t
   48.58                    |> take (the max_facts)
   48.59 +                  |> map fact_of_raw_fact
   48.60                  val ctxt = ctxt |> set_file_name method prob_dir_name
   48.61                  val res as {outcome, ...} =
   48.62                    run_prover_for_mash ctxt params prover facts goal
   48.63 @@ -179,11 +181,11 @@
   48.64            map snd ress
   48.65          end
   48.66        else
   48.67 -        [0, 0, 0, 0, 0, 0]
   48.68 +        zeros
   48.69      fun total_of method ok n =
   48.70        str_of_method method ^ string_of_int ok ^ " (" ^
   48.71        Real.fmt (StringCvt.FIX (SOME 1))
   48.72 -               (100.0 * Real.fromInt ok / Real.fromInt n) ^ "%)"
   48.73 +               (100.0 * Real.fromInt ok / Real.fromInt (Int.max (1, n))) ^ "%)"
   48.74      val inst_inducts = Config.get ctxt instantiate_inducts
   48.75      val options =
   48.76        ["prover = " ^ prover,
   48.77 @@ -199,7 +201,7 @@
   48.78      val n = length oks
   48.79      val [mepo_ok, mash_isar_ok, mash_prover_ok, mesh_isar_ok, mesh_prover_ok,
   48.80           isar_ok] =
   48.81 -      map Integer.sum (map_transpose I oks)
   48.82 +      if n = 0 then zeros else map Integer.sum (map_transpose I oks)
   48.83    in
   48.84      ["Successes (of " ^ string_of_int n ^ " goals)",
   48.85       total_of MePoN mepo_ok n,
    49.1 --- a/src/HOL/TPTP/mash_export.ML	Tue Feb 12 17:39:45 2013 +0100
    49.2 +++ b/src/HOL/TPTP/mash_export.ML	Wed Feb 13 11:46:48 2013 +0100
    49.3 @@ -83,15 +83,31 @@
    49.4        in File.append path s end
    49.5    in List.app do_fact facts end
    49.6  
    49.7 -fun isar_or_prover_dependencies_of ctxt params_opt facts name_tabs th
    49.8 -                                   isar_deps_opt =
    49.9 -  case params_opt of
   49.10 -    SOME (params as {provers = prover :: _, ...}) =>
   49.11 -    prover_dependencies_of ctxt params prover 0 facts name_tabs th |> snd
   49.12 -  | NONE =>
   49.13 -    case isar_deps_opt of
   49.14 -      SOME deps => deps
   49.15 -    | NONE => isar_dependencies_of name_tabs th
   49.16 +val prover_marker = "$a"
   49.17 +val isar_marker = "$i"
   49.18 +val omitted_marker = "$o"
   49.19 +val unprovable_marker = "$u" (* axiom or definition or characteristic theorem *)
   49.20 +val prover_failed_marker = "$x"
   49.21 +
   49.22 +fun smart_dependencies_of ctxt params_opt facts name_tabs th isar_deps_opt =
   49.23 +  let
   49.24 +    val (marker, deps) =
   49.25 +      case params_opt of
   49.26 +        SOME (params as {provers = prover :: _, ...}) =>
   49.27 +        prover_dependencies_of ctxt params prover 0 facts name_tabs th
   49.28 +        |>> (fn true => prover_marker | false => prover_failed_marker)
   49.29 +      | NONE =>
   49.30 +        let
   49.31 +          val deps =
   49.32 +            case isar_deps_opt of
   49.33 +              SOME deps => deps
   49.34 +            | NONE => isar_dependencies_of name_tabs th
   49.35 +        in (if null deps then unprovable_marker else isar_marker, deps) end
   49.36 +  in
   49.37 +    case trim_dependencies th deps of
   49.38 +      SOME deps => (marker, deps)
   49.39 +    | NONE => (omitted_marker, [])
   49.40 +  end
   49.41  
   49.42  fun generate_isar_or_prover_dependencies ctxt params_opt range thys include_thys
   49.43                                           file_name =
   49.44 @@ -105,10 +121,9 @@
   49.45          let
   49.46            val name = nickname_of_thm th
   49.47            val _ = tracing ("Fact " ^ string_of_int j ^ ": " ^ name)
   49.48 -          val deps =
   49.49 -            isar_or_prover_dependencies_of ctxt params_opt facts name_tabs th
   49.50 -                                           NONE
   49.51 -        in encode_str name ^ ": " ^ encode_strs deps ^ "\n" end
   49.52 +          val (marker, deps) =
   49.53 +            smart_dependencies_of ctxt params_opt facts name_tabs th NONE
   49.54 +        in encode_str name ^ ": " ^ marker ^ " " ^ encode_strs deps ^ "\n" end
   49.55        else
   49.56          ""
   49.57      val lines = Par_List.map do_fact (tag_list 1 facts)
   49.58 @@ -123,7 +138,7 @@
   49.59  fun is_bad_query ctxt ho_atp step j th isar_deps =
   49.60    j mod step <> 0 orelse
   49.61    Thm.legacy_get_kind th = "" orelse
   49.62 -  null (these (trim_dependencies th isar_deps)) orelse
   49.63 +  null isar_deps orelse
   49.64    is_blacklisted_or_something ctxt ho_atp (Thm.get_name_hint th)
   49.65  
   49.66  fun generate_isar_or_prover_commands ctxt prover params_opt (range, step) thys
   49.67 @@ -141,9 +156,9 @@
   49.68            val feats =
   49.69              features_of ctxt prover (theory_of_thm th) stature [prop_of th]
   49.70            val isar_deps = isar_dependencies_of name_tabs th
   49.71 -          val deps =
   49.72 -            isar_or_prover_dependencies_of ctxt params_opt facts name_tabs th
   49.73 -                                           (SOME isar_deps)
   49.74 +          val (marker, deps) =
   49.75 +            smart_dependencies_of ctxt params_opt facts name_tabs th
   49.76 +                                  (SOME isar_deps)
   49.77            val core =
   49.78              encode_str name ^ ": " ^ encode_strs prevs ^ "; " ^
   49.79              encode_features (sort_wrt fst feats)
   49.80 @@ -151,8 +166,7 @@
   49.81              if is_bad_query ctxt ho_atp step j th isar_deps then ""
   49.82              else "? " ^ core ^ "\n"
   49.83            val update =
   49.84 -            "! " ^ core ^ "; " ^
   49.85 -            encode_strs (these (trim_dependencies th deps)) ^ "\n"
   49.86 +            "! " ^ core ^ "; " ^ marker ^ " " ^ encode_strs deps ^ "\n"
   49.87          in query ^ update end
   49.88        else
   49.89          ""
    50.1 --- a/src/HOL/TPTP/sledgehammer_tactics.ML	Tue Feb 12 17:39:45 2013 +0100
    50.2 +++ b/src/HOL/TPTP/sledgehammer_tactics.ML	Wed Feb 13 11:46:48 2013 +0100
    50.3 @@ -42,15 +42,15 @@
    50.4        |> relevant_facts ctxt params name
    50.5               (the_default default_max_facts max_facts) fact_override hyp_ts
    50.6               concl_t
    50.7 +      |> hd |> snd
    50.8      val problem =
    50.9        {state = Proof.init ctxt, goal = goal, subgoal = i, subgoal_count = n,
   50.10 -       facts = facts |> map (apfst (apfst (fn name => name ())))
   50.11 -                     |> map Untranslated_Fact}
   50.12 +       factss = [("", facts)]}
   50.13    in
   50.14      (case prover params (K (K (K ""))) problem of
   50.15        {outcome = NONE, used_facts, ...} => used_facts |> map fst |> SOME
   50.16      | _ => NONE)
   50.17 -      handle ERROR message => (warning ("Error: " ^ message ^ "\n"); NONE)
   50.18 +    handle ERROR message => (warning ("Error: " ^ message ^ "\n"); NONE)
   50.19    end
   50.20  
   50.21  fun sledgehammer_with_metis_tac ctxt override_params fact_override i th =
    51.1 --- a/src/HOL/Tools/ATP/atp_proof.ML	Tue Feb 12 17:39:45 2013 +0100
    51.2 +++ b/src/HOL/Tools/ATP/atp_proof.ML	Wed Feb 13 11:46:48 2013 +0100
    51.3 @@ -201,9 +201,7 @@
    51.4  
    51.5  fun step_name_ord p =
    51.6    let val q = pairself fst p in
    51.7 -    (* The "unprefix" part is to cope with remote Vampire's output. The proper
    51.8 -       solution would be to perform a topological sort, e.g. using the nice
    51.9 -       "Graph" functor. *)
   51.10 +    (* The "unprefix" part is to cope with remote Vampire's output. *)
   51.11      case pairself (Int.fromString
   51.12                     o perhaps (try (unprefix vampire_fact_prefix))) q of
   51.13        (NONE, NONE) => string_ord q
   51.14 @@ -229,6 +227,8 @@
   51.15    || Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig
   51.16       >> (fn (ss1, ss2) => implode ss1 ^ implode ss2)
   51.17  
   51.18 +val scan_nat = Scan.repeat1 (Scan.one Symbol.is_ascii_digit) >> implode
   51.19 +
   51.20  val skip_term =
   51.21    let
   51.22      fun skip _ accum [] = (accum, [])
   51.23 @@ -264,9 +264,13 @@
   51.24    (Scan.this_string "inference" |-- $$ "(" |-- scan_general_id
   51.25     --| skip_term --| $$ "," --| skip_term --| $$ "," --| $$ "["
   51.26     -- parse_dependencies --| $$ "]" --| $$ ")") x
   51.27 +and skip_introduced x =
   51.28 +  (Scan.this_string "introduced" |-- $$ "(" |-- skip_term --| $$ ")") x
   51.29  and parse_source x =
   51.30    (parse_file_source >> File_Source
   51.31     || parse_inference_source >> Inference_Source
   51.32 +   || skip_introduced >> K dummy_inference (* for Vampire *)
   51.33 +   || scan_nat >> (fn s => Inference_Source ("", [s])) (* for E *)
   51.34     || skip_term >> K dummy_inference) x
   51.35  
   51.36  fun list_app (f, args) =
   51.37 @@ -508,7 +512,7 @@
   51.38    | atp_proof_from_tstplike_proof problem tstp =
   51.39      tstp ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
   51.40      |> parse_proof problem
   51.41 -    |> sort (step_name_ord o pairself step_name) (* FIXME: needed? *)
   51.42 +    |> sort (step_name_ord o pairself step_name)
   51.43  
   51.44  fun clean_up_dependencies _ [] = []
   51.45    | clean_up_dependencies seen
    52.1 --- a/src/HOL/Tools/ATP/atp_proof_redirect.ML	Tue Feb 12 17:39:45 2013 +0100
    52.2 +++ b/src/HOL/Tools/ATP/atp_proof_redirect.ML	Wed Feb 13 11:46:48 2013 +0100
    52.3 @@ -39,7 +39,8 @@
    52.4    val string_of_ref_graph : ref_graph -> string
    52.5    val redirect_sequent : atom list -> atom -> ref_sequent -> direct_sequent
    52.6    val direct_graph : direct_sequent list -> direct_graph
    52.7 -  val redirect_graph : atom list -> atom list -> ref_graph -> direct_proof
    52.8 +  val redirect_graph :
    52.9 +    atom list -> atom list -> atom -> ref_graph -> direct_proof
   52.10    val succedent_of_cases : (clause * direct_inference list) list -> clause
   52.11    val string_of_direct_proof : direct_proof -> string
   52.12  end;
   52.13 @@ -147,17 +148,11 @@
   52.14    | zones_of n (bs :: bss) =
   52.15      (fold (subtract atom_eq) bss) bs :: zones_of (n - 1) (bss @ [bs])
   52.16  
   52.17 -fun redirect_graph axioms tainted ref_graph =
   52.18 +fun redirect_graph axioms tainted bot ref_graph =
   52.19    let
   52.20 -    val bot =
   52.21 -      case Atom_Graph.maximals ref_graph of
   52.22 -        [bot] => bot
   52.23 -      | bots => raise Fail ("malformed refutation graph with " ^
   52.24 -                            string_of_int (length bots) ^ " maximal nodes")
   52.25      val seqs =
   52.26        map (redirect_sequent tainted bot) (sequents_of_ref_graph ref_graph)
   52.27      val direct_graph = direct_graph seqs
   52.28 -
   52.29      fun redirect c proved seqs =
   52.30        if null seqs then
   52.31          []
    53.1 --- a/src/HOL/Tools/ATP/atp_systems.ML	Tue Feb 12 17:39:45 2013 +0100
    53.2 +++ b/src/HOL/Tools/ATP/atp_systems.ML	Wed Feb 13 11:46:48 2013 +0100
    53.3 @@ -12,7 +12,7 @@
    53.4    type formula_role = ATP_Problem.formula_role
    53.5    type failure = ATP_Proof.failure
    53.6  
    53.7 -  type slice_spec = int * atp_format * string * string * bool
    53.8 +  type slice_spec = (int * string) * atp_format * string * string * bool
    53.9    type atp_config =
   53.10      {exec : string list * string list,
   53.11       arguments :
   53.12 @@ -91,7 +91,7 @@
   53.13  val default_max_mono_iters = 3 (* FUDGE *)
   53.14  val default_max_new_mono_instances = 200 (* FUDGE *)
   53.15  
   53.16 -type slice_spec = int * atp_format * string * string * bool
   53.17 +type slice_spec = (int * string) * atp_format * string * string * bool
   53.18  
   53.19  type atp_config =
   53.20    {exec : string list * string list,
   53.21 @@ -107,18 +107,26 @@
   53.22     best_max_new_mono_instances : int}
   53.23  
   53.24  (* "best_slices" must be found empirically, taking a wholistic approach since
   53.25 -   the ATPs are run in parallel. The "real" component gives the faction of the
   53.26 -   time available given to the slice and should add up to 1.0. The "int"
   53.27 -   component indicates the preferred number of facts to pass; the first
   53.28 -   "string", the preferred type system (which should be sound or quasi-sound);
   53.29 -   the second "string", the preferred lambda translation scheme; the "bool",
   53.30 -   whether uncurried aliased should be generated; the third "string", extra
   53.31 -   information to the prover (e.g., SOS or no SOS).
   53.32 +   the ATPs are run in parallel. Each slice has the format
   53.33 +
   53.34 +     (time_frac, ((max_facts, fact_filter), format, type_enc,
   53.35 +                  lam_trans, uncurried_aliases), extra)
   53.36 +
   53.37 +   where
   53.38 +
   53.39 +     time_frac = faction of the time available given to the slice (which should
   53.40 +       add up to 1.0)
   53.41 +
   53.42 +     extra = extra information to the prover (e.g., SOS or no SOS).
   53.43  
   53.44     The last slice should be the most "normal" one, because it will get all the
   53.45     time available if the other slices fail early and also because it is used if
   53.46     slicing is disabled (e.g., by the minimizer). *)
   53.47  
   53.48 +val mepoN = "mepo"
   53.49 +val mashN = "mash"
   53.50 +val meshN = "mesh"
   53.51 +
   53.52  val known_perl_failures =
   53.53    [(CantConnect, "HTTP error"),
   53.54     (NoPerl, "env: perl"),
   53.55 @@ -209,7 +217,7 @@
   53.56     prem_role = Hypothesis,
   53.57     best_slices = fn _ =>
   53.58       (* FUDGE *)
   53.59 -     [(1.0, ((100, alt_ergo_tff1, "poly_native", liftingN, false), ""))],
   53.60 +     [(1.0, (((100, ""), alt_ergo_tff1, "poly_native", liftingN, false), ""))],
   53.61     best_max_mono_iters = default_max_mono_iters,
   53.62     best_max_new_mono_instances = default_max_new_mono_instances}
   53.63  
   53.64 @@ -328,11 +336,14 @@
   53.65       let val heuristic = effective_e_selection_heuristic ctxt in
   53.66         (* FUDGE *)
   53.67         if heuristic = e_smartN then
   53.68 -         [(0.333, ((500, FOF, "mono_tags??", combsN, false), e_fun_weightN)),
   53.69 -          (0.334, ((50, FOF, "mono_guards??", combsN, false), e_fun_weightN)),
   53.70 -          (0.333, ((1000, FOF, "mono_tags??", combsN, false), e_sym_offset_weightN))]
   53.71 +         [(0.15, (((128, meshN), FOF, "mono_tags??", combsN, false), e_fun_weightN)),
   53.72 +          (0.15, (((128, mashN), FOF, "mono_guards??", combsN, false), e_sym_offset_weightN)),
   53.73 +          (0.15, (((91, mepoN), FOF, "mono_tags??", combsN, false), e_autoN)),
   53.74 +          (0.15, (((64, mashN), FOF, "mono_guards??", combsN, false), e_fun_weightN)),
   53.75 +          (0.15, (((1000, meshN), FOF, "poly_guards??", combsN, false), e_sym_offset_weightN)),
   53.76 +          (0.25, (((256, mepoN), FOF, "mono_tags??", liftingN, false), e_fun_weightN))]
   53.77         else
   53.78 -         [(1.0, ((500, FOF, "mono_tags??", combsN, false), heuristic))]
   53.79 +         [(1.0, (((500, ""), FOF, "mono_tags??", combsN, false), heuristic))]
   53.80       end,
   53.81     best_max_mono_iters = default_max_mono_iters,
   53.82     best_max_new_mono_instances = default_max_new_mono_instances}
   53.83 @@ -351,8 +362,10 @@
   53.84     prem_role = Conjecture,
   53.85     best_slices =
   53.86       (* FUDGE *)
   53.87 -     K [(0.5, ((1000, FOF, "mono_tags??", combsN, false), "")),
   53.88 -        (0.5, ((1000, FOF, "mono_guards??", combsN, false), ""))],
   53.89 +     K [(0.25, (((500, meshN), FOF, "mono_guards??", combs_or_liftingN, false), "")),
   53.90 +        (0.25, (((150, meshN), FOF, "poly_tags??", combs_or_liftingN, false), "")),
   53.91 +        (0.25, (((50, meshN), FOF, "mono_tags??", combs_or_liftingN, false), "")),
   53.92 +        (0.25, (((1000, meshN), FOF, "poly_guards??", combsN, false), ""))],
   53.93     best_max_mono_iters = default_max_mono_iters,
   53.94     best_max_new_mono_instances = default_max_new_mono_instances}
   53.95  
   53.96 @@ -390,7 +403,7 @@
   53.97     prem_role = Hypothesis,
   53.98     best_slices =
   53.99       (* FUDGE *)
  53.100 -     K [(1.0, ((150, FOF, "mono_guards??", liftingN, false), ""))],
  53.101 +     K [(1.0, (((150, ""), FOF, "mono_guards??", liftingN, false), ""))],
  53.102     best_max_mono_iters = default_max_mono_iters,
  53.103     best_max_new_mono_instances = default_max_new_mono_instances}
  53.104  
  53.105 @@ -434,7 +447,7 @@
  53.106     prem_role = Hypothesis,
  53.107     best_slices =
  53.108       (* FUDGE *)
  53.109 -     K [(1.0, ((40, leo2_thf0, "mono_native_higher", keep_lamsN, false), ""))],
  53.110 +     K [(1.0, (((40, ""), leo2_thf0, "mono_native_higher", keep_lamsN, false), ""))],
  53.111     best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *),
  53.112     best_max_new_mono_instances = default_max_new_mono_instances div 2 (* FUDGE *)}
  53.113  
  53.114 @@ -456,7 +469,7 @@
  53.115     prem_role = Hypothesis,
  53.116     best_slices =
  53.117       (* FUDGE *)
  53.118 -     K [(1.0, ((60, satallax_thf0, "mono_native_higher", keep_lamsN, false), ""))],
  53.119 +     K [(1.0, (((60, ""), satallax_thf0, "mono_native_higher", keep_lamsN, false), ""))],
  53.120     best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *),
  53.121     best_max_new_mono_instances = default_max_new_mono_instances div 2 (* FUDGE *)}
  53.122  
  53.123 @@ -497,14 +510,14 @@
  53.124     prem_role = Conjecture,
  53.125     best_slices = fn ctxt =>
  53.126       (* FUDGE *)
  53.127 -     [(0.1667, ((150, DFG Monomorphic, "mono_native", combsN, true), "")),
  53.128 -      (0.1667, ((500, DFG Monomorphic, "mono_native", liftingN, true), spass_H2SOS)),
  53.129 -      (0.1666, ((50, DFG Monomorphic,  "mono_native", liftingN, true), spass_H2LR0LT0)),
  53.130 -      (0.1000, ((250, DFG Monomorphic, "mono_native", combsN, true), spass_H2NuVS0)),
  53.131 -      (0.1000, ((1000, DFG Monomorphic, "mono_native", liftingN, true), spass_H1SOS)),
  53.132 -      (0.1000, ((150, DFG Monomorphic, "poly_guards??", liftingN, false), spass_H2NuVS0Red2)),
  53.133 -      (0.1000, ((300, DFG Monomorphic, "mono_native", combsN, true), spass_H2SOS)),
  53.134 -      (0.1000, ((100, DFG Monomorphic, "mono_native", combs_and_liftingN, true), spass_H2))]
  53.135 +     [(0.1667, (((150, meshN), DFG Monomorphic, "mono_native", combsN, true), "")),
  53.136 +      (0.1667, (((500, meshN), DFG Monomorphic, "mono_native", liftingN, true), spass_H2SOS)),
  53.137 +      (0.1666, (((50, meshN), DFG Monomorphic,  "mono_native", liftingN, true), spass_H2LR0LT0)),
  53.138 +      (0.1000, (((250, meshN), DFG Monomorphic, "mono_native", combsN, true), spass_H2NuVS0)),
  53.139 +      (0.1000, (((1000, mepoN), DFG Monomorphic, "mono_native", liftingN, true), spass_H1SOS)),
  53.140 +      (0.1000, (((150, meshN), DFG Monomorphic, "poly_guards??", liftingN, false), spass_H2NuVS0Red2)),
  53.141 +      (0.1000, (((300, meshN), DFG Monomorphic, "mono_native", combsN, true), spass_H2SOS)),
  53.142 +      (0.1000, (((100, meshN), DFG Monomorphic, "mono_native", combs_and_liftingN, true), spass_H2))]
  53.143       |> (case Config.get ctxt spass_extra_options of
  53.144             "" => I
  53.145           | opts => map (apsnd (apsnd (K opts)))),
  53.146 @@ -551,13 +564,13 @@
  53.147     best_slices = fn ctxt =>
  53.148       (* FUDGE *)
  53.149       (if is_vampire_beyond_1_8 () then
  53.150 -        [(0.333, ((500, vampire_tff0, "mono_guards??", combs_or_liftingN, false), sosN)),
  53.151 -         (0.333, ((150, vampire_tff0, "poly_tags??", combs_or_liftingN, false), sosN)),
  53.152 -         (0.334, ((50, vampire_tff0, "mono_native", combs_or_liftingN, false), no_sosN))]
  53.153 +        [(0.333, (((500, meshN), vampire_tff0, "mono_guards??", combs_or_liftingN, false), sosN)),
  53.154 +         (0.333, (((150, meshN), vampire_tff0, "poly_tags??", combs_or_liftingN, false), sosN)),
  53.155 +         (0.334, (((50, meshN), vampire_tff0, "mono_native", combs_or_liftingN, false), no_sosN))]
  53.156        else
  53.157 -        [(0.333, ((150, FOF, "poly_guards??", combs_or_liftingN, false), sosN)),
  53.158 -         (0.333, ((500, FOF, "mono_tags??", combs_or_liftingN, false), sosN)),
  53.159 -         (0.334, ((50, FOF, "mono_guards??", combs_or_liftingN, false), no_sosN))])
  53.160 +        [(0.333, (((150, meshN), FOF, "poly_guards??", combs_or_liftingN, false), sosN)),
  53.161 +         (0.333, (((500, meshN), FOF, "mono_tags??", combs_or_liftingN, false), sosN)),
  53.162 +         (0.334, (((50, meshN), FOF, "mono_guards??", combs_or_liftingN, false), no_sosN))])
  53.163       |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single
  53.164           else I),
  53.165     best_max_mono_iters = default_max_mono_iters,
  53.166 @@ -580,10 +593,10 @@
  53.167     prem_role = Hypothesis,
  53.168     best_slices =
  53.169       (* FUDGE *)
  53.170 -     K [(0.5, ((250, z3_tff0, "mono_native", combsN, false), "")),
  53.171 -        (0.25, ((125, z3_tff0, "mono_native", combsN, false), "")),
  53.172 -        (0.125, ((62, z3_tff0, "mono_native", combsN, false), "")),
  53.173 -        (0.125, ((31, z3_tff0, "mono_native", combsN, false), ""))],
  53.174 +     K [(0.5, (((250, meshN), z3_tff0, "mono_native", combsN, false), "")),
  53.175 +        (0.25, (((125, mepoN), z3_tff0, "mono_native", combsN, false), "")),
  53.176 +        (0.125, (((62, mashN), z3_tff0, "mono_native", combsN, false), "")),
  53.177 +        (0.125, (((31, meshN), z3_tff0, "mono_native", combsN, false), ""))],
  53.178     best_max_mono_iters = default_max_mono_iters,
  53.179     best_max_new_mono_instances = default_max_new_mono_instances}
  53.180  
  53.181 @@ -599,7 +612,7 @@
  53.182     known_failures = known_szs_status_failures,
  53.183     prem_role = Hypothesis,
  53.184     best_slices =
  53.185 -     K [(1.0, ((200, format, type_enc,
  53.186 +     K [(1.0, (((200, ""), format, type_enc,
  53.187                  if is_format_higher_order format then keep_lamsN
  53.188                  else combsN, false), ""))],
  53.189     best_max_mono_iters = default_max_mono_iters,
  53.190 @@ -692,32 +705,32 @@
  53.191  
  53.192  val remote_e =
  53.193    remotify_atp e "EP" ["1.0", "1.1", "1.2"]
  53.194 -      (K ((750, FOF, "mono_tags??", combsN, false), "") (* FUDGE *))
  53.195 +      (K (((750, ""), FOF, "mono_tags??", combsN, false), "") (* FUDGE *))
  53.196  val remote_iprover =
  53.197    remotify_atp iprover "iProver" []
  53.198 -      (K ((150, FOF, "mono_guards??", liftingN, false), "") (* FUDGE *))
  53.199 +      (K (((150, ""), FOF, "mono_guards??", liftingN, false), "") (* FUDGE *))
  53.200  val remote_iprover_eq =
  53.201    remotify_atp iprover_eq "iProver-Eq" []
  53.202 -      (K ((150, FOF, "mono_guards??", liftingN, false), "") (* FUDGE *))
  53.203 +      (K (((150, ""), FOF, "mono_guards??", liftingN, false), "") (* FUDGE *))
  53.204  val remote_leo2 =
  53.205    remotify_atp leo2 "LEO-II" ["1.2.8", "1.2.6"]
  53.206 -      (K ((100, leo2_thf0, "mono_native_higher", liftingN, false), "") (* FUDGE *))
  53.207 +      (K (((100, ""), leo2_thf0, "mono_native_higher", liftingN, false), "") (* FUDGE *))
  53.208  val remote_satallax =
  53.209    remotify_atp satallax "Satallax" ["2.3", "2.2", "2"]
  53.210 -      (K ((100, satallax_thf0, "mono_native_higher", keep_lamsN, false), "") (* FUDGE *))
  53.211 +      (K (((100, ""), satallax_thf0, "mono_native_higher", keep_lamsN, false), "") (* FUDGE *))
  53.212  val remote_vampire =
  53.213    remotify_atp vampire "Vampire" ["2.6", "2.5", "1.8"]
  53.214 -      (K ((250, vampire_tff0, "mono_native", combs_or_liftingN, false), "") (* FUDGE *))
  53.215 +      (K (((250, ""), vampire_tff0, "mono_native", combs_or_liftingN, false), "") (* FUDGE *))
  53.216  val remote_e_sine =
  53.217    remote_atp e_sineN "SInE" ["0.4"] [] (#known_failures e_config) Conjecture
  53.218 -      (K ((500, FOF, "mono_guards??", combsN, false), "") (* FUDGE *))
  53.219 +      (K (((500, ""), FOF, "mono_guards??", combsN, false), "") (* FUDGE *))
  53.220  val remote_snark =
  53.221    remote_atp snarkN "SNARK" ["20080805r029", "20080805r024"]
  53.222        [("refutation.", "end_refutation.")] [] Hypothesis
  53.223 -      (K ((100, explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *))
  53.224 +      (K (((100, ""), explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *))
  53.225  val remote_e_tofof =
  53.226    remote_atp e_tofofN "ToFoF" ["0.1"] [] (#known_failures e_config) Hypothesis
  53.227 -      (K ((150, explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *))
  53.228 +      (K (((150, ""), explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *))
  53.229  val remote_waldmeister =
  53.230    remote_atp waldmeisterN "Waldmeister" ["710"]
  53.231        [("#START OF PROOF", "Proved Goals:")]
  53.232 @@ -725,7 +738,7 @@
  53.233         (Inappropriate, "****  Unexpected end of file."),
  53.234         (Crashed, "Unrecoverable Segmentation Fault")]
  53.235        Hypothesis
  53.236 -      (K ((50, CNF_UEQ, "raw_mono_tags??", combsN, false), "") (* FUDGE *))
  53.237 +      (K (((50, ""), CNF_UEQ, "raw_mono_tags??", combsN, false), "") (* FUDGE *))
  53.238  
  53.239  (* Setup *)
  53.240  
    54.1 --- a/src/HOL/Tools/Sledgehammer/MaSh/src/sparseNaiveBayes.py	Tue Feb 12 17:39:45 2013 +0100
    54.2 +++ b/src/HOL/Tools/Sledgehammer/MaSh/src/sparseNaiveBayes.py	Wed Feb 13 11:46:48 2013 +0100
    54.3 @@ -34,15 +34,13 @@
    54.4          """
    54.5          for d in trainData:            
    54.6              dFeatureCounts = {}
    54.7 -            # Give p |- p a higher weight
    54.8 +            # Add p proves p with weight self.defaultPriorWeight
    54.9              if not self.defaultPriorWeight == 0:            
   54.10                  for f,_w in dicts.featureDict[d]:
   54.11                      dFeatureCounts[f] = self.defaultPriorWeight
   54.12              self.counts[d] = [self.defaultPriorWeight,dFeatureCounts]
   54.13  
   54.14 -        for key in dicts.dependenciesDict.keys():
   54.15 -            # Add p proves p
   54.16 -            keyDeps = [key]+dicts.dependenciesDict[key]
   54.17 +        for key,keyDeps in dicts.dependenciesDict.iteritems():
   54.18              for dep in keyDeps:
   54.19                  self.counts[dep][0] += 1
   54.20                  depFeatures = dicts.featureDict[key]
   54.21 @@ -105,7 +103,7 @@
   54.22              resultA = log(posA)
   54.23              for f,w in features:
   54.24                  # DEBUG
   54.25 -                #w = 1
   54.26 +                #w = 1.0
   54.27                  if f in fA:
   54.28                      if fWeightsA[f] == 0:
   54.29                          resultA += w*self.defVal
    55.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Tue Feb 12 17:39:45 2013 +0100
    55.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Wed Feb 13 11:46:48 2013 +0100
    55.3 @@ -10,7 +10,8 @@
    55.4    type status = ATP_Problem_Generate.status
    55.5    type stature = ATP_Problem_Generate.stature
    55.6  
    55.7 -  type fact = ((unit -> string) * stature) * thm
    55.8 +  type raw_fact = ((unit -> string) * stature) * thm
    55.9 +  type fact = (string * stature) * thm
   55.10  
   55.11    type fact_override =
   55.12      {add : (Facts.ref * Attrib.src list) list,
   55.13 @@ -33,12 +34,13 @@
   55.14      Proof.context -> term list -> term -> (((unit -> string) * 'a) * thm) list
   55.15      -> (((unit -> string) * 'a) * thm) list
   55.16    val maybe_filter_no_atps : Proof.context -> ('a * thm) list -> ('a * thm) list
   55.17 +  val fact_of_raw_fact : raw_fact -> fact
   55.18    val all_facts :
   55.19      Proof.context -> bool -> bool -> unit Symtab.table -> thm list -> thm list
   55.20 -    -> status Termtab.table -> fact list
   55.21 +    -> status Termtab.table -> raw_fact list
   55.22    val nearly_all_facts :
   55.23      Proof.context -> bool -> fact_override -> unit Symtab.table
   55.24 -    -> status Termtab.table -> thm list -> term list -> term -> fact list
   55.25 +    -> status Termtab.table -> thm list -> term list -> term -> raw_fact list
   55.26  end;
   55.27  
   55.28  structure Sledgehammer_Fact : SLEDGEHAMMER_FACT =
   55.29 @@ -49,7 +51,8 @@
   55.30  open Metis_Tactic
   55.31  open Sledgehammer_Util
   55.32  
   55.33 -type fact = ((unit -> string) * stature) * thm
   55.34 +type raw_fact = ((unit -> string) * stature) * thm
   55.35 +type fact = (string * stature) * thm
   55.36  
   55.37  type fact_override =
   55.38    {add : (Facts.ref * Attrib.src list) list,
   55.39 @@ -419,6 +422,8 @@
   55.40  fun maybe_filter_no_atps ctxt =
   55.41    not (Config.get ctxt ignore_no_atp) ? filter_out (No_ATPs.member ctxt o snd)
   55.42  
   55.43 +fun fact_of_raw_fact ((name, stature), th) = ((name (), stature), th)
   55.44 +
   55.45  fun all_facts ctxt generous ho_atp reserved add_ths chained css =
   55.46    let
   55.47      val thy = Proof_Context.theory_of ctxt
    56.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Tue Feb 12 17:39:45 2013 +0100
    56.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Wed Feb 13 11:46:48 2013 +0100
    56.3 @@ -7,6 +7,7 @@
    56.4  signature SLEDGEHAMMER_MASH =
    56.5  sig
    56.6    type stature = ATP_Problem_Generate.stature
    56.7 +  type raw_fact = Sledgehammer_Fact.raw_fact
    56.8    type fact = Sledgehammer_Fact.fact
    56.9    type fact_override = Sledgehammer_Fact.fact_override
   56.10    type params = Sledgehammer_Provers.params
   56.11 @@ -14,7 +15,10 @@
   56.12    type prover_result = Sledgehammer_Provers.prover_result
   56.13  
   56.14    val trace : bool Config.T
   56.15 +  val snow : bool Config.T
   56.16 +  val MePoN : string
   56.17    val MaShN : string
   56.18 +  val MeShN : string
   56.19    val mepoN : string
   56.20    val mashN : string
   56.21    val meshN : string
   56.22 @@ -62,7 +66,7 @@
   56.23    val isar_dependencies_of :
   56.24      string Symtab.table * string Symtab.table -> thm -> string list
   56.25    val prover_dependencies_of :
   56.26 -    Proof.context -> params -> string -> int -> fact list
   56.27 +    Proof.context -> params -> string -> int -> raw_fact list
   56.28      -> string Symtab.table * string Symtab.table -> thm
   56.29      -> bool * string list
   56.30    val weight_mepo_facts : 'a list -> ('a * real) list
   56.31 @@ -71,8 +75,8 @@
   56.32      int -> string list -> ('b * thm) list -> ('b * thm) list -> ('b * thm) list
   56.33      -> ('b * thm) list * ('b * thm) list
   56.34    val mash_suggested_facts :
   56.35 -    Proof.context -> params -> string -> int -> term list -> term -> fact list
   56.36 -    -> fact list * fact list
   56.37 +    Proof.context -> params -> string -> int -> term list -> term
   56.38 +    -> raw_fact list -> fact list * fact list
   56.39    val mash_learn_proof :
   56.40      Proof.context -> params -> string -> term -> ('a * thm) list -> thm list
   56.41      -> unit
   56.42 @@ -85,7 +89,7 @@
   56.43    val mash_weight : real
   56.44    val relevant_facts :
   56.45      Proof.context -> params -> string -> int -> fact_override -> term list
   56.46 -    -> term -> fact list -> fact list
   56.47 +    -> term -> raw_fact list -> (string * fact list) list
   56.48    val kill_learners : unit -> unit
   56.49    val running_learners : unit -> unit
   56.50  end;
   56.51 @@ -103,9 +107,14 @@
   56.52  
   56.53  val trace =
   56.54    Attrib.setup_config_bool @{binding sledgehammer_mash_trace} (K false)
   56.55 +val snow =
   56.56 +  Attrib.setup_config_bool @{binding sledgehammer_mash_snow} (K false)
   56.57 +
   56.58  fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
   56.59  
   56.60 +val MePoN = "MePo"
   56.61  val MaShN = "MaSh"
   56.62 +val MeShN = "MeSh"
   56.63  
   56.64  val mepoN = "mepo"
   56.65  val mashN = "mash"
   56.66 @@ -146,13 +155,20 @@
   56.67      val sugg_path = Path.explode sugg_file
   56.68      val cmd_file = temp_dir ^ "/mash_commands" ^ serial
   56.69      val cmd_path = Path.explode cmd_file
   56.70 +    val model_dir = File.shell_path (mash_model_dir ())
   56.71      val core =
   56.72        "--inputFile " ^ cmd_file ^ " --predictions " ^ sugg_file ^
   56.73        " --numberOfPredictions " ^ string_of_int max_suggs ^
   56.74        (* " --learnTheories" ^ *) (if save then " --saveModel" else "")
   56.75      val command =
   56.76 -      "\"$ISABELLE_SLEDGEHAMMER_MASH/src/mash.py\" --quiet --outputDir " ^
   56.77 -      File.shell_path (mash_model_dir ()) ^ " --log " ^ log_file ^ " " ^ core ^
   56.78 +      "cd \"$ISABELLE_SLEDGEHAMMER_MASH\"/src; " ^
   56.79 +      "./mash.py --quiet" ^
   56.80 +      (if Config.get ctxt snow then " --snow" else "") ^
   56.81 +      " --outputDir " ^ model_dir ^
   56.82 +      " --modelFile=" ^ model_dir ^ "/model.pickle" ^
   56.83 +      " --dictsFile=" ^ model_dir ^ "/dict.pickle" ^
   56.84 +      " --theoryFile=" ^ model_dir ^ "/theory.pickle" ^
   56.85 +      " --log " ^ log_file ^ " " ^ core ^
   56.86        " >& " ^ err_file
   56.87      fun run_on () =
   56.88        (Isabelle_System.bash command
   56.89 @@ -200,7 +216,7 @@
   56.90  
   56.91  fun encode_feature (name, weight) =
   56.92    encode_str name ^
   56.93 -  (if Real.== (weight, 1.0) then "" else "=" ^ Real.toString weight)
   56.94 +  (if Real.== (weight, 1.0) then "" else "=" ^ smart_string_of_real weight)
   56.95  
   56.96  val encode_features = map encode_feature #> space_implode " "
   56.97  
   56.98 @@ -320,7 +336,7 @@
   56.99  
  56.100  local
  56.101  
  56.102 -val version = "*** MaSh version 20130113a ***"
  56.103 +val version = "*** MaSh version 20130207a ***"
  56.104  
  56.105  exception Too_New of unit
  56.106  
  56.107 @@ -457,9 +473,8 @@
  56.108      map fst (take max_facts sels) @ take (max_facts - length sels) unks
  56.109    | mesh_facts fact_eq max_facts mess =
  56.110      let
  56.111 -      val mess =
  56.112 -        mess |> map (apsnd (apfst (normalize_scores max_facts #> `length)))
  56.113 -      fun score_in fact (global_weight, ((sel_len, sels), unks)) =
  56.114 +      val mess = mess |> map (apsnd (apfst (normalize_scores max_facts)))
  56.115 +      fun score_in fact (global_weight, (sels, unks)) =
  56.116          let
  56.117            fun score_at j =
  56.118              case try (nth sels) j of
  56.119 @@ -474,8 +489,7 @@
  56.120          end
  56.121        fun weight_of fact = mess |> map_filter (score_in fact) |> scaled_avg
  56.122        val facts =
  56.123 -        fold (union fact_eq o map fst o take max_facts o snd o fst o snd) mess
  56.124 -             []
  56.125 +        fold (union fact_eq o map fst o take max_facts o fst o snd) mess []
  56.126      in
  56.127        facts |> map (`weight_of) |> sort (int_ord o swap o pairself fst)
  56.128              |> map snd |> take max_facts
  56.129 @@ -526,8 +540,7 @@
  56.130    let
  56.131      val problem =
  56.132        {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
  56.133 -       facts = facts |> map (apfst (apfst (fn name => name ())))
  56.134 -                     |> map Untranslated_Fact}
  56.135 +       factss = [("", facts)]}
  56.136    in
  56.137      get_minimizing_isar_prover ctxt MaSh (K (K ())) prover params (K (K (K "")))
  56.138                                 problem
  56.139 @@ -659,9 +672,7 @@
  56.140    if length deps > max_dependencies then NONE else SOME deps
  56.141  
  56.142  fun isar_dependencies_of name_tabs th =
  56.143 -  let
  56.144 -    val deps = thms_in_proof (SOME name_tabs) th
  56.145 -  in
  56.146 +  let val deps = thms_in_proof (SOME name_tabs) th in
  56.147      if deps = [typedef_dep] orelse
  56.148         deps = [class_some_dep] orelse
  56.149         exists (member (op =) fundef_ths) deps orelse
  56.150 @@ -682,14 +693,13 @@
  56.151        val goal = goal_of_thm thy th
  56.152        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
  56.153        val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
  56.154 -      fun nickify ((_, stature), th) =
  56.155 -        ((fn () => nickname_of_thm th, stature), th)
  56.156 +      fun nickify ((_, stature), th) = ((nickname_of_thm th, stature), th)
  56.157        fun is_dep dep (_, th) = nickname_of_thm th = dep
  56.158        fun add_isar_dep facts dep accum =
  56.159          if exists (is_dep dep) accum then
  56.160            accum
  56.161          else case find_first (is_dep dep) facts of
  56.162 -          SOME ((name, status), th) => accum @ [((name, status), th)]
  56.163 +          SOME ((_, status), th) => accum @ [(("", status), th)]
  56.164          | NONE => accum (* shouldn't happen *)
  56.165        val facts =
  56.166          facts
  56.167 @@ -820,7 +830,10 @@
  56.168                                      (parents, feats, hints))
  56.169              end)
  56.170      val unknown = facts |> filter_out (is_fact_in_graph access_G snd)
  56.171 -  in find_mash_suggestions max_facts suggs facts chained unknown end
  56.172 +  in
  56.173 +    find_mash_suggestions max_facts suggs facts chained unknown
  56.174 +    |> pairself (map fact_of_raw_fact)
  56.175 +  end
  56.176  
  56.177  fun learn_wrt_access_graph ctxt (name, parents, feats, deps) (learns, graph) =
  56.178    let
  56.179 @@ -1094,9 +1107,11 @@
  56.180    if not (subset (op =) (the_list fact_filter, fact_filters)) then
  56.181      error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".")
  56.182    else if only then
  56.183 -    facts
  56.184 +    let val facts = facts |> map fact_of_raw_fact in
  56.185 +      [("", facts)]
  56.186 +    end
  56.187    else if max_facts <= 0 orelse null facts then
  56.188 -    []
  56.189 +    [("", [])]
  56.190    else
  56.191      let
  56.192        fun maybe_learn () =
  56.193 @@ -1112,21 +1127,22 @@
  56.194            end
  56.195          else
  56.196            ()
  56.197 -      val fact_filter =
  56.198 +      val effective_fact_filter =
  56.199          case fact_filter of
  56.200            SOME ff => (() |> ff <> mepoN ? maybe_learn; ff)
  56.201          | NONE =>
  56.202 -          if is_smt_prover ctxt prover then
  56.203 -            mepoN
  56.204 -          else if is_mash_enabled () then
  56.205 +          if is_mash_enabled () then
  56.206              (maybe_learn ();
  56.207               if mash_can_suggest_facts ctxt then meshN else mepoN)
  56.208            else
  56.209              mepoN
  56.210        val add_ths = Attrib.eval_thms ctxt add
  56.211 -      fun prepend_facts ths accepts =
  56.212 -        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
  56.213 -         (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
  56.214 +      fun in_add (_, th) = member Thm.eq_thm_prop add_ths th
  56.215 +      fun add_and_take accepts =
  56.216 +        (case add_ths of
  56.217 +           [] => accepts
  56.218 +         | _ => (facts |> filter in_add |> map fact_of_raw_fact) @
  56.219 +                (accepts |> filter_out in_add))
  56.220          |> take max_facts
  56.221        fun mepo () =
  56.222          mepo_suggested_facts ctxt params prover max_facts NONE hyp_ts concl_t
  56.223 @@ -1137,13 +1153,24 @@
  56.224              hyp_ts concl_t facts
  56.225          |>> weight_mash_facts
  56.226        val mess =
  56.227 -        [] |> (if fact_filter <> mashN then cons (mepo_weight, (mepo (), []))
  56.228 -               else I)
  56.229 -           |> (if fact_filter <> mepoN then cons (mash_weight, (mash ()))
  56.230 -               else I)
  56.231 +        (* the order is important for the "case" expression below *)
  56.232 +        [] |> (if effective_fact_filter <> mepoN then
  56.233 +                 cons (mash_weight, (mash ()))
  56.234 +               else
  56.235 +                 I)
  56.236 +           |> (if effective_fact_filter <> mashN then
  56.237 +                 cons (mepo_weight, (mepo (), []))
  56.238 +               else
  56.239 +                 I)
  56.240 +      val mesh =
  56.241 +        mesh_facts (Thm.eq_thm_prop o pairself snd) max_facts mess
  56.242 +        |> add_and_take
  56.243      in
  56.244 -      mesh_facts (Thm.eq_thm_prop o pairself snd) max_facts mess
  56.245 -      |> not (null add_ths) ? prepend_facts add_ths
  56.246 +      case (fact_filter, mess) of
  56.247 +        (NONE, [(_, (mepo, _)), (_, (mash, _))]) =>
  56.248 +        [(meshN, mesh), (mepoN, mepo |> map fst |> add_and_take),
  56.249 +         (mashN, mash |> map fst |> add_and_take)]
  56.250 +      | _ => [("", mesh)]
  56.251      end
  56.252  
  56.253  fun kill_learners () = Async_Manager.kill_threads MaShN "learner"
    57.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Tue Feb 12 17:39:45 2013 +0100
    57.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Wed Feb 13 11:46:48 2013 +0100
    57.3 @@ -8,6 +8,7 @@
    57.4  signature SLEDGEHAMMER_MEPO =
    57.5  sig
    57.6    type stature = ATP_Problem_Generate.stature
    57.7 +  type raw_fact = Sledgehammer_Fact.raw_fact
    57.8    type fact = Sledgehammer_Fact.fact
    57.9    type params = Sledgehammer_Provers.params
   57.10    type relevance_fudge = Sledgehammer_Provers.relevance_fudge
   57.11 @@ -20,7 +21,7 @@
   57.12      -> string list
   57.13    val mepo_suggested_facts :
   57.14      Proof.context -> params -> string -> int -> relevance_fudge option
   57.15 -    -> term list -> term -> fact list -> fact list
   57.16 +    -> term list -> term -> raw_fact list -> fact list
   57.17  end;
   57.18  
   57.19  structure Sledgehammer_MePo : SLEDGEHAMMER_MEPO =
   57.20 @@ -337,7 +338,7 @@
   57.21  
   57.22  fun take_most_relevant ctxt max_facts remaining_max
   57.23          ({max_imperfect, max_imperfect_exp, ...} : relevance_fudge)
   57.24 -        (candidates : ((fact * (string * ptype) list) * real) list) =
   57.25 +        (candidates : ((raw_fact * (string * ptype) list) * real) list) =
   57.26    let
   57.27      val max_imperfect =
   57.28        Real.ceil (Math.pow (max_imperfect,
   57.29 @@ -497,10 +498,10 @@
   57.30                      |> chop special_fact_index
   57.31          in bef @ add @ after end
   57.32      fun insert_special_facts accepts =
   57.33 -       (* FIXME: get rid of "ext" here once it is treated as a helper *)
   57.34 -       [] |> could_benefit_from_ext is_built_in_const accepts ? cons @{thm ext}
   57.35 -          |> add_set_const_thms accepts
   57.36 -          |> insert_into_facts accepts
   57.37 +      (* FIXME: get rid of "ext" here once it is treated as a helper *)
   57.38 +      [] |> could_benefit_from_ext is_built_in_const accepts ? cons @{thm ext}
   57.39 +         |> add_set_const_thms accepts
   57.40 +         |> insert_into_facts accepts
   57.41    in
   57.42      facts |> map_filter (pair_consts_fact thy is_built_in_const fudge)
   57.43            |> iter 0 max_facts thres0 goal_const_tab []
   57.44 @@ -533,6 +534,7 @@
   57.45         relevance_filter ctxt thres0 decay max_facts is_built_in_const fudge
   57.46             facts hyp_ts
   57.47             (concl_t |> theory_constify fudge (Context.theory_name thy)))
   57.48 +    |> map fact_of_raw_fact
   57.49    end
   57.50  
   57.51  end;
    58.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Tue Feb 12 17:39:45 2013 +0100
    58.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Wed Feb 13 11:46:48 2013 +0100
    58.3 @@ -68,8 +68,7 @@
    58.4             else
    58.5               "") ^ "...")
    58.6      val {goal, ...} = Proof.goal state
    58.7 -    val facts =
    58.8 -      facts |> maps (fn (n, ths) => ths |> map (Untranslated_Fact o pair n))
    58.9 +    val facts = facts |> maps (fn (n, ths) => map (pair n) ths)
   58.10      val params =
   58.11        {debug = debug, verbose = verbose, overlord = overlord, blocking = true,
   58.12         provers = provers, type_enc = type_enc, strict = strict,
   58.13 @@ -82,7 +81,7 @@
   58.14         preplay_timeout = preplay_timeout, expect = ""}
   58.15      val problem =
   58.16        {state = state, goal = goal, subgoal = i, subgoal_count = n,
   58.17 -       facts = facts}
   58.18 +       factss = [("", facts)]}
   58.19      val result as {outcome, used_facts, run_time, ...} =
   58.20        prover params (K (K (K ""))) problem
   58.21    in
   58.22 @@ -268,8 +267,8 @@
   58.23  
   58.24  fun maybe_minimize ctxt mode do_learn name
   58.25          (params as {verbose, isar_proofs, minimize, ...})
   58.26 -        ({state, subgoal, subgoal_count, facts, ...} : prover_problem)
   58.27 -        (result as {outcome, used_facts, run_time, preplay, message,
   58.28 +        ({state, subgoal, subgoal_count, ...} : prover_problem)
   58.29 +        (result as {outcome, used_facts, used_from, run_time, preplay, message,
   58.30                      message_tail} : prover_result) =
   58.31    if is_some outcome orelse null used_facts then
   58.32      result
   58.33 @@ -309,18 +308,17 @@
   58.34        val (used_facts, (preplay, message, _)) =
   58.35          if minimize then
   58.36            minimize_facts do_learn minimize_name params
   58.37 -                         (mode <> Normal orelse not verbose) subgoal
   58.38 -                         subgoal_count state
   58.39 -                         (filter_used_facts true used_facts
   58.40 -                              (map (apsnd single o untranslated_fact) facts))
   58.41 +              (mode <> Normal orelse not verbose) subgoal subgoal_count state
   58.42 +              (filter_used_facts true used_facts (map (apsnd single) used_from))
   58.43            |>> Option.map (map fst)
   58.44          else
   58.45            (SOME used_facts, (preplay, message, ""))
   58.46      in
   58.47        case used_facts of
   58.48          SOME used_facts =>
   58.49 -        {outcome = NONE, used_facts = used_facts, run_time = run_time,
   58.50 -         preplay = preplay, message = message, message_tail = message_tail}
   58.51 +        {outcome = NONE, used_facts = used_facts, used_from = used_from,
   58.52 +         run_time = run_time, preplay = preplay, message = message,
   58.53 +         message_tail = message_tail}
   58.54        | NONE => result
   58.55      end
   58.56  
    59.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Tue Feb 12 17:39:45 2013 +0100
    59.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Wed Feb 13 11:46:48 2013 +0100
    59.3 @@ -11,6 +11,7 @@
    59.4    type failure = ATP_Proof.failure
    59.5    type stature = ATP_Problem_Generate.stature
    59.6    type type_enc = ATP_Problem_Generate.type_enc
    59.7 +  type fact = Sledgehammer_Fact.fact
    59.8    type reconstructor = Sledgehammer_Reconstruct.reconstructor
    59.9    type play = Sledgehammer_Reconstruct.play
   59.10    type minimize_command = Sledgehammer_Reconstruct.minimize_command
   59.11 @@ -62,20 +63,17 @@
   59.12       threshold_divisor : real,
   59.13       ridiculous_threshold : real}
   59.14  
   59.15 -  datatype prover_fact =
   59.16 -    Untranslated_Fact of (string * stature) * thm |
   59.17 -    SMT_Weighted_Fact of (string * stature) * (int option * thm)
   59.18 -
   59.19    type prover_problem =
   59.20      {state : Proof.state,
   59.21       goal : thm,
   59.22       subgoal : int,
   59.23       subgoal_count : int,
   59.24 -     facts : prover_fact list}
   59.25 +     factss : (string * fact list) list}
   59.26  
   59.27    type prover_result =
   59.28      {outcome : failure option,
   59.29       used_facts : (string * stature) list,
   59.30 +     used_from : fact list,
   59.31       run_time : Time.time,
   59.32       preplay : play Lazy.lazy,
   59.33       message : play -> string,
   59.34 @@ -123,10 +121,6 @@
   59.35    val weight_smt_fact :
   59.36      Proof.context -> int -> ((string * stature) * thm) * int
   59.37      -> (string * stature) * (int option * thm)
   59.38 -  val untranslated_fact : prover_fact -> (string * stature) * thm
   59.39 -  val smt_weighted_fact :
   59.40 -    Proof.context -> int -> prover_fact * int
   59.41 -    -> (string * stature) * (int option * thm)
   59.42    val supported_provers : Proof.context -> unit
   59.43    val kill_provers : unit -> unit
   59.44    val running_provers : unit -> unit
   59.45 @@ -149,6 +143,7 @@
   59.46  open ATP_Proof_Reconstruct
   59.47  open Metis_Tactic
   59.48  open Sledgehammer_Util
   59.49 +open Sledgehammer_Fact
   59.50  open Sledgehammer_Reconstruct
   59.51  
   59.52  
   59.53 @@ -201,7 +196,7 @@
   59.54      if is_reconstructor name then
   59.55        reconstructor_default_max_facts
   59.56      else if is_atp thy name then
   59.57 -      fold (Integer.max o #1 o fst o snd o snd)
   59.58 +      fold (Integer.max o fst o #1 o fst o snd o snd)
   59.59             (get_slices slice (#best_slices (get_atp thy name ()) ctxt)) 0
   59.60      else (* is_smt_prover ctxt name *)
   59.61        SMT_Solver.default_max_relevant ctxt name
   59.62 @@ -355,20 +350,17 @@
   59.63     threshold_divisor : real,
   59.64     ridiculous_threshold : real}
   59.65  
   59.66 -datatype prover_fact =
   59.67 -  Untranslated_Fact of (string * stature) * thm |
   59.68 -  SMT_Weighted_Fact of (string * stature) * (int option * thm)
   59.69 -
   59.70  type prover_problem =
   59.71    {state : Proof.state,
   59.72     goal : thm,
   59.73     subgoal : int,
   59.74     subgoal_count : int,
   59.75 -   facts : prover_fact list}
   59.76 +   factss : (string * fact list) list}
   59.77  
   59.78  type prover_result =
   59.79    {outcome : failure option,
   59.80     used_facts : (string * stature) list,
   59.81 +   used_from : fact list,
   59.82     run_time : Time.time,
   59.83     preplay : play Lazy.lazy,
   59.84     message : play -> string,
   59.85 @@ -430,12 +422,6 @@
   59.86      (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
   59.87    end
   59.88  
   59.89 -fun untranslated_fact (Untranslated_Fact p) = p
   59.90 -  | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th)
   59.91 -fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p
   59.92 -  | smt_weighted_fact ctxt num_facts (fact, j) =
   59.93 -    (untranslated_fact fact, j) |> weight_smt_fact ctxt num_facts
   59.94 -
   59.95  fun overlord_file_location_for_prover prover =
   59.96    (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
   59.97  
   59.98 @@ -544,6 +530,12 @@
   59.99     clearly inconsistent facts such as X = a | X = b, though it by no means
  59.100     guarantees soundness. *)
  59.101  
  59.102 +fun get_facts_for_filter _ [(_, facts)] = facts
  59.103 +  | get_facts_for_filter fact_filter factss =
  59.104 +    case AList.lookup (op =) factss fact_filter of
  59.105 +      SOME facts => facts
  59.106 +    | NONE => snd (hd factss)
  59.107 +
  59.108  (* Unwanted equalities are those between a (bound or schematic) variable that
  59.109     does not properly occur in the second operand. *)
  59.110  val is_exhaustive_finite =
  59.111 @@ -642,11 +634,11 @@
  59.112          ({exec, arguments, proof_delims, known_failures, prem_role, best_slices,
  59.113            best_max_mono_iters, best_max_new_mono_instances, ...} : atp_config)
  59.114          (params as {debug, verbose, overlord, type_enc, strict, lam_trans,
  59.115 -                    uncurried_aliases, max_facts, max_mono_iters,
  59.116 +                    uncurried_aliases, fact_filter, max_facts, max_mono_iters,
  59.117                      max_new_mono_instances, isar_proofs, isar_shrink,
  59.118                      slice, timeout, preplay_timeout, ...})
  59.119          minimize_command
  59.120 -        ({state, goal, subgoal, subgoal_count, facts, ...} : prover_problem) =
  59.121 +        ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
  59.122    let
  59.123      val thy = Proof.theory_of state
  59.124      val ctxt = Proof.context_of state
  59.125 @@ -723,10 +715,14 @@
  59.126            end
  59.127          fun run_slice time_left (cache_key, cache_value)
  59.128                  (slice, (time_frac,
  59.129 -                     (key as (best_max_facts, format, best_type_enc,
  59.130 -                              best_lam_trans, best_uncurried_aliases),
  59.131 +                     (key as ((best_max_facts, best_fact_filter), format,
  59.132 +                              best_type_enc, best_lam_trans,
  59.133 +                              best_uncurried_aliases),
  59.134                        extra))) =
  59.135            let
  59.136 +            val effective_fact_filter =
  59.137 +              fact_filter |> the_default best_fact_filter
  59.138 +            val facts = get_facts_for_filter effective_fact_filter factss
  59.139              val num_facts =
  59.140                length facts |> is_none max_facts ? Integer.min best_max_facts
  59.141              val soundness = if strict then Strict else Non_Strict
  59.142 @@ -773,7 +769,6 @@
  59.143                  cache_value
  59.144                else
  59.145                  facts
  59.146 -                |> map untranslated_fact
  59.147                  |> not sound
  59.148                     ? filter_out (is_dangerous_prop ctxt o prop_of o snd)
  59.149                  |> take num_facts
  59.150 @@ -798,7 +793,7 @@
  59.151                |> lines_for_atp_problem format ord ord_info
  59.152                |> cons ("% " ^ command ^ "\n")
  59.153                |> File.write_list prob_path
  59.154 -            val ((output, run_time), (atp_proof, outcome)) =
  59.155 +            val ((output, run_time), used_from, (atp_proof, outcome)) =
  59.156                time_limit generous_slice_timeout Isabelle_System.bash_output
  59.157                           command
  59.158                |>> (if overlord then
  59.159 @@ -807,14 +802,14 @@
  59.160                       I)
  59.161                |> fst |> split_time
  59.162                |> (fn accum as (output, _) =>
  59.163 -                     (accum,
  59.164 +                     (accum, facts,
  59.165                        extract_tstplike_proof_and_outcome verbose proof_delims
  59.166                                                           known_failures output
  59.167                        |>> atp_proof_from_tstplike_proof atp_problem
  59.168                        handle UNRECOGNIZED_ATP_PROOF () =>
  59.169                               ([], SOME ProofIncomplete)))
  59.170                handle TimeLimit.TimeOut =>
  59.171 -                     (("", the slice_timeout), ([], SOME TimedOut))
  59.172 +                     (("", the slice_timeout), [], ([], SOME TimedOut))
  59.173              val outcome =
  59.174                case outcome of
  59.175                  NONE =>
  59.176 @@ -830,10 +825,12 @@
  59.177                     end
  59.178                   | NONE => NONE)
  59.179                | _ => outcome
  59.180 -          in ((SOME key, value), (output, run_time, atp_proof, outcome)) end
  59.181 +          in
  59.182 +            ((SOME key, value), (output, run_time, facts, atp_proof, outcome))
  59.183 +          end
  59.184          val timer = Timer.startRealTimer ()
  59.185          fun maybe_run_slice slice
  59.186 -                (result as (cache, (_, run_time0, _, SOME _))) =
  59.187 +                (result as (cache, (_, run_time0, _, _, SOME _))) =
  59.188              let
  59.189                val time_left =
  59.190                  Option.map
  59.191 @@ -845,25 +842,26 @@
  59.192                  result
  59.193                else
  59.194                  run_slice time_left cache slice
  59.195 -                |> (fn (cache, (output, run_time, atp_proof, outcome)) =>
  59.196 -                       (cache, (output, Time.+ (run_time0, run_time),
  59.197 +                |> (fn (cache, (output, run_time, used_from, atp_proof,
  59.198 +                                outcome)) =>
  59.199 +                       (cache, (output, Time.+ (run_time0, run_time), used_from,
  59.200                                  atp_proof, outcome)))
  59.201              end
  59.202            | maybe_run_slice _ result = result
  59.203        in
  59.204          ((NONE, ([], Symtab.empty, Vector.fromList [], [], Symtab.empty)),
  59.205 -         ("", Time.zeroTime, [], SOME InternalError))
  59.206 +         ("", Time.zeroTime, [], [], SOME InternalError))
  59.207          |> fold maybe_run_slice actual_slices
  59.208        end
  59.209      (* If the problem file has not been exported, remove it; otherwise, export
  59.210         the proof file too. *)
  59.211      fun clean_up () =
  59.212        if dest_dir = "" then (try File.rm prob_path; ()) else ()
  59.213 -    fun export (_, (output, _, _, _)) =
  59.214 +    fun export (_, (output, _, _, _, _)) =
  59.215        if dest_dir = "" then ()
  59.216        else File.write (Path.explode (Path.implode prob_path ^ "_proof")) output
  59.217      val ((_, (_, pool, fact_names, _, sym_tab)),
  59.218 -         (output, run_time, atp_proof, outcome)) =
  59.219 +         (output, run_time, used_from, atp_proof, outcome)) =
  59.220        with_cleanup clean_up run () |> tap export
  59.221      val important_message =
  59.222        if mode = Normal andalso
  59.223 @@ -887,8 +885,7 @@
  59.224             Lazy.lazy (fn () =>
  59.225               let
  59.226                 val used_pairs =
  59.227 -                 facts |> map untranslated_fact
  59.228 -                       |> filter_used_facts false used_facts
  59.229 +                 used_from |> filter_used_facts false used_facts
  59.230               in
  59.231                 play_one_line_proof mode debug verbose preplay_timeout
  59.232                     used_pairs state subgoal (hd reconstrs) reconstrs
  59.233 @@ -926,10 +923,15 @@
  59.234          ([], Lazy.value (Failed_to_Play plain_metis),
  59.235           fn _ => string_for_failure failure, "")
  59.236    in
  59.237 -    {outcome = outcome, used_facts = used_facts, run_time = run_time,
  59.238 -     preplay = preplay, message = message, message_tail = message_tail}
  59.239 +    {outcome = outcome, used_facts = used_facts, used_from = used_from,
  59.240 +     run_time = run_time, preplay = preplay, message = message,
  59.241 +     message_tail = message_tail}
  59.242    end
  59.243  
  59.244 +fun rotate_one (x :: xs) = xs @ [x]
  59.245 +
  59.246 +fun app_hd f (x :: xs) = f x :: xs
  59.247 +
  59.248  (* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
  59.249     these are sorted out properly in the SMT module, we have to interpret these
  59.250     ourselves. *)
  59.251 @@ -962,11 +964,12 @@
  59.252  val smt_max_slices =
  59.253    Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
  59.254  val smt_slice_fact_frac =
  59.255 -  Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5)
  59.256 +  Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac}
  59.257 +                           (K 0.667)
  59.258  val smt_slice_time_frac =
  59.259 -  Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5)
  59.260 +  Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.333)
  59.261  val smt_slice_min_secs =
  59.262 -  Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 5)
  59.263 +  Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 3)
  59.264  
  59.265  fun smt_filter_loop name
  59.266                      ({debug, verbose, overlord, max_mono_iters,
  59.267 @@ -987,7 +990,8 @@
  59.268      val ctxt = Proof.context_of state |> repair_context
  59.269      val state = state |> Proof.map_context (K ctxt)
  59.270      val max_slices = if slice then Config.get ctxt smt_max_slices else 1
  59.271 -    fun do_slice timeout slice outcome0 time_so_far facts =
  59.272 +    fun do_slice timeout slice outcome0 time_so_far
  59.273 +                 (weighted_factss as (fact_filter, weighted_facts) :: _) =
  59.274        let
  59.275          val timer = Timer.startRealTimer ()
  59.276          val state =
  59.277 @@ -1006,7 +1010,7 @@
  59.278              end
  59.279            else
  59.280              timeout
  59.281 -        val num_facts = length facts
  59.282 +        val num_facts = length weighted_facts
  59.283          val _ =
  59.284            if debug then
  59.285              quote name ^ " slice " ^ string_of_int slice ^ " with " ^
  59.286 @@ -1022,7 +1026,8 @@
  59.287            if debug then Output.urgent_message "Invoking SMT solver..." else ()
  59.288          val state_facts = these (try (#facts o Proof.goal) state)
  59.289          val (outcome, used_facts) =
  59.290 -          SMT_Solver.smt_filter_preprocess ctxt state_facts goal facts i
  59.291 +          SMT_Solver.smt_filter_preprocess ctxt state_facts goal weighted_facts
  59.292 +              i
  59.293            |> SMT_Solver.smt_filter_apply (slice_timeout |> the_default one_day)
  59.294            |> (fn {outcome, used_facts} => (outcome, used_facts))
  59.295            handle exn => if Exn.is_interrupt exn then
  59.296 @@ -1053,36 +1058,48 @@
  59.297              val new_num_facts =
  59.298                Real.ceil (Config.get ctxt smt_slice_fact_frac
  59.299                           * Real.fromInt num_facts)
  59.300 +            val weighted_factss as (new_fact_filter, _) :: _ =
  59.301 +              weighted_factss
  59.302 +              |> rotate_one
  59.303 +              |> app_hd (apsnd (take new_num_facts))
  59.304 +            val show_filter = fact_filter <> new_fact_filter
  59.305 +            fun num_of_facts fact_filter num_facts =
  59.306 +              string_of_int num_facts ^
  59.307 +              (if show_filter then " " ^ quote fact_filter else "") ^
  59.308 +              " fact" ^ plural_s num_facts
  59.309              val _ =
  59.310                if verbose andalso is_some outcome then
  59.311 -                quote name ^ " invoked with " ^ string_of_int num_facts ^
  59.312 -                " fact" ^ plural_s num_facts ^ ": " ^
  59.313 +                quote name ^ " invoked with " ^
  59.314 +                num_of_facts fact_filter num_facts ^ ": " ^
  59.315                  string_for_failure (failure_from_smt_failure (the outcome)) ^
  59.316 -                " Retrying with " ^ string_of_int new_num_facts ^ " fact" ^
  59.317 -                plural_s new_num_facts ^ "..."
  59.318 +                " Retrying with " ^ num_of_facts new_fact_filter new_num_facts ^
  59.319 +                "..."
  59.320                  |> Output.urgent_message
  59.321                else
  59.322                  ()
  59.323            in
  59.324 -            facts |> take new_num_facts
  59.325 -                  |> do_slice timeout (slice + 1) outcome0 time_so_far
  59.326 +            do_slice timeout (slice + 1) outcome0 time_so_far weighted_factss
  59.327            end
  59.328          else
  59.329            {outcome = if is_none outcome then NONE else the outcome0,
  59.330 -           used_facts = used_facts, run_time = time_so_far}
  59.331 +           used_facts = used_facts, used_from = map (apsnd snd) weighted_facts,
  59.332 +           run_time = time_so_far}
  59.333        end
  59.334    in do_slice timeout 1 NONE Time.zeroTime end
  59.335  
  59.336  fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...})
  59.337          minimize_command
  59.338 -        ({state, goal, subgoal, subgoal_count, facts, ...} : prover_problem) =
  59.339 +        ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
  59.340    let
  59.341      val ctxt = Proof.context_of state
  59.342 -    val num_facts = length facts
  59.343 -    val facts = facts ~~ (0 upto num_facts - 1)
  59.344 -                |> map (smt_weighted_fact ctxt num_facts)
  59.345 -    val {outcome, used_facts = used_pairs, run_time} =
  59.346 -      smt_filter_loop name params state goal subgoal facts
  59.347 +    fun weight_facts facts =
  59.348 +      let val num_facts = length facts in
  59.349 +        facts ~~ (0 upto num_facts - 1)
  59.350 +        |> map (weight_smt_fact ctxt num_facts)
  59.351 +      end
  59.352 +    val weighted_factss = factss |> map (apsnd weight_facts)
  59.353 +    val {outcome, used_facts = used_pairs, used_from, run_time} =
  59.354 +      smt_filter_loop name params state goal subgoal weighted_factss
  59.355      val used_facts = used_pairs |> map fst
  59.356      val outcome = outcome |> Option.map failure_from_smt_failure
  59.357      val (preplay, message, message_tail) =
  59.358 @@ -1109,14 +1126,16 @@
  59.359          (Lazy.value (Failed_to_Play plain_metis),
  59.360           fn _ => string_for_failure failure, "")
  59.361    in
  59.362 -    {outcome = outcome, used_facts = used_facts, run_time = run_time,
  59.363 -     preplay = preplay, message = message, message_tail = message_tail}
  59.364 +    {outcome = outcome, used_facts = used_facts, used_from = used_from,
  59.365 +     run_time = run_time, preplay = preplay, message = message,
  59.366 +     message_tail = message_tail}
  59.367    end
  59.368  
  59.369  fun run_reconstructor mode name
  59.370          (params as {debug, verbose, timeout, type_enc, lam_trans, ...})
  59.371          minimize_command
  59.372 -        ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
  59.373 +        ({state, subgoal, subgoal_count, factss = (_, facts) :: _, ...}
  59.374 +         : prover_problem) =
  59.375    let
  59.376      val reconstr =
  59.377        if name = metisN then
  59.378 @@ -1126,15 +1145,14 @@
  59.379          SMT
  59.380        else
  59.381          raise Fail ("unknown reconstructor: " ^ quote name)
  59.382 -    val used_pairs = facts |> map untranslated_fact
  59.383 -    val used_facts = used_pairs |> map fst
  59.384 +    val used_facts = facts |> map fst
  59.385    in
  59.386      case play_one_line_proof (if mode = Minimize then Normal else mode) debug
  59.387 -                             verbose timeout used_pairs state subgoal reconstr
  59.388 +                             verbose timeout facts state subgoal reconstr
  59.389                               [reconstr] of
  59.390        play as Played (_, time) =>
  59.391 -      {outcome = NONE, used_facts = used_facts, run_time = time,
  59.392 -       preplay = Lazy.value play,
  59.393 +      {outcome = NONE, used_facts = used_facts, used_from = facts,
  59.394 +       run_time = time, preplay = Lazy.value play,
  59.395         message =
  59.396           fn play =>
  59.397              let
  59.398 @@ -1150,8 +1168,8 @@
  59.399        let
  59.400          val failure = case play of Failed_to_Play _ => GaveUp | _ => TimedOut
  59.401        in
  59.402 -        {outcome = SOME failure, used_facts = [], run_time = Time.zeroTime,
  59.403 -         preplay = Lazy.value play,
  59.404 +        {outcome = SOME failure, used_facts = [], used_from = [],
  59.405 +         run_time = Time.zeroTime, preplay = Lazy.value play,
  59.406           message = fn _ => string_for_failure failure, message_tail = ""}
  59.407        end
  59.408    end
    60.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Tue Feb 12 17:39:45 2013 +0100
    60.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Wed Feb 13 11:46:48 2013 +0100
    60.3 @@ -131,6 +131,7 @@
    60.4  
    60.5  val is_typed_helper_name =
    60.6    String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
    60.7 +
    60.8  fun is_typed_helper_used_in_atp_proof atp_proof =
    60.9    is_axiom_used_in_proof is_typed_helper_name atp_proof
   60.10  
   60.11 @@ -145,7 +146,7 @@
   60.12  
   60.13  fun ext_name ctxt =
   60.14    if Thm.eq_thm_prop (@{thm ext},
   60.15 -         singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
   60.16 +       singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
   60.17      isa_short_ext
   60.18    else
   60.19      isa_ext
   60.20 @@ -158,8 +159,8 @@
   60.21         insert (op =) (ext_name ctxt, (Global, General))
   60.22       else if rule = leo2_unfold_def_rule then
   60.23         (* LEO 1.3.3 does not record definitions properly, leading to missing
   60.24 -         dependencies in the TSTP proof. Remove the next line once this is
   60.25 -         fixed. *)
   60.26 +          dependencies in the TSTP proof. Remove the next line once this is
   60.27 +          fixed. *)
   60.28         add_non_rec_defs fact_names
   60.29       else if rule = satallax_coreN then
   60.30         (fn [] =>
   60.31 @@ -169,8 +170,7 @@
   60.32           | facts => facts)
   60.33       else
   60.34         I)
   60.35 -    #> (if null deps then union (op =) (resolve_fact fact_names ss)
   60.36 -        else I)
   60.37 +    #> (if null deps then union (op =) (resolve_fact fact_names ss) else I)
   60.38    | add_fact _ _ _ = I
   60.39  
   60.40  fun used_facts_in_atp_proof ctxt fact_names atp_proof =
   60.41 @@ -678,6 +678,7 @@
   60.42          val ref_graph = atp_proof |> map_filter dep_of_step |> make_ref_graph
   60.43          val axioms = axioms_of_ref_graph ref_graph conjs
   60.44          val tainted = tainted_atoms_of_ref_graph ref_graph conjs
   60.45 +        val bot = atp_proof |> List.last |> dep_of_step |> the |> snd
   60.46          val steps =
   60.47            Symtab.empty
   60.48            |> fold (fn Definition_Step _ => I (* FIXME *)
   60.49 @@ -743,7 +744,7 @@
   60.50            map (isar_step_of_direct_inf outer) infs
   60.51          val (isar_proof, (preplay_fail, preplay_time)) =
   60.52            ref_graph
   60.53 -          |> redirect_graph axioms tainted
   60.54 +          |> redirect_graph axioms tainted bot
   60.55            |> map (isar_step_of_direct_inf true)
   60.56            |> append assms
   60.57            |> (if not preplay andalso isar_shrink <= 1.0 then
    61.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Tue Feb 12 17:39:45 2013 +0100
    61.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Wed Feb 13 11:46:48 2013 +0100
    61.3 @@ -8,6 +8,7 @@
    61.4  
    61.5  signature SLEDGEHAMMER_RUN =
    61.6  sig
    61.7 +  type fact = Sledgehammer_Fact.fact
    61.8    type fact_override = Sledgehammer_Fact.fact_override
    61.9    type minimize_command = Sledgehammer_Reconstruct.minimize_command
   61.10    type mode = Sledgehammer_Provers.mode
   61.11 @@ -17,6 +18,7 @@
   61.12    val noneN : string
   61.13    val timeoutN : string
   61.14    val unknownN : string
   61.15 +  val string_of_factss : (string * fact list) list -> string
   61.16    val run_sledgehammer :
   61.17      params -> mode -> int -> fact_override
   61.18      -> ((string * string list) list -> string -> minimize_command)
   61.19 @@ -64,8 +66,8 @@
   61.20  
   61.21  fun launch_prover (params as {debug, verbose, blocking, max_facts, slice,
   61.22                                timeout, expect, ...})
   61.23 -                  mode minimize_command only learn
   61.24 -                  {state, goal, subgoal, subgoal_count, facts} name =
   61.25 +        mode minimize_command only learn
   61.26 +        {state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name =
   61.27    let
   61.28      val ctxt = Proof.context_of state
   61.29      val hard_timeout = time_mult 3.0 (timeout |> the_default one_day)
   61.30 @@ -79,14 +81,15 @@
   61.31      val problem =
   61.32        {state = state, goal = goal, subgoal = subgoal,
   61.33         subgoal_count = subgoal_count,
   61.34 -       facts = facts
   61.35 -               |> not (Sledgehammer_Provers.is_ho_atp ctxt name)
   61.36 -                  ? filter_out (curry (op =) Induction o snd o snd o fst
   61.37 -                                o untranslated_fact)
   61.38 -               |> take num_facts}
   61.39 -    fun print_used_facts used_facts =
   61.40 -      tag_list 1 facts
   61.41 -      |> map (fn (j, fact) => fact |> untranslated_fact |> apsnd (K j))
   61.42 +       factss =
   61.43 +         factss
   61.44 +         |> map (apsnd ((not (is_ho_atp ctxt name)
   61.45 +                         ? filter_out (fn ((_, (_, Induction)), _) => true
   61.46 +                                        | _ => false))
   61.47 +                        #> take num_facts))}
   61.48 +    fun print_used_facts used_facts used_from =
   61.49 +      tag_list 1 used_from
   61.50 +      |> map (fn (j, fact) => fact |> apsnd (K j))
   61.51        |> filter_used_facts false used_facts
   61.52        |> map (fn ((name, _), j) => name ^ "@" ^ string_of_int j)
   61.53        |> commas
   61.54 @@ -96,9 +99,10 @@
   61.55      fun really_go () =
   61.56        problem
   61.57        |> get_minimizing_isar_prover ctxt mode learn name params minimize_command
   61.58 -      |> verbose ? tap (fn {outcome = NONE, used_facts as _ :: _, ...} =>
   61.59 -                           print_used_facts used_facts
   61.60 -                         | _ => ())
   61.61 +      |> verbose
   61.62 +         ? tap (fn {outcome = NONE, used_facts as _ :: _, used_from, ...} =>
   61.63 +                   print_used_facts used_facts used_from
   61.64 +                 | _ => ())
   61.65        |> (fn {outcome, preplay, message, message_tail, ...} =>
   61.66               (if outcome = SOME ATP_Proof.TimedOut then timeoutN
   61.67                else if is_some outcome then noneN
   61.68 @@ -160,11 +164,24 @@
   61.69         (unknownN, state))
   61.70    end
   61.71  
   61.72 -fun class_of_smt_solver ctxt name =
   61.73 -  ctxt |> select_smt_solver name
   61.74 -       |> SMT_Config.solver_class_of |> SMT_Utils.string_of_class
   61.75 +val auto_try_max_facts_divisor = 2 (* FUDGE *)
   61.76  
   61.77 -val auto_try_max_facts_divisor = 2 (* FUDGE *)
   61.78 +fun eq_facts p = eq_list (op = o pairself fst) p
   61.79 +
   61.80 +fun string_of_facts facts =
   61.81 +  "Including " ^ string_of_int (length facts) ^
   61.82 +  " relevant fact" ^ plural_s (length facts) ^ ":\n" ^
   61.83 +  (facts |> map (fst o fst) |> space_implode " ") ^ "."
   61.84 +
   61.85 +fun string_of_factss factss =
   61.86 +  if forall (null o snd) factss then
   61.87 +    "Found no relevant facts."
   61.88 +  else case factss of
   61.89 +    [(_, facts)] => string_of_facts facts
   61.90 +  | _ =>
   61.91 +    factss
   61.92 +    |> map (fn (filter, facts) => quote filter ^ ": " ^ string_of_facts facts)
   61.93 +    |> space_implode "\n\n"
   61.94  
   61.95  fun run_sledgehammer (params as {debug, verbose, blocking, provers, max_facts,
   61.96                                   slice, ...})
   61.97 @@ -182,7 +199,7 @@
   61.98        val ctxt = Proof.context_of state
   61.99        val {facts = chained, goal, ...} = Proof.goal state
  61.100        val (_, hyp_ts, concl_t) = strip_subgoal ctxt goal i
  61.101 -      val ho_atp = exists (Sledgehammer_Provers.is_ho_atp ctxt) provers
  61.102 +      val ho_atp = exists (is_ho_atp ctxt) provers
  61.103        val reserved = reserved_isar_keyword_table ()
  61.104        val css = clasimpset_rule_table_of ctxt
  61.105        val all_facts =
  61.106 @@ -196,15 +213,36 @@
  61.107        val (smts, (ueq_atps, full_atps)) =
  61.108          provers |> List.partition (is_smt_prover ctxt)
  61.109                  ||> List.partition (is_unit_equational_atp ctxt)
  61.110 -      fun launch_provers state get_facts translate provers =
  61.111 +      fun get_factss label is_appropriate_prop provers =
  61.112          let
  61.113 -          val facts = get_facts ()
  61.114 -          val num_facts = length facts
  61.115 -          val facts = facts ~~ (0 upto num_facts - 1)
  61.116 -                      |> map (translate num_facts)
  61.117 +          val max_max_facts =
  61.118 +            case max_facts of
  61.119 +              SOME n => n
  61.120 +            | NONE =>
  61.121 +              0 |> fold (Integer.max o default_max_facts_for_prover ctxt slice)
  61.122 +                        provers
  61.123 +                |> mode = Auto_Try ? (fn n => n div auto_try_max_facts_divisor)
  61.124 +        in
  61.125 +          all_facts
  61.126 +          |> (case is_appropriate_prop of
  61.127 +                SOME is_app => filter (is_app o prop_of o snd)
  61.128 +              | NONE => I)
  61.129 +          |> relevant_facts ctxt params (hd provers) max_max_facts fact_override
  61.130 +                            hyp_ts concl_t
  61.131 +          |> tap (fn factss =>
  61.132 +                     if verbose then
  61.133 +                       label ^ plural_s (length provers) ^ ": " ^
  61.134 +                       string_of_factss factss
  61.135 +                       |> print
  61.136 +                     else
  61.137 +                       ())
  61.138 +        end
  61.139 +      fun launch_provers state label is_appropriate_prop provers =
  61.140 +        let
  61.141 +          val factss = get_factss label is_appropriate_prop provers
  61.142            val problem =
  61.143              {state = state, goal = goal, subgoal = i, subgoal_count = n,
  61.144 -             facts = facts}
  61.145 +             factss = factss}
  61.146            fun learn prover =
  61.147              mash_learn_proof ctxt params prover (prop_of goal) all_facts
  61.148            val launch = launch_prover params mode minimize_command only learn
  61.149 @@ -220,36 +258,6 @@
  61.150              |> (if blocking then Par_List.map else map) (launch problem #> fst)
  61.151              |> max_outcome_code |> rpair state
  61.152          end
  61.153 -      fun get_facts label is_appropriate_prop provers =
  61.154 -        let
  61.155 -          val max_max_facts =
  61.156 -            case max_facts of
  61.157 -              SOME n => n
  61.158 -            | NONE =>
  61.159 -              0 |> fold (Integer.max o default_max_facts_for_prover ctxt slice)
  61.160 -                        provers
  61.161 -                |> mode = Auto_Try ? (fn n => n div auto_try_max_facts_divisor)
  61.162 -        in
  61.163 -          all_facts
  61.164 -          |> (case is_appropriate_prop of
  61.165 -                SOME is_app => filter (is_app o prop_of o snd)
  61.166 -              | NONE => I)
  61.167 -          |> relevant_facts ctxt params (hd provers) max_max_facts fact_override
  61.168 -                            hyp_ts concl_t
  61.169 -          |> map (apfst (apfst (fn name => name ())))
  61.170 -          |> tap (fn facts =>
  61.171 -                     if verbose then
  61.172 -                       label ^ plural_s (length provers) ^ ": " ^
  61.173 -                       (if null facts then
  61.174 -                          "Found no relevant facts."
  61.175 -                        else
  61.176 -                          "Including " ^ string_of_int (length facts) ^
  61.177 -                          " relevant fact" ^ plural_s (length facts) ^ ":\n" ^
  61.178 -                          (facts |> map (fst o fst) |> space_implode " ") ^ ".")
  61.179 -                       |> print
  61.180 -                     else
  61.181 -                       ())
  61.182 -        end
  61.183        fun launch_atps label is_appropriate_prop atps accum =
  61.184          if null atps then
  61.185            accum
  61.186 @@ -263,21 +271,9 @@
  61.187               ();
  61.188             accum)
  61.189          else
  61.190 -          launch_provers state (get_facts label is_appropriate_prop o K atps)
  61.191 -                         (K (Untranslated_Fact o fst)) atps
  61.192 +          launch_provers state label is_appropriate_prop atps
  61.193        fun launch_smts accum =
  61.194 -        if null smts then
  61.195 -          accum
  61.196 -        else
  61.197 -          let
  61.198 -            val facts = get_facts "SMT solver" NONE smts
  61.199 -            val weight = SMT_Weighted_Fact oo weight_smt_fact ctxt
  61.200 -          in
  61.201 -            smts |> map (`(class_of_smt_solver ctxt))
  61.202 -                 |> AList.group (op =)
  61.203 -                 |> map (snd #> launch_provers state (K facts) weight #> fst)
  61.204 -                 |> max_outcome_code |> rpair state
  61.205 -          end
  61.206 +        if null smts then accum else launch_provers state "SMT solver" NONE smts
  61.207        val launch_full_atps = launch_atps "ATP" NONE full_atps
  61.208        val launch_ueq_atps =
  61.209          launch_atps "Unit equational provers" (SOME is_unit_equality) ueq_atps
    62.1 --- a/src/Pure/Isar/attrib.ML	Tue Feb 12 17:39:45 2013 +0100
    62.2 +++ b/src/Pure/Isar/attrib.ML	Wed Feb 13 11:46:48 2013 +0100
    62.3 @@ -420,7 +420,7 @@
    62.4    setup (Binding.name "abs_def")
    62.5      (Scan.succeed (Thm.rule_attribute (fn context =>
    62.6        Local_Defs.meta_rewrite_rule (Context.proof_of context) #> Drule.abs_def)))
    62.7 -    "abstract over free variables of definitionial theorem"));
    62.8 +    "abstract over free variables of definitional theorem"));
    62.9  
   62.10  
   62.11