src/ZF/ex/Bin.ML
changeset 5068 fb28eaa07e01
parent 4446 097004a470fb
child 5137 60205b0de9b9
--- a/src/ZF/ex/Bin.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Bin.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -19,19 +19,19 @@
 
 (** bin_rec -- by Vset recursion **)
 
-goal Bin.thy "bin_rec(Plus,a,b,h) = a";
+Goal "bin_rec(Plus,a,b,h) = a";
 by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac bin.con_defs);
 by (simp_tac rank_ss 1);
 qed "bin_rec_Plus";
 
-goal Bin.thy "bin_rec(Minus,a,b,h) = b";
+Goal "bin_rec(Minus,a,b,h) = b";
 by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac bin.con_defs);
 by (simp_tac rank_ss 1);
 qed "bin_rec_Minus";
 
-goal Bin.thy "bin_rec(Bcons(w,x),a,b,h) = h(w, x, bin_rec(w,a,b,h))";
+Goal "bin_rec(Bcons(w,x),a,b,h) = h(w, x, bin_rec(w,a,b,h))";
 by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac bin.con_defs);
 by (simp_tac rank_ss 1);
@@ -72,23 +72,23 @@
 fun bin_recs def = map standard
         ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]);
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus";
+Goalw [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Plus_0";
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,1) = Bcons(Plus,1)";
+Goalw [norm_Bcons_def] "norm_Bcons(Plus,1) = Bcons(Plus,1)";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Plus_1";
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,0) = Bcons(Minus,0)";
+Goalw [norm_Bcons_def] "norm_Bcons(Minus,0) = Bcons(Minus,0)";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Minus_0";
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,1) = Minus";
+Goalw [norm_Bcons_def] "norm_Bcons(Minus,1) = Minus";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Minus_1";
 
-goalw Bin.thy [norm_Bcons_def]
+Goalw [norm_Bcons_def]
     "norm_Bcons(Bcons(w,x),b) = Bcons(Bcons(w,x),b)";
 by (asm_simp_tac (simpset() addsimps bin.case_eqns) 1);
 qed "norm_Bcons_Bcons";
@@ -101,41 +101,41 @@
 
 val bin_typechecks0 = bin_rec_type :: bin.intrs;
 
-goalw Bin.thy [integ_of_bin_def]
+Goalw [integ_of_bin_def]
     "!!w. w: bin ==> integ_of_bin(w) : integ";
 by (typechk_tac (bin_typechecks0@integ_typechecks@
                  nat_typechecks@[bool_into_nat]));
 qed "integ_of_bin_type";
 
-goalw Bin.thy [norm_Bcons_def]
+Goalw [norm_Bcons_def]
     "!!w. [| w: bin; b: bool |] ==> norm_Bcons(w,b) : bin";
 by (etac bin.elim 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps bin.case_eqns)));
 by (typechk_tac (bin_typechecks0@bool_typechecks));
 qed "norm_Bcons_type";
 
-goalw Bin.thy [bin_succ_def]
+Goalw [bin_succ_def]
     "!!w. w: bin ==> bin_succ(w) : bin";
 by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
 qed "bin_succ_type";
 
-goalw Bin.thy [bin_pred_def]
+Goalw [bin_pred_def]
     "!!w. w: bin ==> bin_pred(w) : bin";
 by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
 qed "bin_pred_type";
 
-goalw Bin.thy [bin_minus_def]
+Goalw [bin_minus_def]
     "!!w. w: bin ==> bin_minus(w) : bin";
 by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks));
 qed "bin_minus_type";
 
-goalw Bin.thy [bin_add_def]
+Goalw [bin_add_def]
     "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin";
 by (typechk_tac ([norm_Bcons_type, bin_succ_type, bin_pred_type]@
                  bin_typechecks0@ bool_typechecks@ZF_typechecks));
 qed "bin_add_type";
 
-goalw Bin.thy [bin_mult_def]
+Goalw [bin_mult_def]
     "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin";
 by (typechk_tac ([norm_Bcons_type, bin_minus_type, bin_add_type]@
                  bin_typechecks0@ bool_typechecks));
@@ -177,7 +177,7 @@
 
 
 (*norm_Bcons preserves the integer value of its argument*)
