| 62479 |      1 | (*  Title:      HOL/Nonstandard_Analysis/Star.thy
 | 
|  |      2 |     Author:     Jacques D. Fleuriot
 | 
|  |      3 |     Copyright:  1998  University of Cambridge
 | 
| 27468 |      4 |     Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 64435 |      7 | section \<open>Star-Transforms in Non-Standard Analysis\<close>
 | 
| 27468 |      8 | 
 | 
|  |      9 | theory Star
 | 
| 64435 |     10 |   imports NSA
 | 
| 27468 |     11 | begin
 | 
|  |     12 | 
 | 
| 64435 |     13 | definition  \<comment> \<open>internal sets\<close>
 | 
| 81142 |     14 |   starset_n :: "(nat \<Rightarrow> 'a set) \<Rightarrow> 'a star set"
 | 
|  |     15 |     (\<open>(\<open>open_block notation=\<open>prefix starset_n\<close>\<close>*sn* _)\<close> [80] 80)
 | 
| 64435 |     16 |   where "*sn* As = Iset (star_n As)"
 | 
| 27468 |     17 | 
 | 
| 64435 |     18 | definition InternalSets :: "'a star set set"
 | 
|  |     19 |   where "InternalSets = {X. \<exists>As. X = *sn* As}"
 | 
| 27468 |     20 | 
 | 
| 64435 |     21 | definition  \<comment> \<open>nonstandard extension of function\<close>
 | 
|  |     22 |   is_starext :: "('a star \<Rightarrow> 'a star) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|  |     23 |   where "is_starext F f \<longleftrightarrow>
 | 
|  |     24 |     (\<forall>x y. \<exists>X \<in> Rep_star x. \<exists>Y \<in> Rep_star y. y = F x \<longleftrightarrow> eventually (\<lambda>n. Y n = f(X n)) \<U>)"
 | 
| 27468 |     25 | 
 | 
| 64435 |     26 | definition  \<comment> \<open>internal functions\<close>
 | 
| 81142 |     27 |   starfun_n :: "(nat \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a star \<Rightarrow> 'b star"
 | 
|  |     28 |     (\<open>(\<open>open_block notation=\<open>prefix starfun_n\<close>\<close>*fn* _)\<close> [80] 80)
 | 
| 64435 |     29 |   where "*fn* F = Ifun (star_n F)"
 | 
| 27468 |     30 | 
 | 
| 64435 |     31 | definition InternalFuns :: "('a star => 'b star) set"
 | 
|  |     32 |   where "InternalFuns = {X. \<exists>F. X = *fn* F}"
 | 
| 27468 |     33 | 
 | 
|  |     34 | 
 | 
| 64435 |     35 | subsection \<open>Preamble - Pulling \<open>\<exists>\<close> over \<open>\<forall>\<close>\<close>
 | 
| 27468 |     36 | 
 | 
| 64435 |     37 | text \<open>This proof does not need AC and was suggested by the
 | 
|  |     38 |    referee for the JCM Paper: let \<open>f x\<close> be least \<open>y\<close> such
 | 
|  |     39 |    that \<open>Q x y\<close>.\<close>
 | 
|  |     40 | lemma no_choice: "\<forall>x. \<exists>y. Q x y \<Longrightarrow> \<exists>f :: 'a \<Rightarrow> nat. \<forall>x. Q x (f x)"
 | 
|  |     41 |   by (rule exI [where x = "\<lambda>x. LEAST y. Q x y"]) (blast intro: LeastI)
 | 
| 27468 |     42 | 
 | 
| 64435 |     43 | 
 | 
|  |     44 | subsection \<open>Properties of the Star-transform Applied to Sets of Reals\<close>
 | 
| 27468 |     45 | 
 | 
| 64435 |     46 | lemma STAR_star_of_image_subset: "star_of ` A \<subseteq> *s* A"
 | 
|  |     47 |   by auto
 | 
| 27468 |     48 | 
 | 
| 64435 |     49 | lemma STAR_hypreal_of_real_Int: "*s* X \<inter> \<real> = hypreal_of_real ` X"
 | 
|  |     50 |   by (auto simp add: SReal_def)
 | 
| 27468 |     51 | 
 | 
| 64435 |     52 | lemma STAR_star_of_Int: "*s* X \<inter> Standard = star_of ` X"
 | 
|  |     53 |   by (auto simp add: Standard_def)
 | 
| 27468 |     54 | 
 | 
| 64435 |     55 | lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A \<Longrightarrow> \<forall>y \<in> A. x \<noteq> hypreal_of_real y"
 | 
|  |     56 |   by auto
 | 
| 27468 |     57 | 
 | 
| 64435 |     58 | lemma lemma_not_starA: "x \<notin> star_of ` A \<Longrightarrow> \<forall>y \<in> A. x \<noteq> star_of y"
 | 
|  |     59 |   by auto
 | 
| 27468 |     60 | 
 | 
| 64435 |     61 | lemma STAR_real_seq_to_hypreal: "\<forall>n. (X n) \<notin> M \<Longrightarrow> star_n X \<notin> *s* M"
 | 
|  |     62 |   by (simp add: starset_def star_of_def Iset_star_n FreeUltrafilterNat.proper)
 | 
| 27468 |     63 | 
 | 
|  |     64 | lemma STAR_singleton: "*s* {x} = {star_of x}"
 | 
| 64435 |     65 |   by simp
 | 
| 27468 |     66 | 
 | 
| 64435 |     67 | lemma STAR_not_mem: "x \<notin> F \<Longrightarrow> star_of x \<notin> *s* F"
 | 
|  |     68 |   by transfer
 | 
| 27468 |     69 | 
 | 
| 64435 |     70 | lemma STAR_subset_closed: "x \<in> *s* A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> x \<in> *s* B"
 | 
|  |     71 |   by (erule rev_subsetD) simp
 | 
| 27468 |     72 | 
 | 
| 64435 |     73 | text \<open>Nonstandard extension of a set (defined using a constant
 | 
|  |     74 |    sequence) as a special case of an internal set.\<close>
 | 
|  |     75 | lemma starset_n_starset: "\<forall>n. As n = A \<Longrightarrow> *sn* As = *s* A"
 | 
|  |     76 |   by (drule fun_eq_iff [THEN iffD2]) (simp add: starset_n_def starset_def star_of_def)
 | 
| 27468 |     77 | 
 | 
|  |     78 | 
 | 
| 64435 |     79 | subsection \<open>Theorems about nonstandard extensions of functions\<close>
 | 
| 27468 |     80 | 
 | 
| 64435 |     81 | text \<open>Nonstandard extension of a function (defined using a
 | 
|  |     82 |   constant sequence) as a special case of an internal function.\<close>
 | 
| 27468 |     83 | 
 | 
| 70218 |     84 | lemma starfun_n_starfun: "F = (\<lambda>n. f) \<Longrightarrow> *fn* F = *f* f"
 | 
|  |     85 |   by (simp add: starfun_n_def starfun_def star_of_def)
 | 
| 27468 |     86 | 
 | 
| 64435 |     87 | text \<open>Prove that \<open>abs\<close> for hypreal is a nonstandard extension of abs for real w/o
 | 
|  |     88 |   use of congruence property (proved after this for general
 | 
|  |     89 |   nonstandard extensions of real valued functions).
 | 
| 27468 |     90 | 
 | 
| 64435 |     91 |   Proof now Uses the ultrafilter tactic!\<close>
 | 
| 27468 |     92 | 
 | 
|  |     93 | lemma hrabs_is_starext_rabs: "is_starext abs abs"
 | 
| 70218 |     94 |   proof -
 | 
|  |     95 |   have "\<exists>f\<in>Rep_star (star_n h). \<exists>g\<in>Rep_star (star_n k). (star_n k = \<bar>star_n h\<bar>) = (\<forall>\<^sub>F n in \<U>. (g n::'a) = \<bar>f n\<bar>)"
 | 
|  |     96 |     for x y :: "'a star" and h k
 | 
|  |     97 |     by (metis (full_types) Rep_star_star_n star_n_abs star_n_eq_iff)
 | 
|  |     98 |   then show ?thesis
 | 
|  |     99 |     unfolding is_starext_def by (metis star_cases)
 | 
|  |    100 | qed
 | 
| 27468 |    101 | 
 | 
|  |    102 | 
 | 
| 64435 |    103 | text \<open>Nonstandard extension of functions.\<close>
 | 
|  |    104 | 
 | 
|  |    105 | lemma starfun: "( *f* f) (star_n X) = star_n (\<lambda>n. f (X n))"
 | 
|  |    106 |   by (rule starfun_star_n)
 | 
| 27468 |    107 | 
 | 
| 64435 |    108 | lemma starfun_if_eq: "\<And>w. w \<noteq> star_of x \<Longrightarrow> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w"
 | 
|  |    109 |   by transfer simp
 | 
| 27468 |    110 | 
 | 
| 64435 |    111 | text \<open>Multiplication: \<open>( *f) x ( *g) = *(f x g)\<close>\<close>
 | 
|  |    112 | lemma starfun_mult: "\<And>x. ( *f* f) x * ( *f* g) x = ( *f* (\<lambda>x. f x * g x)) x"
 | 
|  |    113 |   by transfer (rule refl)
 | 
| 27468 |    114 | declare starfun_mult [symmetric, simp]
 | 
|  |    115 | 
 | 
| 64435 |    116 | text \<open>Addition: \<open>( *f) + ( *g) = *(f + g)\<close>\<close>
 | 
|  |    117 | lemma starfun_add: "\<And>x. ( *f* f) x + ( *f* g) x = ( *f* (\<lambda>x. f x + g x)) x"
 | 
|  |    118 |   by transfer (rule refl)
 | 
| 27468 |    119 | declare starfun_add [symmetric, simp]
 | 
|  |    120 | 
 | 
| 64435 |    121 | text \<open>Subtraction: \<open>( *f) + -( *g) = *(f + -g)\<close>\<close>
 | 
|  |    122 | lemma starfun_minus: "\<And>x. - ( *f* f) x = ( *f* (\<lambda>x. - f x)) x"
 | 
|  |    123 |   by transfer (rule refl)
 | 
| 27468 |    124 | declare starfun_minus [symmetric, simp]
 | 
|  |    125 | 
 | 
|  |    126 | (*FIXME: delete*)
 | 
| 64435 |    127 | lemma starfun_add_minus: "\<And>x. ( *f* f) x + -( *f* g) x = ( *f* (\<lambda>x. f x + -g x)) x"
 | 
|  |    128 |   by transfer (rule refl)
 | 
| 27468 |    129 | declare starfun_add_minus [symmetric, simp]
 | 
|  |    130 | 
 | 
| 64435 |    131 | lemma starfun_diff: "\<And>x. ( *f* f) x  - ( *f* g) x = ( *f* (\<lambda>x. f x - g x)) x"
 | 
|  |    132 |   by transfer (rule refl)
 | 
| 27468 |    133 | declare starfun_diff [symmetric, simp]
 | 
|  |    134 | 
 | 
| 64435 |    135 | text \<open>Composition: \<open>( *f) \<circ> ( *g) = *(f \<circ> g)\<close>\<close>
 | 
|  |    136 | lemma starfun_o2: "(\<lambda>x. ( *f* f) (( *f* g) x)) = *f* (\<lambda>x. f (g x))"
 | 
|  |    137 |   by transfer (rule refl)
 | 
| 27468 |    138 | 
 | 
| 64435 |    139 | lemma starfun_o: "( *f* f) \<circ> ( *f* g) = ( *f* (f \<circ> g))"
 | 
|  |    140 |   by (transfer o_def) (rule refl)
 | 
| 27468 |    141 | 
 | 
| 64435 |    142 | text \<open>NS extension of constant function.\<close>
 | 
|  |    143 | lemma starfun_const_fun [simp]: "\<And>x. ( *f* (\<lambda>x. k)) x = star_of k"
 | 
|  |    144 |   by transfer (rule refl)
 | 
| 27468 |    145 | 
 | 
| 64435 |    146 | text \<open>The NS extension of the identity function.\<close>
 | 
|  |    147 | lemma starfun_Id [simp]: "\<And>x. ( *f* (\<lambda>x. x)) x = x"
 | 
|  |    148 |   by transfer (rule refl)
 | 
| 27468 |    149 | 
 | 
| 64435 |    150 | text \<open>The Star-function is a (nonstandard) extension of the function.\<close>
 | 
| 27468 |    151 | lemma is_starext_starfun: "is_starext ( *f* f) f"
 | 
| 70218 |    152 | proof -
 | 
|  |    153 |   have "\<exists>X\<in>Rep_star x. \<exists>Y\<in>Rep_star y. (y = (*f* f) x) = (\<forall>\<^sub>F n in \<U>. Y n = f (X n))"
 | 
|  |    154 |     for x y
 | 
|  |    155 |     by (metis (mono_tags) Rep_star_star_n star_cases star_n_eq_iff starfun_star_n)
 | 
|  |    156 |   then show ?thesis
 | 
|  |    157 |     by (auto simp: is_starext_def)
 | 
|  |    158 | qed
 | 
| 27468 |    159 | 
 | 
| 64435 |    160 | text \<open>Any nonstandard extension is in fact the Star-function.\<close>
 | 
| 70218 |    161 | lemma is_starfun_starext:
 | 
|  |    162 |   assumes "is_starext F f"
 | 
|  |    163 |   shows "F = *f* f"
 | 
|  |    164 |   proof -
 | 
|  |    165 |   have "F x = (*f* f) x"
 | 
|  |    166 |     if "\<forall>x y. \<exists>X\<in>Rep_star x. \<exists>Y\<in>Rep_star y. (y = F x) = (\<forall>\<^sub>F n in \<U>. Y n = f (X n))" for x
 | 
|  |    167 |     by (metis that mem_Rep_star_iff star_n_eq_iff starfun_star_n)
 | 
|  |    168 |   with assms show ?thesis
 | 
|  |    169 |     by (force simp add: is_starext_def)
 | 
|  |    170 | qed
 | 
| 27468 |    171 | 
 | 
| 64435 |    172 | lemma is_starext_starfun_iff: "is_starext F f \<longleftrightarrow> F = *f* f"
 | 
|  |    173 |   by (blast intro: is_starfun_starext is_starext_starfun)
 | 
| 27468 |    174 | 
 | 
| 70218 |    175 | text \<open>Extended function has same solution as its standard version
 | 
| 64435 |    176 |   for real arguments. i.e they are the same for all real arguments.\<close>
 | 
| 27468 |    177 | lemma starfun_eq: "( *f* f) (star_of a) = star_of (f a)"
 | 
| 64435 |    178 |   by (rule starfun_star_of)
 | 
| 27468 |    179 | 
 | 
| 61982 |    180 | lemma starfun_approx: "( *f* f) (star_of a) \<approx> star_of (f a)"
 | 
| 64435 |    181 |   by simp
 | 
| 27468 |    182 | 
 | 
| 64435 |    183 | text \<open>Useful for NS definition of derivatives.\<close>
 | 
|  |    184 | lemma starfun_lambda_cancel: "\<And>x'. ( *f* (\<lambda>h. f (x + h))) x'  = ( *f* f) (star_of x + x')"
 | 
|  |    185 |   by transfer (rule refl)
 | 
| 27468 |    186 | 
 | 
| 64435 |    187 | lemma starfun_lambda_cancel2: "( *f* (\<lambda>h. f (g (x + h)))) x' = ( *f* (f \<circ> g)) (star_of x + x')"
 | 
|  |    188 |   unfolding o_def by (rule starfun_lambda_cancel)
 | 
| 27468 |    189 | 
 | 
|  |    190 | lemma starfun_mult_HFinite_approx:
 | 
| 64435 |    191 |   "( *f* f) x \<approx> l \<Longrightarrow> ( *f* g) x \<approx> m \<Longrightarrow> l \<in> HFinite \<Longrightarrow> m \<in> HFinite \<Longrightarrow>
 | 
|  |    192 |     ( *f* (\<lambda>x. f x * g x)) x \<approx> l * m"
 | 
|  |    193 |   for l m :: "'a::real_normed_algebra star"
 | 
| 70218 |    194 |   using approx_mult_HFinite by auto
 | 
| 27468 |    195 | 
 | 
| 64435 |    196 | lemma starfun_add_approx: "( *f* f) x \<approx> l \<Longrightarrow> ( *f* g) x \<approx> m \<Longrightarrow> ( *f* (%x. f x + g x)) x \<approx> l + m"
 | 
|  |    197 |   by (auto intro: approx_add)
 | 
| 27468 |    198 | 
 | 
| 64435 |    199 | text \<open>Examples: \<open>hrabs\<close> is nonstandard extension of \<open>rabs\<close>,
 | 
|  |    200 |   \<open>inverse\<close> is nonstandard extension of \<open>inverse\<close>.\<close>
 | 
| 27468 |    201 | 
 | 
| 64435 |    202 | text \<open>Can be proved easily using theorem \<open>starfun\<close> and
 | 
|  |    203 |   properties of ultrafilter as for inverse below we
 | 
|  |    204 |   use the theorem we proved above instead.\<close>
 | 
| 27468 |    205 | 
 | 
|  |    206 | lemma starfun_rabs_hrabs: "*f* abs = abs"
 | 
| 64435 |    207 |   by (simp only: star_abs_def)
 | 
| 27468 |    208 | 
 | 
| 64435 |    209 | lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse x"
 | 
|  |    210 |   by (simp only: star_inverse_def)
 | 
| 27468 |    211 | 
 | 
| 64435 |    212 | lemma starfun_inverse: "\<And>x. inverse (( *f* f) x) = ( *f* (\<lambda>x. inverse (f x))) x"
 | 
|  |    213 |   by transfer (rule refl)
 | 
| 27468 |    214 | declare starfun_inverse [symmetric, simp]
 | 
|  |    215 | 
 | 
| 64435 |    216 | lemma starfun_divide: "\<And>x. ( *f* f) x / ( *f* g) x = ( *f* (\<lambda>x. f x / g x)) x"
 | 
|  |    217 |   by transfer (rule refl)
 | 
| 27468 |    218 | declare starfun_divide [symmetric, simp]
 | 
|  |    219 | 
 | 
| 64435 |    220 | lemma starfun_inverse2: "\<And>x. inverse (( *f* f) x) = ( *f* (\<lambda>x. inverse (f x))) x"
 | 
|  |    221 |   by transfer (rule refl)
 | 
| 27468 |    222 | 
 | 
| 64435 |    223 | text \<open>General lemma/theorem needed for proofs in elementary topology of the reals.\<close>
 | 
| 67613 |    224 | lemma starfun_mem_starset: "\<And>x. ( *f* f) x \<in> *s* A \<Longrightarrow> x \<in> *s* {x. f x \<in> A}"
 | 
| 64435 |    225 |   by transfer simp
 | 
| 27468 |    226 | 
 | 
| 64435 |    227 | text \<open>Alternative definition for \<open>hrabs\<close> with \<open>rabs\<close> function applied
 | 
|  |    228 |   entrywise to equivalence class representative.
 | 
|  |    229 |   This is easily proved using @{thm [source] starfun} and ns extension thm.\<close>
 | 
|  |    230 | lemma hypreal_hrabs: "\<bar>star_n X\<bar> = star_n (\<lambda>n. \<bar>X n\<bar>)"
 | 
|  |    231 |   by (simp only: starfun_rabs_hrabs [symmetric] starfun)
 | 
| 27468 |    232 | 
 | 
| 64435 |    233 | text \<open>Nonstandard extension of set through nonstandard extension
 | 
|  |    234 |    of \<open>rabs\<close> function i.e. \<open>hrabs\<close>. A more general result should be
 | 
|  |    235 |    where we replace \<open>rabs\<close> by some arbitrary function \<open>f\<close> and \<open>hrabs\<close>
 | 
| 61975 |    236 |    by its NS extenson. See second NS set extension below.\<close>
 | 
| 64435 |    237 | lemma STAR_rabs_add_minus: "*s* {x. \<bar>x + - y\<bar> < r} = {x. \<bar>x + -star_of y\<bar> < star_of r}"
 | 
|  |    238 |   by transfer (rule refl)
 | 
| 27468 |    239 | 
 | 
|  |    240 | lemma STAR_starfun_rabs_add_minus:
 | 
| 64435 |    241 |   "*s* {x. \<bar>f x + - y\<bar> < r} = {x. \<bar>( *f* f) x + -star_of y\<bar> < star_of r}"
 | 
|  |    242 |   by transfer (rule refl)
 | 
| 27468 |    243 | 
 | 
| 64435 |    244 | text \<open>Another characterization of Infinitesimal and one of \<open>\<approx>\<close> relation.
 | 
|  |    245 |   In this theory since \<open>hypreal_hrabs\<close> proved here. Maybe move both theorems??\<close>
 | 
| 27468 |    246 | lemma Infinitesimal_FreeUltrafilterNat_iff2:
 | 
| 64435 |    247 |   "star_n X \<in> Infinitesimal \<longleftrightarrow> (\<forall>m. eventually (\<lambda>n. norm (X n) < inverse (real (Suc m))) \<U>)"
 | 
|  |    248 |   by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def hnorm_def
 | 
|  |    249 |       star_of_nat_def starfun_star_n star_n_inverse star_n_less)
 | 
