tuned document;
authorwenzelm
Tue, 27 Feb 2007 00:33:49 +0100
changeset 22367 6860f09242bf
parent 22366 f4840bfffe5d
child 22368 0e0fe77d4768
tuned document;
src/HOL/Library/Coinductive_List.thy
src/HOL/Library/Continuity.thy
src/HOL/Library/GCD.thy
src/HOL/Library/Library/document/root.bib
src/HOL/Library/Ramsey.thy
src/HOL/document/root.bib
--- a/src/HOL/Library/Coinductive_List.thy	Tue Feb 27 00:32:52 2007 +0100
+++ b/src/HOL/Library/Coinductive_List.thy	Tue Feb 27 00:33:49 2007 +0100
@@ -280,7 +280,7 @@
 qed
 
 
-subsection {* Equality as greatest fixed-point; the bisimulation principle. *}
+subsection {* Equality as greatest fixed-point -- the bisimulation principle *}
 
 consts
   EqLList :: "('a Datatype.item \<times> 'a Datatype.item) set \<Rightarrow>
--- a/src/HOL/Library/Continuity.thy	Tue Feb 27 00:32:52 2007 +0100
+++ b/src/HOL/Library/Continuity.thy	Tue Feb 27 00:33:49 2007 +0100
@@ -9,20 +9,22 @@
 imports Main
 begin
 
-subsection{*Continuity for complete lattices*}
+subsection {* Continuity for complete lattices *}
 
-constdefs
- chain :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool"
-"chain M == !i. M i \<le> M(Suc i)"
- continuous :: "('a::order \<Rightarrow> 'a::order) \<Rightarrow> bool"
-"continuous F == !M. chain M \<longrightarrow> F(SUP i. M i) = (SUP i. F(M i))"
+definition
+  chain :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where
+  "chain M \<longleftrightarrow> (\<forall>i. M i \<le> M (Suc i))"
+
+definition
+  continuous :: "('a::order \<Rightarrow> 'a::order) \<Rightarrow> bool" where
+  "continuous F \<longleftrightarrow> (\<forall>M. chain M \<longrightarrow> F (SUP i. M i) = (SUP i. F (M i)))"
 
 abbreviation
   bot :: "'a::order" where
-  "bot == Sup {}"
+  "bot \<equiv> Sup {}"
 
 lemma SUP_nat_conv:
- "(SUP n::nat. M n::'a::comp_lat) = join (M 0) (SUP n. M(Suc n))"
+  "(SUP n::nat. M n::'a::comp_lat) = join (M 0) (SUP n. M(Suc n))"
 apply(rule order_antisym)
  apply(rule SUP_leI)
  apply(case_tac n)
--- a/src/HOL/Library/GCD.thy	Tue Feb 27 00:32:52 2007 +0100
+++ b/src/HOL/Library/GCD.thy	Tue Feb 27 00:33:49 2007 +0100
@@ -180,10 +180,10 @@
 
 lemma gcd_add2 [simp]: "gcd (m, m + n) = gcd (m, n)"
 proof -
-  have "gcd (m, m + n) = gcd (m + n, m)" by (rule gcd_commute) 
+  have "gcd (m, m + n) = gcd (m + n, m)" by (rule gcd_commute)
   also have "... = gcd (n + m, m)" by (simp add: add_commute)
   also have "... = gcd (n, m)" by simp
-  also have  "... = gcd (m, n)" by (rule gcd_commute) 
+  also have  "... = gcd (m, n)" by (rule gcd_commute)
   finally show ?thesis .
 qed
 
@@ -195,110 +195,124 @@
 lemma gcd_add_mult: "gcd (m, k * m + n) = gcd (m, n)"
   by (induct k) (simp_all add: add_assoc)
 
-  (* Division by gcd yields rrelatively primes *)
 
+text {*
+  \medskip Division by gcd yields rrelatively primes.
+*}
 
 lemma div_gcd_relprime:
-  assumes nz:"a\<noteq>0 \<or> b\<noteq>0"
+  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
   shows "gcd (a div gcd(a,b), b div gcd(a,b)) = 1"
-proof-
-  let ?g = "gcd (a,b)"
+proof -
+  let ?g = "gcd (a, b)"
   let ?a' = "a div ?g"
   let ?b' = "b div ?g"
   let ?g' = "gcd (?a', ?b')"
   have dvdg: "?g dvd a" "?g dvd b" by simp_all
   have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
-  from dvdg dvdg' obtain ka kb ka' kb' where 
-   kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'" 
+  from dvdg dvdg' obtain ka kb ka' kb' where
+      kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
     unfolding dvd_def by blast
-  hence	"?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
-  hence dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
-    by (auto simp add: dvd_mult_div_cancel[OF dvdg(1)] 
-      dvd_mult_div_cancel[OF dvdg(2)] dvd_def)
+  then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
+  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
+    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
+      dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
   have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