-goal Bin.thy
+Goal
     "!!w. [| w: bin; b: bool |] ==>     \
 \         integ_of_bin(norm_Bcons(w,b)) = integ_of_bin(Bcons(w,b))";
 by (etac bin.elim 1);
@@ -186,7 +186,7 @@
 by (ALLGOALS (asm_simp_tac (simpset() addsimps (norm_Bcons_simps))));
 qed "integ_of_bin_norm_Bcons";
 
-goal Bin.thy
+Goal
     "!!w. w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)";
 by (etac bin.induct 1);
 by (Simp_tac 1);
@@ -196,7 +196,7 @@
     (asm_simp_tac (simpset() addsimps integ_of_bin_norm_Bcons::zadd_ac)));
 qed "integ_of_bin_succ";
 
-goal Bin.thy
+Goal
     "!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)";
 by (etac bin.induct 1);
 by (Simp_tac 1);
@@ -214,7 +214,7 @@
 Addsimps (bin_recs bin_minus_def @
 	  [integ_of_bin_succ, integ_of_bin_pred]);
 
-goal Bin.thy
+Goal
     "!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)";
 by (etac bin.induct 1);
 by (Simp_tac 1);
@@ -227,23 +227,23 @@
 
 (*** bin_add: binary addition ***)
 
-goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
+Goalw [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
 by (Asm_simp_tac 1);
 qed "bin_add_Plus";
 
-goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
+Goalw [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
 by (Asm_simp_tac 1);
 qed "bin_add_Minus";
 
-goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Plus) = Bcons(v,x)";
+Goalw [bin_add_def] "bin_add(Bcons(v,x),Plus) = Bcons(v,x)";
 by (Simp_tac 1);
 qed "bin_add_Bcons_Plus";
 
-goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Minus) = bin_pred(Bcons(v,x))";
+Goalw [bin_add_def] "bin_add(Bcons(v,x),Minus) = bin_pred(Bcons(v,x))";
 by (Simp_tac 1);
 qed "bin_add_Bcons_Minus";
 
-goalw Bin.thy [bin_add_def]
+Goalw [bin_add_def]
     "!!w y. [| w: bin;  y: bool |] ==> \
 \           bin_add(Bcons(v,x), Bcons(w,y)) = \
 \           norm_Bcons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)";
@@ -257,7 +257,7 @@
 
 Addsimps [bool_subset_nat RS subsetD];
 
-goal Bin.thy
+Goal
     "!!v. v: bin ==> \
 \         ALL w: bin. integ_of_bin(bin_add(v,w)) = \
 \                     integ_of_bin(v) $+ integ_of_bin(w)";
@@ -299,19 +299,19 @@
 val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def;
 val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def;
 
-goal Bin.thy "bin_succ(Bcons(w,1)) = Bcons(bin_succ(w), 0)";
+Goal "bin_succ(Bcons(w,1)) = Bcons(bin_succ(w), 0)";
 by (Simp_tac 1);
 qed "bin_succ_Bcons1";
 
-goal Bin.thy "bin_succ(Bcons(w,0)) = norm_Bcons(w,1)";
+Goal "bin_succ(Bcons(w,0)) = norm_Bcons(w,1)";
 by (Simp_tac 1);
 qed "bin_succ_Bcons0";
 
-goal Bin.thy "bin_pred(Bcons(w,1)) = norm_Bcons(w,0)";
+Goal "bin_pred(Bcons(w,1)) = norm_Bcons(w,0)";
 by (Simp_tac 1);
 qed "bin_pred_Bcons1";
 
-goal Bin.thy "bin_pred(Bcons(w,0)) = Bcons(bin_pred(w), 1)";
+Goal "bin_pred(Bcons(w,0)) = Bcons(bin_pred(w), 1)";
 by (Simp_tac 1);
 qed "bin_pred_Bcons0";
 
@@ -319,29 +319,29 @@
 
 val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def;
 
-goal Bin.thy "bin_minus(Bcons(w,1)) = bin_pred(Bcons(bin_minus(w), 0))";
+Goal "bin_minus(Bcons(w,1)) = bin_pred(Bcons(bin_minus(w), 0))";
 by (Simp_tac 1);
 qed "bin_minus_Bcons1";
 