| 27468 |    250 | 
 | 
|  |    251 | lemma HNatInfinite_inverse_Infinitesimal [simp]:
 | 
| 70218 |    252 |   assumes "n \<in> HNatInfinite"
 | 
|  |    253 |   shows "inverse (hypreal_of_hypnat n) \<in> Infinitesimal"
 | 
|  |    254 | proof (cases n)
 | 
|  |    255 |   case (star_n X)
 | 
|  |    256 |   then have *: "\<And>k. \<forall>\<^sub>F n in \<U>. k < X n"
 | 
|  |    257 |     using HNatInfinite_FreeUltrafilterNat assms by blast
 | 
|  |    258 |   have "\<forall>\<^sub>F n in \<U>. inverse (real (X n)) < inverse (1 + real m)" for m
 | 
|  |    259 |     using * [of "Suc m"] by (auto elim!: eventually_mono)
 | 
|  |    260 |   then show ?thesis
 | 
|  |    261 |     using star_n by (auto simp: of_hypnat_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff2)
 | 
|  |    262 | qed
 | 
| 27468 |    263 | 
 | 
| 64435 |    264 | lemma approx_FreeUltrafilterNat_iff:
 | 
|  |    265 |   "star_n X \<approx> star_n Y \<longleftrightarrow> (\<forall>r>0. eventually (\<lambda>n. norm (X n - Y n) < r) \<U>)"
 | 