-  hence gp: "?g > 0" by simp 
-  from gcd_greatest[OF dvdgg'] have "?g * ?g' dvd ?g" .
-  with dvd_mult_cancel1[OF gp] show "?g' = 1" by simp
+  then have gp: "?g > 0" by simp
+  from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+  with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
 qed
 
-  (* gcd on integers *)
-constdefs igcd:: "int \<Rightarrow> int \<Rightarrow> int"
-  "igcd i j \<equiv> int (gcd (nat (abs i),nat (abs j)))"
-lemma igcd_dvd1[simp]:"igcd i j dvd i"
+
+text {*
+  \medskip Gcd on integers.
+*}
+
+definition
+  igcd :: "int \<Rightarrow> int \<Rightarrow> int" where
+  "igcd i j = int (gcd (nat (abs i), nat (abs j)))"
+
+lemma igcd_dvd1 [simp]: "igcd i j dvd i"
   by (simp add: igcd_def int_dvd_iff)
 
-lemma igcd_dvd2[simp]:"igcd i j dvd j"
-by (simp add: igcd_def int_dvd_iff)
+lemma igcd_dvd2 [simp]: "igcd i j dvd j"
+  by (simp add: igcd_def int_dvd_iff)
 
 lemma igcd_pos: "igcd i j \<ge> 0"
-by (simp add: igcd_def)
-lemma igcd0[simp]: "(igcd i j = 0) = (i = 0 \<and> j = 0)"
-by (simp add: igcd_def gcd_zero) arith
+  by (simp add: igcd_def)
+
+lemma igcd0 [simp]: "(igcd i j = 0) = (i = 0 \<and> j = 0)"
+  by (simp add: igcd_def gcd_zero) arith
 
 lemma igcd_commute: "igcd i j = igcd j i"
   unfolding igcd_def by (simp add: gcd_commute)
-lemma igcd_neg1[simp]: "igcd (- i) j = igcd i j"
+
+lemma igcd_neg1 [simp]: "igcd (- i) j = igcd i j"
   unfolding igcd_def by simp
-lemma igcd_neg2[simp]: "igcd i (- j) = igcd i j"
+
+lemma igcd_neg2 [simp]: "igcd i (- j) = igcd i j"
   unfolding igcd_def by simp
+
 lemma zrelprime_dvd_mult: "igcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
   unfolding igcd_def
-proof-
-  assume H: "int (gcd (nat \<bar>i\<bar>, nat \<bar>j\<bar>)) = 1" "i dvd k * j"
-  hence g: "gcd (nat \<bar>i\<bar>, nat \<bar>j\<bar>) = 1" by simp
-  from H(2) obtain h where h:"k*j = i*h" unfolding dvd_def by blast  
+proof -
+  assume "int (gcd (nat \<bar>i\<bar>, nat \<bar>j\<bar>)) = 1" "i dvd k * j"
+  then have g: "gcd (nat \<bar>i\<bar>, nat \<bar>j\<bar>) = 1" by simp
+  from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast
   have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
-    unfolding dvd_def 
-    by (rule_tac x= "nat \<bar>h\<bar>" in exI,simp add: h nat_abs_mult_distrib[symmetric])
-  from relprime_dvd_mult[OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'" 
+    unfolding dvd_def
+    by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric])
+  from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
     unfolding dvd_def by blast
   from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
-  hence "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
+  then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
   then show ?thesis
-    apply (subst zdvd_abs1[symmetric])
-    apply (subst zdvd_abs2[symmetric])
+    apply (subst zdvd_abs1 [symmetric])
+    apply (subst zdvd_abs2 [symmetric])
     apply (unfold dvd_def)
-    apply (rule_tac x="int h'" in exI, simp)
+    apply (rule_tac x = "int h'" in exI, simp)
     done
 qed
 
 lemma int_nat_abs: "int (nat (abs x)) = abs x"  by arith
-lemma igcd_greatest: assumes km:"k dvd m" and kn:"k dvd n"  shows "k dvd igcd m n"
-proof-
+
+lemma igcd_greatest:
+  assumes "k dvd m" and "k dvd n"
+  shows "k dvd igcd m n"
+proof -
   let ?k' = "nat \<bar>k\<bar>"
   let ?m' = "nat \<bar>m\<bar>"
   let ?n' = "nat \<bar>n\<bar>"
-  from km kn have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
+  from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
     unfolding zdvd_int by (simp_all only: int_nat_abs zdvd_abs1 zdvd_abs2)
-  from gcd_greatest[OF dvd'] have "int (nat \<bar>k\<bar>) dvd igcd m n"
+  from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd igcd m n"
     unfolding igcd_def by (simp only: zdvd_int)
-  hence "\<bar>k\<bar> dvd igcd m n" by (simp only: int_nat_abs)
-  thus "k dvd igcd m n" by (simp add: zdvd_abs1)
+  then have "\<bar>k\<bar> dvd igcd m n" by (simp only: int_nat_abs)
+  then show "k dvd igcd m n" by (simp add: zdvd_abs1)
 qed
 
 lemma div_igcd_relprime:
-  assumes nz:"a\<noteq>0 \<or> b\<noteq>0"
+  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
   shows "igcd (a div (igcd a b)) (b div (igcd a b)) = 1"
-proof-
+proof -
   from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by simp
-  have th1: "(1::int) = int 1" by simp
   let ?g = "igcd a b"
   let ?a' = "a div ?g"
   let ?b' = "b div ?g"
   let ?g' = "igcd ?a' ?b'"
   have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: igcd_dvd1 igcd_dvd2)
   have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: igcd_dvd1 igcd_dvd2)
