src/HOL/UNITY/LessThan.ML
changeset 5069 3ea049f7979d
parent 4776 1f9362e769c1
child 5232 e5a7cdd07ea5
--- a/src/HOL/UNITY/LessThan.ML	Mon Jun 22 17:13:09 1998 +0200
+++ b/src/HOL/UNITY/LessThan.ML	Mon Jun 22 17:26:46 1998 +0200
@@ -12,26 +12,26 @@
 
 (*** lessThan ***)
 
-goalw thy [lessThan_def] "(i: lessThan k) = (i<k)";
+Goalw [lessThan_def] "(i: lessThan k) = (i<k)";
 by (Blast_tac 1);
 qed "lessThan_iff";
 AddIffs [lessThan_iff];
 
-goalw thy [lessThan_def] "lessThan 0 = {}";
+Goalw [lessThan_def] "lessThan 0 = {}";
 by (Simp_tac 1);
 qed "lessThan_0";
 Addsimps [lessThan_0];
 
-goalw thy [lessThan_def] "lessThan (Suc k) = insert k (lessThan k)";
+Goalw [lessThan_def] "lessThan (Suc k) = insert k (lessThan k)";
 by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
 by (Blast_tac 1);
 qed "lessThan_Suc";
 
-goal thy "(UN m. lessThan m) = UNIV";
+Goal "(UN m. lessThan m) = UNIV";
 by (Blast_tac 1);
 qed "UN_lessThan_UNIV";
 
-goalw thy [lessThan_def, atLeast_def, le_def]
+Goalw [lessThan_def, atLeast_def, le_def]
     "Compl(lessThan k) = atLeast k";
 by (Blast_tac 1);
 qed "Compl_lessThan";
@@ -39,38 +39,38 @@
 
 (*** greaterThan ***)
 
-goalw thy [greaterThan_def] "(i: greaterThan k) = (k<i)";
+Goalw [greaterThan_def] "(i: greaterThan k) = (k<i)";
 by (Blast_tac 1);
 qed "greaterThan_iff";
 AddIffs [greaterThan_iff];
 
-goalw thy [greaterThan_def] "greaterThan 0 = range Suc";
+Goalw [greaterThan_def] "greaterThan 0 = range Suc";
 by (blast_tac (claset() addIs [Suc_pred RS sym]) 1);
 qed "greaterThan_0";
 Addsimps [greaterThan_0];
 
-goalw thy [greaterThan_def] "greaterThan (Suc k) = greaterThan k - {Suc k}";
+Goalw [greaterThan_def] "greaterThan (Suc k) = greaterThan k - {Suc k}";
 by Safe_tac;
 by (blast_tac (claset() addIs [less_trans]) 1);
 by (asm_simp_tac (simpset() addsimps [le_Suc_eq, not_le_iff_less RS sym]) 1);
 by (asm_simp_tac (simpset() addsimps [not_le_iff_less]) 1);
 qed "greaterThan_Suc";
 
-goal thy "(INT m. greaterThan m) = {}";
+Goal "(INT m. greaterThan m) = {}";
 by (Blast_tac 1);
 qed "INT_greaterThan_UNIV";
 
-goalw thy [greaterThan_def, atMost_def, le_def]
+Goalw [greaterThan_def, atMost_def, le_def]
     "Compl(greaterThan k) = atMost k";
 by (Blast_tac 1);
 qed "Compl_greaterThan";
 
-goalw thy [greaterThan_def, atMost_def, le_def]
+Goalw [greaterThan_def, atMost_def, le_def]
     "Compl(atMost k) = greaterThan k";
 by (Blast_tac 1);
 qed "Compl_atMost";
 
-goal thy "less_than ^^ {k} = greaterThan k";
+Goal "less_than ^^ {k} = greaterThan k";
 by (Blast_tac 1);
 qed "Image_less_than";
 
@@ -78,32 +78,32 @@
 
 (*** atLeast ***)
 
-goalw thy [atLeast_def] "(i: atLeast k) = (k<=i)";
+Goalw [atLeast_def] "(i: atLeast k) = (k<=i)";
 by (Blast_tac 1);
 qed "atLeast_iff";
 AddIffs [atLeast_iff];
 
-goalw thy [atLeast_def, UNIV_def] "atLeast 0 = UNIV";
+Goalw [atLeast_def, UNIV_def] "atLeast 0 = UNIV";
 by (Simp_tac 1);
 qed "atLeast_0";
 Addsimps [atLeast_0];
 
-goalw thy [atLeast_def] "atLeast (Suc k) = atLeast k - {k}";
+Goalw [atLeast_def] "atLeast (Suc k) = atLeast k - {k}";
 by (simp_tac (simpset() addsimps [Suc_le_eq]) 1);
 by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1);
 by (Blast_tac 1);
 qed "atLeast_Suc";
 
-goal thy "(UN m. atLeast m) = UNIV";
+Goal "(UN m. atLeast m) = UNIV";
 by (Blast_tac 1);
 qed "UN_atLeast_UNIV";
 
-goalw thy [lessThan_def, atLeast_def, le_def]
+Goalw [lessThan_def, atLeast_def, le_def]
     "Compl(atLeast k) = lessThan k";
 by (Blast_tac 1);
 qed "Compl_atLeast";
 
-goal thy "less_than^-1 ^^ {k} = lessThan k";
+Goal "less_than^-1 ^^ {k} = lessThan k";
 by (Blast_tac 1);
 qed "Image_inverse_less_than";
 
@@ -111,26 +111,26 @@
 
 (*** atMost ***)
 
-goalw thy [atMost_def] "(i: atMost k) = (i<=k)";
+Goalw [atMost_def] "(i: atMost k) = (i<=k)";
 by (Blast_tac 1);
 qed "atMost_iff";
 AddIffs [atMost_iff];
 
-goalw thy [atMost_def] "atMost 0 = {0}";
+Goalw [atMost_def] "atMost 0 = {0}";
 by (Simp_tac 1);
 qed "atMost_0";
 Addsimps [atMost_0];
 
-goalw thy [atMost_def] "atMost (Suc k) = insert (Suc k) (atMost k)";
+Goalw [atMost_def] "atMost (Suc k) = insert (Suc k) (atMost k)";
 by (simp_tac (simpset() addsimps [less_Suc_eq, le_eq_less_or_eq]) 1);
 by (Blast_tac 1);
 qed "atMost_Suc";
 
-goal thy "(UN m. atMost m) = UNIV";
+Goal "(UN m. atMost m) = UNIV";
 by (Blast_tac 1);
 qed "UN_atMost_UNIV";
 
-goalw thy [atMost_def, le_def]
+Goalw [atMost_def, le_def]
     "Compl(atMost k) = greaterThan k";
 by (Blast_tac 1);
 qed "Compl_atMost";
@@ -139,7 +139,7 @@
 
 (*** Combined properties ***)
 
-goal thy "atMost n Int atLeast n = {n}";
+Goal "atMost n Int atLeast n = {n}";
 by (blast_tac (claset() addIs [le_anti_sym]) 1);
 qed "atMost_Int_atLeast";