| 70218 |    266 |   (is "?lhs = ?rhs")
 | 
|  |    267 | proof -
 | 
|  |    268 |   have "?lhs = (star_n X - star_n Y \<approx> 0)"
 | 
|  |    269 |     using approx_minus_iff by blast
 | 
|  |    270 |   also have "... = ?rhs"
 | 
|  |    271 |     by (metis (full_types) Infinitesimal_FreeUltrafilterNat_iff mem_infmal_iff star_n_diff)
 | 
|  |    272 |   finally show ?thesis .
 | 
|  |    273 | qed
 | 
| 27468 |    274 | 
 | 
| 64435 |    275 | lemma approx_FreeUltrafilterNat_iff2:
 | 
|  |    276 |   "star_n X \<approx> star_n Y \<longleftrightarrow> (\<forall>m. eventually (\<lambda>n. norm (X n - Y n) < inverse (real (Suc m))) \<U>)"
 | 
| 70218 |    277 |   (is "?lhs = ?rhs")
 | 
|  |    278 | proof -
 | 
|  |    279 |   have "?lhs = (star_n X - star_n Y \<approx> 0)"
 | 
|  |    280 |     using approx_minus_iff by blast
 | 
|  |    281 |   also have "... = ?rhs"
 | 
|  |    282 |     by (metis (full_types) Infinitesimal_FreeUltrafilterNat_iff2 mem_infmal_iff star_n_diff)
 | 
|  |    283 |   finally show ?thesis .
 | 
|  |    284 | qed
 | 
| 27468 |    285 | 
 | 
|  |    286 | lemma inj_starfun: "inj starfun"
 | 
| 70218 |    287 | proof (rule inj_onI)
 | 
|  |    288 |   show "\<phi> = \<psi>" if eq: "*f* \<phi> = *f* \<psi>" for \<phi> \<psi> :: "'a \<Rightarrow> 'b"
 | 
|  |    289 |   proof (rule ext, rule ccontr)
 | 
|  |    290 |     show False
 | 
|  |    291 |       if "\<phi> x \<noteq> \<psi> x" for x
 | 
|  |    292 |       by (metis eq that star_of_inject starfun_eq)
 | 
|  |    293 |   qed
 | 
|  |    294 | qed
 | 
| 27468 |    295 | 
 | 
|  |    296 | end
 |