src/HOL/Hyperreal/NthRoot.thy
changeset 21404 eb85850d3eb7
parent 20898 113c9516a2d7
child 21865 55cc354fd2d9
--- a/src/HOL/Hyperreal/NthRoot.thy	Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/NthRoot.thy	Fri Nov 17 02:20:03 2006 +0100
@@ -11,11 +11,11 @@
 begin
 
 definition
-
-  root :: "[nat, real] \<Rightarrow> real"
+  root :: "[nat, real] \<Rightarrow> real" where
   "root n x = (THE u. (0 < x \<longrightarrow> 0 < u) \<and> (u ^ n = x))"
 
-  sqrt :: "real \<Rightarrow> real"
+definition
+  sqrt :: "real \<Rightarrow> real" where
   "sqrt x = root 2 x"