]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/sqrt.ma
...
[helm.git] / helm / software / matita / library / nat / sqrt.ma
index 733afe641af55679f8d89f44656fa02ea82c178a..d523d634dc5ff3fb503fa34f1ca39eaf432fe040 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/nat/sqrt/".
-
 include "nat/times.ma".
 include "nat/compare.ma".
 include "nat/log.ma".
@@ -204,3 +202,18 @@ apply f_m_to_le_max
         |simplify;apply le_exp_log;assumption]
      |rewrite < H1;simplify;apply le_n]]
 qed.
+
+(* monotonicity *)
+theorem monotonic_sqrt: monotonic nat le sqrt.
+unfold.intros.
+unfold sqrt in ⊢ (? ? (% ?)).
+apply f_m_to_le_max
+  [apply (trans_le ? ? ? ? H).
+   apply le_sqrt_n_n
+  |apply le_to_leb_true.
+   apply (trans_le ? ? ? ? H). 
+   apply leq_sqrt_n
+  ]
+qed.
+   
+