-goal Bin.thy "bin_minus(Bcons(w,0)) = Bcons(bin_minus(w), 0)";
+Goal "bin_minus(Bcons(w,0)) = Bcons(bin_minus(w), 0)";
 by (Simp_tac 1);
 qed "bin_minus_Bcons0";
 
 (** extra rules for bin_add **)
 
-goal Bin.thy 
+Goal 
     "!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,1)) = \
 \                    norm_Bcons(bin_add(v, bin_succ(w)), 0)";
 by (Asm_simp_tac 1);
 qed "bin_add_Bcons_Bcons11";
 
-goal Bin.thy 
+Goal 
     "!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,0)) =  \
 \                    norm_Bcons(bin_add(v,w), 1)";
 by (Asm_simp_tac 1);
 qed "bin_add_Bcons_Bcons10";
 
-goal Bin.thy 
+Goal 
     "!!w y. [| w: bin;  y: bool |] ==> \
 \           bin_add(Bcons(v,0), Bcons(w,y)) = norm_Bcons(bin_add(v,w), y)";
 by (Asm_simp_tac 1);
@@ -351,12 +351,12 @@
 
 val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def;
 
-goal Bin.thy
+Goal
     "bin_mult(Bcons(v,1), w) = bin_add(norm_Bcons(bin_mult(v,w),0), w)";
 by (Simp_tac 1);
 qed "bin_mult_Bcons1";
 
-goal Bin.thy "bin_mult(Bcons(v,0), w) = norm_Bcons(bin_mult(v,w),0)";
+Goal "bin_mult(Bcons(v,0), w) = norm_Bcons(bin_mult(v,w),0)";
 by (Simp_tac 1);
 qed "bin_mult_Bcons0";
 
@@ -387,63 +387,63 @@
 set proof_timing;
 (*All runtimes below are on a SPARCserver 10*)
 
-goal Bin.thy "#13  $+  #19 = #32";
+Goal "#13  $+  #19 = #32";
 by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
 result();
 
 bin_add(binary_of_int 13, binary_of_int 19);
 
-goal Bin.thy "#1234  $+  #5678 = #6912";
+Goal "#1234  $+  #5678 = #6912";
 by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
 result();
 
 bin_add(binary_of_int 1234, binary_of_int 5678);
 
-goal Bin.thy "#1359  $+  #~2468 = #~1109";
+Goal "#1359  $+  #~2468 = #~1109";
 by (simp_tac bin_comp_ss 1);    (*1.2 secs*)
 result();
 
 bin_add(binary_of_int 1359, binary_of_int ~2468);
 
-goal Bin.thy "#93746  $+  #~46375 = #47371";
+Goal "#93746  $+  #~46375 = #47371";
 by (simp_tac bin_comp_ss 1);    (*1.9 secs*)
 result();
 
 bin_add(binary_of_int 93746, binary_of_int ~46375);
 
-goal Bin.thy "$~ #65745 = #~65745";
+Goal "$~ #65745 = #~65745";
 by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
 result();
 
 bin_minus(binary_of_int 65745);
 
 (* negation of ~54321 *)
-goal Bin.thy "$~ #~54321 = #54321";
+Goal "$~ #~54321 = #54321";
 by (simp_tac bin_comp_ss 1);    (*0.5 secs*)
 result();
 
 bin_minus(binary_of_int ~54321);
 
-goal Bin.thy "#13  $*  #19 = #247";
+Goal "#13  $*  #19 = #247";
 by (simp_tac bin_comp_ss 1);    (*0.7 secs*)
 result();
 
 bin_mult(binary_of_int 13, binary_of_int 19);
 
-goal Bin.thy "#~84  $*  #51 = #~4284";
+Goal "#~84  $*  #51 = #~4284";
 by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
 result();
 
 bin_mult(binary_of_int ~84, binary_of_int 51);
 
 (*The worst case for 8-bit operands *)
-goal Bin.thy "#255  $*  #255 = #65025";
+Goal "#255  $*  #255 = #65025";
 by (simp_tac bin_comp_ss 1);    (*4.3 secs*)
 result();
 
 bin_mult(binary_of_int 255, binary_of_int 255);
 
-goal Bin.thy "#1359  $*  #~2468 = #~3354012";
+Goal "#1359  $*  #~2468 = #~3354012";
 by (simp_tac bin_comp_ss 1);    (*6.1 secs*)
 result();