-  from dvdg dvdg' obtain ka kb ka' kb' where 
-   kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'" 
+  from dvdg dvdg' obtain ka kb ka' kb' where
+   kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'"
     unfolding dvd_def by blast
-  hence	"?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
-  hence dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
-    by (auto simp add: zdvd_mult_div_cancel[OF dvdg(1)] 
-      zdvd_mult_div_cancel[OF dvdg(2)] dvd_def)
+  then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
+  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
+    by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
+      zdvd_mult_div_cancel [OF dvdg(2)] dvd_def)
   have "?g \<noteq> 0" using nz by simp
-  hence gp: "?g \<noteq> 0" using igcd_pos[where i="a" and j="b"] by arith
-  from igcd_greatest[OF dvdgg'] have "?g * ?g' dvd ?g" .
-  with zdvd_mult_cancel1[OF gp] have "\<bar>?g'\<bar> = 1" by simp 
+  then have gp: "?g \<noteq> 0" using igcd_pos[where i="a" and j="b"] by arith
+  from igcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+  with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
   with igcd_pos show "?g' = 1" by simp
 qed
 
--- a/src/HOL/Library/Library/document/root.bib	Tue Feb 27 00:32:52 2007 +0100
+++ b/src/HOL/Library/Library/document/root.bib	Tue Feb 27 00:33:49 2007 +0100
@@ -1,5 +1,12 @@
 
- @InProceedings{Avigad-Donnelly,
+@Unpublished{Abrial-Laffitte,
+  author = 	 {Abrial and Laffitte},
+  title = 	 {Towards the Mechanization of the Proofs of
+                  Some Classical Theorems of Set Theory},
+  note = 	 {Unpublished}
+}
+
+@InProceedings{Avigad-Donnelly,
   author = 	 {Jeremy Avigad and Kevin Donnelly},
   title = 	 {Formalizing {O} notation in {Isabelle/HOL}},
   booktitle = 	 {Automated Reasoning: second international conference, IJCAR 2004},
@@ -9,13 +16,6 @@
   publisher =	 {Springer}
 }
 
-@Unpublished{Abrial-Laffitte,
-  author = 	 {Abrial and Laffitte},
-  title = 	 {Towards the Mechanization of the Proofs of
-                  Some Classical Theorems of Set Theory},
-  note = 	 {Unpublished}
-}
-
 @Book{Oberschelp:1993,
   author =	 {Arnold Oberschelp},
   title = 	 {Rekursionstheorie},
@@ -23,6 +23,21 @@
   year = 	 1993
 }
 
+@InProceedings{Podelski-Rybalchenko,
+  author = 	 {Andreas Podelski and Andrey Rybalchenko},
+  title = 	 {Transition Invariants},
+  booktitle = 	 {19th Annual IEEE Symposium on Logic in Computer Science (LICS'04)},
+  pages =	 {32--41},
+  year =	 2004
+}
+
+@Book{davenport92,
+  author =	 {H. Davenport},
+  title = 	 {The Higher Arithmetic},
+  publisher = 	 {Cambridge University Press},
+  year = 	 1992
+}
+
 @InProceedings{paulin-tlca,
   author	= {Christine Paulin-Mohring},
   title		= {Inductive Definitions in the System {Coq}: Rules and
--- a/src/HOL/Library/Ramsey.thy	Tue Feb 27 00:32:52 2007 +0100
+++ b/src/HOL/Library/Ramsey.thy	Tue Feb 27 00:33:49 2007 +0100
@@ -235,10 +235,9 @@
 
 subsection {*Disjunctive Well-Foundedness*}
 
-text{*An application of Ramsey's theorem to program termination. See
-
-Andreas Podelski and Andrey Rybalchenko, Transition Invariants, 19th Annual
-IEEE Symposium on Logic in Computer Science (LICS'04), pages 32--41 (2004).
+text {*
+  An application of Ramsey's theorem to program termination. See
+  \cite{Podelski-Rybalchenko}.
 *}
 
 definition
--- a/src/HOL/document/root.bib	Tue Feb 27 00:32:52 2007 +0100
+++ b/src/HOL/document/root.bib	Tue Feb 27 00:33:49 2007 +0100
@@ -1,9 +1,3 @@
-@Book{davenport92,
-  author =	 {H. Davenport},
-  title = 	 {The Higher Arithmetic},
-  publisher = 	 {Cambridge University Press},
-  year = 	 1992
-}
 
 @book{Birkhoff79,author={Garret Birkhoff},title={Lattice Theory},
 publisher={American Mathematical Society},year=1979}
\ No newline at end of file