]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/arithmetics/minimization.ma
Still porting chebyshev
[helm.git] / matita / matita / lib / arithmetics / minimization.ma
index 143104c50a75b98371aa59b9225ad86834b40c60..5d4fc9bce9a5aa03002dfdd6151cd0bc14324af7 100644 (file)
@@ -122,7 +122,7 @@ qed.
 theorem max_f_g: ∀f,g,n.(∀i. i < n → f i = g i) → 
   max n f = max n g.
 #f #g #n (elim n) //
-#m #Hind #ext normalize >ext normalize {Hind} >Hind //
+#m #Hind #ext normalize >ext normalize in Hind; >Hind //
 #i #ltim @ext /2/
 qed.
 
@@ -140,6 +140,39 @@ theorem f_max_true : ∀ f.∀n.
 #Hall * #x * #ltx #fx @False_ind @(absurd … fx) >Hall /2/
 qed.
 
+theorem f_false_to_le_max: ∀f,n,p. (∃i:nat.i<n∧f i=true) →
+  (∀m. p < m → f m = false) → max n f ≤ p.
+#f #n #p #H1 #H2 @not_lt_to_le % #H3
+@(absurd ?? not_eq_true_false) <(H2 ? H3) @sym_eq
+@(f_max_true ? n H1)
+qed.
+
+theorem exists_forall_lt:∀f,n. 
+(∃i. i < n ∧ f i = true) ∨ (∀i. i < n → f i = false).
+#f #n elim n
+  [%2 #i #lti0 @False_ind @(absurd ? lti0) @le_to_not_lt // 
+  |#n1 *
+    [* #a * #Ha1 #Ha2 %1 %{a} % // @le_S //
+    |#H cases (true_or_false (f n1)) #HfS >HfS
+      [%1 %{n1} /2/
+      |%2 #i #lei 
+       cases (le_to_or_lt_eq ?? lei) #Hi 
+        [@H @le_S_S_to_le @Hi | destruct (Hi) //] 
+      ]
+    ]
+  ]
+qed.
+     
+theorem exists_max_forall_false:∀f,n. 
+((∃i. i < n ∧ f i = true) ∧ (f (max n f) = true))∨
+((∀i. i < n → f i = false) ∧ (max n f) = O).
+#f #n
+cases (exists_forall_lt f n)
+  [#H %1 % // @(f_max_true f n) @H
+  |#H %2 % [@H | @max_not_exists @H 
+  ]
+qed.
+
 (* minimization *)
  
 (* min k b f is the minimun i, b ≤ i < b + k s.t. f i;  
@@ -230,7 +263,7 @@ lemma min_exists: ∀f.∀t,m. m < t → f m = true →
   [#b #lebm #ismin #eqtb @False_ind @(absurd … lebm) <eqtb
    @lt_to_not_le //
   |#d #Hind #b #lebm #ismin #eqt cases(le_to_or_lt_eq …lebm)
-    [#ltbm >false_min /2/ @Hind //
+    [#ltbm >false_min /2 by le_n/ @Hind //
       [#i #H #H1 @ismin /2/ | >eqt normalize //] 
     |#eqbm >true_min //
     ]
@@ -295,6 +328,7 @@ min n b g ≤ min n b f.
   [>ext //
   |(cases (g b)) normalize /2/ @Hind #i #ltb #ltim #fi
     @ext /2/
+  ]
 qed.
 
 theorem f_min_true : ∀ f.∀n,b.