src/HOL/Real/ContNotDenum.thy
author wenzelm
Fri, 02 Jun 2006 23:22:29 +0200
changeset 19765 dfe940911617
parent 19023 5652a536b7e8
child 20635 e95db20977c5
permissions -rw-r--r--
misc cleanup;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19023
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     1
(*  Title       : HOL/Real/ContNonDenum
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     2
    ID          : $Id$
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     3
    Author      : Benjamin Porter, Monash University, NICTA, 2005
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     4
*)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     5
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     6
header {* Non-denumerability of the Continuum. *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     7
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     8
theory ContNotDenum
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
     9
imports RComplete
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    10
begin
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    11
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    12
section {* Abstract *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    13
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    14
text {* The following document presents a proof that the Continuum is
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    15
uncountable. It is formalised in the Isabelle/Isar theorem proving
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    16
system.
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    17
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    18
{\em Theorem:} The Continuum @{text "\<real>"} is not denumerable. In other
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    19
words, there does not exist a function f:@{text "\<nat>\<Rightarrow>\<real>"} such that f is
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    20
surjective.
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    21
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    22
{\em Outline:} An elegant informal proof of this result uses Cantor's
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    23
Diagonalisation argument. The proof presented here is not this
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    24
one. First we formalise some properties of closed intervals, then we
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    25
prove the Nested Interval Property. This property relies on the
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    26
completeness of the Real numbers and is the foundation for our
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    27
argument. Informally it states that an intersection of countable
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    28
closed intervals (where each successive interval is a subset of the
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    29
last) is non-empty. We then assume a surjective function f:@{text
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    30
"\<nat>\<Rightarrow>\<real>"} exists and find a real x such that x is not in the range of f
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    31
by generating a sequence of closed intervals then using the NIP. *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    32
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    33
section {* Closed Intervals *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    34
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    35
text {* This section formalises some properties of closed intervals. *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    36
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    37
subsection {* Definition *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    38
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    39
definition
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    40
  closed_int :: "real \<Rightarrow> real \<Rightarrow> real set"
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    41
  "closed_int x y = {z. x \<le> z \<and> z \<le> y}"
19023
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    42
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    43
subsection {* Properties *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    44
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    45
lemma closed_int_subset:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    46
  assumes xy: "x1 \<ge> x0" "y1 \<le> y0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    47
  shows "closed_int x1 y1 \<subseteq> closed_int x0 y0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    48
proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    49
  {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    50
    fix x::real
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    51
    assume "x \<in> closed_int x1 y1"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    52
    hence "x \<ge> x1 \<and> x \<le> y1" by (simp add: closed_int_def)
19023
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    53
    with xy have "x \<ge> x0 \<and> x \<le> y0" by auto
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    54
    hence "x \<in> closed_int x0 y0" by (simp add: closed_int_def)
19023
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    55
  }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    56
  thus ?thesis by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    57
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    58
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    59
lemma closed_int_least:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    60
  assumes a: "a \<le> b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    61
  shows "a \<in> closed_int a b \<and> (\<forall>x \<in> closed_int a b. a \<le> x)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    62
proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    63
  from a have "a\<in>{x. a\<le>x \<and> x\<le>b}" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    64
  thus "a \<in> closed_int a b" by (unfold closed_int_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    65
next
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    66
  have "\<forall>x\<in>{x. a\<le>x \<and> x\<le>b}. a\<le>x" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    67
  thus "\<forall>x \<in> closed_int a b. a \<le> x" by (unfold closed_int_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    68
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    69
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    70
lemma closed_int_most:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    71
  assumes a: "a \<le> b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    72
  shows "b \<in> closed_int a b \<and> (\<forall>x \<in> closed_int a b. x \<le> b)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    73
proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    74
  from a have "b\<in>{x. a\<le>x \<and> x\<le>b}" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    75
  thus "b \<in> closed_int a b" by (unfold closed_int_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    76
next
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    77
  have "\<forall>x\<in>{x. a\<le>x \<and> x\<le>b}. x\<le>b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    78
  thus "\<forall>x \<in> closed_int a b. x\<le>b" by (unfold closed_int_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    79
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    80
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    81
lemma closed_not_empty:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    82
  shows "a \<le> b \<Longrightarrow> \<exists>x. x \<in> closed_int a b" 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    83
  by (auto dest: closed_int_least)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    84
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    85
lemma closed_mem:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    86
  assumes "a \<le> c" and "c \<le> b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    87
  shows "c \<in> closed_int a b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    88
  by (unfold closed_int_def) auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    89
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    90
lemma closed_subset:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    91
  assumes ac: "a \<le> b"  "c \<le> d" 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    92
  assumes closed: "closed_int a b \<subseteq> closed_int c d"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    93
  shows "b \<ge> c"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    94
proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    95
  from closed have "\<forall>x\<in>closed_int a b. x\<in>closed_int c d" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    96
  hence "\<forall>x. a\<le>x \<and> x\<le>b \<longrightarrow> c\<le>x \<and> x\<le>d" by (unfold closed_int_def, auto)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    97
  with ac have "c\<le>b \<and> b\<le>d" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    98
  thus ?thesis by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
    99
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   100
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   101
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   102
section {* Nested Interval Property *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   103
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   104
theorem NIP:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   105
  fixes f::"nat \<Rightarrow> real set"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   106
  assumes subset: "\<forall>n. f (Suc n) \<subseteq> f n"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   107
  and closed: "\<forall>n. \<exists>a b. f n = closed_int a b \<and> a \<le> b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   108
  shows "(\<Inter>n. f n) \<noteq> {}"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   109
proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   110
  let ?g = "\<lambda>n. (SOME c. c\<in>(f n) \<and> (\<forall>x\<in>(f n). c \<le> x))"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   111
  have ne: "\<forall>n. \<exists>x. x\<in>(f n)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   112
  proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   113
    fix n
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   114
    from closed have "\<exists>a b. f n = closed_int a b \<and> a \<le> b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   115
    then obtain a and b where fn: "f n = closed_int a b \<and> a \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   116
    hence "a \<le> b" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   117
    with closed_not_empty have "\<exists>x. x\<in>closed_int a b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   118
    with fn show "\<exists>x. x\<in>(f n)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   119
  qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   120
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   121
  have gdef: "\<forall>n. (?g n)\<in>(f n) \<and> (\<forall>x\<in>(f n). (?g n)\<le>x)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   122
  proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   123
    fix n
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   124
    from closed have "\<exists>a b. f n = closed_int a b \<and> a \<le> b" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   125
    then obtain a and b where ff: "f n = closed_int a b" and "a \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   126
    hence "a \<le> b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   127
    hence "a\<in>closed_int a b \<and> (\<forall>x\<in>closed_int a b. a \<le> x)" by (rule closed_int_least)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   128
    with ff have "a\<in>(f n) \<and> (\<forall>x\<in>(f n). a \<le> x)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   129
    hence "\<exists>c. c\<in>(f n) \<and> (\<forall>x\<in>(f n). c \<le> x)" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   130
    thus "(?g n)\<in>(f n) \<and> (\<forall>x\<in>(f n). (?g n)\<le>x)" by (rule someI_ex)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   131
  qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   132
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   133
  -- "A denotes the set of all left-most points of all the intervals ..."
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   134
  moreover obtain A where Adef: "A = ?g ` \<nat>" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   135
  ultimately have "\<exists>x. x\<in>A"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   136
  proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   137
    have "(0::nat) \<in> \<nat>" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   138
    moreover have "?g 0 = ?g 0" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   139
    ultimately have "?g 0 \<in> ?g ` \<nat>" by (rule  rev_image_eqI)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   140
    with Adef have "?g 0 \<in> A" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   141
    thus ?thesis ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   142
  qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   143
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   144
  -- "Now show that A is bounded above ..."
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   145
  moreover have "\<exists>y. isUb (UNIV::real set) A y"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   146
  proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   147
    {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   148
      fix n
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   149
      from ne have ex: "\<exists>x. x\<in>(f n)" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   150
      from gdef have "(?g n)\<in>(f n) \<and> (\<forall>x\<in>(f n). (?g n)\<le>x)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   151
      moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   152
      from closed have "\<exists>a b. f n = closed_int a b \<and> a \<le> b" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   153
      then obtain a and b where "f n = closed_int a b \<and> a \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   154
      hence "b\<in>(f n) \<and> (\<forall>x\<in>(f n). x \<le> b)" using closed_int_most by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   155
      ultimately have "\<forall>x\<in>(f n). (?g n) \<le> b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   156
      with ex have "(?g n) \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   157
      hence "\<exists>b. (?g n) \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   158
    }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   159
    hence aux: "\<forall>n. \<exists>b. (?g n) \<le> b" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   160
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   161
    have fs: "\<forall>n::nat. f n \<subseteq> f 0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   162
    proof (rule allI, induct_tac n)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   163
      show "f 0 \<subseteq> f 0" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   164
    next
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   165
      fix n
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   166
      assume "f n \<subseteq> f 0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   167
      moreover from subset have "f (Suc n) \<subseteq> f n" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   168
      ultimately show "f (Suc n) \<subseteq> f 0" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   169
    qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   170
    have "\<forall>n. (?g n)\<in>(f 0)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   171
    proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   172
      fix n
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   173
      from gdef have "(?g n)\<in>(f n) \<and> (\<forall>x\<in>(f n). (?g n)\<le>x)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   174
      hence "?g n \<in> f n" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   175
      with fs show "?g n \<in> f 0" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   176
    qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   177
    moreover from closed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   178
      obtain a and b where "f 0 = closed_int a b" and alb: "a \<le> b" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   179
    ultimately have "\<forall>n. ?g n \<in> closed_int a b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   180
    with alb have "\<forall>n. ?g n \<le> b" using closed_int_most by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   181
    with Adef have "\<forall>y\<in>A. y\<le>b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   182
    hence "A *<= b" by (unfold setle_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   183
    moreover have "b \<in> (UNIV::real set)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   184
    ultimately have "A *<= b \<and> b \<in> (UNIV::real set)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   185
    hence "isUb (UNIV::real set) A b" by (unfold isUb_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   186
    thus ?thesis by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   187
  qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   188
  -- "by the Axiom Of Completeness, A has a least upper bound ..."
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   189
  ultimately have "\<exists>t. isLub UNIV A t" by (rule reals_complete)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   190
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   191
  -- "denote this least upper bound as t ..."
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   192
  then obtain t where tdef: "isLub UNIV A t" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   193
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   194
  -- "and finally show that this least upper bound is in all the intervals..."
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   195
  have "\<forall>n. t \<in> f n"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   196
  proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   197
    fix n::nat
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   198
    from closed obtain a and b where
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   199
      int: "f n = closed_int a b" and alb: "a \<le> b" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   200
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   201
    have "t \<ge> a"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   202
    proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   203
      have "a \<in> A"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   204
      proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   205
          (* by construction *)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   206
        from alb int have ain: "a\<in>f n \<and> (\<forall>x\<in>f n. a \<le> x)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   207
          using closed_int_least by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   208
        moreover have "\<forall>e. e\<in>f n \<and> (\<forall>x\<in>f n. e \<le> x) \<longrightarrow> e = a"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   209
        proof clarsimp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   210
          fix e
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   211
          assume ein: "e \<in> f n" and lt: "\<forall>x\<in>f n. e \<le> x"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   212
          from lt ain have aux: "\<forall>x\<in>f n. a \<le> x \<and> e \<le> x" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   213
  
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   214
          from ein aux have "a \<le> e \<and> e \<le> e" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   215
          moreover from ain aux have "a \<le> a \<and> e \<le> a" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   216
          ultimately show "e = a" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   217
        qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   218
        hence "\<And>e.  e\<in>f n \<and> (\<forall>x\<in>f n. e \<le> x) \<Longrightarrow> e = a" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   219
        ultimately have "(?g n) = a" by (rule some_equality)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   220
        moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   221
        {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   222
          have "n = of_nat n" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   223
          moreover have "of_nat n \<in> \<nat>" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   224
          ultimately have "n \<in> \<nat>"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   225
            apply -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   226
            apply (subst(asm) eq_sym_conv)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   227
            apply (erule subst)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   228
            .
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   229
        }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   230
        with Adef have "(?g n) \<in> A" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   231
        ultimately show ?thesis by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   232
      qed 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   233
      with tdef show "a \<le> t" by (rule isLubD2)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   234
    qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   235
    moreover have "t \<le> b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   236
    proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   237
      have "isUb UNIV A b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   238
      proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   239
        {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   240
          from alb int have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   241
            ain: "b\<in>f n \<and> (\<forall>x\<in>f n. x \<le> b)" using closed_int_most by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   242
          
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   243
          have subsetd: "\<forall>m. \<forall>n. f (n + m) \<subseteq> f n"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   244
          proof (rule allI, induct_tac m)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   245
            show "\<forall>n. f (n + 0) \<subseteq> f n" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   246
          next
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   247
            fix m n
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   248
            assume pp: "\<forall>p. f (p + n) \<subseteq> f p"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   249
            {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   250
              fix p
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   251
              from pp have "f (p + n) \<subseteq> f p" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   252
              moreover from subset have "f (Suc (p + n)) \<subseteq> f (p + n)" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   253
              hence "f (p + (Suc n)) \<subseteq> f (p + n)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   254
              ultimately have "f (p + (Suc n)) \<subseteq> f p" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   255
            }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   256
            thus "\<forall>p. f (p + Suc n) \<subseteq> f p" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   257
          qed 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   258
          have subsetm: "\<forall>\<alpha> \<beta>. \<alpha> \<ge> \<beta> \<longrightarrow> (f \<alpha>) \<subseteq> (f \<beta>)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   259
          proof ((rule allI)+, rule impI)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   260
            fix \<alpha>::nat and \<beta>::nat
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   261
            assume "\<beta> \<le> \<alpha>"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   262
            hence "\<exists>k. \<alpha> = \<beta> + k" by (simp only: le_iff_add)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   263
            then obtain k where "\<alpha> = \<beta> + k" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   264
            moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   265
            from subsetd have "f (\<beta> + k) \<subseteq> f \<beta>" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   266
            ultimately show "f \<alpha> \<subseteq> f \<beta>" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   267
          qed 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   268
          
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   269
          fix m   
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   270
          {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   271
            assume "m \<ge> n"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   272
            with subsetm have "f m \<subseteq> f n" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   273
            with ain have "\<forall>x\<in>f m. x \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   274
            moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   275
            from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   276
            ultimately have "?g m \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   277
          }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   278
          moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   279
          {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   280
            assume "\<not>(m \<ge> n)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   281
            hence "m < n" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   282
            with subsetm have sub: "(f n) \<subseteq> (f m)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   283
            from closed obtain ma and mb where
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   284
              "f m = closed_int ma mb \<and> ma \<le> mb" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   285
            hence one: "ma \<le> mb" and fm: "f m = closed_int ma mb" by auto 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   286
            from one alb sub fm int have "ma \<le> b" using closed_subset by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   287
            moreover have "(?g m) = ma"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   288
            proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   289
              from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" ..
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   290
              moreover from one have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   291
                "ma \<in> closed_int ma mb \<and> (\<forall>x\<in>closed_int ma mb. ma \<le> x)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   292
                by (rule closed_int_least)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   293
              with fm have "ma\<in>f m \<and> (\<forall>x\<in>f m. ma \<le> x)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   294
              ultimately have "ma \<le> ?g m \<and> ?g m \<le> ma" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   295
              thus "?g m = ma" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   296
            qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   297
            ultimately have "?g m \<le> b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   298
          } 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   299
          ultimately have "?g m \<le> b" by (rule case_split)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   300
        }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   301
        with Adef have "\<forall>y\<in>A. y\<le>b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   302
        hence "A *<= b" by (unfold setle_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   303
        moreover have "b \<in> (UNIV::real set)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   304
        ultimately have "A *<= b \<and> b \<in> (UNIV::real set)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   305
        thus "isUb (UNIV::real set) A b" by (unfold isUb_def)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   306
      qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   307
      with tdef show "t \<le> b" by (rule isLub_le_isUb)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   308
    qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   309
    ultimately have "t \<in> closed_int a b" by (rule closed_mem)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   310
    with int show "t \<in> f n" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   311
  qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   312
  hence "t \<in> (\<Inter>n. f n)" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   313
  thus ?thesis by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   314
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   315
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   316
section {* Generating the intervals *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   317
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   318
subsubsection {* Existence of non-singleton closed intervals *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   319
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   320
text {* This lemma asserts that given any non-singleton closed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   321
interval (a,b) and any element c, there exists a closed interval that
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   322
is a subset of (a,b) and that does not contain c and is a
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   323
non-singleton itself. *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   324
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   325
lemma closed_subset_ex:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   326
  fixes c::real
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   327
  assumes alb: "a < b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   328
  shows
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   329
    "\<exists>ka kb. ka < kb \<and> closed_int ka kb \<subseteq> closed_int a b \<and> c \<notin> (closed_int ka kb)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   330
proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   331
  {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   332
    assume clb: "c < b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   333
    {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   334
      assume cla: "c < a"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   335
      from alb cla clb have "c \<notin> closed_int a b" by (unfold closed_int_def, auto)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   336
      with alb have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   337
        "a < b \<and> closed_int a b \<subseteq> closed_int a b \<and> c \<notin> closed_int a b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   338
        by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   339
      hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   340
        "\<exists>ka kb. ka < kb \<and> closed_int ka kb \<subseteq> closed_int a b \<and> c \<notin> (closed_int ka kb)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   341
        by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   342
    }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   343
    moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   344
    {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   345
      assume ncla: "\<not>(c < a)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   346
      with clb have cdef: "a \<le> c \<and> c < b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   347
      obtain ka where kadef: "ka = (c + b)/2" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   348
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   349
      from kadef clb have kalb: "ka < b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   350
      moreover from kadef cdef have kagc: "ka > c" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   351
      ultimately have "c\<notin>(closed_int ka b)" by (unfold closed_int_def, auto)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   352
      moreover from cdef kagc have "ka \<ge> a" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   353
      hence "closed_int ka b \<subseteq> closed_int a b" by (unfold closed_int_def, auto)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   354
      ultimately have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   355
        "ka < b  \<and> closed_int ka b \<subseteq> closed_int a b \<and> c \<notin> closed_int ka b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   356
        using kalb by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   357
      hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   358
        "\<exists>ka kb. ka < kb \<and> closed_int ka kb \<subseteq> closed_int a b \<and> c \<notin> (closed_int ka kb)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   359
        by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   360
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   361
    }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   362
    ultimately have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   363
      "\<exists>ka kb. ka < kb \<and> closed_int ka kb \<subseteq> closed_int a b \<and> c \<notin> (closed_int ka kb)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   364
      by (rule case_split)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   365
  }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   366
  moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   367
  {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   368
    assume "\<not> (c < b)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   369
    hence cgeb: "c \<ge> b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   370
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   371
    obtain kb where kbdef: "kb = (a + b)/2" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   372
    with alb have kblb: "kb < b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   373
    with kbdef cgeb have "a < kb \<and> kb < c" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   374
    moreover hence "c \<notin> (closed_int a kb)" by (unfold closed_int_def, auto)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   375
    moreover from kblb have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   376
      "closed_int a kb \<subseteq> closed_int a b" by (unfold closed_int_def, auto)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   377
    ultimately have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   378
      "a < kb \<and>  closed_int a kb \<subseteq> closed_int a b \<and> c\<notin>closed_int a kb"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   379
      by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   380
    hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   381
      "\<exists>ka kb. ka < kb \<and> closed_int ka kb \<subseteq> closed_int a b \<and> c \<notin> (closed_int ka kb)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   382
      by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   383
  }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   384
  ultimately show ?thesis by (rule case_split)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   385
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   386
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   387
subsection {* newInt: Interval generation *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   388
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   389
text {* Given a function f:@{text "\<nat>\<Rightarrow>\<real>"}, newInt (Suc n) f returns a
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   390
closed interval such that @{text "newInt (Suc n) f \<subseteq> newInt n f"} and
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   391
does not contain @{text "f (Suc n)"}. With the base case defined such
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   392
that @{text "(f 0)\<notin>newInt 0 f"}. *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   393
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   394
subsubsection {* Definition *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   395
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   396
consts newInt :: "nat \<Rightarrow> (nat \<Rightarrow> real) \<Rightarrow> (real set)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   397
primrec
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   398
"newInt 0 f = closed_int (f 0 + 1) (f 0 + 2)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   399
"newInt (Suc n) f =
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   400
  (SOME e. (\<exists>e1 e2.
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   401
   e1 < e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   402
   e = closed_int e1 e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   403
   e \<subseteq> (newInt n f) \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   404
   (f (Suc n)) \<notin> e)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   405
  )"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   406
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   407
subsubsection {* Properties *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   408
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   409
text {* We now show that every application of newInt returns an
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   410
appropriate interval. *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   411
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   412
lemma newInt_ex:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   413
  "\<exists>a b. a < b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   414
   newInt (Suc n) f = closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   415
   newInt (Suc n) f \<subseteq> newInt n f \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   416
   f (Suc n) \<notin> newInt (Suc n) f"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   417
proof (induct n)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   418
  case 0
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   419
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   420
  let ?e = "SOME e. \<exists>e1 e2.
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   421
   e1 < e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   422
   e = closed_int e1 e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   423
   e \<subseteq> closed_int (f 0 + 1) (f 0 + 2) \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   424
   f (Suc 0) \<notin> e"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   425
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   426
  have "newInt (Suc 0) f = ?e" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   427
  moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   428
  have "f 0 + 1 < f 0 + 2" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   429
  with closed_subset_ex have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   430
    "\<exists>ka kb. ka < kb \<and> closed_int ka kb \<subseteq> closed_int (f 0 + 1) (f 0 + 2) \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   431
     f (Suc 0) \<notin> (closed_int ka kb)" .
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   432
  hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   433
    "\<exists>e. \<exists>ka kb. ka < kb \<and> e = closed_int ka kb \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   434
     e \<subseteq> closed_int (f 0 + 1) (f 0 + 2) \<and> f (Suc 0) \<notin> e" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   435
  hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   436
    "\<exists>ka kb. ka < kb \<and> ?e = closed_int ka kb \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   437
     ?e \<subseteq> closed_int (f 0 + 1) (f 0 + 2) \<and> f (Suc 0) \<notin> ?e"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   438
    by (rule someI_ex)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   439
  ultimately have "\<exists>e1 e2. e1 < e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   440
   newInt (Suc 0) f = closed_int e1 e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   441
   newInt (Suc 0) f \<subseteq> closed_int (f 0 + 1) (f 0 + 2) \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   442
   f (Suc 0) \<notin> newInt (Suc 0) f" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   443
  thus
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   444
    "\<exists>a b. a < b \<and> newInt (Suc 0) f = closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   445
     newInt (Suc 0) f \<subseteq> newInt 0 f \<and> f (Suc 0) \<notin> newInt (Suc 0) f"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   446
    by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   447
next
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   448
  case (Suc n)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   449
  hence "\<exists>a b.
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   450
   a < b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   451
   newInt (Suc n) f = closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   452
   newInt (Suc n) f \<subseteq> newInt n f \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   453
   f (Suc n) \<notin> newInt (Suc n) f" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   454
  then obtain a and b where ab: "a < b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   455
   newInt (Suc n) f = closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   456
   newInt (Suc n) f \<subseteq> newInt n f \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   457
   f (Suc n) \<notin> newInt (Suc n) f" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   458
  hence cab: "closed_int a b = newInt (Suc n) f" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   459
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   460
  let ?e = "SOME e. \<exists>e1 e2.
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   461
    e1 < e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   462
    e = closed_int e1 e2 \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   463
    e \<subseteq> closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   464
    f (Suc (Suc n)) \<notin> e"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   465
  from cab have ni: "newInt (Suc (Suc n)) f = ?e" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   466
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   467
  from ab have "a < b" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   468
  with closed_subset_ex have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   469
    "\<exists>ka kb. ka < kb \<and> closed_int ka kb \<subseteq> closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   470
     f (Suc (Suc n)) \<notin> closed_int ka kb" .
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   471
  hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   472
    "\<exists>e. \<exists>ka kb. ka < kb \<and> e = closed_int ka kb \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   473
     closed_int ka kb \<subseteq> closed_int a b \<and> f (Suc (Suc n)) \<notin> closed_int ka kb"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   474
    by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   475
  hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   476
    "\<exists>e.  \<exists>ka kb. ka < kb \<and> e = closed_int ka kb \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   477
     e \<subseteq> closed_int a b \<and> f (Suc (Suc n)) \<notin> e" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   478
  hence
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   479
    "\<exists>ka kb. ka < kb \<and> ?e = closed_int ka kb \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   480
     ?e \<subseteq> closed_int a b \<and> f (Suc (Suc n)) \<notin> ?e" by (rule someI_ex)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   481
  with ab ni show
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   482
    "\<exists>ka kb. ka < kb \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   483
     newInt (Suc (Suc n)) f = closed_int ka kb \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   484
     newInt (Suc (Suc n)) f \<subseteq> newInt (Suc n) f \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   485
     f (Suc (Suc n)) \<notin> newInt (Suc (Suc n)) f" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   486
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   487
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   488
lemma newInt_subset:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   489
  "newInt (Suc n) f \<subseteq> newInt n f"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   490
  using newInt_ex by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   491
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   492
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   493
text {* Another fundamental property is that no element in the range
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   494
of f is in the intersection of all closed intervals generated by
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   495
newInt. *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   496
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   497
lemma newInt_inter:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   498
  "\<forall>n. f n \<notin> (\<Inter>n. newInt n f)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   499
proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   500
  fix n::nat
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   501
  {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   502
    assume n0: "n = 0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   503
    moreover have "newInt 0 f = closed_int (f 0 + 1) (f 0 + 2)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   504
    ultimately have "f n \<notin> newInt n f" by (unfold closed_int_def, simp)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   505
  }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   506
  moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   507
  {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   508
    assume "\<not> n = 0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   509
    hence "n > 0" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   510
    then obtain m where ndef: "n = Suc m" by (auto simp add: gr0_conv_Suc)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   511
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   512
    from newInt_ex have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   513
      "\<exists>a b. a < b \<and> (newInt (Suc m) f) = closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   514
       newInt (Suc m) f \<subseteq> newInt m f \<and> f (Suc m) \<notin> newInt (Suc m) f" .
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   515
    then have "f (Suc m) \<notin> newInt (Suc m) f" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   516
    with ndef have "f n \<notin> newInt n f" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   517
  }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   518
  ultimately have "f n \<notin> newInt n f" by (rule case_split)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   519
  thus "f n \<notin> (\<Inter>n. newInt n f)" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   520
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   521
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   522
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   523
lemma newInt_notempty:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   524
  "(\<Inter>n. newInt n f) \<noteq> {}"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   525
proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   526
  let ?g = "\<lambda>n. newInt n f"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   527
  have "\<forall>n. ?g (Suc n) \<subseteq> ?g n"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   528
  proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   529
    fix n
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   530
    show "?g (Suc n) \<subseteq> ?g n" by (rule newInt_subset)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   531
  qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   532
  moreover have "\<forall>n. \<exists>a b. ?g n = closed_int a b \<and> a \<le> b"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   533
  proof
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   534
    fix n::nat
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   535
    {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   536
      assume "n = 0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   537
      then have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   538
        "?g n = closed_int (f 0 + 1) (f 0 + 2) \<and> (f 0 + 1 \<le> f 0 + 2)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   539
        by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   540
      hence "\<exists>a b. ?g n = closed_int a b \<and> a \<le> b" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   541
    }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   542
    moreover
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   543
    {
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   544
      assume "\<not> n = 0"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   545
      then have "n > 0" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   546
      then obtain m where nd: "n = Suc m" by (auto simp add: gr0_conv_Suc)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   547
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   548
      have
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   549
        "\<exists>a b. a < b \<and> (newInt (Suc m) f) = closed_int a b \<and>
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   550
        (newInt (Suc m) f) \<subseteq> (newInt m f) \<and> (f (Suc m)) \<notin> (newInt (Suc m) f)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   551
        by (rule newInt_ex)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   552
      then obtain a and b where
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   553
        "a < b \<and> (newInt (Suc m) f) = closed_int a b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   554
      with nd have "?g n = closed_int a b \<and> a \<le> b" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   555
      hence "\<exists>a b. ?g n = closed_int a b \<and> a \<le> b" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   556
    }
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   557
    ultimately show "\<exists>a b. ?g n = closed_int a b \<and> a \<le> b" by (rule case_split)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   558
  qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   559
  ultimately show ?thesis by (rule NIP)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   560
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   561
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   562
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   563
section {* Final Theorem *}
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   564
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   565
theorem real_non_denum:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   566
  shows "\<not> (\<exists>f::nat\<Rightarrow>real. surj f)"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   567
proof -- "by contradiction"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   568
  assume "\<exists>f::nat\<Rightarrow>real. surj f"
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   569
  then obtain f::"nat\<Rightarrow>real" where "surj f" by auto
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   570
  hence rangeF: "range f = UNIV" by (rule surj_range)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   571
  -- "We now produce a real number x that is not in the range of f, using the properties of newInt. "
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   572
  have "\<exists>x. x \<in> (\<Inter>n. newInt n f)" using newInt_notempty by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   573
  moreover have "\<forall>n. f n \<notin> (\<Inter>n. newInt n f)" by (rule newInt_inter)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   574
  ultimately obtain x where "x \<in> (\<Inter>n. newInt n f)" and "\<forall>n. f n \<noteq> x" by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   575
  moreover from rangeF have "x \<in> range f" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   576
  ultimately show False by blast
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   577
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents:
diff changeset
   578
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
   579
end