src/HOL/Library/Diagonal_Subsequence.thy
changeset 66447 a1f5c5c26fa6
parent 60500 903bb1495239
child 67091 1393c2340eec
--- a/src/HOL/Library/Diagonal_Subsequence.thy	Thu Aug 17 14:40:42 2017 +0200
+++ b/src/HOL/Library/Diagonal_Subsequence.thy	Thu Aug 17 14:52:56 2017 +0200
@@ -8,28 +8,28 @@
 
 locale subseqs =
   fixes P::"nat\<Rightarrow>(nat\<Rightarrow>nat)\<Rightarrow>bool"
-  assumes ex_subseq: "\<And>n s. subseq s \<Longrightarrow> \<exists>r'. subseq r' \<and> P n (s o r')"
+  assumes ex_subseq: "\<And>n s. strict_mono (s::nat\<Rightarrow>nat) \<Longrightarrow> \<exists>r'. strict_mono r' \<and> P n (s o r')"
 begin
 
-definition reduce where "reduce s n = (SOME r'. subseq r' \<and> P n (s o r'))"
+definition reduce where "reduce s n = (SOME r'::nat\<Rightarrow>nat. strict_mono r' \<and> P n (s o r'))"
 
 lemma subseq_reduce[intro, simp]:
-  "subseq s \<Longrightarrow> subseq (reduce s n)"
+  "strict_mono s \<Longrightarrow> strict_mono (reduce s n)"
   unfolding reduce_def by (rule someI2_ex[OF ex_subseq]) auto
 
 lemma reduce_holds:
-  "subseq s \<Longrightarrow> P n (s o reduce s n)"
+  "strict_mono s \<Longrightarrow> P n (s o reduce s n)"
   unfolding reduce_def by (rule someI2_ex[OF ex_subseq]) (auto simp: o_def)
 
-primrec seqseq where
+primrec seqseq :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
   "seqseq 0 = id"
 | "seqseq (Suc n) = seqseq n o reduce (seqseq n) n"
 
-lemma subseq_seqseq[intro, simp]: "subseq (seqseq n)"
+lemma subseq_seqseq[intro, simp]: "strict_mono (seqseq n)"
 proof (induct n)
-  case 0 thus ?case by (simp add: subseq_def)
+  case 0 thus ?case by (simp add: strict_mono_def)
 next
-  case (Suc n) thus ?case by (subst seqseq.simps) (auto intro!: subseq_o)
+  case (Suc n) thus ?case by (subst seqseq.simps) (auto intro!: strict_mono_o)
 qed
 
 lemma seqseq_holds:
@@ -40,35 +40,29 @@
   thus ?thesis by simp
 qed
 
-definition diagseq where "diagseq i = seqseq i i"
-
-lemma subseq_mono: "subseq f \<Longrightarrow> a \<le> b \<Longrightarrow> f a \<le> f b"
-  by (metis le_eq_less_or_eq subseq_mono)
-
-lemma subseq_strict_mono: "subseq f \<Longrightarrow> a < b \<Longrightarrow> f a < f b"
-  by (simp add: subseq_def)
+definition diagseq :: "nat \<Rightarrow> nat" where "diagseq i = seqseq i i"
 
 lemma diagseq_mono: "diagseq n < diagseq (Suc n)"
 proof -
   have "diagseq n < seqseq n (Suc n)"
-    using subseq_seqseq[of n] by (simp add: diagseq_def subseq_def)
+    using subseq_seqseq[of n] by (simp add: diagseq_def strict_mono_def)
   also have "\<dots> \<le> seqseq n (reduce (seqseq n) n (Suc n))"
-    by (auto intro: subseq_mono seq_suble)
+    using strict_mono_less_eq seq_suble by blast
   also have "\<dots> = diagseq (Suc n)" by (simp add: diagseq_def)
   finally show ?thesis .
 qed
 
-lemma subseq_diagseq: "subseq diagseq"
-  using diagseq_mono by (simp add: subseq_Suc_iff diagseq_def)
+lemma subseq_diagseq: "strict_mono diagseq"
+  using diagseq_mono by (simp add: strict_mono_Suc_iff diagseq_def)
 
 primrec fold_reduce where
   "fold_reduce n 0 = id"
 | "fold_reduce n (Suc k) = fold_reduce n k o reduce (seqseq (n + k)) (n + k)"
 
-lemma subseq_fold_reduce[intro, simp]: "subseq (fold_reduce n k)"
+lemma subseq_fold_reduce[intro, simp]: "strict_mono (fold_reduce n k)"
 proof (induct k)
-  case (Suc k) from subseq_o[OF this subseq_reduce] show ?case by (simp add: o_def)
-qed (simp add: subseq_def)
+  case (Suc k) from strict_mono_o[OF this subseq_reduce] show ?case by (simp add: o_def)
+qed (simp add: strict_mono_def)
 
 lemma ex_subseq_reduce_index: "seqseq (n + k) = seqseq n o fold_reduce n k"
   by (induct k) simp_all
@@ -95,14 +89,14 @@
   assumes "m \<le> n" shows "diagseq n = (seqseq m o (fold_reduce m (n - m))) n"
   using diagseq_add[of m "n - m"] assms by simp
 
-lemma subseq_diagonal_rest: "subseq (\<lambda>x. fold_reduce k x (k + x))"
-  unfolding subseq_Suc_iff fold_reduce.simps o_def
+lemma subseq_diagonal_rest: "strict_mono (\<lambda>x. fold_reduce k x (k + x))"
+  unfolding strict_mono_Suc_iff fold_reduce.simps o_def
 proof
   fix n
   have "fold_reduce k n (k + n) < fold_reduce k n (k + Suc n)" (is "?lhs < _")
-    by (auto intro: subseq_strict_mono)
+    by (auto intro: strict_monoD)
   also have "\<dots> \<le> fold_reduce k n (reduce (seqseq (k + n)) (k + n) (k + Suc n))"
-    by (rule subseq_mono) (auto intro!: seq_suble subseq_mono)
+    by (auto intro: less_mono_imp_le_mono seq_suble strict_monoD)
   finally show "?lhs < \<dots>" .
 qed
 
@@ -110,7 +104,7 @@
   by (auto simp: o_def diagseq_add)
 
 lemma diagseq_holds:
-  assumes subseq_stable: "\<And>r s n. subseq r \<Longrightarrow> P n s \<Longrightarrow> P n (s o r)"
+  assumes subseq_stable: "\<And>r s n. strict_mono r \<Longrightarrow> P n s \<Longrightarrow> P n (s o r)"
   shows "P k (diagseq o (op + (Suc k)))"
   unfolding diagseq_seqseq by (intro subseq_stable subseq_diagonal_rest seqseq_holds)