src/HOL/Library/Landau_Symbols.thy
changeset 80175 200107cdd3ac
parent 77491 9d431c939a7f
--- a/src/HOL/Library/Landau_Symbols.thy	Sun May 05 15:31:08 2024 +0100
+++ b/src/HOL/Library/Landau_Symbols.thy	Mon May 06 14:39:33 2024 +0100
@@ -1447,9 +1447,8 @@
 lemma landau_symbol_if_at_top_eq [simp]:
   assumes "landau_symbol L L' Lr"
   shows   "L at_top (\<lambda>x::'a::linordered_semidom. if x = a then f x else g x) = L at_top (g)"
-apply (rule landau_symbol.cong[OF assms])
-using less_add_one[of a] apply (auto intro: eventually_mono  eventually_ge_at_top[of "a + 1"])
-done
+  apply (rule landau_symbol.cong[OF assms])
+  by (auto simp add: frequently_def eventually_at_top_linorder dest!: spec [where x= "a + 1"])
 
 lemmas landau_symbols_if_at_top_eq [simp] = landau_symbols[THEN landau_symbol_if_at_top_eq]