# HG changeset patch # User wenzelm # Date 898529206 -7200 # Node ID 3ea049f7979dc7ba220e95c7cc75e0a5b73d7f03 # Parent fb28eaa07e01cc0372e17d8a9d8eaebcdd2f756b isatool fixgoal; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Arith.ML --- a/src/HOL/Arith.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Arith.ML Mon Jun 22 17:26:46 1998 +0200 @@ -28,7 +28,7 @@ (* Could be (and is, below) generalized in various ways; However, none of the generalizations are currently in the simpset, and I dread to think what happens if I put them in *) -goal thy "!!n. 0 < n ==> Suc(n-1) = n"; +Goal "!!n. 0 < n ==> Suc(n-1) = n"; by (asm_simp_tac (simpset() addsplits [split_nat_case]) 1); qed "Suc_pred"; Addsimps [Suc_pred]; @@ -63,25 +63,25 @@ (*Addition is an AC-operator*) val add_ac = [add_assoc, add_commute, add_left_commute]; -goal thy "!!k::nat. (k + m = k + n) = (m=n)"; +Goal "!!k::nat. (k + m = k + n) = (m=n)"; by (induct_tac "k" 1); by (Simp_tac 1); by (Asm_simp_tac 1); qed "add_left_cancel"; -goal thy "!!k::nat. (m + k = n + k) = (m=n)"; +Goal "!!k::nat. (m + k = n + k) = (m=n)"; by (induct_tac "k" 1); by (Simp_tac 1); by (Asm_simp_tac 1); qed "add_right_cancel"; -goal thy "!!k::nat. (k + m <= k + n) = (m<=n)"; +Goal "!!k::nat. (k + m <= k + n) = (m<=n)"; by (induct_tac "k" 1); by (Simp_tac 1); by (Asm_simp_tac 1); qed "add_left_cancel_le"; -goal thy "!!k::nat. (k + m < k + n) = (m m+(n-(Suc k)) = (m+n)-(Suc k)" *) -goal thy "!!n. 0 m + (n-1) = (m+n)-1"; +Goal "!!n. 0 m + (n-1) = (m+n)-1"; by (exhaust_tac "m" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_Suc] addsplits [split_nat_case]))); @@ -121,7 +121,7 @@ (**** Additional theorems about "less than" ****) -goal thy "i (EX k. j = Suc(i+k))"; +Goal "i (EX k. j = Suc(i+k))"; by (induct_tac "j" 1); by (Simp_tac 1); by (blast_tac (claset() addSEs [less_SucE] @@ -131,21 +131,21 @@ (* [| i Q |] ==> Q *) bind_thm ("less_natE", lemma RS mp RS exE); -goal thy "!!m. m (? k. n=Suc(m+k))"; +Goal "!!m. m (? k. n=Suc(m+k))"; by (induct_tac "n" 1); by (ALLGOALS (simp_tac (simpset() addsimps [less_Suc_eq]))); by (blast_tac (claset() addSEs [less_SucE] addSIs [add_0_right RS sym, add_Suc_right RS sym]) 1); qed_spec_mp "less_eq_Suc_add"; -goal thy "n <= ((m + n)::nat)"; +Goal "n <= ((m + n)::nat)"; by (induct_tac "m" 1); by (ALLGOALS Simp_tac); by (etac le_trans 1); by (rtac (lessI RS less_imp_le) 1); qed "le_add2"; -goal thy "n <= ((n + m)::nat)"; +Goal "n <= ((n + m)::nat)"; by (simp_tac (simpset() addsimps add_ac) 1); by (rtac le_add2 1); qed "le_add1"; @@ -165,49 +165,49 @@ (*"i < j ==> i < m+j"*) bind_thm ("trans_less_add2", le_add2 RSN (2,less_le_trans)); -goal thy "!!i. i+j < (k::nat) ==> i i m <= n+k"; +Goal "!!k::nat. m <= n ==> m <= n+k"; by (etac le_trans 1); by (rtac le_add1 1); qed "le_imp_add_le"; -goal thy "!!k::nat. m < n ==> m < n+k"; +Goal "!!k::nat. m < n ==> m < n+k"; by (etac less_le_trans 1); by (rtac le_add1 1); qed "less_imp_add_less"; -goal thy "m+k<=n --> m<=(n::nat)"; +Goal "m+k<=n --> m<=(n::nat)"; by (induct_tac "k" 1); by (ALLGOALS Asm_simp_tac); by (blast_tac (claset() addDs [Suc_leD]) 1); qed_spec_mp "add_leD1"; -goal thy "!!n::nat. m+k<=n ==> k<=n"; +Goal "!!n::nat. m+k<=n ==> k<=n"; by (full_simp_tac (simpset() addsimps [add_commute]) 1); by (etac add_leD1 1); qed_spec_mp "add_leD2"; -goal thy "!!n::nat. m+k<=n ==> m<=n & k<=n"; +Goal "!!n::nat. m+k<=n ==> m<=n & k<=n"; by (blast_tac (claset() addDs [add_leD1, add_leD2]) 1); bind_thm ("add_leE", result() RS conjE); -goal thy "!!k l::nat. [| k m m i + k < j + k"; +Goal "!!i j k::nat. i < j ==> i + k < j + k"; by (induct_tac "k" 1); by (ALLGOALS Asm_simp_tac); qed "add_less_mono1"; (*strict, in both arguments*) -goal thy "!!i j k::nat. [|i < j; k < l|] ==> i + k < j + l"; +Goal "!!i j k::nat. [|i < j; k < l|] ==> i + k < j + l"; by (rtac (add_less_mono1 RS less_trans) 1); by (REPEAT (assume_tac 1)); by (induct_tac "j" 1); @@ -244,14 +244,14 @@ qed "less_mono_imp_le_mono"; (*non-strict, in 1st argument*) -goal thy "!!i j k::nat. i<=j ==> i + k <= j + k"; +Goal "!!i j k::nat. i<=j ==> i + k <= j + k"; by (res_inst_tac [("f", "%j. j+k")] less_mono_imp_le_mono 1); by (etac add_less_mono1 1); by (assume_tac 1); qed "add_le_mono1"; (*non-strict, in both arguments*) -goal thy "!!k l::nat. [|i<=j; k<=l |] ==> i + k <= j + l"; +Goal "!!k l::nat. [|i<=j; k<=l |] ==> i + k <= j + l"; by (etac (add_le_mono1 RS le_trans) 1); by (simp_tac (simpset() addsimps [add_commute]) 1); (*j moves to the end because it is free while k, l are bound*) @@ -272,11 +272,11 @@ Addsimps [mult_0_right, mult_Suc_right]; -goal thy "1 * n = n"; +Goal "1 * n = n"; by (Asm_simp_tac 1); qed "mult_1"; -goal thy "n * 1 = n"; +Goal "n * 1 = n"; by (Asm_simp_tac 1); qed "mult_1_right"; @@ -304,14 +304,14 @@ val mult_ac = [mult_assoc,mult_commute,mult_left_commute]; -goal thy "(m*n = 0) = (m=0 | n=0)"; +Goal "(m*n = 0) = (m=0 | n=0)"; by (induct_tac "m" 1); by (induct_tac "n" 2); by (ALLGOALS Asm_simp_tac); qed "mult_is_0"; Addsimps [mult_is_0]; -goal thy "!!m::nat. m <= m*m"; +Goal "!!m::nat. m <= m*m"; by (induct_tac "m" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_assoc RS sym]))); by (etac (le_add2 RSN (2,le_trans)) 1); @@ -326,16 +326,16 @@ Addsimps [diff_self_eq_0]; (*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *) -goal thy "~ m n+(m-n) = (m::nat)"; +Goal "~ m n+(m-n) = (m::nat)"; by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); by (ALLGOALS Asm_simp_tac); qed_spec_mp "add_diff_inverse"; -goal thy "!!m. n<=m ==> n+(m-n) = (m::nat)"; +Goal "!!m. n<=m ==> n+(m-n) = (m::nat)"; by (asm_simp_tac (simpset() addsimps [add_diff_inverse, not_less_iff_le]) 1); qed "le_add_diff_inverse"; -goal thy "!!m. n<=m ==> (m-n)+n = (m::nat)"; +Goal "!!m. n<=m ==> (m-n)+n = (m::nat)"; by (asm_simp_tac (simpset() addsimps [le_add_diff_inverse, add_commute]) 1); qed "le_add_diff_inverse2"; @@ -350,13 +350,13 @@ by (ALLGOALS Asm_simp_tac); qed "Suc_diff_n"; -goal thy "m - n < Suc(m)"; +Goal "m - n < Suc(m)"; by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); by (etac less_SucE 3); by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq]))); qed "diff_less_Suc"; -goal thy "!!m::nat. m - n <= m"; +Goal "!!m::nat. m - n <= m"; by (res_inst_tac [("m","m"), ("n","n")] diff_induct 1); by (ALLGOALS Asm_simp_tac); qed "diff_le_self"; @@ -365,61 +365,61 @@ (* j j-n < k *) bind_thm ("less_imp_diff_less", diff_le_self RS le_less_trans); -goal thy "!!i::nat. i-j-k = i - (j+k)"; +Goal "!!i::nat. i-j-k = i - (j+k)"; by (res_inst_tac [("m","i"),("n","j")] diff_induct 1); by (ALLGOALS Asm_simp_tac); qed "diff_diff_left"; -goal Arith.thy "(Suc m - n) - Suc k = m - n - k"; +Goal "(Suc m - n) - Suc k = m - n - k"; by (simp_tac (simpset() addsimps [diff_diff_left]) 1); qed "Suc_diff_diff"; Addsimps [Suc_diff_diff]; -goal thy "!!n. 0 n - Suc i < n"; +Goal "!!n. 0 n - Suc i < n"; by (res_inst_tac [("n","n")] natE 1); by Safe_tac; by (asm_simp_tac (simpset() addsimps [le_eq_less_Suc RS sym]) 1); qed "diff_Suc_less"; Addsimps [diff_Suc_less]; -goal thy "!!n::nat. m - n <= Suc m - n"; +Goal "!!n::nat. m - n <= Suc m - n"; by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); by (ALLGOALS Asm_simp_tac); qed "diff_le_Suc_diff"; (*This and the next few suggested by Florian Kammueller*) -goal thy "!!i::nat. i-j-k = i-k-j"; +Goal "!!i::nat. i-j-k = i-k-j"; by (simp_tac (simpset() addsimps [diff_diff_left, add_commute]) 1); qed "diff_commute"; -goal thy "!!i j k:: nat. k<=j --> j<=i --> i - (j - k) = i - j + k"; +Goal "!!i j k:: nat. k<=j --> j<=i --> i - (j - k) = i - j + k"; by (res_inst_tac [("m","i"),("n","j")] diff_induct 1); by (ALLGOALS Asm_simp_tac); by (asm_simp_tac (simpset() addsimps [Suc_diff_n, le_imp_less_Suc, le_Suc_eq]) 1); qed_spec_mp "diff_diff_right"; -goal thy "!!i j k:: nat. k<=j --> (i + j) - k = i + (j - k)"; +Goal "!!i j k:: nat. k<=j --> (i + j) - k = i + (j - k)"; by (res_inst_tac [("m","j"),("n","k")] diff_induct 1); by (ALLGOALS Asm_simp_tac); qed_spec_mp "diff_add_assoc"; -goal thy "!!i j k:: nat. k<=j --> (j + i) - k = i + (j - k)"; +Goal "!!i j k:: nat. k<=j --> (j + i) - k = i + (j - k)"; by (asm_simp_tac (simpset() addsimps [add_commute, diff_add_assoc]) 1); qed_spec_mp "diff_add_assoc2"; -goal thy "!!n::nat. (n+m) - n = m"; +Goal "!!n::nat. (n+m) - n = m"; by (induct_tac "n" 1); by (ALLGOALS Asm_simp_tac); qed "diff_add_inverse"; Addsimps [diff_add_inverse]; -goal thy "!!n::nat.(m+n) - n = m"; +Goal "!!n::nat.(m+n) - n = m"; by (simp_tac (simpset() addsimps [diff_add_assoc]) 1); qed "diff_add_inverse2"; Addsimps [diff_add_inverse2]; -goal thy "!!i j k::nat. i<=j ==> (j-i=k) = (j=k+i)"; +Goal "!!i j k::nat. i<=j ==> (j-i=k) = (j=k+i)"; by Safe_tac; by (ALLGOALS Asm_simp_tac); qed "le_imp_diff_is_add"; @@ -442,15 +442,15 @@ by (ALLGOALS Asm_simp_tac); qed "less_imp_diff_positive"; -goal thy "Suc(m)-n = (if m (!n. P(Suc(n))--> P(n)) --> P(k-i)"; +Goal "P(k) --> (!n. P(Suc(n))--> P(n)) --> P(k-i)"; by (res_inst_tac [("m","k"),("n","i")] diff_induct 1); by (ALLGOALS (Clarify_tac THEN' Simp_tac THEN' TRY o Blast_tac)); qed "zero_induct_lemma"; @@ -461,20 +461,20 @@ by (REPEAT (ares_tac ([impI,allI]@prems) 1)); qed "zero_induct"; -goal thy "!!k::nat. (k+m) - (k+n) = m - n"; +Goal "!!k::nat. (k+m) - (k+n) = m - n"; by (induct_tac "k" 1); by (ALLGOALS Asm_simp_tac); qed "diff_cancel"; Addsimps [diff_cancel]; -goal thy "!!m::nat. (m+k) - (n+k) = m - n"; +Goal "!!m::nat. (m+k) - (n+k) = m - n"; val add_commute_k = read_instantiate [("n","k")] add_commute; by (asm_simp_tac (simpset() addsimps ([add_commute_k])) 1); qed "diff_cancel2"; Addsimps [diff_cancel2]; (*From Clemens Ballarin*) -goal thy "!!n::nat. [| k<=n; n<=m |] ==> (m-k) - (n-k) = m-n"; +Goal "!!n::nat. [| k<=n; n<=m |] ==> (m-k) - (n-k) = m-n"; by (subgoal_tac "k<=n --> n<=m --> (m-k) - (n-k) = m-n" 1); by (Asm_full_simp_tac 1); by (induct_tac "k" 1); @@ -489,7 +489,7 @@ addsimps [Suc_diff_n RS sym, le_eq_less_Suc]) 1); qed "diff_right_cancel"; -goal thy "!!n::nat. n - (n+m) = 0"; +Goal "!!n::nat. n - (n+m) = 0"; by (induct_tac "n" 1); by (ALLGOALS Asm_simp_tac); qed "diff_add_0"; @@ -497,12 +497,12 @@ (** Difference distributes over multiplication **) -goal thy "!!m::nat. (m - n) * k = (m * k) - (n * k)"; +Goal "!!m::nat. (m - n) * k = (m * k) - (n * k)"; by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); by (ALLGOALS Asm_simp_tac); qed "diff_mult_distrib" ; -goal thy "!!m::nat. k * (m - n) = (k * m) - (k * n)"; +Goal "!!m::nat. k * (m - n) = (k * m) - (k * n)"; val mult_commute_k = read_instantiate [("m","k")] mult_commute; by (simp_tac (simpset() addsimps [diff_mult_distrib, mult_commute_k]) 1); qed "diff_mult_distrib2" ; @@ -511,13 +511,13 @@ (*** Monotonicity of Multiplication ***) -goal thy "!!i::nat. i<=j ==> i*k<=j*k"; +Goal "!!i::nat. i<=j ==> i*k<=j*k"; by (induct_tac "k" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono]))); qed "mult_le_mono1"; (*<=monotonicity, BOTH arguments*) -goal thy "!!i::nat. [| i<=j; k<=l |] ==> i*k<=j*l"; +Goal "!!i::nat. [| i<=j; k<=l |] ==> i*k<=j*l"; by (etac (mult_le_mono1 RS le_trans) 1); by (rtac le_trans 1); by (stac mult_commute 2); @@ -526,26 +526,26 @@ qed "mult_le_mono"; (*strict, in 1st argument; proof is by induction on k>0*) -goal thy "!!i::nat. [| i k*i < k*j"; +Goal "!!i::nat. [| i k*i < k*j"; by (eres_inst_tac [("i","0")] less_natE 1); by (Asm_simp_tac 1); by (induct_tac "x" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_less_mono]))); qed "mult_less_mono2"; -goal thy "!!i::nat. [| i i*k < j*k"; +Goal "!!i::nat. [| i i*k < j*k"; by (dtac mult_less_mono2 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [mult_commute]))); qed "mult_less_mono1"; -goal thy "(0 < m*n) = (0 (m*k < n*k) = (m (m*k < n*k) = (m (k*m < k*n) = (m (k*m < k*n) = (m (m*k = n*k) = (m=n)"; +Goal "!!k. 0 (m*k = n*k) = (m=n)"; by (cut_facts_tac [less_linear] 1); by Safe_tac; by (assume_tac 2); @@ -584,13 +584,13 @@ by (ALLGOALS Asm_full_simp_tac); qed "mult_cancel2"; -goal thy "!!k. 0 (k*m = k*n) = (m=n)"; +Goal "!!k. 0 (k*m = k*n) = (m=n)"; by (dtac mult_cancel2 1); by (asm_full_simp_tac (simpset() addsimps [mult_commute]) 1); qed "mult_cancel1"; Addsimps [mult_cancel1, mult_cancel2]; -goal thy "(Suc k * m = Suc k * n) = (m = n)"; +Goal "(Suc k * m = Suc k * n) = (m = n)"; by (rtac mult_cancel1 1); by (Simp_tac 1); qed "Suc_mult_cancel1"; @@ -598,7 +598,7 @@ (** Lemma for gcd **) -goal thy "!!m n. m = m*n ==> n=1 | m=0"; +Goal "!!m n. m = m*n ==> n=1 | m=0"; by (dtac sym 1); by (rtac disjCI 1); by (rtac nat_less_cases 1 THEN assume_tac 2); @@ -609,7 +609,7 @@ (*** Subtraction laws -- mostly from Clemens Ballarin ***) -goal thy "!! a b c::nat. [| a < b; c <= a |] ==> a-c < b-c"; +Goal "!! a b c::nat. [| a < b; c <= a |] ==> a-c < b-c"; by (subgoal_tac "c+(a-c) < c+(b-c)" 1); by (Full_simp_tac 1); by (subgoal_tac "c <= b" 1); @@ -617,29 +617,29 @@ by (Asm_simp_tac 1); qed "diff_less_mono"; -goal thy "!! a b c::nat. a+b < c ==> a < c-b"; +Goal "!! a b c::nat. a+b < c ==> a < c-b"; by (dtac diff_less_mono 1); by (rtac le_add2 1); by (Asm_full_simp_tac 1); qed "add_less_imp_less_diff"; -goal thy "!! n. n <= m ==> Suc m - n = Suc (m - n)"; +Goal "!! n. n <= m ==> Suc m - n = Suc (m - n)"; by (asm_full_simp_tac (simpset() addsimps [Suc_diff_n, le_eq_less_Suc]) 1); qed "Suc_diff_le"; -goal thy "!! n. Suc i <= n ==> Suc (n - Suc i) = n - i"; +Goal "!! n. Suc i <= n ==> Suc (n - Suc i) = n - i"; by (asm_full_simp_tac (simpset() addsimps [Suc_diff_n RS sym, le_eq_less_Suc]) 1); qed "Suc_diff_Suc"; -goal thy "!! i::nat. i <= n ==> n - (n - i) = i"; +Goal "!! i::nat. i <= n ==> n - (n - i) = i"; by (etac rev_mp 1); by (res_inst_tac [("m","n"),("n","i")] diff_induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [Suc_diff_le]))); qed "diff_diff_cancel"; Addsimps [diff_diff_cancel]; -goal thy "!!k::nat. k <= n ==> m <= n + m - k"; +Goal "!!k::nat. k <= n ==> m <= n + m - k"; by (etac rev_mp 1); by (res_inst_tac [("m", "k"), ("n", "n")] diff_induct 1); by (Simp_tac 1); @@ -647,7 +647,7 @@ by (Simp_tac 1); qed "le_add_diff"; -goal Arith.thy "!!i::nat. 0 j j+k-i < k"; +Goal "!!i::nat. 0 j j+k-i < k"; by (res_inst_tac [("m","j"),("n","i")] diff_induct 1); by (ALLGOALS Asm_simp_tac); qed_spec_mp "add_diff_less"; @@ -657,14 +657,14 @@ (** (Anti)Monotonicity of subtraction -- by Stefan Merz **) (* Monotonicity of subtraction in first argument *) -goal thy "!!n::nat. m<=n --> (m-l) <= (n-l)"; +Goal "!!n::nat. m<=n --> (m-l) <= (n-l)"; by (induct_tac "n" 1); by (Simp_tac 1); by (simp_tac (simpset() addsimps [le_Suc_eq]) 1); by (blast_tac (claset() addIs [diff_le_Suc_diff, le_trans]) 1); qed_spec_mp "diff_le_mono"; -goal thy "!!n::nat. m<=n ==> (l-n) <= (l-m)"; +Goal "!!n::nat. m<=n ==> (l-n) <= (l-m)"; by (induct_tac "l" 1); by (Simp_tac 1); by (case_tac "n <= l" 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Group/Group.ML --- a/src/HOL/AxClasses/Group/Group.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Group/Group.ML Mon Jun 22 17:26:46 1998 +0200 @@ -9,7 +9,7 @@ fun ssub r = standard (r RS ssubst); -goal thy "x * inverse x = (1::'a::group)"; +Goal "x * inverse x = (1::'a::group)"; by (rtac (sub left_unit) 1); back(); by (rtac (sub assoc) 1); @@ -26,7 +26,7 @@ qed "right_inverse"; -goal thy "x * 1 = (x::'a::group)"; +Goal "x * 1 = (x::'a::group)"; by (rtac (sub left_inverse) 1); by (rtac (sub assoc) 1); by (rtac (ssub right_inverse) 1); @@ -35,7 +35,7 @@ qed "right_unit"; -goal thy "e * x = x --> e = (1::'a::group)"; +Goal "e * x = x --> e = (1::'a::group)"; by (rtac impI 1); by (rtac (sub right_unit) 1); back(); @@ -47,7 +47,7 @@ qed "strong_one_unit"; -goal thy "EX! e. ALL x. e * x = (x::'a::group)"; +Goal "EX! e. ALL x. e * x = (x::'a::group)"; by (rtac ex1I 1); by (rtac allI 1); by (rtac left_unit 1); @@ -58,7 +58,7 @@ qed "ex1_unit"; -goal thy "ALL x. EX! e. e * x = (x::'a::group)"; +Goal "ALL x. EX! e. e * x = (x::'a::group)"; by (rtac allI 1); by (rtac ex1I 1); by (rtac left_unit 1); @@ -67,7 +67,7 @@ qed "ex1_unit'"; -goal thy "inj (op * (x::'a::group))"; +Goal "inj (op * (x::'a::group))"; by (rtac injI 1); by (rtac (sub left_unit) 1); back(); @@ -82,7 +82,7 @@ qed "inj_times"; -goal thy "y * x = 1 --> y = inverse (x::'a::group)"; +Goal "y * x = 1 --> y = inverse (x::'a::group)"; by (rtac impI 1); by (rtac (sub right_unit) 1); back(); @@ -98,7 +98,7 @@ qed "one_inverse"; -goal thy "ALL x. EX! y. y * x = (1::'a::group)"; +Goal "ALL x. EX! y. y * x = (1::'a::group)"; by (rtac allI 1); by (rtac ex1I 1); by (rtac left_inverse 1); @@ -108,7 +108,7 @@ qed "ex1_inverse"; -goal thy "inverse (x * y) = inverse y * inverse (x::'a::group)"; +Goal "inverse (x * y) = inverse y * inverse (x::'a::group)"; by (rtac sym 1); by (rtac mp 1); by (rtac one_inverse 1); @@ -122,7 +122,7 @@ qed "inverse_times"; -goal thy "inverse (inverse x) = (x::'a::group)"; +Goal "inverse (inverse x) = (x::'a::group)"; by (rtac sym 1); by (rtac (one_inverse RS mp) 1); by (rtac (ssub right_inverse) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Group/GroupDefs.ML --- a/src/HOL/AxClasses/Group/GroupDefs.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Group/GroupDefs.ML Mon Jun 22 17:26:46 1998 +0200 @@ -2,24 +2,24 @@ (*this is really overkill - should be rather proven 'inline'*) -goalw thy [times_bool_def] "(x * y) * z = x * (y * (z::bool))"; +Goalw [times_bool_def] "(x * y) * z = x * (y * (z::bool))"; by (Fast_tac 1); qed "bool_assoc"; -goalw thy [times_bool_def, one_bool_def] "1 * x = (x::bool)"; +Goalw [times_bool_def, one_bool_def] "1 * x = (x::bool)"; by (Fast_tac 1); qed "bool_left_unit"; -goalw thy [times_bool_def, one_bool_def] "x * 1 = (x::bool)"; +Goalw [times_bool_def, one_bool_def] "x * 1 = (x::bool)"; by (Fast_tac 1); qed "bool_right_unit"; -goalw thy [times_bool_def, inverse_bool_def, one_bool_def] +Goalw [times_bool_def, inverse_bool_def, one_bool_def] "inverse(x) * x = (1::bool)"; by (Fast_tac 1); qed "bool_left_inverse"; -goalw thy [times_bool_def] "x * y = (y * (x::bool))"; +Goalw [times_bool_def] "x * y = (y * (x::bool))"; by (Fast_tac 1); qed "bool_commut"; @@ -28,59 +28,59 @@ val prod_ss = simpset_of Prod.thy; -goalw thy [times_prod_def] +Goalw [times_prod_def] "(x * y) * z = x * (y * (z::'a::semigroup*'b::semigroup))"; by (simp_tac (prod_ss addsimps [assoc]) 1); qed "prod_assoc"; -goalw thy [times_prod_def, one_prod_def] "1 * x = (x::'a::monoid*'b::monoid)"; +Goalw [times_prod_def, one_prod_def] "1 * x = (x::'a::monoid*'b::monoid)"; by (simp_tac (prod_ss addsimps [Monoid.left_unit]) 1); by (rtac (surjective_pairing RS sym) 1); qed "prod_left_unit"; -goalw thy [times_prod_def, one_prod_def] "x * 1 = (x::'a::monoid*'b::monoid)"; +Goalw [times_prod_def, one_prod_def] "x * 1 = (x::'a::monoid*'b::monoid)"; by (simp_tac (prod_ss addsimps [Monoid.right_unit]) 1); by (rtac (surjective_pairing RS sym) 1); qed "prod_right_unit"; -goalw thy [times_prod_def, inverse_prod_def, one_prod_def] +Goalw [times_prod_def, inverse_prod_def, one_prod_def] "inverse x * x = (1::'a::group*'b::group)"; by (simp_tac (prod_ss addsimps [left_inverse]) 1); qed "prod_left_inverse"; -goalw thy [times_prod_def] "x * y = y * (x::'a::agroup*'b::agroup)"; +Goalw [times_prod_def] "x * y = y * (x::'a::agroup*'b::agroup)"; by (simp_tac (prod_ss addsimps [commut]) 1); qed "prod_commut"; (* function spaces *) -goalw thy [times_fun_def] "(x * y) * z = x * (y * (z::'a => 'b::semigroup))"; +Goalw [times_fun_def] "(x * y) * z = x * (y * (z::'a => 'b::semigroup))"; by (stac expand_fun_eq 1); by (rtac allI 1); by (rtac assoc 1); qed "fun_assoc"; -goalw thy [times_fun_def, one_fun_def] "1 * x = (x::'a => 'b::monoid)"; +Goalw [times_fun_def, one_fun_def] "1 * x = (x::'a => 'b::monoid)"; by (stac expand_fun_eq 1); by (rtac allI 1); by (rtac Monoid.left_unit 1); qed "fun_left_unit"; -goalw thy [times_fun_def, one_fun_def] "x * 1 = (x::'a => 'b::monoid)"; +Goalw [times_fun_def, one_fun_def] "x * 1 = (x::'a => 'b::monoid)"; by (stac expand_fun_eq 1); by (rtac allI 1); by (rtac Monoid.right_unit 1); qed "fun_right_unit"; -goalw thy [times_fun_def, inverse_fun_def, one_fun_def] +Goalw [times_fun_def, inverse_fun_def, one_fun_def] "inverse x * x = (1::'a => 'b::group)"; by (stac expand_fun_eq 1); by (rtac allI 1); by (rtac left_inverse 1); qed "fun_left_inverse"; -goalw thy [times_fun_def] "x * y = y * (x::'a => 'b::agroup)"; +Goalw [times_fun_def] "x * y = y * (x::'a => 'b::agroup)"; by (stac expand_fun_eq 1); by (rtac allI 1); by (rtac commut 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Lattice/CLattice.ML --- a/src/HOL/AxClasses/Lattice/CLattice.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Lattice/CLattice.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,11 +6,11 @@ (* unique existence *) -goalw thy [Inf_def] "is_Inf A (Inf A)"; +Goalw [Inf_def] "is_Inf A (Inf A)"; by (rtac (ex_Inf RS spec RS selectI1) 1); qed "Inf_is_Inf"; -goal thy "is_Inf A inf --> Inf A = inf"; +Goal "is_Inf A inf --> Inf A = inf"; by (rtac impI 1); by (rtac (is_Inf_uniq RS mp) 1); by (rtac conjI 1); @@ -18,7 +18,7 @@ by (assume_tac 1); qed "Inf_uniq"; -goalw thy [Ex1_def] "ALL A. EX! inf::'a::clattice. is_Inf A inf"; +Goalw [Ex1_def] "ALL A. EX! inf::'a::clattice. is_Inf A inf"; by Safe_tac; by (Step_tac 1); by (Step_tac 1); @@ -28,11 +28,11 @@ qed "ex1_Inf"; -goalw thy [Sup_def] "is_Sup A (Sup A)"; +Goalw [Sup_def] "is_Sup A (Sup A)"; by (rtac (ex_Sup RS spec RS selectI1) 1); qed "Sup_is_Sup"; -goal thy "is_Sup A sup --> Sup A = sup"; +Goal "is_Sup A sup --> Sup A = sup"; by (rtac impI 1); by (rtac (is_Sup_uniq RS mp) 1); by (rtac conjI 1); @@ -40,7 +40,7 @@ by (assume_tac 1); qed "Sup_uniq"; -goalw thy [Ex1_def] "ALL A. EX! sup::'a::clattice. is_Sup A sup"; +Goalw [Ex1_def] "ALL A. EX! sup::'a::clattice. is_Sup A sup"; by Safe_tac; by (Step_tac 1); by (Step_tac 1); @@ -94,7 +94,7 @@ (** minorized Infs / majorized Sups **) -goal thy "(x [= Inf A) = (ALL y:A. x [= y)"; +Goal "(x [= Inf A) = (ALL y:A. x [= y)"; by (rtac iffI 1); (*==>*) by (rtac ballI 1); @@ -106,7 +106,7 @@ by (Fast_tac 1); qed "le_Inf_eq"; -goal thy "(Sup A [= x) = (ALL y:A. y [= x)"; +Goal "(Sup A [= x) = (ALL y:A. y [= x)"; by (rtac iffI 1); (*==>*) by (rtac ballI 1); @@ -123,7 +123,7 @@ (** Subsets and limits **) -goal thy "A <= B --> Inf B [= Inf A"; +Goal "A <= B --> Inf B [= Inf A"; by (rtac impI 1); by (stac le_Inf_eq 1); by (rewtac Ball_def); @@ -133,7 +133,7 @@ by (etac Inf_lb 1); qed "Inf_subset_antimon"; -goal thy "A <= B --> Sup A [= Sup B"; +Goal "A <= B --> Sup A [= Sup B"; by (rtac impI 1); by (stac ge_Sup_eq 1); by (rewtac Ball_def); @@ -146,7 +146,7 @@ (** singleton / empty limits **) -goal thy "Inf {x} = x"; +Goal "Inf {x} = x"; by (rtac (Inf_uniq RS mp) 1); by (rewtac is_Inf_def); by Safe_tac; @@ -154,7 +154,7 @@ by (Fast_tac 1); qed "sing_Inf_eq"; -goal thy "Sup {x} = x"; +Goal "Sup {x} = x"; by (rtac (Sup_uniq RS mp) 1); by (rewtac is_Sup_def); by Safe_tac; @@ -163,7 +163,7 @@ qed "sing_Sup_eq"; -goal thy "Inf {} = Sup {x. True}"; +Goal "Inf {} = Sup {x. True}"; by (rtac (Inf_uniq RS mp) 1); by (rewtac is_Inf_def); by Safe_tac; @@ -173,7 +173,7 @@ by (Fast_tac 1); qed "empty_Inf_eq"; -goal thy "Sup {} = Inf {x. True}"; +Goal "Sup {} = Inf {x. True}"; by (rtac (Sup_uniq RS mp) 1); by (rewtac is_Sup_def); by Safe_tac; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Lattice/LatInsts.ML --- a/src/HOL/AxClasses/Lattice/LatInsts.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Lattice/LatInsts.ML Mon Jun 22 17:26:46 1998 +0200 @@ -2,13 +2,13 @@ open LatInsts; -goal thy "Inf {x, y} = x && y"; +Goal "Inf {x, y} = x && y"; by (rtac (Inf_uniq RS mp) 1); by (stac bin_is_Inf_eq 1); by (rtac inf_is_inf 1); qed "bin_Inf_eq"; -goal thy "Sup {x, y} = x || y"; +Goal "Sup {x, y} = x || y"; by (rtac (Sup_uniq RS mp) 1); by (stac bin_is_Sup_eq 1); by (rtac sup_is_sup 1); @@ -18,7 +18,7 @@ (* Unions and limits *) -goal thy "Inf (A Un B) = Inf A && Inf B"; +Goal "Inf (A Un B) = Inf A && Inf B"; by (rtac (Inf_uniq RS mp) 1); by (rewtac is_Inf_def); by Safe_tac; @@ -39,7 +39,7 @@ by (Fast_tac 1); qed "Inf_Un_eq"; -goal thy "Inf (UN i:A. B i) = Inf {Inf (B i) |i. i:A}"; +Goal "Inf (UN i:A. B i) = Inf {Inf (B i) |i. i:A}"; by (rtac (Inf_uniq RS mp) 1); by (rewtac is_Inf_def); by Safe_tac; @@ -58,7 +58,7 @@ -goal thy "Sup (A Un B) = Sup A || Sup B"; +Goal "Sup (A Un B) = Sup A || Sup B"; by (rtac (Sup_uniq RS mp) 1); by (rewtac is_Sup_def); by Safe_tac; @@ -79,7 +79,7 @@ by (Fast_tac 1); qed "Sup_Un_eq"; -goal thy "Sup (UN i:A. B i) = Sup {Sup (B i) |i. i:A}"; +Goal "Sup (UN i:A. B i) = Sup {Sup (B i) |i. i:A}"; by (rtac (Sup_uniq RS mp) 1); by (rewtac is_Sup_def); by Safe_tac; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Lattice/LatMorph.ML --- a/src/HOL/AxClasses/Lattice/LatMorph.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Lattice/LatMorph.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,7 +4,7 @@ (** monotone functions vs. "&&"- / "||"-semi-morphisms **) -goalw thy [is_mono_def] "is_mono f = (ALL x y. f (x && y) [= f x && f y)"; +Goalw [is_mono_def] "is_mono f = (ALL x y. f (x && y) [= f x && f y)"; by Safe_tac; (*==> (level 1)*) by (stac le_inf_eq 1); @@ -27,7 +27,7 @@ by (assume_tac 1); qed "mono_inf_eq"; -goalw thy [is_mono_def] "is_mono f = (ALL x y. f x || f y [= f (x || y))"; +Goalw [is_mono_def] "is_mono f = (ALL x y. f x || f y [= f (x || y))"; by Safe_tac; (*==> (level 1)*) by (stac ge_sup_eq 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Lattice/LatPreInsts.ML --- a/src/HOL/AxClasses/Lattice/LatPreInsts.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Lattice/LatPreInsts.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,12 +4,12 @@ (** complete lattices **) -goal thy "is_inf x y (Inf {x, y})"; +Goal "is_inf x y (Inf {x, y})"; by (rtac (bin_is_Inf_eq RS subst) 1); by (rtac Inf_is_Inf 1); qed "Inf_is_inf"; -goal thy "is_sup x y (Sup {x, y})"; +Goal "is_sup x y (Sup {x, y})"; by (rtac (bin_is_Sup_eq RS subst) 1); by (rtac Sup_is_Sup 1); qed "Sup_is_sup"; @@ -20,13 +20,13 @@ (* pairs *) -goalw thy [is_inf_def, le_prod_def] "is_inf p q (fst p && fst q, snd p && snd q)"; +Goalw [is_inf_def, le_prod_def] "is_inf p q (fst p && fst q, snd p && snd q)"; by (Simp_tac 1); by Safe_tac; by (REPEAT_FIRST (fn i => resolve_tac [inf_lb1, inf_lb2, inf_ub_lbs] i ORELSE atac i)); qed "prod_is_inf"; -goalw thy [is_sup_def, le_prod_def] "is_sup p q (fst p || fst q, snd p || snd q)"; +Goalw [is_sup_def, le_prod_def] "is_sup p q (fst p || fst q, snd p || snd q)"; by (Simp_tac 1); by Safe_tac; by (REPEAT_FIRST (fn i => resolve_tac [sup_ub1, sup_ub2, sup_lb_ubs] i ORELSE atac i)); @@ -35,7 +35,7 @@ (* functions *) -goalw thy [is_inf_def, le_fun_def] "is_inf f g (%x. f x && g x)"; +Goalw [is_inf_def, le_fun_def] "is_inf f g (%x. f x && g x)"; by Safe_tac; by (rtac inf_lb1 1); by (rtac inf_lb2 1); @@ -43,7 +43,7 @@ by (REPEAT_FIRST (Fast_tac)); qed "fun_is_inf"; -goalw thy [is_sup_def, le_fun_def] "is_sup f g (%x. f x || g x)"; +Goalw [is_sup_def, le_fun_def] "is_sup f g (%x. f x || g x)"; by Safe_tac; by (rtac sup_ub1 1); by (rtac sup_ub2 1); @@ -55,7 +55,7 @@ (** dual lattices **) -goalw thy [is_inf_def, le_dual_def] "is_inf x y (Abs_dual (Rep_dual x || Rep_dual y))"; +Goalw [is_inf_def, le_dual_def] "is_inf x y (Abs_dual (Rep_dual x || Rep_dual y))"; by (stac Abs_dual_inverse' 1); by Safe_tac; by (rtac sup_ub1 1); @@ -65,7 +65,7 @@ by (assume_tac 1); qed "dual_is_inf"; -goalw thy [is_sup_def, le_dual_def] "is_sup x y (Abs_dual (Rep_dual x && Rep_dual y))"; +Goalw [is_sup_def, le_dual_def] "is_sup x y (Abs_dual (Rep_dual x && Rep_dual y))"; by (stac Abs_dual_inverse' 1); by Safe_tac; by (rtac inf_lb1 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Lattice/Lattice.ML --- a/src/HOL/AxClasses/Lattice/Lattice.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Lattice/Lattice.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,11 +6,11 @@ (* unique existence *) -goalw thy [inf_def] "is_inf x y (x && y)"; +Goalw [inf_def] "is_inf x y (x && y)"; by (rtac (ex_inf RS spec RS spec RS selectI1) 1); qed "inf_is_inf"; -goal thy "is_inf x y inf --> x && y = inf"; +Goal "is_inf x y inf --> x && y = inf"; by (rtac impI 1); by (rtac (is_inf_uniq RS mp) 1); by (rtac conjI 1); @@ -18,7 +18,7 @@ by (assume_tac 1); qed "inf_uniq"; -goalw thy [Ex1_def] "ALL x y. EX! inf::'a::lattice. is_inf x y inf"; +Goalw [Ex1_def] "ALL x y. EX! inf::'a::lattice. is_inf x y inf"; by Safe_tac; by (Step_tac 1); by (Step_tac 1); @@ -28,11 +28,11 @@ qed "ex1_inf"; -goalw thy [sup_def] "is_sup x y (x || y)"; +Goalw [sup_def] "is_sup x y (x || y)"; by (rtac (ex_sup RS spec RS spec RS selectI1) 1); qed "sup_is_sup"; -goal thy "is_sup x y sup --> x || y = sup"; +Goal "is_sup x y sup --> x || y = sup"; by (rtac impI 1); by (rtac (is_sup_uniq RS mp) 1); by (rtac conjI 1); @@ -40,7 +40,7 @@ by (assume_tac 1); qed "sup_uniq"; -goalw thy [Ex1_def] "ALL x y. EX! sup::'a::lattice. is_sup x y sup"; +Goalw [Ex1_def] "ALL x y. EX! sup::'a::lattice. is_sup x y sup"; by Safe_tac; by (Step_tac 1); by (Step_tac 1); @@ -57,11 +57,11 @@ rewrite_goals_tac [inf_def, is_inf_def] THEN Fast_tac 1; -goal thy "x && y [= x"; +Goal "x && y [= x"; by tac; qed "inf_lb1"; -goal thy "x && y [= y"; +Goal "x && y [= y"; by tac; qed "inf_lb2"; @@ -76,11 +76,11 @@ rewrite_goals_tac [sup_def, is_sup_def] THEN Fast_tac 1; -goal thy "x [= x || y"; +Goal "x [= x || y"; by tac; qed "sup_ub1"; -goal thy "y [= x || y"; +Goal "y [= x || y"; by tac; qed "sup_ub2"; @@ -95,7 +95,7 @@ (* the Connection Theorems: "[=" expressed via "&&" or "||" *) -goal thy "(x && y = x) = (x [= y)"; +Goal "(x && y = x) = (x [= y)"; by (rtac iffI 1); (*==>*) by (etac subst 1); @@ -109,7 +109,7 @@ by (Fast_tac 1); qed "inf_connect"; -goal thy "(x || y = y) = (x [= y)"; +Goal "(x || y = y) = (x [= y)"; by (rtac iffI 1); (*==>*) by (etac subst 1); @@ -126,7 +126,7 @@ (* minorized infs / majorized sups *) -goal thy "(x [= y && z) = (x [= y & x [= z)"; +Goal "(x [= y && z) = (x [= y & x [= z)"; by (rtac iffI 1); (*==> (level 1)*) by (cut_facts_tac [inf_lb1, inf_lb2] 1); @@ -143,7 +143,7 @@ by (assume_tac 1); qed "le_inf_eq"; -goal thy "(x || y [= z) = (x [= z & y [= z)"; +Goal "(x || y [= z) = (x [= z & y [= z)"; by (rtac iffI 1); (*==> (level 1)*) by (cut_facts_tac [sup_ub1, sup_ub2] 1); @@ -165,7 +165,7 @@ (* associativity *) -goal thy "(x && y) && z = x && (y && z)"; +Goal "(x && y) && z = x && (y && z)"; by (stac (inf_uniq RS mp RS sym) 1); back(); back(); @@ -181,7 +181,7 @@ by (REPEAT_FIRST (rtac inf_is_inf)); qed "inf_assoc"; -goal thy "(x || y) || z = x || (y || z)"; +Goal "(x || y) || z = x || (y || z)"; by (stac (sup_uniq RS mp RS sym) 1); back(); back(); @@ -200,12 +200,12 @@ (* commutativity *) -goalw thy [inf_def] "x && y = y && x"; +Goalw [inf_def] "x && y = y && x"; by (stac (is_inf_commut RS ext) 1); by (rtac refl 1); qed "inf_commut"; -goalw thy [sup_def] "x || y = y || x"; +Goalw [sup_def] "x || y = y || x"; by (stac (is_sup_commut RS ext) 1); by (rtac refl 1); qed "sup_commut"; @@ -213,12 +213,12 @@ (* idempotency *) -goal thy "x && x = x"; +Goal "x && x = x"; by (stac inf_connect 1); by (rtac le_refl 1); qed "inf_idemp"; -goal thy "x || x = x"; +Goal "x || x = x"; by (stac sup_connect 1); by (rtac le_refl 1); qed "sup_idemp"; @@ -226,13 +226,13 @@ (* absorption *) -goal thy "x || (x && y) = x"; +Goal "x || (x && y) = x"; by (stac sup_commut 1); by (stac sup_connect 1); by (rtac inf_lb1 1); qed "sup_inf_absorb"; -goal thy "x && (x || y) = x"; +Goal "x && (x || y) = x"; by (stac inf_connect 1); by (rtac sup_ub1 1); qed "inf_sup_absorb"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Lattice/OrdDefs.ML --- a/src/HOL/AxClasses/Lattice/OrdDefs.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Lattice/OrdDefs.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,13 +6,13 @@ (* pairs *) -goalw thy [le_prod_def] "x [= (x::'a::quasi_order*'b::quasi_order)"; +Goalw [le_prod_def] "x [= (x::'a::quasi_order*'b::quasi_order)"; by (rtac conjI 1); by (rtac le_refl 1); by (rtac le_refl 1); qed "le_prod_refl"; -goalw thy [le_prod_def] "x [= y & y [= z --> x [= (z::'a::quasi_order*'b::quasi_order)"; +Goalw [le_prod_def] "x [= y & y [= z --> x [= (z::'a::quasi_order*'b::quasi_order)"; by Safe_tac; by (etac (conjI RS (le_trans RS mp)) 1); by (assume_tac 1); @@ -20,7 +20,7 @@ by (assume_tac 1); qed "le_prod_trans"; -goalw thy [le_prod_def] "x [= y & y [= x --> x = (y::'a::partial_order*'b::partial_order)"; +Goalw [le_prod_def] "x [= y & y [= x --> x = (y::'a::partial_order*'b::partial_order)"; by Safe_tac; by (stac Pair_fst_snd_eq 1); by (rtac conjI 1); @@ -33,18 +33,18 @@ (* functions *) -goalw thy [le_fun_def] "f [= (f::'a=>'b::quasi_order)"; +Goalw [le_fun_def] "f [= (f::'a=>'b::quasi_order)"; by (rtac allI 1); by (rtac le_refl 1); qed "le_fun_refl"; -goalw thy [le_fun_def] "f [= g & g [= h --> f [= (h::'a=>'b::quasi_order)"; +Goalw [le_fun_def] "f [= g & g [= h --> f [= (h::'a=>'b::quasi_order)"; by Safe_tac; by (rtac (le_trans RS mp) 1); by (Fast_tac 1); qed "le_fun_trans"; -goalw thy [le_fun_def] "f [= g & g [= f --> f = (g::'a=>'b::partial_order)"; +Goalw [le_fun_def] "f [= g & g [= f --> f = (g::'a=>'b::partial_order)"; by Safe_tac; by (rtac ext 1); by (rtac (le_antisym RS mp) 1); @@ -56,23 +56,23 @@ (** duals **) (*"'a dual" is even an isotype*) -goal thy "Rep_dual (Abs_dual y) = y"; +Goal "Rep_dual (Abs_dual y) = y"; by (rtac Abs_dual_inverse 1); by (rewtac dual_def); by (Fast_tac 1); qed "Abs_dual_inverse'"; -goalw thy [le_dual_def] "x [= (x::'a::quasi_order dual)"; +Goalw [le_dual_def] "x [= (x::'a::quasi_order dual)"; by (rtac le_refl 1); qed "le_dual_refl"; -goalw thy [le_dual_def] "x [= y & y [= z --> x [= (z::'a::quasi_order dual)"; +Goalw [le_dual_def] "x [= y & y [= z --> x [= (z::'a::quasi_order dual)"; by (stac conj_commut 1); by (rtac le_trans 1); qed "le_dual_trans"; -goalw thy [le_dual_def] "x [= y & y [= x --> x = (y::'a::partial_order dual)"; +Goalw [le_dual_def] "x [= y & y [= x --> x = (y::'a::partial_order dual)"; by Safe_tac; by (rtac (Rep_dual_inverse RS subst) 1); by (rtac sym 1); @@ -83,6 +83,6 @@ by (assume_tac 1); qed "le_dual_antisym"; -goalw thy [le_dual_def] "x [= y | y [= (x::'a::linear_order dual)"; +Goalw [le_dual_def] "x [= y | y [= (x::'a::linear_order dual)"; by (rtac le_linear 1); qed "le_dual_linear"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Lattice/Order.ML --- a/src/HOL/AxClasses/Lattice/Order.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Lattice/Order.ML Mon Jun 22 17:26:46 1998 +0200 @@ -12,20 +12,20 @@ Fast_tac 1; -goalw thy [is_inf_def] "is_inf x y inf & is_inf x y inf' --> inf = inf'"; +Goalw [is_inf_def] "is_inf x y inf & is_inf x y inf' --> inf = inf'"; by tac; qed "is_inf_uniq"; -goalw thy [is_sup_def] "is_sup x y sup & is_sup x y sup' --> sup = sup'"; +Goalw [is_sup_def] "is_sup x y sup & is_sup x y sup' --> sup = sup'"; by tac; qed "is_sup_uniq"; -goalw thy [is_Inf_def] "is_Inf A inf & is_Inf A inf' --> inf = inf'"; +Goalw [is_Inf_def] "is_Inf A inf & is_Inf A inf' --> inf = inf'"; by tac; qed "is_Inf_uniq"; -goalw thy [is_Sup_def] "is_Sup A sup & is_Sup A sup' --> sup = sup'"; +Goalw [is_Sup_def] "is_Sup A sup & is_Sup A sup' --> sup = sup'"; by tac; qed "is_Sup_uniq"; @@ -33,18 +33,18 @@ (* commutativity *) -goalw thy [is_inf_def] "is_inf x y inf = is_inf y x inf"; +Goalw [is_inf_def] "is_inf x y inf = is_inf y x inf"; by (Fast_tac 1); qed "is_inf_commut"; -goalw thy [is_sup_def] "is_sup x y sup = is_sup y x sup"; +Goalw [is_sup_def] "is_sup x y sup = is_sup y x sup"; by (Fast_tac 1); qed "is_sup_commut"; (* associativity *) -goalw thy [is_inf_def] "is_inf x y xy & is_inf y z yz & is_inf xy z xyz --> is_inf x yz xyz"; +Goalw [is_inf_def] "is_inf x y xy & is_inf y z yz & is_inf xy z xyz --> is_inf x yz xyz"; by Safe_tac; (*level 1*) by (rtac (le_trans RS mp) 1); @@ -78,7 +78,7 @@ qed "is_inf_assoc"; -goalw thy [is_sup_def] "is_sup x y xy & is_sup y z yz & is_sup xy z xyz --> is_sup x yz xyz"; +Goalw [is_sup_def] "is_sup x y xy & is_sup y z yz & is_sup xy z xyz --> is_sup x yz xyz"; by Safe_tac; (*level 1*) by (rtac (le_trans RS mp) 1); @@ -114,7 +114,7 @@ (** limits in linear orders **) -goalw thy [minimum_def, is_inf_def] "is_inf (x::'a::linear_order) y (minimum x y)"; +Goalw [minimum_def, is_inf_def] "is_inf (x::'a::linear_order) y (minimum x y)"; by (stac split_if 1); by (REPEAT_FIRST (resolve_tac [conjI, impI])); (*case "x [= y"*) @@ -130,7 +130,7 @@ by (Fast_tac 1); qed "min_is_inf"; -goalw thy [maximum_def, is_sup_def] "is_sup (x::'a::linear_order) y (maximum x y)"; +Goalw [maximum_def, is_sup_def] "is_sup (x::'a::linear_order) y (maximum x y)"; by (stac split_if 1); by (REPEAT_FIRST (resolve_tac [conjI, impI])); (*case "x [= y"*) @@ -150,7 +150,7 @@ (** general vs. binary limits **) -goalw thy [is_inf_def, is_Inf_def] "is_Inf {x, y} inf = is_inf x y inf"; +Goalw [is_inf_def, is_Inf_def] "is_Inf {x, y} inf = is_inf x y inf"; by (rtac iffI 1); (*==>*) by (Fast_tac 1); @@ -161,7 +161,7 @@ by (Fast_tac 1); qed "bin_is_Inf_eq"; -goalw thy [is_sup_def, is_Sup_def] "is_Sup {x, y} sup = is_sup x y sup"; +Goalw [is_sup_def, is_Sup_def] "is_Sup {x, y} sup = is_sup x y sup"; by (rtac iffI 1); (*==>*) by (Fast_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/AxClasses/Tutorial/Group.ML --- a/src/HOL/AxClasses/Tutorial/Group.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/AxClasses/Tutorial/Group.ML Mon Jun 22 17:26:46 1998 +0200 @@ -11,7 +11,7 @@ fun ssub r = standard (r RS ssubst); -goal Group.thy "x <*> inverse x = (1::'a::group)"; +Goal "x <*> inverse x = (1::'a::group)"; by (rtac (sub left_unit) 1); back(); by (rtac (sub assoc) 1); @@ -28,7 +28,7 @@ qed "right_inverse"; -goal Group.thy "x <*> 1 = (x::'a::group)"; +Goal "x <*> 1 = (x::'a::group)"; by (rtac (sub left_inverse) 1); by (rtac (sub assoc) 1); by (rtac (ssub right_inverse) 1); @@ -37,7 +37,7 @@ qed "right_unit"; -goal Group.thy "e <*> x = x --> e = (1::'a::group)"; +Goal "e <*> x = x --> e = (1::'a::group)"; by (rtac impI 1); by (rtac (sub right_unit) 1); back(); @@ -49,7 +49,7 @@ qed "strong_one_unit"; -goal Group.thy "EX! e. ALL x. e <*> x = (x::'a::group)"; +Goal "EX! e. ALL x. e <*> x = (x::'a::group)"; by (rtac ex1I 1); by (rtac allI 1); by (rtac left_unit 1); @@ -60,7 +60,7 @@ qed "ex1_unit"; -goal Group.thy "ALL x. EX! e. e <*> x = (x::'a::group)"; +Goal "ALL x. EX! e. e <*> x = (x::'a::group)"; by (rtac allI 1); by (rtac ex1I 1); by (rtac left_unit 1); @@ -69,7 +69,7 @@ qed "ex1_unit'"; -goal Group.thy "y <*> x = 1 --> y = inverse (x::'a::group)"; +Goal "y <*> x = 1 --> y = inverse (x::'a::group)"; by (rtac impI 1); by (rtac (sub right_unit) 1); back(); @@ -85,7 +85,7 @@ qed "one_inverse"; -goal Group.thy "ALL x. EX! y. y <*> x = (1::'a::group)"; +Goal "ALL x. EX! y. y <*> x = (1::'a::group)"; by (rtac allI 1); by (rtac ex1I 1); by (rtac left_inverse 1); @@ -95,7 +95,7 @@ qed "ex1_inverse"; -goal Group.thy "inverse (x <*> y) = inverse y <*> inverse (x::'a::group)"; +Goal "inverse (x <*> y) = inverse y <*> inverse (x::'a::group)"; by (rtac sym 1); by (rtac mp 1); by (rtac one_inverse 1); @@ -109,7 +109,7 @@ qed "inverse_product"; -goal Group.thy "inverse (inverse x) = (x::'a::group)"; +Goal "inverse (inverse x) = (x::'a::group)"; by (rtac sym 1); by (rtac (one_inverse RS mp) 1); by (rtac (ssub right_inverse) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Divides.ML --- a/src/HOL/Divides.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Divides.ML Mon Jun 22 17:26:46 1998 +0200 @@ -22,58 +22,58 @@ (*** Remainder ***) -goal thy "(%m. m mod n) = wfrec (trancl pred_nat) \ +Goal "(%m. m mod n) = wfrec (trancl pred_nat) \ \ (%f j. if j m mod n = m"; +Goal "!!m. m m mod n = m"; by (rtac (mod_eq RS wf_less_trans) 1); by (Asm_simp_tac 1); qed "mod_less"; -goal thy "!!m. [| 0 m mod n = (m-n) mod n"; +Goal "!!m. [| 0 m mod n = (m-n) mod n"; by (rtac (mod_eq RS wf_less_trans) 1); by (asm_simp_tac (simpset() addsimps [diff_less, cut_apply, less_eq]) 1); qed "mod_geq"; (*NOT suitable for rewriting: loops*) -goal thy "!!m. 0 m mod n = (if m m mod n = (if m n mod n = 0"; +Goal "!!n. 0 n mod n = 0"; by (asm_simp_tac (simpset() addsimps [mod_less, mod_geq]) 1); qed "mod_self"; -goal thy "!!n. 0 (m+n) mod n = m mod n"; +Goal "!!n. 0 (m+n) mod n = m mod n"; by (subgoal_tac "(n + m) mod n = (n+m-n) mod n" 1); by (stac (mod_geq RS sym) 2); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_commute]))); qed "mod_add_self2"; -goal thy "!!k. 0 (n+m) mod n = m mod n"; +Goal "!!k. 0 (n+m) mod n = m mod n"; by (asm_simp_tac (simpset() addsimps [add_commute, mod_add_self2]) 1); qed "mod_add_self1"; -goal thy "!!n. 0 (m + k*n) mod n = m mod n"; +Goal "!!n. 0 (m + k*n) mod n = m mod n"; by (induct_tac "k" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps (add_ac @ [mod_add_self1])))); qed "mod_mult_self1"; -goal thy "!!m. 0 (m + n*k) mod n = m mod n"; +Goal "!!m. 0 (m + n*k) mod n = m mod n"; by (asm_simp_tac (simpset() addsimps [mult_commute, mod_mult_self1]) 1); qed "mod_mult_self2"; Addsimps [mod_mult_self1, mod_mult_self2]; -goal thy "!!k. [| 0 (m mod n)*k = (m*k) mod (n*k)"; +Goal "!!k. [| 0 (m mod n)*k = (m*k) mod (n*k)"; by (res_inst_tac [("n","m")] less_induct 1); by (stac mod_if 1); by (Asm_simp_tac 1); @@ -81,7 +81,7 @@ diff_less, diff_mult_distrib]) 1); qed "mod_mult_distrib"; -goal thy "!!k. [| 0 k*(m mod n) = (k*m) mod (k*n)"; +Goal "!!k. [| 0 k*(m mod n) = (k*m) mod (k*n)"; by (res_inst_tac [("n","m")] less_induct 1); by (stac mod_if 1); by (Asm_simp_tac 1); @@ -89,7 +89,7 @@ diff_less, diff_mult_distrib2]) 1); qed "mod_mult_distrib2"; -goal thy "!!n. 0 m*n mod n = 0"; +Goal "!!n. 0 m*n mod n = 0"; by (induct_tac "m" 1); by (asm_simp_tac (simpset() addsimps [mod_less]) 1); by (dres_inst_tac [("m","m*n")] mod_add_self2 1); @@ -99,28 +99,28 @@ (*** Quotient ***) -goal thy "(%m. m div n) = wfrec (trancl pred_nat) \ +Goal "(%m. m div n) = wfrec (trancl pred_nat) \ \ (%f j. if j m div n = 0"; +Goal "!!m. m m div n = 0"; by (rtac (div_eq RS wf_less_trans) 1); by (Asm_simp_tac 1); qed "div_less"; -goal thy "!!M. [| 0 m div n = Suc((m-n) div n)"; +Goal "!!M. [| 0 m div n = Suc((m-n) div n)"; by (rtac (div_eq RS wf_less_trans) 1); by (asm_simp_tac (simpset() addsimps [diff_less, cut_apply, less_eq]) 1); qed "div_geq"; (*NOT suitable for rewriting: loops*) -goal thy "!!m. 0 m div n = (if m m div n = (if m (m div n)*n + m mod n = m"; +Goal "!!m. 0 (m div n)*n + m mod n = m"; by (res_inst_tac [("n","m")] less_induct 1); by (stac mod_if 1); by (ALLGOALS (asm_simp_tac @@ -130,39 +130,39 @@ qed "mod_div_equality"; (* a simple rearrangement of mod_div_equality: *) -goal thy "!!k. 0 k*(m div k) = m - (m mod k)"; +Goal "!!k. 0 k*(m div k) = m - (m mod k)"; by (dres_inst_tac [("m","m")] mod_div_equality 1); by (EVERY1[etac subst, simp_tac (simpset() addsimps mult_ac), K(IF_UNSOLVED no_tac)]); qed "mult_div_cancel"; -goal thy "m div 1 = m"; +Goal "m div 1 = m"; by (induct_tac "m" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [div_less, div_geq]))); qed "div_1"; Addsimps [div_1]; -goal thy "!!n. 0 n div n = 1"; +Goal "!!n. 0 n div n = 1"; by (asm_simp_tac (simpset() addsimps [div_less, div_geq]) 1); qed "div_self"; -goal thy "!!n. 0 (m+n) div n = Suc (m div n)"; +Goal "!!n. 0 (m+n) div n = Suc (m div n)"; by (subgoal_tac "(n + m) div n = Suc ((n+m-n) div n)" 1); by (stac (div_geq RS sym) 2); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_commute]))); qed "div_add_self2"; -goal thy "!!k. 0 (n+m) div n = Suc (m div n)"; +Goal "!!k. 0 (n+m) div n = Suc (m div n)"; by (asm_simp_tac (simpset() addsimps [add_commute, div_add_self2]) 1); qed "div_add_self1"; -goal thy "!!n. 0 (m + k*n) div n = k + m div n"; +Goal "!!n. 0 (m + k*n) div n = k + m div n"; by (induct_tac "k" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps (add_ac @ [div_add_self1])))); qed "div_mult_self1"; -goal thy "!!m. 0 (m + n*k) div n = k + m div n"; +Goal "!!m. 0 (m + n*k) div n = k + m div n"; by (asm_simp_tac (simpset() addsimps [mult_commute, div_mult_self1]) 1); qed "div_mult_self2"; @@ -170,7 +170,7 @@ (* Monotonicity of div in first argument *) -goal thy "!!n. 0 ALL m. m <= n --> (m div k) <= (n div k)"; +Goal "!!n. 0 ALL m. m <= n --> (m div k) <= (n div k)"; by (res_inst_tac [("n","n")] less_induct 1); by (Clarify_tac 1); by (case_tac "na (k div n) <= (k div m)"; +Goal "!!k m n. [| 0 (k div n) <= (k div m)"; by (subgoal_tac "0 m div n <= m"; +Goal "!!m n. 0 m div n <= m"; by (subgoal_tac "m div n <= m div 1" 1); by (Asm_full_simp_tac 1); by (rtac div_le_mono2 1); @@ -214,7 +214,7 @@ Addsimps [div_le_dividend]; (* Similar for "less than" *) -goal thy "!!m n. 1 (0 < m) --> (m div n < m)"; +Goal "!!m n. 1 (0 < m) --> (m div n < m)"; by (res_inst_tac [("n","m")] less_induct 1); by (Simp_tac 1); by (rename_tac "m" 1); @@ -238,7 +238,7 @@ (*** Further facts about mod (mainly for the mutilated chess board ***) -goal thy +Goal "!!m n. 0 \ \ Suc(m) mod n = (if Suc(m mod n) = n then 0 else Suc(m mod n))"; by (res_inst_tac [("n","m")] less_induct 1); @@ -255,7 +255,7 @@ by (asm_simp_tac (simpset() addsimps [mod_less]) 1); qed "mod_Suc"; -goal thy "!!m n. 0 m mod n < n"; +Goal "!!m n. 0 m mod n < n"; by (res_inst_tac [("n","m")] less_induct 1); by (excluded_middle_tac "na k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)"; +Goal "!!k b. b<2 ==> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)"; by (subgoal_tac "k mod 2 < 2" 1); by (asm_simp_tac (simpset() addsimps [mod_less_divisor]) 2); by (Asm_simp_tac 1); by Safe_tac; qed "mod2_cases"; -goal thy "Suc(Suc(m)) mod 2 = m mod 2"; +Goal "Suc(Suc(m)) mod 2 = m mod 2"; by (subgoal_tac "m mod 2 < 2" 1); by (asm_simp_tac (simpset() addsimps [mod_less_divisor]) 2); by Safe_tac; @@ -285,21 +285,21 @@ qed "mod2_Suc_Suc"; Addsimps [mod2_Suc_Suc]; -goal Divides.thy "(0 < m mod 2) = (m mod 2 = 1)"; +Goal "(0 < m mod 2) = (m mod 2 = 1)"; by (subgoal_tac "m mod 2 < 2" 1); by (asm_simp_tac (simpset() addsimps [mod_less_divisor]) 2); by Auto_tac; qed "mod2_gr_0"; Addsimps [mod2_gr_0]; -goal thy "(m+m) mod 2 = 0"; +Goal "(m+m) mod 2 = 0"; by (induct_tac "m" 1); by (simp_tac (simpset() addsimps [mod_less]) 1); by (Asm_simp_tac 1); qed "mod2_add_self_eq_0"; Addsimps [mod2_add_self_eq_0]; -goal thy "((m+m)+n) mod 2 = n mod 2"; +Goal "((m+m)+n) mod 2 = n mod 2"; by (induct_tac "m" 1); by (simp_tac (simpset() addsimps [mod_less]) 1); by (Asm_simp_tac 1); @@ -311,7 +311,7 @@ (*** More division laws ***) -goal thy "!!n. 0 m*n div n = m"; +Goal "!!n. 0 m*n div n = m"; by (cut_inst_tac [("m", "m*n")] mod_div_equality 1); by (assume_tac 1); by (asm_full_simp_tac (simpset() addsimps [mod_mult_self_is_0]) 1); @@ -319,7 +319,7 @@ Addsimps [div_mult_self_is_m]; (*Cancellation law for division*) -goal thy "!!k. [| 0 (k*m) div (k*n) = m div n"; +Goal "!!k. [| 0 (k*m) div (k*n) = m div n"; by (res_inst_tac [("n","m")] less_induct 1); by (case_tac "na (k*m) mod (k*n) = k * (m mod n)"; +Goal "!!k. [| 0 (k*m) mod (k*n) = k * (m mod n)"; by (res_inst_tac [("n","m")] less_induct 1); by (case_tac "na m = 0"; +Goalw [dvd_def] "!!m. 0 dvd m ==> m = 0"; by (fast_tac (claset() addss simpset()) 1); qed "dvd_0_left"; -goalw thy [dvd_def] "1 dvd k"; +Goalw [dvd_def] "1 dvd k"; by (Simp_tac 1); qed "dvd_1_left"; AddIffs [dvd_1_left]; -goalw thy [dvd_def] "m dvd m"; +Goalw [dvd_def] "m dvd m"; by (blast_tac (claset() addIs [mult_1_right RS sym]) 1); qed "dvd_refl"; Addsimps [dvd_refl]; -goalw thy [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p"; +Goalw [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p"; by (blast_tac (claset() addIs [mult_assoc] ) 1); qed "dvd_trans"; -goalw thy [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n"; +Goalw [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n"; by (fast_tac (claset() addDs [mult_eq_self_implies_10] addss (simpset() addsimps [mult_assoc, mult_eq_1_iff])) 1); qed "dvd_anti_sym"; -goalw thy [dvd_def] "!!k. [| k dvd m; k dvd n |] ==> k dvd (m + n)"; +Goalw [dvd_def] "!!k. [| k dvd m; k dvd n |] ==> k dvd (m + n)"; by (blast_tac (claset() addIs [add_mult_distrib2 RS sym]) 1); qed "dvd_add"; -goalw thy [dvd_def] "!!k. [| k dvd m; k dvd n |] ==> k dvd (m-n)"; +Goalw [dvd_def] "!!k. [| k dvd m; k dvd n |] ==> k dvd (m-n)"; by (blast_tac (claset() addIs [diff_mult_distrib2 RS sym]) 1); qed "dvd_diff"; -goal thy "!!k. [| k dvd (m-n); k dvd n; n<=m |] ==> k dvd m"; +Goal "!!k. [| k dvd (m-n); k dvd n; n<=m |] ==> k dvd m"; by (etac (not_less_iff_le RS iffD2 RS add_diff_inverse RS subst) 1); by (blast_tac (claset() addIs [dvd_add]) 1); qed "dvd_diffD"; -goalw thy [dvd_def] "!!k. k dvd n ==> k dvd (m*n)"; +Goalw [dvd_def] "!!k. k dvd n ==> k dvd (m*n)"; by (blast_tac (claset() addIs [mult_left_commute]) 1); qed "dvd_mult"; -goal thy "!!k. k dvd m ==> k dvd (m*n)"; +Goal "!!k. k dvd m ==> k dvd (m*n)"; by (stac mult_commute 1); by (etac dvd_mult 1); qed "dvd_mult2"; @@ -404,7 +404,7 @@ (* k dvd (m*k) *) AddIffs [dvd_refl RS dvd_mult, dvd_refl RS dvd_mult2]; -goalw thy [dvd_def] "!!m. [| f dvd m; f dvd n; 0 f dvd (m mod n)"; +Goalw [dvd_def] "!!m. [| f dvd m; f dvd n; 0 f dvd (m mod n)"; by (Clarify_tac 1); by (full_simp_tac (simpset() addsimps [zero_less_mult_iff]) 1); by (res_inst_tac @@ -414,30 +414,30 @@ mult_mod_distrib, add_mult_distrib2]) 1); qed "dvd_mod"; -goal thy "!!k. [| k dvd (m mod n); k dvd n; 0 k dvd m"; +Goal "!!k. [| k dvd (m mod n); k dvd n; 0 k dvd m"; by (subgoal_tac "k dvd ((m div n)*n + m mod n)" 1); by (asm_simp_tac (simpset() addsimps [dvd_add, dvd_mult]) 2); by (asm_full_simp_tac (simpset() addsimps [mod_div_equality]) 1); qed "dvd_mod_imp_dvd"; -goalw thy [dvd_def] "!!k m n. [| (k*m) dvd (k*n); 0 m dvd n"; +Goalw [dvd_def] "!!k m n. [| (k*m) dvd (k*n); 0 m dvd n"; by (etac exE 1); by (asm_full_simp_tac (simpset() addsimps mult_ac) 1); by (Blast_tac 1); qed "dvd_mult_cancel"; -goalw thy [dvd_def] "!!i j. [| i dvd m; j dvd n|] ==> (i*j) dvd (m*n)"; +Goalw [dvd_def] "!!i j. [| i dvd m; j dvd n|] ==> (i*j) dvd (m*n)"; by (Clarify_tac 1); by (res_inst_tac [("x","k*ka")] exI 1); by (asm_simp_tac (simpset() addsimps mult_ac) 1); qed "mult_dvd_mono"; -goalw thy [dvd_def] "!!c. (i*j) dvd k ==> i dvd k"; +Goalw [dvd_def] "!!c. (i*j) dvd k ==> i dvd k"; by (full_simp_tac (simpset() addsimps [mult_assoc]) 1); by (Blast_tac 1); qed "dvd_mult_left"; -goalw thy [dvd_def] "!!n. [| k dvd n; 0 < n |] ==> k <= n"; +Goalw [dvd_def] "!!n. [| k dvd n; 0 < n |] ==> k <= n"; by (Clarify_tac 1); by (ALLGOALS (full_simp_tac (simpset() addsimps [zero_less_mult_iff]))); by (etac conjE 1); @@ -447,7 +447,7 @@ by (Simp_tac 1); qed "dvd_imp_le"; -goalw thy [dvd_def] "!!k. 0 (k dvd n) = (n mod k = 0)"; +Goalw [dvd_def] "!!k. 0 (k dvd n) = (n mod k = 0)"; by Safe_tac; by (stac mult_commute 1); by (Asm_simp_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Finite.ML --- a/src/HOL/Finite.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Finite.ML Mon Jun 22 17:26:46 1998 +0200 @@ -42,7 +42,7 @@ qed "finite_UnI"; (*Every subset of a finite set is finite*) -goal Finite.thy "!!B. finite B ==> ALL A. A<=B --> finite A"; +Goal "!!B. finite B ==> ALL A. A<=B --> finite A"; by (etac finite_induct 1); by (Simp_tac 1); by (safe_tac (claset() addSDs [subset_insert_iff RS iffD1])); @@ -50,19 +50,19 @@ by (ALLGOALS Asm_simp_tac); val lemma = result(); -goal Finite.thy "!!A. [| A<=B; finite B |] ==> finite A"; +Goal "!!A. [| A<=B; finite B |] ==> finite A"; by (dtac lemma 1); by (Blast_tac 1); qed "finite_subset"; -goal Finite.thy "finite(F Un G) = (finite F & finite G)"; +Goal "finite(F Un G) = (finite F & finite G)"; by (blast_tac (claset() addIs [read_instantiate [("B", "?AA Un ?BB")] finite_subset, finite_UnI]) 1); qed "finite_Un"; AddIffs[finite_Un]; -goal Finite.thy "finite(insert a A) = finite A"; +Goal "finite(insert a A) = finite A"; by (stac insert_is_Un 1); by (simp_tac (HOL_ss addsimps [finite_Un]) 1); by (Blast_tac 1); @@ -70,7 +70,7 @@ Addsimps[finite_insert]; (*The image of a finite set is finite *) -goal Finite.thy "!!F. finite F ==> finite(h``F)"; +Goal "!!F. finite F ==> finite(h``F)"; by (etac finite_induct 1); by (Simp_tac 1); by (Asm_simp_tac 1); @@ -102,7 +102,7 @@ bind_thm ("finite_Diff", Diff_subset RS finite_subset); Addsimps [finite_Diff]; -goal Finite.thy "finite(A-{a}) = finite(A)"; +Goal "finite(A-{a}) = finite(A)"; by (case_tac "a:A" 1); by (rtac (finite_insert RS sym RS trans) 1); by (stac insert_Diff 1); @@ -111,7 +111,7 @@ AddIffs [finite_Diff_singleton]; (*Lemma for proving finite_imageD*) -goal Finite.thy "!!A. finite B ==> !A. f``A = B --> inj_on f A --> finite A"; +Goal "!!A. finite B ==> !A. f``A = B --> inj_on f A --> finite A"; by (etac finite_induct 1); by (ALLGOALS Asm_simp_tac); by (Clarify_tac 1); @@ -127,7 +127,7 @@ (asm_full_simp_tac (simpset() addsimps [inj_on_image_set_diff]))); val lemma = result(); -goal Finite.thy "!!A. [| finite(f``A); inj_on f A |] ==> finite A"; +Goal "!!A. [| finite(f``A); inj_on f A |] ==> finite A"; by (dtac lemma 1); by (Blast_tac 1); qed "finite_imageD"; @@ -143,7 +143,7 @@ (** Sigma of finite sets **) -goalw Finite.thy [Sigma_def] +Goalw [Sigma_def] "!!A. [| finite A; !a:A. finite(B a) |] ==> finite(SIGMA a:A. B a)"; by (blast_tac (claset() addSIs [finite_UnionI]) 1); bind_thm("finite_SigmaI", ballI RSN (2,result())); @@ -151,7 +151,7 @@ (** The powerset of a finite set **) -goal Finite.thy "!!A. finite(Pow A) ==> finite A"; +Goal "!!A. finite(Pow A) ==> finite A"; by (subgoal_tac "finite ((%x.{x})``A)" 1); by (rtac finite_subset 2); by (assume_tac 3); @@ -159,7 +159,7 @@ (fast_tac (claset() addSDs [rewrite_rule [inj_on_def] finite_imageD]))); val lemma = result(); -goal Finite.thy "finite(Pow A) = finite A"; +Goal "finite(Pow A) = finite A"; by (rtac iffI 1); by (etac lemma 1); (*Opposite inclusion: finite A ==> finite (Pow A) *) @@ -170,7 +170,7 @@ qed "finite_Pow_iff"; AddIffs [finite_Pow_iff]; -goal Finite.thy "finite(r^-1) = finite r"; +Goal "finite(r^-1) = finite r"; by (subgoal_tac "r^-1 = (%(x,y).(y,x))``r" 1); by (Asm_simp_tac 1); by (rtac iffI 1); @@ -191,7 +191,7 @@ by (Blast_tac 1); val Collect_conv_insert = result(); -goalw Finite.thy [card_def] "card {} = 0"; +Goalw [card_def] "card {} = 0"; by (rtac Least_equality 1); by (ALLGOALS Asm_full_simp_tac); qed "card_empty"; @@ -211,7 +211,7 @@ addcongs [rev_conj_cong]) 1); qed "finite_has_card"; -goal Finite.thy +Goal "!!A.[| x ~: A; insert x A = {f i|i. i \ \ ? m::nat. m \ +Goal "!!A. [| finite A; x ~: A |] ==> \ \ (LEAST n. ? f. insert x A = {f i|i. i card(insert x A) = Suc(card A)"; by (etac lemma 1); by (assume_tac 1); qed "card_insert_disjoint"; Addsimps [card_insert_disjoint]; -goal Finite.thy "!!A. finite A ==> card A <= card (insert x A)"; +Goal "!!A. finite A ==> card A <= card (insert x A)"; by (case_tac "x: A" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [insert_absorb]))); qed "card_insert_le"; -goal Finite.thy "!!A. finite A ==> !B. B <= A --> card(B) <= card(A)"; +Goal "!!A. finite A ==> !B. B <= A --> card(B) <= card(A)"; by (etac finite_induct 1); by (Simp_tac 1); by (Clarify_tac 1); @@ -325,13 +325,13 @@ (simpset() addsimps [subset_insert_iff, finite_subset])) 1); qed_spec_mp "card_mono"; -goal Finite.thy "!!A B. [| finite A; finite B |]\ +Goal "!!A B. [| finite A; finite B |]\ \ ==> A Int B = {} --> card(A Un B) = card A + card B"; by (etac finite_induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [Int_insert_left]))); qed_spec_mp "card_Un_disjoint"; -goal Finite.thy "!!A. [| finite A; B<=A |] ==> card A - card B = card (A - B)"; +Goal "!!A. [| finite A; B<=A |] ==> card A - card B = card (A - B)"; by (subgoal_tac "(A-B) Un B = A" 1); by (Blast_tac 2); by (rtac (add_right_cancel RS iffD1) 1); @@ -344,18 +344,18 @@ add_diff_inverse, card_mono, finite_subset]))); qed "card_Diff_subset"; -goal Finite.thy "!!A. [| finite A; x: A |] ==> Suc(card(A-{x})) = card A"; +Goal "!!A. [| finite A; x: A |] ==> Suc(card(A-{x})) = card A"; by (res_inst_tac [("t", "A")] (insert_Diff RS subst) 1); by (assume_tac 1); by (Asm_simp_tac 1); qed "card_Suc_Diff"; -goal Finite.thy "!!A. [| finite A; x: A |] ==> card(A-{x}) < card A"; +Goal "!!A. [| finite A; x: A |] ==> card(A-{x}) < card A"; by (rtac Suc_less_SucD 1); by (asm_simp_tac (simpset() addsimps [card_Suc_Diff]) 1); qed "card_Diff"; -goal Finite.thy "!!A. finite A ==> card(A-{x}) <= card A"; +Goal "!!A. finite A ==> card(A-{x}) <= card A"; by (case_tac "x: A" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [card_Diff, less_imp_le]))); qed "card_Diff_le"; @@ -363,14 +363,14 @@ (*** Cardinality of the Powerset ***) -goal Finite.thy "!!A. finite A ==> card(insert x A) = Suc(card(A-{x}))"; +Goal "!!A. finite A ==> card(insert x A) = Suc(card(A-{x}))"; by (case_tac "x:A" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [card_Suc_Diff, insert_absorb]))); qed "card_insert"; Addsimps [card_insert]; -goal Finite.thy "!!A. finite(A) ==> inj_on f A --> card (f `` A) = card A"; +Goal "!!A. finite(A) ==> inj_on f A --> card (f `` A) = card A"; by (etac finite_induct 1); by (ALLGOALS Asm_simp_tac); by Safe_tac; @@ -382,7 +382,7 @@ by (Blast_tac 1); qed_spec_mp "card_image"; -goal thy "!!A. finite A ==> card (Pow A) = 2 ^ card A"; +Goal "!!A. finite A ==> card (Pow A) = 2 ^ card A"; by (etac finite_induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [Pow_insert]))); by (stac card_Un_disjoint 1); @@ -396,7 +396,7 @@ (*Proper subsets*) -goalw Finite.thy [psubset_def] +Goalw [psubset_def] "!!B. finite B ==> !A. A < B --> card(A) < card(B)"; by (etac finite_induct 1); by (Simp_tac 1); @@ -418,7 +418,7 @@ (*Relates to equivalence classes. Based on a theorem of F. Kammueller's. The "finite C" premise is redundant*) -goal thy "!!C. finite C ==> finite (Union C) --> \ +Goal "!!C. finite C ==> finite (Union C) --> \ \ (! c : C. k dvd card c) --> \ \ (! c1: C. ! c2: C. c1 ~= c2 --> c1 Int c2 = {}) \ \ --> k dvd card(Union C)"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Fun.ML --- a/src/HOL/Fun.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Fun.ML Mon Jun 22 17:26:46 1998 +0200 @@ -7,7 +7,7 @@ *) -goal thy "(f = g) = (!x. f(x)=g(x))"; +Goal "(f = g) = (!x. f(x)=g(x))"; by (rtac iffI 1); by (Asm_simp_tac 1); by (rtac ext 1 THEN Asm_simp_tac 1); @@ -97,7 +97,7 @@ by (REPEAT (resolve_tac prems 1)); qed "inj_onD"; -goal thy "!!x y.[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)"; +Goal "!!x y.[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)"; by (blast_tac (claset() addSDs [inj_onD]) 1); qed "inj_on_iff"; @@ -108,7 +108,7 @@ by (REPEAT (resolve_tac prems 1)); qed "inj_on_contraD"; -goalw thy [inj_on_def] +Goalw [inj_on_def] "!!A B. [| A<=B; inj_on f B |] ==> inj_on f A"; by (Blast_tac 1); qed "subset_inj_on"; @@ -116,7 +116,7 @@ (*** Lemmas about inj ***) -goalw thy [o_def] +Goalw [o_def] "!!f g. [| inj(f); inj_on g (range f) |] ==> inj(g o f)"; by (fast_tac (claset() addIs [injI] addEs [injD, inj_onD]) 1); qed "comp_inj"; @@ -135,26 +135,26 @@ by (REPEAT (resolve_tac (prems @ [f_inv_f]) 1)); qed "inv_injective"; -goal thy "!!f. [| inj(f); A<=range(f) |] ==> inj_on (inv f) A"; +Goal "!!f. [| inj(f); A<=range(f) |] ==> inj_on (inv f) A"; by (fast_tac (claset() addIs [inj_onI] addEs [inv_injective,injD]) 1); qed "inj_on_inv"; -goalw thy [inj_on_def] +Goalw [inj_on_def] "!!f. [| inj_on f C; A<=C; B<=C |] ==> f``(A Int B) = f``A Int f``B"; by (Blast_tac 1); qed "inj_on_image_Int"; -goalw thy [inj_on_def] +Goalw [inj_on_def] "!!f. [| inj_on f C; A<=C; B<=C |] ==> f``(A-B) = f``A - f``B"; by (Blast_tac 1); qed "inj_on_image_set_diff"; -goalw thy [inj_def] "!!f. inj f ==> f``(A Int B) = f``A Int f``B"; +Goalw [inj_def] "!!f. inj f ==> f``(A Int B) = f``A Int f``B"; by (Blast_tac 1); qed "image_Int"; -goalw thy [inj_def] "!!f. inj f ==> f``(A-B) = f``A - f``B"; +Goalw [inj_def] "!!f. inj f ==> f``(A-B) = f``A - f``B"; by (Blast_tac 1); qed "image_set_diff"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Gfp.ML --- a/src/HOL/Gfp.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Gfp.ML Mon Jun 22 17:26:46 1998 +0200 @@ -56,7 +56,7 @@ qed "coinduct_lemma"; (*strong version, thanks to Coen & Frost*) -goal Gfp.thy +Goal "!!X. [| mono(f); a: X; X <= f(X Un gfp(f)) |] ==> a : gfp(f)"; by (rtac (coinduct_lemma RSN (2, weak_coinduct)) 1); by (REPEAT (ares_tac [UnI1, Un_least] 1)); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/HOL.thy --- a/src/HOL/HOL.thy Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/HOL.thy Mon Jun 22 17:26:46 1998 +0200 @@ -178,11 +178,13 @@ True_or_False "(P=True) | (P=False)" defs - (* Misc Definitions *) - + (*misc definitions*) Let_def "Let s f == f(s)" o_def "(f::'b=>'c) o g == (%(x::'a). f(g(x)))" if_def "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)" + + (*arbitrary is completely unspecified, but is made to appear as a + definition syntactically*) arbitrary_def "False ==> arbitrary == (@x. False)" diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Hoare/Arith2.ML --- a/src/HOL/Hoare/Arith2.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Hoare/Arith2.ML Mon Jun 22 17:26:46 1998 +0200 @@ -49,7 +49,7 @@ (*** gcd ***) -goalw thy [gcd_def] "!!n. 0 n = gcd n n"; +Goalw [gcd_def] "!!n. 0 n = gcd n n"; by (forward_tac [cd_nnn] 1); by (rtac (select_equality RS sym) 1); by (blast_tac (claset() addDs [cd_le]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Hoare/Examples.ML --- a/src/HOL/Hoare/Examples.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Hoare/Examples.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,7 +10,7 @@ (*** multiplication by successive addition ***) -goal thy +Goal "{m=0 & s=0} \ \ WHILE m ~= a DO {s = m*b} s := s+b; m := Suc(m) END\ \ {s = a*b}"; @@ -21,7 +21,7 @@ (*** Euclid's algorithm for GCD ***) -goal thy +Goal " {0 [Fast_tac 1]); -goal Denotation.thy "C(WHILE b DO c) = C(IF b THEN c;WHILE b DO c ELSE SKIP)"; +Goal "C(WHILE b DO c) = C(IF b THEN c;WHILE b DO c ELSE SKIP)"; by (Simp_tac 1); by (EVERY[stac (Gamma_mono RS lfp_Tarski) 1, stac (Gamma_def RS meta_eq_to_obj_eq RS fun_cong) 1, @@ -25,7 +25,7 @@ (* Operational Semantics implies Denotational Semantics *) -goal Denotation.thy "!!c s t. -c-> t ==> (s,t) : C(c)"; +Goal "!!c s t. -c-> t ==> (s,t) : C(c)"; (* start with rule induction *) by (etac evalc.induct 1); @@ -41,7 +41,7 @@ (* Denotational Semantics implies Operational Semantics *) -goal Denotation.thy "!s t. (s,t):C(c) --> -c-> t"; +Goal "!s t. (s,t):C(c) --> -c-> t"; by (com.induct_tac "c" 1); by (ALLGOALS Full_simp_tac); @@ -58,6 +58,6 @@ (**** Proof of Equivalence ****) -goal Denotation.thy "(s,t) : C(c) = ( -c-> t)"; +Goal "(s,t) : C(c) = ( -c-> t)"; by (fast_tac (claset() addEs [com2] addDs [com1]) 1); qed "denotational_is_natural"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IMP/Expr.ML --- a/src/HOL/IMP/Expr.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IMP/Expr.ML Mon Jun 22 17:26:46 1998 +0200 @@ -32,14 +32,14 @@ "((b0 ori b1,sigma) -b-> w) = \ \ (? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x|y)))"]; -goal Expr.thy "!n. ((a,s) -a-> n) = (A a s = n)"; +Goal "!n. ((a,s) -a-> n) = (A a s = n)"; by (aexp.induct_tac "a" 1); by (ALLGOALS (fast_tac (claset() addSIs evala.intrs addSEs evala_elim_cases addss (simpset())))); qed_spec_mp "aexp_iff"; -goal Expr.thy "!w. ((b,s) -b-> w) = (B b s = w)"; +Goal "!w. ((b,s) -b-> w) = (B b s = w)"; by (bexp.induct_tac "b" 1); by (ALLGOALS (fast_tac (claset() diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IMP/Hoare.ML --- a/src/HOL/IMP/Hoare.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IMP/Hoare.ML Mon Jun 22 17:26:46 1998 +0200 @@ -9,7 +9,7 @@ open Hoare; -goalw Hoare.thy [hoare_valid_def] "!!P c Q. |- {P}c{Q} ==> |= {P}c{Q}"; +Goalw [hoare_valid_def] "!!P c Q. |- {P}c{Q} ==> |= {P}c{Q}"; by (etac hoare.induct 1); by (ALLGOALS Asm_simp_tac); by (Fast_tac 1); @@ -23,21 +23,21 @@ by (Fast_tac 1); qed "hoare_sound"; -goalw Hoare.thy [wp_def] "wp SKIP Q = Q"; +Goalw [wp_def] "wp SKIP Q = Q"; by (Simp_tac 1); qed "wp_SKIP"; -goalw Hoare.thy [wp_def] "wp (x:=a) Q = (%s. Q(s[x:=a s]))"; +Goalw [wp_def] "wp (x:=a) Q = (%s. Q(s[x:=a s]))"; by (Simp_tac 1); qed "wp_Ass"; -goalw Hoare.thy [wp_def] "wp (c;d) Q = wp c (wp d Q)"; +Goalw [wp_def] "wp (c;d) Q = wp c (wp d Q)"; by (Simp_tac 1); by (rtac ext 1); by (Fast_tac 1); qed "wp_Semi"; -goalw Hoare.thy [wp_def] +Goalw [wp_def] "wp (IF b THEN c ELSE d) Q = (%s. (b s --> wp c Q s) & \ \ (~b s --> wp d Q s))"; by (Simp_tac 1); @@ -45,13 +45,13 @@ by (Fast_tac 1); qed "wp_If"; -goalw Hoare.thy [wp_def] +Goalw [wp_def] "!!s. b s ==> wp (WHILE b DO c) Q s = wp (c;WHILE b DO c) Q s"; by (stac C_While_If 1); by (Asm_simp_tac 1); qed "wp_While_True"; -goalw Hoare.thy [wp_def] "!!s. ~b s ==> wp (WHILE b DO c) Q s = Q s"; +Goalw [wp_def] "!!s. ~b s ==> wp (WHILE b DO c) Q s = Q s"; by (stac C_While_If 1); by (Asm_simp_tac 1); qed "wp_While_False"; @@ -59,12 +59,12 @@ Addsimps [wp_SKIP,wp_Ass,wp_Semi,wp_If,wp_While_True,wp_While_False]; (*Not suitable for rewriting: LOOPS!*) -goal Hoare.thy "wp (WHILE b DO c) Q s = \ +Goal "wp (WHILE b DO c) Q s = \ \ (if b s then wp (c;WHILE b DO c) Q s else Q s)"; by (Simp_tac 1); qed "wp_While_if"; -goal thy +Goal "wp (WHILE b DO c) Q s = \ \ (s : gfp(%S.{s. if b s then wp c (%s. s:S) s else Q s}))"; by (Simp_tac 1); @@ -91,7 +91,7 @@ AddSIs [hoare.skip, hoare.ass, hoare.semi, hoare.If]; -goal Hoare.thy "!Q. |- {wp c Q} c {Q}"; +Goal "!Q. |- {wp c Q} c {Q}"; by (com.induct_tac "c" 1); by (ALLGOALS Simp_tac); by (REPEAT_FIRST Fast_tac); @@ -114,7 +114,7 @@ by (Asm_full_simp_tac 1); qed_spec_mp "wp_is_pre"; -goal Hoare.thy "!!c. |= {P}c{Q} ==> |- {P}c{Q}"; +Goal "!!c. |= {P}c{Q} ==> |- {P}c{Q}"; by (rtac (wp_is_pre RSN (2,hoare.conseq)) 1); by (Fast_tac 2); by (rewrite_goals_tac [hoare_valid_def,wp_def]); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IMP/Natural.ML --- a/src/HOL/IMP/Natural.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IMP/Natural.ML Mon Jun 22 17:26:46 1998 +0200 @@ -17,7 +17,7 @@ AddSEs evalc_elim_cases; (* evaluation of com is deterministic *) -goal Natural.thy "!!c s t. -c-> t ==> (!u. -c-> u --> u=t)"; +Goal "!!c s t. -c-> t ==> (!u. -c-> u --> u=t)"; by (etac evalc.induct 1); by (thin_tac " -c-> s1" 7); (*blast_tac needs Unify.search_bound, trace_bound := 40*) diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IMP/Transition.ML --- a/src/HOL/IMP/Transition.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IMP/Transition.ML Mon Jun 22 17:26:46 1998 +0200 @@ -22,14 +22,14 @@ AddIs evalc1.intrs; -goal Transition.thy "!!s. (SKIP,s) -m-> (SKIP,t) ==> s = t & m = 0"; +Goal "!!s. (SKIP,s) -m-> (SKIP,t) ==> s = t & m = 0"; by (etac rel_pow_E2 1); by (Asm_full_simp_tac 1); by (Fast_tac 1); val hlemma = result(); -goal Transition.thy +Goal "!s t u c d. (c,s) -n-> (SKIP,t) --> (d,t) -*-> (SKIP,u) --> \ \ (c;d, s) -*-> (SKIP, u)"; by (nat_ind_tac "n" 1); @@ -38,7 +38,7 @@ qed_spec_mp "lemma1"; -goal Transition.thy "!!c s s1. -c-> s1 ==> (c,s) -*-> (SKIP,s1)"; +Goal "!!c s s1. -c-> s1 ==> (c,s) -*-> (SKIP,s1)"; by (etac evalc.induct 1); (* SKIP *) @@ -62,7 +62,7 @@ qed "evalc_impl_evalc1"; -goal Transition.thy +Goal "!c d s u. (c;d,s) -n-> (SKIP,u) --> \ \ (? t m. (c,s) -*-> (SKIP,t) & (d,t) -m-> (SKIP,u) & m <= n)"; by (induct_tac "n" 1); @@ -74,7 +74,7 @@ addSEs [rel_pow_imp_rtrancl,rtrancl_into_rtrancl2]) 1); qed_spec_mp "lemma2"; -goal Transition.thy "!s t. (c,s) -*-> (SKIP,t) --> -c-> t"; +Goal "!s t. (c,s) -*-> (SKIP,t) --> -c-> t"; by (induct_tac "c" 1); by (safe_tac (claset() addSDs [rtrancl_imp_UN_rel_pow])); @@ -112,14 +112,14 @@ (**** proof of the equivalence of evalc and evalc1 ****) -goal Transition.thy "((c, s) -*-> (SKIP, t)) = ( -c-> t)"; +Goal "((c, s) -*-> (SKIP, t)) = ( -c-> t)"; by (fast_tac (HOL_cs addSEs [evalc1_impl_evalc,evalc_impl_evalc1]) 1); qed "evalc1_eq_evalc"; section "A Proof Without -n->"; -goal Transition.thy +Goal "!!c1. (c1,s1) -*-> (SKIP,s2) ==> \ \ (c2,s2) -*-> cs3 --> (c1;c2,s1) -*-> cs3"; by (etac converse_rtrancl_induct2 1); @@ -128,7 +128,7 @@ qed_spec_mp "my_lemma1"; -goal Transition.thy "!!c s s1. -c-> s1 ==> (c,s) -*-> (SKIP,s1)"; +Goal "!!c s s1. -c-> s1 ==> (c,s) -*-> (SKIP,s1)"; by (etac evalc.induct 1); (* SKIP *) @@ -185,7 +185,7 @@ *) (*Delsimps [update_apply];*) -goal Transition.thy +Goal "!!c s. ((c,s) -1-> (c',s')) ==> (!t. -c-> t --> -c-> t)"; by (etac evalc1.induct 1); by Auto_tac; @@ -199,6 +199,6 @@ by (fast_tac (claset() addIs [FB_lemma3]) 1); qed_spec_mp "FB_lemma2"; -goal Transition.thy "!!c. (c,s) -*-> (SKIP,t) ==> -c-> t"; +Goal "!!c. (c,s) -*-> (SKIP,t) ==> -c-> t"; by (fast_tac (claset() addEs [FB_lemma2]) 1); qed "evalc1_impl_evalc"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IMP/VC.ML --- a/src/HOL/IMP/VC.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IMP/VC.ML Mon Jun 22 17:26:46 1998 +0200 @@ -12,7 +12,7 @@ val lemma = prove_goal HOL.thy "!s. P s --> P s" (K[Fast_tac 1]); -goal VC.thy "!Q. (!s. vc c Q s) --> |- {awp c Q} astrip c {Q}"; +Goal "!Q. (!s. vc c Q s) --> |- {awp c Q} astrip c {Q}"; by (acom.induct_tac "c" 1); by (ALLGOALS Simp_tac); by (Fast_tac 1); @@ -29,7 +29,7 @@ by (ALLGOALS(fast_tac HOL_cs)); qed "vc_sound"; -goal VC.thy "!P Q. (!s. P s --> Q s) --> (!s. awp c P s --> awp c Q s)"; +Goal "!P Q. (!s. P s --> Q s) --> (!s. awp c P s --> awp c Q s)"; by (acom.induct_tac "c" 1); by (ALLGOALS Asm_simp_tac); by (EVERY1[rtac allI, rtac allI, rtac impI]); @@ -37,7 +37,7 @@ by (EVERY1[etac allE, etac allE, etac mp, atac]); qed_spec_mp "awp_mono"; -goal VC.thy "!P Q. (!s. P s --> Q s) --> (!s. vc c P s --> vc c Q s)"; +Goal "!P Q. (!s. P s --> Q s) --> (!s. vc c P s --> vc c Q s)"; by (acom.induct_tac "c" 1); by (ALLGOALS Asm_simp_tac); by (safe_tac HOL_cs); @@ -45,7 +45,7 @@ by (fast_tac (HOL_cs addEs [awp_mono]) 1); qed_spec_mp "vc_mono"; -goal VC.thy +Goal "!!P c Q. |- {P}c{Q} ==> \ \ (? ac. astrip ac = c & (!s. vc ac Q s) & (!s. P s --> awp ac Q s))"; by (etac hoare.induct 1); @@ -69,7 +69,7 @@ by (fast_tac (HOL_cs addSEs [awp_mono,vc_mono]) 1); qed "vc_complete"; -goal VC.thy "!Q. vcawp c Q = (vc c Q, awp c Q)"; +Goal "!Q. vcawp c Q = (vc c Q, awp c Q)"; by (acom.induct_tac "c" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [Let_def]))); qed "vcawp_vc_awp"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IOA/Asig.ML --- a/src/HOL/IOA/Asig.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IOA/Asig.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,10 +10,10 @@ val asig_projections = [asig_inputs_def, asig_outputs_def, asig_internals_def]; -goal Asig.thy "!!a.[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)"; +Goal "!!a.[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)"; by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1); qed"int_and_ext_is_act"; -goal Asig.thy "!!a.[|a:externals(S)|] ==> a:actions(S)"; +Goal "!!a.[|a:externals(S)|] ==> a:actions(S)"; by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1); qed"ext_is_act"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IOA/IOA.ML --- a/src/HOL/IOA/IOA.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IOA/IOA.ML Mon Jun 22 17:26:46 1998 +0200 @@ -14,12 +14,12 @@ val exec_rws = [executions_def,is_execution_fragment_def]; -goal IOA.thy +Goal "asig_of((x,y,z)) = x & starts_of((x,y,z)) = y & trans_of((x,y,z)) = z"; by (simp_tac (simpset() addsimps ioa_projections) 1); qed "ioa_triple_proj"; -goalw IOA.thy [ioa_def,state_trans_def,actions_def, is_asig_def] +Goalw [ioa_def,state_trans_def,actions_def, is_asig_def] "!!A. [| IOA(A); (s1,a,s2):trans_of(A) |] ==> a:actions(asig_of(A))"; by (REPEAT(etac conjE 1)); by (EVERY1[etac allE, etac impE, atac]); @@ -27,7 +27,7 @@ qed "trans_in_actions"; -goal IOA.thy "filter_oseq p (filter_oseq p s) = filter_oseq p s"; +Goal "filter_oseq p (filter_oseq p s) = filter_oseq p s"; by (simp_tac (simpset() addsimps [filter_oseq_def]) 1); by (rtac ext 1); by (exhaust_tac "s(i)" 1); @@ -35,7 +35,7 @@ by (Asm_simp_tac 1); qed "filter_oseq_idemp"; -goalw IOA.thy [mk_trace_def,filter_oseq_def] +Goalw [mk_trace_def,filter_oseq_def] "(mk_trace A s n = None) = \ \ (s(n)=None | (? a. s(n)=Some(a) & a ~: externals(asig_of(A)))) \ \ & \ @@ -46,13 +46,13 @@ by (Fast_tac 1); qed "mk_trace_thm"; -goalw IOA.thy [reachable_def] "!!A. s:starts_of(A) ==> reachable A s"; +Goalw [reachable_def] "!!A. s:starts_of(A) ==> reachable A s"; by (res_inst_tac [("x","(%i. None,%i. s)")] bexI 1); by (Simp_tac 1); by (asm_simp_tac (simpset() addsimps exec_rws) 1); qed "reachable_0"; -goalw IOA.thy (reachable_def::exec_rws) +Goalw (reachable_def::exec_rws) "!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t"; by (asm_full_simp_tac (simpset() delsimps bex_simps) 1); by (split_all_tac 1); @@ -104,26 +104,26 @@ by (rtac (p2 RS (p1 RS spec RS mp)) 1); qed "invariantE"; -goal IOA.thy +Goal "actions(asig_comp a b) = actions(a) Un actions(b)"; by (simp_tac (simpset() addsimps ([actions_def,asig_comp_def]@asig_projections)) 1); by (Fast_tac 1); qed "actions_asig_comp"; -goal IOA.thy +Goal "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}"; by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1); qed "starts_of_par"; (* Every state in an execution is reachable *) -goalw IOA.thy [reachable_def] +Goalw [reachable_def] "!!A. ex:executions(A) ==> !n. reachable A (snd ex n)"; by (Fast_tac 1); qed "states_of_exec_reachable"; -goal IOA.thy +Goal "(s,a,t) : trans_of(A || B || C || D) = \ \ ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) | \ \ a:actions(asig_of(D))) & \ @@ -141,32 +141,32 @@ ioa_projections)) 1); qed "trans_of_par4"; -goal IOA.thy "starts_of(restrict ioa acts) = starts_of(ioa) & \ +Goal "starts_of(restrict ioa acts) = starts_of(ioa) & \ \ trans_of(restrict ioa acts) = trans_of(ioa) & \ \ reachable (restrict ioa acts) s = reachable ioa s"; by (simp_tac (simpset() addsimps ([is_execution_fragment_def,executions_def, reachable_def,restrict_def]@ioa_projections)) 1); qed "cancel_restrict"; -goal IOA.thy "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)"; +Goal "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)"; by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1); qed "asig_of_par"; -goal IOA.thy "externals(asig_of(A1||A2)) = \ +Goal "externals(asig_of(A1||A2)) = \ \ (externals(asig_of(A1)) Un externals(asig_of(A2)))"; by (asm_full_simp_tac (simpset() addsimps [externals_def,asig_of_par,asig_comp_def,asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1); by (rtac set_ext 1); by (Fast_tac 1); qed"externals_of_par"; -goalw IOA.thy [externals_def,actions_def,compat_ioas_def,compat_asigs_def] +Goalw [externals_def,actions_def,compat_ioas_def,compat_asigs_def] "!! a. [| compat_ioas A1 A2; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))"; by (Asm_full_simp_tac 1); by (best_tac (claset() addEs [equalityCE]) 1); qed"ext1_is_not_int2"; -goalw IOA.thy [externals_def,actions_def,compat_ioas_def,compat_asigs_def] +Goalw [externals_def,actions_def,compat_ioas_def,compat_asigs_def] "!! a. [| compat_ioas A2 A1 ; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))"; by (Asm_full_simp_tac 1); by (best_tac (claset() addEs [equalityCE]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/IOA/Solve.ML --- a/src/HOL/IOA/Solve.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/IOA/Solve.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,7 +10,7 @@ Addsimps [mk_trace_thm,trans_in_actions]; -goalw Solve.thy [is_weak_pmap_def,traces_def] +Goalw [is_weak_pmap_def,traces_def] "!!f. [| IOA(C); IOA(A); externals(asig_of(C)) = externals(asig_of(A)); \ \ is_weak_pmap f C A |] ==> traces(C) <= traces(A)"; @@ -65,7 +65,7 @@ by (Fast_tac 1); val externals_of_par_extra = result(); -goal Solve.thy "!!s.[| reachable (C1||C2) s |] ==> reachable C1 (fst s)"; +Goal "!!s.[| reachable (C1||C2) s |] ==> reachable C1 (fst s)"; by (asm_full_simp_tac (simpset() addsimps [reachable_def]) 1); by (etac bexE 1); by (res_inst_tac [("x", @@ -85,7 +85,7 @@ (* Exact copy of proof of comp1_reachable for the second component of a parallel composition. *) -goal Solve.thy "!!s.[| reachable (C1||C2) s|] ==> reachable C2 (snd s)"; +Goal "!!s.[| reachable (C1||C2) s|] ==> reachable C2 (snd s)"; by (asm_full_simp_tac (simpset() addsimps [reachable_def]) 1); by (etac bexE 1); by (res_inst_tac [("x", @@ -105,7 +105,7 @@ Delsplits [split_if]; (* Composition of possibility-mappings *) -goalw Solve.thy [is_weak_pmap_def] "!!f g.[| is_weak_pmap f C1 A1 & \ +Goalw [is_weak_pmap_def] "!!f g.[| is_weak_pmap f C1 A1 & \ \ externals(asig_of(A1))=externals(asig_of(C1)) &\ \ is_weak_pmap g C2 A2 & \ \ externals(asig_of(A2))=externals(asig_of(C2)) & \ @@ -149,7 +149,7 @@ qed"fxg_is_weak_pmap_of_product_IOA"; -goal Solve.thy "!!s.[| reachable (rename C g) s |] ==> reachable C s"; +Goal "!!s.[| reachable (rename C g) s |] ==> reachable C s"; by (asm_full_simp_tac (simpset() addsimps [reachable_def]) 1); by (etac bexE 1); by (res_inst_tac [("x", @@ -166,7 +166,7 @@ qed"reachable_rename_ioa"; -goal Solve.thy "!!f.[| is_weak_pmap f C A |]\ +Goal "!!f.[| is_weak_pmap f C A |]\ \ ==> (is_weak_pmap f (rename C g) (rename A g))"; by (asm_full_simp_tac (simpset() addsimps [is_weak_pmap_def]) 1); by (rtac conjI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Acc.ML --- a/src/HOL/Induct/Acc.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Acc.ML Mon Jun 22 17:26:46 1998 +0200 @@ -18,7 +18,7 @@ map (rewrite_rule [pred_def]) acc.intrs)) 1); qed "accI"; -goal Acc.thy "!!a b r. [| b: acc(r); (a,b): r |] ==> a: acc(r)"; +Goal "!!a b r. [| b: acc(r); (a,b): r |] ==> a: acc(r)"; by (etac acc.elim 1); by (rewtac pred_def); by (Fast_tac 1); @@ -58,6 +58,6 @@ major RS acc_wfD_lemma RS spec RS mp]) 1); qed "acc_wfD"; -goal Acc.thy "wf(r) = (r <= (acc r) Times (acc r))"; +Goal "wf(r) = (r <= (acc r) Times (acc r))"; by (EVERY1 [rtac iffI, etac acc_wfD, etac acc_wfI]); qed "wf_acc_iff"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Com.ML --- a/src/HOL/Induct/Com.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Com.ML Mon Jun 22 17:26:46 1998 +0200 @@ -19,7 +19,7 @@ AddSEs exec_elim_cases; (*This theorem justifies using "exec" in the inductive definition of "eval"*) -goalw thy exec.defs "!!A B. A<=B ==> exec(A) <= exec(B)"; +Goalw exec.defs "!!A B. A<=B ==> exec(A) <= exec(B)"; by (rtac lfp_mono 1); by (REPEAT (ares_tac basic_monos 1)); qed "exec_mono"; @@ -29,7 +29,7 @@ Unify.search_bound := 60; (*Command execution is functional (deterministic) provided evaluation is*) -goal thy "!!x. Function ev ==> Function(exec ev)"; +Goal "!!x. Function ev ==> Function(exec ev)"; by (simp_tac (simpset() addsimps [Function_def, Unique_def]) 1); by (REPEAT (rtac allI 1)); by (rtac impI 1); @@ -44,15 +44,15 @@ qed "Function_exec"; -goalw thy [assign_def] "(s[v/x])x = v"; +Goalw [assign_def] "(s[v/x])x = v"; by (Simp_tac 1); qed "assign_same"; -goalw thy [assign_def] "!!y. y~=x ==> (s[v/x])y = s y"; +Goalw [assign_def] "!!y. y~=x ==> (s[v/x])y = s y"; by (Asm_simp_tac 1); qed "assign_different"; -goalw thy [assign_def] "s[s x/x] = s"; +Goalw [assign_def] "s[s x/x] = s"; by (rtac ext 1); by (Simp_tac 1); qed "assign_triv"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Comb.ML --- a/src/HOL/Induct/Comb.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Comb.ML Mon Jun 22 17:26:46 1998 +0200 @@ -24,7 +24,7 @@ val [_, spec_mp] = [spec] RL [mp]; (*Strip lemma. The induction hyp is all but the last diamond of the strip.*) -goalw Comb.thy [diamond_def] +Goalw [diamond_def] "!!r. [| diamond(r); (x,y):r^* |] ==> \ \ ALL y'. (x,y'):r --> (EX z. (y',z): r^* & (y,z): r)"; by (etac rtrancl_induct 1); @@ -55,41 +55,41 @@ AddIs [contract.Ap1, contract.Ap2]; AddSEs [K_contractE, S_contractE, Ap_contractE]; -goalw Comb.thy [I_def] "!!z. I -1-> z ==> P"; +Goalw [I_def] "!!z. I -1-> z ==> P"; by (Blast_tac 1); qed "I_contract_E"; AddSEs [I_contract_E]; -goal Comb.thy "!!x z. K#x -1-> z ==> (EX x'. z = K#x' & x -1-> x')"; +Goal "!!x z. K#x -1-> z ==> (EX x'. z = K#x' & x -1-> x')"; by (Blast_tac 1); qed "K1_contractD"; AddSEs [K1_contractD]; -goal Comb.thy "!!x z. x ---> y ==> x#z ---> y#z"; +Goal "!!x z. x ---> y ==> x#z ---> y#z"; by (etac rtrancl_induct 1); by (ALLGOALS (blast_tac (claset() addIs [r_into_rtrancl, rtrancl_trans]))); qed "Ap_reduce1"; -goal Comb.thy "!!x z. x ---> y ==> z#x ---> z#y"; +Goal "!!x z. x ---> y ==> z#x ---> z#y"; by (etac rtrancl_induct 1); by (ALLGOALS (blast_tac (claset() addIs [r_into_rtrancl, rtrancl_trans]))); qed "Ap_reduce2"; (** Counterexample to the diamond property for -1-> **) -goal Comb.thy "K#I#(I#I) -1-> I"; +Goal "K#I#(I#I) -1-> I"; by (rtac contract.K 1); qed "KIII_contract1"; -goalw Comb.thy [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))"; +Goalw [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))"; by (Blast_tac 1); qed "KIII_contract2"; -goal Comb.thy "K#I#((K#I)#(K#I)) -1-> I"; +Goal "K#I#((K#I)#(K#I)) -1-> I"; by (Blast_tac 1); qed "KIII_contract3"; -goalw Comb.thy [diamond_def] "~ diamond(contract)"; +Goalw [diamond_def] "~ diamond(contract)"; by (blast_tac (claset() addIs [KIII_contract1,KIII_contract2,KIII_contract3]) 1); qed "not_diamond_contract"; @@ -108,17 +108,17 @@ (*** Basic properties of parallel contraction ***) -goal Comb.thy "!!x z. K#x =1=> z ==> (EX x'. z = K#x' & x =1=> x')"; +Goal "!!x z. K#x =1=> z ==> (EX x'. z = K#x' & x =1=> x')"; by (Blast_tac 1); qed "K1_parcontractD"; AddSDs [K1_parcontractD]; -goal Comb.thy "!!x z. S#x =1=> z ==> (EX x'. z = S#x' & x =1=> x')"; +Goal "!!x z. S#x =1=> z ==> (EX x'. z = S#x' & x =1=> x')"; by (Blast_tac 1); qed "S1_parcontractD"; AddSDs [S1_parcontractD]; -goal Comb.thy +Goal "!!x y z. S#x#y =1=> z ==> (EX x' y'. z = S#x'#y' & x =1=> x' & y =1=> y')"; by (Blast_tac 1); qed "S2_parcontractD"; @@ -128,7 +128,7 @@ (*Church-Rosser property for parallel contraction*) -goalw Comb.thy [diamond_def] "diamond parcontract"; +Goalw [diamond_def] "diamond parcontract"; by (rtac (impI RS allI RS allI) 1); by (etac parcontract.induct 1 THEN prune_params_tac); by Safe_tac; @@ -138,7 +138,7 @@ (*** Equivalence of x--->y and x===>y ***) -goal Comb.thy "contract <= parcontract"; +Goal "contract <= parcontract"; by (rtac subsetI 1); by (split_all_tac 1); by (etac contract.induct 1); @@ -151,25 +151,25 @@ AddIs [Ap_reduce1, Ap_reduce2, r_into_rtrancl, rtrancl_trans]; (*Example only: not used*) -goalw Comb.thy [I_def] "I#x ---> x"; +Goalw [I_def] "I#x ---> x"; by (Blast_tac 1); qed "reduce_I"; -goal Comb.thy "parcontract <= contract^*"; +Goal "parcontract <= contract^*"; by (rtac subsetI 1); by (split_all_tac 1); by (etac parcontract.induct 1 THEN prune_params_tac); by (ALLGOALS Blast_tac); qed "parcontract_subset_reduce"; -goal Comb.thy "contract^* = parcontract^*"; +Goal "contract^* = parcontract^*"; by (REPEAT (resolve_tac [equalityI, contract_subset_parcontract RS rtrancl_mono, parcontract_subset_reduce RS rtrancl_subset_rtrancl] 1)); qed "reduce_eq_parreduce"; -goal Comb.thy "diamond(contract^*)"; +Goal "diamond(contract^*)"; by (simp_tac (simpset() addsimps [reduce_eq_parreduce, diamond_rtrancl, diamond_parcontract]) 1); qed "diamond_reduce"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Exp.ML --- a/src/HOL/Induct/Exp.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Exp.ML Mon Jun 22 17:26:46 1998 +0200 @@ -20,7 +20,7 @@ AddSEs eval_elim_cases; -goal thy "(X x, s[n/x]) -|-> (n, s[n/x])"; +Goal "(X x, s[n/x]) -|-> (n, s[n/x])"; by (rtac (assign_same RS subst) 1 THEN resolve_tac eval.intrs 1); qed "var_assign_eval"; @@ -30,7 +30,7 @@ (** Make the induction rule look nicer -- though eta_contract makes the new version look worse than it is...**) -goal thy "{((e,s),(n,s')). P e s n s'} = \ +Goal "{((e,s),(n,s')). P e s n s'} = \ \ Collect (split (%v. split (split P v)))"; by (rtac Collect_cong 1); by (split_all_tac 1); @@ -67,7 +67,7 @@ the execution (c,s) -[eval Int {...}]-> s1 assures us that execution is functional on the argument (c,s). *) -goal thy +Goal "!!x. (c,s) -[eval Int {((e,s),(n,s')). Unique (e,s) (n,s') eval}]-> s1 \ \ ==> (ALL s2. (c,s) -[eval]-> s2 --> s2=s1)"; by (etac exec.induct 1); @@ -87,7 +87,7 @@ (*Expression evaluation is functional, or deterministic*) -goalw thy [Function_def] "Function eval"; +Goalw [Function_def] "Function eval"; by (strip_tac 1); by (split_all_tac 1); by (etac eval_induct 1); @@ -97,7 +97,7 @@ qed "Function_eval"; -goal thy "!!x. (e,s) -|-> (v,s') ==> (e = N n) --> (v=n & s'=s)"; +Goal "!!x. (e,s) -|-> (v,s') ==> (e = N n) --> (v=n & s'=s)"; by (etac eval_induct 1); by (ALLGOALS Asm_simp_tac); val lemma = result(); @@ -107,7 +107,7 @@ (*This theorem says that "WHILE TRUE DO c" cannot terminate*) -goal thy "!!x. (c', s) -[eval]-> t ==> (c' = WHILE (N 0) DO c) --> False"; +Goal "!!x. (c', s) -[eval]-> t ==> (c' = WHILE (N 0) DO c) --> False"; by (etac exec.induct 1); by Auto_tac; bind_thm ("while_true_E", refl RSN (2, result() RS mp)); @@ -115,7 +115,7 @@ (** Equivalence of IF e THEN c;;(WHILE e DO c) ELSE SKIP and WHILE e DO c **) -goal thy "!!x. (c',s) -[eval]-> t ==> \ +Goal "!!x. (c',s) -[eval]-> t ==> \ \ (c' = WHILE e DO c) --> \ \ (IF e THEN c;;c' ELSE SKIP, s) -[eval]-> t"; by (etac exec.induct 1); @@ -124,7 +124,7 @@ bind_thm ("while_if1", refl RSN (2, result() RS mp)); -goal thy "!!x. (c',s) -[eval]-> t ==> \ +Goal "!!x. (c',s) -[eval]-> t ==> \ \ (c' = IF e THEN c;;(WHILE e DO c) ELSE SKIP) --> \ \ (WHILE e DO c, s) -[eval]-> t"; by (etac exec.induct 1); @@ -133,7 +133,7 @@ bind_thm ("while_if2", refl RSN (2, result() RS mp)); -goal thy "((IF e THEN c;;(WHILE e DO c) ELSE SKIP, s) -[eval]-> t) = \ +Goal "((IF e THEN c;;(WHILE e DO c) ELSE SKIP, s) -[eval]-> t) = \ \ ((WHILE e DO c, s) -[eval]-> t)"; by (blast_tac (claset() addIs [while_if1, while_if2]) 1); qed "while_if"; @@ -143,7 +143,7 @@ (** Equivalence of (IF e THEN c1 ELSE c2);;c and IF e THEN (c1;;c) ELSE (c2;;c) **) -goal thy "!!x. (c',s) -[eval]-> t ==> \ +Goal "!!x. (c',s) -[eval]-> t ==> \ \ (c' = (IF e THEN c1 ELSE c2);;c) --> \ \ (IF e THEN (c1;;c) ELSE (c2;;c), s) -[eval]-> t"; by (etac exec.induct 1); @@ -152,7 +152,7 @@ bind_thm ("if_semi1", refl RSN (2, result() RS mp)); -goal thy "!!x. (c',s) -[eval]-> t ==> \ +Goal "!!x. (c',s) -[eval]-> t ==> \ \ (c' = IF e THEN (c1;;c) ELSE (c2;;c)) --> \ \ ((IF e THEN c1 ELSE c2);;c, s) -[eval]-> t"; by (etac exec.induct 1); @@ -161,7 +161,7 @@ bind_thm ("if_semi2", refl RSN (2, result() RS mp)); -goal thy "(((IF e THEN c1 ELSE c2);;c, s) -[eval]-> t) = \ +Goal "(((IF e THEN c1 ELSE c2);;c, s) -[eval]-> t) = \ \ ((IF e THEN (c1;;c) ELSE (c2;;c), s) -[eval]-> t)"; by (blast_tac (claset() addIs [if_semi1, if_semi2]) 1); qed "if_semi"; @@ -172,7 +172,7 @@ and VALOF c1;;c2 RESULTIS e **) -goal thy "!!x. (e',s) -|-> (v,s') ==> \ +Goal "!!x. (e',s) -|-> (v,s') ==> \ \ (e' = VALOF c1 RESULTIS (VALOF c2 RESULTIS e)) --> \ \ (VALOF c1;;c2 RESULTIS e, s) -|-> (v,s')"; by (etac eval_induct 1); @@ -181,7 +181,7 @@ bind_thm ("valof_valof1", refl RSN (2, result() RS mp)); -goal thy "!!x. (e',s) -|-> (v,s') ==> \ +Goal "!!x. (e',s) -|-> (v,s') ==> \ \ (e' = VALOF c1;;c2 RESULTIS e) --> \ \ (VALOF c1 RESULTIS (VALOF c2 RESULTIS e), s) -|-> (v,s')"; by (etac eval_induct 1); @@ -190,7 +190,7 @@ bind_thm ("valof_valof2", refl RSN (2, result() RS mp)); -goal thy "((VALOF c1 RESULTIS (VALOF c2 RESULTIS e), s) -|-> (v,s')) = \ +Goal "((VALOF c1 RESULTIS (VALOF c2 RESULTIS e), s) -|-> (v,s')) = \ \ ((VALOF c1;;c2 RESULTIS e, s) -|-> (v,s'))"; by (blast_tac (claset() addIs [valof_valof1, valof_valof2]) 1); qed "valof_valof"; @@ -198,7 +198,7 @@ (** Equivalence of VALOF SKIP RESULTIS e and e **) -goal thy "!!x. (e',s) -|-> (v,s') ==> \ +Goal "!!x. (e',s) -|-> (v,s') ==> \ \ (e' = VALOF SKIP RESULTIS e) --> \ \ (e, s) -|-> (v,s')"; by (etac eval_induct 1); @@ -206,18 +206,18 @@ by (Blast_tac 1); bind_thm ("valof_skip1", refl RSN (2, result() RS mp)); -goal thy "!!x. (e,s) -|-> (v,s') ==> (VALOF SKIP RESULTIS e, s) -|-> (v,s')"; +Goal "!!x. (e,s) -|-> (v,s') ==> (VALOF SKIP RESULTIS e, s) -|-> (v,s')"; by (Blast_tac 1); qed "valof_skip2"; -goal thy "((VALOF SKIP RESULTIS e, s) -|-> (v,s')) = ((e, s) -|-> (v,s'))"; +Goal "((VALOF SKIP RESULTIS e, s) -|-> (v,s')) = ((e, s) -|-> (v,s'))"; by (blast_tac (claset() addIs [valof_skip1, valof_skip2]) 1); qed "valof_skip"; (** Equivalence of VALOF x:=e RESULTIS x and e **) -goal thy "!!x. (e',s) -|-> (v,s'') ==> \ +Goal "!!x. (e',s) -|-> (v,s'') ==> \ \ (e' = VALOF x:=e RESULTIS X x) --> \ \ (EX s'. (e, s) -|-> (v,s') & (s'' = s'[v/x]))"; by (etac eval_induct 1); @@ -229,7 +229,7 @@ bind_thm ("valof_assign1", refl RSN (2, result() RS mp)); -goal thy "!!x. (e,s) -|-> (v,s') ==> \ +Goal "!!x. (e,s) -|-> (v,s') ==> \ \ (VALOF x:=e RESULTIS X x, s) -|-> (v,s'[v/x])"; by (Blast_tac 1); qed "valof_assign2"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/LFilter.ML --- a/src/HOL/Induct/LFilter.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/LFilter.ML Mon Jun 22 17:26:46 1998 +0200 @@ -18,19 +18,19 @@ AddSEs [findRel_LConsE]; -goal thy "!!p. (l,l'): findRel p ==> (l,l''): findRel p --> l'' = l'"; +Goal "!!p. (l,l'): findRel p ==> (l,l''): findRel p --> l'' = l'"; by (etac findRel.induct 1); by (Blast_tac 1); by (Blast_tac 1); qed_spec_mp "findRel_functional"; -goal thy "!!p. (l,l'): findRel p ==> EX x l''. l' = LCons x l'' & p x"; +Goal "!!p. (l,l'): findRel p ==> EX x l''. l' = LCons x l'' & p x"; by (etac findRel.induct 1); by (Blast_tac 1); by (Blast_tac 1); qed_spec_mp "findRel_imp_LCons"; -goal thy "!!p. (LNil,l): findRel p ==> R"; +Goal "!!p. (LNil,l): findRel p ==> R"; by (blast_tac (claset() addEs [findRel.elim]) 1); qed "findRel_LNil"; @@ -39,7 +39,7 @@ (*** Properties of Domain (findRel p) ***) -goal thy "!!p. LCons x l : Domain(findRel p) = (p x | l : Domain(findRel p))"; +Goal "!!p. LCons x l : Domain(findRel p) = (p x | l : Domain(findRel p))"; by (case_tac "p x" 1); by (ALLGOALS (blast_tac (claset() addIs findRel.intrs))); qed "LCons_Domain_findRel"; @@ -47,7 +47,7 @@ Addsimps [LCons_Domain_findRel]; val major::prems = -goal thy "[| l: Domain (findRel p); \ +Goal "[| l: Domain (findRel p); \ \ !!x l'. [| (l, LCons x l') : findRel p; p x |] ==> Q \ \ |] ==> Q"; by (rtac (major RS DomainE) 1); @@ -68,27 +68,27 @@ (*** find: basic equations ***) -goalw thy [find_def] "find p LNil = LNil"; +Goalw [find_def] "find p LNil = LNil"; by (Blast_tac 1); qed "find_LNil"; Addsimps [find_LNil]; -goalw thy [find_def] "!!p. (l,l') : findRel p ==> find p l = l'"; +Goalw [find_def] "!!p. (l,l') : findRel p ==> find p l = l'"; by (blast_tac (claset() addDs [findRel_functional]) 1); qed "findRel_imp_find"; Addsimps [findRel_imp_find]; -goal thy "!!p. p x ==> find p (LCons x l) = LCons x l"; +Goal "!!p. p x ==> find p (LCons x l) = LCons x l"; by (blast_tac (claset() addIs (findRel_imp_find::findRel.intrs)) 1); qed "find_LCons_found"; Addsimps [find_LCons_found]; -goalw thy [find_def] "!!p. l ~: Domain(findRel p) ==> find p l = LNil"; +Goalw [find_def] "!!p. l ~: Domain(findRel p) ==> find p l = LNil"; by (Blast_tac 1); qed "diverge_find_LNil"; Addsimps [diverge_find_LNil]; -goal thy "!!p. ~ (p x) ==> find p (LCons x l) = find p l"; +Goal "!!p. ~ (p x) ==> find p (LCons x l) = find p l"; by (case_tac "LCons x l : Domain(findRel p)" 1); by (Asm_full_simp_tac 2); by (Clarify_tac 1); @@ -97,7 +97,7 @@ qed "find_LCons_seek"; Addsimps [find_LCons_seek]; -goal thy "find p (LCons x l) = (if p x then LCons x l else find p l)"; +Goal "find p (LCons x l) = (if p x then LCons x l else find p l)"; by (Asm_simp_tac 1); qed "find_LCons"; @@ -105,13 +105,13 @@ (*** lfilter: basic equations ***) -goal thy "lfilter p LNil = LNil"; +Goal "lfilter p LNil = LNil"; by (rtac (lfilter_def RS def_llist_corec RS trans) 1); by (Simp_tac 1); qed "lfilter_LNil"; Addsimps [lfilter_LNil]; -goal thy "!!p. l ~: Domain(findRel p) ==> lfilter p l = LNil"; +Goal "!!p. l ~: Domain(findRel p) ==> lfilter p l = LNil"; by (rtac (lfilter_def RS def_llist_corec RS trans) 1); by (Asm_simp_tac 1); qed "diverge_lfilter_LNil"; @@ -119,7 +119,7 @@ Addsimps [diverge_lfilter_LNil]; -goal thy "!!p. p x ==> lfilter p (LCons x l) = LCons x (lfilter p l)"; +Goal "!!p. p x ==> lfilter p (LCons x l) = LCons x (lfilter p l)"; by (rtac (lfilter_def RS def_llist_corec RS trans) 1); by (Asm_simp_tac 1); qed "lfilter_LCons_found"; @@ -127,7 +127,7 @@ subsumes both*) -goal thy "!!p. (l, LCons x l') : findRel p \ +Goal "!!p. (l, LCons x l') : findRel p \ \ ==> lfilter p l = LCons x (lfilter p l')"; by (rtac (lfilter_def RS def_llist_corec RS trans) 1); by (Asm_simp_tac 1); @@ -136,7 +136,7 @@ Addsimps [findRel_imp_lfilter]; -goal thy "!!p. ~ (p x) ==> lfilter p (LCons x l) = lfilter p l"; +Goal "!!p. ~ (p x) ==> lfilter p (LCons x l) = lfilter p l"; by (rtac (lfilter_def RS def_llist_corec RS trans) 1); by (case_tac "LCons x l : Domain(findRel p)" 1); by (Asm_full_simp_tac 2); @@ -146,7 +146,7 @@ qed "lfilter_LCons_seek"; -goal thy "lfilter p (LCons x l) = \ +Goal "lfilter p (LCons x l) = \ \ (if p x then LCons x (lfilter p l) else lfilter p l)"; by (asm_simp_tac (simpset() addsimps [lfilter_LCons_found, lfilter_LCons_seek]) 1); qed "lfilter_LCons"; @@ -155,7 +155,7 @@ Addsimps [lfilter_LCons]; -goal thy "!!p. lfilter p l = LNil ==> l ~: Domain(findRel p)"; +Goal "!!p. lfilter p l = LNil ==> l ~: Domain(findRel p)"; by (rtac notI 1); by (etac Domain_findRelE 1); by (etac rev_mp 1); @@ -163,7 +163,7 @@ qed "lfilter_eq_LNil"; -goal thy "!!p. lfilter p l = LCons x l' --> \ +Goal "!!p. lfilter p l = LCons x l' --> \ \ (EX l''. l' = lfilter p l'' & (l, LCons x l'') : findRel p)"; by (stac (lfilter_def RS def_llist_corec) 1); by (case_tac "l : Domain(findRel p)" 1); @@ -174,7 +174,7 @@ qed_spec_mp "lfilter_eq_LCons"; -goal thy "lfilter p l = LNil | \ +Goal "lfilter p l = LNil | \ \ (EX y l'. lfilter p l = LCons y (lfilter p l') & p y)"; by (case_tac "l : Domain(findRel p)" 1); by (Asm_simp_tac 2); @@ -185,12 +185,12 @@ (*** lfilter: simple facts by coinduction ***) -goal thy "lfilter (%x. True) l = l"; +Goal "lfilter (%x. True) l = l"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); qed "lfilter_K_True"; -goal thy "lfilter p (lfilter p l) = lfilter p l"; +Goal "lfilter p (lfilter p l) = lfilter p l"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); by Safe_tac; @@ -205,7 +205,7 @@ lfilter p (lfilter q l) = lfilter (%x. p x & q x) l ***) -goal thy "!!p. (l,l') : findRel q \ +Goal "!!p. (l,l') : findRel q \ \ ==> l' = LCons x l'' --> p x --> (l,l') : findRel (%x. p x & q x)"; by (etac findRel.induct 1); by (blast_tac (claset() addIs findRel.intrs) 1); @@ -214,7 +214,7 @@ val findRel_conj = refl RSN (2, findRel_conj_lemma); -goal thy "!!p. (l,l'') : findRel (%x. p x & q x) \ +Goal "!!p. (l,l'') : findRel (%x. p x & q x) \ \ ==> (l, LCons x l') : findRel q --> ~ p x \ \ --> l' : Domain (findRel (%x. p x & q x))"; by (etac findRel.induct 1); @@ -222,7 +222,7 @@ qed_spec_mp "findRel_not_conj_Domain"; -goal thy "!!p. (l,lxx) : findRel q ==> \ +Goal "!!p. (l,lxx) : findRel q ==> \ \ lxx = LCons x lx --> (lx,lz) : findRel(%x. p x & q x) --> ~ p x \ \ --> (l,lz) : findRel (%x. p x & q x)"; by (etac findRel.induct 1); @@ -230,7 +230,7 @@ qed_spec_mp "findRel_conj2"; -goal thy "!!p. (lx,ly) : findRel p \ +Goal "!!p. (lx,ly) : findRel p \ \ ==> ALL l. lx = lfilter q l \ \ --> l : Domain (findRel(%x. p x & q x))"; by (etac findRel.induct 1); @@ -245,7 +245,7 @@ qed_spec_mp "findRel_lfilter_Domain_conj"; -goal thy "!!p. (l,l'') : findRel(%x. p x & q x) \ +Goal "!!p. (l,l'') : findRel(%x. p x & q x) \ \ ==> l'' = LCons y l' --> \ \ (lfilter q l, LCons y (lfilter q l')) : findRel p"; by (etac findRel.induct 1); @@ -255,7 +255,7 @@ -goal thy "(lfilter p (lfilter q l), lfilter (%x. p x & q x) l) \ +Goal "(lfilter p (lfilter q l), lfilter (%x. p x & q x) l) \ \ : llistD_Fun (range \ \ (%u. (lfilter p (lfilter q u), \ \ lfilter (%x. p x & q x) u)))"; @@ -288,7 +288,7 @@ val lemma = result(); -goal thy "lfilter p (lfilter q l) = lfilter (%x. p x & q x) l"; +Goal "lfilter p (lfilter q l) = lfilter (%x. p x & q x) l"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); by (blast_tac (claset() addIs [lemma, impOfSubs llistD_Fun_mono]) 1); @@ -299,13 +299,13 @@ lfilter p (lmap f l) = lmap f (lfilter (%x. p(f x)) l) ***) -goal thy "!!p. (l,l') : findRel(%x. p (f x)) ==> lmap f l : Domain(findRel p)"; +Goal "!!p. (l,l') : findRel(%x. p (f x)) ==> lmap f l : Domain(findRel p)"; by (etac findRel.induct 1); by (ALLGOALS Asm_full_simp_tac); qed "findRel_lmap_Domain"; -goal thy "!!p. lmap f l = LCons x l' --> \ +Goal "!!p. lmap f l = LCons x l' --> \ \ (EX y l''. x = f y & l' = lmap f l'' & l = LCons y l'')"; by (stac (lmap_def RS def_llist_corec) 1); by (res_inst_tac [("l", "l")] llistE 1); @@ -313,7 +313,7 @@ qed_spec_mp "lmap_eq_LCons"; -goal thy "!!p. (lx,ly) : findRel p ==> \ +Goal "!!p. (lx,ly) : findRel p ==> \ \ ALL l. lmap f l = lx --> ly = LCons x l' --> \ \ (EX y l''. x = f y & l' = lmap f l'' & \ \ (l, LCons y l'') : findRel(%x. p(f x)))"; @@ -326,7 +326,7 @@ val lmap_LCons_findRel = refl RSN (2, refl RSN (2, lmap_LCons_findRel_lemma)); -goal thy "lfilter p (lmap f l) = lmap f (lfilter (p o f) l)"; +Goal "lfilter p (lmap f l) = lmap f (lfilter (p o f) l)"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); by Safe_tac; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/LList.ML --- a/src/HOL/Induct/LList.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/LList.ML Mon Jun 22 17:26:46 1998 +0200 @@ -14,13 +14,13 @@ (*This justifies using llist in other recursive type definitions*) -goalw LList.thy llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)"; +Goalw llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)"; by (rtac gfp_mono 1); by (REPEAT (ares_tac basic_monos 1)); qed "llist_mono"; -goal LList.thy "llist(A) = {Numb(0)} <+> (A <*> llist(A))"; +Goal "llist(A) = {Numb(0)} <+> (A <*> llist(A))"; let val rew = rewrite_rule [NIL_def, CONS_def] in by (fast_tac (claset() addSIs (map rew llist.intrs) addEs [rew llist.elim]) 1) @@ -32,19 +32,19 @@ THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! ***) -goalw LList.thy [list_Fun_def] +Goalw [list_Fun_def] "!!M. [| M : X; X <= list_Fun A (X Un llist(A)) |] ==> M : llist(A)"; by (etac llist.coinduct 1); by (etac (subsetD RS CollectD) 1); by (assume_tac 1); qed "llist_coinduct"; -goalw LList.thy [list_Fun_def, NIL_def] "NIL: list_Fun A X"; +Goalw [list_Fun_def, NIL_def] "NIL: list_Fun A X"; by (Fast_tac 1); qed "list_Fun_NIL_I"; AddIffs [list_Fun_NIL_I]; -goalw LList.thy [list_Fun_def,CONS_def] +Goalw [list_Fun_def,CONS_def] "!!M N. [| M: A; N: X |] ==> CONS M N : list_Fun A X"; by (Fast_tac 1); qed "list_Fun_CONS_I"; @@ -52,7 +52,7 @@ AddSIs [list_Fun_CONS_I]; (*Utilise the "strong" part, i.e. gfp(f)*) -goalw LList.thy (llist.defs @ [list_Fun_def]) +Goalw (llist.defs @ [list_Fun_def]) "!!M N. M: llist(A) ==> M : list_Fun A (X Un llist(A))"; by (etac (llist.mono RS gfp_fun_UnI2) 1); qed "list_Fun_llist_I"; @@ -60,7 +60,7 @@ (*** LList_corec satisfies the desired recurion equation ***) (*A continuity result?*) -goalw LList.thy [CONS_def] "CONS M (UN x. f(x)) = (UN x. CONS M (f x))"; +Goalw [CONS_def] "CONS M (UN x. f(x)) = (UN x. CONS M (f x))"; by (simp_tac (simpset() addsimps [In1_UN1, Scons_UN1_y]) 1); qed "CONS_UN1"; @@ -84,7 +84,7 @@ (** The directions of the equality are proved separately **) -goalw LList.thy [LList_corec_def] +Goalw [LList_corec_def] "LList_corec a f <= sum_case (%u. NIL) \ \ (split(%z w. CONS z (LList_corec w f))) (f a)"; by (rtac UN_least 1); @@ -94,7 +94,7 @@ UNIV_I RS UN_upper] 1)); qed "LList_corec_subset1"; -goalw LList.thy [LList_corec_def] +Goalw [LList_corec_def] "sum_case (%u. NIL) (split(%z w. CONS z (LList_corec w f))) (f a) <= \ \ LList_corec a f"; by (simp_tac (simpset() addsimps [CONS_UN1]) 1); @@ -104,7 +104,7 @@ qed "LList_corec_subset2"; (*the recursion equation for LList_corec -- NOT SUITABLE FOR REWRITING!*) -goal LList.thy +Goal "LList_corec a f = sum_case (%u. NIL) \ \ (split(%z w. CONS z (LList_corec w f))) (f a)"; by (REPEAT (resolve_tac [equalityI, LList_corec_subset1, @@ -121,7 +121,7 @@ (*A typical use of co-induction to show membership in the gfp. Bisimulation is range(%x. LList_corec x f) *) -goal LList.thy "LList_corec a f : llist({u. True})"; +Goal "LList_corec a f : llist({u. True})"; by (res_inst_tac [("X", "range(%x. LList_corec x ?g)")] llist_coinduct 1); by (rtac rangeI 1); by Safe_tac; @@ -130,7 +130,7 @@ qed "LList_corec_type"; (*Lemma for the proof of llist_corec*) -goal LList.thy +Goal "LList_corec a (%z. sum_case Inl (split(%v w. Inr((Leaf(v),w)))) (f z)) : \ \ llist(range Leaf)"; by (res_inst_tac [("X", "range(%x. LList_corec x ?g)")] llist_coinduct 1); @@ -144,14 +144,14 @@ (**** llist equality as a gfp; the bisimulation principle ****) (*This theorem is actually used, unlike the many similar ones in ZF*) -goal LList.thy "LListD(r) = diag({Numb(0)}) <++> (r <**> LListD(r))"; +Goal "LListD(r) = diag({Numb(0)}) <++> (r <**> LListD(r))"; let val rew = rewrite_rule [NIL_def, CONS_def] in by (fast_tac (claset() addSIs (map rew LListD.intrs) addEs [rew LListD.elim]) 1) end; qed "LListD_unfold"; -goal LList.thy "!M N. (M,N) : LListD(diag(A)) --> ntrunc k M = ntrunc k N"; +Goal "!M N. (M,N) : LListD(diag(A)) --> ntrunc k M = ntrunc k N"; by (res_inst_tac [("n", "k")] less_induct 1); by (safe_tac ((claset_of Fun.thy) delrules [equalityI])); by (etac LListD.elim 1); @@ -165,7 +165,7 @@ qed "LListD_implies_ntrunc_equality"; (*The domain of the LListD relation*) -goalw LList.thy (llist.defs @ [NIL_def, CONS_def]) +Goalw (llist.defs @ [NIL_def, CONS_def]) "fst``LListD(diag(A)) <= llist(A)"; by (rtac gfp_upperbound 1); (*avoids unfolding LListD on the rhs*) @@ -175,7 +175,7 @@ qed "fst_image_LListD"; (*This inclusion justifies the use of coinduction to show M=N*) -goal LList.thy "LListD(diag(A)) <= diag(llist(A))"; +Goal "LListD(diag(A)) <= diag(llist(A))"; by (rtac subsetI 1); by (res_inst_tac [("p","x")] PairE 1); by Safe_tac; @@ -191,35 +191,35 @@ THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! **) -goalw thy [LListD_Fun_def] "!!A B. A<=B ==> LListD_Fun r A <= LListD_Fun r B"; +Goalw [LListD_Fun_def] "!!A B. A<=B ==> LListD_Fun r A <= LListD_Fun r B"; by (REPEAT (ares_tac basic_monos 1)); qed "LListD_Fun_mono"; -goalw LList.thy [LListD_Fun_def] +Goalw [LListD_Fun_def] "!!M. [| M : X; X <= LListD_Fun r (X Un LListD(r)) |] ==> M : LListD(r)"; by (etac LListD.coinduct 1); by (etac (subsetD RS CollectD) 1); by (assume_tac 1); qed "LListD_coinduct"; -goalw LList.thy [LListD_Fun_def,NIL_def] "(NIL,NIL) : LListD_Fun r s"; +Goalw [LListD_Fun_def,NIL_def] "(NIL,NIL) : LListD_Fun r s"; by (Fast_tac 1); qed "LListD_Fun_NIL_I"; -goalw LList.thy [LListD_Fun_def,CONS_def] +Goalw [LListD_Fun_def,CONS_def] "!!x. [| x:A; (M,N):s |] ==> (CONS x M, CONS x N) : LListD_Fun (diag A) s"; by (Fast_tac 1); qed "LListD_Fun_CONS_I"; (*Utilise the "strong" part, i.e. gfp(f)*) -goalw LList.thy (LListD.defs @ [LListD_Fun_def]) +Goalw (LListD.defs @ [LListD_Fun_def]) "!!M N. M: LListD(r) ==> M : LListD_Fun r (X Un LListD(r))"; by (etac (LListD.mono RS gfp_fun_UnI2) 1); qed "LListD_Fun_LListD_I"; (*This converse inclusion helps to strengthen LList_equalityI*) -goal LList.thy "diag(llist(A)) <= LListD(diag(A))"; +Goal "diag(llist(A)) <= LListD(diag(A))"; by (rtac subsetI 1); by (etac LListD_coinduct 1); by (rtac subsetI 1); @@ -231,12 +231,12 @@ LListD_Fun_CONS_I]))); qed "diag_subset_LListD"; -goal LList.thy "LListD(diag(A)) = diag(llist(A))"; +Goal "LListD(diag(A)) = diag(llist(A))"; by (REPEAT (resolve_tac [equalityI, LListD_subset_diag, diag_subset_LListD] 1)); qed "LListD_eq_diag"; -goal LList.thy +Goal "!!M N. M: llist(A) ==> (M,M) : LListD_Fun (diag A) (X Un diag(llist(A)))"; by (rtac (LListD_eq_diag RS subst) 1); by (rtac LListD_Fun_LListD_I 1); @@ -247,7 +247,7 @@ (** To show two LLists are equal, exhibit a bisimulation! [also admits true equality] Replace "A" by some particular set, like {x.True}??? *) -goal LList.thy +Goal "!!r. [| (M,N) : r; r <= LListD_Fun (diag A) (r Un diag(llist(A))) \ \ |] ==> M=N"; by (rtac (LListD_subset_diag RS subsetD RS diagE) 1); @@ -291,11 +291,11 @@ (** Obsolete LList_corec_unique proof: complete induction, not coinduction **) -goalw LList.thy [CONS_def] "ntrunc (Suc 0) (CONS M N) = {}"; +Goalw [CONS_def] "ntrunc (Suc 0) (CONS M N) = {}"; by (rtac ntrunc_one_In1 1); qed "ntrunc_one_CONS"; -goalw LList.thy [CONS_def] +Goalw [CONS_def] "ntrunc (Suc(Suc(k))) (CONS M N) = CONS (ntrunc k M) (ntrunc k N)"; by (Simp_tac 1); qed "ntrunc_CONS"; @@ -327,7 +327,7 @@ (*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***) -goal LList.thy "mono(CONS(M))"; +Goal "mono(CONS(M))"; by (REPEAT (ares_tac [monoI, subset_refl, CONS_mono] 1)); qed "Lconst_fun_mono"; @@ -336,21 +336,21 @@ (*A typical use of co-induction to show membership in the gfp. The containing set is simply the singleton {Lconst(M)}. *) -goal LList.thy "!!M A. M:A ==> Lconst(M): llist(A)"; +Goal "!!M A. M:A ==> Lconst(M): llist(A)"; by (rtac (singletonI RS llist_coinduct) 1); by Safe_tac; by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1); by (REPEAT (ares_tac [list_Fun_CONS_I, singletonI, UnI1] 1)); qed "Lconst_type"; -goal LList.thy "Lconst(M) = LList_corec M (%x. Inr((x,x)))"; +Goal "Lconst(M) = LList_corec M (%x. Inr((x,x)))"; by (rtac (equals_LList_corec RS fun_cong) 1); by (Simp_tac 1); by (rtac Lconst 1); qed "Lconst_eq_LList_corec"; (*Thus we could have used gfp in the definition of Lconst*) -goal LList.thy "gfp(%N. CONS M N) = LList_corec M (%x. Inr((x,x)))"; +Goal "gfp(%N. CONS M N) = LList_corec M (%x. Inr((x,x)))"; by (rtac (equals_LList_corec RS fun_cong) 1); by (Simp_tac 1); by (rtac (Lconst_fun_mono RS gfp_Tarski) 1); @@ -359,19 +359,19 @@ (*** Isomorphisms ***) -goal LList.thy "inj(Rep_llist)"; +Goal "inj(Rep_llist)"; by (rtac inj_inverseI 1); by (rtac Rep_llist_inverse 1); qed "inj_Rep_llist"; -goal LList.thy "inj_on Abs_llist (llist(range Leaf))"; +Goal "inj_on Abs_llist (llist(range Leaf))"; by (rtac inj_on_inverseI 1); by (etac Abs_llist_inverse 1); qed "inj_on_Abs_llist"; (** Distinctness of constructors **) -goalw LList.thy [LNil_def,LCons_def] "~ LCons x xs = LNil"; +Goalw [LNil_def,LCons_def] "~ LCons x xs = LNil"; by (rtac (CONS_not_NIL RS (inj_on_Abs_llist RS inj_on_contraD)) 1); by (REPEAT (resolve_tac (llist.intrs @ [rangeI, Rep_llist]) 1)); qed "LCons_not_LNil"; @@ -383,12 +383,12 @@ (** llist constructors **) -goalw LList.thy [LNil_def] +Goalw [LNil_def] "Rep_llist(LNil) = NIL"; by (rtac (llist.NIL_I RS Abs_llist_inverse) 1); qed "Rep_llist_LNil"; -goalw LList.thy [LCons_def] +Goalw [LCons_def] "Rep_llist(LCons x l) = CONS (Leaf x) (Rep_llist l)"; by (REPEAT (resolve_tac [llist.CONS_I RS Abs_llist_inverse, rangeI, Rep_llist] 1)); @@ -396,7 +396,7 @@ (** Injectiveness of CONS and LCons **) -goalw LList.thy [CONS_def] "(CONS M N=CONS M' N') = (M=M' & N=N')"; +Goalw [CONS_def] "(CONS M N=CONS M' N') = (M=M' & N=N')"; by (fast_tac (claset() addSEs [Scons_inject]) 1); qed "CONS_CONS_eq2"; @@ -409,7 +409,7 @@ AddSDs [inj_on_Abs_llist RS inj_onD, inj_Rep_llist RS injD, Leaf_inject]; -goalw LList.thy [LCons_def] "(LCons x xs=LCons y ys) = (x=y & xs=ys)"; +Goalw [LCons_def] "(LCons x xs=LCons y ys) = (x=y & xs=ys)"; by (Fast_tac 1); qed "LCons_LCons_eq"; @@ -449,12 +449,12 @@ (*** The functional "Lmap" ***) -goal LList.thy "Lmap f NIL = NIL"; +Goal "Lmap f NIL = NIL"; by (rtac (Lmap_def RS def_LList_corec RS trans) 1); by (Simp_tac 1); qed "Lmap_NIL"; -goal LList.thy "Lmap f (CONS M N) = CONS (f M) (Lmap f N)"; +Goal "Lmap f (CONS M N) = CONS (f M) (Lmap f N)"; by (rtac (Lmap_def RS def_LList_corec RS trans) 1); by (Simp_tac 1); qed "Lmap_CONS"; @@ -502,18 +502,18 @@ (*** Lappend -- its two arguments cause some complications! ***) -goalw LList.thy [Lappend_def] "Lappend NIL NIL = NIL"; +Goalw [Lappend_def] "Lappend NIL NIL = NIL"; by (rtac (LList_corec RS trans) 1); by (Simp_tac 1); qed "Lappend_NIL_NIL"; -goalw LList.thy [Lappend_def] +Goalw [Lappend_def] "Lappend NIL (CONS N N') = CONS N (Lappend NIL N')"; by (rtac (LList_corec RS trans) 1); by (Simp_tac 1); qed "Lappend_NIL_CONS"; -goalw LList.thy [Lappend_def] +Goalw [Lappend_def] "Lappend (CONS M M') N = CONS M (Lappend M' N)"; by (rtac (LList_corec RS trans) 1); by (Simp_tac 1); @@ -523,12 +523,12 @@ Lappend_CONS, LListD_Fun_CONS_I, range_eqI, image_eqI]; -goal LList.thy "!!M. M: llist(A) ==> Lappend NIL M = M"; +Goal "!!M. M: llist(A) ==> Lappend NIL M = M"; by (etac LList_fun_equalityI 1); by (ALLGOALS Asm_simp_tac); qed "Lappend_NIL"; -goal LList.thy "!!M. M: llist(A) ==> Lappend M NIL = M"; +Goal "!!M. M: llist(A) ==> Lappend M NIL = M"; by (etac LList_fun_equalityI 1); by (ALLGOALS Asm_simp_tac); qed "Lappend_NIL2"; @@ -539,7 +539,7 @@ (** Alternative type-checking proofs for Lappend **) (*weak co-induction: bisimulation and case analysis on both variables*) -goal LList.thy +Goal "!!M N. [| M: llist(A); N: llist(A) |] ==> Lappend M N: llist(A)"; by (res_inst_tac [("X", "UN u:llist(A). UN v: llist(A). {Lappend u v}")] llist_coinduct 1); @@ -552,7 +552,7 @@ qed "Lappend_type"; (*strong co-induction: bisimulation and case analysis on one variable*) -goal LList.thy +Goal "!!M N. [| M: llist(A); N: llist(A) |] ==> Lappend M N: llist(A)"; by (res_inst_tac [("X", "(%u. Lappend u N)``llist(A)")] llist_coinduct 1); by (etac imageI 1); @@ -569,11 +569,11 @@ Addsimps ([Abs_llist_inverse, Rep_llist_inverse, Rep_llist, rangeI, inj_Leaf, inv_f_f] @ llist.intrs); -goalw LList.thy [llist_case_def,LNil_def] "llist_case c d LNil = c"; +Goalw [llist_case_def,LNil_def] "llist_case c d LNil = c"; by (Simp_tac 1); qed "llist_case_LNil"; -goalw LList.thy [llist_case_def,LCons_def] +Goalw [llist_case_def,LCons_def] "llist_case c d (LCons M N) = d M N"; by (Simp_tac 1); qed "llist_case_LCons"; @@ -596,7 +596,7 @@ (** llist_corec: corecursion for 'a llist **) -goalw LList.thy [llist_corec_def,LNil_def,LCons_def] +Goalw [llist_corec_def,LNil_def,LCons_def] "llist_corec a f = sum_case (%u. LNil) \ \ (split(%z w. LCons z (llist_corec w f))) (f a)"; by (stac LList_corec 1); @@ -620,7 +620,7 @@ (*** Deriving llist_equalityI -- llist equality is a bisimulation ***) -goalw LList.thy [LListD_Fun_def] +Goalw [LListD_Fun_def] "!!r A. r <= (llist A) Times (llist A) ==> \ \ LListD_Fun (diag A) r <= (llist A) Times (llist A)"; by (stac llist_unfold 1); @@ -628,7 +628,7 @@ by (Fast_tac 1); qed "LListD_Fun_subset_Sigma_llist"; -goal LList.thy +Goal "prod_fun Rep_llist Rep_llist `` r <= \ \ (llist(range Leaf)) Times (llist(range Leaf))"; by (fast_tac (claset() delrules [image_subsetI] @@ -644,7 +644,7 @@ by (asm_simp_tac (simpset() addsimps [Abs_llist_inverse]) 1); qed "prod_fun_lemma"; -goal LList.thy +Goal "prod_fun Rep_llist Rep_llist `` range(%x. (x, x)) = \ \ diag(llist(range Leaf))"; by (rtac equalityI 1); @@ -654,7 +654,7 @@ qed "prod_fun_range_eq_diag"; (*Surprisingly hard to prove. Used with lfilter*) -goalw thy [llistD_Fun_def, prod_fun_def] +Goalw [llistD_Fun_def, prod_fun_def] "!!A B. A<=B ==> llistD_Fun A <= llistD_Fun B"; by Auto_tac; by (rtac image_eqI 1); @@ -682,7 +682,7 @@ qed "llist_equalityI"; (** Rules to prove the 2nd premise of llist_equalityI **) -goalw LList.thy [llistD_Fun_def,LNil_def] "(LNil,LNil) : llistD_Fun(r)"; +Goalw [llistD_Fun_def,LNil_def] "(LNil,LNil) : llistD_Fun(r)"; by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1); qed "llistD_Fun_LNil_I"; @@ -693,7 +693,7 @@ qed "llistD_Fun_LCons_I"; (*Utilise the "strong" part, i.e. gfp(f)*) -goalw LList.thy [llistD_Fun_def] +Goalw [llistD_Fun_def] "!!l. (l,l) : llistD_Fun(r Un range(%x.(x,x)))"; by (rtac (Rep_llist_inverse RS subst) 1); by (rtac prod_fun_imageI 1); @@ -728,12 +728,12 @@ (*** The functional "lmap" ***) -goal LList.thy "lmap f LNil = LNil"; +Goal "lmap f LNil = LNil"; by (rtac (lmap_def RS def_llist_corec RS trans) 1); by (Simp_tac 1); qed "lmap_LNil"; -goal LList.thy "lmap f (LCons M N) = LCons (f M) (lmap f N)"; +Goal "lmap f (LCons M N) = LCons (f M) (lmap f N)"; by (rtac (lmap_def RS def_llist_corec RS trans) 1); by (Simp_tac 1); qed "lmap_LCons"; @@ -743,12 +743,12 @@ (** Two easy results about lmap **) -goal LList.thy "lmap (f o g) l = lmap f (lmap g l)"; +Goal "lmap (f o g) l = lmap f (lmap g l)"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); qed "lmap_compose"; -goal LList.thy "lmap (%x. x) l = l"; +Goal "lmap (%x. x) l = l"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); qed "lmap_ident"; @@ -756,12 +756,12 @@ (*** iterates -- llist_fun_equalityI cannot be used! ***) -goal LList.thy "iterates f x = LCons x (iterates f (f x))"; +Goal "iterates f x = LCons x (iterates f (f x))"; by (rtac (iterates_def RS def_llist_corec RS trans) 1); by (Simp_tac 1); qed "iterates"; -goal LList.thy "lmap f (iterates f x) = iterates f (f x)"; +Goal "lmap f (iterates f x) = iterates f (f x)"; by (res_inst_tac [("r", "range(%u.(lmap f (iterates f u),iterates f (f u)))")] llist_equalityI 1); by (rtac rangeI 1); @@ -771,7 +771,7 @@ by (Simp_tac 1); qed "lmap_iterates"; -goal LList.thy "iterates f x = LCons x (lmap f (iterates f x))"; +Goal "iterates f x = LCons x (lmap f (iterates f x))"; by (stac lmap_iterates 1); by (rtac iterates 1); qed "iterates_lmap"; @@ -780,7 +780,7 @@ (** Two lemmas about natrec n x (%m.g), which is essentially (g^n)(x) **) -goal LList.thy +Goal "nat_rec (LCons b l) (%m. lmap(f)) n = \ \ LCons (nat_rec b (%m. f) n) (nat_rec l (%m. lmap(f)) n)"; by (nat_ind_tac "n" 1); @@ -821,18 +821,18 @@ (*** lappend -- its two arguments cause some complications! ***) -goalw LList.thy [lappend_def] "lappend LNil LNil = LNil"; +Goalw [lappend_def] "lappend LNil LNil = LNil"; by (rtac (llist_corec RS trans) 1); by (Simp_tac 1); qed "lappend_LNil_LNil"; -goalw LList.thy [lappend_def] +Goalw [lappend_def] "lappend LNil (LCons l l') = LCons l (lappend LNil l')"; by (rtac (llist_corec RS trans) 1); by (Simp_tac 1); qed "lappend_LNil_LCons"; -goalw LList.thy [lappend_def] +Goalw [lappend_def] "lappend (LCons l l') N = LCons l (lappend l' N)"; by (rtac (llist_corec RS trans) 1); by (Simp_tac 1); @@ -840,12 +840,12 @@ Addsimps [lappend_LNil_LNil, lappend_LNil_LCons, lappend_LCons]; -goal LList.thy "lappend LNil l = l"; +Goal "lappend LNil l = l"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); qed "lappend_LNil"; -goal LList.thy "lappend l LNil = l"; +Goal "lappend l LNil = l"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (ALLGOALS Simp_tac); qed "lappend_LNil2"; @@ -853,7 +853,7 @@ Addsimps [lappend_LNil, lappend_LNil2]; (*The infinite first argument blocks the second*) -goal LList.thy "lappend (iterates f x) N = iterates f x"; +Goal "lappend (iterates f x) N = iterates f x"; by (res_inst_tac [("r", "range(%u.(lappend (iterates f u) N,iterates f u))")] llist_equalityI 1); by (rtac rangeI 1); @@ -865,7 +865,7 @@ (** Two proofs that lmap distributes over lappend **) (*Long proof requiring case analysis on both both arguments*) -goal LList.thy "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"; +Goal "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"; by (res_inst_tac [("r", "UN n. range(%l.(lmap f (lappend l n),lappend (lmap f l) (lmap f n)))")] @@ -880,14 +880,14 @@ qed "lmap_lappend_distrib"; (*Shorter proof of theorem above using llist_equalityI as strong coinduction*) -goal LList.thy "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"; +Goal "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"; by (res_inst_tac [("l","l")] llist_fun_equalityI 1); by (Simp_tac 1); by (Simp_tac 1); qed "lmap_lappend_distrib"; (*Without strong coinduction, three case analyses might be needed*) -goal LList.thy "lappend (lappend l1 l2) l3 = lappend l1 (lappend l2 l3)"; +Goal "lappend (lappend l1 l2) l3 = lappend l1 (lappend l2 l3)"; by (res_inst_tac [("l","l1")] llist_fun_equalityI 1); by (Simp_tac 1); by (Simp_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Mutil.ML --- a/src/HOL/Induct/Mutil.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Mutil.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,7 +10,7 @@ (** The union of two disjoint tilings is a tiling **) -goal thy "!!t. t: tiling A ==> \ +Goal "!!t. t: tiling A ==> \ \ u: tiling A --> t <= Compl u --> t Un u : tiling A"; by (etac tiling.induct 1); by (Simp_tac 1); @@ -21,29 +21,29 @@ (*** Chess boards ***) -goalw thy [below_def] "(i: below k) = (i finite(evnodd X b) *) bind_thm("finite_evnodd", evnodd_subset RS finite_subset); -goalw thy [evnodd_def] "evnodd (A Un B) b = evnodd A b Un evnodd B b"; +Goalw [evnodd_def] "evnodd (A Un B) b = evnodd A b Un evnodd B b"; by (Blast_tac 1); qed "evnodd_Un"; -goalw thy [evnodd_def] "evnodd (A - B) b = evnodd A b - evnodd B b"; +Goalw [evnodd_def] "evnodd (A - B) b = evnodd A b - evnodd B b"; by (Blast_tac 1); qed "evnodd_Diff"; -goalw thy [evnodd_def] "evnodd {} b = {}"; +Goalw [evnodd_def] "evnodd {} b = {}"; by (Simp_tac 1); qed "evnodd_empty"; -goalw thy [evnodd_def] +Goalw [evnodd_def] "evnodd (insert (i,j) C) b = \ \ (if (i+j) mod 2 = b then insert (i,j) (evnodd C b) else evnodd C b)"; by (Simp_tac 1); @@ -101,7 +101,7 @@ (*** Dominoes ***) -goal thy "!!d. [| d:domino; b<2 |] ==> EX i j. evnodd d b = {(i,j)}"; +Goal "!!d. [| d:domino; b<2 |] ==> EX i j. evnodd d b = {(i,j)}"; by (eresolve_tac [domino.elim] 1); by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 2); by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 1); @@ -111,20 +111,20 @@ THEN Blast_tac 1)); qed "domino_singleton"; -goal thy "!!d. d:domino ==> finite d"; +Goal "!!d. d:domino ==> finite d"; by (blast_tac (claset() addSEs [domino.elim]) 1); qed "domino_finite"; (*** Tilings of dominoes ***) -goal thy "!!t. t:tiling domino ==> finite t"; +Goal "!!t. t:tiling domino ==> finite t"; by (eresolve_tac [tiling.induct] 1); by (rtac Finites.emptyI 1); by (blast_tac (claset() addSIs [finite_UnI] addIs [domino_finite]) 1); qed "tiling_domino_finite"; -goal thy "!!t. t: tiling domino ==> card(evnodd t 0) = card(evnodd t 1)"; +Goal "!!t. t: tiling domino ==> card(evnodd t 0) = card(evnodd t 1)"; by (eresolve_tac [tiling.induct] 1); by (simp_tac (simpset() addsimps [evnodd_def]) 1); by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1); @@ -140,7 +140,7 @@ (*Final argument is surprisingly complex. Note the use of small simpsets to avoid moving Sucs, etc.*) -goal thy "!!m n. [| t = below(Suc m + Suc m) Times below(Suc n + Suc n); \ +Goal "!!m n. [| t = below(Suc m + Suc m) Times below(Suc n + Suc n); \ \ t' = t - {(0,0)} - {(Suc(m+m), Suc(n+n))} \ \ |] ==> t' ~: tiling domino"; by (rtac notI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Perm.ML --- a/src/HOL/Induct/Perm.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Perm.ML Mon Jun 22 17:26:46 1998 +0200 @@ -13,7 +13,7 @@ open Perm; -goal Perm.thy "l <~~> l"; +Goal "l <~~> l"; by (list.induct_tac "l" 1); by (REPEAT (ares_tac perm.intrs 1)); qed "perm_refl"; @@ -22,23 +22,23 @@ (** Some examples of rule induction on permutations **) (*The form of the premise lets the induction bind xs and ys.*) -goal Perm.thy "!!xs. xs <~~> ys ==> xs=[] --> ys=[]"; +Goal "!!xs. xs <~~> ys ==> xs=[] --> ys=[]"; by (etac perm.induct 1); by (ALLGOALS Asm_simp_tac); qed "perm_Nil_lemma"; (*A more general version is actually easier to understand!*) -goal Perm.thy "!!xs. xs <~~> ys ==> length(xs) = length(ys)"; +Goal "!!xs. xs <~~> ys ==> length(xs) = length(ys)"; by (etac perm.induct 1); by (ALLGOALS Asm_simp_tac); qed "perm_length"; -goal Perm.thy "!!xs. xs <~~> ys ==> ys <~~> xs"; +Goal "!!xs. xs <~~> ys ==> ys <~~> xs"; by (etac perm.induct 1); by (REPEAT (ares_tac perm.intrs 1)); qed "perm_sym"; -goal Perm.thy "!!xs. [| xs <~~> ys |] ==> x mem xs --> x mem ys"; +Goal "!!xs. [| xs <~~> ys |] ==> x mem xs --> x mem ys"; by (etac perm.induct 1); by (Fast_tac 4); by (ALLGOALS Asm_simp_tac); @@ -49,7 +49,7 @@ (** Ways of making new permutations **) (*We can insert the head anywhere in the list*) -goal Perm.thy "a # xs @ ys <~~> xs @ a # ys"; +Goal "a # xs @ ys <~~> xs @ a # ys"; by (list.induct_tac "xs" 1); by (simp_tac (simpset() addsimps [perm_refl]) 1); by (Simp_tac 1); @@ -62,7 +62,7 @@ by (rtac perm.Cons 1); *) -goal Perm.thy "xs@ys <~~> ys@xs"; +Goal "xs@ys <~~> ys@xs"; by (list.induct_tac "xs" 1); by (simp_tac (simpset() addsimps [perm_refl]) 1); by (Simp_tac 1); @@ -70,13 +70,13 @@ qed "perm_append_swap"; -goal Perm.thy "a # xs <~~> xs @ [a]"; +Goal "a # xs <~~> xs @ [a]"; by (rtac perm.trans 1); by (rtac perm_append_swap 2); by (simp_tac (simpset() addsimps [perm_refl]) 1); qed "perm_append_single"; -goal Perm.thy "rev xs <~~> xs"; +Goal "rev xs <~~> xs"; by (list.induct_tac "xs" 1); by (simp_tac (simpset() addsimps [perm_refl]) 1); by (Simp_tac 1); @@ -84,13 +84,13 @@ by (etac perm.Cons 1); qed "perm_rev"; -goal Perm.thy "!!xs. xs <~~> ys ==> l@xs <~~> l@ys"; +Goal "!!xs. xs <~~> ys ==> l@xs <~~> l@ys"; by (list.induct_tac "l" 1); by (Simp_tac 1); by (asm_simp_tac (simpset() addsimps [perm.Cons]) 1); qed "perm_append1"; -goal Perm.thy "!!xs. xs <~~> ys ==> xs@l <~~> ys@l"; +Goal "!!xs. xs <~~> ys ==> xs@l <~~> ys@l"; by (rtac (perm_append_swap RS perm.trans) 1); by (etac (perm_append1 RS perm.trans) 1); by (rtac perm_append_swap 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/PropLog.ML --- a/src/HOL/Induct/PropLog.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/PropLog.ML Mon Jun 22 17:26:46 1998 +0200 @@ -11,13 +11,13 @@ open PropLog; (** Monotonicity **) -goalw PropLog.thy thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)"; +Goalw thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)"; by (rtac lfp_mono 1); by (REPEAT (ares_tac basic_monos 1)); qed "thms_mono"; (*Rule is called I for Identity Combinator, not for Introduction*) -goal PropLog.thy "H |- p->p"; +Goal "H |- p->p"; by (best_tac (claset() addIs [thms.K,thms.S,thms.MP]) 1); qed "thms_I"; @@ -34,12 +34,12 @@ val weaken_left_Un1 = Un_upper1 RS weaken_left; val weaken_left_Un2 = Un_upper2 RS weaken_left; -goal PropLog.thy "!!H. H |- q ==> H |- p->q"; +Goal "!!H. H |- q ==> H |- p->q"; by (fast_tac (claset() addIs [thms.K,thms.MP]) 1); qed "weaken_right"; (*The deduction theorem*) -goal PropLog.thy "!!H. insert p H |- q ==> H |- p->q"; +Goal "!!H. insert p H |- q ==> H |- p->q"; by (etac thms.induct 1); by (ALLGOALS (fast_tac (claset() addIs [thms_I, thms.H, thms.K, thms.S, thms.DN, @@ -60,22 +60,22 @@ (** The function eval **) -goalw PropLog.thy [eval_def] "tt[false] = False"; +Goalw [eval_def] "tt[false] = False"; by (Simp_tac 1); qed "eval_false"; -goalw PropLog.thy [eval_def] "tt[#v] = (v:tt)"; +Goalw [eval_def] "tt[#v] = (v:tt)"; by (Simp_tac 1); qed "eval_var"; -goalw PropLog.thy [eval_def] "tt[p->q] = (tt[p]-->tt[q])"; +Goalw [eval_def] "tt[p->q] = (tt[p]-->tt[q])"; by (Simp_tac 1); qed "eval_imp"; Addsimps [eval_false, eval_var, eval_imp]; (*Soundness of the rules wrt truth-table semantics*) -goalw PropLog.thy [sat_def] "!!H. H |- p ==> H |= p"; +Goalw [sat_def] "!!H. H |- p ==> H |= p"; by (etac thms.induct 1); by (fast_tac (claset() addSDs [eval_imp RS iffD1 RS mp]) 5); by (ALLGOALS Asm_simp_tac); @@ -83,7 +83,7 @@ (*** Towards the completeness proof ***) -goal PropLog.thy "!!H. H |- p->false ==> H |- p->q"; +Goal "!!H. H |- p->false ==> H |- p->q"; by (rtac deduction 1); by (etac (weaken_left_insert RS thms_notE) 1); by (rtac thms.H 1); @@ -100,7 +100,7 @@ qed "imp_false"; (*This formulation is required for strong induction hypotheses*) -goal PropLog.thy "hyps p tt |- (if tt[p] then p else p->false)"; +Goal "hyps p tt |- (if tt[p] then p else p->false)"; by (rtac (split_if RS iffD2) 1); by (PropLog.pl.induct_tac "p" 1); by (ALLGOALS (simp_tac (simpset() addsimps [thms_I, thms.H]))); @@ -122,7 +122,7 @@ AddIs [thms.H, thms.H RS thms.MP]; (*The excluded middle in the form of an elimination rule*) -goal PropLog.thy "H |- (p->q) -> ((p->false)->q) -> q"; +Goal "H |- (p->q) -> ((p->false)->q) -> q"; by (rtac (deduction RS deduction) 1); by (rtac (thms.DN RS thms.MP) 1); by (ALLGOALS (best_tac (claset() addSIs prems))); @@ -139,7 +139,7 @@ (*For the case hyps(p,t)-insert(#v,Y) |- p; we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) -goal PropLog.thy "hyps p (t-{v}) <= insert (#v->false) ((hyps p t)-{#v})"; +Goal "hyps p (t-{v}) <= insert (#v->false) ((hyps p t)-{#v})"; by (PropLog.pl.induct_tac "p" 1); by (ALLGOALS Simp_tac); by (Fast_tac 1); @@ -147,7 +147,7 @@ (*For the case hyps(p,t)-insert(#v -> false,Y) |- p; we also have hyps(p,t)-{#v->false} <= hyps(p, insert(v,t)) *) -goal PropLog.thy "hyps p (insert v t) <= insert (#v) (hyps p t-{#v->false})"; +Goal "hyps p (insert v t) <= insert (#v) (hyps p t-{#v->false})"; by (PropLog.pl.induct_tac "p" 1); by (ALLGOALS Simp_tac); by (Fast_tac 1); @@ -163,12 +163,12 @@ by (Fast_tac 1); qed "insert_Diff_subset2"; -goal PropLog.thy "finite(hyps p t)"; +Goal "finite(hyps p t)"; by (induct_tac "p" 1); by (ALLGOALS Asm_simp_tac); qed "hyps_finite"; -goal PropLog.thy "hyps p t <= (UN v. {#v, #v->false})"; +Goal "hyps p t <= (UN v. {#v, #v->false})"; by (PropLog.pl.induct_tac "p" 1); by (ALLGOALS Simp_tac); by (Blast_tac 1); @@ -206,12 +206,12 @@ qed "completeness_0"; (*A semantic analogue of the Deduction Theorem*) -goalw PropLog.thy [sat_def] "!!p H. insert p H |= q ==> H |= p->q"; +Goalw [sat_def] "!!p H. insert p H |= q ==> H |= p->q"; by (Simp_tac 1); by (Fast_tac 1); qed "sat_imp"; -goal PropLog.thy "!!H. finite H ==> !p. H |= p --> H |- p"; +Goal "!!H. finite H ==> !p. H |= p --> H |- p"; by (etac finite_induct 1); by (safe_tac ((claset_of Fun.thy) addSIs [completeness_0])); by (rtac (weaken_left_insert RS thms.MP) 1); @@ -221,6 +221,6 @@ val completeness = completeness_lemma RS spec RS mp; -goal PropLog.thy "!!H. finite H ==> (H |- p) = (H |= p)"; +Goal "!!H. finite H ==> (H |- p) = (H |= p)"; by (fast_tac (claset() addSEs [soundness, completeness]) 1); qed "syntax_iff_semantics"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/SList.ML --- a/src/HOL/Induct/SList.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/SList.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,7 +10,7 @@ val list_con_defs = [NIL_def, CONS_def]; -goal SList.thy "list(A) = {Numb(0)} <+> (A <*> list(A))"; +Goal "list(A) = {Numb(0)} <+> (A <*> list(A))"; let val rew = rewrite_rule list_con_defs in by (fast_tac (claset() addSIs (equalityI :: map rew list.intrs) addEs [rew list.elim]) 1) @@ -18,13 +18,13 @@ qed "list_unfold"; (*This justifies using list in other recursive type definitions*) -goalw SList.thy list.defs "!!A B. A<=B ==> list(A) <= list(B)"; +Goalw list.defs "!!A B. A<=B ==> list(A) <= list(B)"; by (rtac lfp_mono 1); by (REPEAT (ares_tac basic_monos 1)); qed "list_mono"; (*Type checking -- list creates well-founded sets*) -goalw SList.thy (list_con_defs @ list.defs) "list(sexp) <= sexp"; +Goalw (list_con_defs @ list.defs) "list(sexp) <= sexp"; by (rtac lfp_lowerbound 1); by (fast_tac (claset() addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1); qed "list_sexp"; @@ -49,19 +49,19 @@ (*** Isomorphisms ***) -goal SList.thy "inj(Rep_list)"; +Goal "inj(Rep_list)"; by (rtac inj_inverseI 1); by (rtac Rep_list_inverse 1); qed "inj_Rep_list"; -goal SList.thy "inj_on Abs_list (list(range Leaf))"; +Goal "inj_on Abs_list (list(range Leaf))"; by (rtac inj_on_inverseI 1); by (etac Abs_list_inverse 1); qed "inj_on_Abs_list"; (** Distinctness of constructors **) -goalw SList.thy list_con_defs "CONS M N ~= NIL"; +Goalw list_con_defs "CONS M N ~= NIL"; by (rtac In1_not_In0 1); qed "CONS_not_NIL"; bind_thm ("NIL_not_CONS", (CONS_not_NIL RS not_sym)); @@ -69,7 +69,7 @@ bind_thm ("CONS_neq_NIL", (CONS_not_NIL RS notE)); val NIL_neq_CONS = sym RS CONS_neq_NIL; -goalw SList.thy [Nil_def,Cons_def] "x # xs ~= Nil"; +Goalw [Nil_def,Cons_def] "x # xs ~= Nil"; by (rtac (CONS_not_NIL RS (inj_on_Abs_list RS inj_on_contraD)) 1); by (REPEAT (resolve_tac (list.intrs @ [rangeI, Rep_list]) 1)); qed "Cons_not_Nil"; @@ -78,7 +78,7 @@ (** Injectiveness of CONS and Cons **) -goalw SList.thy [CONS_def] "(CONS K M=CONS L N) = (K=L & M=N)"; +Goalw [CONS_def] "(CONS K M=CONS L N) = (K=L & M=N)"; by (fast_tac (claset() addSEs [Scons_inject, make_elim In1_inject]) 1); qed "CONS_CONS_eq"; @@ -90,7 +90,7 @@ AddSDs [inj_on_Abs_list RS inj_onD, inj_Rep_list RS injD, Leaf_inject]; -goalw SList.thy [Cons_def] "(x#xs=y#ys) = (x=y & xs=ys)"; +Goalw [Cons_def] "(x#xs=y#ys) = (x=y & xs=ys)"; by (Fast_tac 1); qed "Cons_Cons_eq"; bind_thm ("Cons_inject2", (Cons_Cons_eq RS iffD1 RS conjE)); @@ -113,18 +113,18 @@ AddIffs [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq]; -goal SList.thy "!!N. N: list(A) ==> !M. N ~= CONS M N"; +Goal "!!N. N: list(A) ==> !M. N ~= CONS M N"; by (etac list.induct 1); by (ALLGOALS Asm_simp_tac); qed "not_CONS_self"; -goal SList.thy "!x. l ~= x#l"; +Goal "!x. l ~= x#l"; by (list_ind_tac "l" 1); by (ALLGOALS Asm_simp_tac); qed "not_Cons_self2"; -goal SList.thy "(xs ~= []) = (? y ys. xs = y#ys)"; +Goal "(xs ~= []) = (? y ys. xs = y#ys)"; by (list_ind_tac "xs" 1); by (Simp_tac 1); by (Asm_simp_tac 1); @@ -133,11 +133,11 @@ (** Conversion rules for List_case: case analysis operator **) -goalw SList.thy [List_case_def,NIL_def] "List_case c h NIL = c"; +Goalw [List_case_def,NIL_def] "List_case c h NIL = c"; by (rtac Case_In0 1); qed "List_case_NIL"; -goalw SList.thy [List_case_def,CONS_def] "List_case c h (CONS M N) = h M N"; +Goalw [List_case_def,CONS_def] "List_case c h (CONS M N) = h M N"; by (Simp_tac 1); qed "List_case_CONS"; @@ -149,7 +149,7 @@ (* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not hold if pred_sexp^+ were changed to pred_sexp. *) -goal SList.thy +Goal "(%M. List_rec M c d) = wfrec (trancl pred_sexp) \ \ (%g. List_case c (%x y. d x y (g y)))"; by (simp_tac (HOL_ss addsimps [List_rec_def]) 1); @@ -164,12 +164,12 @@ (** pred_sexp lemmas **) -goalw SList.thy [CONS_def,In1_def] +Goalw [CONS_def,In1_def] "!!M. [| M: sexp; N: sexp |] ==> (M, CONS M N) : pred_sexp^+"; by (Asm_simp_tac 1); qed "pred_sexp_CONS_I1"; -goalw SList.thy [CONS_def,In1_def] +Goalw [CONS_def,In1_def] "!!M. [| M: sexp; N: sexp |] ==> (N, CONS M N) : pred_sexp^+"; by (Asm_simp_tac 1); qed "pred_sexp_CONS_I2"; @@ -186,12 +186,12 @@ (** Conversion rules for List_rec **) -goal SList.thy "List_rec NIL c h = c"; +Goal "List_rec NIL c h = c"; by (rtac (List_rec_unfold RS trans) 1); by (Simp_tac 1); qed "List_rec_NIL"; -goal SList.thy "!!M. [| M: sexp; N: sexp |] ==> \ +Goal "!!M. [| M: sexp; N: sexp |] ==> \ \ List_rec (CONS M N) c h = h M N (List_rec N c h)"; by (rtac (List_rec_unfold RS trans) 1); by (asm_simp_tac (simpset() addsimps [pred_sexp_CONS_I2]) 1); @@ -237,21 +237,21 @@ (** Generalized map functionals **) -goalw SList.thy [Rep_map_def] "Rep_map f Nil = NIL"; +Goalw [Rep_map_def] "Rep_map f Nil = NIL"; by (rtac list_rec_Nil 1); qed "Rep_map_Nil"; -goalw SList.thy [Rep_map_def] +Goalw [Rep_map_def] "Rep_map f (x#xs) = CONS (f x) (Rep_map f xs)"; by (rtac list_rec_Cons 1); qed "Rep_map_Cons"; -goalw SList.thy [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map f xs: list(A)"; +Goalw [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map f xs: list(A)"; by (rtac list_induct2 1); by (ALLGOALS Asm_simp_tac); qed "Rep_map_type"; -goalw SList.thy [Abs_map_def] "Abs_map g NIL = Nil"; +Goalw [Abs_map_def] "Abs_map g NIL = Nil"; by (rtac List_rec_NIL 1); qed "Abs_map_NIL"; @@ -303,24 +303,24 @@ (** @ - append **) -goal SList.thy "(xs@ys)@zs = xs@(ys@zs)"; +Goal "(xs@ys)@zs = xs@(ys@zs)"; by (list_ind_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "append_assoc2"; -goal SList.thy "xs @ [] = xs"; +Goal "xs @ [] = xs"; by (list_ind_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "append_Nil4"; (** mem **) -goal SList.thy "x mem (xs@ys) = (x mem xs | x mem ys)"; +Goal "x mem (xs@ys) = (x mem xs | x mem ys)"; by (list_ind_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "mem_append2"; -goal SList.thy "x mem [x:xs. P(x)] = (x mem xs & P(x))"; +Goal "x mem [x:xs. P(x)] = (x mem xs & P(x))"; by (list_ind_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "mem_filter2"; @@ -341,7 +341,7 @@ (** list_case **) -goal SList.thy +Goal "P(list_case a f xs) = ((xs=[] --> P(a)) & \ \ (!y ys. xs=y#ys --> P(f y ys)))"; by (list_ind_tac "xs" 1); @@ -351,22 +351,22 @@ (** Additional mapping lemmas **) -goal SList.thy "map (%x. x) xs = xs"; +Goal "map (%x. x) xs = xs"; by (list_ind_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "map_ident2"; -goal SList.thy "map f (xs@ys) = map f xs @ map f ys"; +Goal "map f (xs@ys) = map f xs @ map f ys"; by (list_ind_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "map_append2"; -goalw SList.thy [o_def] "map (f o g) xs = map f (map g xs)"; +Goalw [o_def] "map (f o g) xs = map f (map g xs)"; by (list_ind_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "map_compose2"; -goal SList.thy "!!f. (!!x. f(x): sexp) ==> \ +Goal "!!f. (!!x. f(x): sexp) ==> \ \ Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs"; by (list_ind_tac "xs" 1); by (ALLGOALS (asm_simp_tac(simpset() addsimps diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Simult.ML --- a/src/HOL/Induct/Simult.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Simult.ML Mon Jun 22 17:26:46 1998 +0200 @@ -15,7 +15,7 @@ (*** Monotonicity and unfolding of the function ***) -goal Simult.thy "mono(%Z. A <*> Part Z In1 \ +Goal "mono(%Z. A <*> Part Z In1 \ \ <+> ({Numb(0)} <+> Part Z In0 <*> Part Z In1))"; by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono, Part_mono] 1)); @@ -23,11 +23,11 @@ val TF_unfold = TF_fun_mono RS (TF_def RS def_lfp_Tarski); -goalw Simult.thy [TF_def] "!!A B. A<=B ==> TF(A) <= TF(B)"; +Goalw [TF_def] "!!A B. A<=B ==> TF(A) <= TF(B)"; by (REPEAT (ares_tac [lfp_mono, subset_refl, usum_mono, uprod_mono] 1)); qed "TF_mono"; -goalw Simult.thy [TF_def] "TF(sexp) <= sexp"; +Goalw [TF_def] "TF(sexp) <= sexp"; by (rtac lfp_lowerbound 1); by (blast_tac (claset() addIs sexp.intrs@[sexp_In0I, sexp_In1I] addSEs [PartE]) 1); @@ -55,7 +55,7 @@ qed "TF_induct"; (*This lemma replaces a use of subgoal_tac to prove tree_forest_induct*) -goalw Simult.thy [Part_def] +Goalw [Part_def] "!!A. ! M: TF(A). (M: Part (TF A) In0 --> P(M)) & \ \ (M: Part (TF A) In1 --> Q(M)) \ \ ==> (! M: Part (TF A) In0. P(M)) & (! M: Part (TF A) In1. Q(M))"; @@ -104,22 +104,22 @@ (*** Isomorphisms ***) -goal Simult.thy "inj(Rep_Tree)"; +Goal "inj(Rep_Tree)"; by (rtac inj_inverseI 1); by (rtac Rep_Tree_inverse 1); qed "inj_Rep_Tree"; -goal Simult.thy "inj_on Abs_Tree (Part (TF(range Leaf)) In0)"; +Goal "inj_on Abs_Tree (Part (TF(range Leaf)) In0)"; by (rtac inj_on_inverseI 1); by (etac Abs_Tree_inverse 1); qed "inj_on_Abs_Tree"; -goal Simult.thy "inj(Rep_Forest)"; +Goal "inj(Rep_Forest)"; by (rtac inj_inverseI 1); by (rtac Rep_Forest_inverse 1); qed "inj_Rep_Forest"; -goal Simult.thy "inj_on Abs_Forest (Part (TF(range Leaf)) In1)"; +Goal "inj_on Abs_Forest (Part (TF(range Leaf)) In1)"; by (rtac inj_on_inverseI 1); by (etac Abs_Forest_inverse 1); qed "inj_on_Abs_Forest"; @@ -134,17 +134,17 @@ AddIs [TF_I, uprodI, usum_In0I, usum_In1I]; AddSEs [Scons_inject]; -goalw Simult.thy TF_Rep_defs +Goalw TF_Rep_defs "!!A. [| a: A; M: Part (TF A) In1 |] ==> TCONS a M : Part (TF A) In0"; by (Blast_tac 1); qed "TCONS_I"; (* FNIL is a TF(A) -- this also justifies the type definition*) -goalw Simult.thy TF_Rep_defs "FNIL: Part (TF A) In1"; +Goalw TF_Rep_defs "FNIL: Part (TF A) In1"; by (Blast_tac 1); qed "FNIL_I"; -goalw Simult.thy TF_Rep_defs +Goalw TF_Rep_defs "!!A. [| M: Part (TF A) In0; N: Part (TF A) In1 |] ==> \ \ FCONS M N : Part (TF A) In1"; by (Blast_tac 1); @@ -152,19 +152,19 @@ (** Injectiveness of TCONS and FCONS **) -goalw Simult.thy TF_Rep_defs "(TCONS K M=TCONS L N) = (K=L & M=N)"; +Goalw TF_Rep_defs "(TCONS K M=TCONS L N) = (K=L & M=N)"; by (Blast_tac 1); qed "TCONS_TCONS_eq"; bind_thm ("TCONS_inject", (TCONS_TCONS_eq RS iffD1 RS conjE)); -goalw Simult.thy TF_Rep_defs "(FCONS K M=FCONS L N) = (K=L & M=N)"; +Goalw TF_Rep_defs "(FCONS K M=FCONS L N) = (K=L & M=N)"; by (Blast_tac 1); qed "FCONS_FCONS_eq"; bind_thm ("FCONS_inject", (FCONS_FCONS_eq RS iffD1 RS conjE)); (** Distinctness of TCONS, FNIL and FCONS **) -goalw Simult.thy TF_Rep_defs "TCONS M N ~= FNIL"; +Goalw TF_Rep_defs "TCONS M N ~= FNIL"; by (Blast_tac 1); qed "TCONS_not_FNIL"; bind_thm ("FNIL_not_TCONS", (TCONS_not_FNIL RS not_sym)); @@ -172,7 +172,7 @@ bind_thm ("TCONS_neq_FNIL", (TCONS_not_FNIL RS notE)); val FNIL_neq_TCONS = sym RS TCONS_neq_FNIL; -goalw Simult.thy TF_Rep_defs "FCONS M N ~= FNIL"; +Goalw TF_Rep_defs "FCONS M N ~= FNIL"; by (Blast_tac 1); qed "FCONS_not_FNIL"; bind_thm ("FNIL_not_FCONS", (FCONS_not_FNIL RS not_sym)); @@ -180,7 +180,7 @@ bind_thm ("FCONS_neq_FNIL", (FCONS_not_FNIL RS notE)); val FNIL_neq_FCONS = sym RS FCONS_neq_FNIL; -goalw Simult.thy TF_Rep_defs "TCONS M N ~= FCONS K L"; +Goalw TF_Rep_defs "TCONS M N ~= FCONS K L"; by (Blast_tac 1); qed "TCONS_not_FCONS"; bind_thm ("FCONS_not_TCONS", (TCONS_not_FCONS RS not_sym)); @@ -204,12 +204,12 @@ inj_Rep_Tree RS injD, inj_Rep_Forest RS injD, Leaf_inject]; -goalw Simult.thy [Tcons_def] "(Tcons x xs=Tcons y ys) = (x=y & xs=ys)"; +Goalw [Tcons_def] "(Tcons x xs=Tcons y ys) = (x=y & xs=ys)"; by (Blast_tac 1); qed "Tcons_Tcons_eq"; bind_thm ("Tcons_inject", (Tcons_Tcons_eq RS iffD1 RS conjE)); -goalw Simult.thy [Fcons_def,Fnil_def] "Fcons x xs ~= Fnil"; +Goalw [Fcons_def,Fnil_def] "Fcons x xs ~= Fnil"; by (Blast_tac 1); qed "Fcons_not_Fnil"; @@ -219,7 +219,7 @@ (** Injectiveness of Fcons **) -goalw Simult.thy [Fcons_def] "(Fcons x xs=Fcons y ys) = (x=y & xs=ys)"; +Goalw [Fcons_def] "(Fcons x xs=Fcons y ys) = (x=y & xs=ys)"; by (Blast_tac 1); qed "Fcons_Fcons_eq"; bind_thm ("Fcons_inject", Fcons_Fcons_eq RS iffD1 RS conjE); @@ -227,7 +227,7 @@ (*** TF_rec -- by wf recursion on pred_sexp ***) -goal Simult.thy +Goal "(%M. TF_rec M b c d) = wfrec (trancl pred_sexp) \ \ (%g. Case (Split(%x y. b x y (g y))) \ \ (List_case c (%x y. d x y (g x) (g y))))"; @@ -243,7 +243,7 @@ (** conversion rules for TF_rec **) -goalw Simult.thy [TCONS_def] +Goalw [TCONS_def] "!!M N. [| M: sexp; N: sexp |] ==> \ \ TF_rec (TCONS M N) b c d = b M N (TF_rec N b c d)"; by (rtac (TF_rec_unfold RS trans) 1); @@ -251,12 +251,12 @@ by (asm_simp_tac (simpset() addsimps [In0_def]) 1); qed "TF_rec_TCONS"; -goalw Simult.thy [FNIL_def] "TF_rec FNIL b c d = c"; +Goalw [FNIL_def] "TF_rec FNIL b c d = c"; by (rtac (TF_rec_unfold RS trans) 1); by (simp_tac (HOL_ss addsimps [Case_In1, List_case_NIL]) 1); qed "TF_rec_FNIL"; -goalw Simult.thy [FCONS_def] +Goalw [FCONS_def] "!!M N. [| M: sexp; N: sexp |] ==> \ \ TF_rec (FCONS M N) b c d = d M N (TF_rec M b c d) (TF_rec N b c d)"; by (rtac (TF_rec_unfold RS trans) 1); @@ -282,16 +282,16 @@ inj_Leaf, inv_f_f, sexp.LeafI, range_eqI, Rep_Tree_in_sexp, Rep_Forest_in_sexp]; -goalw Simult.thy [tree_rec_def, forest_rec_def, Tcons_def] +Goalw [tree_rec_def, forest_rec_def, Tcons_def] "tree_rec (Tcons a tf) b c d = b a tf (forest_rec tf b c d)"; by (simp_tac tf_rec_ss 1); qed "tree_rec_Tcons"; -goalw Simult.thy [forest_rec_def, Fnil_def] "forest_rec Fnil b c d = c"; +Goalw [forest_rec_def, Fnil_def] "forest_rec Fnil b c d = c"; by (simp_tac tf_rec_ss 1); qed "forest_rec_Fnil"; -goalw Simult.thy [tree_rec_def, forest_rec_def, Fcons_def] +Goalw [tree_rec_def, forest_rec_def, Fcons_def] "forest_rec (Fcons t tf) b c d = \ \ d t tf (tree_rec t b c d) (forest_rec tf b c d)"; by (simp_tac tf_rec_ss 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Induct/Term.ML --- a/src/HOL/Induct/Term.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Induct/Term.ML Mon Jun 22 17:26:46 1998 +0200 @@ -11,19 +11,19 @@ (*** Monotonicity and unfolding of the function ***) -goal Term.thy "term(A) = A <*> list(term(A))"; +Goal "term(A) = A <*> list(term(A))"; by (fast_tac (claset() addSIs term.intrs addEs [term.elim]) 1); qed "term_unfold"; (*This justifies using term in other recursive type definitions*) -goalw Term.thy term.defs "!!A B. A<=B ==> term(A) <= term(B)"; +Goalw term.defs "!!A B. A<=B ==> term(A) <= term(B)"; by (REPEAT (ares_tac ([lfp_mono, list_mono] @ basic_monos) 1)); qed "term_mono"; (** Type checking -- term creates well-founded sets **) -goalw Term.thy term.defs "term(sexp) <= sexp"; +Goalw term.defs "term(sexp) <= sexp"; by (rtac lfp_lowerbound 1); by (fast_tac (claset() addIs [sexp.SconsI, list_sexp RS subsetD]) 1); qed "term_sexp"; @@ -110,7 +110,7 @@ (*** Term_rec -- by wf recursion on pred_sexp ***) -goal Term.thy +Goal "(%M. Term_rec M d) = wfrec (trancl pred_sexp) \ \ (%g. Split(%x y. d x y (Abs_map g y)))"; by (simp_tac (HOL_ss addsimps [Term_rec_def]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/Bin.ML --- a/src/HOL/Integ/Bin.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/Bin.ML Mon Jun 22 17:26:46 1998 +0200 @@ -138,7 +138,7 @@ integ_of_bin_norm_Bcons]; val bin_simps = [iob_Plus,iob_Minus,iob_Bcons]; -goal Bin.thy +Goal "! w. integ_of_bin(bin_add v w) = integ_of_bin v + integ_of_bin w"; by (bin.induct_tac "v" 1); by (simp_tac (simpset() addsimps bin_add_simps) 1); @@ -153,7 +153,7 @@ by (asm_simp_tac (simpset() addsimps (bin_add_simps @ zadd_ac)) 1); val integ_of_bin_add_lemma = result(); -goal Bin.thy "integ_of_bin(bin_add v w) = integ_of_bin v + integ_of_bin w"; +Goal "integ_of_bin(bin_add v w) = integ_of_bin v + integ_of_bin w"; by (cut_facts_tac [integ_of_bin_add_lemma] 1); by (Fast_tac 1); qed "integ_of_bin_add"; @@ -161,7 +161,7 @@ val bin_mult_simps = [integ_of_bin_minus, integ_of_bin_add, integ_of_bin_norm_Bcons]; -goal Bin.thy "integ_of_bin(bin_mult v w) = integ_of_bin v * integ_of_bin w"; +Goal "integ_of_bin(bin_mult v w) = integ_of_bin v * integ_of_bin w"; by (bin.induct_tac "v" 1); by (simp_tac (simpset() addsimps bin_mult_simps) 1); by (simp_tac (simpset() addsimps bin_mult_simps) 1); @@ -195,7 +195,7 @@ Addsimps [zadd_assoc]; -goal Bin.thy +Goal "(integ_of_bin x = integ_of_bin y) \ \ = (integ_of_bin (bin_add x (bin_minus y)) = $# 0)"; by (simp_tac (HOL_ss addsimps @@ -209,16 +209,16 @@ zadd_zminus_inverse2]) 1); val iob_eq_eq_diff_0 = result(); -goal Bin.thy "(integ_of_bin PlusSign = $# 0) = True"; +Goal "(integ_of_bin PlusSign = $# 0) = True"; by (stac iob_Plus 1); by (Simp_tac 1); val iob_Plus_eq_0 = result(); -goal Bin.thy "(integ_of_bin MinusSign = $# 0) = False"; +Goal "(integ_of_bin MinusSign = $# 0) = False"; by (stac iob_Minus 1); by (Simp_tac 1); val iob_Minus_not_eq_0 = result(); -goal Bin.thy +Goal "(integ_of_bin (Bcons w x) = $# 0) = (~x & integ_of_bin w = $# 0)"; by (stac iob_Bcons 1); by (case_tac "x" 1); @@ -233,21 +233,21 @@ by (Asm_full_simp_tac 1); val iob_Bcons_eq_0 = result(); -goalw Bin.thy [zless_def,zdiff_def] +Goalw [zless_def,zdiff_def] "integ_of_bin x < integ_of_bin y \ \ = (integ_of_bin (bin_add x (bin_minus y)) < $# 0)"; by (Simp_tac 1); val iob_less_eq_diff_lt_0 = result(); -goal Bin.thy "(integ_of_bin PlusSign < $# 0) = False"; +Goal "(integ_of_bin PlusSign < $# 0) = False"; by (stac iob_Plus 1); by (Simp_tac 1); val iob_Plus_not_lt_0 = result(); -goal Bin.thy "(integ_of_bin MinusSign < $# 0) = True"; +Goal "(integ_of_bin MinusSign < $# 0) = True"; by (stac iob_Minus 1); by (Simp_tac 1); val iob_Minus_lt_0 = result(); -goal Bin.thy +Goal "(integ_of_bin (Bcons w x) < $# 0) = (integ_of_bin w < $# 0)"; by (stac iob_Bcons 1); by (case_tac "x" 1); @@ -262,7 +262,7 @@ by (Simp_tac 1); val iob_Bcons_lt_0 = result(); -goal Bin.thy +Goal "integ_of_bin x <= integ_of_bin y \ \ = ( integ_of_bin (bin_add x (bin_minus y)) < $# 0 \ \ | integ_of_bin (bin_add x (bin_minus y)) = $# 0)"; @@ -283,66 +283,66 @@ (*** Examples of performing binary arithmetic by simplification ***) -goal Bin.thy "#13 + #19 = #32"; +Goal "#13 + #19 = #32"; by (Simp_tac 1); result(); -goal Bin.thy "#1234 + #5678 = #6912"; +Goal "#1234 + #5678 = #6912"; by (Simp_tac 1); result(); -goal Bin.thy "#1359 + #~2468 = #~1109"; +Goal "#1359 + #~2468 = #~1109"; by (Simp_tac 1); result(); -goal Bin.thy "#93746 + #~46375 = #47371"; +Goal "#93746 + #~46375 = #47371"; by (Simp_tac 1); result(); -goal Bin.thy "$~ #65745 = #~65745"; +Goal "$~ #65745 = #~65745"; by (Simp_tac 1); result(); -goal Bin.thy "$~ #~54321 = #54321"; +Goal "$~ #~54321 = #54321"; by (Simp_tac 1); result(); -goal Bin.thy "#13 * #19 = #247"; +Goal "#13 * #19 = #247"; by (Simp_tac 1); result(); -goal Bin.thy "#~84 * #51 = #~4284"; +Goal "#~84 * #51 = #~4284"; by (Simp_tac 1); result(); -goal Bin.thy "#255 * #255 = #65025"; +Goal "#255 * #255 = #65025"; by (Simp_tac 1); result(); -goal Bin.thy "#1359 * #~2468 = #~3354012"; +Goal "#1359 * #~2468 = #~3354012"; by (Simp_tac 1); result(); -goal Bin.thy "#89 * #10 ~= #889"; +Goal "#89 * #10 ~= #889"; by (Simp_tac 1); result(); -goal Bin.thy "#13 < #18 - #4"; +Goal "#13 < #18 - #4"; by (Simp_tac 1); result(); -goal Bin.thy "#~345 < #~242 + #~100"; +Goal "#~345 < #~242 + #~100"; by (Simp_tac 1); result(); -goal Bin.thy "#13557456 < #18678654"; +Goal "#13557456 < #18678654"; by (Simp_tac 1); result(); -goal Bin.thy "#999999 <= (#1000001 + #1)-#2"; +Goal "#999999 <= (#1000001 + #1)-#2"; by (Simp_tac 1); result(); -goal Bin.thy "#1234567 <= #1234567"; +Goal "#1234567 <= #1234567"; by (Simp_tac 1); result(); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/Equiv.ML --- a/src/HOL/Integ/Equiv.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/Equiv.ML Mon Jun 22 17:26:46 1998 +0200 @@ -16,17 +16,17 @@ (** first half: equiv A r ==> r^-1 O r = r **) -goalw Equiv.thy [trans_def,sym_def,converse_def] +Goalw [trans_def,sym_def,converse_def] "!!r. [| sym(r); trans(r) |] ==> r^-1 O r <= r"; by (Blast_tac 1); qed "sym_trans_comp_subset"; -goalw Equiv.thy [refl_def] +Goalw [refl_def] "!!A r. refl A r ==> r <= r^-1 O r"; by (Blast_tac 1); qed "refl_comp_subset"; -goalw Equiv.thy [equiv_def] +Goalw [equiv_def] "!!A r. equiv A r ==> r^-1 O r = r"; by (Clarify_tac 1); by (rtac equalityI 1); @@ -34,7 +34,7 @@ qed "equiv_comp_eq"; (*second half*) -goalw Equiv.thy [equiv_def,refl_def,sym_def,trans_def] +Goalw [equiv_def,refl_def,sym_def,trans_def] "!!A r. [| r^-1 O r = r; Domain(r) = A |] ==> equiv A r"; by (etac equalityE 1); by (subgoal_tac "ALL x y. (x,y) : r --> (y,x) : r" 1); @@ -44,35 +44,35 @@ (** Equivalence classes **) (*Lemma for the next result*) -goalw Equiv.thy [equiv_def,trans_def,sym_def] +Goalw [equiv_def,trans_def,sym_def] "!!A r. [| equiv A r; (a,b): r |] ==> r^^{a} <= r^^{b}"; by (Blast_tac 1); qed "equiv_class_subset"; -goal Equiv.thy "!!A r. [| equiv A r; (a,b): r |] ==> r^^{a} = r^^{b}"; +Goal "!!A r. [| equiv A r; (a,b): r |] ==> r^^{a} = r^^{b}"; by (REPEAT (ares_tac [equalityI, equiv_class_subset] 1)); by (rewrite_goals_tac [equiv_def,sym_def]); by (Blast_tac 1); qed "equiv_class_eq"; -goalw Equiv.thy [equiv_def,refl_def] +Goalw [equiv_def,refl_def] "!!A r. [| equiv A r; a: A |] ==> a: r^^{a}"; by (Blast_tac 1); qed "equiv_class_self"; (*Lemma for the next result*) -goalw Equiv.thy [equiv_def,refl_def] +Goalw [equiv_def,refl_def] "!!A r. [| equiv A r; r^^{b} <= r^^{a}; b: A |] ==> (a,b): r"; by (Blast_tac 1); qed "subset_equiv_class"; -goal Equiv.thy +Goal "!!A r. [| r^^{a} = r^^{b}; equiv A r; b: A |] ==> (a,b): r"; by (REPEAT (ares_tac [equalityD2, subset_equiv_class] 1)); qed "eq_equiv_class"; (*thus r^^{a} = r^^{b} as well*) -goalw Equiv.thy [equiv_def,trans_def,sym_def] +Goalw [equiv_def,trans_def,sym_def] "!!A r. [| equiv A r; x: (r^^{a} Int r^^{b}) |] ==> (a,b): r"; by (Blast_tac 1); qed "equiv_class_nondisjoint"; @@ -82,13 +82,13 @@ by (rtac (major RS conjunct1 RS conjunct1) 1); qed "equiv_type"; -goal Equiv.thy +Goal "!!A r. equiv A r ==> ((x,y): r) = (r^^{x} = r^^{y} & x:A & y:A)"; by (blast_tac (claset() addSIs [equiv_class_eq] addDs [eq_equiv_class, equiv_type]) 1); qed "equiv_class_eq_iff"; -goal Equiv.thy +Goal "!!A r. [| equiv A r; x: A; y: A |] ==> (r^^{x} = r^^{y}) = ((x,y): r)"; by (blast_tac (claset() addSIs [equiv_class_eq] addDs [eq_equiv_class, equiv_type]) 1); @@ -98,7 +98,7 @@ (** Introduction/elimination rules -- needed? **) -goalw Equiv.thy [quotient_def] "!!A. x:A ==> r^^{x}: A/r"; +Goalw [quotient_def] "!!A. x:A ==> r^^{x}: A/r"; by (Blast_tac 1); qed "quotientI"; @@ -111,12 +111,12 @@ by (Fast_tac 1); (*Blast_tac FAILS to prove it*) qed "quotientE"; -goalw Equiv.thy [equiv_def,refl_def,quotient_def] +Goalw [equiv_def,refl_def,quotient_def] "!!A r. equiv A r ==> Union(A/r) = A"; by (blast_tac (claset() addSIs [equalityI]) 1); qed "Union_quotient"; -goalw Equiv.thy [quotient_def] +Goalw [quotient_def] "!!A r. [| equiv A r; X: A/r; Y: A/r |] ==> X=Y | (X Int Y = {})"; by (safe_tac (claset() addSIs [equiv_class_eq])); by (assume_tac 1); @@ -139,7 +139,7 @@ **) (*Conversion rule*) -goal Equiv.thy "!!A r. [| equiv A r; congruent r b; a: A |] \ +Goal "!!A r. [| equiv A r; congruent r b; a: A |] \ \ ==> (UN x:r^^{a}. b(x)) = b(a)"; by (rtac (equiv_class_self RS UN_singleton) 1 THEN REPEAT (assume_tac 1)); by (rewrite_goals_tac [equiv_def,congruent_def,sym_def]); @@ -177,12 +177,12 @@ (**** Defining binary operations upon equivalence classes ****) -goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def] +Goalw [congruent_def,congruent2_def,equiv_def,refl_def] "!!A r. [| equiv A r; congruent2 r b; a: A |] ==> congruent r (b a)"; by (Blast_tac 1); qed "congruent2_implies_congruent"; -goalw Equiv.thy [congruent_def] +Goalw [congruent_def] "!!A r. [| equiv A r; congruent2 r b; a: A |] ==> \ \ congruent r (%x1. UN x2:r^^{a}. b x1 x2)"; by (Clarify_tac 1); @@ -193,7 +193,7 @@ by (Blast_tac 1); qed "congruent2_implies_congruent_UN"; -goal Equiv.thy +Goal "!!A r. [| equiv A r; congruent2 r b; a1: A; a2: A |] \ \ ==> (UN x1:r^^{a1}. UN x2:r^^{a2}. b x1 x2) = b a1 a2"; by (asm_simp_tac (simpset() addsimps [UN_equiv_class, @@ -244,21 +244,21 @@ (*** Cardinality results suggested by Florian Kammueller ***) (*Recall that equiv A r implies r <= A Times A (equiv_type) *) -goal thy "!!A. [| finite A; r <= A Times A |] ==> finite (A/r)"; +Goal "!!A. [| finite A; r <= A Times A |] ==> finite (A/r)"; by (rtac finite_subset 1); by (etac (finite_Pow_iff RS iffD2) 2); by (rewtac quotient_def); by (Blast_tac 1); qed "finite_quotient"; -goalw thy [quotient_def] +Goalw [quotient_def] "!!A. [| finite A; r <= A Times A; X : A/r |] ==> finite X"; by (rtac finite_subset 1); by (assume_tac 2); by (Blast_tac 1); qed "finite_equiv_class"; -goal thy "!!A. [| finite A; equiv A r; ! X : A/r. k dvd card(X) |] \ +Goal "!!A. [| finite A; equiv A r; ! X : A/r. k dvd card(X) |] \ \ ==> k dvd card(A)"; by (rtac (Union_quotient RS subst) 1); by (assume_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/Group.ML --- a/src/HOL/Integ/Group.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/Group.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,21 +10,21 @@ based on the unary inverse zero-x. *) -goal Group.thy "!!x::'a::add_group. (zero-x)+(x+y) = y"; +Goal "!!x::'a::add_group. (zero-x)+(x+y) = y"; by (rtac trans 1); by (rtac (plus_assoc RS sym) 1); by (stac left_inv 1); by (rtac zeroL 1); qed "left_inv2"; -goal Group.thy "!!x::'a::add_group. (zero-(zero-x)) = x"; +Goal "!!x::'a::add_group. (zero-(zero-x)) = x"; by (rtac trans 1); by (res_inst_tac [("x","zero-x")] left_inv2 2); by (stac left_inv 1); by (rtac (zeroR RS sym) 1); qed "inv_inv"; -goal Group.thy "zero-zero = (zero::'a::add_group)"; +Goal "zero-zero = (zero::'a::add_group)"; by (rtac trans 1); by (rtac (zeroR RS sym) 1); by (rtac trans 1); @@ -32,14 +32,14 @@ by (simp_tac (simpset() addsimps [zeroR]) 1); qed "inv_zero"; -goal Group.thy "!!x::'a::add_group. x+(zero-x) = zero"; +Goal "!!x::'a::add_group. x+(zero-x) = zero"; by (rtac trans 1); by (res_inst_tac [("x","zero-x")] left_inv 2); by (stac inv_inv 1); by (rtac refl 1); qed "right_inv"; -goal Group.thy "!!x::'a::add_group. x+((zero-x)+y) = y"; +Goal "!!x::'a::add_group. x+((zero-x)+y) = y"; by (rtac trans 1); by (res_inst_tac [("x","zero-x")] left_inv2 2); by (stac inv_inv 1); @@ -48,7 +48,7 @@ val plus_cong = read_instantiate [("f1","op +")] (arg_cong RS cong); -goal Group.thy "!!x::'a::add_group. zero-(x+y) = (zero-y)+(zero-x)"; +Goal "!!x::'a::add_group. zero-(x+y) = (zero-y)+(zero-x)"; by (rtac trans 1); by (rtac zeroR 2); by (rtac trans 1); @@ -102,19 +102,19 @@ (* The following two equations are not used in any of the decision procedures, but are still very useful. They also demonstrate mk_group1_tac. *) -goal Group.thy "x-x = (zero::'a::add_group)"; +Goal "x-x = (zero::'a::add_group)"; by (mk_group1_tac 1); by (group1_tac 1); qed "minus_self_zero"; -goal Group.thy "x-zero = (x::'a::add_group)"; +Goal "x-zero = (x::'a::add_group)"; by (mk_group1_tac 1); by (group1_tac 1); qed "minus_zero"; (*** Abelian Groups ***) -goal Group.thy "x+(y+z)=y+(x+z::'a::add_agroup)"; +Goal "x+(y+z)=y+(x+z::'a::add_agroup)"; by (rtac trans 1); by (rtac plus_commute 1); by (rtac trans 1); @@ -143,32 +143,32 @@ 2. and for moving `-' over to the other side: (x-y = z) = (x = z+y) and (x = y-z) = (x+z = y) *) -goal Group.thy "x + (y - z) = (x + y) - (z::'a::add_group)"; +Goal "x + (y - z) = (x + y) - (z::'a::add_group)"; by (mk_group1_tac 1); by (group1_tac 1); qed "plus_minusR"; -goal Group.thy "(x - y) + z = (x + z) - (y::'a::add_agroup)"; +Goal "(x - y) + z = (x + z) - (y::'a::add_agroup)"; by (mk_group1_tac 1); by (agroup1_tac 1); qed "plus_minusL"; -goal Group.thy "(x - y) - z = x - (y + (z::'a::add_agroup))"; +Goal "(x - y) - z = x - (y + (z::'a::add_agroup))"; by (mk_group1_tac 1); by (agroup1_tac 1); qed "minus_minusL"; -goal Group.thy "x - (y - z) = (x + z) - (y::'a::add_agroup)"; +Goal "x - (y - z) = (x + z) - (y::'a::add_agroup)"; by (mk_group1_tac 1); by (agroup1_tac 1); qed "minus_minusR"; -goal Group.thy "!!x::'a::add_group. (x-y = z) = (x = z+y)"; +Goal "!!x::'a::add_group. (x-y = z) = (x = z+y)"; by (stac minus_inv 1); by (fast_tac (claset() addss (HOL_basic_ss addsimps group1_simps)) 1); qed "minusL_iff"; -goal Group.thy "!!x::'a::add_group. (x = y-z) = (x+z = y)"; +Goal "!!x::'a::add_group. (x = y-z) = (x+z = y)"; by (stac minus_inv 1); by (fast_tac (claset() addss (HOL_basic_ss addsimps group1_simps)) 1); qed "minusR_iff"; @@ -182,34 +182,34 @@ it merely decides equality. (* Phase 1 *) -goal Group.thy "!!x::'a::add_agroup. (x+(zero-y))+z = (x+z)+(zero-y)"; +Goal "!!x::'a::add_agroup. (x+(zero-y))+z = (x+z)+(zero-y)"; by (Simp_tac 1); val lemma = result(); bind_thm("plus_minusL",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma); -goal Group.thy "!!x::'a::add_agroup. x+(zero-(y+z)) = (x+(zero-y))+(zero-z)"; +Goal "!!x::'a::add_agroup. x+(zero-(y+z)) = (x+(zero-y))+(zero-z)"; by (Simp_tac 1); val lemma = result(); bind_thm("minus_plusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma); -goal Group.thy "!!x::'a::add_agroup. x+(zero-(y+(zero-z))) = (x+z)+(zero-y)"; +Goal "!!x::'a::add_agroup. x+(zero-(y+(zero-z))) = (x+z)+(zero-y)"; by (Simp_tac 1); val lemma = result(); bind_thm("minus_minusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma); -goal Group.thy "!!x::'a::add_agroup. x+(y+(zero-z)) = (x+y)+(zero-z)"; +Goal "!!x::'a::add_agroup. x+(y+(zero-z)) = (x+y)+(zero-z)"; by (Simp_tac 1); val lemma = result(); bind_thm("plus_minusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma); (* Phase 2 *) -goal Group.thy "!!x::'a::add_agroup. (x+y)+(zero-z) = x+(y+(zero-z))"; +Goal "!!x::'a::add_agroup. (x+y)+(zero-z) = x+(y+(zero-z))"; by (Simp_tac 1); val lemma = result(); bind_thm("minus_plusL2",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma); -goal Group.thy "!!x::'a::add_agroup. (x+y)+(zero-x) = y"; +Goal "!!x::'a::add_agroup. (x+y)+(zero-x) = y"; by (rtac (plus_assoc RS trans) 1); by (rtac trans 1); by (rtac plus_cong 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/IntRing.ML --- a/src/HOL/Integ/IntRing.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/IntRing.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,7 +8,7 @@ open IntRing; -goal thy "!!i1::int. \ +Goal "!!i1::int. \ \ (sq i1 + sq i2 + sq i3 + sq i4) * (sq j1 + sq j2 + sq j3 + sq j4) = \ \ sq(i1*j1 - i2*j2 - i3*j3 - i4*j4) + \ \ sq(i1*j2 + i2*j1 + i3*j4 - i4*j3) + \ diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/IntRingDefs.ML --- a/src/HOL/Integ/IntRingDefs.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/IntRingDefs.ML Mon Jun 22 17:26:46 1998 +0200 @@ -7,10 +7,10 @@ open IntRingDefs; -goalw thy [zero_int_def,zdiff_def] "(zero-x)+(x::int) = zero"; +Goalw [zero_int_def,zdiff_def] "(zero-x)+(x::int) = zero"; by (Simp_tac 1); qed "left_inv_int"; -goalw thy [zero_int_def,zdiff_def] "x-y = (x::int) + (zero-y)"; +Goalw [zero_int_def,zdiff_def] "x-y = (x::int) + (zero-y)"; by (Simp_tac 1); qed "minus_inv_int"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/Integ.ML --- a/src/HOL/Integ/Integ.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/Integ.ML Mon Jun 22 17:26:46 1998 +0200 @@ -39,7 +39,7 @@ qed "intrelI"; (*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*) -goalw Integ.thy [intrel_def] +Goalw [intrel_def] "p: intrel --> (EX x1 y1 x2 y2. \ \ p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1)"; by (Fast_tac 1); @@ -56,15 +56,15 @@ AddSIs [intrelI]; AddSEs [intrelE]; -goal Integ.thy "((x1,y1),(x2,y2)): intrel = (x1+y2 = x2+y1)"; +Goal "((x1,y1),(x2,y2)): intrel = (x1+y2 = x2+y1)"; by (Fast_tac 1); qed "intrel_iff"; -goal Integ.thy "(x,x): intrel"; +Goal "(x,x): intrel"; by (stac surjective_pairing 1 THEN rtac (refl RS intrelI) 1); qed "intrel_refl"; -goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def] +Goalw [equiv_def, refl_def, sym_def, trans_def] "equiv {x::(nat*nat).True} intrel"; by (fast_tac (claset() addSIs [intrel_refl] addSEs [sym, integ_trans_lemma]) 1); @@ -75,11 +75,11 @@ ([CollectI, CollectI] MRS (equiv_intrel RS eq_equiv_class_iff)); -goalw Integ.thy [Integ_def,intrel_def,quotient_def] "intrel^^{(x,y)}:Integ"; +Goalw [Integ_def,intrel_def,quotient_def] "intrel^^{(x,y)}:Integ"; by (Fast_tac 1); qed "intrel_in_integ"; -goal Integ.thy "inj_on Abs_Integ Integ"; +Goal "inj_on Abs_Integ Integ"; by (rtac inj_on_inverseI 1); by (etac Abs_Integ_inverse 1); qed "inj_on_Abs_Integ"; @@ -87,7 +87,7 @@ Addsimps [equiv_intrel_iff, inj_on_Abs_Integ RS inj_on_iff, intrel_iff, intrel_in_integ, Abs_Integ_inverse]; -goal Integ.thy "inj(Rep_Integ)"; +Goal "inj(Rep_Integ)"; by (rtac inj_inverseI 1); by (rtac Rep_Integ_inverse 1); qed "inj_Rep_Integ"; @@ -97,7 +97,7 @@ (** znat: the injection from nat to Integ **) -goal Integ.thy "inj(znat)"; +Goal "inj(znat)"; by (rtac injI 1); by (rewtac znat_def); by (dtac (inj_on_Abs_Integ RS inj_onD) 1); @@ -112,7 +112,7 @@ (**** zminus: unary negation on Integ ****) -goalw Integ.thy [congruent_def] +Goalw [congruent_def] "congruent intrel (%p. split (%x y. intrel^^{(y,x)}) p)"; by Safe_tac; by (asm_simp_tac (simpset() addsimps add_ac) 1); @@ -122,7 +122,7 @@ (*Resolve th against the corresponding facts for zminus*) val zminus_ize = RSLIST [equiv_intrel, zminus_congruent]; -goalw Integ.thy [zminus_def] +Goalw [zminus_def] "$~ Abs_Integ(intrel^^{(x,y)}) = Abs_Integ(intrel ^^ {(y,x)})"; by (res_inst_tac [("f","Abs_Integ")] arg_cong 1); by (simp_tac (simpset() addsimps @@ -140,18 +140,18 @@ by (asm_full_simp_tac (simpset() addsimps [Rep_Integ_inverse]) 1); qed "eq_Abs_Integ"; -goal Integ.thy "$~ ($~ z) = z"; +Goal "$~ ($~ z) = z"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps [zminus]) 1); qed "zminus_zminus"; -goal Integ.thy "inj(zminus)"; +Goal "inj(zminus)"; by (rtac injI 1); by (dres_inst_tac [("f","zminus")] arg_cong 1); by (asm_full_simp_tac (simpset() addsimps [zminus_zminus]) 1); qed "inj_zminus"; -goalw Integ.thy [znat_def] "$~ ($#0) = $#0"; +Goalw [znat_def] "$~ ($#0) = $#0"; by (simp_tac (simpset() addsimps [zminus]) 1); qed "zminus_0"; @@ -159,13 +159,13 @@ (**** znegative: the test for negative integers ****) -goalw Integ.thy [znegative_def, znat_def] +Goalw [znegative_def, znat_def] "~ znegative($# n)"; by (Simp_tac 1); by Safe_tac; qed "not_znegative_znat"; -goalw Integ.thy [znegative_def, znat_def] "znegative($~ $# Suc(n))"; +Goalw [znegative_def, znat_def] "znegative($~ $# Suc(n))"; by (simp_tac (simpset() addsimps [zminus]) 1); qed "znegative_zminus_znat"; @@ -182,7 +182,7 @@ by (ALLGOALS Asm_simp_tac); qed "diff_Suc_add_inverse"; -goalw Integ.thy [congruent_def] +Goalw [congruent_def] "congruent intrel (split (%x y. intrel^^{((y-x) + (x-(y::nat)),0)}))"; by Safe_tac; by (Asm_simp_tac 1); @@ -201,18 +201,18 @@ val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent]; -goalw Integ.thy [zmagnitude_def] +Goalw [zmagnitude_def] "zmagnitude (Abs_Integ(intrel^^{(x,y)})) = \ \ Abs_Integ(intrel^^{((y - x) + (x - y),0)})"; by (res_inst_tac [("f","Abs_Integ")] arg_cong 1); by (asm_simp_tac (simpset() addsimps [zmagnitude_ize UN_equiv_class]) 1); qed "zmagnitude"; -goalw Integ.thy [znat_def] "zmagnitude($# n) = $#n"; +Goalw [znat_def] "zmagnitude($# n) = $#n"; by (asm_simp_tac (simpset() addsimps [zmagnitude]) 1); qed "zmagnitude_znat"; -goalw Integ.thy [znat_def] "zmagnitude($~ $# n) = $#n"; +Goalw [znat_def] "zmagnitude($~ $# n) = $#n"; by (asm_simp_tac (simpset() addsimps [zmagnitude, zminus]) 1); qed "zmagnitude_zminus_znat"; @@ -221,7 +221,7 @@ (** Congruence property for addition **) -goalw Integ.thy [congruent2_def] +Goalw [congruent2_def] "congruent2 intrel (%p1 p2. \ \ split (%x1 y1. split (%x2 y2. intrel^^{(x1+x2, y1+y2)}) p2) p1)"; (*Proof via congruent2_commuteI seems longer*) @@ -235,31 +235,31 @@ (*Resolve th against the corresponding facts for zadd*) val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2]; -goalw Integ.thy [zadd_def] +Goalw [zadd_def] "Abs_Integ(intrel^^{(x1,y1)}) + Abs_Integ(intrel^^{(x2,y2)}) = \ \ Abs_Integ(intrel^^{(x1+x2, y1+y2)})"; by (asm_simp_tac (simpset() addsimps [zadd_ize UN_equiv_class2]) 1); qed "zadd"; -goalw Integ.thy [znat_def] "$#0 + z = z"; +Goalw [znat_def] "$#0 + z = z"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps [zadd]) 1); qed "zadd_0"; -goal Integ.thy "$~ (z + w) = $~ z + $~ w"; +Goal "$~ (z + w) = $~ z + $~ w"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (res_inst_tac [("z","w")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps [zminus,zadd]) 1); qed "zminus_zadd_distrib"; -goal Integ.thy "(z::int) + w = w + z"; +Goal "(z::int) + w = w + z"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (res_inst_tac [("z","w")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps (add_ac @ [zadd])) 1); qed "zadd_commute"; -goal Integ.thy "((z1::int) + z2) + z3 = z1 + (z2 + z3)"; +Goal "((z1::int) + z2) + z3 = z1 + (z2 + z3)"; by (res_inst_tac [("z","z1")] eq_Abs_Integ 1); by (res_inst_tac [("z","z2")] eq_Abs_Integ 1); by (res_inst_tac [("z","z3")] eq_Abs_Integ 1); @@ -267,7 +267,7 @@ qed "zadd_assoc"; (*For AC rewriting*) -goal Integ.thy "(x::int)+(y+z)=y+(x+z)"; +Goal "(x::int)+(y+z)=y+(x+z)"; by (rtac (zadd_commute RS trans) 1); by (rtac (zadd_assoc RS trans) 1); by (rtac (zadd_commute RS arg_cong) 1); @@ -276,21 +276,21 @@ (*Integer addition is an AC operator*) val zadd_ac = [zadd_assoc,zadd_commute,zadd_left_commute]; -goalw Integ.thy [znat_def] "$# (m + n) = ($#m) + ($#n)"; +Goalw [znat_def] "$# (m + n) = ($#m) + ($#n)"; by (asm_simp_tac (simpset() addsimps [zadd]) 1); qed "znat_add"; -goalw Integ.thy [znat_def] "z + ($~ z) = $#0"; +Goalw [znat_def] "z + ($~ z) = $#0"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps [zminus, zadd, add_commute]) 1); qed "zadd_zminus_inverse"; -goal Integ.thy "($~ z) + z = $#0"; +Goal "($~ z) + z = $#0"; by (rtac (zadd_commute RS trans) 1); by (rtac zadd_zminus_inverse 1); qed "zadd_zminus_inverse2"; -goal Integ.thy "z + $#0 = z"; +Goal "z + $#0 = z"; by (rtac (zadd_commute RS trans) 1); by (rtac zadd_0 1); qed "zadd_0_right"; @@ -312,11 +312,11 @@ (** Congruence property for multiplication **) -goal Integ.thy "((k::nat) + l) + (m + n) = (k + m) + (n + l)"; +Goal "((k::nat) + l) + (m + n) = (k + m) + (n + l)"; by (simp_tac (simpset() addsimps add_ac) 1); qed "zmult_congruent_lemma"; -goal Integ.thy +Goal "congruent2 intrel (%p1 p2. \ \ split (%x1 y1. split (%x2 y2. \ \ intrel^^{(x1*x2 + y1*y2, x1*y2 + y1*x2)}) p2) p1)"; @@ -340,50 +340,50 @@ (*Resolve th against the corresponding facts for zmult*) val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2]; -goalw Integ.thy [zmult_def] +Goalw [zmult_def] "Abs_Integ((intrel^^{(x1,y1)})) * Abs_Integ((intrel^^{(x2,y2)})) = \ \ Abs_Integ(intrel ^^ {(x1*x2 + y1*y2, x1*y2 + y1*x2)})"; by (simp_tac (simpset() addsimps [zmult_ize UN_equiv_class2]) 1); qed "zmult"; -goalw Integ.thy [znat_def] "$#0 * z = $#0"; +Goalw [znat_def] "$#0 * z = $#0"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps [zmult]) 1); qed "zmult_0"; -goalw Integ.thy [znat_def] "$#Suc(0) * z = z"; +Goalw [znat_def] "$#Suc(0) * z = z"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps [zmult]) 1); qed "zmult_1"; -goal Integ.thy "($~ z) * w = $~ (z * w)"; +Goal "($~ z) * w = $~ (z * w)"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (res_inst_tac [("z","w")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1); qed "zmult_zminus"; -goal Integ.thy "($~ z) * ($~ w) = (z * w)"; +Goal "($~ z) * ($~ w) = (z * w)"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (res_inst_tac [("z","w")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1); qed "zmult_zminus_zminus"; -goal Integ.thy "(z::int) * w = w * z"; +Goal "(z::int) * w = w * z"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (res_inst_tac [("z","w")] eq_Abs_Integ 1); by (asm_simp_tac (simpset() addsimps ([zmult] @ add_ac @ mult_ac)) 1); qed "zmult_commute"; -goal Integ.thy "z * $# 0 = $#0"; +Goal "z * $# 0 = $#0"; by (rtac ([zmult_commute, zmult_0] MRS trans) 1); qed "zmult_0_right"; -goal Integ.thy "z * $#Suc(0) = z"; +Goal "z * $#Suc(0) = z"; by (rtac ([zmult_commute, zmult_1] MRS trans) 1); qed "zmult_1_right"; -goal Integ.thy "((z1::int) * z2) * z3 = z1 * (z2 * z3)"; +Goal "((z1::int) * z2) * z3 = z1 * (z2 * z3)"; by (res_inst_tac [("z","z1")] eq_Abs_Integ 1); by (res_inst_tac [("z","z2")] eq_Abs_Integ 1); by (res_inst_tac [("z","z3")] eq_Abs_Integ 1); @@ -400,7 +400,7 @@ (*Integer multiplication is an AC operator*) val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute]; -goal Integ.thy "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"; +Goal "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"; by (res_inst_tac [("z","z1")] eq_Abs_Integ 1); by (res_inst_tac [("z","z2")] eq_Abs_Integ 1); by (res_inst_tac [("z","w")] eq_Abs_Integ 1); @@ -411,11 +411,11 @@ val zmult_commute'= read_instantiate [("z","w")] zmult_commute; -goal Integ.thy "w * ($~ z) = $~ (w * z)"; +Goal "w * ($~ z) = $~ (w * z)"; by (simp_tac (simpset() addsimps [zmult_commute', zmult_zminus]) 1); qed "zmult_zminus_right"; -goal Integ.thy "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"; +Goal "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"; by (simp_tac (simpset() addsimps [zmult_commute',zadd_zmult_distrib]) 1); qed "zadd_zmult_distrib2"; @@ -434,53 +434,53 @@ (**** Additional Theorems (by Mattolini; proofs mainly by lcp) ****) (* Some Theorems about zsuc and zpred *) -goalw Integ.thy [zsuc_def] "$#(Suc(n)) = zsuc($# n)"; +Goalw [zsuc_def] "$#(Suc(n)) = zsuc($# n)"; by (simp_tac (simpset() addsimps [znat_add RS sym]) 1); qed "znat_Suc"; -goalw Integ.thy [zpred_def,zsuc_def,zdiff_def] "$~ zsuc(z) = zpred($~ z)"; +Goalw [zpred_def,zsuc_def,zdiff_def] "$~ zsuc(z) = zpred($~ z)"; by (Simp_tac 1); qed "zminus_zsuc"; -goalw Integ.thy [zpred_def,zsuc_def,zdiff_def] "$~ zpred(z) = zsuc($~ z)"; +Goalw [zpred_def,zsuc_def,zdiff_def] "$~ zpred(z) = zsuc($~ z)"; by (Simp_tac 1); qed "zminus_zpred"; -goalw Integ.thy [zsuc_def,zpred_def,zdiff_def] +Goalw [zsuc_def,zpred_def,zdiff_def] "zpred(zsuc(z)) = z"; by (simp_tac (simpset() addsimps [zadd_assoc]) 1); qed "zpred_zsuc"; -goalw Integ.thy [zsuc_def,zpred_def,zdiff_def] +Goalw [zsuc_def,zpred_def,zdiff_def] "zsuc(zpred(z)) = z"; by (simp_tac (simpset() addsimps [zadd_assoc]) 1); qed "zsuc_zpred"; -goal Integ.thy "(zpred(z)=w) = (z=zsuc(w))"; +Goal "(zpred(z)=w) = (z=zsuc(w))"; by Safe_tac; by (rtac (zsuc_zpred RS sym) 1); by (rtac zpred_zsuc 1); qed "zpred_to_zsuc"; -goal Integ.thy "(zsuc(z)=w)=(z=zpred(w))"; +Goal "(zsuc(z)=w)=(z=zpred(w))"; by Safe_tac; by (rtac (zpred_zsuc RS sym) 1); by (rtac zsuc_zpred 1); qed "zsuc_to_zpred"; -goal Integ.thy "($~ z = w) = (z = $~ w)"; +Goal "($~ z = w) = (z = $~ w)"; by Safe_tac; by (rtac (zminus_zminus RS sym) 1); by (rtac zminus_zminus 1); qed "zminus_exchange"; -goal Integ.thy"(zsuc(z)=zsuc(w)) = (z=w)"; +Goal"(zsuc(z)=zsuc(w)) = (z=w)"; by Safe_tac; by (dres_inst_tac [("f","zpred")] arg_cong 1); by (asm_full_simp_tac (simpset() addsimps [zpred_zsuc]) 1); qed "bijective_zsuc"; -goal Integ.thy"(zpred(z)=zpred(w)) = (z=w)"; +Goal"(zpred(z)=zpred(w)) = (z=w)"; by Safe_tac; by (dres_inst_tac [("f","zsuc")] arg_cong 1); by (asm_full_simp_tac (simpset() addsimps [zsuc_zpred]) 1); @@ -488,58 +488,58 @@ (* Additional Theorems about zadd *) -goalw Integ.thy [zsuc_def] "zsuc(z) + w = zsuc(z+w)"; +Goalw [zsuc_def] "zsuc(z) + w = zsuc(z+w)"; by (simp_tac (simpset() addsimps zadd_ac) 1); qed "zadd_zsuc"; -goalw Integ.thy [zsuc_def] "w + zsuc(z) = zsuc(w+z)"; +Goalw [zsuc_def] "w + zsuc(z) = zsuc(w+z)"; by (simp_tac (simpset() addsimps zadd_ac) 1); qed "zadd_zsuc_right"; -goalw Integ.thy [zpred_def,zdiff_def] "zpred(z) + w = zpred(z+w)"; +Goalw [zpred_def,zdiff_def] "zpred(z) + w = zpred(z+w)"; by (simp_tac (simpset() addsimps zadd_ac) 1); qed "zadd_zpred"; -goalw Integ.thy [zpred_def,zdiff_def] "w + zpred(z) = zpred(w+z)"; +Goalw [zpred_def,zdiff_def] "w + zpred(z) = zpred(w+z)"; by (simp_tac (simpset() addsimps zadd_ac) 1); qed "zadd_zpred_right"; (* Additional Theorems about zmult *) -goalw Integ.thy [zsuc_def] "zsuc(w) * z = z + w * z"; +Goalw [zsuc_def] "zsuc(w) * z = z + w * z"; by (simp_tac (simpset() addsimps [zadd_zmult_distrib, zadd_commute]) 1); qed "zmult_zsuc"; -goalw Integ.thy [zsuc_def] "z * zsuc(w) = z + w * z"; +Goalw [zsuc_def] "z * zsuc(w) = z + w * z"; by (simp_tac (simpset() addsimps [zadd_zmult_distrib2, zadd_commute, zmult_commute]) 1); qed "zmult_zsuc_right"; -goalw Integ.thy [zpred_def, zdiff_def] "zpred(w) * z = w * z - z"; +Goalw [zpred_def, zdiff_def] "zpred(w) * z = w * z - z"; by (simp_tac (simpset() addsimps [zadd_zmult_distrib]) 1); qed "zmult_zpred"; -goalw Integ.thy [zpred_def, zdiff_def] "z * zpred(w) = w * z - z"; +Goalw [zpred_def, zdiff_def] "z * zpred(w) = w * z - z"; by (simp_tac (simpset() addsimps [zadd_zmult_distrib2, zmult_commute]) 1); qed "zmult_zpred_right"; (* Further Theorems about zsuc and zpred *) -goal Integ.thy "$#Suc(m) ~= $#0"; +Goal "$#Suc(m) ~= $#0"; by (simp_tac (simpset() addsimps [inj_znat RS inj_eq]) 1); qed "znat_Suc_not_znat_Zero"; bind_thm ("znat_Zero_not_znat_Suc", (znat_Suc_not_znat_Zero RS not_sym)); -goalw Integ.thy [zsuc_def,znat_def] "w ~= zsuc(w)"; +Goalw [zsuc_def,znat_def] "w ~= zsuc(w)"; by (res_inst_tac [("z","w")] eq_Abs_Integ 1); by (asm_full_simp_tac (simpset() addsimps [zadd]) 1); qed "n_not_zsuc_n"; val zsuc_n_not_n = n_not_zsuc_n RS not_sym; -goal Integ.thy "w ~= zpred(w)"; +Goal "w ~= zpred(w)"; by Safe_tac; by (dres_inst_tac [("x","w"),("f","zsuc")] arg_cong 1); by (asm_full_simp_tac (simpset() addsimps [zsuc_zpred,zsuc_n_not_n]) 1); @@ -550,7 +550,7 @@ (* Theorems about less and less_equal *) -goalw Integ.thy [zless_def, znegative_def, zdiff_def, znat_def] +Goalw [zless_def, znegative_def, zdiff_def, znat_def] "!!w. w ? n. z = w + $#(Suc(n))"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (res_inst_tac [("z","w")] eq_Abs_Integ 1); @@ -561,26 +561,26 @@ by (asm_full_simp_tac (simpset() addsimps add_ac) 1); qed "zless_eq_zadd_Suc"; -goalw Integ.thy [zless_def, znegative_def, zdiff_def, znat_def] +Goalw [zless_def, znegative_def, zdiff_def, znat_def] "z < z + $#(Suc(n))"; by (res_inst_tac [("z","z")] eq_Abs_Integ 1); by (Clarify_tac 1); by (simp_tac (simpset() addsimps [zadd, zminus]) 1); qed "zless_zadd_Suc"; -goal Integ.thy "!!z1 z2 z3. [| z1 z1 < (z3::int)"; +Goal "!!z1 z2 z3. [| z1 z1 < (z3::int)"; by (safe_tac (claset() addSDs [zless_eq_zadd_Suc])); by (simp_tac (simpset() addsimps [zadd_assoc, zless_zadd_Suc, znat_add RS sym]) 1); qed "zless_trans"; -goalw Integ.thy [zsuc_def] "z ~w ~w R *) bind_thm ("zless_asym", (zless_not_sym RS notE)); -goal Integ.thy "!!z::int. ~ z R *) bind_thm ("zless_irrefl", (zless_not_refl RS notE)); -goal Integ.thy "!!w. z w ~= (z::int)"; +Goal "!!w. z w ~= (z::int)"; by (fast_tac (claset() addEs [zless_irrefl]) 1); qed "zless_not_refl2"; (*"Less than" is a linear ordering*) -goalw Integ.thy [zless_def, znegative_def, zdiff_def] +Goalw [zless_def, znegative_def, zdiff_def] "z z<=(w::int)"; +Goalw [zle_def] "!!w. ~(w z<=(w::int)"; by (assume_tac 1); qed "zleI"; -goalw Integ.thy [zle_def] "!!w. z<=w ==> ~(w<(z::int))"; +Goalw [zle_def] "!!w. z<=w ==> ~(w<(z::int))"; by (assume_tac 1); qed "zleD"; val zleE = make_elim zleD; -goalw Integ.thy [zle_def] "!!z. ~ z <= w ==> w<(z::int)"; +Goalw [zle_def] "!!z. ~ z <= w ==> w<(z::int)"; by (Fast_tac 1); qed "not_zleE"; -goalw Integ.thy [zle_def] "!!z. z < w ==> z <= (w::int)"; +Goalw [zle_def] "!!z. z < w ==> z <= (w::int)"; by (fast_tac (claset() addEs [zless_asym]) 1); qed "zless_imp_zle"; -goalw Integ.thy [zle_def] "!!z. z <= w ==> z < w | z=(w::int)"; +Goalw [zle_def] "!!z. z <= w ==> z < w | z=(w::int)"; by (cut_facts_tac [zless_linear] 1); by (fast_tac (claset() addEs [zless_irrefl,zless_asym]) 1); qed "zle_imp_zless_or_eq"; -goalw Integ.thy [zle_def] "!!z. z z <=(w::int)"; +Goalw [zle_def] "!!z. z z <=(w::int)"; by (cut_facts_tac [zless_linear] 1); by (fast_tac (claset() addEs [zless_irrefl,zless_asym]) 1); qed "zless_or_eq_imp_zle"; -goal Integ.thy "(x <= (y::int)) = (x < y | x=y)"; +Goal "(x <= (y::int)) = (x < y | x=y)"; by (REPEAT(ares_tac [iffI, zless_or_eq_imp_zle, zle_imp_zless_or_eq] 1)); qed "zle_eq_zless_or_eq"; -goal Integ.thy "w <= (w::int)"; +Goal "w <= (w::int)"; by (simp_tac (simpset() addsimps [zle_eq_zless_or_eq]) 1); qed "zle_refl"; @@ -670,18 +670,18 @@ by (fast_tac (claset() addIs [zless_trans]) 1); qed "zle_zless_trans"; -goal Integ.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::int)"; +Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::int)"; by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, rtac zless_or_eq_imp_zle, fast_tac (claset() addIs [zless_trans])]); qed "zle_trans"; -goal Integ.thy "!!z. [| z <= w; w <= z |] ==> z = (w::int)"; +Goal "!!z. [| z <= w; w <= z |] ==> z = (w::int)"; by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, fast_tac (claset() addEs [zless_irrefl,zless_asym])]); qed "zle_anti_sym"; -goal Integ.thy "!!w w' z::int. z + w' = z + w ==> w' = w"; +Goal "!!w w' z::int. z + w' = z + w ==> w' = w"; by (dres_inst_tac [("f", "%x. x + $~z")] arg_cong 1); by (asm_full_simp_tac (simpset() addsimps zadd_ac) 1); qed "zadd_left_cancel"; @@ -689,19 +689,19 @@ (*** Monotonicity results ***) -goal Integ.thy "!!v w z::int. v < w ==> v + z < w + z"; +Goal "!!v w z::int. v < w ==> v + z < w + z"; by (safe_tac (claset() addSDs [zless_eq_zadd_Suc])); by (simp_tac (simpset() addsimps zadd_ac) 1); by (simp_tac (simpset() addsimps [zadd_assoc RS sym, zless_zadd_Suc]) 1); qed "zadd_zless_mono1"; -goal Integ.thy "!!v w z::int. (v+z < w+z) = (v < w)"; +Goal "!!v w z::int. (v+z < w+z) = (v < w)"; by (safe_tac (claset() addSEs [zadd_zless_mono1])); by (dres_inst_tac [("z", "$~z")] zadd_zless_mono1 1); by (asm_full_simp_tac (simpset() addsimps [zadd_assoc]) 1); qed "zadd_left_cancel_zless"; -goal Integ.thy "!!v w z::int. (v+z <= w+z) = (v <= w)"; +Goal "!!v w z::int. (v+z <= w+z) = (v <= w)"; by (asm_full_simp_tac (simpset() addsimps [zle_def, zadd_left_cancel_zless]) 1); qed "zadd_left_cancel_zle"; @@ -710,14 +710,14 @@ bind_thm ("zadd_zle_mono1", zadd_left_cancel_zle RS iffD2); -goal Integ.thy "!!z' z::int. [| w'<=w; z'<=z |] ==> w' + z' <= w + z"; +Goal "!!z' z::int. [| w'<=w; z'<=z |] ==> w' + z' <= w + z"; by (etac (zadd_zle_mono1 RS zle_trans) 1); by (simp_tac (simpset() addsimps [zadd_commute]) 1); (*w moves to the end because it is free while z', z are bound*) by (etac zadd_zle_mono1 1); qed "zadd_zle_mono"; -goal Integ.thy "!!w z::int. z<=$#0 ==> w+z <= w"; +Goal "!!w z::int. z<=$#0 ==> w+z <= w"; by (dres_inst_tac [("z", "w")] zadd_zle_mono1 1); by (asm_full_simp_tac (simpset() addsimps [zadd_commute]) 1); qed "zadd_zle_self"; @@ -736,67 +736,67 @@ Addsimps [zless_eq_less, zle_eq_le, znegative_zminus_znat, not_znegative_znat]; -goal Integ.thy "!! x. (x::int) = y ==> x <= y"; +Goal "!! x. (x::int) = y ==> x <= y"; by (etac subst 1); by (rtac zle_refl 1); qed "zequalD1"; -goal Integ.thy "($~ x < $~ y) = (y < x)"; +Goal "($~ x < $~ y) = (y < x)"; by (rewrite_goals_tac [zless_def,zdiff_def]); by (simp_tac (simpset() addsimps zadd_ac ) 1); qed "zminus_zless_zminus"; -goal Integ.thy "($~ x <= $~ y) = (y <= x)"; +Goal "($~ x <= $~ y) = (y <= x)"; by (simp_tac (HOL_ss addsimps[zle_def, zminus_zless_zminus]) 1); qed "zminus_zle_zminus"; -goal Integ.thy "(x < $~ y) = (y < $~ x)"; +Goal "(x < $~ y) = (y < $~ x)"; by (rewrite_goals_tac [zless_def,zdiff_def]); by (simp_tac (simpset() addsimps zadd_ac ) 1); qed "zless_zminus"; -goal Integ.thy "($~ x < y) = ($~ y < x)"; +Goal "($~ x < y) = ($~ y < x)"; by (rewrite_goals_tac [zless_def,zdiff_def]); by (simp_tac (simpset() addsimps zadd_ac ) 1); qed "zminus_zless"; -goal Integ.thy "(x <= $~ y) = (y <= $~ x)"; +Goal "(x <= $~ y) = (y <= $~ x)"; by (simp_tac (HOL_ss addsimps[zle_def, zminus_zless]) 1); qed "zle_zminus"; -goal Integ.thy "($~ x <= y) = ($~ y <= x)"; +Goal "($~ x <= y) = ($~ y <= x)"; by (simp_tac (HOL_ss addsimps[zle_def, zless_zminus]) 1); qed "zminus_zle"; -goal Integ.thy " $#0 < $# Suc n"; +Goal " $#0 < $# Suc n"; by (rtac (zero_less_Suc RS (zless_eq_less RS iffD2)) 1); qed "zero_zless_Suc_pos"; -goal Integ.thy "($# n= $# m) = (n = m)"; +Goal "($# n= $# m) = (n = m)"; by (fast_tac (HOL_cs addSEs[inj_znat RS injD]) 1); qed "znat_znat_eq"; AddIffs[znat_znat_eq]; -goal Integ.thy "$~ $# Suc n < $#0"; +Goal "$~ $# Suc n < $#0"; by (stac (zminus_0 RS sym) 1); by (rtac (zminus_zless_zminus RS iffD2) 1); by (rtac (zero_less_Suc RS (zless_eq_less RS iffD2)) 1); qed "negative_zless_0"; Addsimps [zero_zless_Suc_pos, negative_zless_0]; -goal Integ.thy "$~ $# n <= $#0"; +Goal "$~ $# n <= $#0"; by (rtac zless_or_eq_imp_zle 1); by (nat_ind_tac "n" 1); by (ALLGOALS Asm_simp_tac); qed "negative_zle_0"; Addsimps[negative_zle_0]; -goal Integ.thy "~($#0 <= $~ $# Suc n)"; +Goal "~($#0 <= $~ $# Suc n)"; by (stac zle_zminus 1); by (Simp_tac 1); qed "not_zle_0_negative"; Addsimps[not_zle_0_negative]; -goal Integ.thy "($# n <= $~ $# m) = (n = 0 & m = 0)"; +Goal "($# n <= $~ $# m) = (n = 0 & m = 0)"; by (safe_tac HOL_cs); by (Simp_tac 3); by (dtac (zle_zminus RS iffD1) 2); @@ -804,7 +804,7 @@ by (ALLGOALS Asm_full_simp_tac); qed "znat_zle_znegative"; -goal Integ.thy "~($# n < $~ $# Suc m)"; +Goal "~($# n < $~ $# Suc m)"; by (rtac notI 1); by (forward_tac [zless_imp_zle] 1); by (dtac (znat_zle_znegative RS iffD1) 1); by (safe_tac HOL_cs); @@ -812,7 +812,7 @@ by (Asm_full_simp_tac 1); qed "not_znat_zless_negative"; -goal Integ.thy "($~ $# n = $# m) = (n = 0 & m = 0)"; +Goal "($~ $# n = $# m) = (n = 0 & m = 0)"; by (rtac iffI 1); by (rtac (znat_zle_znegative RS iffD1) 1); by (dtac sym 1); @@ -822,16 +822,16 @@ Addsimps [zminus_zless_zminus, zminus_zle_zminus, negative_eq_positive, not_znat_zless_negative]; -goalw Integ.thy [zdiff_def,zless_def] "!! x. znegative x = (x < $# 0)"; +Goalw [zdiff_def,zless_def] "!! x. znegative x = (x < $# 0)"; by Auto_tac; qed "znegative_less_0"; -goalw Integ.thy [zdiff_def,zless_def] "!! x. (~znegative x) = ($# 0 <= x)"; +Goalw [zdiff_def,zless_def] "!! x. (~znegative x) = ($# 0 <= x)"; by (stac znegative_less_0 1); by (safe_tac (HOL_cs addSDs[zleD,not_zleE,zleI]) ); qed "not_znegative_ge_0"; -goal Integ.thy "!! x. znegative x ==> ? n. x = $~ $# Suc n"; +Goal "!! x. znegative x ==> ? n. x = $~ $# Suc n"; by (dtac (znegative_less_0 RS iffD1 RS zless_eq_zadd_Suc) 1); by (etac exE 1); by (rtac exI 1); @@ -839,7 +839,7 @@ by (auto_tac(claset(), simpset() addsimps [zadd_assoc])); qed "znegativeD"; -goal Integ.thy "!! x. ~znegative x ==> ? n. x = $# n"; +Goal "!! x. ~znegative x ==> ? n. x = $# n"; by (dtac (not_znegative_ge_0 RS iffD1) 1); by (dtac zle_imp_zless_or_eq 1); by (etac disjE 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/Lagrange.ML --- a/src/HOL/Integ/Lagrange.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/Lagrange.ML Mon Jun 22 17:26:46 1998 +0200 @@ -9,7 +9,7 @@ abstract thm about commutative rings and has a priori nothing to do with nat. *) -goalw Lagrange.thy [Lagrange.sq_def] "!!x1::'a::cring. \ +Goalw [Lagrange.sq_def] "!!x1::'a::cring. \ \ (sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) = \ \ sq(x1*y1 - x2*y2 - x3*y3 - x4*y4) + \ \ sq(x1*y2 + x2*y1 + x3*y4 - x4*y3) + \ @@ -22,7 +22,7 @@ (* A challenge by John Harrison. Takes forever because of the naive bottom-up strategy of the rewriter. -goalw Lagrange.thy [Lagrange.sq_def] "!!p1::'a::cring.\ +Goalw [Lagrange.sq_def] "!!p1::'a::cring.\ \ (sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) * \ \ (sq p2 + sq q2 + sq r2 + sq s2 + sq t2 + sq u2 + sq v2 + sq w2) \ \ = sq (p1*p2 - q1*q2 - r1*r2 - s1*s2 - t1*t2 - u1*u2 - v1*v2 - w1*w2) + \ diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Integ/Ring.ML --- a/src/HOL/Integ/Ring.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Integ/Ring.ML Mon Jun 22 17:26:46 1998 +0200 @@ -7,7 +7,7 @@ and defines cring_simpl, a simplification tactic for commutative rings. *) -goal Ring.thy "!!x::'a::cring. x*(y*z)=y*(x*z)"; +Goal "!!x::'a::cring. x*(y*z)=y*(x*z)"; by (rtac trans 1); by (rtac times_commute 1); by (rtac trans 1); @@ -17,7 +17,7 @@ val times_cong = read_instantiate [("f1","op *")] (arg_cong RS cong); -goal Ring.thy "!!x::'a::ring. zero*x = zero"; +Goal "!!x::'a::ring. zero*x = zero"; by (rtac trans 1); by (rtac right_inv 2); by (rtac trans 1); @@ -37,7 +37,7 @@ by (rtac (zeroR RS sym) 1); qed "mult_zeroL"; -goal Ring.thy "!!x::'a::ring. x*zero = zero"; +Goal "!!x::'a::ring. x*zero = zero"; by (rtac trans 1); by (rtac right_inv 2); by (rtac trans 1); @@ -57,7 +57,7 @@ by (rtac (zeroR RS sym) 1); qed "mult_zeroR"; -goal Ring.thy "!!x::'a::ring. (zero-x)*y = zero-(x*y)"; +Goal "!!x::'a::ring. (zero-x)*y = zero-(x*y)"; by (rtac trans 1); by (rtac zeroL 2); by (rtac trans 1); @@ -83,7 +83,7 @@ by (rtac (zeroR RS sym) 1); qed "mult_invL"; -goal Ring.thy "!!x::'a::ring. x*(zero-y) = zero-(x*y)"; +Goal "!!x::'a::ring. x*(zero-y) = zero-(x*y)"; by (rtac trans 1); by (rtac zeroL 2); by (rtac trans 1); @@ -109,12 +109,12 @@ by (rtac (zeroR RS sym) 1); qed "mult_invR"; -goal Ring.thy "x*(y-z) = (x*y - x*z::'a::ring)"; +Goal "x*(y-z) = (x*y - x*z::'a::ring)"; by (mk_group1_tac 1); by (simp_tac (HOL_basic_ss addsimps [distribL,mult_invR]) 1); qed "minus_distribL"; -goal Ring.thy "(x-y)*z = (x*z - y*z::'a::ring)"; +Goal "(x-y)*z = (x*z - y*z::'a::ring)"; by (mk_group1_tac 1); by (simp_tac (HOL_basic_ss addsimps [distribR,mult_invL]) 1); qed "minus_distribR"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lambda/Commutation.ML --- a/src/HOL/Lambda/Commutation.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lambda/Commutation.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,21 +10,21 @@ (*** square ***) -goalw Commutation.thy [square_def] "!!R. square R S T U ==> square S R U T"; +Goalw [square_def] "!!R. square R S T U ==> square S R U T"; by (Blast_tac 1); qed "square_sym"; -goalw Commutation.thy [square_def] +Goalw [square_def] "!!R. [| square R S T U; T <= T' |] ==> square R S T' U"; by (Blast_tac 1); qed "square_subset"; -goalw Commutation.thy [square_def] +Goalw [square_def] "!!R. [| square R S T (R^=); S <= T |] ==> square (R^=) S T (R^=)"; by (Blast_tac 1); qed "square_reflcl"; -goalw Commutation.thy [square_def] +Goalw [square_def] "!!R. square R S S T ==> square (R^*) S S (T^*)"; by (strip_tac 1); by (etac rtrancl_induct 1); @@ -32,7 +32,7 @@ by (blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); qed "square_rtrancl"; -goalw Commutation.thy [commute_def] +Goalw [commute_def] "!!R. square R S (S^*) (R^=) ==> commute (R^*) (S^*)"; by (fast_tac (claset() addDs [square_reflcl,square_sym RS square_rtrancl] addEs [r_into_rtrancl] @@ -41,44 +41,44 @@ (*** commute ***) -goalw Commutation.thy [commute_def] "!!R. commute R S ==> commute S R"; +Goalw [commute_def] "!!R. commute R S ==> commute S R"; by (blast_tac (claset() addIs [square_sym]) 1); qed "commute_sym"; -goalw Commutation.thy [commute_def] "!!R. commute R S ==> commute (R^*) (S^*)"; +Goalw [commute_def] "!!R. commute R S ==> commute (R^*) (S^*)"; by (blast_tac (claset() addIs [square_rtrancl,square_sym]) 1); qed "commute_rtrancl"; -goalw Commutation.thy [commute_def,square_def] +Goalw [commute_def,square_def] "!!R. [| commute R T; commute S T |] ==> commute (R Un S) T"; by (Blast_tac 1); qed "commute_Un"; (*** diamond, confluence and union ***) -goalw Commutation.thy [diamond_def] +Goalw [diamond_def] "!!R. [| diamond R; diamond S; commute R S |] ==> diamond (R Un S)"; by (REPEAT(ares_tac [commute_Un,commute_sym] 1)); qed "diamond_Un"; -goalw Commutation.thy [diamond_def] "!!R. diamond R ==> confluent (R)"; +Goalw [diamond_def] "!!R. diamond R ==> confluent (R)"; by (etac commute_rtrancl 1); qed "diamond_confluent"; -goalw Commutation.thy [diamond_def] +Goalw [diamond_def] "!!R. square R R (R^=) (R^=) ==> confluent R"; by (fast_tac (claset() addIs [square_rtrancl_reflcl_commute, r_into_rtrancl] addEs [square_subset]) 1); qed "square_reflcl_confluent"; -goal Commutation.thy +Goal "!!R. [| confluent R; confluent S; commute (R^*) (S^*) |] \ \ ==> confluent(R Un S)"; by (rtac (rtrancl_Un_rtrancl RS subst) 1); by (blast_tac (claset() addDs [diamond_Un] addIs [diamond_confluent]) 1); qed "confluent_Un"; -goal Commutation.thy +Goal "!!R.[| diamond(R); T <= R; R <= T^* |] ==> confluent(T)"; by (fast_tac (claset() addIs [diamond_confluent] addDs [rtrancl_subset RS sym] addss simpset()) 1); @@ -86,7 +86,7 @@ (*** Church_Rosser ***) -goalw Commutation.thy [square_def,commute_def,diamond_def,Church_Rosser_def] +Goalw [square_def,commute_def,diamond_def,Church_Rosser_def] "Church_Rosser(R) = confluent(R)"; by (safe_tac HOL_cs); by (blast_tac (HOL_cs addIs diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lambda/Eta.ML --- a/src/HOL/Lambda/Eta.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lambda/Eta.ML Mon Jun 22 17:26:46 1998 +0200 @@ -24,7 +24,7 @@ section "eta, subst and free"; -goal Eta.thy "!i t u. ~free s i --> s[t/i] = s[u/i]"; +Goal "!i t u. ~free s i --> s[t/i] = s[u/i]"; by (dB.induct_tac "s" 1); by (ALLGOALS(simp_tac (addsplit (simpset())))); by (Blast_tac 1); @@ -32,7 +32,7 @@ qed_spec_mp "subst_not_free"; Addsimps [subst_not_free RS eqTrueI]; -goal Eta.thy "!i k. free (lift t k) i = \ +Goal "!i k. free (lift t k) i = \ \ (i < k & free t i | k < i & free t (i-1))"; by (dB.induct_tac "t" 1); by (ALLGOALS(asm_full_simp_tac (addsplit (simpset()) addcongs [conj_cong]))); @@ -43,7 +43,7 @@ qed "free_lift"; Addsimps [free_lift]; -goal Eta.thy "!i k t. free (s[t/k]) i = \ +Goal "!i k t. free (s[t/k]) i = \ \ (free s k & free t i | free s (if i t ==> !i. free t i = free s i"; +Goal "!!s. s -e> t ==> !i. free t i = free s i"; by (etac eta.induct 1); by (ALLGOALS(asm_simp_tac (simpset() addcongs [conj_cong]))); qed_spec_mp "free_eta"; -goal Eta.thy "!!s. [| s -e> t; ~free s i |] ==> ~free t i"; +Goal "!!s. [| s -e> t; ~free s i |] ==> ~free t i"; by (asm_simp_tac (simpset() addsimps [free_eta]) 1); qed "not_free_eta"; -goal Eta.thy "!!s. s -e> t ==> !u i. s[u/i] -e> t[u/i]"; +Goal "!!s. s -e> t ==> !u i. s[u/i] -e> t[u/i]"; by (etac eta.induct 1); by (ALLGOALS(asm_simp_tac (simpset() addsimps [subst_subst RS sym]))); qed_spec_mp "eta_subst"; @@ -73,7 +73,7 @@ section "Confluence of -e>"; -goalw Eta.thy [square_def,id_def] "square eta eta (eta^=) (eta^=)"; +Goalw [square_def,id_def] "square eta eta (eta^=) (eta^=)"; by (rtac (impI RS allI RS allI) 1); by (Simp_tac 1); by (etac eta.induct 1); @@ -84,46 +84,46 @@ by (ALLGOALS Blast_tac); qed "square_eta"; -goal Eta.thy "confluent(eta)"; +Goal "confluent(eta)"; by (rtac (square_eta RS square_reflcl_confluent) 1); qed "eta_confluent"; section "Congruence rules for -e>>"; -goal Eta.thy "!!s. s -e>> s' ==> Abs s -e>> Abs s'"; +Goal "!!s. s -e>> s' ==> Abs s -e>> Abs s'"; by (etac rtrancl_induct 1); by (ALLGOALS(blast_tac (claset() addIs [rtrancl_refl,rtrancl_into_rtrancl]))); qed "rtrancl_eta_Abs"; -goal Eta.thy "!!s. s -e>> s' ==> s @ t -e>> s' @ t"; +Goal "!!s. s -e>> s' ==> s @ t -e>> s' @ t"; by (etac rtrancl_induct 1); by (ALLGOALS(blast_tac (claset() addIs [rtrancl_refl,rtrancl_into_rtrancl]))); qed "rtrancl_eta_AppL"; -goal Eta.thy "!!s. t -e>> t' ==> s @ t -e>> s @ t'"; +Goal "!!s. t -e>> t' ==> s @ t -e>> s @ t'"; by (etac rtrancl_induct 1); by (ALLGOALS(blast_tac (claset() addIs [rtrancl_refl,rtrancl_into_rtrancl]))); qed "rtrancl_eta_AppR"; -goal Eta.thy "!!s. [| s -e>> s'; t -e>> t' |] ==> s @ t -e>> s' @ t'"; +Goal "!!s. [| s -e>> s'; t -e>> t' |] ==> s @ t -e>> s' @ t'"; by (blast_tac (claset() addSIs [rtrancl_eta_AppL, rtrancl_eta_AppR] addIs [rtrancl_trans]) 1); qed "rtrancl_eta_App"; section "Commutation of -> and -e>"; -goal Eta.thy "!!s t. s -> t ==> (!i. free t i --> free s i)"; +Goal "!!s t. s -> t ==> (!i. free t i --> free s i)"; by (etac beta.induct 1); by (ALLGOALS(Asm_full_simp_tac)); qed_spec_mp "free_beta"; -goal Eta.thy "!!s t. s -> t ==> !u i. s[u/i] -> t[u/i]"; +Goal "!!s t. s -> t ==> !u i. s[u/i] -> t[u/i]"; by (etac beta.induct 1); by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [subst_subst RS sym]))); qed_spec_mp "beta_subst"; AddIs [beta_subst]; -goal Eta.thy "!i. t[Var i/i] = t[Var(i)/Suc i]"; +Goal "!i. t[Var i/i] = t[Var(i)/Suc i]"; by (dB.induct_tac "t" 1); by (ALLGOALS (asm_simp_tac (addsplit (simpset())))); by (safe_tac (HOL_cs addSEs [nat_neqE])); @@ -131,13 +131,13 @@ qed_spec_mp "subst_Var_Suc"; Addsimps [subst_Var_Suc]; -goal Eta.thy "!!s t. s -e> t ==> (!i. lift s i -e> lift t i)"; +Goal "!!s t. s -e> t ==> (!i. lift s i -e> lift t i)"; by (etac eta.induct 1); by (ALLGOALS(asm_simp_tac (addsplit (simpset())))); qed_spec_mp "eta_lift"; Addsimps [eta_lift]; -goal Eta.thy "!s t i. s -e> t --> u[s/i] -e>> u[t/i]"; +Goal "!s t i. s -e> t --> u[s/i] -e>> u[t/i]"; by (dB.induct_tac "u" 1); by (ALLGOALS(asm_simp_tac (addsplit (simpset())))); by (blast_tac (claset() addIs [r_into_rtrancl]) 1); @@ -145,7 +145,7 @@ by (blast_tac (claset() addSIs [rtrancl_eta_Abs,eta_lift]) 1); qed_spec_mp "rtrancl_eta_subst"; -goalw Eta.thy [square_def] "square beta eta (eta^*) (beta^=)"; +Goalw [square_def] "square beta eta (eta^*) (beta^=)"; by (rtac (impI RS allI RS allI) 1); by (etac beta.induct 1); by (slow_tac (claset() addIs [r_into_rtrancl,rtrancl_eta_subst,eta_subst] @@ -159,14 +159,14 @@ addss simpset()) 1); qed "square_beta_eta"; -goal Eta.thy "confluent(beta Un eta)"; +Goal "confluent(beta Un eta)"; by (REPEAT(ares_tac [square_rtrancl_reflcl_commute,confluent_Un, beta_confluent,eta_confluent,square_beta_eta] 1)); qed "confluent_beta_eta"; section "Implicit definition of -e>: Abs(lift s 0 @ Var 0) -e> s"; -goal Eta.thy "!i. (~free s i) = (? t. s = lift t i)"; +Goal "!i. (~free s i) = (? t. s = lift t i)"; by (dB.induct_tac "s" 1); by (Simp_tac 1); by (SELECT_GOAL(safe_tac HOL_cs)1); @@ -215,7 +215,7 @@ by (Blast_tac 1); qed_spec_mp "not_free_iff_lifted"; -goal Eta.thy "(!s u. (~free s 0) --> R(Abs(s @ Var 0))(s[u/0])) = \ +Goal "(!s u. (~free s 0) --> R(Abs(s @ Var 0))(s[u/0])) = \ \ (!s. R(Abs(lift s 0 @ Var 0))(s))"; by (fast_tac (HOL_cs addss (simpset() addsimps [not_free_iff_lifted])) 1); qed "explicit_is_implicit"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lambda/Lambda.ML --- a/src/HOL/Lambda/Lambda.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lambda/Lambda.ML Mon Jun 22 17:26:46 1998 +0200 @@ -21,23 +21,23 @@ (*** Congruence rules for ->> ***) -goal Lambda.thy "!!s. s ->> s' ==> Abs s ->> Abs s'"; +Goal "!!s. s ->> s' ==> Abs s ->> Abs s'"; by (etac rtrancl_induct 1); by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_rtrancl]))); qed "rtrancl_beta_Abs"; AddSIs [rtrancl_beta_Abs]; -goal Lambda.thy "!!s. s ->> s' ==> s @ t ->> s' @ t"; +Goal "!!s. s ->> s' ==> s @ t ->> s' @ t"; by (etac rtrancl_induct 1); by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_rtrancl]))); qed "rtrancl_beta_AppL"; -goal Lambda.thy "!!s. t ->> t' ==> s @ t ->> s @ t'"; +Goal "!!s. t ->> t' ==> s @ t ->> s @ t'"; by (etac rtrancl_induct 1); by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_rtrancl]))); qed "rtrancl_beta_AppR"; -goal Lambda.thy "!!s. [| s ->> s'; t ->> t' |] ==> s @ t ->> s' @ t'"; +Goal "!!s. [| s ->> s'; t ->> t' |] ==> s @ t ->> s' @ t'"; by (blast_tac (claset() addSIs [rtrancl_beta_AppL, rtrancl_beta_AppR] addIs [rtrancl_trans]) 1); qed "rtrancl_beta_App"; @@ -49,22 +49,22 @@ delsplits [split_if] setloop (split_inside_tac [split_if]); -goal Lambda.thy "(Var k)[u/k] = u"; +Goal "(Var k)[u/k] = u"; by (asm_full_simp_tac(addsplit(simpset())) 1); qed "subst_eq"; -goal Lambda.thy "!!s. i (Var j)[u/i] = Var(j-1)"; +Goal "!!s. i (Var j)[u/i] = Var(j-1)"; by (asm_full_simp_tac(addsplit(simpset())) 1); qed "subst_gt"; -goal Lambda.thy "!!s. j (Var j)[u/i] = Var(j)"; +Goal "!!s. j (Var j)[u/i] = Var(j)"; by (asm_full_simp_tac (addsplit(simpset()) addsimps [less_not_refl2 RS not_sym,less_SucI]) 1); qed "subst_lt"; Addsimps [subst_eq,subst_gt,subst_lt]; -goal Lambda.thy +Goal "!i k. i < Suc k --> lift (lift t i) (Suc k) = lift (lift t k) i"; by (dB.induct_tac "t" 1); by (ALLGOALS(asm_simp_tac (simpset() addSolver cut_trans_tac))); @@ -72,7 +72,7 @@ by (ALLGOALS trans_tac); qed_spec_mp "lift_lift"; -goal Lambda.thy "!i j s. j < Suc i --> \ +Goal "!i j s. j < Suc i --> \ \ lift (t[s/j]) i = (lift t (Suc i)) [lift s i / j]"; by (dB.induct_tac "t" 1); by (ALLGOALS(asm_simp_tac (simpset() addsimps [diff_Suc,subst_Var,lift_lift] @@ -83,7 +83,7 @@ qed "lift_subst"; Addsimps [lift_subst]; -goal Lambda.thy +Goal "!i j s. i < Suc j -->\ \ lift (t[s/j]) i = (lift t i) [lift s i / Suc j]"; by (dB.induct_tac "t" 1); @@ -93,14 +93,14 @@ by (ALLGOALS trans_tac); qed "lift_subst_lt"; -goal Lambda.thy "!k s. (lift t k)[s/k] = t"; +Goal "!k s. (lift t k)[s/k] = t"; by (dB.induct_tac "t" 1); by (ALLGOALS Asm_full_simp_tac); qed "subst_lift"; Addsimps [subst_lift]; -goal Lambda.thy "!i j u v. i < Suc j --> \ +Goal "!i j u v. i < Suc j --> \ \ t[lift v i / Suc j][u[v/j]/i] = t[u/i][v/j]"; by (dB.induct_tac "t" 1); by (ALLGOALS(asm_simp_tac @@ -116,25 +116,25 @@ (*** Equivalence proof for optimized substitution ***) -goal Lambda.thy "!k. liftn 0 t k = t"; +Goal "!k. liftn 0 t k = t"; by (dB.induct_tac "t" 1); by (ALLGOALS(asm_simp_tac(addsplit(simpset())))); qed "liftn_0"; Addsimps [liftn_0]; -goal Lambda.thy "!k. liftn (Suc n) t k = lift (liftn n t k) k"; +Goal "!k. liftn (Suc n) t k = lift (liftn n t k) k"; by (dB.induct_tac "t" 1); by (ALLGOALS(asm_simp_tac(addsplit(simpset())))); by (blast_tac (claset() addDs [add_lessD1]) 1); qed "liftn_lift"; Addsimps [liftn_lift]; -goal Lambda.thy "!n. substn t s n = t[liftn n s 0 / n]"; +Goal "!n. substn t s n = t[liftn n s 0 / n]"; by (dB.induct_tac "t" 1); by (ALLGOALS(asm_simp_tac(addsplit(simpset())))); qed "substn_subst_n"; Addsimps [substn_subst_n]; -goal Lambda.thy "substn t s 0 = t[s/0]"; +Goal "substn t s 0 = t[s/0]"; by (Simp_tac 1); qed "substn_subst_0"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lambda/ParRed.ML --- a/src/HOL/Lambda/ParRed.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lambda/ParRed.ML Mon Jun 22 17:26:46 1998 +0200 @@ -20,26 +20,26 @@ (*** beta <= par_beta <= beta^* ***) -goal ParRed.thy "(Var n => t) = (t = Var n)"; +Goal "(Var n => t) = (t = Var n)"; by (Blast_tac 1); qed "par_beta_varL"; Addsimps [par_beta_varL]; -goal ParRed.thy "t => t"; +Goal "t => t"; by (dB.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed"par_beta_refl"; Addsimps [par_beta_refl]; (* AddSIs [par_beta_refl]; causes search to blow up *) -goal ParRed.thy "beta <= par_beta"; +Goal "beta <= par_beta"; by (rtac subsetI 1); by (split_all_tac 1); by (etac beta.induct 1); by (ALLGOALS(blast_tac (claset() addSIs [par_beta_refl]))); qed "beta_subset_par_beta"; -goal ParRed.thy "par_beta <= beta^*"; +Goal "par_beta <= beta^*"; by (rtac subsetI 1); by (split_all_tac 1); by (etac par_beta.induct 1); @@ -51,13 +51,13 @@ (*** => ***) -goal ParRed.thy "!t' n. t => t' --> lift t n => lift t' n"; +Goal "!t' n. t => t' --> lift t n => lift t' n"; by (dB.induct_tac "t" 1); by (ALLGOALS(fast_tac (claset() addss (simpset())))); qed_spec_mp "par_beta_lift"; Addsimps [par_beta_lift]; -goal ParRed.thy +Goal "!s s' t' n. s => s' --> t => t' --> t[s/n] => t'[s'/n]"; by (dB.induct_tac "t" 1); by (asm_simp_tac (addsplit(simpset())) 1); @@ -71,7 +71,7 @@ (*** Confluence (directly) ***) -goalw ParRed.thy [diamond_def,commute_def,square_def] "diamond(par_beta)"; +Goalw [diamond_def,commute_def,square_def] "diamond(par_beta)"; by (rtac (impI RS allI RS allI) 1); by (etac par_beta.induct 1); by (ALLGOALS(blast_tac (claset() addSIs [par_beta_subst]))); @@ -80,7 +80,7 @@ (*** cd ***) -goal ParRed.thy "!t. s => t --> t => cd s"; +Goal "!t. s => t --> t => cd s"; by (dB.induct_tac "s" 1); by (Simp_tac 1); by (etac rev_mp 1); @@ -91,11 +91,11 @@ (*** Confluence (via cd) ***) -goalw ParRed.thy [diamond_def,commute_def,square_def] "diamond(par_beta)"; +Goalw [diamond_def,commute_def,square_def] "diamond(par_beta)"; by (blast_tac (claset() addIs [par_beta_cd]) 1); qed "diamond_par_beta2"; -goal ParRed.thy "confluent(beta)"; +Goal "confluent(beta)"; by (blast_tac (HOL_cs addIs [diamond_par_beta2, diamond_to_confluence, par_beta_subset_beta, beta_subset_par_beta]) 1); qed"beta_confluent"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/AutoChopper.ML --- a/src/HOL/Lex/AutoChopper.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/AutoChopper.ML Mon Jun 22 17:26:46 1998 +0200 @@ -13,12 +13,12 @@ Addsimps [Let_def]; -goalw thy [acc_prefix_def] "~acc_prefix A [] s"; +Goalw [acc_prefix_def] "~acc_prefix A [] s"; by (Simp_tac 1); qed"acc_prefix_Nil"; Addsimps [acc_prefix_Nil]; -goalw thy [acc_prefix_def] +Goalw [acc_prefix_def] "acc_prefix A (x#xs) s = (fin A (next A x s) | acc_prefix A xs (next A x s))"; by (simp_tac (simpset() addsimps [prefix_Cons]) 1); by Safe_tac; @@ -34,14 +34,14 @@ qed"acc_prefix_Cons"; Addsimps [acc_prefix_Cons]; -goal thy "!st us p y ys. acc A st p (ys@[y]) xs us ~= ([],zs)"; +Goal "!st us p y ys. acc A st p (ys@[y]) xs us ~= ([],zs)"; by (induct_tac "xs" 1); by (Simp_tac 1); by (Asm_simp_tac 1); val accept_not_Nil = result() repeat_RS spec; Addsimps [accept_not_Nil]; -goal thy +Goal "!st us. acc A st ([],ys) [] xs us = ([], zs) --> \ \ zs = ys & (!ys. ys ~= [] & ys<=xs --> ~fin A (delta A ys st))"; by (induct_tac "xs" 1); @@ -67,7 +67,7 @@ val ex_special = result(); -goal thy +Goal "! r erk l rst st ys yss zs::'a list. \ \ acc A st (l,rst) erk xs r = (ys#yss, zs) --> \ \ ys@concat(yss)@zs = (if acc_prefix A xs st then r@xs else erk@concat(l)@rst)"; @@ -86,7 +86,7 @@ val step2_a = (result() repeat_RS spec) RS mp; -goal thy +Goal "! st erk r l rest ys yss zs.\ \ acc A st (l,rest) erk xs r = (ys#yss, zs) --> \ \ (if acc_prefix A xs st \ @@ -109,7 +109,7 @@ val step2_b = (result() repeat_RS spec) RS mp; -goal thy +Goal "! st erk r l rest ys yss zs. \ \ acc A st (l,rest) erk xs r = (ys#yss, zs) --> \ \ (if acc_prefix A xs st \ @@ -148,7 +148,7 @@ val step2_c = (result() repeat_RS spec) RS mp; -goal thy +Goal "! st erk r l rest ys yss zs. \ \ acc A st (l,rest) erk xs r = (ys#yss, zs) --> \ \ (if acc_prefix A xs st \ @@ -187,7 +187,7 @@ val step2_d = (result() repeat_RS spec) RS mp; Delsimps [split_paired_All]; -goal thy +Goal "! st erk r p ys yss zs. \ \ acc A st p erk xs r = (ys#yss, zs) --> \ \ (if acc_prefix A xs st \ @@ -238,7 +238,7 @@ val step2_e = (result() repeat_RS spec) RS mp; Addsimps[split_paired_All]; -goalw thy [accepts_def, is_auto_chopper_def, auto_chopper_def, +Goalw [accepts_def, is_auto_chopper_def, auto_chopper_def, Chopper.is_longest_prefix_chopper_def] "is_auto_chopper(auto_chopper)"; by (REPEAT(ares_tac [no_acc,allI,impI,conjI] 1)); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/AutoMaxChop.ML --- a/src/HOL/Lex/AutoMaxChop.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/AutoMaxChop.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,11 +4,11 @@ Copyright 1998 TUM *) -goal thy "delta A (xs@[y]) q = next A y (delta A xs q)"; +Goal "delta A (xs@[y]) q = next A y (delta A xs q)"; by(Simp_tac 1); qed "delta_snoc"; -goal thy +Goal "!q ps res. auto_split A (delta A ps q) res ps xs = \ \ maxsplit (%ys. fin A (delta A ys q)) res ps xs"; by(induct_tac "xs" 1); @@ -17,20 +17,20 @@ delsimps [delta_append]) 1); qed_spec_mp "auto_split_lemma"; -goalw thy [accepts_def] +Goalw [accepts_def] "auto_split A (start A) res [] xs = maxsplit (accepts A) res [] xs"; by(stac ((read_instantiate [("s","start A")] delta_Nil) RS sym) 1); by(stac auto_split_lemma 1); by(Simp_tac 1); qed_spec_mp "auto_split_is_maxsplit"; -goal thy +Goal "is_maxsplitter (accepts A) (%xs. auto_split A (start A) ([],xs) [] xs)"; by(simp_tac (simpset() addsimps [auto_split_is_maxsplit,is_maxsplitter_maxsplit]) 1); qed "is_maxsplitter_auto_split"; -goalw thy [auto_chop_def] +Goalw [auto_chop_def] "is_maxchopper (accepts A) (auto_chop A)"; br is_maxchopper_chop 1; br is_maxsplitter_auto_split 1; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/AutoProj.ML --- a/src/HOL/Lex/AutoProj.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/AutoProj.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,15 +4,15 @@ Copyright 1998 TUM *) -goalw thy [start_def] "start(q,d,f) = q"; +Goalw [start_def] "start(q,d,f) = q"; by(Simp_tac 1); qed "start_conv"; -goalw thy [next_def] "next(q,d,f) = d"; +Goalw [next_def] "next(q,d,f) = d"; by(Simp_tac 1); qed "next_conv"; -goalw thy [fin_def] "fin(q,d,f) = f"; +Goalw [fin_def] "fin(q,d,f) = f"; by(Simp_tac 1); qed "fin_conv"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/Automata.ML --- a/src/HOL/Lex/Automata.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/Automata.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,14 +6,14 @@ (*** Equivalence of NA and DA --- redundant ***) -goalw thy [na2da_def] +Goalw [na2da_def] "!Q. DA.delta (na2da A) w Q = Union(NA.delta A w `` Q)"; by(induct_tac "w" 1); by(ALLGOALS Asm_simp_tac); by(ALLGOALS Blast_tac); qed_spec_mp "DA_delta_is_lift_NA_delta"; -goalw thy [DA.accepts_def,NA.accepts_def] +Goalw [DA.accepts_def,NA.accepts_def] "NA.accepts A w = DA.accepts (na2da A) w"; by(simp_tac (simpset() addsimps [DA_delta_is_lift_NA_delta]) 1); by(simp_tac (simpset() addsimps [na2da_def]) 1); @@ -21,7 +21,7 @@ (*** Direct equivalence of NAe and DA ***) -goalw thy [nae2da_def] +Goalw [nae2da_def] "!Q. (eps A)^* ^^ (DA.delta (nae2da A) w Q) = steps A w ^^ Q"; by(induct_tac "w" 1); by (Simp_tac 1); @@ -29,12 +29,12 @@ by(Blast_tac 1); qed_spec_mp "espclosure_DA_delta_is_steps"; -goalw thy [nae2da_def] +Goalw [nae2da_def] "fin (nae2da A) Q = (? q : (eps A)^* ^^ Q. fin A q)"; by(Simp_tac 1); val lemma = result(); -goalw thy [NAe.accepts_def,DA.accepts_def] +Goalw [NAe.accepts_def,DA.accepts_def] "DA.accepts (nae2da A) w = NAe.accepts A w"; by(simp_tac (simpset() addsimps [lemma,espclosure_DA_delta_is_steps]) 1); by(simp_tac (simpset() addsimps [nae2da_def]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/DA.ML --- a/src/HOL/Lex/DA.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/DA.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,27 +4,27 @@ Copyright 1998 TUM *) -goalw thy [foldl2_def] "foldl2 f [] a = a"; +Goalw [foldl2_def] "foldl2 f [] a = a"; by(Simp_tac 1); qed "foldl2_Nil"; -goalw thy [foldl2_def] "foldl2 f (x#xs) a = foldl2 f xs (f x a)"; +Goalw [foldl2_def] "foldl2 f (x#xs) a = foldl2 f xs (f x a)"; by(Simp_tac 1); qed "foldl2_Cons"; Addsimps [foldl2_Nil,foldl2_Cons]; -goalw thy [delta_def] "delta A [] s = s"; +Goalw [delta_def] "delta A [] s = s"; by(Simp_tac 1); qed "delta_Nil"; -goalw thy [delta_def] "delta A (a#w) s = delta A w (next A a s)"; +Goalw [delta_def] "delta A (a#w) s = delta A w (next A a s)"; by(Simp_tac 1); qed "delta_Cons"; Addsimps [delta_Nil,delta_Cons]; -goal thy "!q ys. delta A (xs@ys) q = delta A ys (delta A xs q)"; +Goal "!q ys. delta A (xs@ys) q = delta A ys (delta A xs q)"; by(induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "delta_append"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/MaxChop.ML --- a/src/HOL/Lex/MaxChop.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/MaxChop.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,7 +6,7 @@ (* Termination of chop *) -goalw thy [reducing_def] "reducing(%qs. maxsplit P ([],qs) [] qs)"; +Goalw [reducing_def] "reducing(%qs. maxsplit P ([],qs) [] qs)"; by(asm_full_simp_tac (simpset() addsimps [maxsplit_eq]) 1); qed "reducing_maxsplit"; @@ -17,7 +17,7 @@ val chopr_rule = let val [chopr_rule] = chopr.rules in lemma RS chopr_rule end; -goalw thy [chop_def] "!!splitf. reducing splitf ==> \ +Goalw [chop_def] "!!splitf. reducing splitf ==> \ \ chop splitf xs = (let (pre,post) = splitf xs \ \ in if pre=[] then ([],xs) \ \ else let (xss,zs) = chop splitf post \ @@ -26,12 +26,12 @@ by(simp_tac (simpset() addsimps [Let_def] addsplits [split_split]) 1); qed "chop_rule"; -goalw thy [is_maxsplitter_def,reducing_def] +Goalw [is_maxsplitter_def,reducing_def] "!!splitf. is_maxsplitter P splitf ==> reducing splitf"; by(Asm_full_simp_tac 1); qed "is_maxsplitter_reducing"; -goal thy "!!P. is_maxsplitter P splitf ==> \ +Goal "!!P. is_maxsplitter P splitf ==> \ \ !yss zs. chop splitf xs = (yss,zs) --> xs = concat yss @ zs"; by(res_inst_tac [("xs","xs")] length_induct 1); by(asm_simp_tac (simpset() delsplits [split_if] @@ -41,7 +41,7 @@ addsplits [split_split]) 1); qed_spec_mp "chop_concat"; -goal thy "!!P. is_maxsplitter P splitf ==> \ +Goal "!!P. is_maxsplitter P splitf ==> \ \ !yss zs. chop splitf xs = (yss,zs) --> (!ys : set yss. ys ~= [])"; by(res_inst_tac [("xs","xs")] length_induct 1); by(asm_simp_tac (simpset() addsimps [chop_rule,is_maxsplitter_reducing] diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/MaxPrefix.ML --- a/src/HOL/Lex/MaxPrefix.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/MaxPrefix.ML Mon Jun 22 17:26:46 1998 +0200 @@ -5,7 +5,7 @@ *) Delsplits [split_if]; -goalw thy [is_maxpref_def] "!(ps::'a list) res. \ +Goalw [is_maxpref_def] "!(ps::'a list) res. \ \ (maxsplit P res ps qs = (xs,ys)) = \ \ (if ? us. us <= qs & P(ps@us) then xs@ys=ps@qs & is_maxpref P xs (ps@qs) \ \ else (xs,ys)=res)"; @@ -49,13 +49,13 @@ qed_spec_mp "maxsplit_lemma"; Addsplits [split_if]; -goalw thy [is_maxpref_def] +Goalw [is_maxpref_def] "!!P. ~(? us. us<=xs & P us) ==> is_maxpref P ps xs = (ps = [])"; by(Blast_tac 1); qed "is_maxpref_Nil"; Addsimps [is_maxpref_Nil]; -goalw thy [is_maxsplitter_def] +Goalw [is_maxsplitter_def] "is_maxsplitter P (%xs. maxsplit P ([],xs) [] xs)"; by(simp_tac (simpset() addsimps [maxsplit_lemma]) 1); by(fast_tac (claset() addss simpset()) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/NAe.ML --- a/src/HOL/Lex/NAe.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/NAe.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,39 +4,39 @@ Copyright 1998 TUM *) -goal thy +Goal "steps A w O (eps A)^* = steps A w"; by (exhaust_tac "w" 1); by(ALLGOALS(asm_simp_tac (simpset() addsimps [O_assoc]))); qed_spec_mp "steps_epsclosure"; Addsimps [steps_epsclosure]; -goal thy +Goal "!! A. [| (p,q) : (eps A)^*; (q,r) : steps A w |] ==> (p,r) : steps A w"; by(rtac (steps_epsclosure RS equalityE) 1); by(Blast_tac 1); qed "in_steps_epsclosure"; -goal thy "(eps A)^* O steps A w = steps A w"; +Goal "(eps A)^* O steps A w = steps A w"; by (induct_tac "w" 1); by (Simp_tac 1); by (asm_simp_tac (simpset() addsimps [O_assoc RS sym]) 1); qed "epsclosure_steps"; -goal thy +Goal "!!A. [| (p,q) : steps A w; (q,r) : (eps A)^* |] ==> (p,r) : steps A w"; by(rtac (epsclosure_steps RS equalityE) 1); by(Blast_tac 1); qed "in_epsclosure_steps"; -goal thy +Goal "steps A (v@w) = steps A w O steps A v"; by (induct_tac "v" 1); by(ALLGOALS(asm_simp_tac (simpset() addsimps [O_assoc]))); qed "steps_append"; Addsimps [steps_append]; -goal thy "(p,r) : steps A (v@w) = ((p,r) : (steps A w O steps A v))"; +Goal "(p,r) : steps A (v@w) = ((p,r) : (steps A w O steps A v))"; by(rtac (steps_append RS equalityE) 1); by(Blast_tac 1); qed "in_steps_append"; @@ -44,7 +44,7 @@ (* Equivalence of steps and delta (* Use "(? x : f `` A. P x) = (? a:A. P(f x))" ?? *) -goal thy "!p. (p,q) : steps A w = (q : delta A w p)"; +Goal "!p. (p,q) : steps A w = (q : delta A w p)"; by (induct_tac "w" 1); by (Simp_tac 1); by (asm_simp_tac (simpset() addsimps [comp_def,step_def]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/Prefix.ML --- a/src/HOL/Lex/Prefix.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/Prefix.ML Mon Jun 22 17:26:46 1998 +0200 @@ -14,36 +14,36 @@ (** <= is a partial order: **) -goalw thy [prefix_def] "xs <= (xs::'a list)"; +Goalw [prefix_def] "xs <= (xs::'a list)"; by(Simp_tac 1); qed "prefix_refl"; AddIffs[prefix_refl]; -goalw thy [prefix_def] "!!xs::'a list. [| xs <= ys; ys <= zs |] ==> xs <= zs"; +Goalw [prefix_def] "!!xs::'a list. [| xs <= ys; ys <= zs |] ==> xs <= zs"; by(Clarify_tac 1); by(Simp_tac 1); qed "prefix_trans"; -goalw thy [prefix_def] "!!xs::'a list. [| xs <= ys; ys <= xs |] ==> xs = ys"; +Goalw [prefix_def] "!!xs::'a list. [| xs <= ys; ys <= xs |] ==> xs = ys"; by(Clarify_tac 1); by(Asm_full_simp_tac 1); qed "prefix_antisym"; (** recursion equations **) -goalw Prefix.thy [prefix_def] "[] <= xs"; +Goalw [prefix_def] "[] <= xs"; by (simp_tac (simpset() addsimps [eq_sym_conv]) 1); qed "Nil_prefix"; AddIffs[Nil_prefix]; -goalw Prefix.thy [prefix_def] "(xs <= []) = (xs = [])"; +Goalw [prefix_def] "(xs <= []) = (xs = [])"; by (list.induct_tac "xs" 1); by (Simp_tac 1); by (Simp_tac 1); qed "prefix_Nil"; Addsimps [prefix_Nil]; -goalw thy [prefix_def] "(xs <= ys@[y]) = (xs = ys@[y] | xs <= ys)"; +Goalw [prefix_def] "(xs <= ys@[y]) = (xs = ys@[y] | xs <= ys)"; br iffI 1; be exE 1; by(rename_tac "zs" 1); @@ -59,13 +59,13 @@ qed "prefix_snoc"; Addsimps [prefix_snoc]; -goalw Prefix.thy [prefix_def] "(x#xs <= y#ys) = (x=y & xs<=ys)"; +Goalw [prefix_def] "(x#xs <= y#ys) = (x=y & xs<=ys)"; by (Simp_tac 1); by (Fast_tac 1); qed"Cons_prefix_Cons"; Addsimps [Cons_prefix_Cons]; -goal thy "(xs@ys <= xs@zs) = (ys <= zs)"; +Goal "(xs@ys <= xs@zs) = (ys <= zs)"; by (induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); qed "same_prefix_prefix"; @@ -74,14 +74,14 @@ AddIffs [simplify (simpset()) (read_instantiate [("zs","[]")] same_prefix_prefix)]; -goalw thy [prefix_def] "!!xs. xs <= ys ==> xs <= ys@zs"; +Goalw [prefix_def] "!!xs. xs <= ys ==> xs <= ys@zs"; by(Clarify_tac 1); by (Simp_tac 1); qed "prefix_prefix"; Addsimps [prefix_prefix]; (* nicht sehr elegant bewiesen - Induktion eigentlich ueberfluessig *) -goalw Prefix.thy [prefix_def] +Goalw [prefix_def] "(xs <= y#ys) = (xs=[] | (? zs. xs=y#zs & zs <= ys))"; by (list.induct_tac "xs" 1); by (Simp_tac 1); @@ -89,7 +89,7 @@ by (Fast_tac 1); qed "prefix_Cons"; -goal thy "(xs <= ys@zs) = (xs <= ys | (? us. xs = ys@us & us <= zs))"; +Goal "(xs <= ys@zs) = (xs <= ys | (? us. xs = ys@us & us <= zs))"; by(res_inst_tac [("xs","zs")] rev_induct 1); by(Simp_tac 1); by(Blast_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/RegExp2NAe.ML --- a/src/HOL/Lex/RegExp2NAe.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/RegExp2NAe.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,37 +8,37 @@ (* atom *) (******************************************************) -goalw thy [atom_def] "(fin (atom a) q) = (q = [False])"; +Goalw [atom_def] "(fin (atom a) q) = (q = [False])"; by(Simp_tac 1); qed "fin_atom"; -goalw thy [atom_def] "start (atom a) = [True]"; +Goalw [atom_def] "start (atom a) = [True]"; by(Simp_tac 1); qed "start_atom"; (* Use {x. False} = {}? *) -goalw thy [atom_def,step_def] +Goalw [atom_def,step_def] "eps(atom a) = {}"; by(Simp_tac 1); by (Blast_tac 1); qed "eps_atom"; Addsimps [eps_atom]; -goalw thy [atom_def,step_def] +Goalw [atom_def,step_def] "(p,q) : step (atom a) (Some b) = (p=[True] & q=[False] & b=a)"; by(Simp_tac 1); qed "in_step_atom_Some"; Addsimps [in_step_atom_Some]; -goal thy +Goal "([False],[False]) : steps (atom a) w = (w = [])"; by (induct_tac "w" 1); by(Simp_tac 1); by(asm_simp_tac (simpset() addsimps [comp_def]) 1); qed "False_False_in_steps_atom"; -goal thy +Goal "(start (atom a), [False]) : steps (atom a) w = (w = [a])"; by (induct_tac "w" 1); by(asm_simp_tac (simpset() addsimps [start_atom,rtrancl_empty]) 1); @@ -46,7 +46,7 @@ addsimps [False_False_in_steps_atom,comp_def,start_atom]) 1); qed "start_fin_in_steps_atom"; -goal thy +Goal "accepts (atom a) w = (w = [a])"; by(simp_tac(simpset() addsimps [accepts_def,start_fin_in_steps_atom,fin_atom]) 1); @@ -59,12 +59,12 @@ (***** True/False ueber fin anheben *****) -goalw thy [union_def] +Goalw [union_def] "!L R. fin (union L R) (True#p) = fin L p"; by (Simp_tac 1); qed_spec_mp "fin_union_True"; -goalw thy [union_def] +Goalw [union_def] "!L R. fin (union L R) (False#p) = fin R p"; by (Simp_tac 1); qed_spec_mp "fin_union_False"; @@ -73,13 +73,13 @@ (***** True/False ueber step anheben *****) -goalw thy [union_def,step_def] +Goalw [union_def,step_def] "!L R. (True#p,q) : step (union L R) a = (? r. q = True#r & (p,r) : step L a)"; by (Simp_tac 1); by(Blast_tac 1); qed_spec_mp "True_in_step_union"; -goalw thy [union_def,step_def] +Goalw [union_def,step_def] "!L R. (False#p,q) : step (union L R) a = (? r. q = False#r & (p,r) : step R a)"; by (Simp_tac 1); by(Blast_tac 1); @@ -89,7 +89,7 @@ (***** True/False ueber epsclosure anheben *****) -goal thy +Goal "!!d. (tp,tq) : (eps(union L R))^* ==> \ \ !p. tp = True#p --> (? q. (p,q) : (eps L)^* & tq = True#q)"; be rtrancl_induct 1; @@ -99,7 +99,7 @@ by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma1a = result(); -goal thy +Goal "!!d. (tp,tq) : (eps(union L R))^* ==> \ \ !p. tp = False#p --> (? q. (p,q) : (eps R)^* & tq = False#q)"; be rtrancl_induct 1; @@ -109,26 +109,26 @@ by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma1b = result(); -goal thy +Goal "!!p. (p,q) : (eps L)^* ==> (True#p, True#q) : (eps(union L R))^*"; be rtrancl_induct 1; by(Blast_tac 1); by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma2a = result(); -goal thy +Goal "!!p. (p,q) : (eps R)^* ==> (False#p, False#q) : (eps(union L R))^*"; be rtrancl_induct 1; by(Blast_tac 1); by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma2b = result(); -goal thy +Goal "(True#p,q) : (eps(union L R))^* = (? r. q = True#r & (p,r) : (eps L)^*)"; by(blast_tac (claset() addDs [lemma1a,lemma2a]) 1); qed "True_epsclosure_union"; -goal thy +Goal "(False#p,q) : (eps(union L R))^* = (? r. q = False#r & (p,r) : (eps R)^*)"; by(blast_tac (claset() addDs [lemma1b,lemma2b]) 1); qed "False_epsclosure_union"; @@ -137,7 +137,7 @@ (***** True/False ueber steps anheben *****) -goal thy +Goal "!p. (True#p,q):steps (union L R) w = (? r. q = True # r & (p,r):steps L w)"; by (induct_tac "w" 1); by (ALLGOALS Asm_simp_tac); @@ -145,7 +145,7 @@ by(Fast_tac 1); qed_spec_mp "lift_True_over_steps_union"; -goal thy +Goal "!p. (False#p,q):steps (union L R) w = (? r. q = False#r & (p,r):steps R w)"; by (induct_tac "w" 1); by (ALLGOALS Asm_simp_tac); @@ -158,7 +158,7 @@ (***** Epsilonhuelle des Startzustands *****) -goal thy +Goal "R^* = id Un (R^* O R)"; by(rtac set_ext 1); by(split_all_tac 1); @@ -169,7 +169,7 @@ by(blast_tac (claset() addIs [rtrancl_into_rtrancl2]) 1); qed "unfold_rtrancl2"; -goal thy +Goal "(p,q) : R^* = (q = p | (? r. (p,r) : R & (r,q) : R^*))"; by(rtac (unfold_rtrancl2 RS equalityE) 1); by(Blast_tac 1); @@ -179,20 +179,20 @@ read_instantiate [("p","start(union L R)")] in_unfold_rtrancl2; AddIffs [epsclosure_start_step_union]; -goalw thy [union_def,step_def] +Goalw [union_def,step_def] "!L R. (start(union L R),q) : eps(union L R) = \ \ (q = True#start L | q = False#start R)"; by(Simp_tac 1); qed_spec_mp "start_eps_union"; AddIffs [start_eps_union]; -goalw thy [union_def,step_def] +Goalw [union_def,step_def] "!L R. (start(union L R),q) ~: step (union L R) (Some a)"; by(Simp_tac 1); qed_spec_mp "not_start_step_union_Some"; AddIffs [not_start_step_union_Some]; -goal thy +Goal "(start(union L R), q) : steps (union L R) w = \ \ ( (w = [] & q = start(union L R)) | \ \ (? p. q = True # p & (start L,p) : steps L w | \ @@ -238,13 +238,13 @@ ba 1; qed "steps_union"; -goalw thy [union_def] +Goalw [union_def] "!L R. ~ fin (union L R) (start(union L R))"; by(Simp_tac 1); qed_spec_mp "start_union_not_final"; AddIffs [start_union_not_final]; -goalw thy [accepts_def] +Goalw [accepts_def] "accepts (union L R) w = (accepts L w | accepts R w)"; by (simp_tac (simpset() addsimps [steps_union]) 1); auto(); @@ -257,12 +257,12 @@ (** True/False in fin **) -goalw thy [conc_def] +Goalw [conc_def] "!L R. fin (conc L R) (True#p) = False"; by (Simp_tac 1); qed_spec_mp "fin_conc_True"; -goalw thy [conc_def] +Goalw [conc_def] "!L R. fin (conc L R) (False#p) = fin R p"; by (Simp_tac 1); qed "fin_conc_False"; @@ -271,7 +271,7 @@ (** True/False in step **) -goalw thy [conc_def,step_def] +Goalw [conc_def,step_def] "!L R. (True#p,q) : step (conc L R) a = \ \ ((? r. q=True#r & (p,r): step L a) | \ \ (fin L p & a=None & q=False#start R))"; @@ -279,7 +279,7 @@ by(Blast_tac 1); qed_spec_mp "True_step_conc"; -goalw thy [conc_def,step_def] +Goalw [conc_def,step_def] "!L R. (False#p,q) : step (conc L R) a = \ \ (? r. q = False#r & (p,r) : step R a)"; by (Simp_tac 1); @@ -290,7 +290,7 @@ (** False in epsclosure **) -goal thy +Goal "!!d. (tp,tq) : (eps(conc L R))^* ==> \ \ !p. tp = False#p --> (? q. (p,q) : (eps R)^* & tq = False#q)"; by(etac rtrancl_induct 1); @@ -298,14 +298,14 @@ by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); qed "lemma1b"; -goal thy +Goal "!!p. (p,q) : (eps R)^* ==> (False#p, False#q) : (eps(conc L R))^*"; by(etac rtrancl_induct 1); by(Blast_tac 1); by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma2b = result(); -goal thy +Goal "((False # p, q) : (eps (conc L R))^*) = \ \ (? r. q = False # r & (p, r) : (eps R)^*)"; by (rtac iffI 1); @@ -316,7 +316,7 @@ (** False in steps **) -goal thy +Goal "!p. (False#p,q): steps (conc L R) w = (? r. q=False#r & (p,r): steps R w)"; by (induct_tac "w" 1); by (Simp_tac 1); @@ -328,14 +328,14 @@ (** True in epsclosure **) -goal thy +Goal "!!L R. (p,q): (eps L)^* ==> (True#p,True#q) : (eps(conc L R))^*"; be rtrancl_induct 1; by(Blast_tac 1); by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); qed "True_True_eps_concI"; -goal thy +Goal "!!L R. !p. (p,q) : steps L w --> (True#p,True#q) : steps (conc L R) w"; by(induct_tac "w" 1); by (simp_tac (simpset() addsimps [True_True_eps_concI]) 1); @@ -343,7 +343,7 @@ by(blast_tac (claset() addIs [True_True_eps_concI]) 1); qed_spec_mp "True_True_steps_concI"; -goal thy +Goal "!!d. (tp,tq) : (eps(conc L R))^* ==> \ \ !p. tp = True#p --> \ \ (? q. tq = True#q & (p,q) : (eps L)^*) | \ @@ -353,20 +353,20 @@ by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma1a = result(); -goal thy +Goal "!!p. (p, q) : (eps L)^* ==> (True#p, True#q) : (eps(conc L R))^*"; by(etac rtrancl_induct 1); by(Blast_tac 1); by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma2a = result(); -goalw thy [conc_def,step_def] +Goalw [conc_def,step_def] "!!L R. (p,q) : step R None ==> (False#p, False#q) : step (conc L R) None"; by(split_all_tac 1); by (Asm_full_simp_tac 1); val lemma = result(); -goal thy +Goal "!!L R. (p,q) : (eps R)^* ==> (False#p, False#q) : (eps(conc L R))^*"; by(etac rtrancl_induct 1); by(Blast_tac 1); @@ -374,13 +374,13 @@ by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma2b = result(); -goalw thy [conc_def,step_def] +Goalw [conc_def,step_def] "!!L R. fin L p ==> (True#p, False#start R) : eps(conc L R)"; by(split_all_tac 1); by(Asm_full_simp_tac 1); qed "True_False_eps_concI"; -goal thy +Goal "((True#p,q) : (eps(conc L R))^*) = \ \ ((? r. (p,r) : (eps L)^* & q = True#r) | \ \ (? r. (p,r) : (eps L)^* & fin L r & \ @@ -400,7 +400,7 @@ (** True in steps **) -goal thy +Goal "!p. (True#p,q) : steps (conc L R) w --> \ \ ((? r. (p,r) : steps L w & q = True#r) | \ \ (? u v. w = u@v & (? r. (p,r) : steps L u & fin L r & \ @@ -431,7 +431,7 @@ by (Blast_tac 1); qed_spec_mp "True_steps_concD"; -goal thy +Goal "(True#p,q) : steps (conc L R) w = \ \ ((? r. (p,r) : steps L w & q = True#r) | \ \ (? u v. w = u@v & (? r. (p,r) : steps L u & fin L r & \ @@ -442,17 +442,17 @@ (** starting from the start **) -goalw thy [conc_def] +Goalw [conc_def] "!L R. start(conc L R) = True#start L"; by(Simp_tac 1); qed_spec_mp "start_conc"; -goalw thy [conc_def] +Goalw [conc_def] "!L R. fin(conc L R) p = (? s. p = False#s & fin R s)"; by (simp_tac (simpset() addsplits [split_list_case]) 1); qed_spec_mp "final_conc"; -goal thy +Goal "accepts (conc L R) w = (? u v. w = u@v & accepts L u & accepts R v)"; by (simp_tac (simpset() addsimps [accepts_def,True_steps_conc,final_conc,start_conc]) 1); @@ -463,7 +463,7 @@ (* star *) (******************************************************) -goalw thy [star_def,step_def] +Goalw [star_def,step_def] "!A. (True#p,q) : eps(star A) = \ \ ( (? r. q = True#r & (p,r) : eps A) | (fin A p & q = True#start A) )"; by(Simp_tac 1); @@ -471,24 +471,24 @@ qed_spec_mp "True_in_eps_star"; AddIffs [True_in_eps_star]; -goalw thy [star_def,step_def] +Goalw [star_def,step_def] "!A. (p,q) : step A a --> (True#p, True#q) : step (star A) a"; by(Simp_tac 1); qed_spec_mp "True_True_step_starI"; -goal thy +Goal "!!A. (p,r) : (eps A)^* ==> (True#p, True#r) : (eps(star A))^*"; be rtrancl_induct 1; by(Blast_tac 1); by(blast_tac (claset() addIs [True_True_step_starI,rtrancl_into_rtrancl]) 1); qed_spec_mp "True_True_eps_starI"; -goalw thy [star_def,step_def] +Goalw [star_def,step_def] "!A. fin A p --> (True#p,True#start A) : eps(star A)"; by(Simp_tac 1); qed_spec_mp "True_start_eps_starI"; -goal thy +Goal "!!dummy. (tp,s) : (eps(star A))^* ==> (! p. tp = True#p --> \ \ (? r. ((p,r) : (eps A)^* | \ \ (? q. (p,q) : (eps A)^* & fin A q & (start A,r) : (eps A)^*)) & \ @@ -500,7 +500,7 @@ by(blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); val lemma = result(); -goal thy +Goal "((True#p,s) : (eps(star A))^*) = \ \ (? r. ((p,r) : (eps A)^* | \ \ (? q. (p,q) : (eps A)^* & fin A q & (start A,r) : (eps A)^*)) & \ @@ -524,7 +524,7 @@ (** True in step Some **) -goalw thy [star_def,step_def] +Goalw [star_def,step_def] "!A. (True#p,r): step (star A) (Some a) = \ \ (? q. (p,q): step A (Some a) & r=True#q)"; by(Simp_tac 1); @@ -536,7 +536,7 @@ (** True in steps **) (* reverse list induction! Complicates matters for conc? *) -goal thy +Goal "!rr. (True#start A,rr) : steps (star A) w --> \ \ (? us v. w = concat us @ v & \ \ (!u:set us. accepts A u) & \ @@ -565,14 +565,14 @@ by(Blast_tac 1); qed_spec_mp "True_start_steps_starD"; -goal thy "!p. (p,q) : steps A w --> (True#p,True#q) : steps (star A) w"; +Goal "!p. (p,q) : steps A w --> (True#p,True#q) : steps (star A) w"; by(induct_tac "w" 1); by(Simp_tac 1); by(Simp_tac 1); by(blast_tac (claset() addIs [True_True_eps_starI,True_True_step_starI]) 1); qed_spec_mp "True_True_steps_starI"; -goalw thy [accepts_def] +Goalw [accepts_def] "(!u : set us. accepts A u) --> \ \ (True#start A,True#start A) : steps (star A) (concat us)"; by(induct_tac "us" 1); @@ -582,7 +582,7 @@ qed_spec_mp "steps_star_cycle"; (* Better stated directly with start(star A)? Loop in star A back to start(star A)?*) -goal thy +Goal "(True#start A,rr) : steps (star A) w = \ \ (? us v. w = concat us @ v & \ \ (!u:set us. accepts A u) & \ @@ -596,7 +596,7 @@ (** the start state **) -goalw thy [star_def,step_def] +Goalw [star_def,step_def] "!A. (start(star A),r) : step (star A) a = (a=None & r = True#start A)"; by(Simp_tac 1); qed_spec_mp "start_step_star"; @@ -605,7 +605,7 @@ val epsclosure_start_step_star = read_instantiate [("p","start(star A)")] in_unfold_rtrancl2; -goal thy +Goal "(start(star A),r) : steps (star A) w = \ \ ((w=[] & r= start(star A)) | (True#start A,r) : steps (star A) w)"; br iffI 1; @@ -622,18 +622,18 @@ by(blast_tac (claset() addIs [in_steps_epsclosure,r_into_rtrancl]) 1); qed "start_steps_star"; -goalw thy [star_def] "!A. fin (star A) (True#p) = fin A p"; +Goalw [star_def] "!A. fin (star A) (True#p) = fin A p"; by(Simp_tac 1); qed_spec_mp "fin_star_True"; AddIffs [fin_star_True]; -goalw thy [star_def] "!A. fin (star A) (start(star A))"; +Goalw [star_def] "!A. fin (star A) (start(star A))"; by(Simp_tac 1); qed_spec_mp "fin_star_start"; AddIffs [fin_star_start]; (* too complex! Simpler if loop back to start(star A)? *) -goalw thy [accepts_def] +Goalw [accepts_def] "accepts (star A) w = \ \ (? us. (!u : set(us). accepts A u) & (w = concat us) )"; by(simp_tac (simpset() addsimps [start_steps_star,True_start_steps_star]) 1); @@ -660,7 +660,7 @@ (***** Correctness of r2n *****) -goal thy +Goal "!w. accepts (rexp2nae r) w = (w : lang r)"; by(induct_tac "r" 1); by(simp_tac (simpset() addsimps [accepts_def]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/RegSet.ML --- a/src/HOL/Lex/RegSet.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/RegSet.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,12 +8,12 @@ Addsimps [star.ConsI]; AddIs [star.ConsI]; -goal thy "(!xs: set xss. xs:S) --> concat xss : star S"; +Goal "(!xs: set xss. xs:S) --> concat xss : star S"; by (induct_tac "xss" 1); by (ALLGOALS Asm_full_simp_tac); qed_spec_mp "concat_in_star"; -goal thy +Goal "w : star U = (? us. (!u : set(us). u : U) & (w = concat us))"; br iffI 1; be star.induct 1; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/RegSet_of_nat_DA.ML --- a/src/HOL/Lex/RegSet_of_nat_DA.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/RegSet_of_nat_DA.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,13 +10,13 @@ (* Lists *) -goal thy "(butlast xs = []) = (case xs of [] => True | y#ys => ys=[])"; +Goal "(butlast xs = []) = (case xs of [] => True | y#ys => ys=[])"; by (exhaust_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "butlast_empty"; AddIffs [butlast_empty]; -goal thy "x:set(butlast xs) --> xs:set xss --> x:set(butlast(concat xss))"; +Goal "x:set(butlast xs) --> xs:set xss --> x:set(butlast(concat xss))"; by (induct_tac "xss" 1); by (Simp_tac 1); by (asm_simp_tac (simpset() addsimps [butlast_append] delsimps ball_simps) 1); @@ -37,7 +37,7 @@ (* The main lemma: how to decompose a trace into a prefix, a list of loops and a suffix. *) -goal thy "!i. k : set(trace d i xs) --> (? pref mids suf. \ +Goal "!i. k : set(trace d i xs) --> (? pref mids suf. \ \ xs = pref @ concat mids @ suf & \ \ deltas d pref i = k & (!n:set(butlast(trace d i pref)). n ~= k) & \ \ (!mid:set mids. (deltas d mid k = k) & \ @@ -77,38 +77,38 @@ by (Asm_simp_tac 1); qed_spec_mp "decompose"; -goal thy "!i. length(trace d i xs) = length xs"; +Goal "!i. length(trace d i xs) = length xs"; by (induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "length_trace"; Addsimps [length_trace]; -goal thy "!i. deltas d (xs@ys) i = deltas d ys (deltas d xs i)"; +Goal "!i. deltas d (xs@ys) i = deltas d ys (deltas d xs i)"; by (induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "deltas_append"; Addsimps [deltas_append]; -goal thy "!i. trace d i (xs@ys) = trace d i xs @ trace d (deltas d xs i) ys"; +Goal "!i. trace d i (xs@ys) = trace d i xs @ trace d (deltas d xs i) ys"; by (induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "trace_append"; Addsimps [trace_append]; -goal thy "(!xs: set xss. deltas d xs i = i) --> \ +Goal "(!xs: set xss. deltas d xs i = i) --> \ \ trace d i (concat xss) = concat (map (trace d i) xss)"; by (induct_tac "xss" 1); by (ALLGOALS Asm_simp_tac); qed_spec_mp "trace_concat"; Addsimps [trace_concat]; -goal thy "!i. (trace d i xs = []) = (xs = [])"; +Goal "!i. (trace d i xs = []) = (xs = [])"; by (exhaust_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "trace_is_Nil"; Addsimps [trace_is_Nil]; -goal thy "(trace d i xs = n#ns) = \ +Goal "(trace d i xs = n#ns) = \ \ (case xs of [] => False | y#ys => n = d y i & ns = trace d n ys)"; by (exhaust_tac "xs" 1); by (ALLGOALS Asm_simp_tac); @@ -116,14 +116,14 @@ qed_spec_mp "trace_is_Cons_conv"; Addsimps [trace_is_Cons_conv]; -goal thy "!i. set(trace d i xs) = \ +Goal "!i. set(trace d i xs) = \ \ (if xs=[] then {} else insert(deltas d xs i)(set(butlast(trace d i xs))))"; by (induct_tac "xs" 1); by (Simp_tac 1); by (asm_simp_tac (simpset() addsimps [insert_commute]) 1); qed "set_trace_conv"; -goal thy +Goal "(!mid:set mids. deltas d mid k = k) --> deltas d (concat mids) k = k"; by (induct_tac "mids" 1); by (ALLGOALS Asm_simp_tac); @@ -136,7 +136,7 @@ val lemma = result(); -goal thy +Goal "!i j xs. xs : regset d i j k = \ \ ((!n:set(butlast(trace d i xs)). n < k) & deltas d xs i = j)"; by (induct_tac "k" 1); @@ -186,28 +186,28 @@ by (Asm_simp_tac 1); qed_spec_mp "regset_spec"; -goalw thy [bounded_def] +Goalw [bounded_def] "!!d. bounded d k ==> !i. i < k --> (!n:set(trace d i xs). n < k)"; by (induct_tac "xs" 1); by (ALLGOALS Simp_tac); by (Blast_tac 1); qed_spec_mp "trace_below"; -goal thy "!!d. [| bounded d k; i < k; j < k |] ==> \ +Goal "!!d. [| bounded d k; i < k; j < k |] ==> \ \ regset d i j k = {xs. deltas d xs i = j}"; by (rtac set_ext 1); by (simp_tac (simpset() addsimps [regset_spec]) 1); by (blast_tac (claset() addDs [trace_below,in_set_butlastD]) 1); qed "regset_below"; -goalw thy [bounded_def] +Goalw [bounded_def] "!!d. bounded d k ==> !i. i < k --> deltas d w i < k"; by (induct_tac "w" 1); by (ALLGOALS Simp_tac); by (Blast_tac 1); qed_spec_mp "deltas_below"; -goalw thy [regset_of_DA_def] +Goalw [regset_of_DA_def] "!!d. [| bounded (next A) k; start A < k; j < k |] ==> \ \ w : regset_of_DA A k = accepts A w"; by(asm_simp_tac (simpset() addcongs [conj_cong] addsimps diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Lex/Scanner.ML --- a/src/HOL/Lex/Scanner.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Lex/Scanner.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,7 +4,7 @@ Copyright 1998 TUM *) -goal thy +Goal "DA.accepts (nae2da(rexp2nae r)) w = (w : lang r)"; by(simp_tac (simpset() addsimps [NAe_DA_equiv,accepts_rexp2nae]) 1); qed "accepts_nae2da_rexp2nae"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Map.ML --- a/src/HOL/Map.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Map.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,12 +6,12 @@ Map lemmas *) -goalw thy [empty_def] "empty k = None"; +Goalw [empty_def] "empty k = None"; by (Simp_tac 1); qed "empty_def2"; Addsimps [empty_def2]; -goalw thy [update_def] "(m[a|->b])x = (if x=a then Some b else m x)"; +Goalw [update_def] "(m[a|->b])x = (if x=a then Some b else m x)"; by (Simp_tac 1); qed "update_def2"; Addsimps [update_def2]; @@ -36,12 +36,12 @@ section "++"; -goalw thy [override_def] "m ++ empty = m"; +Goalw [override_def] "m ++ empty = m"; by (Simp_tac 1); qed "override_empty"; Addsimps [override_empty]; -goalw thy [override_def] "empty ++ m = m"; +Goalw [override_def] "empty ++ m = m"; by (Simp_tac 1); by (rtac ext 1); by (split_tac [split_option_case] 1); @@ -49,20 +49,20 @@ qed "empty_override"; Addsimps [empty_override]; -goalw thy [override_def] +Goalw [override_def] "((m ++ n) k = Some x) = (n k = Some x | n k = None & m k = Some x)"; by (simp_tac (simpset() addsplits [split_option_case]) 1); qed_spec_mp "override_Some_iff"; bind_thm ("override_SomeD", standard(override_Some_iff RS iffD1)); -goalw thy [override_def] +Goalw [override_def] "((m ++ n) k = None) = (n k = None & m k = None)"; by (simp_tac (simpset() addsplits [split_option_case]) 1); qed "override_None"; AddIffs [override_None]; -goalw thy [override_def] "map_of(xs@ys) = map_of ys ++ map_of xs"; +Goalw [override_def] "map_of(xs@ys) = map_of ys ++ map_of xs"; by (induct_tac "xs" 1); by (Simp_tac 1); by (rtac ext 1); @@ -70,7 +70,7 @@ qed "map_of_append"; Addsimps [map_of_append]; -goal thy "map_of xs k = Some y --> (k,y):set xs"; +Goal "map_of xs k = Some y --> (k,y):set xs"; by (list.induct_tac "xs" 1); by (Simp_tac 1); by (split_all_tac 1); @@ -79,12 +79,12 @@ section "dom"; -goalw thy [dom_def] "dom empty = {}"; +Goalw [dom_def] "dom empty = {}"; by (Simp_tac 1); qed "dom_empty"; Addsimps [dom_empty]; -goalw thy [dom_def] "dom(m[a|->b]) = insert a (dom m)"; +Goalw [dom_def] "dom(m[a|->b]) = insert a (dom m)"; by (Simp_tac 1); by (Blast_tac 1); qed "dom_update"; @@ -96,19 +96,19 @@ stac (insert_Collect RS sym) 1, Asm_simp_tac 1]); -goalw thy [dom_def] "dom(m++n) = dom n Un dom m"; +Goalw [dom_def] "dom(m++n) = dom n Un dom m"; by (Blast_tac 1); qed "dom_override"; Addsimps [dom_override]; section "ran"; -goalw thy [ran_def] "ran empty = {}"; +Goalw [ran_def] "ran empty = {}"; by (Simp_tac 1); qed "ran_empty"; Addsimps [ran_empty]; -goalw thy [ran_def] "!!X. m a = None ==> ran(m[a|->b]) = insert b (ran m)"; +Goalw [ran_def] "!!X. m a = None ==> ran(m[a|->b]) = insert b (ran m)"; by Auto_tac; by (subgoal_tac "~(aa = a)" 1); by Auto_tac; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/MiniML/Generalize.ML --- a/src/HOL/MiniML/Generalize.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/MiniML/Generalize.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,12 +6,12 @@ AddSEs [equalityE]; -goal thy "!!A B. free_tv A = free_tv B ==> gen A t = gen B t"; +Goal "!!A B. free_tv A = free_tv B ==> gen A t = gen B t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "gen_eq_on_free_tv"; -goal thy "!!t.(free_tv t) <= (free_tv sch) --> gen sch t = (mk_scheme t)"; +Goal "!!t.(free_tv t) <= (free_tv sch) --> gen sch t = (mk_scheme t)"; by (typ.induct_tac "t" 1); by (Asm_simp_tac 1); by (Simp_tac 1); @@ -20,7 +20,7 @@ Addsimps [gen_without_effect]; -goal thy "!!A. free_tv (gen ($ S A) t) = free_tv t Int free_tv ($ S A)"; +Goal "!!A. free_tv (gen ($ S A) t) = free_tv t Int free_tv ($ S A)"; by (typ.induct_tac "t" 1); by (Simp_tac 1); by (case_tac "nat : free_tv ($ S A)" 1); @@ -34,14 +34,14 @@ Addsimps [free_tv_gen]; -goal thy "!!A. free_tv (gen ($ S A) t # $ S A) = free_tv ($ S A)"; +Goal "!!A. free_tv (gen ($ S A) t # $ S A) = free_tv ($ S A)"; by (Simp_tac 1); by (Fast_tac 1); qed "free_tv_gen_cons"; Addsimps [free_tv_gen_cons]; -goal thy "!!A. bound_tv (gen A t1) = (free_tv t1) - (free_tv A)"; +Goal "!!A. bound_tv (gen A t1) = (free_tv t1) - (free_tv A)"; by (typ.induct_tac "t1" 1); by (Simp_tac 1); by (case_tac "nat : free_tv A" 1); @@ -54,7 +54,7 @@ Addsimps [bound_tv_gen]; -goal thy "!!A. new_tv n t --> new_tv n (gen A t)"; +Goal "!!A. new_tv n t --> new_tv n (gen A t)"; by (typ.induct_tac "t" 1); by (Simp_tac 1); by (case_tac "nat : free_tv A" 1); @@ -62,13 +62,13 @@ by (Asm_simp_tac 1); qed_spec_mp "new_tv_compatible_gen"; -goalw thy [gen_ML_def] "!!A. gen A t = gen_ML A t"; +Goalw [gen_ML_def] "!!A. gen A t = gen_ML A t"; by (typ.induct_tac "t" 1); by (simp_tac (simpset() addsimps [free_tv_ML_scheme_list]) 1); by (asm_simp_tac (simpset() addsimps [free_tv_ML_scheme_list]) 1); qed "gen_eq_gen_ML"; -goal thy "!!S. (free_tv S) Int ((free_tv t) - (free_tv A)) = {} \ +Goal "!!S. (free_tv S) Int ((free_tv t) - (free_tv A)) = {} \ \ --> gen ($ S A) ($ S t) = $ S (gen A t)"; by (induct_tac "t" 1); by (strip_tac 1); @@ -84,14 +84,14 @@ by (Blast_tac 1); qed_spec_mp "gen_subst_commutes"; -goal Generalize.thy "free_tv(t::typ) <= free_tv(A) --> bound_typ_inst S (gen A t) = t"; +Goal "free_tv(t::typ) <= free_tv(A) --> bound_typ_inst S (gen A t) = t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); by (Fast_tac 1); qed_spec_mp "bound_typ_inst_gen"; Addsimps [bound_typ_inst_gen]; -goalw Generalize.thy [le_type_scheme_def,is_bound_typ_instance] +Goalw [le_type_scheme_def,is_bound_typ_instance] "gen ($ S A) ($ S t) <= $ S (gen A t)"; by Safe_tac; by (rename_tac "R" 1); @@ -101,7 +101,7 @@ by (Asm_simp_tac 1); qed "gen_bound_typ_instance"; -goalw Generalize.thy [le_type_scheme_def,is_bound_typ_instance] +Goalw [le_type_scheme_def,is_bound_typ_instance] "!!A B. free_tv B <= free_tv A ==> gen A t <= gen B t"; by Safe_tac; by (rename_tac "S" 1); @@ -112,7 +112,7 @@ by (Asm_simp_tac 1); qed "free_tv_subset_gen_le"; -goalw thy [le_type_scheme_def,is_bound_typ_instance] +Goalw [le_type_scheme_def,is_bound_typ_instance] "!!A. new_tv n A --> gen A t <= gen A ($ (%x. TVar (if x : free_tv A then x else n + x)) t)"; by (strip_tac 1); by (etac exE 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/MiniML/Instance.ML --- a/src/HOL/MiniML/Instance.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/MiniML/Instance.ML Mon Jun 22 17:26:46 1998 +0200 @@ -9,21 +9,21 @@ (* lemmatas for bound_typ_inst *) -goal thy "bound_typ_inst S (mk_scheme t) = t"; +Goal "bound_typ_inst S (mk_scheme t) = t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "bound_typ_inst_mk_scheme"; Addsimps [bound_typ_inst_mk_scheme]; -goal thy "!!S. bound_typ_inst ($S o R) ($S sch) = $S (bound_typ_inst R sch)"; +Goal "!!S. bound_typ_inst ($S o R) ($S sch) = $S (bound_typ_inst R sch)"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS Asm_full_simp_tac); qed "bound_typ_inst_composed_subst"; Addsimps [bound_typ_inst_composed_subst]; -goal thy "!!S. S = S' ==> sch = sch' ==> bound_typ_inst S sch = bound_typ_inst S' sch'"; +Goal "!!S. S = S' ==> sch = sch' ==> bound_typ_inst S sch = bound_typ_inst S' sch'"; by (Asm_full_simp_tac 1); qed "bound_typ_inst_eq"; @@ -31,7 +31,7 @@ (* lemmatas for bound_scheme_inst *) -goal thy "!!t. bound_scheme_inst B (mk_scheme t) = mk_scheme t"; +Goal "!!t. bound_scheme_inst B (mk_scheme t) = mk_scheme t"; by (typ.induct_tac "t" 1); by (Simp_tac 1); by (Asm_simp_tac 1); @@ -39,14 +39,14 @@ Addsimps [bound_scheme_inst_mk_scheme]; -goal thy "!!S. $S (bound_scheme_inst B sch) = (bound_scheme_inst ($S o B) ($ S sch))"; +Goal "!!S. $S (bound_scheme_inst B sch) = (bound_scheme_inst ($S o B) ($ S sch))"; by (type_scheme.induct_tac "sch" 1); by (Simp_tac 1); by (Simp_tac 1); by (Asm_simp_tac 1); qed "substitution_lemma"; -goal thy "!t. mk_scheme t = bound_scheme_inst B sch --> \ +Goal "!t. mk_scheme t = bound_scheme_inst B sch --> \ \ (? S. !x:bound_tv sch. B x = mk_scheme (S x))"; by (type_scheme.induct_tac "sch" 1); by (Simp_tac 1); @@ -79,7 +79,7 @@ (* lemmatas for subst_to_scheme *) -goal thy "!!sch. new_tv n sch --> subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) \ +Goal "!!sch. new_tv n sch --> subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) \ \ (bound_typ_inst (%k. TVar (k + n)) sch) = sch"; by (type_scheme.induct_tac "sch" 1); by (simp_tac (simpset() addsimps [leD]) 1); @@ -87,12 +87,12 @@ by (Asm_simp_tac 1); qed_spec_mp "subst_to_scheme_inverse"; -goal thy "!!t t'. t = t' ==> subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) t = \ +Goal "!!t t'. t = t' ==> subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) t = \ \ subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) t'"; by (Fast_tac 1); val aux = result (); -goal thy "new_tv n sch --> \ +Goal "new_tv n sch --> \ \ (subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) (bound_typ_inst S sch) = \ \ bound_scheme_inst ((subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k)) o S) sch)"; by (type_scheme.induct_tac "sch" 1); @@ -104,7 +104,7 @@ (* lemmata for <= *) -goalw thy [le_type_scheme_def,is_bound_typ_instance] +Goalw [le_type_scheme_def,is_bound_typ_instance] "!!(sch::type_scheme) sch'. (sch' <= sch) = (? B. sch' = bound_scheme_inst B sch)"; by (rtac iffI 1); by (cut_inst_tac [("sch","sch")] fresh_variable_type_schemes 1); @@ -131,7 +131,7 @@ by (Asm_simp_tac 1); qed "le_type_scheme_def2"; -goalw thy [is_bound_typ_instance] "(mk_scheme t) <= sch = t <| sch"; +Goalw [is_bound_typ_instance] "(mk_scheme t) <= sch = t <| sch"; by (simp_tac (simpset() addsimps [le_type_scheme_def2]) 1); by (rtac iffI 1); by (etac exE 1); @@ -158,7 +158,7 @@ by (Asm_full_simp_tac 1); qed_spec_mp "le_type_eq_is_bound_typ_instance"; -goalw thy [le_env_def] +Goalw [le_env_def] "(sch # A <= sch' # B) = (sch <= (sch'::type_scheme) & A <= B)"; by (Simp_tac 1); by (rtac iffI 1); @@ -175,7 +175,7 @@ qed "le_env_Cons"; AddIffs [le_env_Cons]; -goalw thy [is_bound_typ_instance]"!!t. t <| sch ==> $S t <| $S sch"; +Goalw [is_bound_typ_instance]"!!t. t <| sch ==> $S t <| $S sch"; by (etac exE 1); by (rename_tac "SA" 1); by (hyp_subst_tac 1); @@ -183,40 +183,40 @@ by (Simp_tac 1); qed "is_bound_typ_instance_closed_subst"; -goal thy "!!(sch::type_scheme) sch'. sch' <= sch ==> $S sch' <= $ S sch"; +Goal "!!(sch::type_scheme) sch'. sch' <= sch ==> $S sch' <= $ S sch"; by (asm_full_simp_tac (simpset() addsimps [le_type_scheme_def2]) 1); by (etac exE 1); by (asm_full_simp_tac (simpset() addsimps [substitution_lemma]) 1); by (Fast_tac 1); qed "S_compatible_le_scheme"; -goalw thy [le_env_def,app_subst_list] "!!(A::type_scheme list) A'. A' <= A ==> $S A' <= $ S A"; +Goalw [le_env_def,app_subst_list] "!!(A::type_scheme list) A'. A' <= A ==> $S A' <= $ S A"; by (simp_tac (simpset() addcongs [conj_cong]) 1); by (fast_tac (claset() addSIs [S_compatible_le_scheme]) 1); qed "S_compatible_le_scheme_lists"; -goalw thy [le_type_scheme_def] "!!t.[| t <| sch; sch <= sch' |] ==> t <| sch'"; +Goalw [le_type_scheme_def] "!!t.[| t <| sch; sch <= sch' |] ==> t <| sch'"; by (Fast_tac 1); qed "bound_typ_instance_trans"; -goalw thy [le_type_scheme_def] "sch <= (sch::type_scheme)"; +Goalw [le_type_scheme_def] "sch <= (sch::type_scheme)"; by (Fast_tac 1); qed "le_type_scheme_refl"; AddIffs [le_type_scheme_refl]; -goalw thy [le_env_def] "A <= (A::type_scheme list)"; +Goalw [le_env_def] "A <= (A::type_scheme list)"; by (Fast_tac 1); qed "le_env_refl"; AddIffs [le_env_refl]; -goalw thy [le_type_scheme_def,is_bound_typ_instance] "sch <= BVar n"; +Goalw [le_type_scheme_def,is_bound_typ_instance] "sch <= BVar n"; by (strip_tac 1); by (res_inst_tac [("x","%a. t")]exI 1); by (Simp_tac 1); qed "bound_typ_instance_BVar"; AddIffs [bound_typ_instance_BVar]; -goalw thy [le_type_scheme_def,is_bound_typ_instance] "(sch <= FVar n) = (sch = FVar n)"; +Goalw [le_type_scheme_def,is_bound_typ_instance] "(sch <= FVar n) = (sch = FVar n)"; by (type_scheme.induct_tac "sch" 1); by (Simp_tac 1); by (Simp_tac 1); @@ -226,12 +226,12 @@ qed "le_FVar"; Addsimps [le_FVar]; -goalw thy [le_type_scheme_def,is_bound_typ_instance] "~(FVar n <= sch1 =-> sch2)"; +Goalw [le_type_scheme_def,is_bound_typ_instance] "~(FVar n <= sch1 =-> sch2)"; by (Simp_tac 1); qed "not_FVar_le_Fun"; AddIffs [not_FVar_le_Fun]; -goalw thy [le_type_scheme_def,is_bound_typ_instance] "~(BVar n <= sch1 =-> sch2)"; +Goalw [le_type_scheme_def,is_bound_typ_instance] "~(BVar n <= sch1 =-> sch2)"; by (Simp_tac 1); by (res_inst_tac [("x","TVar n")] exI 1); by (Simp_tac 1); @@ -239,19 +239,19 @@ qed "not_BVar_le_Fun"; AddIffs [not_BVar_le_Fun]; -goalw thy [le_type_scheme_def,is_bound_typ_instance] +Goalw [le_type_scheme_def,is_bound_typ_instance] "!!sch1. (sch1 =-> sch2 <= sch1' =-> sch2') ==> sch1 <= sch1' & sch2 <= sch2'"; by (fast_tac (claset() addss simpset()) 1); qed "Fun_le_FunD"; -goal thy "(sch' <= sch1 =-> sch2) --> (? sch'1 sch'2. sch' = sch'1 =-> sch'2)"; +Goal "(sch' <= sch1 =-> sch2) --> (? sch'1 sch'2. sch' = sch'1 =-> sch'2)"; by (type_scheme.induct_tac "sch'" 1); by (Asm_simp_tac 1); by (Asm_simp_tac 1); by (Fast_tac 1); qed_spec_mp "scheme_le_Fun"; -goal thy "!sch'::type_scheme. sch <= sch' --> free_tv sch' <= free_tv sch"; +Goal "!sch'::type_scheme. sch <= sch' --> free_tv sch' <= free_tv sch"; by (type_scheme.induct_tac "sch" 1); by (rtac allI 1); by (type_scheme.induct_tac "sch'" 1); @@ -273,7 +273,7 @@ by (Fast_tac 1); qed_spec_mp "le_type_scheme_free_tv"; -goal thy "!A::type_scheme list. A <= B --> free_tv B <= free_tv A"; +Goal "!A::type_scheme list. A <= B --> free_tv B <= free_tv A"; by (list.induct_tac "B" 1); by (Simp_tac 1); by (rtac allI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/MiniML/Maybe.ML --- a/src/HOL/MiniML/Maybe.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/MiniML/Maybe.ML Mon Jun 22 17:26:46 1998 +0200 @@ -5,23 +5,23 @@ *) (* constructor laws for option_bind *) -goalw thy [option_bind_def] "option_bind (Some s) f = (f s)"; +Goalw [option_bind_def] "option_bind (Some s) f = (f s)"; by (Simp_tac 1); qed "option_bind_Some"; -goalw thy [option_bind_def] "option_bind None f = None"; +Goalw [option_bind_def] "option_bind None f = None"; by (Simp_tac 1); qed "option_bind_None"; Addsimps [option_bind_Some,option_bind_None]; (* expansion of option_bind *) -goalw thy [option_bind_def] "P(option_bind res f) = \ +Goalw [option_bind_def] "P(option_bind res f) = \ \ ((res = None --> P None) & (!s. res = Some s --> P(f s)))"; by (rtac split_option_case 1); qed "split_option_bind"; -goal thy +Goal "((option_bind m f) = None) = ((m=None) | (? p. m = Some p & f p = None))"; by (simp_tac (simpset() addsplits [split_option_bind]) 1); qed "option_bind_eq_None"; @@ -29,6 +29,6 @@ Addsimps [option_bind_eq_None]; (* auxiliary lemma *) -goal Maybe.thy "(y = Some x) = (Some x = y)"; +Goal "(y = Some x) = (Some x = y)"; by ( simp_tac (simpset() addsimps [eq_sym_conv]) 1); qed "rotate_Some"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/MiniML/MiniML.ML --- a/src/HOL/MiniML/MiniML.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/MiniML/MiniML.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,21 +10,21 @@ Addsimps [is_bound_typ_instance_closed_subst]; -goal thy "!!t::typ. $(%n. if n : (free_tv A) Un (free_tv t) then (S n) else (TVar n)) t = $S t"; +Goal "!!t::typ. $(%n. if n : (free_tv A) Un (free_tv t) then (S n) else (TVar n)) t = $S t"; by (rtac typ_substitutions_only_on_free_variables 1); by (asm_full_simp_tac (simpset() addsimps [Ball_def]) 1); qed "s'_t_equals_s_t"; Addsimps [s'_t_equals_s_t]; -goal thy "!!A::type_scheme list. $(%n. if n : (free_tv A) Un (free_tv t) then (S n) else (TVar n)) A = $S A"; +Goal "!!A::type_scheme list. $(%n. if n : (free_tv A) Un (free_tv t) then (S n) else (TVar n)) A = $S A"; by (rtac scheme_list_substitutions_only_on_free_variables 1); by (asm_full_simp_tac (simpset() addsimps [Ball_def]) 1); qed "s'_a_equals_s_a"; Addsimps [s'_a_equals_s_a]; -goal thy "!!S.($ (%n. if n : (free_tv A) Un (free_tv t) then (S n) else (TVar n)) A |- e :: \ +Goal "!!S.($ (%n. if n : (free_tv A) Un (free_tv t) then (S n) else (TVar n)) A |- e :: \ \ $ (%n. if n : (free_tv A) Un (free_tv t) then (S n) else (TVar n)) t) ==> \ \ $S A |- e :: $S t"; @@ -35,38 +35,38 @@ by (Asm_full_simp_tac 1); qed "replace_s_by_s'"; -goal thy "!!A::type_scheme list. $ (%x. TVar (if x : free_tv A then x else n + x)) A = $ id_subst A"; +Goal "!!A::type_scheme list. $ (%x. TVar (if x : free_tv A then x else n + x)) A = $ id_subst A"; by (rtac scheme_list_substitutions_only_on_free_variables 1); by (asm_full_simp_tac (simpset() addsimps [id_subst_def]) 1); qed "alpha_A'"; -goal thy "!!A::type_scheme list. $ (%x. TVar (if x : free_tv A then x else n + x)) A = A"; +Goal "!!A::type_scheme list. $ (%x. TVar (if x : free_tv A then x else n + x)) A = A"; by (rtac (alpha_A' RS ssubst) 1); by (Asm_full_simp_tac 1); qed "alpha_A"; -goal thy "!!alpha. $ (S o alpha) (t::typ) = $ S ($ (%x. TVar (alpha x)) t)"; +Goal "!!alpha. $ (S o alpha) (t::typ) = $ S ($ (%x. TVar (alpha x)) t)"; by (typ.induct_tac "t" 1); by (ALLGOALS (Asm_full_simp_tac)); qed "S_o_alpha_typ"; -goal thy "!!alpha. $ (%u. (S o alpha) u) (t::typ) = $ S ($ (%x. TVar (alpha x)) t)"; +Goal "!!alpha. $ (%u. (S o alpha) u) (t::typ) = $ S ($ (%x. TVar (alpha x)) t)"; by (typ.induct_tac "t" 1); by (ALLGOALS (Asm_full_simp_tac)); val S_o_alpha_typ' = result (); -goal thy "!!alpha. $ (S o alpha) (sch::type_scheme) = $ S ($ (%x. TVar (alpha x)) sch)"; +Goal "!!alpha. $ (S o alpha) (sch::type_scheme) = $ S ($ (%x. TVar (alpha x)) sch)"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS (Asm_full_simp_tac)); qed "S_o_alpha_type_scheme"; -goal thy "!!alpha. $ (S o alpha) (A::type_scheme list) = $ S ($ (%x. TVar (alpha x)) A)"; +Goal "!!alpha. $ (S o alpha) (A::type_scheme list) = $ S ($ (%x. TVar (alpha x)) A)"; by (list.induct_tac "A" 1); by (ALLGOALS (Asm_full_simp_tac)); by (rtac (rewrite_rule [o_def] S_o_alpha_type_scheme) 1); qed "S_o_alpha_type_scheme_list"; -goal thy "!!A::type_scheme list. \ +Goal "!!A::type_scheme list. \ \ ($ (%n. if n : free_tv A Un free_tv t \ \ then S n else TVar n) \ \ A) = \ @@ -79,14 +79,14 @@ by (rtac refl 1); qed "S'_A_eq_S'_alpha_A"; -goalw thy [free_tv_subst,dom_def] +Goalw [free_tv_subst,dom_def] "!!A. dom (%n. if n : free_tv A Un free_tv t then S n else TVar n) <= \ \ free_tv A Un free_tv t"; by (Simp_tac 1); by (Fast_tac 1); qed "dom_S'"; -goalw thy [free_tv_subst,cod_def,subset_def] +Goalw [free_tv_subst,cod_def,subset_def] "!!(A::type_scheme list) (t::typ). \ \ cod (%n. if n : free_tv A Un free_tv t then S n else TVar n) <= \ \ free_tv ($ S A) Un free_tv ($ S t)"; @@ -99,14 +99,14 @@ addIs [free_tv_of_substitutions_extend_to_types RS subsetD]) 1); qed "cod_S'"; -goalw thy [free_tv_subst] +Goalw [free_tv_subst] "!!(A::type_scheme list) (t::typ). \ \ free_tv (%n. if n : free_tv A Un free_tv t then S n else TVar n) <= \ \ free_tv A Un free_tv ($ S A) Un free_tv t Un free_tv ($ S t)"; by (fast_tac (claset() addDs [dom_S' RS subsetD,cod_S' RS subsetD]) 1); qed "free_tv_S'"; -goal thy "!!t1::typ. \ +Goal "!!t1::typ. \ \ (free_tv ($ (%x. TVar (if x : free_tv A then x else n + x)) t1) - free_tv A) <= \ \ {x. ? y. x = n + y}"; by (typ.induct_tac "t1" 1); @@ -116,7 +116,7 @@ by (Fast_tac 1); qed "free_tv_alpha"; -goal thy "!!(k::nat). n <= n + k"; +Goal "!!(k::nat). n <= n + k"; by (res_inst_tac [("n","k")] nat_induct 1); by (Simp_tac 1); by (Asm_simp_tac 1); @@ -124,7 +124,7 @@ Addsimps [aux_plus]; -goal thy "!!t::typ. new_tv n t ==> {x. ? y. x = n + y} Int free_tv t = {}"; +Goal "!!t::typ. new_tv n t ==> {x. ? y. x = n + y} Int free_tv t = {}"; by Safe_tac; by (cut_facts_tac [aux_plus] 1); by (dtac new_tv_le 1); @@ -134,7 +134,7 @@ by (Fast_tac 1); val new_tv_Int_free_tv_empty_type = result (); -goal thy "!!sch::type_scheme. new_tv n sch ==> {x. ? y. x = n + y} Int free_tv sch = {}"; +Goal "!!sch::type_scheme. new_tv n sch ==> {x. ? y. x = n + y} Int free_tv sch = {}"; by Safe_tac; by (cut_facts_tac [aux_plus] 1); by (dtac new_tv_le 1); @@ -144,7 +144,7 @@ by (Fast_tac 1); val new_tv_Int_free_tv_empty_scheme = result (); -goal thy "!!n. !A::type_scheme list. new_tv n A --> {x. ? y. x = n + y} Int free_tv A = {}"; +Goal "!!n. !A::type_scheme list. new_tv n A --> {x. ? y. x = n + y} Int free_tv A = {}"; by (rtac allI 1); by (list.induct_tac "A" 1); by (Simp_tac 1); @@ -152,7 +152,7 @@ by (fast_tac (claset() addDs [new_tv_Int_free_tv_empty_scheme ]) 1); val new_tv_Int_free_tv_empty_scheme_list = result (); -goalw thy [le_type_scheme_def,is_bound_typ_instance] +Goalw [le_type_scheme_def,is_bound_typ_instance] "!!A. new_tv n A --> gen A t <= gen A ($ (%x. TVar (if x : free_tv A then x else n + x)) t)"; by (strip_tac 1); by (etac exE 1); @@ -174,7 +174,7 @@ AddSIs has_type.intrs; -goal thy "!!e. A |- e::t ==> !B. A <= B --> B |- e::t"; +Goal "!!e. A |- e::t ==> !B. A <= B --> B |- e::t"; by (etac has_type.induct 1); by (simp_tac (simpset() addsimps [le_env_def]) 1); by (fast_tac (claset() addEs [bound_typ_instance_trans] addss simpset()) 1); @@ -184,7 +184,7 @@ qed_spec_mp "has_type_le_env"; (* has_type is closed w.r.t. substitution *) -goal thy "!!a. A |- e :: t ==> !S. $S A |- e :: $S t"; +Goal "!!a. A |- e :: t ==> !S. $S A |- e :: $S t"; by (etac has_type.induct 1); (* case VarI *) by (rtac allI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/MiniML/Type.ML --- a/src/HOL/MiniML/Type.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/MiniML/Type.ML Mon Jun 22 17:26:46 1998 +0200 @@ -9,14 +9,14 @@ (* lemmata for make scheme *) -goal thy "mk_scheme t = sch1 =-> sch2 --> (? t1 t2. sch1 = mk_scheme t1 & sch2 = mk_scheme t2)"; +Goal "mk_scheme t = sch1 =-> sch2 --> (? t1 t2. sch1 = mk_scheme t1 & sch2 = mk_scheme t2)"; by (typ.induct_tac "t" 1); by (Simp_tac 1); by (Asm_full_simp_tac 1); by (Fast_tac 1); qed_spec_mp "mk_scheme_Fun"; -goal thy "!t'. mk_scheme t = mk_scheme t' --> t=t'"; +Goal "!t'. mk_scheme t = mk_scheme t' --> t=t'"; by (typ.induct_tac "t" 1); by (rtac allI 1); by (typ.induct_tac "t'" 1); @@ -28,14 +28,14 @@ by (Asm_full_simp_tac 1); qed_spec_mp "mk_scheme_injective"; -goal thy "!!t. free_tv (mk_scheme t) = free_tv t"; +Goal "!!t. free_tv (mk_scheme t) = free_tv t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "free_tv_mk_scheme"; Addsimps [free_tv_mk_scheme]; -goal thy "!!t. $ S (mk_scheme t) = mk_scheme ($ S t)"; +Goal "!!t. $ S (mk_scheme t) = mk_scheme ($ S t)"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "subst_mk_scheme"; @@ -45,12 +45,12 @@ (* constructor laws for app_subst *) -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ S [] = []"; by (Simp_tac 1); qed "app_subst_Nil"; -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ S (x#l) = ($ S x)#($ S l)"; by (Simp_tac 1); qed "app_subst_Cons"; @@ -60,62 +60,62 @@ (* constructor laws for new_tv *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n (TVar m) = (m t2) = (new_tv n t1 & new_tv n t2)"; by (fast_tac (HOL_cs addss simpset()) 1); qed "new_tv_Fun"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n (t1 =-> t2) = (new_tv n t1 & new_tv n t2)"; by (fast_tac (HOL_cs addss simpset()) 1); qed "new_tv_Fun2"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n []"; by (Simp_tac 1); qed "new_tv_Nil"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n (x#l) = (new_tv n x & new_tv n l)"; by (fast_tac (HOL_cs addss simpset()) 1); qed "new_tv_Cons"; -goalw thy [new_tv_def] "!!n. new_tv n TVar"; +Goalw [new_tv_def] "!!n. new_tv n TVar"; by (strip_tac 1); by (asm_full_simp_tac (simpset() addsimps [free_tv_subst,dom_def,cod_def]) 1); qed "new_tv_TVar_subst"; Addsimps [new_tv_TVar,new_tv_FVar,new_tv_BVar,new_tv_Fun,new_tv_Fun2,new_tv_Nil,new_tv_Cons,new_tv_TVar_subst]; -goalw thy [id_subst_def,new_tv_def,free_tv_subst,dom_def,cod_def] +Goalw [id_subst_def,new_tv_def,free_tv_subst,dom_def,cod_def] "new_tv n id_subst"; by (Simp_tac 1); qed "new_tv_id_subst"; Addsimps[new_tv_id_subst]; -goal thy "new_tv n (sch::type_scheme) --> \ +Goal "new_tv n (sch::type_scheme) --> \ \ $(%k. if k \ +Goal "new_tv n (A::type_scheme list) --> \ \ $(%k. if k x : free_tv (A!n) --> x : free_tv A"; +Goal "!!A. !n. n < length A --> x : free_tv (A!n) --> x : free_tv A"; by (list.induct_tac "A" 1); by (Asm_full_simp_tac 1); by (rtac allI 1); @@ -157,7 +157,7 @@ Addsimps [free_tv_nth_A_impl_free_tv_A]; -goal thy "!!A. !nat. nat < length A --> x : free_tv (A!nat) --> x : free_tv A"; +Goal "!!A. !nat. nat < length A --> x : free_tv (A!nat) --> x : free_tv A"; by (list.induct_tac "A" 1); by (Asm_full_simp_tac 1); by (rtac allI 1); @@ -172,32 +172,32 @@ structure the substitutions coincide on the free type variables occurring in the type structure *) -goal thy "!!S S'. (!x:free_tv t. (S x) = (S' x)) --> $ S (t::typ) = $ S' t"; +Goal "!!S S'. (!x:free_tv t. (S x) = (S' x)) --> $ S (t::typ) = $ S' t"; by (typ.induct_tac "t" 1); by (Simp_tac 1); by (Asm_full_simp_tac 1); qed_spec_mp "typ_substitutions_only_on_free_variables"; -goal thy +Goal "!!t. (!n. n:(free_tv t) --> S1 n = S2 n) ==> $ S1 (t::typ) = $ S2 t"; by (rtac typ_substitutions_only_on_free_variables 1); by (simp_tac (simpset() addsimps [Ball_def]) 1); qed "eq_free_eq_subst_te"; -goal thy "!!S S'. (!x:free_tv sch. (S x) = (S' x)) --> $ S (sch::type_scheme) = $ S' sch"; +Goal "!!S S'. (!x:free_tv sch. (S x) = (S' x)) --> $ S (sch::type_scheme) = $ S' sch"; by (type_scheme.induct_tac "sch" 1); by (Simp_tac 1); by (Simp_tac 1); by (Asm_full_simp_tac 1); qed_spec_mp "scheme_substitutions_only_on_free_variables"; -goal thy +Goal "!!sch. (!n. n:(free_tv sch) --> S1 n = S2 n) ==> $ S1 (sch::type_scheme) = $ S2 sch"; by (rtac scheme_substitutions_only_on_free_variables 1); by (simp_tac (simpset() addsimps [Ball_def]) 1); qed "eq_free_eq_subst_type_scheme"; -goal thy +Goal "(!n. n:(free_tv A) --> S1 n = S2 n) --> $S1 (A::type_scheme list) = $S2 A"; by (list.induct_tac "A" 1); (* case [] *) @@ -206,11 +206,11 @@ by (fast_tac (HOL_cs addIs [eq_free_eq_subst_type_scheme] addss (simpset())) 1); qed_spec_mp "eq_free_eq_subst_scheme_list"; -goal thy "!!P Q. ((!x:A. (P x)) --> Q) ==> ((!x:A Un B. (P x)) --> Q)"; +Goal "!!P Q. ((!x:A. (P x)) --> Q) ==> ((!x:A Un B. (P x)) --> Q)"; by (Fast_tac 1); val weaken_asm_Un = result (); -goal thy "!!S S'. (!x:free_tv A. (S x) = (S' x)) --> $ S (A::type_scheme list) = $ S' A"; +Goal "!!S S'. (!x:free_tv A. (S x) = (S' x)) --> $ S (A::type_scheme list) = $ S' A"; by (list.induct_tac "A" 1); by (Simp_tac 1); by (Asm_full_simp_tac 1); @@ -219,7 +219,7 @@ by (etac scheme_substitutions_only_on_free_variables 1); qed_spec_mp "scheme_list_substitutions_only_on_free_variables"; -goal thy +Goal "$ S1 (t::typ) = $ S2 t --> n:(free_tv t) --> S1 n = S2 n"; by (typ.induct_tac "t" 1); (* case TVar n *) @@ -228,7 +228,7 @@ by (fast_tac (HOL_cs addss simpset()) 1); qed_spec_mp "eq_subst_te_eq_free"; -goal thy +Goal "$ S1 (sch::type_scheme) = $ S2 sch --> n:(free_tv sch) --> S1 n = S2 n"; by (type_scheme.induct_tac "sch" 1); (* case TVar n *) @@ -241,7 +241,7 @@ by (Asm_full_simp_tac 1); qed_spec_mp "eq_subst_type_scheme_eq_free"; -goal thy +Goal "$S1 (A::type_scheme list) = $S2 A --> n:(free_tv A) --> S1 n = S2 n"; by (list.induct_tac "A" 1); (* case [] *) @@ -250,22 +250,22 @@ by (fast_tac (HOL_cs addIs [eq_subst_type_scheme_eq_free] addss (simpset())) 1); qed_spec_mp "eq_subst_scheme_list_eq_free"; -goalw thy [free_tv_subst] +Goalw [free_tv_subst] "!!v. v : cod S ==> v : free_tv S"; by (fast_tac set_cs 1); qed "codD"; -goalw thy [free_tv_subst,dom_def] +Goalw [free_tv_subst,dom_def] "!! x. x ~: free_tv S ==> S x = TVar x"; by (fast_tac set_cs 1); qed "not_free_impl_id"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "!! n. [| new_tv n t; m:free_tv t |] ==> m v : cod S"; by (Simp_tac 1); by (safe_tac (empty_cs addSIs [exI,conjI,notI])); @@ -275,14 +275,14 @@ qed "cod_app_subst"; Addsimps [cod_app_subst]; -goal thy +Goal "free_tv (S (v::nat)) <= insert v (cod S)"; by (expand_case_tac "v:dom S" 1); by (fast_tac (set_cs addss (simpset() addsimps [cod_def])) 1); by (fast_tac (set_cs addss (simpset() addsimps [dom_def])) 1); qed "free_tv_subst_var"; -goal thy +Goal "free_tv ($ S (t::typ)) <= cod S Un free_tv t"; by (typ.induct_tac "t" 1); (* case TVar n *) @@ -291,7 +291,7 @@ by (fast_tac (set_cs addss simpset()) 1); qed "free_tv_app_subst_te"; -goal thy +Goal "free_tv ($ S (sch::type_scheme)) <= cod S Un free_tv sch"; by (type_scheme.induct_tac "sch" 1); (* case FVar n *) @@ -302,7 +302,7 @@ by (fast_tac (set_cs addss simpset()) 1); qed "free_tv_app_subst_type_scheme"; -goal thy +Goal "free_tv ($ S (A::type_scheme list)) <= cod S Un free_tv A"; by (list.induct_tac "A" 1); (* case [] *) @@ -312,7 +312,7 @@ by (fast_tac (set_cs addss simpset()) 1); qed "free_tv_app_subst_scheme_list"; -goalw thy [free_tv_subst,dom_def] +Goalw [free_tv_subst,dom_def] "free_tv (%u::nat. $ s1 (s2 u) :: typ) <= \ \ free_tv s1 Un free_tv s2"; by (fast_tac (set_cs addSDs [free_tv_app_subst_te RS subsetD, @@ -321,19 +321,19 @@ addsimps [cod_def,dom_def])) 1); qed "free_tv_comp_subst"; -goalw thy [o_def] +Goalw [o_def] "free_tv ($ S1 o S2) <= free_tv S1 Un free_tv (S2 :: nat => typ)"; by (rtac free_tv_comp_subst 1); qed "free_tv_o_subst"; -goal thy "!!n. n : free_tv t --> free_tv (S n) <= free_tv ($ S t::typ)"; +Goal "!!n. n : free_tv t --> free_tv (S n) <= free_tv ($ S t::typ)"; by (typ.induct_tac "t" 1); by (Simp_tac 1); by (Simp_tac 1); by (Fast_tac 1); qed_spec_mp "free_tv_of_substitutions_extend_to_types"; -goal thy "!!n. n : free_tv sch --> free_tv (S n) <= free_tv ($ S sch::type_scheme)"; +Goal "!!n. n : free_tv sch --> free_tv (S n) <= free_tv ($ S sch::type_scheme)"; by (type_scheme.induct_tac "sch" 1); by (Simp_tac 1); by (Simp_tac 1); @@ -341,7 +341,7 @@ by (Fast_tac 1); qed_spec_mp "free_tv_of_substitutions_extend_to_schemes"; -goal thy "!!n. n : free_tv A --> free_tv (S n) <= free_tv ($ S A::type_scheme list)"; +Goal "!!n. n : free_tv A --> free_tv (S n) <= free_tv ($ S A::type_scheme list)"; by (list.induct_tac "A" 1); by (Simp_tac 1); by (Simp_tac 1); @@ -350,14 +350,14 @@ Addsimps [free_tv_of_substitutions_extend_to_scheme_lists]; -goal thy "!!sch::type_scheme. (n : free_tv sch) = (n mem free_tv_ML sch)"; +Goal "!!sch::type_scheme. (n : free_tv sch) = (n mem free_tv_ML sch)"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS Asm_simp_tac); by (strip_tac 1); by (Fast_tac 1); qed "free_tv_ML_scheme"; -goal thy "!!A::type_scheme list. (n : free_tv A) = (n mem free_tv_ML A)"; +Goal "!!A::type_scheme list. (n : free_tv A) = (n mem free_tv_ML A)"; by (list.induct_tac "A" 1); by (Simp_tac 1); by (asm_simp_tac (simpset() addsimps [free_tv_ML_scheme]) 1); @@ -366,21 +366,21 @@ (* lemmata for bound_tv *) -goal thy "!!t. bound_tv (mk_scheme t) = {}"; +Goal "!!t. bound_tv (mk_scheme t) = {}"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "bound_tv_mk_scheme"; Addsimps [bound_tv_mk_scheme]; -goal thy "!!sch::type_scheme. bound_tv ($ S sch) = bound_tv sch"; +Goal "!!sch::type_scheme. bound_tv ($ S sch) = bound_tv sch"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS Asm_simp_tac); qed "bound_tv_subst_scheme"; Addsimps [bound_tv_subst_scheme]; -goal thy "!!A::type_scheme list. bound_tv ($ S A) = bound_tv A"; +Goal "!!A::type_scheme list. bound_tv ($ S A) = bound_tv A"; by (rtac list.induct 1); by (Simp_tac 1); by (Asm_simp_tac 1); @@ -391,7 +391,7 @@ (* lemmata for new_tv *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n S = ((!m. n <= m --> (S m = TVar m)) & \ \ (! l. l < n --> new_tv n (S l) ))"; by (safe_tac HOL_cs ); @@ -412,7 +412,7 @@ by (fast_tac (HOL_cs addSIs [less_or_eq_imp_le]) 1); qed "new_tv_subst"; -goal thy +Goal "new_tv n = list_all (new_tv n)"; by (rtac ext 1); by (list.induct_tac "x" 1); @@ -420,21 +420,21 @@ qed "new_tv_list"; (* substitution affects only variables occurring freely *) -goal thy +Goal "new_tv n (t::typ) --> $(%x. if x=n then t' else S x) t = $S t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "subst_te_new_tv"; Addsimps [subst_te_new_tv]; -goal thy +Goal "new_tv n (sch::type_scheme) --> $(%x. if x=n then sch' else S x) sch = $S sch"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS Asm_simp_tac); qed_spec_mp "subst_te_new_type_scheme"; Addsimps [subst_te_new_type_scheme]; -goal thy +Goal "new_tv n (A::type_scheme list) --> $(%x. if x=n then t else S x) A = $S A"; by (list.induct_tac "A" 1); by (ALLGOALS Asm_full_simp_tac); @@ -442,7 +442,7 @@ Addsimps [subst_tel_new_scheme_list]; (* all greater variables are also new *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "!!n m. n<=m ==> new_tv n t ==> new_tv m t"; by Safe_tac; by (dtac spec 1); @@ -451,32 +451,32 @@ qed "new_tv_le"; Addsimps [lessI RS less_imp_le RS new_tv_le]; -goal thy "!!n m. n<=m ==> new_tv n (t::typ) ==> new_tv m t"; +Goal "!!n m. n<=m ==> new_tv n (t::typ) ==> new_tv m t"; by (asm_simp_tac (simpset() addsimps [new_tv_le]) 1); qed "new_tv_typ_le"; -goal thy "!!n m. n<=m ==> new_tv n (A::type_scheme list) ==> new_tv m A"; +Goal "!!n m. n<=m ==> new_tv n (A::type_scheme list) ==> new_tv m A"; by (asm_simp_tac (simpset() addsimps [new_tv_le]) 1); qed "new_scheme_list_le"; -goal thy "!!n m. n<=m ==> new_tv n (S::subst) ==> new_tv m S"; +Goal "!!n m. n<=m ==> new_tv n (S::subst) ==> new_tv m S"; by (asm_simp_tac (simpset() addsimps [new_tv_le]) 1); qed "new_tv_subst_le"; (* new_tv property remains if a substitution is applied *) -goal thy +Goal "!!n. [| n new_tv m (S n)"; by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); qed "new_tv_subst_var"; -goal thy +Goal "new_tv n S --> new_tv n (t::typ) --> new_tv n ($ S t)"; by (typ.induct_tac "t" 1); by (ALLGOALS(fast_tac (HOL_cs addss (simpset() addsimps [new_tv_subst])))); qed_spec_mp "new_tv_subst_te"; Addsimps [new_tv_subst_te]; -goal thy "new_tv n S --> new_tv n (sch::type_scheme) --> new_tv n ($ S sch)"; +Goal "new_tv n S --> new_tv n (sch::type_scheme) --> new_tv n ($ S sch)"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS (Asm_full_simp_tac)); by (rewtac new_tv_def); @@ -490,29 +490,29 @@ qed_spec_mp "new_tv_subst_type_scheme"; Addsimps [new_tv_subst_type_scheme]; -goal thy +Goal "new_tv n S --> new_tv n (A::type_scheme list) --> new_tv n ($ S A)"; by (list.induct_tac "A" 1); by (ALLGOALS(fast_tac (HOL_cs addss (simpset())))); qed_spec_mp "new_tv_subst_scheme_list"; Addsimps [new_tv_subst_scheme_list]; -goal thy +Goal "new_tv n A --> new_tv (Suc n) ((TVar n)#A)"; by (simp_tac (simpset() addsimps [new_tv_list]) 1); by (list.induct_tac "A" 1); by (ALLGOALS Asm_full_simp_tac); qed "new_tv_Suc_list"; -goalw thy [new_tv_def] "!!sch::type_scheme. (free_tv sch) = (free_tv sch') --> (new_tv n sch --> new_tv n sch')"; +Goalw [new_tv_def] "!!sch::type_scheme. (free_tv sch) = (free_tv sch') --> (new_tv n sch --> new_tv n sch')"; by (Asm_simp_tac 1); qed_spec_mp "new_tv_only_depends_on_free_tv_type_scheme"; -goalw thy [new_tv_def] "!!A::type_scheme list. (free_tv A) = (free_tv A') --> (new_tv n A --> new_tv n A')"; +Goalw [new_tv_def] "!!A::type_scheme list. (free_tv A) = (free_tv A') --> (new_tv n A --> new_tv n A')"; by (Asm_simp_tac 1); qed_spec_mp "new_tv_only_depends_on_free_tv_scheme_list"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "!!A. !nat. nat < length A --> new_tv n A --> (new_tv n (A!nat))"; by (list.induct_tac "A" 1); by (Asm_full_simp_tac 1); @@ -525,13 +525,13 @@ (* composition of substitutions preserves new_tv proposition *) -goal thy +Goal "!! n. [| new_tv n (S::subst); new_tv n R |] ==> \ \ new_tv n (($ R) o S)"; by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); qed "new_tv_subst_comp_1"; -goal thy +Goal "!! n. [| new_tv n (S::subst); new_tv n R |] ==> \ \ new_tv n (%v.$ R (S v))"; by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); @@ -540,24 +540,24 @@ Addsimps [new_tv_subst_comp_1,new_tv_subst_comp_2]; (* new type variables do not occur freely in a type structure *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "!!n. new_tv n A ==> n~:(free_tv A)"; by (fast_tac (HOL_cs addEs [less_irrefl]) 1); qed "new_tv_not_free_tv"; Addsimps [new_tv_not_free_tv]; -goalw thy [max_def] "!!n::nat. m < n ==> m < max n n'"; +Goalw [max_def] "!!n::nat. m < n ==> m < max n n'"; by (Simp_tac 1); by Safe_tac; by (trans_tac 1); qed "less_maxL"; -goalw thy [max_def] "!!n::nat. m < n' ==> m < max n n'"; +Goalw [max_def] "!!n::nat. m < n' ==> m < max n n'"; by (Simp_tac 1); by (fast_tac (claset() addDs [not_leE] addIs [less_trans]) 1); val less_maxR = result(); -goalw thy [new_tv_def] "!!t::typ. ? n. (new_tv n t)"; +Goalw [new_tv_def] "!!t::typ. ? n. (new_tv n t)"; by (typ.induct_tac "t" 1); by (res_inst_tac [("x","Suc nat")] exI 1); by (Asm_simp_tac 1); @@ -570,7 +570,7 @@ Addsimps [fresh_variable_types]; -goalw thy [new_tv_def] "!!sch::type_scheme. ? n. (new_tv n sch)"; +Goalw [new_tv_def] "!!sch::type_scheme. ? n. (new_tv n sch)"; by (type_scheme.induct_tac "sch" 1); by (res_inst_tac [("x","Suc nat")] exI 1); by (Simp_tac 1); @@ -585,16 +585,16 @@ Addsimps [fresh_variable_type_schemes]; -goalw thy [max_def] "!!n::nat. n <= (max n n')"; +Goalw [max_def] "!!n::nat. n <= (max n n')"; by (Simp_tac 1); val le_maxL = result(); -goalw thy [max_def] "!!n'::nat. n' <= (max n n')"; +Goalw [max_def] "!!n'::nat. n' <= (max n n')"; by (Simp_tac 1); by (fast_tac (claset() addDs [not_leE] addIs [less_or_eq_imp_le]) 1); val le_maxR = result(); -goal thy "!!A::type_scheme list. ? n. (new_tv n A)"; +Goal "!!A::type_scheme list. ? n. (new_tv n A)"; by (list.induct_tac "A" 1); by (Simp_tac 1); by (Simp_tac 1); @@ -611,7 +611,7 @@ Addsimps [fresh_variable_type_scheme_lists]; -goal thy "!!x y. [| ? n1. (new_tv n1 x); ? n2. (new_tv n2 y)|] ==> ? n. (new_tv n x) & (new_tv n y)"; +Goal "!!x y. [| ? n1. (new_tv n1 x); ? n2. (new_tv n2 y)|] ==> ? n. (new_tv n x) & (new_tv n y)"; by (REPEAT (etac exE 1)); by (rename_tac "n1 n2" 1); by (res_inst_tac [("x","(max n1 n2)")] exI 1); @@ -621,7 +621,7 @@ by (ALLGOALS (simp_tac (simpset() addsimps [le_maxL,le_maxR]))); qed "make_one_new_out_of_two"; -goal thy "!!(A::type_scheme list) (A'::type_scheme list) (t::typ) (t'::typ). \ +Goal "!!(A::type_scheme list) (A'::type_scheme list) (t::typ) (t'::typ). \ \ ? n. (new_tv n A) & (new_tv n A') & (new_tv n t) & (new_tv n t')" ; by (cut_inst_tac [("t","t")] fresh_variable_types 1); by (cut_inst_tac [("t","t'")] fresh_variable_types 1); @@ -642,7 +642,7 @@ qed "ex_fresh_variable"; (* mgu does not introduce new type variables *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "!! n. [|mgu t1 t2 = Some u; new_tv n t1; new_tv n t2|] ==> \ \ new_tv n u"; by (fast_tac (set_cs addDs [mgu_free]) 1); @@ -651,20 +651,20 @@ (* lemmata for substitutions *) -goalw Type.thy [app_subst_list] +Goalw [app_subst_list] "!!A:: ('a::type_struct) list. length ($ S A) = length A"; by (Simp_tac 1); qed "length_app_subst_list"; Addsimps [length_app_subst_list]; -goal thy "!!sch::type_scheme. $ TVar sch = sch"; +Goal "!!sch::type_scheme. $ TVar sch = sch"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS Asm_simp_tac); qed "subst_TVar_scheme"; Addsimps [subst_TVar_scheme]; -goal thy "!!A::type_scheme list. $ TVar A = A"; +Goal "!!A::type_scheme list. $ TVar A = A"; by (rtac list.induct 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [app_subst_list]))); qed "subst_TVar_scheme_list"; @@ -672,7 +672,7 @@ Addsimps [subst_TVar_scheme_list]; (* application of id_subst does not change type expression *) -goalw thy [id_subst_def] +Goalw [id_subst_def] "$ id_subst = (%t::typ. t)"; by (rtac ext 1); by (typ.induct_tac "t" 1); @@ -680,7 +680,7 @@ qed "app_subst_id_te"; Addsimps [app_subst_id_te]; -goalw thy [id_subst_def] +Goalw [id_subst_def] "$ id_subst = (%sch::type_scheme. sch)"; by (rtac ext 1); by (type_scheme.induct_tac "t" 1); @@ -689,7 +689,7 @@ Addsimps [app_subst_id_type_scheme]; (* application of id_subst does not change list of type expressions *) -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ id_subst = (%A::type_scheme list. A)"; by (rtac ext 1); by (list.induct_tac "A" 1); @@ -697,14 +697,14 @@ qed "app_subst_id_tel"; Addsimps [app_subst_id_tel]; -goal thy "!!sch::type_scheme. $ id_subst sch = sch"; +Goal "!!sch::type_scheme. $ id_subst sch = sch"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [id_subst_def]))); qed "id_subst_sch"; Addsimps [id_subst_sch]; -goal thy "!!A::type_scheme list. $ id_subst A = A"; +Goal "!!A::type_scheme list. $ id_subst A = A"; by (list.induct_tac "A" 1); by (Asm_full_simp_tac 1); by (Asm_full_simp_tac 1); @@ -713,25 +713,25 @@ Addsimps [id_subst_A]; (* composition of substitutions *) -goalw Type.thy [id_subst_def,o_def] "$S o id_subst = S"; +Goalw [id_subst_def,o_def] "$S o id_subst = S"; by (rtac ext 1); by (Simp_tac 1); qed "o_id_subst"; Addsimps[o_id_subst]; -goal thy +Goal "$ R ($ S t::typ) = $ (%x. $ R (S x) ) t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "subst_comp_te"; -goal thy +Goal "$ R ($ S sch::type_scheme) = $ (%x. $ R (S x) ) sch"; by (type_scheme.induct_tac "sch" 1); by (ALLGOALS Asm_full_simp_tac); qed "subst_comp_type_scheme"; -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ R ($ S A::type_scheme list) = $ (%x. $ R (S x)) A"; by (list.induct_tac "A" 1); (* case [] *) @@ -740,18 +740,18 @@ by (asm_full_simp_tac (simpset() addsimps [subst_comp_type_scheme]) 1); qed "subst_comp_scheme_list"; -goal thy "!!A::type_scheme list. !x : free_tv A. S x = (TVar x) ==> $ S A = $ id_subst A"; +Goal "!!A::type_scheme list. !x : free_tv A. S x = (TVar x) ==> $ S A = $ id_subst A"; by (rtac scheme_list_substitutions_only_on_free_variables 1); by (asm_full_simp_tac (simpset() addsimps [id_subst_def]) 1); qed "subst_id_on_type_scheme_list'"; -goal thy "!!A::type_scheme list. !x : free_tv A. S x = (TVar x) ==> $ S A = A"; +Goal "!!A::type_scheme list. !x : free_tv A. S x = (TVar x) ==> $ S A = A"; by (stac subst_id_on_type_scheme_list' 1); by (assume_tac 1); by (Simp_tac 1); qed "subst_id_on_type_scheme_list"; -goal thy "!n. n < length A --> ($ S A)!n = $S (A!n)"; +Goal "!n. n < length A --> ($ S A)!n = $S (A!n)"; by (list.induct_tac "A" 1); by (Asm_full_simp_tac 1); by (rtac allI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/MiniML/W.ML --- a/src/HOL/MiniML/W.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/MiniML/W.ML Mon Jun 22 17:26:46 1998 +0200 @@ -14,7 +14,7 @@ [" A |- Var n :: t"," A |- Abs e :: t","A |- App e1 e2 ::t","A |- LET e1 e2 ::t" ]; (* the resulting type variable is always greater or equal than the given one *) -goal thy +Goal "!A n S t m. W e A n = Some (S,t,m) --> n<=m"; by (expr.induct_tac "e" 1); (* case Var(n) *) @@ -32,19 +32,19 @@ Addsimps [W_var_ge]; -goal thy +Goal "!! s. Some (S,t,m) = W e A n ==> n<=m"; by (asm_full_simp_tac (simpset() addsimps [eq_sym_conv]) 1); qed "W_var_geD"; -goal thy "!! s. new_tv n A ==> Some (S,t,m) = W e A n ==> new_tv m A"; +Goal "!! s. new_tv n A ==> Some (S,t,m) = W e A n ==> new_tv m A"; by (dtac W_var_geD 1); by (rtac new_scheme_list_le 1); by (assume_tac 1); by (assume_tac 1); qed "new_tv_compatible_W"; -goal thy "!!sch. new_tv n sch --> new_tv (n + (min_new_bound_tv sch)) (bound_typ_inst (%b. TVar (b + n)) sch)"; +Goal "!!sch. new_tv n sch --> new_tv (n + (min_new_bound_tv sch)) (bound_typ_inst (%b. TVar (b + n)) sch)"; by (type_scheme.induct_tac "sch" 1); by (Asm_full_simp_tac 1); by (asm_full_simp_tac (simpset() addsimps [add_commute]) 1); @@ -71,7 +71,7 @@ Addsimps [new_tv_bound_typ_inst_sch]; (* resulting type variable is new *) -goal thy +Goal "!n A S t m. new_tv n A --> W e A n = Some (S,t,m) --> \ \ new_tv m S & new_tv m t"; by (expr.induct_tac "e" 1); @@ -155,7 +155,7 @@ by (assume_tac 1); qed_spec_mp "new_tv_W"; -goal thy "!!sch. (v ~: free_tv sch) --> (v : free_tv (bound_typ_inst (TVar o S) sch)) --> (? x. v = S x)"; +Goal "!!sch. (v ~: free_tv sch) --> (v : free_tv (bound_typ_inst (TVar o S) sch)) --> (? x. v = S x)"; by (type_scheme.induct_tac "sch" 1); by (Asm_full_simp_tac 1); by (Asm_full_simp_tac 1); @@ -167,7 +167,7 @@ Addsimps [free_tv_bound_typ_inst1]; -goal thy +Goal "!n A S t m v. W e A n = Some (S,t,m) --> \ \ (v:free_tv S | v:free_tv t) --> v v:free_tv A"; by (expr.induct_tac "e" 1); @@ -254,16 +254,16 @@ addDs [less_le_trans]) 1); qed_spec_mp "free_tv_W"; -goal thy "!!A. (!x. x : A --> x ~: B) ==> A Int B = {}"; +Goal "!!A. (!x. x : A --> x ~: B) ==> A Int B = {}"; by (Fast_tac 1); val weaken_A_Int_B_eq_empty = result(); -goal thy "!!A. x ~: A | x : B ==> x ~: A - B"; +Goal "!!A. x ~: A | x : B ==> x ~: A - B"; by (Fast_tac 1); val weaken_not_elem_A_minus_B = result(); (* correctness of W with respect to has_type *) -goal W.thy +Goal "!A S t m n . new_tv n A --> Some (S,t,m) = W e A n --> $S A |- e :: t"; by (expr.induct_tac "e" 1); (* case Var n *) @@ -384,7 +384,7 @@ qed_spec_mp "W_correct_lemma"; (* Completeness of W w.r.t. has_type *) -goal thy +Goal "!S' A t' n. $S' A |- e :: t' --> new_tv n A --> \ \ (? S t. (? m. W e A n = Some (S,t,m)) & \ \ (? R. $S' A = $R ($S A) & t' = $R t))"; @@ -558,7 +558,7 @@ by (simp_tac (simpset() addsimps [o_def,subst_comp_scheme_list RS sym]) 1); qed_spec_mp "W_complete_lemma"; -goal W.thy +Goal "!!e. [] |- e :: t' ==> (? S t. (? m. W e [] n = Some(S,t,m)) & \ \ (? R. t' = $ R t))"; by (cut_inst_tac [("A","[]"),("S'","id_subst"),("e","e"),("t'","t'")] diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Modelcheck/Example.ML --- a/src/HOL/Modelcheck/Example.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Modelcheck/Example.ML Mon Jun 22 17:26:46 1998 +0200 @@ -7,7 +7,7 @@ val reach_rws = [reach_def,INIT_def,N_def]; -goal thy "reach (True,True,True)"; +Goal "reach (True,True,True)"; by (simp_tac (MC_ss addsimps reach_rws) 1); (*show the current proof state using the model checker syntax*) @@ -23,6 +23,6 @@ (* just to make a proof in this file :-) *) -goalw thy [INIT_def] "INIT (a,b,c) = (~a & ~b &~c)"; +Goalw [INIT_def] "INIT (a,b,c) = (~a & ~b &~c)"; by (Simp_tac 1); qed"init_state"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Modelcheck/MCSyn.ML --- a/src/HOL/Modelcheck/MCSyn.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Modelcheck/MCSyn.ML Mon Jun 22 17:26:46 1998 +0200 @@ -18,7 +18,7 @@ end; -goalw thy [split_def] "(f::'a*'b=>'c) = (%(x, y). f (x, y))"; +Goalw [split_def] "(f::'a*'b=>'c) = (%(x, y). f (x, y))"; by (rtac ext 1); by (stac (surjective_pairing RS sym) 1); by (rtac refl 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Nat.ML --- a/src/HOL/Nat.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Nat.ML Mon Jun 22 17:26:46 1998 +0200 @@ -4,17 +4,17 @@ Copyright 1997 TU Muenchen *) -goal thy "min 0 n = 0"; +Goal "min 0 n = 0"; by (rtac min_leastL 1); by (trans_tac 1); qed "min_0L"; -goal thy "min n 0 = 0"; +Goal "min n 0 = 0"; by (rtac min_leastR 1); by (trans_tac 1); qed "min_0R"; -goalw thy [min_def] "min (Suc m) (Suc n) = Suc(min m n)"; +Goalw [min_def] "min (Suc m) (Suc n) = Suc(min m n)"; by (Simp_tac 1); qed "min_Suc_Suc"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/NatDef.ML --- a/src/HOL/NatDef.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/NatDef.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,14 +8,14 @@ because the type will be promoted to type class "order". *) -goal thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))"; +Goal "mono(%X. {Zero_Rep} Un (Suc_Rep``X))"; by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1)); qed "Nat_fun_mono"; val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski); (* Zero is a natural number -- this also justifies the type definition*) -goal thy "Zero_Rep: Nat"; +Goal "Zero_Rep: Nat"; by (stac Nat_unfold 1); by (rtac (singletonI RS UnI1) 1); qed "Zero_RepI"; @@ -80,19 +80,19 @@ (*We can't take these properties as axioms, or take Abs_Nat==Inv(Rep_Nat), since we assume the isomorphism equations will one day be given by Isabelle*) -goal thy "inj(Rep_Nat)"; +Goal "inj(Rep_Nat)"; by (rtac inj_inverseI 1); by (rtac Rep_Nat_inverse 1); qed "inj_Rep_Nat"; -goal thy "inj_on Abs_Nat Nat"; +Goal "inj_on Abs_Nat Nat"; by (rtac inj_on_inverseI 1); by (etac Abs_Nat_inverse 1); qed "inj_on_Abs_Nat"; (*** Distinctness of constructors ***) -goalw thy [Zero_def,Suc_def] "Suc(m) ~= 0"; +Goalw [Zero_def,Suc_def] "Suc(m) ~= 0"; by (rtac (inj_on_Abs_Nat RS inj_on_contraD) 1); by (rtac Suc_Rep_not_Zero_Rep 1); by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1)); @@ -107,7 +107,7 @@ (** Injectiveness of Suc **) -goalw thy [Suc_def] "inj(Suc)"; +Goalw [Suc_def] "inj(Suc)"; by (rtac injI 1); by (dtac (inj_on_Abs_Nat RS inj_onD) 1); by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1)); @@ -117,20 +117,20 @@ val Suc_inject = inj_Suc RS injD; -goal thy "(Suc(m)=Suc(n)) = (m=n)"; +Goal "(Suc(m)=Suc(n)) = (m=n)"; by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); qed "Suc_Suc_eq"; AddIffs [Suc_Suc_eq]; -goal thy "n ~= Suc(n)"; +Goal "n ~= Suc(n)"; by (nat_ind_tac "n" 1); by (ALLGOALS Asm_simp_tac); qed "n_not_Suc_n"; bind_thm ("Suc_n_not_n", n_not_Suc_n RS not_sym); -goal thy "!!n. n ~= 0 ==> EX m. n = Suc m"; +Goal "!!n. n ~= 0 ==> EX m. n = Suc m"; by (rtac natE 1); by (REPEAT (Blast_tac 1)); qed "not0_implies_Suc"; @@ -138,15 +138,15 @@ (*** nat_case -- the selection operator for nat ***) -goalw thy [nat_case_def] "nat_case a f 0 = a"; +Goalw [nat_case_def] "nat_case a f 0 = a"; by (Blast_tac 1); qed "nat_case_0"; -goalw thy [nat_case_def] "nat_case a f (Suc k) = f(k)"; +Goalw [nat_case_def] "nat_case a f (Suc k) = f(k)"; by (Blast_tac 1); qed "nat_case_Suc"; -goalw thy [wf_def, pred_nat_def] "wf(pred_nat)"; +Goalw [wf_def, pred_nat_def] "wf(pred_nat)"; by (Clarify_tac 1); by (nat_ind_tac "x" 1); by (ALLGOALS Blast_tac); @@ -156,7 +156,7 @@ (*** nat_rec -- by wf recursion on pred_nat ***) (* The unrolling rule for nat_rec *) -goal thy +Goal "nat_rec c d = wfrec pred_nat (%f. nat_case c (%m. d m (f m)))"; by (simp_tac (HOL_ss addsimps [nat_rec_def]) 1); bind_thm("nat_rec_unfold", wf_pred_nat RS @@ -169,12 +169,12 @@ (** conversion rules **) -goal thy "nat_rec c h 0 = c"; +Goal "nat_rec c h 0 = c"; by (rtac (nat_rec_unfold RS trans) 1); by (simp_tac (simpset() addsimps [nat_case_0]) 1); qed "nat_rec_0"; -goal thy "nat_rec c h (Suc n) = h n (nat_rec c h n)"; +Goal "nat_rec c h (Suc n) = h n (nat_rec c h n)"; by (rtac (nat_rec_unfold RS trans) 1); by (simp_tac (simpset() addsimps [nat_case_Suc, pred_nat_def, cut_apply]) 1); qed "nat_rec_Suc"; @@ -200,7 +200,7 @@ (*** Basic properties of "less than" ***) (*Used in TFL/post.sml*) -goalw thy [less_def] "(m,n) : pred_nat^+ = (m R *) bind_thm ("less_asym", (less_not_sym RS notE)); -goalw thy [less_def] "~ n<(n::nat)"; +Goalw [less_def] "~ n<(n::nat)"; by (rtac notI 1); by (etac (wf_pred_nat RS wf_trancl RS wf_irrefl) 1); qed "less_not_refl"; @@ -246,7 +246,7 @@ (* n R *) bind_thm ("less_irrefl", (less_not_refl RS notE)); -goal thy "!!m. n m ~= (n::nat)"; +Goal "!!m. n m ~= (n::nat)"; by (blast_tac (claset() addSEs [less_irrefl]) 1); qed "less_not_refl2"; @@ -260,7 +260,7 @@ eresolve_tac (prems@[asm_rl, Pair_inject]))); qed "lessE"; -goal thy "~ n<0"; +Goal "~ n<0"; by (rtac notI 1); by (etac lessE 1); by (etac Zero_neq_Suc 1); @@ -281,11 +281,11 @@ by (Blast_tac 1); qed "less_SucE"; -goal thy "(m < Suc(n)) = (m < n | m = n)"; +Goal "(m < Suc(n)) = (m < n | m = n)"; by (blast_tac (claset() addSEs [less_SucE] addIs [less_trans]) 1); qed "less_Suc_eq"; -goal thy "(n<1) = (n=0)"; +Goal "(n<1) = (n=0)"; by (simp_tac (simpset() addsimps [less_Suc_eq]) 1); qed "less_one"; AddIffs [less_one]; @@ -296,7 +296,7 @@ by (ALLGOALS Asm_full_simp_tac); qed "gr_implies_not0"; -goal thy "(n ~= 0) = (0 < n)"; +Goal "(n ~= 0) = (0 < n)"; by (rtac natE 1); by (Blast_tac 1); by (Blast_tac 1); @@ -306,28 +306,28 @@ (*This theorem is useful with blast_tac: (n=0 ==> False) ==> 0 0 < n"; +Goal "!!m. m 0 < n"; by (dtac gr_implies_not0 1); by (Asm_full_simp_tac 1); qed "gr_implies_gr0"; Addsimps [gr_implies_gr0]; -goal thy "!!m n. m Suc(m) < Suc(n)"; +Goal "!!m n. m Suc(m) < Suc(n)"; by (etac rev_mp 1); by (nat_ind_tac "n" 1); by (ALLGOALS (fast_tac (claset() addEs [less_trans, lessE]))); qed "Suc_mono"; (*"Less than" is a linear ordering*) -goal thy "m Suc(m) < n"; +Goal "!!m. [| m Suc(m) < n"; by (full_simp_tac (simpset() addsimps [nat_neq_iff]) 1); by (blast_tac (claset() addSEs [less_irrefl, less_SucE] addEs [less_asym]) 1); qed "Suc_lessI"; @@ -375,22 +375,22 @@ by (assume_tac 1); qed "Suc_lessE"; -goal thy "!!m n. Suc(m) < Suc(n) ==> m m j Suc i < k"; +Goal "!!i. i j Suc i < k"; by (nat_ind_tac "k" 1); by (ALLGOALS (asm_simp_tac (simpset()))); by (asm_simp_tac (simpset() addsimps [less_Suc_eq]) 1); @@ -398,7 +398,7 @@ qed_spec_mp "less_trans_Suc"; (*Can be used with less_Suc_eq to get n=m | n m < Suc n *) bind_thm ("le_imp_less_Suc", le_eq_less_Suc RS iffD1); -goalw thy [le_def] "0 <= n"; +Goalw [le_def] "0 <= n"; by (rtac not_less0 1); qed "le0"; -goalw thy [le_def] "~ Suc n <= n"; +Goalw [le_def] "~ Suc n <= n"; by (Simp_tac 1); qed "Suc_n_not_le_n"; -goalw thy [le_def] "(i <= 0) = (i = 0)"; +Goalw [le_def] "(i <= 0) = (i = 0)"; by (nat_ind_tac "i" 1); by (ALLGOALS Asm_simp_tac); qed "le_0_eq"; @@ -456,7 +456,7 @@ n_not_Suc_n, Suc_n_not_n, nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; -goal thy "!!m. (m <= Suc(n)) = (m<=n | m = Suc n)"; +Goal "!!m. (m <= Suc(n)) = (m<=n | m = Suc n)"; by (simp_tac (simpset() addsimps [le_eq_less_Suc]) 1); by (blast_tac (claset() addSEs [less_SucE] addIs [less_SucI]) 1); qed "le_Suc_eq"; @@ -465,7 +465,7 @@ bind_thm ("le_SucE", le_Suc_eq RS iffD1 RS disjE); (* -goal thy "(Suc m < n | Suc m = n) = (m < n)"; +Goal "(Suc m < n | Suc m = n) = (m < n)"; by (stac (less_Suc_eq RS sym) 1); by (rtac Suc_less_eq 1); qed "Suc_le_eq"; @@ -497,108 +497,108 @@ val leE = make_elim leD; -goal thy "(~n n<(m::nat)"; +Goalw [le_def] "!!m. ~ m <= n ==> n<(m::nat)"; by (Blast_tac 1); qed "not_leE"; -goalw thy [le_def] "(~n<=m) = (m<(n::nat))"; +Goalw [le_def] "(~n<=m) = (m<(n::nat))"; by (Simp_tac 1); qed "not_le_iff_less"; -goalw thy [le_def] "!!m. m < n ==> Suc(m) <= n"; +Goalw [le_def] "!!m. m < n ==> Suc(m) <= n"; by (simp_tac (simpset() addsimps [less_Suc_eq]) 1); by (blast_tac (claset() addSEs [less_irrefl,less_asym]) 1); qed "Suc_leI"; (*formerly called lessD*) -goalw thy [le_def] "!!m. Suc(m) <= n ==> m <= n"; +Goalw [le_def] "!!m. Suc(m) <= n ==> m <= n"; by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1); qed "Suc_leD"; (* stronger version of Suc_leD *) -goalw thy [le_def] +Goalw [le_def] "!!m. Suc m <= n ==> m < n"; by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq]) 1); by (cut_facts_tac [less_linear] 1); by (Blast_tac 1); qed "Suc_le_lessD"; -goal thy "(Suc m <= n) = (m < n)"; +Goal "(Suc m <= n) = (m < n)"; by (blast_tac (claset() addIs [Suc_leI, Suc_le_lessD]) 1); qed "Suc_le_eq"; -goalw thy [le_def] "!!m. m <= n ==> m <= Suc n"; +Goalw [le_def] "!!m. m <= n ==> m <= Suc n"; by (blast_tac (claset() addDs [Suc_lessD]) 1); qed "le_SucI"; Addsimps[le_SucI]; bind_thm ("le_Suc", not_Suc_n_less_n RS leI); -goalw thy [le_def] "!!m. m < n ==> m <= (n::nat)"; +Goalw [le_def] "!!m. m < n ==> m <= (n::nat)"; by (blast_tac (claset() addEs [less_asym]) 1); qed "less_imp_le"; (** Equivalence of m<=n and m m < n | m=(n::nat)"; +Goalw [le_def] "!!m. m <= n ==> m < n | m=(n::nat)"; by (cut_facts_tac [less_linear] 1); by (blast_tac (claset() addEs [less_irrefl,less_asym]) 1); qed "le_imp_less_or_eq"; -goalw thy [le_def] "!!m. m m <=(n::nat)"; +Goalw [le_def] "!!m. m m <=(n::nat)"; by (cut_facts_tac [less_linear] 1); by (blast_tac (claset() addSEs [less_irrefl] addEs [less_asym]) 1); qed "less_or_eq_imp_le"; -goal thy "(m <= (n::nat)) = (m < n | m=n)"; +Goal "(m <= (n::nat)) = (m < n | m=n)"; by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1)); qed "le_eq_less_or_eq"; (*Useful with Blast_tac. m=n ==> m<=n *) bind_thm ("eq_imp_le", disjI2 RS less_or_eq_imp_le); -goal thy "n <= (n::nat)"; +Goal "n <= (n::nat)"; by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); qed "le_refl"; -goal thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)"; +Goal "!!i. [| i <= j; j < k |] ==> i < (k::nat)"; by (blast_tac (claset() addSDs [le_imp_less_or_eq] addIs [less_trans]) 1); qed "le_less_trans"; -goal thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)"; +Goal "!!i. [| i < j; j <= k |] ==> i < (k::nat)"; by (blast_tac (claset() addSDs [le_imp_less_or_eq] addIs [less_trans]) 1); qed "less_le_trans"; -goal thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)"; +Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)"; by (blast_tac (claset() addSDs [le_imp_less_or_eq] addIs [less_or_eq_imp_le, less_trans]) 1); qed "le_trans"; -goal thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)"; +Goal "!!m. [| m <= n; n <= m |] ==> m = (n::nat)"; (*order_less_irrefl could make this proof fail*) by (blast_tac (claset() addSDs [le_imp_less_or_eq] addSEs [less_irrefl] addEs [less_asym]) 1); qed "le_anti_sym"; -goal thy "(Suc(n) <= Suc(m)) = (n <= m)"; +Goal "(Suc(n) <= Suc(m)) = (n <= m)"; by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); qed "Suc_le_mono"; AddIffs [Suc_le_mono]; (* Axiom 'order_le_less' of class 'order': *) -goal thy "(m::nat) < n = (m <= n & m ~= n)"; +Goal "(m::nat) < n = (m <= n & m ~= n)"; by (simp_tac (simpset() addsimps [le_def, nat_neq_iff]) 1); by (blast_tac (claset() addSEs [less_asym]) 1); qed "nat_less_le"; (* Axiom 'linorder_linear' of class 'linorder': *) -goal thy "(m::nat) <= n | n <= m"; +Goal "(m::nat) <= n | n <= m"; by (simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); by (cut_facts_tac [less_linear] 1); by(Blast_tac 1); @@ -607,12 +607,12 @@ (** max -goalw thy [max_def] "!!z::nat. (z <= max x y) = (z <= x | z <= y)"; +Goalw [max_def] "!!z::nat. (z <= max x y) = (z <= x | z <= y)"; by (simp_tac (simpset() addsimps [not_le_iff_less]) 1); by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); qed "le_max_iff_disj"; -goalw thy [max_def] "!!z::nat. (max x y <= z) = (x <= z & y <= z)"; +Goalw [max_def] "!!z::nat. (max x y <= z) = (x <= z & y <= z)"; by (simp_tac (simpset() addsimps [not_le_iff_less]) 1); by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); qed "max_le_iff_conj"; @@ -620,12 +620,12 @@ (** min **) -goalw thy [min_def] "!!z::nat. (z <= min x y) = (z <= x & z <= y)"; +Goalw [min_def] "!!z::nat. (z <= min x y) = (z <= x & z <= y)"; by (simp_tac (simpset() addsimps [not_le_iff_less]) 1); by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); qed "le_min_iff_conj"; -goalw thy [min_def] "!!z::nat. (min x y <= z) = (x <= z | y <= z)"; +Goalw [min_def] "!!z::nat. (min x y <= z) = (x <= z | y <= z)"; by (simp_tac (simpset() addsimps [not_le_iff_less] addsplits) 1); by (blast_tac (claset() addIs [less_imp_le, le_trans]) 1); qed "min_le_iff_disj"; @@ -633,12 +633,12 @@ (** LEAST -- the least number operator **) -goal thy "(! m::nat. P m --> n <= m) = (! m. m < n --> ~ P m)"; +Goal "(! m::nat. P m --> n <= m) = (! m. m < n --> ~ P m)"; by (blast_tac (claset() addIs [leI] addEs [leE]) 1); val lemma = result(); (* This is an old def of Least for nat, which is derived for compatibility *) -goalw thy [Least_def] +Goalw [Least_def] "(LEAST n::nat. P n) == (@n. P(n) & (ALL m. m < n --> ~P(m)))"; by (simp_tac (simpset() addsimps [lemma]) 1); qed "Least_nat_def"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Ord.ML --- a/src/HOL/Ord.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Ord.ML Mon Jun 22 17:26:46 1998 +0200 @@ -25,24 +25,24 @@ AddIffs [order_refl]; (*This form is useful with the classical reasoner*) -goal Ord.thy "!!x::'a::order. x = y ==> x <= y"; +Goal "!!x::'a::order. x = y ==> x <= y"; by (etac ssubst 1); by (rtac order_refl 1); qed "order_eq_refl"; -goal Ord.thy "~ x < (x::'a::order)"; +Goal "~ x < (x::'a::order)"; by (simp_tac (simpset() addsimps [order_less_le]) 1); qed "order_less_irrefl"; AddIffs [order_less_irrefl]; -goal thy "(x::'a::order) <= y = (x < y | x = y)"; +Goal "(x::'a::order) <= y = (x < y | x = y)"; by (simp_tac (simpset() addsimps [order_less_le]) 1); by (Blast_tac 1); qed "order_le_less"; (** min **) -goalw thy [min_def] "!!least. (!!x. least <= x) ==> min least x = least"; +Goalw [min_def] "!!least. (!!x. least <= x) ==> min least x = least"; by (Asm_simp_tac 1); qed "min_leastL"; @@ -56,31 +56,31 @@ section "Linear/Total Orders"; -goal thy "!!x::'a::linorder. x 0 < i^n"; +Goal "!! i. 0 < i ==> 0 < i^n"; by (induct_tac "n" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [zero_less_mult_iff]))); qed "zero_less_power"; -goalw thy [dvd_def] "!!m::nat. m<=n ==> i^m dvd i^n"; +Goalw [dvd_def] "!!m::nat. m<=n ==> i^m dvd i^n"; by (etac (not_less_iff_le RS iffD2 RS add_diff_inverse RS subst) 1); by (asm_simp_tac (simpset() addsimps [power_add]) 1); by (Blast_tac 1); qed "le_imp_power_dvd"; -goal thy "!! i r. [| 0 < i; i^m < i^n |] ==> m m i k^i dvd n"; +Goal "!!n. k^j dvd n --> i k^i dvd n"; by (induct_tac "j" 1); by (ALLGOALS (simp_tac (simpset() addsimps [less_Suc_eq]))); by (stac mult_commute 1); @@ -50,16 +50,16 @@ (*** Binomial Coefficients, following Andy Gordon and Florian Kammueller ***) -goal thy "(n choose 0) = 1"; +Goal "(n choose 0) = 1"; by (exhaust_tac "n" 1); by (ALLGOALS Asm_simp_tac); qed "binomial_n_0"; -goal thy "(0 choose Suc k) = 0"; +Goal "(0 choose Suc k) = 0"; by (Simp_tac 1); qed "binomial_0_Suc"; -goal thy "(Suc n choose Suc k) = (n choose k) + (n choose Suc k)"; +Goal "(Suc n choose Suc k) = (n choose k) + (n choose Suc k)"; by (Simp_tac 1); qed "binomial_Suc_Suc"; @@ -67,25 +67,25 @@ Delsimps [binomial_0, binomial_Suc]; -goal thy "!k. n (n choose k) = 0"; +Goal "!k. n (n choose k) = 0"; by (induct_tac "n" 1); by (ALLGOALS (rtac allI THEN' exhaust_tac "k")); by (ALLGOALS Asm_simp_tac); qed_spec_mp "less_imp_binomial_eq_0"; -goal thy "(n choose n) = 1"; +Goal "(n choose n) = 1"; by (induct_tac "n" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_imp_binomial_eq_0]))); qed "binomial_n_n"; Addsimps [binomial_n_n]; -goal thy "(Suc n choose n) = Suc n"; +Goal "(Suc n choose n) = Suc n"; by (induct_tac "n" 1); by (ALLGOALS Asm_simp_tac); qed "binomial_Suc_n"; Addsimps [binomial_Suc_n]; -goal thy "(n choose 1) = n"; +Goal "(n choose 1) = n"; by (induct_tac "n" 1); by (ALLGOALS Asm_simp_tac); qed "binomial_1"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Prod.ML --- a/src/HOL/Prod.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Prod.ML Mon Jun 22 17:26:46 1998 +0200 @@ -9,7 +9,7 @@ open Prod; (*This counts as a non-emptiness result for admitting 'a * 'b as a type*) -goalw Prod.thy [Prod_def] "Pair_Rep a b : Prod"; +Goalw [Prod_def] "Pair_Rep a b : Prod"; by (EVERY1 [rtac CollectI, rtac exI, rtac exI, rtac refl]); qed "ProdI"; @@ -19,7 +19,7 @@ rtac conjI, rtac refl, rtac refl]); qed "Pair_Rep_inject"; -goal Prod.thy "inj_on Abs_Prod Prod"; +Goal "inj_on Abs_Prod Prod"; by (rtac inj_on_inverseI 1); by (etac Abs_Prod_inverse 1); qed "inj_on_Abs_Prod"; @@ -30,20 +30,20 @@ by (REPEAT (ares_tac (prems@[ProdI]) 1)); qed "Pair_inject"; -goal Prod.thy "((a,b) = (a',b')) = (a=a' & b=b')"; +Goal "((a,b) = (a',b')) = (a=a' & b=b')"; by (blast_tac (claset() addSEs [Pair_inject]) 1); qed "Pair_eq"; AddIffs [Pair_eq]; -goalw Prod.thy [fst_def] "fst((a,b)) = a"; +Goalw [fst_def] "fst((a,b)) = a"; by (Blast_tac 1); qed "fst_conv"; -goalw Prod.thy [snd_def] "snd((a,b)) = b"; +Goalw [snd_def] "snd((a,b)) = b"; by (Blast_tac 1); qed "snd_conv"; Addsimps [fst_conv, snd_conv]; -goalw Prod.thy [Pair_def] "? x y. p = (x,y)"; +Goalw [Pair_def] "? x y. p = (x,y)"; by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1); by (EVERY1[etac exE, etac exE, rtac exI, rtac exI, rtac (Rep_Prod_inverse RS sym RS trans), etac arg_cong]); @@ -58,7 +58,7 @@ K prune_params_tac]; (* Do not add as rewrite rule: invalidates some proofs in IMP *) -goal Prod.thy "p = (fst(p),snd(p))"; +Goal "p = (fst(p),snd(p))"; by (pair_tac "p" 1); by (Asm_simp_tac 1); qed "surjective_pairing"; @@ -107,29 +107,29 @@ claset_ref() := claset() addSWrapper ("split_all_tac", fn tac2 => split_all_tac ORELSE' tac2); -goal Prod.thy "(!x. P x) = (!a b. P(a,b))"; +Goal "(!x. P x) = (!a b. P(a,b))"; by (Fast_tac 1); qed "split_paired_All"; Addsimps [split_paired_All]; (* AddIffs is not a good idea because it makes Blast_tac loop *) -goal Prod.thy "(? x. P x) = (? a b. P(a,b))"; +Goal "(? x. P x) = (? a b. P(a,b))"; by (Fast_tac 1); qed "split_paired_Ex"; Addsimps [split_paired_Ex]; -goalw Prod.thy [split_def] "split c (a,b) = c a b"; +Goalw [split_def] "split c (a,b) = c a b"; by (Simp_tac 1); qed "split"; Addsimps [split]; -goal Prod.thy "split Pair p = p"; +Goal "split Pair p = p"; by (pair_tac "p" 1); by (Simp_tac 1); qed "split_Pair"; (*unused: val surjective_pairing2 = split_Pair RS sym;*) -goal Prod.thy "!!s t. (s=t) = (fst(s)=fst(t) & snd(s)=snd(t))"; +Goal "!!s t. (s=t) = (fst(s)=fst(t) & snd(s)=snd(t))"; by (split_all_tac 1); by (Asm_simp_tac 1); qed "Pair_fst_snd_eq"; @@ -162,7 +162,7 @@ (K [stac surjective_pairing 1, stac split 1, rtac refl 1]); (*For use with split_tac and the simplifier*) -goal Prod.thy "R (split c p) = (! x y. p = (x,y) --> R (c x y))"; +Goal "R (split c p) = (! x y. p = (x,y) --> R (c x y))"; by (stac surjective_pairing 1); by (stac split 1); by (Blast_tac 1); @@ -174,7 +174,7 @@ the current goal contains one of those constants *) -goal Prod.thy "R (split c p) = (~(? x y. p = (x,y) & (~R (c x y))))"; +Goal "R (split c p) = (~(? x y. p = (x,y) & (~R (c x y))))"; by (stac split_split 1); by (Simp_tac 1); qed "expand_split_asm"; @@ -184,12 +184,12 @@ (*These rules are for use with blast_tac. Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*) -goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> c a b |] ==> split c p"; +Goal "!!p. [| !!a b. p=(a,b) ==> c a b |] ==> split c p"; by (split_all_tac 1); by (Asm_simp_tac 1); qed "splitI2"; -goal Prod.thy "!!a b c. c a b ==> split c (a,b)"; +Goal "!!a b c. c a b ==> split c (a,b)"; by (Asm_simp_tac 1); qed "splitI"; @@ -204,15 +204,15 @@ rtac (split_beta RS subst) 1, rtac (hd prems) 1]); -goal Prod.thy "!!R a b. split R (a,b) ==> R a b"; +Goal "!!R a b. split R (a,b) ==> R a b"; by (etac (split RS iffD1) 1); qed "splitD"; -goal Prod.thy "!!a b c. z: c a b ==> z: split c (a,b)"; +Goal "!!a b c. z: c a b ==> z: split c (a,b)"; by (Asm_simp_tac 1); qed "mem_splitI"; -goal Prod.thy "!!p. [| !!a b. p=(a,b) ==> z: c a b |] ==> z: split c p"; +Goal "!!p. [| !!a b. p=(a,b) ==> z: c a b |] ==> z: split c p"; by (split_all_tac 1); by (Asm_simp_tac 1); qed "mem_splitI2"; @@ -226,13 +226,13 @@ AddSEs [splitE, mem_splitE]; (* allows simplifications of nested splits in case of independent predicates *) -goal Prod.thy "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"; +Goal "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"; by (rtac ext 1); by (Blast_tac 1); qed "split_part"; Addsimps [split_part]; -goal Prod.thy "(@(x',y'). x = x' & y = y') = (x,y)"; +Goal "(@(x',y'). x = x' & y = y') = (x,y)"; by (Blast_tac 1); qed "Eps_split_eq"; Addsimps [Eps_split_eq]; @@ -241,7 +241,7 @@ but cannot be used as rewrite rule: ### Cannot add premise as rewrite rule because it contains (type) unknowns: ### ?y = .x -goal Prod.thy "!!P. [| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)"; +Goal "!!P. [| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)"; by (rtac select_equality 1); by ( Simp_tac 1); by (split_all_tac 1); @@ -251,19 +251,19 @@ (*** prod_fun -- action of the product functor upon functions ***) -goalw Prod.thy [prod_fun_def] "prod_fun f g (a,b) = (f(a),g(b))"; +Goalw [prod_fun_def] "prod_fun f g (a,b) = (f(a),g(b))"; by (rtac split 1); qed "prod_fun"; Addsimps [prod_fun]; -goal Prod.thy +Goal "prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))"; by (rtac ext 1); by (pair_tac "x" 1); by (Asm_simp_tac 1); qed "prod_fun_compose"; -goal Prod.thy "prod_fun (%x. x) (%y. y) = (%z. z)"; +Goal "prod_fun (%x. x) (%y. y) = (%z. z)"; by (rtac ext 1); by (pair_tac "z" 1); by (Asm_simp_tac 1); @@ -340,7 +340,7 @@ Addsimps [Sigma_empty1,Sigma_empty2]; -goal Prod.thy "((a,b): Sigma A B) = (a:A & b:B(a))"; +Goal "((a,b): Sigma A B) = (a:A & b:B(a))"; by (Blast_tac 1); qed "mem_Sigma_iff"; AddIffs [mem_Sigma_iff]; @@ -350,7 +350,7 @@ Addsimps [Collect_split]; (*Suggested by Pierre Chartier*) -goal Prod.thy +Goal "(UN (a,b):(A Times B). E a Times F b) = (UNION A E) Times (UNION B F)"; by (Blast_tac 1); qed "UNION_Times_distrib"; @@ -393,7 +393,7 @@ (** Exhaustion rule for unit -- a degenerate form of induction **) -goalw Prod.thy [Unity_def] +Goalw [Unity_def] "u = ()"; by (stac (rewrite_rule [unit_def] Rep_unit RS singletonD RS sym) 1); by (rtac (Rep_unit_inverse RS sym) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Quot/FRACT.ML --- a/src/HOL/Quot/FRACT.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Quot/FRACT.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,13 +6,13 @@ *) open FRACT; -goalw thy [per_def,per_NP_def] +Goalw [per_def,per_NP_def] "(op ===)=(%x y. fst(rep_NP x)*snd(rep_NP y)=fst(rep_NP y)*snd(rep_NP x))"; fr refl; qed "inst_NP_per"; -goalw thy [half_def] "half = <[abs_NP(n,2*n)]>"; +Goalw [half_def] "half = <[abs_NP(n,2*n)]>"; fr per_class_eqI; by (simp_tac (simpset() addsimps [inst_NP_per]) 1); qed "test"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Quot/HQUOT.ML --- a/src/HOL/Quot/HQUOT.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Quot/HQUOT.ML Mon Jun 22 17:26:46 1998 +0200 @@ -7,7 +7,7 @@ open HQUOT; (* first prove some helpful lemmas *) -goalw thy [quot_def] "{y. y===x}:quot"; +Goalw [quot_def] "{y. y===x}:quot"; by (Asm_simp_tac 1); by (fast_tac (set_cs addIs [per_sym]) 1); qed "per_class_rep_quot"; @@ -28,7 +28,7 @@ by (asm_full_simp_tac (HOL_ss addsimps [Rep_quot,Rep_quot_inverse]) 1); qed "all_q"; -goal thy "? s. s:quot & x=Abs_quot s"; +Goal "? s. s:quot & x=Abs_quot s"; by (res_inst_tac [("x","Rep_quot x")] exI 1); by (asm_full_simp_tac (HOL_ss addsimps [Rep_quot,Rep_quot_inverse]) 1); qed "exh_q"; @@ -64,7 +64,7 @@ by (safe_tac set_cs); qed "per_class_eqE"; -goal thy "<[(x::'a::er)]>=<[y]>=x===y"; +Goal "<[(x::'a::er)]>=<[y]>=x===y"; by (rtac iffI 1); by (etac er_class_eqE 1); by (etac per_class_eqI 1); @@ -106,14 +106,14 @@ by (etac per_class_neqI 1);by (assume_tac 1); qed "per_class_not"; -goal thy "<[(x::'a::er)]>~=<[y]>=(~x===y)"; +Goal "<[(x::'a::er)]>~=<[y]>=(~x===y)"; by (rtac iffI 1); by (etac per_class_neqE 1); by (etac er_class_neqI 1); qed "er_class_not"; (* exhaustiveness and induction *) -goalw thy [peclass_def] "? s. x=<[s]>"; +Goalw [peclass_def] "? s. x=<[s]>"; by (rtac all_q 1); by (strip_tac 1); by (asm_full_simp_tac (HOL_ss addsimps [mem_Collect_eq,quot_def]) 1); @@ -135,7 +135,7 @@ qed "per_class_all"; (* theorems for any_in *) -goalw thy [any_in_def,peclass_def] "any_in<[(s::'a::er)]>===s"; +Goalw [any_in_def,peclass_def] "any_in<[(s::'a::er)]>===s"; fr selectI2; fr refl; by (fold_goals_tac [peclass_def]); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Quot/NPAIR.ML --- a/src/HOL/Quot/NPAIR.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Quot/NPAIR.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,7 +6,7 @@ *) open NPAIR; -goalw thy [rep_NP_def] "rep_NP(abs_NP np) = np"; +Goalw [rep_NP_def] "rep_NP(abs_NP np) = np"; by Auto_tac; qed "rep_abs_NP"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Quot/PER.ML --- a/src/HOL/Quot/PER.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Quot/PER.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,12 +6,12 @@ *) open PER; -goalw thy [fun_per_def,per_def] "f===g=(!x y. x:D&y:D&x===y-->f x===g y)"; +Goalw [fun_per_def,per_def] "f===g=(!x y. x:D&y:D&x===y-->f x===g y)"; by (rtac refl 1); qed "inst_fun_per"; (* Witness that quot is not empty *) -goal thy "?z:{s.? r.!y. y:s=y===r}"; +Goal "?z:{s.? r.!y. y:s=y===r}"; fr CollectI; by (res_inst_tac [("x","x")] exI 1); by (rtac allI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Quot/PER0.ML --- a/src/HOL/Quot/PER0.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Quot/PER0.ML Mon Jun 22 17:26:46 1998 +0200 @@ -18,13 +18,13 @@ by (assume_tac 1); qed "per_trans"; -goalw thy [per_def] "(x::'a::er) === x"; +Goalw [per_def] "(x::'a::er) === x"; by (rtac ax_er_refl 1); qed "er_refl"; (* now everything works without axclasses *) -goal thy "x===y=y===x"; +Goal "x===y=y===x"; by (rtac iffI 1); by (etac per_sym 1); by (etac per_sym 1); @@ -52,13 +52,13 @@ by (fast_tac set_cs 1); qed "DomainI"; -goal thy "x:D = x===x"; +Goal "x:D = x===x"; by (rtac iffI 1); by (etac DomainD 1); by (etac DomainI 1); qed "DomainEq"; -goal thy "(~ x === y) = (~ y === x)"; +Goal "(~ x === y) = (~ y === x)"; by (rtac (per_sym2 RS arg_cong) 1); qed "per_not_sym"; @@ -88,7 +88,7 @@ qed "notDomainE2"; (* theorems for equivalence relations *) -goal thy "(x::'a::er) : D"; +Goal "(x::'a::er) : D"; by (rtac DomainI 1); by (rtac er_refl 1); qed "er_Domain"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/RelPow.ML --- a/src/HOL/RelPow.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/RelPow.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,28 +6,28 @@ open RelPow; -goal RelPow.thy "!!R:: ('a*'a)set. R^1 = R"; +Goal "!!R:: ('a*'a)set. R^1 = R"; by (Simp_tac 1); qed "rel_pow_1"; Addsimps [rel_pow_1]; -goal RelPow.thy "(x,x) : R^0"; +Goal "(x,x) : R^0"; by (Simp_tac 1); qed "rel_pow_0_I"; -goal RelPow.thy "!!R. [| (x,y) : R^n; (y,z):R |] ==> (x,z):R^(Suc n)"; +Goal "!!R. [| (x,y) : R^n; (y,z):R |] ==> (x,z):R^(Suc n)"; by (Simp_tac 1); by (Blast_tac 1); qed "rel_pow_Suc_I"; -goal RelPow.thy "!z. (x,y) : R --> (y,z):R^n --> (x,z):R^(Suc n)"; +Goal "!z. (x,y) : R --> (y,z):R^n --> (x,z):R^(Suc n)"; by (nat_ind_tac "n" 1); by (Simp_tac 1); by (Asm_full_simp_tac 1); by (Blast_tac 1); qed_spec_mp "rel_pow_Suc_I2"; -goal RelPow.thy "!!R. [| (x,y) : R^0; x=y ==> P |] ==> P"; +Goal "!!R. [| (x,y) : R^0; x=y ==> P |] ==> P"; by (Asm_full_simp_tac 1); qed "rel_pow_0_E"; @@ -51,14 +51,14 @@ by (REPEAT(ares_tac [p3] 1)); qed "rel_pow_E"; -goal RelPow.thy "!x z. (x,z):R^(Suc n) --> (? y. (x,y):R & (y,z):R^n)"; +Goal "!x z. (x,z):R^(Suc n) --> (? y. (x,y):R & (y,z):R^n)"; by (nat_ind_tac "n" 1); by (blast_tac (claset() addIs [rel_pow_0_I] addEs [rel_pow_0_E,rel_pow_Suc_E]) 1); by (blast_tac (claset() addIs [rel_pow_Suc_I] addEs[rel_pow_0_E,rel_pow_Suc_E]) 1); qed_spec_mp "rel_pow_Suc_D2"; -goal RelPow.thy +Goal "!x y z. (x,y) : R^n & (y,z) : R --> (? w. (x,w) : R & (w,z) : R^n)"; by (nat_ind_tac "n" 1); by (fast_tac (claset() addss (simpset())) 1); @@ -83,30 +83,30 @@ by (etac conjunct2 1); qed "rel_pow_E2"; -goal RelPow.thy "!!p. p:R^* ==> p : (UN n. R^n)"; +Goal "!!p. p:R^* ==> p : (UN n. R^n)"; by (split_all_tac 1); by (etac rtrancl_induct 1); by (ALLGOALS (blast_tac (claset() addIs [rel_pow_0_I,rel_pow_Suc_I]))); qed "rtrancl_imp_UN_rel_pow"; -goal RelPow.thy "!y. (x,y):R^n --> (x,y):R^*"; +Goal "!y. (x,y):R^n --> (x,y):R^*"; by (nat_ind_tac "n" 1); by (blast_tac (claset() addIs [rtrancl_refl] addEs [rel_pow_0_E]) 1); by (blast_tac (claset() addEs [rel_pow_Suc_E] addIs [rtrancl_into_rtrancl]) 1); val lemma = result() RS spec RS mp; -goal RelPow.thy "!!p. p:R^n ==> p:R^*"; +Goal "!!p. p:R^n ==> p:R^*"; by (split_all_tac 1); by (etac lemma 1); qed "rel_pow_imp_rtrancl"; -goal RelPow.thy "R^* = (UN n. R^n)"; +Goal "R^* = (UN n. R^n)"; by (blast_tac (claset() addIs [rtrancl_imp_UN_rel_pow, rel_pow_imp_rtrancl]) 1); qed "rtrancl_is_UN_rel_pow"; -goal RelPow.thy "!!r::('a * 'a)set. Univalent r ==> Univalent (r^n)"; +Goal "!!r::('a * 'a)set. Univalent r ==> Univalent (r^n)"; by (rtac UnivalentI 1); by (induct_tac "n" 1); by (Simp_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Relation.ML --- a/src/HOL/Relation.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Relation.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,7 +8,7 @@ (** Identity relation **) -goalw thy [id_def] "(a,a) : id"; +Goalw [id_def] "(a,a) : id"; by (Blast_tac 1); qed "idI"; @@ -20,7 +20,7 @@ by (eresolve_tac prems 1); qed "idE"; -goalw thy [id_def] "(a,b):id = (a=b)"; +Goalw [id_def] "(a,b):id = (a=b)"; by (Blast_tac 1); qed "pair_in_id_conv"; Addsimps [pair_in_id_conv]; @@ -28,7 +28,7 @@ (** Composition of two relations **) -goalw thy [comp_def] +Goalw [comp_def] "!!r s. [| (a,b):s; (b,c):r |] ==> (a,c) : r O s"; by (Blast_tac 1); qed "compI"; @@ -54,25 +54,25 @@ AddIs [compI, idI]; AddSEs [compE, idE]; -goal thy "R O id = R"; +Goal "R O id = R"; by (Fast_tac 1); qed "R_O_id"; -goal thy "id O R = R"; +Goal "id O R = R"; by (Fast_tac 1); qed "id_O_R"; Addsimps [R_O_id,id_O_R]; -goal thy "(R O S) O T = R O (S O T)"; +Goal "(R O S) O T = R O (S O T)"; by (Blast_tac 1); qed "O_assoc"; -goal thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +Goal "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; by (Blast_tac 1); qed "comp_mono"; -goal thy +Goal "!!r s. [| s <= A Times B; r <= B Times C |] ==> (r O s) <= A Times C"; by (Blast_tac 1); qed "comp_subset_Sigma"; @@ -84,24 +84,24 @@ by (REPEAT (ares_tac (prems@[allI,impI]) 1)); qed "transI"; -goalw thy [trans_def] +Goalw [trans_def] "!!r. [| trans(r); (a,b):r; (b,c):r |] ==> (a,c):r"; by (Blast_tac 1); qed "transD"; (** Natural deduction for r^-1 **) -goalw thy [converse_def] "!!a b r. ((a,b): r^-1) = ((b,a):r)"; +Goalw [converse_def] "!!a b r. ((a,b): r^-1) = ((b,a):r)"; by (Simp_tac 1); qed "converse_iff"; AddIffs [converse_iff]; -goalw thy [converse_def] "!!a b r. (a,b):r ==> (b,a): r^-1"; +Goalw [converse_def] "!!a b r. (a,b):r ==> (b,a): r^-1"; by (Simp_tac 1); qed "converseI"; -goalw thy [converse_def] "!!a b r. (a,b) : r^-1 ==> (b,a) : r"; +Goalw [converse_def] "!!a b r. (a,b) : r^-1 ==> (b,a) : r"; by (Blast_tac 1); qed "converseD"; @@ -117,16 +117,16 @@ AddSEs [converseE]; -goalw thy [converse_def] "(r^-1)^-1 = r"; +Goalw [converse_def] "(r^-1)^-1 = r"; by (Blast_tac 1); qed "converse_converse"; Addsimps [converse_converse]; -goal thy "(r O s)^-1 = s^-1 O r^-1"; +Goal "(r O s)^-1 = s^-1 O r^-1"; by (Blast_tac 1); qed "converse_comp"; -goal thy "id^-1 = id"; +Goal "id^-1 = id"; by (Blast_tac 1); qed "converse_id"; Addsimps [converse_id]; @@ -149,7 +149,7 @@ AddIs [DomainI]; AddSEs [DomainE]; -goal thy "Domain id = UNIV"; +Goal "Domain id = UNIV"; by (Blast_tac 1); qed "Domain_id"; Addsimps [Domain_id]; @@ -169,7 +169,7 @@ AddIs [RangeI]; AddSEs [RangeE]; -goal thy "Range id = UNIV"; +Goal "Range id = UNIV"; by (Blast_tac 1); qed "Range_id"; Addsimps [Range_id]; @@ -213,7 +213,7 @@ Addsimps [Image_empty]; -goal thy "id ^^ A = A"; +Goal "id ^^ A = A"; by (Blast_tac 1); qed "Image_id"; @@ -232,7 +232,7 @@ (REPEAT (eresolve_tac [asm_rl, ImageE, subsetD RS SigmaD2] 1)) ]); (*NOT suitable for rewriting*) -goal thy "r^^B = (UN y: B. r^^{y})"; +Goal "r^^B = (UN y: B. r^^{y})"; by (Blast_tac 1); qed "Image_eq_UN"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Set.ML --- a/src/HOL/Set.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Set.ML Mon Jun 22 17:26:46 1998 +0200 @@ -13,7 +13,7 @@ Addsimps [Collect_mem_eq]; AddIffs [mem_Collect_eq]; -goal Set.thy "!!a. P(a) ==> a : {x. P(x)}"; +Goal "!!a. P(a) ==> a : {x. P(x)}"; by (Asm_simp_tac 1); qed "CollectI"; @@ -82,12 +82,12 @@ AddSEs [bexE]; (*Trival rewrite rule*) -goal Set.thy "(! x:A. P) = ((? x. x:A) --> P)"; +Goal "(! x:A. P) = ((? x. x:A) --> P)"; by (simp_tac (simpset() addsimps [Ball_def]) 1); qed "ball_triv"; (*Dual form for existentials*) -goal Set.thy "(? x:A. P) = ((? x. x:A) & P)"; +Goal "(? x:A. P) = ((? x. x:A) & P)"; by (simp_tac (simpset() addsimps [Bex_def]) 1); qed "bex_triv"; @@ -227,11 +227,11 @@ (** Eta-contracting these two rules (to remove P) causes them to be ignored because of their interaction with congruence rules. **) -goalw Set.thy [Ball_def] "Ball UNIV P = All P"; +Goalw [Ball_def] "Ball UNIV P = All P"; by (Simp_tac 1); qed "ball_UNIV"; -goalw Set.thy [Bex_def] "Bex UNIV P = Ex P"; +Goalw [Bex_def] "Bex UNIV P = Ex P"; by (Simp_tac 1); qed "bex_UNIV"; Addsimps [ball_UNIV, bex_UNIV]; @@ -259,16 +259,16 @@ qed_goal "equals0D" Set.thy "!!a. [| A={}; a:A |] ==> P" (fn _ => [ (Blast_tac 1) ]); -goalw Set.thy [Ball_def] "Ball {} P = True"; +Goalw [Ball_def] "Ball {} P = True"; by (Simp_tac 1); qed "ball_empty"; -goalw Set.thy [Bex_def] "Bex {} P = False"; +Goalw [Bex_def] "Bex {} P = False"; by (Simp_tac 1); qed "bex_empty"; Addsimps [ball_empty, bex_empty]; -goal thy "UNIV ~= {}"; +Goal "UNIV ~= {}"; by (blast_tac (claset() addEs [equalityE]) 1); qed "UNIV_not_empty"; AddIffs [UNIV_not_empty]; @@ -325,11 +325,11 @@ Addsimps [Un_iff]; -goal Set.thy "!!c. c:A ==> c : A Un B"; +Goal "!!c. c:A ==> c : A Un B"; by (Asm_simp_tac 1); qed "UnI1"; -goal Set.thy "!!c. c:B ==> c : A Un B"; +Goal "!!c. c:B ==> c : A Un B"; by (Asm_simp_tac 1); qed "UnI2"; @@ -356,15 +356,15 @@ Addsimps [Int_iff]; -goal Set.thy "!!c. [| c:A; c:B |] ==> c : A Int B"; +Goal "!!c. [| c:A; c:B |] ==> c : A Int B"; by (Asm_simp_tac 1); qed "IntI"; -goal Set.thy "!!c. c : A Int B ==> c:A"; +Goal "!!c. c : A Int B ==> c:A"; by (Asm_full_simp_tac 1); qed "IntD1"; -goal Set.thy "!!c. c : A Int B ==> c:B"; +Goal "!!c. c : A Int B ==> c:B"; by (Asm_full_simp_tac 1); qed "IntD2"; @@ -436,7 +436,7 @@ qed_goal "singletonI" Set.thy "a : {a}" (fn _=> [ (rtac insertI1 1) ]); -goal Set.thy "!!a. b : {a} ==> b=a"; +Goal "!!a. b : {a} ==> b=a"; by (Blast_tac 1); qed "singletonD"; @@ -445,7 +445,7 @@ qed_goal "singleton_iff" thy "(b : {a}) = (b=a)" (fn _ => [Blast_tac 1]); -goal Set.thy "!!a b. {a}={b} ==> a=b"; +Goal "!!a b. {a}={b} ==> a=b"; by (blast_tac (claset() addEs [equalityE]) 1); qed "singleton_inject"; @@ -454,7 +454,7 @@ AddSDs [singleton_inject]; AddSEs [singletonE]; -goal Set.thy "{x. x=a} = {a}"; +Goal "{x. x=a} = {a}"; by (Blast_tac 1); qed "singleton_conv"; Addsimps [singleton_conv]; @@ -462,14 +462,14 @@ section "Unions of families -- UNION x:A. B(x) is Union(B``A)"; -goalw Set.thy [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))"; +Goalw [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))"; by (Blast_tac 1); qed "UN_iff"; Addsimps [UN_iff]; (*The order of the premises presupposes that A is rigid; b may be flexible*) -goal Set.thy "!!b. [| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; +Goal "!!b. [| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; by Auto_tac; qed "UN_I"; @@ -493,7 +493,7 @@ section "Intersections of families -- INTER x:A. B(x) is Inter(B``A)"; -goalw Set.thy [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))"; +Goalw [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))"; by Auto_tac; qed "INT_iff"; @@ -504,7 +504,7 @@ by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); qed "INT_I"; -goal Set.thy "!!b. [| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; +Goal "!!b. [| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; by Auto_tac; qed "INT_D"; @@ -529,14 +529,14 @@ section "Union"; -goalw Set.thy [Union_def] "(A : Union(C)) = (EX X:C. A:X)"; +Goalw [Union_def] "(A : Union(C)) = (EX X:C. A:X)"; by (Blast_tac 1); qed "Union_iff"; Addsimps [Union_iff]; (*The order of the premises presupposes that C is rigid; A may be flexible*) -goal Set.thy "!!X. [| X:C; A:X |] ==> A : Union(C)"; +Goal "!!X. [| X:C; A:X |] ==> A : Union(C)"; by Auto_tac; qed "UnionI"; @@ -552,7 +552,7 @@ section "Inter"; -goalw Set.thy [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)"; +Goalw [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)"; by (Blast_tac 1); qed "Inter_iff"; @@ -565,7 +565,7 @@ (*A "destruct" rule -- every X in C contains A as an element, but A:X can hold when X:C does not! This rule is analogous to "spec". *) -goal Set.thy "!!X. [| A : Inter(C); X:C |] ==> A:X"; +Goal "!!X. [| A : Inter(C); X:C |] ==> A:X"; by Auto_tac; qed "InterD"; @@ -600,20 +600,20 @@ AddIs [image_eqI]; AddSEs [imageE]; -goalw thy [o_def] "(f o g)``r = f``(g``r)"; +Goalw [o_def] "(f o g)``r = f``(g``r)"; by (Blast_tac 1); qed "image_compose"; -goal thy "f``(A Un B) = f``A Un f``B"; +Goal "f``(A Un B) = f``A Un f``B"; by (Blast_tac 1); qed "image_Un"; -goal thy "(z : f``A) = (EX x:A. z = f x)"; +Goal "(z : f``A) = (EX x:A. z = f x)"; by (Blast_tac 1); qed "image_iff"; (*This rewrite rule would confuse users if made default.*) -goal thy "(f``A <= B) = (ALL x:A. f(x): B)"; +Goal "(f``A <= B) = (ALL x:A. f(x): B)"; by (Blast_tac 1); qed "image_subset_iff"; @@ -626,7 +626,7 @@ (*** Range of a function -- just a translation for image! ***) -goal thy "!!b. b=f(x) ==> b : range(f)"; +Goal "!!b. b=f(x) ==> b : range(f)"; by (EVERY1 [etac image_eqI, rtac UNIV_I]); bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI)); @@ -663,7 +663,7 @@ mem_Collect_eq, UN_iff, Union_iff, INT_iff, Inter_iff]; (*Not for Addsimps -- it can cause goals to blow up!*) -goal Set.thy "(a : (if Q then x else y)) = ((Q --> a:x) & (~Q --> a : y))"; +Goal "(a : (if Q then x else y)) = ((Q --> a:x) & (~Q --> a : y))"; by (Simp_tac 1); qed "mem_if"; @@ -677,11 +677,11 @@ (*** < ***) -goalw Set.thy [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A A (x ~: A) & A<=B | x:A & A-{x} (M, M$N) : pred_sexp"; by (Blast_tac 1); qed "pred_sexpI1"; -goalw Sexp.thy [pred_sexp_def] +Goalw [pred_sexp_def] "!!M. [| M: sexp; N: sexp |] ==> (N, M$N) : pred_sexp"; by (Blast_tac 1); qed "pred_sexpI2"; @@ -89,7 +89,7 @@ by (REPEAT (eresolve_tac ([asm_rl,emptyE,insertE,UN_E]@prems) 1)); qed "pred_sexpE"; -goal Sexp.thy "wf(pred_sexp)"; +Goal "wf(pred_sexp)"; by (rtac (pred_sexp_subset_Sigma RS wfI) 1); by (etac sexp.induct 1); by (ALLGOALS (blast_tac (claset() addSEs [allE, pred_sexpE]))); @@ -98,7 +98,7 @@ (*** sexp_rec -- by wf recursion on pred_sexp ***) -goal Sexp.thy +Goal "(%M. sexp_rec M c d e) = wfrec pred_sexp \ \ (%g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2)))"; by (simp_tac (HOL_ss addsimps [sexp_rec_def]) 1); @@ -114,17 +114,17 @@ (** conversion rules **) -goal Sexp.thy "sexp_rec (Leaf a) c d h = c(a)"; +Goal "sexp_rec (Leaf a) c d h = c(a)"; by (stac sexp_rec_unfold 1); by (rtac sexp_case_Leaf 1); qed "sexp_rec_Leaf"; -goal Sexp.thy "sexp_rec (Numb k) c d h = d(k)"; +Goal "sexp_rec (Numb k) c d h = d(k)"; by (stac sexp_rec_unfold 1); by (rtac sexp_case_Numb 1); qed "sexp_rec_Numb"; -goal Sexp.thy "!!M. [| M: sexp; N: sexp |] ==> \ +Goal "!!M. [| M: sexp; N: sexp |] ==> \ \ sexp_rec (M$N) c d h = h M N (sexp_rec M c d h) (sexp_rec N c d h)"; by (rtac (sexp_rec_unfold RS trans) 1); by (asm_simp_tac (simpset() addsimps [pred_sexpI1, pred_sexpI2]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Subst/Subst.ML --- a/src/HOL/Subst/Subst.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Subst/Subst.ML Mon Jun 22 17:26:46 1998 +0200 @@ -11,19 +11,19 @@ (**** Substitutions ****) -goal Subst.thy "t <| [] = t"; +Goal "t <| [] = t"; by (induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "subst_Nil"; Addsimps [subst_Nil]; -goal Subst.thy "t <: u --> t <| s <: u <| s"; +Goal "t <: u --> t <| s <: u <| s"; by (induct_tac "u" 1); by (ALLGOALS Asm_simp_tac); qed_spec_mp "subst_mono"; -goal Subst.thy "~ (Var(v) <: t) --> t <| (v,t <| s) # s = t <| s"; +Goal "~ (Var(v) <: t) --> t <| (v,t <| s) # s = t <| s"; by (case_tac "t = Var(v)" 1); by (etac rev_mp 2); by (res_inst_tac [("P", @@ -33,14 +33,14 @@ by (Blast_tac 1); qed_spec_mp "Var_not_occs"; -goal Subst.thy +Goal "(t <|r = t <|s) = (! v. v : vars_of(t) --> Var(v) <|r = Var(v) <|s)"; by (induct_tac "t" 1); by (ALLGOALS Asm_full_simp_tac); by (ALLGOALS Blast_tac); qed "agreement"; -goal Subst.thy "~ v: vars_of(t) --> t <| (v,u)#s = t <| s"; +Goal "~ v: vars_of(t) --> t <| (v,u)#s = t <| s"; by (simp_tac (simpset() addsimps [agreement]) 1); qed_spec_mp"repl_invariance"; @@ -53,7 +53,7 @@ (**** Equality between Substitutions ****) -goalw Subst.thy [subst_eq_def] "r =$= s = (! t. t <| r = t <| s)"; +Goalw [subst_eq_def] "r =$= s = (! t. t <| r = t <| s)"; by (Simp_tac 1); qed "subst_eq_iff"; @@ -96,18 +96,18 @@ end; -goal Subst.thy "s <> [] = s"; +Goal "s <> [] = s"; by (alist_ind_tac "s" 1); by (ALLGOALS Asm_simp_tac); qed "comp_Nil"; Addsimps [comp_Nil]; -goal Subst.thy "s =$= s <> []"; +Goal "s =$= s <> []"; by (Simp_tac 1); qed "subst_comp_Nil"; -goal Subst.thy "(t <| r <> s) = (t <| r <| s)"; +Goal "(t <| r <> s) = (t <| r <| s)"; by (induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); by (alist_ind_tac "r" 1); @@ -116,17 +116,17 @@ Addsimps [subst_comp]; -goal Subst.thy "(q <> r) <> s =$= q <> (r <> s)"; +Goal "(q <> r) <> s =$= q <> (r <> s)"; by (simp_tac (simpset() addsimps [subst_eq_iff]) 1); qed "comp_assoc"; -goal Subst.thy "!!s. [| theta =$= theta1; sigma =$= sigma1|] ==> \ +Goal "!!s. [| theta =$= theta1; sigma =$= sigma1|] ==> \ \ (theta <> sigma) =$= (theta1 <> sigma1)"; by (asm_full_simp_tac (simpset() addsimps [subst_eq_def]) 1); qed "subst_cong"; -goal Subst.thy "(w, Var(w) <| s) # s =$= s"; +Goal "(w, Var(w) <| s) # s =$= s"; by (simp_tac (simpset() addsimps [subst_eq_iff]) 1); by (rtac allI 1); by (induct_tac "t" 1); @@ -134,49 +134,49 @@ qed "Cons_trivial"; -goal Subst.thy "!!s. q <> r =$= s ==> t <| q <| r = t <| s"; +Goal "!!s. q <> r =$= s ==> t <| q <| r = t <| s"; by (asm_full_simp_tac (simpset() addsimps [subst_eq_iff]) 1); qed "comp_subst_subst"; (**** Domain and range of Substitutions ****) -goal Subst.thy "(v : sdom(s)) = (Var(v) <| s ~= Var(v))"; +Goal "(v : sdom(s)) = (Var(v) <| s ~= Var(v))"; by (alist_ind_tac "s" 1); by (ALLGOALS Asm_simp_tac); by (Blast_tac 1); qed "sdom_iff"; -goalw Subst.thy [srange_def] +Goalw [srange_def] "v : srange(s) = (? w. w : sdom(s) & v : vars_of(Var(w) <| s))"; by (Blast_tac 1); qed "srange_iff"; -goalw thy [empty_def] "(A = {}) = (ALL a.~ a:A)"; +Goalw [empty_def] "(A = {}) = (ALL a.~ a:A)"; by (Blast_tac 1); qed "empty_iff_all_not"; -goal Subst.thy "(t <| s = t) = (sdom(s) Int vars_of(t) = {})"; +Goal "(t <| s = t) = (sdom(s) Int vars_of(t) = {})"; by (induct_tac "t" 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [empty_iff_all_not, sdom_iff]))); by (ALLGOALS Blast_tac); qed "invariance"; -goal Subst.thy "v : sdom(s) --> v : vars_of(t <| s) --> v : srange(s)"; +Goal "v : sdom(s) --> v : vars_of(t <| s) --> v : srange(s)"; by (induct_tac "t" 1); by (case_tac "a : sdom(s)" 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [sdom_iff, srange_iff]))); by (ALLGOALS Blast_tac); qed_spec_mp "Var_in_srange"; -goal Subst.thy +Goal "!!v. [| v : sdom(s); v ~: srange(s) |] ==> v ~: vars_of(t <| s)"; by (blast_tac (claset() addIs [Var_in_srange]) 1); qed "Var_elim"; -goal Subst.thy "v : vars_of(t <| s) --> v : srange(s) | v : vars_of(t)"; +Goal "v : vars_of(t <| s) --> v : srange(s) | v : vars_of(t)"; by (induct_tac "t" 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [sdom_iff,srange_iff]))); by (Blast_tac 2); @@ -184,24 +184,24 @@ by Auto_tac; qed_spec_mp "Var_intro"; -goal Subst.thy +Goal "v : srange(s) --> (? w. w : sdom(s) & v : vars_of(Var(w) <| s))"; by (simp_tac (simpset() addsimps [srange_iff]) 1); qed_spec_mp "srangeD"; -goal Subst.thy +Goal "sdom(s) Int srange(s) = {} = (! t. sdom(s) Int vars_of(t <| s) = {})"; by (simp_tac (simpset() addsimps [empty_iff_all_not]) 1); by (fast_tac (claset() addIs [Var_in_srange] addDs [srangeD]) 1); qed "dom_range_disjoint"; -goal Subst.thy "!!u. ~ u <| s = u ==> (? x. x : sdom(s))"; +Goal "!!u. ~ u <| s = u ==> (? x. x : sdom(s))"; by (full_simp_tac (simpset() addsimps [empty_iff_all_not, invariance]) 1); by (Blast_tac 1); qed "subst_not_empty"; -goal Subst.thy "(M <| [(x, Var x)]) = M"; +Goal "(M <| [(x, Var x)]) = M"; by (induct_tac "M" 1); by (ALLGOALS Asm_simp_tac); qed "id_subst_lemma"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Subst/UTerm.ML --- a/src/HOL/Subst/UTerm.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Subst/UTerm.ML Mon Jun 22 17:26:46 1998 +0200 @@ -12,12 +12,12 @@ (**** vars_of lemmas ****) -goal UTerm.thy "(v : vars_of(Var(w))) = (w=v)"; +Goal "(v : vars_of(Var(w))) = (w=v)"; by (Simp_tac 1); by (fast_tac HOL_cs 1); qed "vars_var_iff"; -goal UTerm.thy "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)"; +Goal "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)"; by (induct_tac "t" 1); by (ALLGOALS Simp_tac); by (fast_tac HOL_cs 1); @@ -25,18 +25,18 @@ (* Not used, but perhaps useful in other proofs *) -goal UTerm.thy "M<:N --> vars_of(M) <= vars_of(N)"; +Goal "M<:N --> vars_of(M) <= vars_of(N)"; by (induct_tac "N" 1); by (ALLGOALS Asm_simp_tac); by (fast_tac set_cs 1); val occs_vars_subset = result(); -goal UTerm.thy "vars_of M Un vars_of N <= (vars_of M Un A) Un (vars_of N Un B)"; +Goal "vars_of M Un vars_of N <= (vars_of M Un A) Un (vars_of N Un B)"; by (Blast_tac 1); val monotone_vars_of = result(); -goal UTerm.thy "finite(vars_of M)"; +Goal "finite(vars_of M)"; by (induct_tac"M" 1); by (ALLGOALS Simp_tac); by (forward_tac [finite_UnI] 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Subst/Unifier.ML --- a/src/HOL/Subst/Unifier.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Subst/Unifier.ML Mon Jun 22 17:26:46 1998 +0200 @@ -15,14 +15,14 @@ * Unifiers *---------------------------------------------------------------------------*) -goal Unifier.thy +Goal "Unifier s (Comb t u) (Comb v w) = (Unifier s t v & Unifier s u w)"; by (simp_tac (simpset() addsimps [Unifier_def]) 1); qed "Unifier_Comb"; AddIffs [Unifier_Comb]; -goal Unifier.thy +Goal "!!v. [| v ~: vars_of t; v ~: vars_of u; Unifier s t u |] ==> \ \ Unifier ((v,r)#s) t u"; by (asm_full_simp_tac (simpset() addsimps [Unifier_def, repl_invariance]) 1); @@ -33,25 +33,25 @@ * Most General Unifiers *---------------------------------------------------------------------------*) -goalw Unifier.thy unify_defs "MGUnifier s t u = MGUnifier s u t"; +Goalw unify_defs "MGUnifier s t u = MGUnifier s u t"; by (blast_tac (claset() addIs [sym]) 1); qed "mgu_sym"; -goal Unifier.thy "[] >> s"; +Goal "[] >> s"; by (simp_tac (simpset() addsimps [MoreGeneral_def]) 1); by (Blast_tac 1); qed "MoreGen_Nil"; AddIffs [MoreGen_Nil]; -goalw Unifier.thy unify_defs +Goalw unify_defs "MGUnifier s t u = (ALL r. Unifier r t u = s >> r)"; by (auto_tac (claset() addIs [ssubst_subst2, subst_comp_Nil], simpset())); qed "MGU_iff"; -goal Unifier.thy +Goal "!!v. ~ Var v <: t ==> MGUnifier [(v,t)] (Var v) t"; by (simp_tac(simpset() addsimps [MGU_iff, Unifier_def, MoreGeneral_def] delsimps [subst_Var]) 1); @@ -69,25 +69,25 @@ * Idempotence. *---------------------------------------------------------------------------*) -goalw Unifier.thy [Idem_def] "Idem([])"; +Goalw [Idem_def] "Idem([])"; by (Simp_tac 1); qed "Idem_Nil"; AddIffs [Idem_Nil]; -goalw Unifier.thy [Idem_def] "Idem(s) = (sdom(s) Int srange(s) = {})"; +Goalw [Idem_def] "Idem(s) = (sdom(s) Int srange(s) = {})"; by (simp_tac (simpset() addsimps [subst_eq_iff, invariance, dom_range_disjoint]) 1); qed "Idem_iff"; -goal Unifier.thy "~ (Var(v) <: t) --> Idem([(v,t)])"; +Goal "~ (Var(v) <: t) --> Idem([(v,t)])"; by (simp_tac (simpset() addsimps [vars_iff_occseq, Idem_iff, srange_iff, empty_iff_all_not]) 1); qed_spec_mp "Var_Idem"; AddSIs [Var_Idem]; -goalw Unifier.thy [Idem_def] +Goalw [Idem_def] "!!r. [| Idem(r); Unifier s (t<|r) (u<|r) |] \ \ ==> Unifier (r <> s) (t <| r) (u <| r)"; by (asm_full_simp_tac (simpset() addsimps [Unifier_def, comp_subst_subst]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Subst/Unify.ML --- a/src/HOL/Subst/Unify.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Subst/Unify.ML Mon Jun 22 17:26:46 1998 +0200 @@ -57,7 +57,7 @@ * Termination proof. *---------------------------------------------------------------------------*) -goalw Unify.thy [unifyRel_def, measure_def] "trans unifyRel"; +Goalw [unifyRel_def, measure_def] "trans unifyRel"; by (REPEAT (resolve_tac [trans_inv_image, trans_lex_prod, trans_finite_psubset, trans_less_than, trans_inv_image] 1)); @@ -69,7 +69,7 @@ * the nested call in Unify. Loosely, it says that unifyRel doesn't care * about term structure. *---------------------------------------------------------------------------*) -goalw Unify.thy [unifyRel_def,lex_prod_def, inv_image_def] +Goalw [unifyRel_def,lex_prod_def, inv_image_def] "!!x. ((X,Y), (Comb A (Comb B C), Comb D (Comb E F))) : unifyRel ==> \ \ ((X,Y), (Comb (Comb A B) C, Comb (Comb D E) F)) : unifyRel"; by (asm_full_simp_tac (simpset() addsimps [measure_def, @@ -87,7 +87,7 @@ * This lemma proves the nested termination condition for the base cases * 3, 4, and 6. *---------------------------------------------------------------------------*) -goal Unify.thy +Goal "!!x. ~(Var x <: M) ==> \ \ ((N1 <| [(x,M)], N2 <| [(x,M)]), (Comb M N1, Comb(Var x) N2)) : unifyRel \ \ & ((N1 <| [(x,M)], N2 <| [(x,M)]), (Comb(Var x) N1, Comb M N2)) : unifyRel"; @@ -175,7 +175,7 @@ *---------------------------------------------------------------------------*) (*Desired rule, copied from the theory file. Could it be made available?*) -goal Unify.thy +Goal "unify(Comb M1 N1, Comb M2 N2) = \ \ (case unify(M1,M2) \ \ of None => None \ @@ -204,7 +204,7 @@ * approach of M&W, who used idempotence and MGU-ness in the termination proof. *---------------------------------------------------------------------------*) -goal Unify.thy "!theta. unify(M,N) = Some theta --> MGUnifier theta M N"; +Goal "!theta. unify(M,N) = Some theta --> MGUnifier theta M N"; by (res_inst_tac [("u","M"),("v","N")] unifyInduct 1); by (ALLGOALS Asm_simp_tac); (*Const-Const case*) @@ -240,7 +240,7 @@ (*--------------------------------------------------------------------------- * Unify returns idempotent substitutions, when it succeeds. *---------------------------------------------------------------------------*) -goal Unify.thy "!theta. unify(M,N) = Some theta --> Idem theta"; +Goal "!theta. unify(M,N) = Some theta --> Idem theta"; by (res_inst_tac [("u","M"),("v","N")] unifyInduct 1); by (ALLGOALS (asm_simp_tac diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Sum.ML --- a/src/HOL/Sum.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Sum.ML Mon Jun 22 17:26:46 1998 +0200 @@ -11,29 +11,29 @@ (** Inl_Rep and Inr_Rep: Representations of the constructors **) (*This counts as a non-emptiness result for admitting 'a+'b as a type*) -goalw Sum.thy [Sum_def] "Inl_Rep(a) : Sum"; +Goalw [Sum_def] "Inl_Rep(a) : Sum"; by (EVERY1 [rtac CollectI, rtac disjI1, rtac exI, rtac refl]); qed "Inl_RepI"; -goalw Sum.thy [Sum_def] "Inr_Rep(b) : Sum"; +Goalw [Sum_def] "Inr_Rep(b) : Sum"; by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]); qed "Inr_RepI"; -goal Sum.thy "inj_on Abs_Sum Sum"; +Goal "inj_on Abs_Sum Sum"; by (rtac inj_on_inverseI 1); by (etac Abs_Sum_inverse 1); qed "inj_on_Abs_Sum"; (** Distinctness of Inl and Inr **) -goalw Sum.thy [Inl_Rep_def, Inr_Rep_def] "Inl_Rep(a) ~= Inr_Rep(b)"; +Goalw [Inl_Rep_def, Inr_Rep_def] "Inl_Rep(a) ~= Inr_Rep(b)"; by (EVERY1 [rtac notI, etac (fun_cong RS fun_cong RS fun_cong RS iffE), rtac (notE RS ccontr), etac (mp RS conjunct2), REPEAT o (ares_tac [refl,conjI]) ]); qed "Inl_Rep_not_Inr_Rep"; -goalw Sum.thy [Inl_def,Inr_def] "Inl(a) ~= Inr(b)"; +Goalw [Inl_def,Inr_def] "Inl(a) ~= Inr(b)"; by (rtac (inj_on_Abs_Sum RS inj_on_contraD) 1); by (rtac Inl_Rep_not_Inr_Rep 1); by (rtac Inl_RepI 1); @@ -61,7 +61,7 @@ by (Blast_tac 1); qed "Inr_Rep_inject"; -goalw Sum.thy [Inl_def] "inj(Inl)"; +Goalw [Inl_def] "inj(Inl)"; by (rtac injI 1); by (etac (inj_on_Abs_Sum RS inj_onD RS Inl_Rep_inject) 1); by (rtac Inl_RepI 1); @@ -69,7 +69,7 @@ qed "inj_Inl"; val Inl_inject = inj_Inl RS injD; -goalw Sum.thy [Inr_def] "inj(Inr)"; +Goalw [Inr_def] "inj(Inr)"; by (rtac injI 1); by (etac (inj_on_Abs_Sum RS inj_onD RS Inr_Rep_inject) 1); by (rtac Inr_RepI 1); @@ -77,11 +77,11 @@ qed "inj_Inr"; val Inr_inject = inj_Inr RS injD; -goal Sum.thy "(Inl(x)=Inl(y)) = (x=y)"; +Goal "(Inl(x)=Inl(y)) = (x=y)"; by (blast_tac (claset() addSDs [Inl_inject]) 1); qed "Inl_eq"; -goal Sum.thy "(Inr(x)=Inr(y)) = (x=y)"; +Goal "(Inr(x)=Inr(y)) = (x=y)"; by (blast_tac (claset() addSDs [Inr_inject]) 1); qed "Inr_eq"; @@ -91,11 +91,11 @@ (** Introduction rules for the injections **) -goalw Sum.thy [sum_def] "!!a A B. a : A ==> Inl(a) : A Plus B"; +Goalw [sum_def] "!!a A B. a : A ==> Inl(a) : A Plus B"; by (Blast_tac 1); qed "InlI"; -goalw Sum.thy [sum_def] "!!b A B. b : B ==> Inr(b) : A Plus B"; +Goalw [sum_def] "!!b A B. b : B ==> Inr(b) : A Plus B"; by (Blast_tac 1); qed "InrI"; @@ -118,11 +118,11 @@ (** sum_case -- the selection operator for sums **) -goalw Sum.thy [sum_case_def] "sum_case f g (Inl x) = f(x)"; +Goalw [sum_case_def] "sum_case f g (Inl x) = f(x)"; by (Blast_tac 1); qed "sum_case_Inl"; -goalw Sum.thy [sum_case_def] "sum_case f g (Inr x) = g(x)"; +Goalw [sum_case_def] "sum_case f g (Inr x) = g(x)"; by (Blast_tac 1); qed "sum_case_Inr"; @@ -140,13 +140,13 @@ rtac (Rep_Sum_inverse RS sym)])); qed "sumE"; -goal Sum.thy "sum_case (%x::'a. f(Inl x)) (%y::'b. f(Inr y)) s = f(s)"; +Goal "sum_case (%x::'a. f(Inl x)) (%y::'b. f(Inr y)) s = f(s)"; by (EVERY1 [res_inst_tac [("s","s")] sumE, etac ssubst, rtac sum_case_Inl, etac ssubst, rtac sum_case_Inr]); qed "surjective_sum"; -goal Sum.thy "R(sum_case f g s) = \ +Goal "R(sum_case f g s) = \ \ ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))"; by (res_inst_tac [("s","s")] sumE 1); by Auto_tac; @@ -166,7 +166,7 @@ (** Rules for the Part primitive **) -goalw Sum.thy [Part_def] +Goalw [Part_def] "!!a b A h. [| a : A; a=h(b) |] ==> a : Part A h"; by (Blast_tac 1); qed "Part_eqI"; @@ -185,29 +185,29 @@ AddIs [Part_eqI]; AddSEs [PartE]; -goalw Sum.thy [Part_def] "Part A h <= A"; +Goalw [Part_def] "Part A h <= A"; by (rtac Int_lower1 1); qed "Part_subset"; -goal Sum.thy "!!A B. A<=B ==> Part A h <= Part B h"; +Goal "!!A B. A<=B ==> Part A h <= Part B h"; by (Blast_tac 1); qed "Part_mono"; val basic_monos = basic_monos @ [Part_mono]; -goalw Sum.thy [Part_def] "!!a. a : Part A h ==> a : A"; +Goalw [Part_def] "!!a. a : Part A h ==> a : A"; by (etac IntD1 1); qed "PartD1"; -goal Sum.thy "Part A (%x. x) = A"; +Goal "Part A (%x. x) = A"; by (Blast_tac 1); qed "Part_id"; -goal Sum.thy "Part (A Int B) h = (Part A h) Int (Part B h)"; +Goal "Part (A Int B) h = (Part A h) Int (Part B h)"; by (Blast_tac 1); qed "Part_Int"; (*For inductive definitions*) -goal Sum.thy "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}"; +Goal "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}"; by (Blast_tac 1); qed "Part_Collect"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/TLA/Buffer/Buffer.ML --- a/src/HOL/TLA/Buffer/Buffer.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/TLA/Buffer/Buffer.ML Mon Jun 22 17:26:46 1998 +0200 @@ -21,18 +21,18 @@ (* ---------------------------- Action lemmas ---------------------------- *) (* Dequeue is visible *) -goal Buffer.thy "_ .= Deq ic q oc"; +Goal "_ .= Deq ic q oc"; by (auto_tac (claset(), simpset() addsimps [angle_def,Deq_def])); qed "Deq_visible"; (* Enabling condition for dequeue -- NOT NEEDED *) -goalw Buffer.thy [temp_rewrite Deq_visible] +Goalw [temp_rewrite Deq_visible] "!!q. base_var ==> $Enabled (_) .= ($q .~= .[])"; by (auto_tac (claset() addSEs [base_enabled,enabledE], simpset() addsimps [Deq_def])); qed "Deq_enabled"; (* For the left-to-right implication, we don't need the base variable stuff *) -goalw Buffer.thy [temp_rewrite Deq_visible] +Goalw [temp_rewrite Deq_visible] "$Enabled (_) .-> ($q .~= .[])"; by (auto_tac (claset() addSEs [enabledE], simpset() addsimps [Deq_def])); qed "Deq_enabledE"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/TLA/Buffer/DBuffer.ML --- a/src/HOL/TLA/Buffer/DBuffer.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/TLA/Buffer/DBuffer.ML Mon Jun 22 17:26:46 1998 +0200 @@ -14,13 +14,13 @@ (*** Proper initialization ***) -goal DBuffer.thy "Init DBInit .-> Init (BInit inp qc out)"; +Goal "Init DBInit .-> Init (BInit inp qc out)"; by (auto_tac (db_css addsimps2 [Init_def,DBInit_def,BInit_def])); qed "DBInit"; (*** Step simulation ***) -goal DBuffer.thy "[DBNext]_ .-> [Next inp qc out]_"; +Goal "[DBNext]_ .-> [Next inp qc out]_"; by (rtac square_simulation 1); by (Action_simp_tac 1); by (action_simp_tac (simpset() addsimps hd_append::db_defs) [] [] 1); @@ -30,22 +30,22 @@ (*** Simulation of fairness ***) (* Compute enabledness predicates for DBDeq and DBPass actions *) -goal DBuffer.thy "_ .= DBDeq"; +Goal "_ .= DBDeq"; by (auto_tac (db_css addsimps2 [angle_def,DBDeq_def,Deq_def])); qed "DBDeq_visible"; -goal DBuffer.thy "$Enabled (_) .= ($q2 .~= .[])"; +Goal "$Enabled (_) .= ($q2 .~= .[])"; by (rewtac (action_rewrite DBDeq_visible)); by (cut_facts_tac [DB_base] 1); by (old_auto_tac (db_css addSEs2 [base_enabled,enabledE] addsimps2 [angle_def,DBDeq_def,Deq_def])); qed "DBDeq_enabled"; -goal DBuffer.thy "_ .= DBPass"; +Goal "_ .= DBPass"; by (auto_tac (db_css addsimps2 [angle_def,DBPass_def,Deq_def])); qed "DBPass_visible"; -goal DBuffer.thy "$Enabled (_) .= ($q1 .~= .[])"; +Goal "$Enabled (_) .= ($q1 .~= .[])"; by (rewtac (action_rewrite DBPass_visible)); by (cut_facts_tac [DB_base] 1); by (auto_tac (db_css addSEs2 [base_enabled,enabledE] @@ -73,7 +73,7 @@ *) (* Condition (1a) *) -goal DBuffer.thy +Goal "[][DBNext]_ .& WF(DBPass)_ \ \ .-> ($qc .~= .[] .& $q2 .= .[] ~> $q2 .~= .[])"; by (rtac WF1 1); @@ -83,7 +83,7 @@ qed "DBFair_1a"; (* Condition (1) *) -goal DBuffer.thy +Goal "[][DBNext]_ .& WF(DBPass)_ \ \ .-> ($Enabled (_) ~> $q2 .~= .[])"; by (auto_tac (temp_css addSIs2 [leadsto_classical] addSEs2 [temp_conjimpE DBFair_1a])); @@ -92,7 +92,7 @@ qed "DBFair_1"; (* Condition (2) *) -goal DBuffer.thy +Goal "[][DBNext]_ .& WF(DBDeq)_ \ \ .-> ($q2 .~= .[] ~> DBDeq)"; by (rtac WF_leadsto 1); @@ -102,7 +102,7 @@ qed "DBFair_2"; (* High-level fairness *) -goal DBuffer.thy +Goal "[][DBNext]_ .& WF(DBPass)_ \ \ .& WF(DBDeq)_ \ \ .-> WF(Deq inp qc out)_"; @@ -114,6 +114,6 @@ qed "DBFair"; (*** Main theorem ***) -goalw DBuffer.thy [DBuffer_def,Buffer_def,IBuffer_def] "DBuffer .-> Buffer inp out"; +Goalw [DBuffer_def,Buffer_def,IBuffer_def] "DBuffer .-> Buffer inp out"; by (auto_tac (db_css addSIs2 (map temp_mp [eexI,DBInit,DB_step_simulation RS STL4,DBFair]))); qed "DBuffer_impl_Buffer"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/TLA/Inc/Inc.ML --- a/src/HOL/TLA/Inc/Inc.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/TLA/Inc/Inc.ML Mon Jun 22 17:26:46 1998 +0200 @@ -54,7 +54,7 @@ (* Automatic proof works too, but it make take a while on a slow machine. More substantial examples require manual guidance anyway. -goal Inc.thy "Psi .-> []PsiInv"; +Goal "Psi .-> []PsiInv"; by (auto_inv_tac (simpset() addsimps PsiInv_defs @ Psi_defs @ pcount.simps) 1); *) diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/TLA/TLA.ML --- a/src/HOL/TLA/TLA.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/TLA/TLA.ML Mon Jun 22 17:26:46 1998 +0200 @@ -1046,7 +1046,7 @@ (* ---------------------------------------------------------------------- example of a history variable: existence of a clock -goal TLA.thy "(EEX h. Init($h .= #True) .& [](h$ .~= $h))"; +Goal "(EEX h. Init($h .= #True) .& [](h$ .~= $h))"; by (rtac tempI 1); by (rtac historyI 1); by (rewrite_goals_tac action_rews); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/TLA/hypsubst.ML --- a/src/HOL/TLA/hypsubst.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/TLA/hypsubst.ML Mon Jun 22 17:26:46 1998 +0200 @@ -17,16 +17,16 @@ Test data: -goal thy "!!x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)"; -goal thy "!!x.[| Q(x,y,z); z=f(x); x=z |] ==> P(z)"; -goal thy "!!y. [| ?x=y; P(?x) |] ==> y = a"; -goal thy "!!z. [| ?x=y; P(?x) |] ==> y = a"; +Goal "!!x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)"; +Goal "!!x.[| Q(x,y,z); z=f(x); x=z |] ==> P(z)"; +Goal "!!y. [| ?x=y; P(?x) |] ==> y = a"; +Goal "!!z. [| ?x=y; P(?x) |] ==> y = a"; by (hyp_subst_tac 1); by (bound_hyp_subst_tac 1); Here hyp_subst_tac goes wrong; harder still to prove P(f(f(a))) & P(f(a)) -goal thy "P(a) --> (EX y. a=y --> P(f(a)))"; +Goal "P(a) --> (EX y. a=y --> P(f(a)))"; *) (*** Signatures unchanged (but renamed) from the original hypsubst.ML ***) diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Trancl.ML --- a/src/HOL/Trancl.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Trancl.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,7 +10,7 @@ (** The relation rtrancl **) -goal Trancl.thy "mono(%s. id Un (r O s))"; +Goal "mono(%s. id Un (r O s))"; by (rtac monoI 1); by (REPEAT (ares_tac [monoI, subset_refl, comp_mono, Un_mono] 1)); qed "rtrancl_fun_mono"; @@ -18,7 +18,7 @@ val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski); (*Reflexivity of rtrancl*) -goal Trancl.thy "(a,a) : r^*"; +Goal "(a,a) : r^*"; by (stac rtrancl_unfold 1); by (Blast_tac 1); qed "rtrancl_refl"; @@ -28,19 +28,19 @@ (*Closure under composition with r*) -goal Trancl.thy "!!r. [| (a,b) : r^*; (b,c) : r |] ==> (a,c) : r^*"; +Goal "!!r. [| (a,b) : r^*; (b,c) : r |] ==> (a,c) : r^*"; by (stac rtrancl_unfold 1); by (Blast_tac 1); qed "rtrancl_into_rtrancl"; (*rtrancl of r contains r*) -goal Trancl.thy "!!p. p : r ==> p : r^*"; +Goal "!!p. p : r ==> p : r^*"; by (split_all_tac 1); by (etac (rtrancl_refl RS rtrancl_into_rtrancl) 1); qed "r_into_rtrancl"; (*monotonicity of rtrancl*) -goalw Trancl.thy [rtrancl_def] "!!r s. r <= s ==> r^* <= s^*"; +Goalw [rtrancl_def] "!!r s. r <= s ==> r^* <= s^*"; by (REPEAT(ares_tac [lfp_mono,Un_mono,comp_mono,subset_refl] 1)); qed "rtrancl_mono"; @@ -77,7 +77,7 @@ (read_instantiate [("a","(ax,ay)"), ("b","(bx,by)")] rtrancl_induct)); (*transitivity of transitive closure!! -- by induction.*) -goalw Trancl.thy [trans_def] "trans(r^*)"; +Goalw [trans_def] "trans(r^*)"; by Safe_tac; by (eres_inst_tac [("b","z")] rtrancl_induct 1); by (ALLGOALS(blast_tac (claset() addIs [rtrancl_into_rtrancl]))); @@ -103,7 +103,7 @@ (*** More r^* equations and inclusions ***) -goal Trancl.thy "(r^*)^* = r^*"; +Goal "(r^*)^* = r^*"; by (rtac set_ext 1); by (res_inst_tac [("p","x")] PairE 1); by (hyp_subst_tac 1); @@ -115,51 +115,51 @@ qed "rtrancl_idemp"; Addsimps [rtrancl_idemp]; -goal thy "R^* O R^* = R^*"; +Goal "R^* O R^* = R^*"; br set_ext 1; by(split_all_tac 1); by(blast_tac (claset() addIs [rtrancl_trans]) 1); qed "rtrancl_idemp_self_comp"; Addsimps [rtrancl_idemp_self_comp]; -goal Trancl.thy "!!r s. r <= s^* ==> r^* <= s^*"; +Goal "!!r s. r <= s^* ==> r^* <= s^*"; by (dtac rtrancl_mono 1); by (Asm_full_simp_tac 1); qed "rtrancl_subset_rtrancl"; -goal Trancl.thy "!!R. [| R <= S; S <= R^* |] ==> S^* = R^*"; +Goal "!!R. [| R <= S; S <= R^* |] ==> S^* = R^*"; by (dtac rtrancl_mono 1); by (dtac rtrancl_mono 1); by (Asm_full_simp_tac 1); by (Blast_tac 1); qed "rtrancl_subset"; -goal Trancl.thy "!!R. (R^* Un S^*)^* = (R Un S)^*"; +Goal "!!R. (R^* Un S^*)^* = (R Un S)^*"; by (blast_tac (claset() addSIs [rtrancl_subset] addIs [r_into_rtrancl, rtrancl_mono RS subsetD]) 1); qed "rtrancl_Un_rtrancl"; -goal Trancl.thy "(R^=)^* = R^*"; +Goal "(R^=)^* = R^*"; by (blast_tac (claset() addSIs [rtrancl_subset] addIs [rtrancl_refl, r_into_rtrancl]) 1); qed "rtrancl_reflcl"; Addsimps [rtrancl_reflcl]; -goal Trancl.thy "!!r. (x,y) : (r^-1)^* ==> (x,y) : (r^*)^-1"; +Goal "!!r. (x,y) : (r^-1)^* ==> (x,y) : (r^*)^-1"; by (rtac converseI 1); by (etac rtrancl_induct 1); by (rtac rtrancl_refl 1); by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 1); qed "rtrancl_converseD"; -goal Trancl.thy "!!r. (x,y) : (r^*)^-1 ==> (x,y) : (r^-1)^*"; +Goal "!!r. (x,y) : (r^*)^-1 ==> (x,y) : (r^-1)^*"; by (dtac converseD 1); by (etac rtrancl_induct 1); by (rtac rtrancl_refl 1); by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 1); qed "rtrancl_converseI"; -goal Trancl.thy "(r^-1)^* = (r^*)^-1"; +Goal "(r^-1)^* = (r^*)^-1"; by (safe_tac (claset() addSDs [rtrancl_converseD] addSIs [rtrancl_converseI])); qed "rtrancl_converse"; @@ -198,7 +198,7 @@ by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1)); qed "rtranclE2"; -goal Trancl.thy "r O r^* = r^* O r"; +Goal "r O r^* = r^* O r"; by (blast_tac (claset() addEs [rtranclE, rtranclE2] addIs [rtrancl_into_rtrancl, rtrancl_into_rtrancl2]) 1); qed "r_comp_rtrancl_eq"; @@ -206,13 +206,13 @@ (**** The relation trancl ****) -goalw Trancl.thy [trancl_def] "!!p.[| p:r^+; r <= s |] ==> p:s^+"; +Goalw [trancl_def] "!!p.[| p:r^+; r <= s |] ==> p:s^+"; by (blast_tac (claset() addIs [rtrancl_mono RS subsetD]) 1); qed "trancl_mono"; (** Conversions between trancl and rtrancl **) -goalw Trancl.thy [trancl_def] +Goalw [trancl_def] "!!p. p : r^+ ==> p : r^*"; by (split_all_tac 1); by (etac compEpair 1); @@ -220,7 +220,7 @@ qed "trancl_into_rtrancl"; (*r^+ contains r*) -goalw Trancl.thy [trancl_def] +Goalw [trancl_def] "!!p. p : r ==> p : r^+"; by (split_all_tac 1); by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); @@ -273,7 +273,7 @@ (*Transitivity of r^+. Proved by unfolding since it uses transitivity of rtrancl. *) -goalw Trancl.thy [trancl_def] "trans(r^+)"; +Goalw [trancl_def] "trans(r^+)"; by (rtac transI 1); by (REPEAT (etac compEpair 1)); by (rtac (rtrancl_into_rtrancl RS (rtrancl_trans RS compI)) 1); @@ -282,7 +282,7 @@ bind_thm ("trancl_trans", trans_trancl RS transD); -goalw Trancl.thy [trancl_def] +Goalw [trancl_def] "!!r. [| (x,y):r^*; (y,z):r^+ |] ==> (x,z):r^+"; by (blast_tac (claset() addIs [rtrancl_trans,r_into_rtrancl]) 1); qed "rtrancl_trancl_trancl"; @@ -295,7 +295,7 @@ qed "trancl_into_trancl2"; (* primitive recursion for trancl over finite relations: *) -goal Trancl.thy "(insert (y,x) r)^+ = r^+ Un {(a,b). (a,y):r^* & (x,b):r^*}"; +Goal "(insert (y,x) r)^+ = r^+ Un {(a,b). (a,y):r^* & (x,b):r^*}"; by (rtac equalityI 1); by (rtac subsetI 1); by (split_all_tac 1); @@ -309,7 +309,7 @@ impOfSubs rtrancl_mono, trancl_mono]) 1); qed "trancl_insert"; -goalw Trancl.thy [trancl_def] "(r^-1)^+ = (r^+)^-1"; +Goalw [trancl_def] "(r^-1)^+ = (r^+)^-1"; by (simp_tac (simpset() addsimps [rtrancl_converse,converse_comp]) 1); by (simp_tac (simpset() addsimps [rtrancl_converse RS sym,r_comp_rtrancl_eq])1); qed "trancl_converse"; @@ -331,13 +331,13 @@ by (Blast_tac 1); val lemma = result(); -goalw Trancl.thy [trancl_def] +Goalw [trancl_def] "!!r. r <= A Times A ==> r^+ <= A Times A"; by (blast_tac (claset() addSDs [lemma]) 1); qed "trancl_subset_Sigma"; -goal Trancl.thy "(r^+)^= = r^*"; +Goal "(r^+)^= = r^*"; by Safe_tac; by (etac trancl_into_rtrancl 1); by (etac rtranclE 1); @@ -347,7 +347,7 @@ qed "reflcl_trancl"; Addsimps[reflcl_trancl]; -goal Trancl.thy "(r^=)^+ = r^*"; +Goal "(r^=)^+ = r^*"; by Safe_tac; by (dtac trancl_into_rtrancl 1); by (Asm_full_simp_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/Channel.ML --- a/src/HOL/UNITY/Channel.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/Channel.ML Mon Jun 22 17:26:46 1998 +0200 @@ -14,22 +14,22 @@ (*None represents "infinity" while Some represents proper integers*) -goalw thy [minSet_def] "!!A. minSet A = Some x --> x : A"; +Goalw [minSet_def] "!!A. minSet A = Some x --> x : A"; by (Simp_tac 1); by (fast_tac (claset() addIs [LeastI]) 1); qed_spec_mp "minSet_eq_SomeD"; -goalw thy [minSet_def] " minSet{} = None"; +Goalw [minSet_def] " minSet{} = None"; by (Asm_simp_tac 1); qed_spec_mp "minSet_empty"; Addsimps [minSet_empty]; -goalw thy [minSet_def] "!!A. x:A ==> minSet A = Some (LEAST x. x: A)"; +Goalw [minSet_def] "!!A. x:A ==> minSet A = Some (LEAST x. x: A)"; by (ALLGOALS Asm_simp_tac); by (Blast_tac 1); qed_spec_mp "minSet_nonempty"; -goal thy +Goal "leadsTo Acts (minSet -`` {Some x}) (minSet -`` (Some``greaterThan x))"; by (rtac leadsTo_weaken 1); by (rtac ([UC2, UC1] MRS PSP) 1); @@ -42,7 +42,7 @@ (*The induction*) -goal thy "leadsTo Acts (UNIV-{{}}) (minSet -`` (Some``atLeast y))"; +Goal "leadsTo Acts (UNIV-{{}}) (minSet -`` (Some``atLeast y))"; by (rtac leadsTo_weaken_R 1); by (res_inst_tac [("l", "y"), ("f", "the o minSet"), ("B", "{}")] greaterThan_bounded_induct 1); @@ -58,7 +58,7 @@ val lemma = result(); -goal thy "!!y::nat. leadsTo Acts (UNIV-{{}}) {s. y ~: s}"; +Goal "!!y::nat. leadsTo Acts (UNIV-{{}}) {s. y ~: s}"; by (rtac (lemma RS leadsTo_weaken_R) 1); by (Clarify_tac 1); by (forward_tac [minSet_nonempty] 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/Common.ML --- a/src/HOL/UNITY/Common.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/Common.ML Mon Jun 22 17:26:46 1998 +0200 @@ -14,7 +14,7 @@ open Common; (*Misra's property CMT4: t exceeds no common meeting time*) -goal thy +Goal "!!Acts. [| ALL m. constrains Acts {m} (maxfg m); n: common |] \ \ ==> stable Acts (atMost n)"; by (dres_inst_tac [("P", "%t. t<=n")] elimination_sing 1); @@ -28,7 +28,7 @@ addIs [order_eq_refl, fmono, gmono, le_trans]) 1); qed "common_stable"; -goal thy +Goal "!!Acts. [| ALL m. constrains Acts {m} (maxfg m); n: common |] \ \ ==> reachable {0} Acts <= atMost n"; by (rtac strongest_invariant 1); @@ -40,13 +40,13 @@ (*** Some programs that implement the safety property above ***) (*This one is just Skip*) -goal thy "constrains {id} {m} (maxfg m)"; +Goal "constrains {id} {m} (maxfg m)"; by (simp_tac (simpset() addsimps [constrains_def, maxfg_def, le_max_iff_disj, fasc, gasc]) 1); result(); (*This one is t := ftime t || t := gtime t really needs Skip too*) -goal thy "constrains {range(%t.(t,ftime t)), range(%t.(t,gtime t))} \ +Goal "constrains {range(%t.(t,ftime t)), range(%t.(t,gtime t))} \ \ {m} (maxfg m)"; by (simp_tac (simpset() addsimps [constrains_def, maxfg_def, le_max_iff_disj]) 1); @@ -54,14 +54,14 @@ result(); (*This one is t := max (ftime t) (gtime t) really needs Skip too*) -goal thy "constrains {range(%t.(t, max (ftime t) (gtime t)))} \ +Goal "constrains {range(%t.(t, max (ftime t) (gtime t)))} \ \ {m} (maxfg m)"; by (simp_tac (simpset() addsimps [constrains_def, maxfg_def]) 1); by (Blast_tac 1); result(); (*This one is t := t+1 if t stable Acts (A Int B)"; by (Blast_tac 1); result(); -goal thy "{i. i < Suc n} = insert n {i. i < n}"; +Goal "{i. i < Suc n} = insert n {i. i < n}"; by (blast_tac (claset() addSEs [less_SucE] addIs [less_SucI]) 1); qed "Collect_less_Suc_insert"; -goal thy "{i. i <= Suc n} = insert (Suc n) {i. i <= n}"; +Goal "{i. i <= Suc n} = insert (Suc n) {i. i <= n}"; by (blast_tac (claset() addSEs [le_SucE] addIs [le_SucI]) 1); qed "Collect_le_Suc_insert"; (*a simplification step*) -goal thy "(INT i:{i. i <= n}. A(Suc i) Int A i) = \ +Goal "(INT i:{i. i <= n}. A(Suc i) Int A i) = \ \ (INT i:{i. i <= Suc n}. A i)"; by (induct_tac "n" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [Collect_le_Suc_insert]))); @@ -30,7 +30,7 @@ (*Dual of the required property. Converse inclusion fails.*) -goal thy "(UN i:{i. i < n}. A i) Int Compl (A n) <= \ +Goal "(UN i:{i. i < n}. A i) Int Compl (A n) <= \ \ (UN i:{i. i < n}. (A i) Int Compl (A(Suc i)))"; by (induct_tac "n" 1); by (Asm_simp_tac 1); @@ -40,7 +40,7 @@ (*Converse inclusion fails.*) -goal thy "(INT i:{i. i < n}. Compl(A i) Un A (Suc i)) <= \ +Goal "(INT i:{i. i < n}. Compl(A i) Un A (Suc i)) <= \ \ (INT i:{i. i < n}. Compl(A i)) Un A n"; by (induct_tac "n" 1); by (Asm_simp_tac 1); @@ -50,14 +50,14 @@ (*Specialized rewriting*) -goal thy "A 0 Int (Compl (A n) Int \ +Goal "A 0 Int (Compl (A n) Int \ \ (INT i:{i. i < n}. Compl(A i) Un A (Suc i))) = {}"; by (blast_tac (claset() addIs [gr0I] addDs [impOfSubs INT_Un_Compl_subset]) 1); val lemma = result(); (*Reverse direction makes it harder to invoke the ind hyp*) -goal thy "(INT i:{i. i <= n}. A i) = \ +Goal "(INT i:{i. i <= n}. A i) = \ \ A 0 Int (INT i:{i. i < n}. Compl(A i) Un A(Suc i))"; by (induct_tac "n" 1); by (Asm_simp_tac 1); @@ -67,7 +67,7 @@ Collect_less_Suc_insert, Collect_le_Suc_insert])) 1); qed "INT_le_equals_Int"; -goal thy "(INT i:{i. i <= Suc n}. A i) = \ +Goal "(INT i:{i. i <= Suc n}. A i) = \ \ A 0 Int (INT i:{i. i <= n}. Compl(A i) Un A(Suc i))"; by (simp_tac (simpset() addsimps [le_eq_less_Suc RS sym, INT_le_equals_Int]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/FP.ML --- a/src/HOL/UNITY/FP.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/FP.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,13 +8,13 @@ From Misra, "A Logic for Concurrent Programming", 1994 *) -goal thy "Union(B) Int A = (UN C:B. C Int A)"; +Goal "Union(B) Int A = (UN C:B. C Int A)"; by (Blast_tac 1); qed "Int_Union2"; open FP; -goalw thy [FP_Orig_def, stable_def] "stable Acts (FP_Orig Acts Int B)"; +Goalw [FP_Orig_def, stable_def] "stable Acts (FP_Orig Acts Int B)"; by (stac Int_Union2 1); by (rtac ball_constrains_UN 1); by (Simp_tac 1); @@ -27,7 +27,7 @@ qed "FP_Orig_weakest"; -goal thy "stable Acts (FP Acts Int B)"; +Goal "stable Acts (FP Acts Int B)"; by (subgoal_tac "FP Acts Int B = (UN x:B. FP Acts Int {x})" 1); by (Blast_tac 2); by (asm_simp_tac (simpset() addsimps [Int_insert_right]) 1); @@ -36,17 +36,17 @@ by (Simp_tac 1); qed "stable_FP_Int"; -goal thy "FP Acts <= FP_Orig Acts"; +Goal "FP Acts <= FP_Orig Acts"; by (rtac (stable_FP_Int RS FP_Orig_weakest) 1); val lemma1 = result(); -goalw thy [FP_Orig_def, FP_def] "FP_Orig Acts <= FP Acts"; +Goalw [FP_Orig_def, FP_def] "FP_Orig Acts <= FP Acts"; by (Clarify_tac 1); by (dres_inst_tac [("x", "{x}")] spec 1); by (asm_full_simp_tac (simpset() addsimps [Int_insert_right]) 1); val lemma2 = result(); -goal thy "FP Acts = FP_Orig Acts"; +Goal "FP Acts = FP_Orig Acts"; by (rtac ([lemma1,lemma2] MRS equalityI) 1); qed "FP_equivalence"; @@ -55,12 +55,12 @@ by (simp_tac (simpset() addsimps [FP_equivalence, prem RS FP_Orig_weakest]) 1); qed "FP_weakest"; -goalw thy [FP_def, stable_def, constrains_def] +Goalw [FP_def, stable_def, constrains_def] "Compl (FP Acts) = (UN act:Acts. Compl{s. act^^{s} <= {s}})"; by (Blast_tac 1); qed "Compl_FP"; -goal thy "A - (FP Acts) = (UN act:Acts. A - {s. act^^{s} <= {s}})"; +Goal "A - (FP Acts) = (UN act:Acts. A - {s. act^^{s} <= {s}})"; by (simp_tac (simpset() addsimps [Diff_eq, Compl_FP]) 1); qed "Diff_FP"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/LessThan.ML --- 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 Suc (metric (s[x:=True])) = metric s"; +Goalw [metric_def] "!!s. ~ s x ==> Suc (metric (s[x:=True])) = metric s"; by (subgoal_tac "{v. ~ (s[x:=True]) v} = {v. ~ s v} - {x}" 1); by Auto_tac; by (asm_simp_tac (simpset() addsimps [card_Suc_Diff]) 1); qed "Suc_metric"; -goal thy "!!s. ~ s x ==> metric (s[x:=True]) < metric s"; +Goal "!!s. ~ s x ==> metric (s[x:=True]) < metric s"; by (etac (Suc_metric RS subst) 1); by (Blast_tac 1); qed "metric_less"; AddSIs [metric_less]; -goal thy "metric (s[y:=s x | s y]) <= metric s"; +Goal "metric (s[y:=s x | s y]) <= metric s"; by (case_tac "s x --> s y" 1); by (auto_tac (claset() addIs [less_imp_le], simpset() addsimps [update_idem])); qed "metric_le"; -goal thy "!!m. (u,v): edges ==> \ +Goal "!!m. (u,v): edges ==> \ \ ensures racts ((metric-``{m}) Int {s. s u & ~ s v}) \ \ (metric-``(lessThan m))"; by (ensures_tac "asgt u v" 1); @@ -160,27 +160,27 @@ by (fast_tac (claset() addSDs [le_imp_less_or_eq]) 1); qed "edges_ensures"; -goal thy "leadsTo racts ((metric-``{m}) - fixedpoint) (metric-``(lessThan m))"; +Goal "leadsTo racts ((metric-``{m}) - fixedpoint) (metric-``(lessThan m))"; by (simp_tac (simpset() addsimps [Diff_fixedpoint]) 1); by (rtac leadsTo_UN 1); by (split_all_tac 1); by (asm_simp_tac (simpset() addsimps [edges_ensures RS leadsTo_Basis]) 1); qed "leadsTo_Diff_fixedpoint"; -goal thy "leadsTo racts (metric-``{m}) (metric-``(lessThan m) Un fixedpoint)"; +Goal "leadsTo racts (metric-``{m}) (metric-``(lessThan m) Un fixedpoint)"; by (rtac (leadsTo_Diff_fixedpoint RS leadsTo_weaken_R RS leadsTo_Diff) 1); by (ALLGOALS (blast_tac (claset() addIs [subset_imp_leadsTo]))); qed "leadsTo_Un_fixedpoint"; (*Execution in any state leads to a fixedpoint (i.e. can terminate)*) -goal thy "leadsTo racts UNIV fixedpoint"; +Goal "leadsTo racts UNIV fixedpoint"; by (rtac lessThan_induct 1); by Auto_tac; by (rtac leadsTo_Un_fixedpoint 1); qed "leadsTo_fixedpoint"; -goal thy "LeadsTo rinit racts UNIV { %v. (init, v) : edges^* }"; +Goal "LeadsTo rinit racts UNIV { %v. (init, v) : edges^* }"; by (stac (fixedpoint_invariant_correct RS sym) 1); by (rtac (leadsTo_fixedpoint RS leadsTo_imp_LeadsTo RS LeadsTo_weaken_R) 1); by (cut_facts_tac [reachable_subset_invariant] 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/SubstAx.ML --- a/src/HOL/UNITY/SubstAx.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/SubstAx.ML Mon Jun 22 17:26:46 1998 +0200 @@ -18,12 +18,12 @@ (*** Introduction rules: Basis, Trans, Union ***) -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. leadsTo Acts A B ==> LeadsTo Init Acts A B"; by (blast_tac (claset() addIs [PSP_stable2, stable_reachable]) 1); qed "leadsTo_imp_LeadsTo"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| constrains Acts (reachable Init Acts Int (A - A')) \ \ (A Un A'); \ \ transient Acts (reachable Init Acts Int (A-A')) |] \ @@ -35,7 +35,7 @@ stable_constrains_Int]) 1); qed "LeadsTo_Basis"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| LeadsTo Init Acts A B; LeadsTo Init Acts B C |] \ \ ==> LeadsTo Init Acts A C"; by (blast_tac (claset() addIs [leadsTo_Trans]) 1); @@ -50,18 +50,18 @@ (*** Derived rules ***) -goal thy "!!Acts. id: Acts ==> LeadsTo Init Acts A UNIV"; +Goal "!!Acts. id: Acts ==> LeadsTo Init Acts A UNIV"; by (asm_simp_tac (simpset() addsimps [LeadsTo_def, Int_lower1 RS subset_imp_leadsTo]) 1); qed "LeadsTo_UNIV"; Addsimps [LeadsTo_UNIV]; (*Useful with cancellation, disjunction*) -goal thy "!!Acts. LeadsTo Init Acts A (A' Un A') ==> LeadsTo Init Acts A A'"; +Goal "!!Acts. LeadsTo Init Acts A (A' Un A') ==> LeadsTo Init Acts A A'"; by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); qed "LeadsTo_Un_duplicate"; -goal thy "!!Acts. LeadsTo Init Acts A (A' Un C Un C) ==> LeadsTo Init Acts A (A' Un C)"; +Goal "!!Acts. LeadsTo Init Acts A (A' Un C Un C) ==> LeadsTo Init Acts A (A' Un C)"; by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); qed "LeadsTo_Un_duplicate2"; @@ -73,20 +73,20 @@ qed "LeadsTo_UN"; (*Binary union introduction rule*) -goal thy +Goal "!!C. [| LeadsTo Init Acts A C; LeadsTo Init Acts B C |] ==> LeadsTo Init Acts (A Un B) C"; by (stac Un_eq_Union 1); by (blast_tac (claset() addIs [LeadsTo_Union]) 1); qed "LeadsTo_Un"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!A B. [| reachable Init Acts Int A <= B; id: Acts |] \ \ ==> LeadsTo Init Acts A B"; by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1); qed "Int_subset_imp_LeadsTo"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!A B. [| A <= B; id: Acts |] \ \ ==> LeadsTo Init Acts A B"; by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1); @@ -95,14 +95,14 @@ bind_thm ("empty_LeadsTo", empty_subsetI RS subset_imp_LeadsTo); Addsimps [empty_LeadsTo]; -goal thy +Goal "!!A B. [| reachable Init Acts Int A = {}; id: Acts |] \ \ ==> LeadsTo Init Acts A B"; by (asm_simp_tac (simpset() addsimps [Int_subset_imp_LeadsTo]) 1); qed "Int_empty_LeadsTo"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| LeadsTo Init Acts A A'; \ \ reachable Init Acts Int A' <= B' |] \ \ ==> LeadsTo Init Acts A B'"; @@ -110,7 +110,7 @@ qed_spec_mp "LeadsTo_weaken_R"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| LeadsTo Init Acts A A'; \ \ reachable Init Acts Int B <= A; id: Acts |] \ \ ==> LeadsTo Init Acts B A'"; @@ -119,21 +119,21 @@ (*Distributes over binary unions*) -goal thy +Goal "!!C. id: Acts ==> \ \ LeadsTo Init Acts (A Un B) C = \ \ (LeadsTo Init Acts A C & LeadsTo Init Acts B C)"; by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken_L]) 1); qed "LeadsTo_Un_distrib"; -goal thy +Goal "!!C. id: Acts ==> \ \ LeadsTo Init Acts (UN i:I. A i) B = \ \ (ALL i : I. LeadsTo Init Acts (A i) B)"; by (blast_tac (claset() addIs [LeadsTo_UN, LeadsTo_weaken_L]) 1); qed "LeadsTo_UN_distrib"; -goal thy +Goal "!!C. id: Acts ==> \ \ LeadsTo Init Acts (Union S) B = \ \ (ALL A : S. LeadsTo Init Acts A B)"; @@ -141,7 +141,7 @@ qed "LeadsTo_Union_distrib"; -goal thy +Goal "!!Acts. [| LeadsTo Init Acts A A'; id: Acts; \ \ reachable Init Acts Int B <= A; \ \ reachable Init Acts Int A' <= B' |] \ @@ -153,7 +153,7 @@ (*Set difference: maybe combine with leadsTo_weaken_L??*) -goal thy +Goal "!!C. [| LeadsTo Init Acts (A-B) C; LeadsTo Init Acts B C; id: Acts |] \ \ ==> LeadsTo Init Acts A C"; by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken]) 1); @@ -181,7 +181,7 @@ qed "LeadsTo_UN_UN_noindex"; (*Version with no index set*) -goal thy +Goal "!!Acts. ALL i. LeadsTo Init Acts (A i) (A' i) \ \ ==> LeadsTo Init Acts (UN i. A i) (UN i. A' i)"; by (blast_tac (claset() addIs [LeadsTo_UN_UN]) 1); @@ -189,7 +189,7 @@ (*Binary union version*) -goal thy "!!Acts. [| LeadsTo Init Acts A A'; LeadsTo Init Acts B B' |] \ +Goal "!!Acts. [| LeadsTo Init Acts A A'; LeadsTo Init Acts B B' |] \ \ ==> LeadsTo Init Acts (A Un B) (A' Un B')"; by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken_R]) 1); @@ -198,7 +198,7 @@ (** The cancellation law **) -goal thy +Goal "!!Acts. [| LeadsTo Init Acts A (A' Un B); LeadsTo Init Acts B B'; \ \ id: Acts |] \ \ ==> LeadsTo Init Acts A (A' Un B')"; @@ -206,7 +206,7 @@ subset_imp_LeadsTo, LeadsTo_Trans]) 1); qed "LeadsTo_cancel2"; -goal thy +Goal "!!Acts. [| LeadsTo Init Acts A (A' Un B); LeadsTo Init Acts (B-A') B'; id: Acts |] \ \ ==> LeadsTo Init Acts A (A' Un B')"; by (rtac LeadsTo_cancel2 1); @@ -214,14 +214,14 @@ by (ALLGOALS Asm_simp_tac); qed "LeadsTo_cancel_Diff2"; -goal thy +Goal "!!Acts. [| LeadsTo Init Acts A (B Un A'); LeadsTo Init Acts B B'; id: Acts |] \ \ ==> LeadsTo Init Acts A (B' Un A')"; by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1); by (blast_tac (claset() addSIs [LeadsTo_cancel2]) 1); qed "LeadsTo_cancel1"; -goal thy +Goal "!!Acts. [| LeadsTo Init Acts A (B Un A'); LeadsTo Init Acts (B-A') B'; id: Acts |] \ \ ==> LeadsTo Init Acts A (B' Un A')"; by (rtac LeadsTo_cancel1 1); @@ -233,7 +233,7 @@ (** The impossibility law **) -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. LeadsTo Init Acts A {} ==> reachable Init Acts Int A = {}"; by (Full_simp_tac 1); by (etac leadsTo_empty 1); @@ -243,20 +243,20 @@ (** PSP: Progress-Safety-Progress **) (*Special case of PSP: Misra's "stable conjunction". Doesn't need id:Acts. *) -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| LeadsTo Init Acts A A'; stable Acts B |] \ \ ==> LeadsTo Init Acts (A Int B) (A' Int B)"; by (asm_simp_tac (simpset() addsimps [Int_assoc RS sym, PSP_stable]) 1); qed "R_PSP_stable"; -goal thy +Goal "!!Acts. [| LeadsTo Init Acts A A'; stable Acts B |] \ \ ==> LeadsTo Init Acts (B Int A) (B Int A')"; by (asm_simp_tac (simpset() addsimps (R_PSP_stable::Int_ac)) 1); qed "R_PSP_stable2"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| LeadsTo Init Acts A A'; constrains Acts B B'; id: Acts |] \ \ ==> LeadsTo Init Acts (A Int B) ((A' Int B) Un (B' - B))"; by (dtac PSP 1); @@ -265,13 +265,13 @@ by (ALLGOALS Blast_tac); qed "R_PSP"; -goal thy +Goal "!!Acts. [| LeadsTo Init Acts A A'; constrains Acts B B'; id: Acts |] \ \ ==> LeadsTo Init Acts (B Int A) ((B Int A') Un (B' - B))"; by (asm_simp_tac (simpset() addsimps (R_PSP::Int_ac)) 1); qed "R_PSP2"; -goalw thy [unless_def] +Goalw [unless_def] "!!Acts. [| LeadsTo Init Acts A A'; unless Acts B B'; id: Acts |] \ \ ==> LeadsTo Init Acts (A Int B) ((A' Int B) Un B')"; by (dtac R_PSP 1); @@ -287,7 +287,7 @@ (*** Induction rules ***) (** Meta or object quantifier ????? **) -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| wf r; \ \ ALL m. LeadsTo Init Acts (A Int f-``{m}) \ \ ((A Int f-``(r^-1 ^^ {m})) Un B); \ @@ -299,7 +299,7 @@ qed "LeadsTo_wf_induct"; -goal thy +Goal "!!Acts. [| wf r; \ \ ALL m:I. LeadsTo Init Acts (A Int f-``{m}) \ \ ((A Int f-``(r^-1 ^^ {m})) Un B); \ @@ -313,7 +313,7 @@ qed "R_bounded_induct"; -goal thy +Goal "!!Acts. [| ALL m. LeadsTo Init Acts (A Int f-``{m}) \ \ ((A Int f-``(lessThan m)) Un B); \ \ id: Acts |] \ @@ -323,7 +323,7 @@ by (Asm_simp_tac 1); qed "R_lessThan_induct"; -goal thy +Goal "!!Acts. [| ALL m:(greaterThan l). LeadsTo Init Acts (A Int f-``{m}) \ \ ((A Int f-``(lessThan m)) Un B); \ \ id: Acts |] \ @@ -334,7 +334,7 @@ by (Asm_simp_tac 1); qed "R_lessThan_bounded_induct"; -goal thy +Goal "!!Acts. [| ALL m:(lessThan l). LeadsTo Init Acts (A Int f-``{m}) \ \ ((A Int f-``(greaterThan m)) Un B); \ \ id: Acts |] \ @@ -353,7 +353,7 @@ (*** Completion: Binary and General Finite versions ***) -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| LeadsTo Init Acts A A'; stable Acts A'; \ \ LeadsTo Init Acts B B'; stable Acts B'; id: Acts |] \ \ ==> LeadsTo Init Acts (A Int B) (A' Int B')"; @@ -362,7 +362,7 @@ qed "R_stable_completion"; -goal thy +Goal "!!Acts. [| finite I; id: Acts |] \ \ ==> (ALL i:I. LeadsTo Init Acts (A i) (A' i)) --> \ \ (ALL i:I. stable Acts (A' i)) --> \ @@ -375,7 +375,7 @@ qed_spec_mp "R_finite_stable_completion"; -goalw thy [LeadsTo_def] +Goalw [LeadsTo_def] "!!Acts. [| LeadsTo Init Acts A (A' Un C); constrains Acts A' (A' Un C); \ \ LeadsTo Init Acts B (B' Un C); constrains Acts B' (B' Un C); \ \ id: Acts |] \ @@ -391,7 +391,7 @@ qed "R_completion"; -goal thy +Goal "!!Acts. [| finite I; id: Acts |] \ \ ==> (ALL i:I. LeadsTo Init Acts (A i) (A' i Un C)) --> \ \ (ALL i:I. constrains Acts (A' i) (A' i Un C)) --> \ diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/Token.ML --- a/src/HOL/UNITY/Token.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/Token.ML Mon Jun 22 17:26:46 1998 +0200 @@ -16,11 +16,11 @@ AddIffs [N_positive, skip]; -goalw thy [HasTok_def] "!!i. [| s: HasTok i; s: HasTok j |] ==> i=j"; +Goalw [HasTok_def] "!!i. [| s: HasTok i; s: HasTok j |] ==> i=j"; by Auto_tac; qed "HasToK_partition"; -goalw thy Token_defs "!!s. s : WellTyped ==> (s ~: E i) = (s : H i | s : T i)"; +Goalw Token_defs "!!s. s : WellTyped ==> (s ~: E i) = (s : H i | s : T i)"; by (Simp_tac 1); by (exhaust_tac "s (Suc i)" 1); by Auto_tac; @@ -32,7 +32,7 @@ we'd not have to mention WellTyped in the statement of this theorem. **) -goalw thy [stable_def] +Goalw [stable_def] "stable Acts (WellTyped Int {s. s : E i --> s : HasTok i})"; by (rtac (stable_WT RS stable_constrains_Int) 1); by (cut_facts_tac [TR2,TR4,TR5] 1); @@ -44,7 +44,7 @@ qed "token_stable"; (*This proof is in the "massaging" style and is much faster! *) -goalw thy [stable_def] +Goalw [stable_def] "stable Acts (WellTyped Int (Compl(E i) Un (HasTok i)))"; by (rtac (stable_WT RS stable_constrains_Int) 1); by (rtac constrains_weaken 1); @@ -56,17 +56,17 @@ (*** Progress under weak fairness ***) -goalw thy [nodeOrder_def] "wf(nodeOrder j)"; +Goalw [nodeOrder_def] "wf(nodeOrder j)"; by (rtac (wf_less_than RS wf_inv_image RS wf_subset) 1); by (Blast_tac 1); qed"wf_nodeOrder"; - goal thy "!!m. [| m Suc(m) < n"; +Goal "!!m. [| m Suc(m) < n"; by (full_simp_tac (simpset() addsimps [nat_neq_iff]) 1); by (blast_tac (claset() addEs [less_asym, less_SucE]) 1); qed "Suc_lessI"; -goalw thy [nodeOrder_def, next_def, inv_image_def] +Goalw [nodeOrder_def, next_def, inv_image_def] "!!x. [| i ((next i, i) : nodeOrder j) = (i ~= j)"; by (auto_tac (claset(), simpset() addsimps [Suc_lessI, mod_Suc, mod_less, mod_geq])); @@ -99,7 +99,7 @@ (*From "A Logic for Concurrent Programming", but not used in Chapter 4. Note the use of case_tac. Reasoning about leadsTo takes practice!*) -goal thy +Goal "!!i. [| i \ \ leadsTo Acts (HasTok i) ({s. (Token s, i) : nodeOrder j} Un HasTok j)"; by (case_tac "i=j" 1); @@ -111,7 +111,7 @@ (*Chapter 4 variant, the one actually used below.*) -goal thy +Goal "!!i. [| i \ \ leadsTo Acts (HasTok i) {s. (Token s, i) : nodeOrder j}"; by (rtac (TR7 RS leadsTo_weaken_R) 1); @@ -119,14 +119,14 @@ by (asm_full_simp_tac (simpset() addsimps [HasTok_def, nodeOrder_eq]) 1); qed "TR7_aux"; -goal thy "({s. Token s < N} Int Token -`` {m}) = \ +Goal "({s. Token s < N} Int Token -`` {m}) = \ \ (if m leadsTo Acts {s. Token s < N} (HasTok j)"; by (rtac leadsTo_weaken_R 1); by (res_inst_tac [("I", "Compl{j}"), ("f", "Token"), ("B", "{}")] @@ -141,7 +141,7 @@ qed "leadsTo_j"; (*Misra's TR8: a hungry process eventually eats*) -goal thy "!!i. j leadsTo Acts ({s. Token s < N} Int H j) (E j)"; +Goal "!!i. j leadsTo Acts ({s. Token s < N} Int H j) (E j)"; by (rtac (leadsTo_cancel1 RS leadsTo_Un_duplicate) 1); by (rtac TR6 2); by (rtac leadsTo_weaken_R 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/Traces.ML --- a/src/HOL/UNITY/Traces.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/Traces.ML Mon Jun 22 17:26:46 1998 +0200 @@ -1,39 +1,39 @@ open Traces; -goal thy "reachable Init Acts <= {s. EX evs. s#evs: traces Init Acts}"; +Goal "reachable Init Acts <= {s. EX evs. s#evs: traces Init Acts}"; by (rtac subsetI 1); be reachable.induct 1; by (REPEAT (blast_tac (claset() addIs traces.intrs) 1)); val lemma1 = result(); -goal thy "!!evs. evs : traces Init Acts \ +Goal "!!evs. evs : traces Init Acts \ \ ==> ALL s evs'. evs = s#evs' --> s: reachable Init Acts"; be traces.induct 1; by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1)); val lemma2 = normalize_thm [RSmp, RSspec] (result()); -goal thy "{s. EX evs. s#evs: traces Init Acts} <= reachable Init Acts"; +Goal "{s. EX evs. s#evs: traces Init Acts} <= reachable Init Acts"; by (blast_tac (claset() addIs [lemma2]) 1); val lemma3 = result(); -goal thy "reachable Init Acts = {s. EX evs. s#evs: traces Init Acts}"; +Goal "reachable Init Acts = {s. EX evs. s#evs: traces Init Acts}"; by (rtac ([lemma1,lemma3] MRS equalityI) 1); qed "reachable_eq_traces"; (*Could/should this be proved?*) -goal thy "reachable Init Acts = (UN evs: traces Init Acts. set evs)"; +Goal "reachable Init Acts = (UN evs: traces Init Acts. set evs)"; (*The strongest invariant is the set of all reachable states!*) -goalw thy [stable_def, constrains_def] +Goalw [stable_def, constrains_def] "!!A. [| Init<=A; stable Acts A |] ==> reachable Init Acts <= A"; by (rtac subsetI 1); be reachable.induct 1; by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1)); qed "strongest_invariant"; -goal thy "stable Acts (reachable Init Acts)"; +Goal "stable Acts (reachable Init Acts)"; by (REPEAT (blast_tac (claset() addIs [stableI, constrainsI] addIs reachable.intrs) 1)); qed "stable_reachable"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/UNITY.ML --- a/src/HOL/UNITY/UNITY.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/UNITY.ML Mon Jun 22 17:26:46 1998 +0200 @@ -18,36 +18,36 @@ be any set including A Int C and contained in A Un C, such as B=A or B=C. **) -goal thy "!!x. [| A Int C <= B; B <= A Un C |] \ +Goal "!!x. [| A Int C <= B; B <= A Un C |] \ \ ==> (A Int B) Un (Compl A Int C) = B Un C"; by (Blast_tac 1); -goal thy "!!x. [| A Int C <= B; B <= A Un C |] \ +Goal "!!x. [| A Int C <= B; B <= A Un C |] \ \ ==> (A Un B) Int (Compl A Un C) = B Int C"; by (Blast_tac 1); (*The base B=A*) -goal thy "A Un (Compl A Int C) = A Un C"; +Goal "A Un (Compl A Int C) = A Un C"; by (Blast_tac 1); -goal thy "A Int (Compl A Un C) = A Int C"; +Goal "A Int (Compl A Un C) = A Int C"; by (Blast_tac 1); (*The base B=C*) -goal thy "(A Int C) Un (Compl A Int C) = C"; +Goal "(A Int C) Un (Compl A Int C) = C"; by (Blast_tac 1); -goal thy "(A Un C) Int (Compl A Un C) = C"; +Goal "(A Un C) Int (Compl A Un C) = C"; by (Blast_tac 1); (** More ad-hoc rules **) -goal thy "A Un B - (A - B) = B"; +Goal "A Un B - (A - B) = B"; by (Blast_tac 1); qed "Un_Diff_Diff"; -goal thy "A Int (B - C) Un C = A Int B Un C"; +Goal "A Int (B - C) Un C = A Int B Un C"; by (Blast_tac 1); qed "Int_Diff_Un"; @@ -63,38 +63,38 @@ by (blast_tac (claset() addIs prems) 1); qed "constrainsI"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A A'; act: Acts; (s,s'): act; s: A |] \ \ ==> s': A'"; by (Blast_tac 1); qed "constrainsD"; -goalw thy [constrains_def] "constrains Acts {} B"; +Goalw [constrains_def] "constrains Acts {} B"; by (Blast_tac 1); qed "constrains_empty"; -goalw thy [constrains_def] "constrains Acts A UNIV"; +Goalw [constrains_def] "constrains Acts A UNIV"; by (Blast_tac 1); qed "constrains_UNIV"; AddIffs [constrains_empty, constrains_UNIV]; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A A'; A'<=B' |] ==> constrains Acts A B'"; by (Blast_tac 1); qed "constrains_weaken_R"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A A'; B<=A |] ==> constrains Acts B A'"; by (Blast_tac 1); qed "constrains_weaken_L"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A A'; B<=A; A'<=B' |] ==> constrains Acts B B'"; by (Blast_tac 1); qed "constrains_weaken"; (*Set difference: UNUSED*) -goalw thy [constrains_def] +Goalw [constrains_def] "!!C. [| constrains Acts (A-B) C; constrains Acts B C |] \ \ ==> constrains Acts A C"; by (Blast_tac 1); @@ -102,19 +102,19 @@ (** Union **) -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A A'; constrains Acts B B' |] \ \ ==> constrains Acts (A Un B) (A' Un B')"; by (Blast_tac 1); qed "constrains_Un"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. ALL i:I. constrains Acts (A i) (A' i) \ \ ==> constrains Acts (UN i:I. A i) (UN i:I. A' i)"; by (Blast_tac 1); qed "ball_constrains_UN"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| ALL i. constrains Acts (A i) (A' i) |] \ \ ==> constrains Acts (UN i. A i) (UN i. A' i)"; by (Blast_tac 1); @@ -122,31 +122,31 @@ (** Intersection **) -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A A'; constrains Acts B B' |] \ \ ==> constrains Acts (A Int B) (A' Int B')"; by (Blast_tac 1); qed "constrains_Int"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. ALL i:I. constrains Acts (A i) (A' i) \ \ ==> constrains Acts (INT i:I. A i) (INT i:I. A' i)"; by (Blast_tac 1); qed "ball_constrains_INT"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| ALL i. constrains Acts (A i) (A' i) |] \ \ ==> constrains Acts (INT i. A i) (INT i. A' i)"; by (Blast_tac 1); qed "all_constrains_INT"; -goalw thy [stable_def, constrains_def] +Goalw [stable_def, constrains_def] "!!Acts. [| stable Acts C; constrains Acts A (C Un A') |] \ \ ==> constrains Acts (C Un A) (C Un A')"; by (Blast_tac 1); qed "stable_constrains_Un"; -goalw thy [stable_def, constrains_def] +Goalw [stable_def, constrains_def] "!!Acts. [| stable Acts C; constrains Acts (C Int A) A' |] \ \ ==> constrains Acts (C Int A) (C Int A')"; by (Blast_tac 1); @@ -155,35 +155,35 @@ (*** stable ***) -goalw thy [stable_def] +Goalw [stable_def] "!!Acts. constrains Acts A A ==> stable Acts A"; by (assume_tac 1); qed "stableI"; -goalw thy [stable_def] +Goalw [stable_def] "!!Acts. stable Acts A ==> constrains Acts A A"; by (assume_tac 1); qed "stableD"; -goalw thy [stable_def] +Goalw [stable_def] "!!Acts. [| stable Acts A; stable Acts A' |] \ \ ==> stable Acts (A Un A')"; by (blast_tac (claset() addIs [constrains_Un]) 1); qed "stable_Un"; -goalw thy [stable_def] +Goalw [stable_def] "!!Acts. [| stable Acts A; stable Acts A' |] \ \ ==> stable Acts (A Int A')"; by (blast_tac (claset() addIs [constrains_Int]) 1); qed "stable_Int"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A A'; id: Acts |] ==> A<=A'"; by (Blast_tac 1); qed "constrains_imp_subset"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| id: Acts; constrains Acts A B; constrains Acts B C |] \ \ ==> constrains Acts A C"; by (Blast_tac 1); @@ -193,7 +193,7 @@ (*The Elimination Theorem. The "free" m has become universally quantified! Should the premise be !!m instead of ALL m ? Would make it harder to use in forward proof.*) -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| ALL m. constrains Acts {s. s x = m} (B m) |] \ \ ==> constrains Acts {s. P(s x)} (UN m. {s. P(m)} Int B m)"; by (Blast_tac 1); @@ -201,14 +201,14 @@ (*As above, but for the trivial case of a one-variable state, in which the state is identified with its one variable.*) -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| ALL m. constrains Acts {m} (B m) |] \ \ ==> constrains Acts {s. P s} (UN m. {s. P(m)} Int B m)"; by (Blast_tac 1); qed "elimination_sing"; -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| constrains Acts A (A' Un B); constrains Acts B B'; id: Acts |] \ \ ==> constrains Acts A (A' Un B')"; by (Blast_tac 1); @@ -218,12 +218,12 @@ (*** Theoretical Results from Section 6 ***) -goalw thy [constrains_def, strongest_rhs_def] +Goalw [constrains_def, strongest_rhs_def] "constrains Acts A (strongest_rhs Acts A )"; by (Blast_tac 1); qed "constrains_strongest_rhs"; -goalw thy [constrains_def, strongest_rhs_def] +Goalw [constrains_def, strongest_rhs_def] "!!Acts. constrains Acts A B ==> strongest_rhs Acts A <= B"; by (Blast_tac 1); qed "strongest_rhs_is_strongest"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/UNITY/WFair.ML --- a/src/HOL/UNITY/WFair.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/UNITY/WFair.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,26 +10,26 @@ open WFair; -goal thy "Union(B) Int A = Union((%C. C Int A)``B)"; +Goal "Union(B) Int A = Union((%C. C Int A)``B)"; by (Blast_tac 1); qed "Int_Union_Union"; (*** transient ***) -goalw thy [stable_def, constrains_def, transient_def] +Goalw [stable_def, constrains_def, transient_def] "!!A. [| stable Acts A; transient Acts A |] ==> A = {}"; by (Blast_tac 1); qed "stable_transient_empty"; -goalw thy [transient_def] +Goalw [transient_def] "!!A. [| transient Acts A; B<=A |] ==> transient Acts B"; by (Clarify_tac 1); by (rtac bexI 1 THEN assume_tac 2); by (Blast_tac 1); qed "transient_strengthen"; -goalw thy [transient_def] +Goalw [transient_def] "!!A. [| act:Acts; A <= Domain act; act^^A <= Compl A |] \ \ ==> transient Acts A"; by (Blast_tac 1); @@ -38,31 +38,31 @@ (*** ensures ***) -goalw thy [ensures_def] +Goalw [ensures_def] "!!Acts. [| constrains Acts (A-B) (A Un B); transient Acts (A-B) |] \ \ ==> ensures Acts A B"; by (Blast_tac 1); qed "ensuresI"; -goalw thy [ensures_def] +Goalw [ensures_def] "!!Acts. ensures Acts A B \ \ ==> constrains Acts (A-B) (A Un B) & transient Acts (A-B)"; by (Blast_tac 1); qed "ensuresD"; (*The L-version (precondition strengthening) doesn't hold for ENSURES*) -goalw thy [ensures_def] +Goalw [ensures_def] "!!Acts. [| ensures Acts A A'; A'<=B' |] ==> ensures Acts A B'"; by (blast_tac (claset() addIs [constrains_weaken, transient_strengthen]) 1); qed "ensures_weaken_R"; -goalw thy [ensures_def, constrains_def, transient_def] +Goalw [ensures_def, constrains_def, transient_def] "!!Acts. Acts ~= {} ==> ensures Acts A UNIV"; by (Asm_simp_tac 1); (*omitting this causes PROOF FAILED, even with Safe_tac*) by (Blast_tac 1); qed "ensures_UNIV"; -goalw thy [ensures_def] +Goalw [ensures_def] "!!Acts. [| stable Acts C; \ \ constrains Acts (C Int (A - A')) (A Un A'); \ \ transient Acts (C Int (A-A')) |] \ @@ -79,17 +79,17 @@ bind_thm ("leadsTo_Basis", leadsto.Basis); bind_thm ("leadsTo_Trans", leadsto.Trans); -goal thy "!!Acts. act: Acts ==> leadsTo Acts A UNIV"; +Goal "!!Acts. act: Acts ==> leadsTo Acts A UNIV"; by (blast_tac (claset() addIs [ensures_UNIV RS leadsTo_Basis]) 1); qed "leadsTo_UNIV"; Addsimps [leadsTo_UNIV]; (*Useful with cancellation, disjunction*) -goal thy "!!Acts. leadsTo Acts A (A' Un A') ==> leadsTo Acts A A'"; +Goal "!!Acts. leadsTo Acts A (A' Un A') ==> leadsTo Acts A A'"; by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); qed "leadsTo_Un_duplicate"; -goal thy "!!Acts. leadsTo Acts A (A' Un C Un C) ==> leadsTo Acts A (A' Un C)"; +Goal "!!Acts. leadsTo Acts A (A' Un C Un C) ==> leadsTo Acts A (A' Un C)"; by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); qed "leadsTo_Un_duplicate2"; @@ -106,7 +106,7 @@ qed "leadsTo_UN"; (*Binary union introduction rule*) -goal thy +Goal "!!C. [| leadsTo Acts A C; leadsTo Acts B C |] ==> leadsTo Acts (A Un B) C"; by (stac Un_eq_Union 1); by (blast_tac (claset() addIs [leadsTo_Union]) 1); @@ -126,7 +126,7 @@ qed "leadsTo_induct"; -goal thy "!!A B. [| A<=B; id: Acts |] ==> leadsTo Acts A B"; +Goal "!!A B. [| A<=B; id: Acts |] ==> leadsTo Acts A B"; by (rtac leadsTo_Basis 1); by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]); by (Blast_tac 1); @@ -138,7 +138,7 @@ (*There's a direct proof by leadsTo_Trans and subset_imp_leadsTo, but it needs the extra premise id:Acts*) -goal thy "!!Acts. leadsTo Acts A A' ==> A'<=B' --> leadsTo Acts A B'"; +Goal "!!Acts. leadsTo Acts A A' ==> A'<=B' --> leadsTo Acts A B'"; by (etac leadsTo_induct 1); by (Clarify_tac 3); by (blast_tac (claset() addIs [leadsTo_Union]) 3); @@ -147,33 +147,33 @@ qed_spec_mp "leadsTo_weaken_R"; -goal thy "!!Acts. [| leadsTo Acts A A'; B<=A; id: Acts |] ==> \ +Goal "!!Acts. [| leadsTo Acts A A'; B<=A; id: Acts |] ==> \ \ leadsTo Acts B A'"; by (blast_tac (claset() addIs [leadsTo_Basis, leadsTo_Trans, subset_imp_leadsTo]) 1); qed_spec_mp "leadsTo_weaken_L"; (*Distributes over binary unions*) -goal thy +Goal "!!C. id: Acts ==> \ \ leadsTo Acts (A Un B) C = (leadsTo Acts A C & leadsTo Acts B C)"; by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken_L]) 1); qed "leadsTo_Un_distrib"; -goal thy +Goal "!!C. id: Acts ==> \ \ leadsTo Acts (UN i:I. A i) B = (ALL i : I. leadsTo Acts (A i) B)"; by (blast_tac (claset() addIs [leadsTo_UN, leadsTo_weaken_L]) 1); qed "leadsTo_UN_distrib"; -goal thy +Goal "!!C. id: Acts ==> \ \ leadsTo Acts (Union S) B = (ALL A : S. leadsTo Acts A B)"; by (blast_tac (claset() addIs [leadsTo_Union, leadsTo_weaken_L]) 1); qed "leadsTo_Union_distrib"; -goal thy +Goal "!!Acts. [| leadsTo Acts A A'; id: Acts; B<=A; A'<=B' |] \ \ ==> leadsTo Acts B B'"; (*PROOF FAILED: why?*) @@ -183,7 +183,7 @@ (*Set difference: maybe combine with leadsTo_weaken_L??*) -goal thy +Goal "!!C. [| leadsTo Acts (A-B) C; leadsTo Acts B C; id: Acts |] \ \ ==> leadsTo Acts A C"; by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken]) 1); @@ -211,7 +211,7 @@ qed "leadsTo_UN_UN_noindex"; (*Version with no index set*) -goal thy +Goal "!!Acts. ALL i. leadsTo Acts (A i) (A' i) \ \ ==> leadsTo Acts (UN i. A i) (UN i. A' i)"; by (blast_tac (claset() addIs [leadsTo_UN_UN]) 1); @@ -219,7 +219,7 @@ (*Binary union version*) -goal thy "!!Acts. [| leadsTo Acts A A'; leadsTo Acts B B' |] \ +Goal "!!Acts. [| leadsTo Acts A A'; leadsTo Acts B B' |] \ \ ==> leadsTo Acts (A Un B) (A' Un B')"; by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken_R]) 1); @@ -228,14 +228,14 @@ (** The cancellation law **) -goal thy +Goal "!!Acts. [| leadsTo Acts A (A' Un B); leadsTo Acts B B'; id: Acts |] \ \ ==> leadsTo Acts A (A' Un B')"; by (blast_tac (claset() addIs [leadsTo_Un_Un, subset_imp_leadsTo, leadsTo_Trans]) 1); qed "leadsTo_cancel2"; -goal thy +Goal "!!Acts. [| leadsTo Acts A (A' Un B); leadsTo Acts (B-A') B'; id: Acts |] \ \ ==> leadsTo Acts A (A' Un B')"; by (rtac leadsTo_cancel2 1); @@ -243,14 +243,14 @@ by (ALLGOALS Asm_simp_tac); qed "leadsTo_cancel_Diff2"; -goal thy +Goal "!!Acts. [| leadsTo Acts A (B Un A'); leadsTo Acts B B'; id: Acts |] \ \ ==> leadsTo Acts A (B' Un A')"; by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1); by (blast_tac (claset() addSIs [leadsTo_cancel2]) 1); qed "leadsTo_cancel1"; -goal thy +Goal "!!Acts. [| leadsTo Acts A (B Un A'); leadsTo Acts (B-A') B'; id: Acts |] \ \ ==> leadsTo Acts A (B' Un A')"; by (rtac leadsTo_cancel1 1); @@ -262,14 +262,14 @@ (** The impossibility law **) -goal thy "!!Acts. leadsTo Acts A B ==> B={} --> A={}"; +Goal "!!Acts. leadsTo Acts A B ==> B={} --> A={}"; by (etac leadsTo_induct 1); by (ALLGOALS Asm_simp_tac); by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]); by (Blast_tac 1); val lemma = result() RS mp; -goal thy "!!Acts. leadsTo Acts A {} ==> A={}"; +Goal "!!Acts. leadsTo Acts A {} ==> A={}"; by (blast_tac (claset() addSIs [lemma]) 1); qed "leadsTo_empty"; @@ -277,7 +277,7 @@ (** PSP: Progress-Safety-Progress **) (*Special case of PSP: Misra's "stable conjunction". Doesn't need id:Acts. *) -goalw thy [stable_def] +Goalw [stable_def] "!!Acts. [| leadsTo Acts A A'; stable Acts B |] \ \ ==> leadsTo Acts (A Int B) (A' Int B)"; by (etac leadsTo_induct 1); @@ -291,14 +291,14 @@ by (blast_tac (claset() addIs [transient_strengthen, constrains_Int]) 1); qed "PSP_stable"; -goal thy +Goal "!!Acts. [| leadsTo Acts A A'; stable Acts B |] \ \ ==> leadsTo Acts (B Int A) (B Int A')"; by (asm_simp_tac (simpset() addsimps (PSP_stable::Int_ac)) 1); qed "PSP_stable2"; -goalw thy [ensures_def] +Goalw [ensures_def] "!!Acts. [| ensures Acts A A'; constrains Acts B B' |] \ \ ==> ensures Acts (A Int B) ((A' Int B) Un (B' - B))"; by Safe_tac; @@ -308,7 +308,7 @@ qed "PSP_ensures"; -goal thy +Goal "!!Acts. [| leadsTo Acts A A'; constrains Acts B B'; id: Acts |] \ \ ==> leadsTo Acts (A Int B) ((A' Int B) Un (B' - B))"; by (etac leadsTo_induct 1); @@ -323,14 +323,14 @@ by (blast_tac (claset() addIs [leadsTo_Basis, PSP_ensures]) 1); qed "PSP"; -goal thy +Goal "!!Acts. [| leadsTo Acts A A'; constrains Acts B B'; id: Acts |] \ \ ==> leadsTo Acts (B Int A) ((B Int A') Un (B' - B))"; by (asm_simp_tac (simpset() addsimps (PSP::Int_ac)) 1); qed "PSP2"; -goalw thy [unless_def] +Goalw [unless_def] "!!Acts. [| leadsTo Acts A A'; unless Acts B B'; id: Acts |] \ \ ==> leadsTo Acts (A Int B) ((A' Int B) Un B')"; by (dtac PSP 1); @@ -345,7 +345,7 @@ (*** Proving the induction rules ***) -goal thy +Goal "!!Acts. [| wf r; \ \ ALL m. leadsTo Acts (A Int f-``{m}) \ \ ((A Int f-``(r^-1 ^^ {m})) Un B); \ @@ -361,7 +361,7 @@ (** Meta or object quantifier ????? **) -goal thy +Goal "!!Acts. [| wf r; \ \ ALL m. leadsTo Acts (A Int f-``{m}) \ \ ((A Int f-``(r^-1 ^^ {m})) Un B); \ @@ -375,7 +375,7 @@ qed "leadsTo_wf_induct"; -goal thy +Goal "!!Acts. [| wf r; \ \ ALL m:I. leadsTo Acts (A Int f-``{m}) \ \ ((A Int f-``(r^-1 ^^ {m})) Un B); \ @@ -390,7 +390,7 @@ (*Alternative proof is via the lemma leadsTo Acts (A Int f-``(lessThan m)) B*) -goal thy +Goal "!!Acts. [| ALL m. leadsTo Acts (A Int f-``{m}) \ \ ((A Int f-``(lessThan m)) Un B); \ \ id: Acts |] \ @@ -400,7 +400,7 @@ by (Asm_simp_tac 1); qed "lessThan_induct"; -goal thy +Goal "!!Acts. [| ALL m:(greaterThan l). leadsTo Acts (A Int f-``{m}) \ \ ((A Int f-``(lessThan m)) Un B); \ \ id: Acts |] \ @@ -411,7 +411,7 @@ by (Asm_simp_tac 1); qed "lessThan_bounded_induct"; -goal thy +Goal "!!Acts. [| ALL m:(lessThan l). leadsTo Acts (A Int f-``{m}) \ \ ((A Int f-``(greaterThan m)) Un B); \ \ id: Acts |] \ @@ -431,29 +431,29 @@ (*** wlt ****) (*Misra's property W3*) -goalw thy [wlt_def] "leadsTo Acts (wlt Acts B) B"; +Goalw [wlt_def] "leadsTo Acts (wlt Acts B) B"; by (blast_tac (claset() addSIs [leadsTo_Union]) 1); qed "wlt_leadsTo"; -goalw thy [wlt_def] "!!Acts. leadsTo Acts A B ==> A <= wlt Acts B"; +Goalw [wlt_def] "!!Acts. leadsTo Acts A B ==> A <= wlt Acts B"; by (blast_tac (claset() addSIs [leadsTo_Union]) 1); qed "leadsTo_subset"; (*Misra's property W2*) -goal thy "!!Acts. id: Acts ==> leadsTo Acts A B = (A <= wlt Acts B)"; +Goal "!!Acts. id: Acts ==> leadsTo Acts A B = (A <= wlt Acts B)"; by (blast_tac (claset() addSIs [leadsTo_subset, wlt_leadsTo RS leadsTo_weaken_L]) 1); qed "leadsTo_eq_subset_wlt"; (*Misra's property W4*) -goal thy "!!Acts. id: Acts ==> B <= wlt Acts B"; +Goal "!!Acts. id: Acts ==> B <= wlt Acts B"; by (asm_simp_tac (simpset() addsimps [leadsTo_eq_subset_wlt RS sym, subset_imp_leadsTo]) 1); qed "wlt_increasing"; (*Used in the Trans case below*) -goalw thy [constrains_def] +Goalw [constrains_def] "!!Acts. [| B <= A2; \ \ constrains Acts (A1 - B) (A1 Un B); \ \ constrains Acts (A2 - C) (A2 Un C) |] \ @@ -464,7 +464,7 @@ (*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*) -goal thy +Goal "!!Acts. [| leadsTo Acts A A'; id: Acts |] ==> \ \ EX B. A<=B & leadsTo Acts B A' & constrains Acts (B-A') (B Un A')"; by (etac leadsTo_induct 1); @@ -485,7 +485,7 @@ (*Misra's property W5*) -goal thy "!!Acts. id: Acts ==> constrains Acts (wlt Acts B - B) (wlt Acts B)"; +Goal "!!Acts. id: Acts ==> constrains Acts (wlt Acts B - B) (wlt Acts B)"; by (forward_tac [wlt_leadsTo RS leadsTo_123] 1); by (Clarify_tac 1); by (subgoal_tac "Ba = wlt Acts B" 1); @@ -497,7 +497,7 @@ (*** Completion: Binary and General Finite versions ***) -goal thy +Goal "!!Acts. [| leadsTo Acts A A'; stable Acts A'; \ \ leadsTo Acts B B'; stable Acts B'; id: Acts |] \ \ ==> leadsTo Acts (A Int B) (A' Int B')"; @@ -519,7 +519,7 @@ qed "stable_completion"; -goal thy +Goal "!!Acts. [| finite I; id: Acts |] \ \ ==> (ALL i:I. leadsTo Acts (A i) (A' i)) --> \ \ (ALL i:I. stable Acts (A' i)) --> \ @@ -532,7 +532,7 @@ qed_spec_mp "finite_stable_completion"; -goal thy +Goal "!!Acts. [| W = wlt Acts (B' Un C); \ \ leadsTo Acts A (A' Un C); constrains Acts A' (A' Un C); \ \ leadsTo Acts B (B' Un C); constrains Acts B' (B' Un C); \ @@ -562,7 +562,7 @@ bind_thm("completion", refl RS result()); -goal thy +Goal "!!Acts. [| finite I; id: Acts |] \ \ ==> (ALL i:I. leadsTo Acts (A i) (A' i Un C)) --> \ \ (ALL i:I. constrains Acts (A' i) (A' i Un C)) --> \ diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Univ.ML --- a/src/HOL/Univ.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Univ.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,7 +10,7 @@ (** apfst -- can be used in similar type definitions **) -goalw Univ.thy [apfst_def] "apfst f (a,b) = (f(a),b)"; +Goalw [apfst_def] "apfst f (a,b) = (f(a),b)"; by (rtac split 1); qed "apfst_conv"; @@ -53,12 +53,12 @@ (*** Isomorphisms ***) -goal Univ.thy "inj(Rep_Node)"; +Goal "inj(Rep_Node)"; by (rtac inj_inverseI 1); (*cannot combine by RS: multiple unifiers*) by (rtac Rep_Node_inverse 1); qed "inj_Rep_Node"; -goal Univ.thy "inj_on Abs_Node Node"; +Goal "inj_on Abs_Node Node"; by (rtac inj_on_inverseI 1); by (etac Abs_Node_inverse 1); qed "inj_on_Abs_Node"; @@ -68,11 +68,11 @@ (*** Introduction rules for Node ***) -goalw Univ.thy [Node_def] "(%k. 0,a) : Node"; +Goalw [Node_def] "(%k. 0,a) : Node"; by (Blast_tac 1); qed "Node_K0_I"; -goalw Univ.thy [Node_def,Push_def] +Goalw [Node_def,Push_def] "!!p. p: Node ==> apfst (Push i) p : Node"; by (blast_tac (claset() addSIs [apfst_conv, nat_case_Suc RS trans]) 1); qed "Node_Push_I"; @@ -82,7 +82,7 @@ (** Scons vs Atom **) -goalw Univ.thy [Atom_def,Scons_def,Push_Node_def] "(M$N) ~= Atom(a)"; +Goalw [Atom_def,Scons_def,Push_Node_def] "(M$N) ~= Atom(a)"; by (rtac notI 1); by (etac (equalityD2 RS subsetD RS UnE) 1); by (rtac singletonI 1); @@ -97,17 +97,17 @@ (** Atomic nodes **) -goalw Univ.thy [Atom_def, inj_def] "inj(Atom)"; +Goalw [Atom_def, inj_def] "inj(Atom)"; by (blast_tac (claset() addSIs [Node_K0_I] addSDs [Abs_Node_inject]) 1); qed "inj_Atom"; val Atom_inject = inj_Atom RS injD; -goal Univ.thy "(Atom(a)=Atom(b)) = (a=b)"; +Goal "(Atom(a)=Atom(b)) = (a=b)"; by (blast_tac (claset() addSDs [Atom_inject]) 1); qed "Atom_Atom_eq"; AddIffs [Atom_Atom_eq]; -goalw Univ.thy [Leaf_def,o_def] "inj(Leaf)"; +Goalw [Leaf_def,o_def] "inj(Leaf)"; by (rtac injI 1); by (etac (Atom_inject RS Inl_inject) 1); qed "inj_Leaf"; @@ -115,7 +115,7 @@ val Leaf_inject = inj_Leaf RS injD; AddSDs [Leaf_inject]; -goalw Univ.thy [Numb_def,o_def] "inj(Numb)"; +Goalw [Numb_def,o_def] "inj(Numb)"; by (rtac injI 1); by (etac (Atom_inject RS Inr_inject) 1); qed "inj_Numb"; @@ -142,11 +142,11 @@ (** Injectiveness of Scons **) -goalw Univ.thy [Scons_def] "!!M. M$N <= M'$N' ==> M<=M'"; +Goalw [Scons_def] "!!M. M$N <= M'$N' ==> M<=M'"; by (blast_tac (claset() addSDs [Push_Node_inject]) 1); qed "Scons_inject_lemma1"; -goalw Univ.thy [Scons_def] "!!M. M$N <= M'$N' ==> N<=N'"; +Goalw [Scons_def] "!!M. M$N <= M'$N' ==> N<=N'"; by (blast_tac (claset() addSDs [Push_Node_inject]) 1); qed "Scons_inject_lemma2"; @@ -166,7 +166,7 @@ by (rtac ((major RS Scons_inject2) RS ((major RS Scons_inject1) RS minor)) 1); qed "Scons_inject"; -goal Univ.thy "(M$N = M'$N') = (M=M' & N=N')"; +Goal "(M$N = M'$N') = (M=M' & N=N')"; by (blast_tac (claset() addSEs [Scons_inject]) 1); qed "Scons_Scons_eq"; @@ -174,7 +174,7 @@ (** Scons vs Leaf **) -goalw Univ.thy [Leaf_def,o_def] "(M$N) ~= Leaf(a)"; +Goalw [Leaf_def,o_def] "(M$N) ~= Leaf(a)"; by (rtac Scons_not_Atom 1); qed "Scons_not_Leaf"; bind_thm ("Leaf_not_Scons", Scons_not_Leaf RS not_sym); @@ -184,7 +184,7 @@ (** Scons vs Numb **) -goalw Univ.thy [Numb_def,o_def] "(M$N) ~= Numb(k)"; +Goalw [Numb_def,o_def] "(M$N) ~= Numb(k)"; by (rtac Scons_not_Atom 1); qed "Scons_not_Numb"; bind_thm ("Numb_not_Scons", Scons_not_Numb RS not_sym); @@ -194,7 +194,7 @@ (** Leaf vs Numb **) -goalw Univ.thy [Leaf_def,Numb_def] "Leaf(a) ~= Numb(k)"; +Goalw [Leaf_def,Numb_def] "Leaf(a) ~= Numb(k)"; by (simp_tac (simpset() addsimps [Inl_not_Inr]) 1); qed "Leaf_not_Numb"; bind_thm ("Numb_not_Leaf", Leaf_not_Numb RS not_sym); @@ -208,14 +208,14 @@ AddIffs [Scons_not_Atom, Atom_not_Scons, Scons_Scons_eq]; -goalw Univ.thy [ndepth_def] "ndepth (Abs_Node((%k.0, x))) = 0"; +Goalw [ndepth_def] "ndepth (Abs_Node((%k.0, x))) = 0"; by (EVERY1[stac (Node_K0_I RS Abs_Node_inverse), stac split]); by (rtac Least_equality 1); by (rtac refl 1); by (etac less_zeroE 1); qed "ndepth_K0"; -goal Univ.thy "k < Suc(LEAST x. f(x)=0) --> 0 < nat_case (Suc i) f k"; +Goal "k < Suc(LEAST x. f(x)=0) --> 0 < nat_case (Suc i) f k"; by (nat_ind_tac "k" 1); by (ALLGOALS Simp_tac); by (rtac impI 1); @@ -223,7 +223,7 @@ by (Asm_full_simp_tac 1); val lemma = result(); -goalw Univ.thy [ndepth_def,Push_Node_def] +Goalw [ndepth_def,Push_Node_def] "ndepth (Push_Node i n) = Suc(ndepth(n))"; by (stac (Rep_Node RS Node_Push_I RS Abs_Node_inverse) 1); by (cut_facts_tac [rewrite_rule [Node_def] Rep_Node] 1); @@ -240,23 +240,23 @@ (*** ntrunc applied to the various node sets ***) -goalw Univ.thy [ntrunc_def] "ntrunc 0 M = {}"; +Goalw [ntrunc_def] "ntrunc 0 M = {}"; by (Blast_tac 1); qed "ntrunc_0"; -goalw Univ.thy [Atom_def,ntrunc_def] "ntrunc (Suc k) (Atom a) = Atom(a)"; +Goalw [Atom_def,ntrunc_def] "ntrunc (Suc k) (Atom a) = Atom(a)"; by (fast_tac (claset() addss (simpset() addsimps [ndepth_K0])) 1); qed "ntrunc_Atom"; -goalw Univ.thy [Leaf_def,o_def] "ntrunc (Suc k) (Leaf a) = Leaf(a)"; +Goalw [Leaf_def,o_def] "ntrunc (Suc k) (Leaf a) = Leaf(a)"; by (rtac ntrunc_Atom 1); qed "ntrunc_Leaf"; -goalw Univ.thy [Numb_def,o_def] "ntrunc (Suc k) (Numb i) = Numb(i)"; +Goalw [Numb_def,o_def] "ntrunc (Suc k) (Numb i) = Numb(i)"; by (rtac ntrunc_Atom 1); qed "ntrunc_Numb"; -goalw Univ.thy [Scons_def,ntrunc_def] +Goalw [Scons_def,ntrunc_def] "ntrunc (Suc k) (M$N) = ntrunc k M $ ntrunc k N"; by (safe_tac (claset() addSIs [imageI])); by (REPEAT (stac ndepth_Push_Node 3 THEN etac Suc_mono 3)); @@ -270,24 +270,24 @@ (** Injection nodes **) -goalw Univ.thy [In0_def] "ntrunc (Suc 0) (In0 M) = {}"; +Goalw [In0_def] "ntrunc (Suc 0) (In0 M) = {}"; by (Simp_tac 1); by (rewtac Scons_def); by (Blast_tac 1); qed "ntrunc_one_In0"; -goalw Univ.thy [In0_def] +Goalw [In0_def] "ntrunc (Suc (Suc k)) (In0 M) = In0 (ntrunc (Suc k) M)"; by (Simp_tac 1); qed "ntrunc_In0"; -goalw Univ.thy [In1_def] "ntrunc (Suc 0) (In1 M) = {}"; +Goalw [In1_def] "ntrunc (Suc 0) (In1 M) = {}"; by (Simp_tac 1); by (rewtac Scons_def); by (Blast_tac 1); qed "ntrunc_one_In1"; -goalw Univ.thy [In1_def] +Goalw [In1_def] "ntrunc (Suc (Suc k)) (In1 M) = In1 (ntrunc (Suc k) M)"; by (Simp_tac 1); qed "ntrunc_In1"; @@ -297,7 +297,7 @@ (*** Cartesian Product ***) -goalw Univ.thy [uprod_def] "!!M N. [| M:A; N:B |] ==> (M$N) : A<*>B"; +Goalw [uprod_def] "!!M N. [| M:A; N:B |] ==> (M$N) : A<*>B"; by (REPEAT (ares_tac [singletonI,UN_I] 1)); qed "uprodI"; @@ -322,11 +322,11 @@ (*** Disjoint Sum ***) -goalw Univ.thy [usum_def] "!!M. M:A ==> In0(M) : A<+>B"; +Goalw [usum_def] "!!M. M:A ==> In0(M) : A<+>B"; by (Blast_tac 1); qed "usum_In0I"; -goalw Univ.thy [usum_def] "!!N. N:B ==> In1(N) : A<+>B"; +Goalw [usum_def] "!!N. N:B ==> In1(N) : A<+>B"; by (Blast_tac 1); qed "usum_In1I"; @@ -343,7 +343,7 @@ (** Injection **) -goalw Univ.thy [In0_def,In1_def] "In0(M) ~= In1(N)"; +Goalw [In0_def,In1_def] "In0(M) ~= In1(N)"; by (rtac notI 1); by (etac (Scons_inject1 RS Numb_inject RS Zero_neq_Suc) 1); qed "In0_not_In1"; @@ -360,28 +360,28 @@ by (rtac (major RS Scons_inject2) 1); qed "In1_inject"; -goal Univ.thy "(In0 M = In0 N) = (M=N)"; +Goal "(In0 M = In0 N) = (M=N)"; by (blast_tac (claset() addSDs [In0_inject]) 1); qed "In0_eq"; -goal Univ.thy "(In1 M = In1 N) = (M=N)"; +Goal "(In1 M = In1 N) = (M=N)"; by (blast_tac (claset() addSDs [In1_inject]) 1); qed "In1_eq"; AddIffs [In0_eq, In1_eq]; -goalw Univ.thy [inj_def] "inj In0"; +Goalw [inj_def] "inj In0"; by (Blast_tac 1); qed "inj_In0"; -goalw Univ.thy [inj_def] "inj In1"; +Goalw [inj_def] "inj In1"; by (Blast_tac 1); qed "inj_In1"; (*** proving equality of sets and functions using ntrunc ***) -goalw Univ.thy [ntrunc_def] "ntrunc k M <= M"; +Goalw [ntrunc_def] "ntrunc k M <= M"; by (Blast_tac 1); qed "ntrunc_subsetI"; @@ -408,38 +408,38 @@ (*** Monotonicity ***) -goalw Univ.thy [uprod_def] "!!A B. [| A<=A'; B<=B' |] ==> A<*>B <= A'<*>B'"; +Goalw [uprod_def] "!!A B. [| A<=A'; B<=B' |] ==> A<*>B <= A'<*>B'"; by (Blast_tac 1); qed "uprod_mono"; -goalw Univ.thy [usum_def] "!!A B. [| A<=A'; B<=B' |] ==> A<+>B <= A'<+>B'"; +Goalw [usum_def] "!!A B. [| A<=A'; B<=B' |] ==> A<+>B <= A'<+>B'"; by (Blast_tac 1); qed "usum_mono"; -goalw Univ.thy [Scons_def] "!!M N. [| M<=M'; N<=N' |] ==> M$N <= M'$N'"; +Goalw [Scons_def] "!!M N. [| M<=M'; N<=N' |] ==> M$N <= M'$N'"; by (Blast_tac 1); qed "Scons_mono"; -goalw Univ.thy [In0_def] "!!M N. M<=N ==> In0(M) <= In0(N)"; +Goalw [In0_def] "!!M N. M<=N ==> In0(M) <= In0(N)"; by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); qed "In0_mono"; -goalw Univ.thy [In1_def] "!!M N. M<=N ==> In1(M) <= In1(N)"; +Goalw [In1_def] "!!M N. M<=N ==> In1(M) <= In1(N)"; by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); qed "In1_mono"; (*** Split and Case ***) -goalw Univ.thy [Split_def] "Split c (M$N) = c M N"; +Goalw [Split_def] "Split c (M$N) = c M N"; by (Blast_tac 1); qed "Split"; -goalw Univ.thy [Case_def] "Case c d (In0 M) = c(M)"; +Goalw [Case_def] "Case c d (In0 M) = c(M)"; by (Blast_tac 1); qed "Case_In0"; -goalw Univ.thy [Case_def] "Case c d (In1 N) = d(N)"; +Goalw [Case_def] "Case c d (In1 N) = d(N)"; by (Blast_tac 1); qed "Case_In1"; @@ -448,30 +448,30 @@ (**** UN x. B(x) rules ****) -goalw Univ.thy [ntrunc_def] "ntrunc k (UN x. f(x)) = (UN x. ntrunc k (f x))"; +Goalw [ntrunc_def] "ntrunc k (UN x. f(x)) = (UN x. ntrunc k (f x))"; by (Blast_tac 1); qed "ntrunc_UN1"; -goalw Univ.thy [Scons_def] "(UN x. f(x)) $ M = (UN x. f(x) $ M)"; +Goalw [Scons_def] "(UN x. f(x)) $ M = (UN x. f(x) $ M)"; by (Blast_tac 1); qed "Scons_UN1_x"; -goalw Univ.thy [Scons_def] "M $ (UN x. f(x)) = (UN x. M $ f(x))"; +Goalw [Scons_def] "M $ (UN x. f(x)) = (UN x. M $ f(x))"; by (Blast_tac 1); qed "Scons_UN1_y"; -goalw Univ.thy [In0_def] "In0(UN x. f(x)) = (UN x. In0(f(x)))"; +Goalw [In0_def] "In0(UN x. f(x)) = (UN x. In0(f(x)))"; by (rtac Scons_UN1_y 1); qed "In0_UN1"; -goalw Univ.thy [In1_def] "In1(UN x. f(x)) = (UN x. In1(f(x)))"; +Goalw [In1_def] "In1(UN x. f(x)) = (UN x. In1(f(x)))"; by (rtac Scons_UN1_y 1); qed "In1_UN1"; (*** Equality : the diagonal relation ***) -goalw Univ.thy [diag_def] "!!a A. [| a=b; a:A |] ==> (a,b) : diag(A)"; +Goalw [diag_def] "!!a A. [| a=b; a:A |] ==> (a,b) : diag(A)"; by (Blast_tac 1); qed "diag_eqI"; @@ -488,7 +488,7 @@ (*** Equality for Cartesian Product ***) -goalw Univ.thy [dprod_def] +Goalw [dprod_def] "!!r s. [| (M,M'):r; (N,N'):s |] ==> (M$N, M'$N') : r<**>s"; by (Blast_tac 1); qed "dprodI"; @@ -506,11 +506,11 @@ (*** Equality for Disjoint Sum ***) -goalw Univ.thy [dsum_def] "!!r. (M,M'):r ==> (In0(M), In0(M')) : r<++>s"; +Goalw [dsum_def] "!!r. (M,M'):r ==> (In0(M), In0(M')) : r<++>s"; by (Blast_tac 1); qed "dsum_In0I"; -goalw Univ.thy [dsum_def] "!!r. (N,N'):s ==> (In1(N), In1(N')) : r<++>s"; +Goalw [dsum_def] "!!r. (N,N'):s ==> (In1(N), In1(N')) : r<++>s"; by (Blast_tac 1); qed "dsum_In1I"; @@ -531,36 +531,36 @@ (*** Monotonicity ***) -goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<**>s <= r'<**>s'"; +Goal "!!r s. [| r<=r'; s<=s' |] ==> r<**>s <= r'<**>s'"; by (Blast_tac 1); qed "dprod_mono"; -goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<++>s <= r'<++>s'"; +Goal "!!r s. [| r<=r'; s<=s' |] ==> r<++>s <= r'<++>s'"; by (Blast_tac 1); qed "dsum_mono"; (*** Bounding theorems ***) -goal Univ.thy "diag(A) <= A Times A"; +Goal "diag(A) <= A Times A"; by (Blast_tac 1); qed "diag_subset_Sigma"; -goal Univ.thy "((A Times B) <**> (C Times D)) <= (A<*>C) Times (B<*>D)"; +Goal "((A Times B) <**> (C Times D)) <= (A<*>C) Times (B<*>D)"; by (Blast_tac 1); qed "dprod_Sigma"; val dprod_subset_Sigma = [dprod_mono, dprod_Sigma] MRS subset_trans |>standard; (*Dependent version*) -goal Univ.thy +Goal "(Sigma A B <**> Sigma C D) <= Sigma (A<*>C) (Split(%x y. B(x)<*>D(y)))"; by Safe_tac; by (stac Split 1); by (Blast_tac 1); qed "dprod_subset_Sigma2"; -goal Univ.thy "(A Times B <++> C Times D) <= (A<+>C) Times (B<+>D)"; +Goal "(A Times B <++> C Times D) <= (A<+>C) Times (B<+>D)"; by (Blast_tac 1); qed "dsum_Sigma"; @@ -569,15 +569,15 @@ (*** Domain ***) -goal Univ.thy "fst `` diag(A) = A"; +Goal "fst `` diag(A) = A"; by Auto_tac; qed "fst_image_diag"; -goal Univ.thy "fst `` (r<**>s) = (fst``r) <*> (fst``s)"; +Goal "fst `` (r<**>s) = (fst``r) <*> (fst``s)"; by Auto_tac; qed "fst_image_dprod"; -goal Univ.thy "fst `` (r<++>s) = (fst``r) <+> (fst``s)"; +Goal "fst `` (r<++>s) = (fst``r) <+> (fst``s)"; by Auto_tac; qed "fst_image_dsum"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Update.ML --- a/src/HOL/Update.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Update.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,7 +8,7 @@ open Update; -goalw thy [update_def] "(f[x:=y] = f) = (f x = y)"; +Goalw [update_def] "(f[x:=y] = f) = (f x = y)"; by Safe_tac; by (etac subst 1); by (rtac ext 2); @@ -20,7 +20,7 @@ AddIffs [refl RS update_idem]; -goal thy "(f[x:=y])z = (if x=z then y else f z)"; +Goal "(f[x:=y])z = (if x=z then y else f z)"; by (simp_tac (simpset() addsimps [update_def]) 1); qed "update_apply"; Addsimps [update_apply]; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/Vimage.ML --- a/src/HOL/Vimage.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/Vimage.ML Mon Jun 22 17:26:46 1998 +0200 @@ -36,60 +36,60 @@ (*** Simple equalities ***) -goal thy "(%x. x) -`` B = B"; +Goal "(%x. x) -`` B = B"; by (Blast_tac 1); qed "ident_vimage"; Addsimps [ident_vimage]; -goal thy "f-``{} = {}"; +Goal "f-``{} = {}"; by (Blast_tac 1); qed "vimage_empty"; -goal thy "f-``(Compl A) = Compl (f-``A)"; +Goal "f-``(Compl A) = Compl (f-``A)"; by (Blast_tac 1); qed "vimage_Compl"; -goal thy "f-``(A Un B) = (f-``A) Un (f-``B)"; +Goal "f-``(A Un B) = (f-``A) Un (f-``B)"; by (Blast_tac 1); qed "vimage_Un"; -goal thy "f-``(UN x:A. B x) = (UN x:A. f -`` B x)"; +Goal "f-``(UN x:A. B x) = (UN x:A. f -`` B x)"; by (Blast_tac 1); qed "vimage_UN"; Addsimps [vimage_empty, vimage_Un]; (*NOT suitable for rewriting because of the recurrence of {a}*) -goal thy "f-``(insert a B) = (f-``{a}) Un (f-``B)"; +Goal "f-``(insert a B) = (f-``{a}) Un (f-``B)"; by (Blast_tac 1); qed "vimage_insert"; -goal thy "f-``(A Int B) = (f-``A) Int (f-``B)"; +Goal "f-``(A Int B) = (f-``A) Int (f-``B)"; by (Blast_tac 1); qed "vimage_Int"; -goal thy "f-``(A-B) = (f-``A) - (f-``B)"; +Goal "f-``(A-B) = (f-``A) - (f-``B)"; by (Blast_tac 1); qed "vimage_Diff"; -goal thy "f-``UNIV = UNIV"; +Goal "f-``UNIV = UNIV"; by (Blast_tac 1); qed "vimage_UNIV"; Addsimps [vimage_UNIV]; -goal thy "(UN x:A. f -`` B x) = f -`` (UN x:A. B x)"; +Goal "(UN x:A. f -`` B x) = f -`` (UN x:A. B x)"; by (Blast_tac 1); qed "UN_vimage"; Addsimps [UN_vimage]; (*NOT suitable for rewriting*) -goal thy "f-``B = (UN y: B. f-``{y})"; +Goal "f-``B = (UN y: B. f-``{y})"; by (Blast_tac 1); qed "vimage_eq_UN"; (** monotonicity **) -goal thy "!!f. A<=B ==> f-``A <= f-``B"; +Goal "!!f. A<=B ==> f-``A <= f-``B"; by (Blast_tac 1); qed "vimage_mono"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/W0/I.ML --- a/src/HOL/W0/I.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/W0/I.ML Mon Jun 22 17:26:46 1998 +0200 @@ -8,7 +8,7 @@ open I; -goal thy +Goal "! a m s s' t n. \ \ (new_tv m a & new_tv m s) --> I e a m s = Ok(s',t,n) --> \ \ ( ? r. W e ($ s a) m = Ok(r, $ s' t, n) & s' = ($ r o s) )"; @@ -123,7 +123,7 @@ (*** We actually want the corollary -goal I.thy +Goal "I e [] m id_subst = Ok(s,t,n) --> W e [] m = Ok(s, $s t, n)"; by (cut_facts_tac [(read_instantiate[("x","id_subst")] (read_instantiate[("x","[]")](thm RS spec) @@ -141,7 +141,7 @@ Addsimps [split_paired_Ex]; -goal I.thy "!a m s. \ +Goal "!a m s. \ \ new_tv m a & new_tv m s --> I e a m s = Fail --> W e ($s a) m = Fail"; by (expr.induct_tac "e" 1); by (simp_tac (simpset() addsimps [app_subst_list]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/W0/Maybe.ML --- a/src/HOL/W0/Maybe.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/W0/Maybe.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,25 +6,25 @@ (* constructor laws for bind *) -goalw thy [bind_def] "(Ok s) bind f = (f s)"; +Goalw [bind_def] "(Ok s) bind f = (f s)"; by (Simp_tac 1); qed "bind_Ok"; -goalw thy [bind_def] "Fail bind f = Fail"; +Goalw [bind_def] "Fail bind f = Fail"; by (Simp_tac 1); qed "bind_Fail"; Addsimps [bind_Ok,bind_Fail]; (* case splitting of bind *) -goal thy +Goal "P(res bind f) = ((res = Fail --> P Fail) & (!s. res = Ok s --> P(f s)))"; by (induct_tac "res" 1); by (fast_tac (HOL_cs addss simpset()) 1); by (Asm_simp_tac 1); qed "split_bind"; -goal Maybe.thy +Goal "((m bind f) = Fail) = ((m=Fail) | (? p. m = Ok p & f p = Fail))"; by (simp_tac (simpset() addsplits [split_bind]) 1); qed "bind_eq_Fail"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/W0/MiniML.ML --- a/src/HOL/W0/MiniML.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/W0/MiniML.ML Mon Jun 22 17:26:46 1998 +0200 @@ -11,7 +11,7 @@ (* has_type is closed w.r.t. substitution *) -goal MiniML.thy "!!a e t. a |- e :: t ==> $s a |- e :: $s t"; +Goal "!!a e t. a |- e :: t ==> $s a |- e :: $s t"; by (etac has_type.induct 1); (* case VarI *) by (asm_full_simp_tac (simpset() addsimps [app_subst_list]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/W0/Type.ML --- a/src/HOL/W0/Type.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/W0/Type.ML Mon Jun 22 17:26:46 1998 +0200 @@ -7,14 +7,14 @@ Addsimps [mgu_eq,mgu_mg,mgu_free]; (* mgu does not introduce new type variables *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "!! n. [|mgu t1 t2 = Ok u; new_tv n t1; new_tv n t2|] ==> \ \ new_tv n u"; by (fast_tac (set_cs addDs [mgu_free]) 1); qed "mgu_new"; (* application of id_subst does not change type expression *) -goalw thy [id_subst_def] +Goalw [id_subst_def] "$ id_subst = (%t::typ. t)"; by (rtac ext 1); by (typ.induct_tac "t" 1); @@ -23,7 +23,7 @@ Addsimps [app_subst_id_te]; (* application of id_subst does not change list of type expressions *) -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ id_subst = (%ts::typ list. ts)"; by (rtac ext 1); by (list.induct_tac "ts" 1); @@ -31,31 +31,31 @@ qed "app_subst_id_tel"; Addsimps [app_subst_id_tel]; -goalw Type.thy [id_subst_def,o_def] "$s o id_subst = s"; +Goalw [id_subst_def,o_def] "$s o id_subst = s"; by (rtac ext 1); by (Simp_tac 1); qed "o_id_subst"; Addsimps[o_id_subst]; -goalw thy [dom_def,id_subst_def,empty_def] +Goalw [dom_def,id_subst_def,empty_def] "dom id_subst = {}"; by (Simp_tac 1); qed "dom_id_subst"; Addsimps [dom_id_subst]; -goalw thy [cod_def] +Goalw [cod_def] "cod id_subst = {}"; by (Simp_tac 1); qed "cod_id_subst"; Addsimps [cod_id_subst]; -goalw thy [free_tv_subst] +Goalw [free_tv_subst] "free_tv id_subst = {}"; by (Simp_tac 1); qed "free_tv_id_subst"; Addsimps [free_tv_id_subst]; -goalw thy [dom_def,cod_def,UNION_def,Bex_def] +Goalw [dom_def,cod_def,UNION_def,Bex_def] "!!v. [| v : free_tv(s n); v~=n |] ==> v : cod s"; by (Simp_tac 1); by (safe_tac (empty_cs addSIs [exI,conjI,notI])); @@ -67,13 +67,13 @@ (* composition of substitutions *) -goal thy +Goal "$ g ($ f t::typ) = $ (%x. $ g (f x) ) t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "subst_comp_te"; -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ g ($ f ts::typ list) = $ (%x. $ g (f x)) ts"; by (list.induct_tac "ts" 1); (* case [] *) @@ -84,12 +84,12 @@ (* constructor laws for app_subst *) -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ s [] = []"; by (Simp_tac 1); qed "app_subst_Nil"; -goalw thy [app_subst_list] +Goalw [app_subst_list] "$ s (t#ts) = ($ s t)#($ s ts)"; by (Simp_tac 1); qed "app_subst_Cons"; @@ -97,35 +97,35 @@ Addsimps [app_subst_Nil,app_subst_Cons]; (* constructor laws for new_tv *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n (TVar m) = (m t2) = (new_tv n t1 & new_tv n t2)"; by (fast_tac (HOL_cs addss simpset()) 1); qed "new_tv_Fun"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n []"; by (Simp_tac 1); qed "new_tv_Nil"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n (t#ts) = (new_tv n t & new_tv n ts)"; by (fast_tac (HOL_cs addss simpset()) 1); qed "new_tv_Cons"; Addsimps [new_tv_TVar,new_tv_Fun,new_tv_Nil,new_tv_Cons]; -goalw Type.thy [id_subst_def,new_tv_def,free_tv_subst,dom_def,cod_def] +Goalw [id_subst_def,new_tv_def,free_tv_subst,dom_def,cod_def] "new_tv n id_subst"; by (Simp_tac 1); qed "new_tv_id_subst"; Addsimps[new_tv_id_subst]; -goalw thy [new_tv_def] +Goalw [new_tv_def] "new_tv n s = ((!m. n <= m --> (s m = TVar m)) & \ \ (! l. l < n --> new_tv n (s l) ))"; by (safe_tac HOL_cs ); @@ -146,7 +146,7 @@ by (fast_tac (HOL_cs addSIs [less_or_eq_imp_le]) 1); qed "new_tv_subst"; -goal thy +Goal "new_tv n = list_all (new_tv n)"; by (rtac ext 1); by (list.induct_tac "x" 1); @@ -154,14 +154,14 @@ qed "new_tv_list"; (* substitution affects only variables occurring freely *) -goal thy +Goal "new_tv n (t::typ) --> $(%x. if x=n then t' else s x) t = $s t"; by (typ.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "subst_te_new_tv"; Addsimps [subst_te_new_tv]; -goal thy +Goal "new_tv n (ts::typ list) --> $(%x. if x=n then t else s x) ts = $s ts"; by (list.induct_tac "ts" 1); by (ALLGOALS Asm_full_simp_tac); @@ -169,7 +169,7 @@ Addsimps [subst_tel_new_tv]; (* all greater variables are also new *) -goal thy +Goal "n<=m --> new_tv n (t::typ) --> new_tv m t"; by (typ.induct_tac "t" 1); (* case TVar n *) @@ -179,7 +179,7 @@ qed_spec_mp "new_tv_le"; Addsimps [lessI RS less_imp_le RS new_tv_le]; -goal thy +Goal "n<=m --> new_tv n (ts::typ list) --> new_tv m ts"; by (list.induct_tac "ts" 1); (* case [] *) @@ -189,7 +189,7 @@ qed_spec_mp "new_tv_list_le"; Addsimps [lessI RS less_imp_le RS new_tv_list_le]; -goal thy +Goal "!! n. [| n<=m; new_tv n (s::subst) |] ==> new_tv m s"; by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); by (rtac conjI 1); @@ -202,19 +202,19 @@ Addsimps [lessI RS less_imp_le RS new_tv_subst_le]; (* new_tv property remains if a substitution is applied *) -goal thy +Goal "!!n. [| n new_tv m (s n)"; by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); qed "new_tv_subst_var"; -goal thy +Goal "new_tv n s --> new_tv n (t::typ) --> new_tv n ($ s t)"; by (typ.induct_tac "t" 1); by (ALLGOALS(fast_tac (HOL_cs addss (simpset() addsimps [new_tv_subst])))); qed_spec_mp "new_tv_subst_te"; Addsimps [new_tv_subst_te]; -goal thy +Goal "new_tv n s --> new_tv n (ts::typ list) --> new_tv n ($ s ts)"; by (simp_tac (simpset() addsimps [new_tv_list]) 1); by (list.induct_tac "ts" 1); @@ -223,7 +223,7 @@ Addsimps [new_tv_subst_tel]; (* auxilliary lemma *) -goal thy +Goal "new_tv n ts --> new_tv (Suc n) ((TVar n)#ts)"; by (simp_tac (simpset() addsimps [new_tv_list]) 1); by (list.induct_tac "ts" 1); @@ -232,13 +232,13 @@ (* composition of substitutions preserves new_tv proposition *) -goal thy +Goal "!! n. [| new_tv n (s::subst); new_tv n r |] ==> \ \ new_tv n (($ r) o s)"; by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); qed "new_tv_subst_comp_1"; -goal thy +Goal "!! n. [| new_tv n (s::subst); new_tv n r |] ==> \ \ new_tv n (%v.$ r (s v))"; by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); @@ -247,13 +247,13 @@ Addsimps [new_tv_subst_comp_1,new_tv_subst_comp_2]; (* new type variables do not occur freely in a type structure *) -goalw thy [new_tv_def] +Goalw [new_tv_def] "!!n. new_tv n ts ==> n~:(free_tv ts)"; by (fast_tac (HOL_cs addEs [less_irrefl]) 1); qed "new_tv_not_free_tv"; Addsimps [new_tv_not_free_tv]; -goal thy +Goal "(t::typ) mem ts --> free_tv t <= free_tv ts"; by (list.induct_tac "ts" 1); (* case [] *) @@ -267,7 +267,7 @@ (* if two substitutions yield the same result if applied to a type structure the substitutions coincide on the free type variables occurring in the type structure *) -goal thy +Goal "$ s1 (t::typ) = $ s2 t --> n:(free_tv t) --> s1 n = s2 n"; by (typ.induct_tac "t" 1); (* case TVar n *) @@ -276,7 +276,7 @@ by (fast_tac (HOL_cs addss simpset()) 1); qed_spec_mp "eq_subst_te_eq_free"; -goal thy +Goal "(!n. n:(free_tv t) --> s1 n = s2 n) --> $ s1 (t::typ) = $ s2 t"; by (typ.induct_tac "t" 1); (* case TVar n *) @@ -285,7 +285,7 @@ by (fast_tac (HOL_cs addss simpset()) 1); qed_spec_mp "eq_free_eq_subst_te"; -goal thy +Goal "$s1 (ts::typ list) = $s2 ts --> n:(free_tv ts) --> s1 n = s2 n"; by (list.induct_tac "ts" 1); (* case [] *) @@ -294,7 +294,7 @@ by (fast_tac (HOL_cs addIs [eq_subst_te_eq_free] addss (simpset())) 1); qed_spec_mp "eq_subst_tel_eq_free"; -goal thy +Goal "(!n. n:(free_tv ts) --> s1 n = s2 n) --> $s1 (ts::typ list) = $s2 ts"; by (list.induct_tac "ts" 1); (* case [] *) @@ -304,29 +304,29 @@ qed_spec_mp "eq_free_eq_subst_tel"; (* some useful theorems *) -goalw thy [free_tv_subst] +Goalw [free_tv_subst] "!!v. v : cod s ==> v : free_tv s"; by (fast_tac set_cs 1); qed "codD"; -goalw thy [free_tv_subst,dom_def] +Goalw [free_tv_subst,dom_def] "!! x. x ~: free_tv s ==> s x = TVar x"; by (fast_tac set_cs 1); qed "not_free_impl_id"; -goalw thy [new_tv_def] +Goalw [new_tv_def] "!! n. [| new_tv n t; m:free_tv t |] ==> m $s a |- e :: t"; by (expr.induct_tac "e" 1); (* case Var n *) @@ -41,7 +41,7 @@ (* the resulting type variable is always greater or equal than the given one *) -goal thy +Goal "!a n s t m. W e a n = Ok (s,t,m) --> n<=m"; by (expr.induct_tac "e" 1); (* case Var(n) *) @@ -72,7 +72,7 @@ Addsimps [W_var_ge]; -goal thy +Goal "!! s. Ok (s,t,m) = W e a n ==> n<=m"; by (asm_full_simp_tac (simpset() addsimps [eq_sym_conv]) 1); qed "W_var_geD"; @@ -85,7 +85,7 @@ (* resulting type variable is new *) -goal thy +Goal "!n a s t m. new_tv n a --> W e a n = Ok (s,t,m) --> \ \ new_tv m s & new_tv m t"; by (expr.induct_tac "e" 1); @@ -152,7 +152,7 @@ qed_spec_mp "new_tv_W"; -goal thy +Goal "!n a s t m v. W e a n = Ok (s,t,m) --> \ \ (v:free_tv s | v:free_tv t) --> v v:free_tv a"; by (expr.induct_tac "e" 1); @@ -217,7 +217,7 @@ qed_spec_mp "free_tv_W"; (* Completeness of W w.r.t. has_type *) -goal thy +Goal "!s' a t' n. $s' a |- e :: t' --> new_tv n a --> \ \ (? s t. (? m. W e a n = Ok (s,t,m)) & \ \ (? r. $s' a = $r ($s a) & t' = $r t))"; @@ -361,7 +361,7 @@ by (Fast_tac 1); qed_spec_mp "W_complete_lemma"; -goal W.thy +Goal "!!e. [] |- e :: t' ==> (? s t. (? m. W e [] n = Ok(s,t,m)) & \ \ (? r. t' = $r t))"; by (cut_inst_tac [("a","[]"),("s'","id_subst"),("e","e"),("t'","t'")] diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/WF.ML --- a/src/HOL/WF.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/WF.ML Mon Jun 22 17:26:46 1998 +0200 @@ -78,14 +78,14 @@ by (Blast_tac 1); val lemma1 = result(); -goalw WF.thy [wf_def] +Goalw [wf_def] "!!r. (! Q x. x:Q --> (? z:Q. ! y. (y,z):r --> y~:Q)) ==> wf r"; by (Clarify_tac 1); by (dres_inst_tac [("x", "{x. ~ P x}")] spec 1); by (Blast_tac 1); val lemma2 = result(); -goal WF.thy "wf r = (! Q x. x:Q --> (? z:Q. ! y. (y,z):r --> y~:Q))"; +Goal "wf r = (! Q x. x:Q --> (? z:Q. ! y. (y,z):r --> y~:Q))"; by (blast_tac (claset() addSIs [lemma1, lemma2]) 1); qed "wf_eq_minimal"; @@ -93,7 +93,7 @@ * Wellfoundedness of subsets *---------------------------------------------------------------------------*) -goal thy "!!r. [| wf(r); p<=r |] ==> wf(p)"; +Goal "!!r. [| wf(r); p<=r |] ==> wf(p)"; by (full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1); by (Fast_tac 1); qed "wf_subset"; @@ -102,7 +102,7 @@ * Wellfoundedness of the empty relation. *---------------------------------------------------------------------------*) -goal thy "wf({})"; +Goal "wf({})"; by (simp_tac (simpset() addsimps [wf_def]) 1); qed "wf_empty"; AddSIs [wf_empty]; @@ -111,7 +111,7 @@ * Wellfoundedness of `insert' *---------------------------------------------------------------------------*) -goal WF.thy "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)"; +Goal "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)"; by (rtac iffI 1); by (blast_tac (claset() addEs [wf_trancl RS wf_irrefl] addIs [rtrancl_into_trancl1,wf_subset,impOfSubs rtrancl_mono]) 1); @@ -139,19 +139,19 @@ val acyclicI = prove_goalw WF.thy [acyclic_def] "!!r. !x. (x, x) ~: r^+ ==> acyclic r" (K [atac 1]); -goalw WF.thy [acyclic_def] +Goalw [acyclic_def] "!!r. wf r ==> acyclic r"; by (blast_tac (claset() addEs [wf_trancl RS wf_irrefl]) 1); qed "wf_acyclic"; -goalw WF.thy [acyclic_def] +Goalw [acyclic_def] "acyclic(insert (y,x) r) = (acyclic r & (x,y) ~: r^*)"; by (simp_tac (simpset() addsimps [trancl_insert]) 1); by (blast_tac (claset() addEs [make_elim rtrancl_trans]) 1); qed "acyclic_insert"; AddIffs [acyclic_insert]; -goalw WF.thy [acyclic_def] "acyclic(r^-1) = acyclic r"; +Goalw [acyclic_def] "acyclic(r^-1) = acyclic r"; by (simp_tac (simpset() addsimps [trancl_converse]) 1); qed "acyclic_converse"; @@ -159,18 +159,18 @@ (*This rewrite rule works upon formulae; thus it requires explicit use of H_cong to expose the equality*) -goalw WF.thy [cut_def] +Goalw [cut_def] "(cut f r x = cut g r x) = (!y. (y,x):r --> f(y)=g(y))"; by (simp_tac (HOL_ss addsimps [expand_fun_eq]) 1); qed "cuts_eq"; -goalw WF.thy [cut_def] "!!x. (x,a):r ==> (cut f r a)(x) = f(x)"; +Goalw [cut_def] "!!x. (x,a):r ==> (cut f r a)(x) = f(x)"; by (asm_simp_tac HOL_ss 1); qed "cut_apply"; (*** is_recfun ***) -goalw WF.thy [is_recfun_def,cut_def] +Goalw [is_recfun_def,cut_def] "!!f. [| is_recfun r H a f; ~(b,a):r |] ==> f(b) = arbitrary"; by (etac ssubst 1); by (asm_simp_tac HOL_ss 1); @@ -260,7 +260,7 @@ val th = rewrite_rule[is_recfun_def] (trans_trancl RSN (2,(wf_trancl RS unfold_the_recfun))); -goalw WF.thy [wfrec_def] +Goalw [wfrec_def] "!!r. wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"; by (rtac H_cong 1); by (rtac refl 2); @@ -296,7 +296,7 @@ qed "wfrec"; (*--------------Old proof----------------------------------------------------- -goalw WF.thy [wfrec_def] +Goalw [wfrec_def] "!!r. wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"; by (etac (wf_trancl RS wftrec RS ssubst) 1); by (rtac trans_trancl 1); @@ -317,7 +317,7 @@ (**** TFL variants ****) -goal WF.thy +Goal "!R. wf R --> (!P. (!x. (!y. (y,x):R --> P y) --> P x) --> (!x. P x))"; by (Clarify_tac 1); by (res_inst_tac [("r","R"),("P","P"), ("a","x")] wf_induct 1); @@ -325,13 +325,13 @@ by (Blast_tac 1); qed"tfl_wf_induct"; -goal WF.thy "!f R. (x,a):R --> (cut f R a)(x) = f(x)"; +Goal "!f R. (x,a):R --> (cut f R a)(x) = f(x)"; by (Clarify_tac 1); by (rtac cut_apply 1); by (assume_tac 1); qed"tfl_cut_apply"; -goal WF.thy "!M R f. (f=wfrec R M) --> wf R --> (!x. f x = M (cut f R x) x)"; +Goal "!M R f. (f=wfrec R M) --> wf R --> (!x. f x = M (cut f R x) x)"; by (Clarify_tac 1); by (etac wfrec 1); qed "tfl_wfrec"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/WF_Rel.ML --- a/src/HOL/WF_Rel.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/WF_Rel.ML Mon Jun 22 17:26:46 1998 +0200 @@ -13,17 +13,17 @@ * "Less than" on the natural numbers *---------------------------------------------------------------------------*) -goalw thy [less_than_def] "wf less_than"; +Goalw [less_than_def] "wf less_than"; by (rtac (wf_pred_nat RS wf_trancl) 1); qed "wf_less_than"; AddIffs [wf_less_than]; -goalw thy [less_than_def] "trans less_than"; +Goalw [less_than_def] "trans less_than"; by (rtac trans_trancl 1); qed "trans_less_than"; AddIffs [trans_less_than]; -goalw thy [less_than_def, less_def] "((x,y): less_than) = (x wf(inv_image r (f::'a=>'b))"; +Goal "!!r. wf(r) ==> wf(inv_image r (f::'a=>'b))"; by (full_simp_tac (simpset() addsimps [inv_image_def, wf_eq_minimal]) 1); by (Clarify_tac 1); by (subgoal_tac "? (w::'b). w : {w. ? (x::'a). x: Q & (f x = w)}" 1); @@ -43,7 +43,7 @@ qed "wf_inv_image"; AddSIs [wf_inv_image]; -goalw thy [trans_def,inv_image_def] +Goalw [trans_def,inv_image_def] "!!r. trans r ==> trans (inv_image r f)"; by (Simp_tac 1); by (Blast_tac 1); @@ -54,7 +54,7 @@ * All measures are wellfounded. *---------------------------------------------------------------------------*) -goalw thy [measure_def] "wf (measure f)"; +Goalw [measure_def] "wf (measure f)"; by (rtac (wf_less_than RS wf_inv_image) 1); qed "wf_measure"; AddIffs [wf_measure]; @@ -82,7 +82,7 @@ (*--------------------------------------------------------------------------- * Transitivity of WF combinators. *---------------------------------------------------------------------------*) -goalw thy [trans_def, lex_prod_def] +Goalw [trans_def, lex_prod_def] "!!R1 R2. [| trans R1; trans R2 |] ==> trans (R1 ** R2)"; by (Simp_tac 1); by (Blast_tac 1); @@ -93,14 +93,14 @@ (*--------------------------------------------------------------------------- * Wellfoundedness of proper subset on finite sets. *---------------------------------------------------------------------------*) -goalw thy [finite_psubset_def] "wf(finite_psubset)"; +Goalw [finite_psubset_def] "wf(finite_psubset)"; by (rtac (wf_measure RS wf_subset) 1); by (simp_tac (simpset() addsimps [measure_def, inv_image_def, less_than_def, symmetric less_def])1); by (fast_tac (claset() addSIs [psubset_card]) 1); qed "wf_finite_psubset"; -goalw thy [finite_psubset_def, trans_def] "trans finite_psubset"; +Goalw [finite_psubset_def, trans_def] "trans finite_psubset"; by (simp_tac (simpset() addsimps [psubset_def]) 1); by (Blast_tac 1); qed "trans_finite_psubset"; @@ -110,7 +110,7 @@ * Cannot go into WF because it needs Finite *---------------------------------------------------------------------------*) -goal thy "!!r. finite r ==> acyclic r --> wf r"; +Goal "!!r. finite r ==> acyclic r --> wf r"; by (etac finite_induct 1); by (Blast_tac 1); by (split_all_tac 1); @@ -122,7 +122,7 @@ etac (finite_converse RS iffD2 RS finite_acyclic_wf) 1, etac (acyclic_converse RS iffD2) 1]); -goal thy "!!r. finite r ==> wf r = acyclic r"; +Goal "!!r. finite r ==> wf r = acyclic r"; by (blast_tac (claset() addIs [finite_acyclic_wf,wf_acyclic]) 1); qed "wf_iff_acyclic_if_finite"; @@ -131,7 +131,7 @@ * A relation is wellfounded iff it has no infinite descending chain *---------------------------------------------------------------------------*) -goalw thy [wf_eq_minimal RS eq_reflection] +Goalw [wf_eq_minimal RS eq_reflection] "wf r = (~(? f. !i. (f(Suc i),f i) : r))"; by (rtac iffI 1); by (rtac notI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/equalities.ML --- a/src/HOL/equalities.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/equalities.ML Mon Jun 22 17:26:46 1998 +0200 @@ -12,26 +12,26 @@ section "{}"; -goal thy "{x. False} = {}"; +Goal "{x. False} = {}"; by (Blast_tac 1); qed "Collect_False_empty"; Addsimps [Collect_False_empty]; -goal thy "(A <= {}) = (A = {})"; +Goal "(A <= {}) = (A = {})"; by (Blast_tac 1); qed "subset_empty"; Addsimps [subset_empty]; -goalw thy [psubset_def] "~ (A < {})"; +Goalw [psubset_def] "~ (A < {})"; by (Blast_tac 1); qed "not_psubset_empty"; AddIffs [not_psubset_empty]; -goal thy "{x. P x | Q x} = {x. P x} Un {x. Q x}"; +Goal "{x. P x | Q x} = {x. P x} Un {x. Q x}"; by (Blast_tac 1); qed "Collect_disj_eq"; -goal thy "{x. P x & Q x} = {x. P x} Int {x. Q x}"; +Goal "{x. P x & Q x} = {x. P x} Int {x. Q x}"; by (Blast_tac 1); qed "Collect_conj_eq"; @@ -39,11 +39,11 @@ section "insert"; (*NOT SUITABLE FOR REWRITING since {a} == insert a {}*) -goal thy "insert a A = {a} Un A"; +Goal "insert a A = {a} Un A"; by (Blast_tac 1); qed "insert_is_Un"; -goal thy "insert a A ~= {}"; +Goal "insert a A ~= {}"; by (blast_tac (claset() addEs [equalityCE]) 1); qed"insert_not_empty"; Addsimps[insert_not_empty]; @@ -51,33 +51,33 @@ bind_thm("empty_not_insert",insert_not_empty RS not_sym); Addsimps[empty_not_insert]; -goal thy "!!a. a:A ==> insert a A = A"; +Goal "!!a. a:A ==> insert a A = A"; by (Blast_tac 1); qed "insert_absorb"; (* Addsimps [insert_absorb] causes recursive (ie quadtratic) calls in case of nested inserts! *) -goal thy "insert x (insert x A) = insert x A"; +Goal "insert x (insert x A) = insert x A"; by (Blast_tac 1); qed "insert_absorb2"; Addsimps [insert_absorb2]; -goal thy "insert x (insert y A) = insert y (insert x A)"; +Goal "insert x (insert y A) = insert y (insert x A)"; by (Blast_tac 1); qed "insert_commute"; -goal thy "(insert x A <= B) = (x:B & A <= B)"; +Goal "(insert x A <= B) = (x:B & A <= B)"; by (Blast_tac 1); qed "insert_subset"; Addsimps[insert_subset]; -goal thy "!!a. insert a A ~= insert a B ==> A ~= B"; +Goal "!!a. insert a A ~= insert a B ==> A ~= B"; by (Blast_tac 1); qed "insert_lim"; (* use new B rather than (A-{a}) to avoid infinite unfolding *) -goal thy "!!a. a:A ==> ? B. A = insert a B & a ~: B"; +Goal "!!a. a:A ==> ? B. A = insert a B & a ~: B"; by (res_inst_tac [("x","A-{a}")] exI 1); by (Blast_tac 1); qed "mk_disjoint_insert"; @@ -85,46 +85,46 @@ bind_thm ("insert_Collect", prove_goal thy "insert a (Collect P) = {u. u ~= a --> P u}" (K [Auto_tac])); -goal thy +Goal "!!A. A~={} ==> (UN x:A. insert a (B x)) = insert a (UN x:A. B x)"; by (Blast_tac 1); qed "UN_insert_distrib"; section "``"; -goal thy "f``{} = {}"; +Goal "f``{} = {}"; by (Blast_tac 1); qed "image_empty"; Addsimps[image_empty]; -goal thy "f``insert a B = insert (f a) (f``B)"; +Goal "f``insert a B = insert (f a) (f``B)"; by (Blast_tac 1); qed "image_insert"; Addsimps[image_insert]; -goal thy "(f `` (UNION A B)) = (UN x:A.(f `` (B x)))"; +Goal "(f `` (UNION A B)) = (UN x:A.(f `` (B x)))"; by (Blast_tac 1); qed "image_UNION"; -goal thy "(%x. x) `` Y = Y"; +Goal "(%x. x) `` Y = Y"; by (Blast_tac 1); qed "image_id"; -goal thy "f``(g``A) = (%x. f (g x)) `` A"; +Goal "f``(g``A) = (%x. f (g x)) `` A"; by (Blast_tac 1); qed "image_image"; -goal thy "!!x. x:A ==> insert (f x) (f``A) = f``A"; +Goal "!!x. x:A ==> insert (f x) (f``A) = f``A"; by (Blast_tac 1); qed "insert_image"; Addsimps [insert_image]; -goal thy "(f``A = {}) = (A = {})"; +Goal "(f``A = {}) = (A = {})"; by (blast_tac (claset() addSEs [equalityCE]) 1); qed "image_is_empty"; AddIffs [image_is_empty]; -goalw thy [image_def] +Goalw [image_def] "(%x. if P x then f x else g x) `` S \ \ = (f `` (S Int {x. P x})) Un (g `` (S Int {x. ~(P x)}))"; by (Simp_tac 1); @@ -140,218 +140,218 @@ section "Int"; -goal thy "A Int A = A"; +Goal "A Int A = A"; by (Blast_tac 1); qed "Int_absorb"; Addsimps[Int_absorb]; -goal thy " A Int (A Int B) = A Int B"; +Goal " A Int (A Int B) = A Int B"; by (Blast_tac 1); qed "Int_left_absorb"; -goal thy "A Int B = B Int A"; +Goal "A Int B = B Int A"; by (Blast_tac 1); qed "Int_commute"; -goal thy "A Int (B Int C) = B Int (A Int C)"; +Goal "A Int (B Int C) = B Int (A Int C)"; by (Blast_tac 1); qed "Int_left_commute"; -goal thy "(A Int B) Int C = A Int (B Int C)"; +Goal "(A Int B) Int C = A Int (B Int C)"; by (Blast_tac 1); qed "Int_assoc"; (*Intersection is an AC-operator*) val Int_ac = [Int_assoc, Int_left_absorb, Int_commute, Int_left_commute]; -goal thy "!!A B. B<=A ==> A Int B = B"; +Goal "!!A B. B<=A ==> A Int B = B"; by (Blast_tac 1); qed "Int_absorb1"; -goal thy "!!A B. A<=B ==> A Int B = A"; +Goal "!!A B. A<=B ==> A Int B = A"; by (Blast_tac 1); qed "Int_absorb2"; -goal thy "{} Int B = {}"; +Goal "{} Int B = {}"; by (Blast_tac 1); qed "Int_empty_left"; Addsimps[Int_empty_left]; -goal thy "A Int {} = {}"; +Goal "A Int {} = {}"; by (Blast_tac 1); qed "Int_empty_right"; Addsimps[Int_empty_right]; -goal thy "(A Int B = {}) = (A <= Compl B)"; +Goal "(A Int B = {}) = (A <= Compl B)"; by (blast_tac (claset() addSEs [equalityCE]) 1); qed "disjoint_eq_subset_Compl"; -goal thy "UNIV Int B = B"; +Goal "UNIV Int B = B"; by (Blast_tac 1); qed "Int_UNIV_left"; Addsimps[Int_UNIV_left]; -goal thy "A Int UNIV = A"; +Goal "A Int UNIV = A"; by (Blast_tac 1); qed "Int_UNIV_right"; Addsimps[Int_UNIV_right]; -goal thy "A Int B = Inter{A,B}"; +Goal "A Int B = Inter{A,B}"; by (Blast_tac 1); qed "Int_eq_Inter"; -goal thy "A Int (B Un C) = (A Int B) Un (A Int C)"; +Goal "A Int (B Un C) = (A Int B) Un (A Int C)"; by (Blast_tac 1); qed "Int_Un_distrib"; -goal thy "(B Un C) Int A = (B Int A) Un (C Int A)"; +Goal "(B Un C) Int A = (B Int A) Un (C Int A)"; by (Blast_tac 1); qed "Int_Un_distrib2"; -goal thy "(A<=B) = (A Int B = A)"; +Goal "(A<=B) = (A Int B = A)"; by (blast_tac (claset() addSEs [equalityCE]) 1); qed "subset_Int_eq"; -goal thy "(A Int B = UNIV) = (A = UNIV & B = UNIV)"; +Goal "(A Int B = UNIV) = (A = UNIV & B = UNIV)"; by (blast_tac (claset() addEs [equalityCE]) 1); qed "Int_UNIV"; Addsimps[Int_UNIV]; section "Un"; -goal thy "A Un A = A"; +Goal "A Un A = A"; by (Blast_tac 1); qed "Un_absorb"; Addsimps[Un_absorb]; -goal thy " A Un (A Un B) = A Un B"; +Goal " A Un (A Un B) = A Un B"; by (Blast_tac 1); qed "Un_left_absorb"; -goal thy "A Un B = B Un A"; +Goal "A Un B = B Un A"; by (Blast_tac 1); qed "Un_commute"; -goal thy "A Un (B Un C) = B Un (A Un C)"; +Goal "A Un (B Un C) = B Un (A Un C)"; by (Blast_tac 1); qed "Un_left_commute"; -goal thy "(A Un B) Un C = A Un (B Un C)"; +Goal "(A Un B) Un C = A Un (B Un C)"; by (Blast_tac 1); qed "Un_assoc"; (*Union is an AC-operator*) val Un_ac = [Un_assoc, Un_left_absorb, Un_commute, Un_left_commute]; -goal thy "!!A B. A<=B ==> A Un B = B"; +Goal "!!A B. A<=B ==> A Un B = B"; by (Blast_tac 1); qed "Un_absorb1"; -goal thy "!!A B. B<=A ==> A Un B = A"; +Goal "!!A B. B<=A ==> A Un B = A"; by (Blast_tac 1); qed "Un_absorb2"; -goal thy "{} Un B = B"; +Goal "{} Un B = B"; by (Blast_tac 1); qed "Un_empty_left"; Addsimps[Un_empty_left]; -goal thy "A Un {} = A"; +Goal "A Un {} = A"; by (Blast_tac 1); qed "Un_empty_right"; Addsimps[Un_empty_right]; -goal thy "UNIV Un B = UNIV"; +Goal "UNIV Un B = UNIV"; by (Blast_tac 1); qed "Un_UNIV_left"; Addsimps[Un_UNIV_left]; -goal thy "A Un UNIV = UNIV"; +Goal "A Un UNIV = UNIV"; by (Blast_tac 1); qed "Un_UNIV_right"; Addsimps[Un_UNIV_right]; -goal thy "A Un B = Union{A,B}"; +Goal "A Un B = Union{A,B}"; by (Blast_tac 1); qed "Un_eq_Union"; -goal thy "(insert a B) Un C = insert a (B Un C)"; +Goal "(insert a B) Un C = insert a (B Un C)"; by (Blast_tac 1); qed "Un_insert_left"; Addsimps[Un_insert_left]; -goal thy "A Un (insert a B) = insert a (A Un B)"; +Goal "A Un (insert a B) = insert a (A Un B)"; by (Blast_tac 1); qed "Un_insert_right"; Addsimps[Un_insert_right]; -goal thy "(insert a B) Int C = (if a:C then insert a (B Int C) \ +Goal "(insert a B) Int C = (if a:C then insert a (B Int C) \ \ else B Int C)"; by (Simp_tac 1); by (Blast_tac 1); qed "Int_insert_left"; -goal thy "A Int (insert a B) = (if a:A then insert a (A Int B) \ +Goal "A Int (insert a B) = (if a:A then insert a (A Int B) \ \ else A Int B)"; by (Simp_tac 1); by (Blast_tac 1); qed "Int_insert_right"; -goal thy "A Un (B Int C) = (A Un B) Int (A Un C)"; +Goal "A Un (B Int C) = (A Un B) Int (A Un C)"; by (Blast_tac 1); qed "Un_Int_distrib"; -goal thy "(B Int C) Un A = (B Un A) Int (C Un A)"; +Goal "(B Int C) Un A = (B Un A) Int (C Un A)"; by (Blast_tac 1); qed "Un_Int_distrib2"; -goal thy +Goal "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; by (Blast_tac 1); qed "Un_Int_crazy"; -goal thy "(A<=B) = (A Un B = B)"; +Goal "(A<=B) = (A Un B = B)"; by (blast_tac (claset() addSEs [equalityCE]) 1); qed "subset_Un_eq"; -goal thy "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)"; +Goal "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)"; by (Blast_tac 1); qed "subset_insert_iff"; -goal thy "(A Un B = {}) = (A = {} & B = {})"; +Goal "(A Un B = {}) = (A = {} & B = {})"; by (blast_tac (claset() addEs [equalityCE]) 1); qed "Un_empty"; Addsimps[Un_empty]; section "Compl"; -goal thy "A Int Compl(A) = {}"; +Goal "A Int Compl(A) = {}"; by (Blast_tac 1); qed "Compl_disjoint"; Addsimps[Compl_disjoint]; -goal thy "A Un Compl(A) = UNIV"; +Goal "A Un Compl(A) = UNIV"; by (Blast_tac 1); qed "Compl_partition"; -goal thy "Compl(Compl(A)) = A"; +Goal "Compl(Compl(A)) = A"; by (Blast_tac 1); qed "double_complement"; Addsimps[double_complement]; -goal thy "Compl(A Un B) = Compl(A) Int Compl(B)"; +Goal "Compl(A Un B) = Compl(A) Int Compl(B)"; by (Blast_tac 1); qed "Compl_Un"; -goal thy "Compl(A Int B) = Compl(A) Un Compl(B)"; +Goal "Compl(A Int B) = Compl(A) Un Compl(B)"; by (Blast_tac 1); qed "Compl_Int"; -goal thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"; +Goal "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"; by (Blast_tac 1); qed "Compl_UN"; -goal thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"; +Goal "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"; by (Blast_tac 1); qed "Compl_INT"; @@ -359,68 +359,68 @@ (*Halmos, Naive Set Theory, page 16.*) -goal thy "((A Int B) Un C = A Int (B Un C)) = (C<=A)"; +Goal "((A Int B) Un C = A Int (B Un C)) = (C<=A)"; by (blast_tac (claset() addSEs [equalityCE]) 1); qed "Un_Int_assoc_eq"; section "Union"; -goal thy "Union({}) = {}"; +Goal "Union({}) = {}"; by (Blast_tac 1); qed "Union_empty"; Addsimps[Union_empty]; -goal thy "Union(UNIV) = UNIV"; +Goal "Union(UNIV) = UNIV"; by (Blast_tac 1); qed "Union_UNIV"; Addsimps[Union_UNIV]; -goal thy "Union(insert a B) = a Un Union(B)"; +Goal "Union(insert a B) = a Un Union(B)"; by (Blast_tac 1); qed "Union_insert"; Addsimps[Union_insert]; -goal thy "Union(A Un B) = Union(A) Un Union(B)"; +Goal "Union(A Un B) = Union(A) Un Union(B)"; by (Blast_tac 1); qed "Union_Un_distrib"; Addsimps[Union_Un_distrib]; -goal thy "Union(A Int B) <= Union(A) Int Union(B)"; +Goal "Union(A Int B) <= Union(A) Int Union(B)"; by (Blast_tac 1); qed "Union_Int_subset"; -goal thy "(Union M = {}) = (! A : M. A = {})"; +Goal "(Union M = {}) = (! A : M. A = {})"; by (blast_tac (claset() addEs [equalityCE]) 1); qed "Union_empty_conv"; AddIffs [Union_empty_conv]; -goal thy "(Union(C) Int A = {}) = (! B:C. B Int A = {})"; +Goal "(Union(C) Int A = {}) = (! B:C. B Int A = {})"; by (blast_tac (claset() addSEs [equalityCE]) 1); qed "Union_disjoint"; section "Inter"; -goal thy "Inter({}) = UNIV"; +Goal "Inter({}) = UNIV"; by (Blast_tac 1); qed "Inter_empty"; Addsimps[Inter_empty]; -goal thy "Inter(UNIV) = {}"; +Goal "Inter(UNIV) = {}"; by (Blast_tac 1); qed "Inter_UNIV"; Addsimps[Inter_UNIV]; -goal thy "Inter(insert a B) = a Int Inter(B)"; +Goal "Inter(insert a B) = a Int Inter(B)"; by (Blast_tac 1); qed "Inter_insert"; Addsimps[Inter_insert]; -goal thy "Inter(A) Un Inter(B) <= Inter(A Int B)"; +Goal "Inter(A) Un Inter(B) <= Inter(A Int B)"; by (Blast_tac 1); qed "Inter_Un_subset"; -goal thy "Inter(A Un B) = Inter(A) Int Inter(B)"; +Goal "Inter(A Un B) = Inter(A) Int Inter(B)"; by (Blast_tac 1); qed "Inter_Un_distrib"; @@ -431,138 +431,138 @@ val not_empty = prove_goal Set.thy "(A ~= {}) = (? x. x:A)" (K [Blast_tac 1]); (*Addsimps[not_empty];*) -goal thy "(UN x:{}. B x) = {}"; +Goal "(UN x:{}. B x) = {}"; by (Blast_tac 1); qed "UN_empty"; Addsimps[UN_empty]; -goal thy "(UN x:A. {}) = {}"; +Goal "(UN x:A. {}) = {}"; by (Blast_tac 1); qed "UN_empty2"; Addsimps[UN_empty2]; -goal thy "(UN x:A. {x}) = A"; +Goal "(UN x:A. {x}) = A"; by (Blast_tac 1); qed "UN_singleton"; Addsimps [UN_singleton]; -goal thy "!!k. k:I ==> A k Un (UN i:I. A i) = (UN i:I. A i)"; +Goal "!!k. k:I ==> A k Un (UN i:I. A i) = (UN i:I. A i)"; by (Blast_tac 1); qed "UN_absorb"; -goal thy "(INT x:{}. B x) = UNIV"; +Goal "(INT x:{}. B x) = UNIV"; by (Blast_tac 1); qed "INT_empty"; Addsimps[INT_empty]; -goal thy "!!k. k:I ==> A k Int (INT i:I. A i) = (INT i:I. A i)"; +Goal "!!k. k:I ==> A k Int (INT i:I. A i) = (INT i:I. A i)"; by (Blast_tac 1); qed "INT_absorb"; -goal thy "(UN x:insert a A. B x) = B a Un UNION A B"; +Goal "(UN x:insert a A. B x) = B a Un UNION A B"; by (Blast_tac 1); qed "UN_insert"; Addsimps[UN_insert]; -goal thy "(UN i: A Un B. M i) = ((UN i: A. M i) Un (UN i:B. M i))"; +Goal "(UN i: A Un B. M i) = ((UN i: A. M i) Un (UN i:B. M i))"; by (Blast_tac 1); qed "UN_Un"; -goal thy "(UN x : (UN y:A. B y). C x) = (UN y:A. UN x: B y. C x)"; +Goal "(UN x : (UN y:A. B y). C x) = (UN y:A. UN x: B y. C x)"; by (Blast_tac 1); qed "UN_UN_flatten"; -goal thy "(INT x:insert a A. B x) = B a Int INTER A B"; +Goal "(INT x:insert a A. B x) = B a Int INTER A B"; by (Blast_tac 1); qed "INT_insert"; Addsimps[INT_insert]; -goal thy +Goal "!!A. A~={} ==> (INT x:A. insert a (B x)) = insert a (INT x:A. B x)"; by (Blast_tac 1); qed "INT_insert_distrib"; -goal thy "Union(B``A) = (UN x:A. B(x))"; +Goal "Union(B``A) = (UN x:A. B(x))"; by (Blast_tac 1); qed "Union_image_eq"; -goal thy "Inter(B``A) = (INT x:A. B(x))"; +Goal "Inter(B``A) = (INT x:A. B(x))"; by (Blast_tac 1); qed "Inter_image_eq"; -goal thy "!!A. A~={} ==> (UN y:A. c) = c"; +Goal "!!A. A~={} ==> (UN y:A. c) = c"; by (Blast_tac 1); qed "UN_constant"; Addsimps[UN_constant]; -goal thy "!!A. A~={} ==> (INT y:A. c) = c"; +Goal "!!A. A~={} ==> (INT y:A. c) = c"; by (Blast_tac 1); qed "INT_constant"; Addsimps[INT_constant]; -goal thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})"; +Goal "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})"; by (Blast_tac 1); qed "UN_eq"; (*Look: it has an EXISTENTIAL quantifier*) -goal thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})"; +Goal "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})"; by (Blast_tac 1); qed "INT_eq"; -goalw thy [o_def] "UNION A (g o f) = UNION (f``A) g"; +Goalw [o_def] "UNION A (g o f) = UNION (f``A) g"; by (Blast_tac 1); qed "UNION_o"; (*Distributive laws...*) -goal thy "A Int Union(B) = (UN C:B. A Int C)"; +Goal "A Int Union(B) = (UN C:B. A Int C)"; by (Blast_tac 1); qed "Int_Union"; -goal thy "Union(B) Int A = (UN C:B. C Int A)"; +Goal "Union(B) Int A = (UN C:B. C Int A)"; by (Blast_tac 1); qed "Int_Union2"; (* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: Union of a family of unions **) -goal thy "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)"; +Goal "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)"; by (Blast_tac 1); qed "Un_Union_image"; (*Equivalent version*) -goal thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; +Goal "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; by (Blast_tac 1); qed "UN_Un_distrib"; -goal thy "A Un Inter(B) = (INT C:B. A Un C)"; +Goal "A Un Inter(B) = (INT C:B. A Un C)"; by (Blast_tac 1); qed "Un_Inter"; -goal thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)"; +Goal "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)"; by (Blast_tac 1); qed "Int_Inter_image"; (*Equivalent version*) -goal thy "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; +Goal "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; by (Blast_tac 1); qed "INT_Int_distrib"; (*Halmos, Naive Set Theory, page 35.*) -goal thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; +Goal "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; by (Blast_tac 1); qed "Int_UN_distrib"; -goal thy "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; +Goal "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; by (Blast_tac 1); qed "Un_INT_distrib"; -goal thy +Goal "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; by (Blast_tac 1); qed "Int_UN_distrib2"; -goal thy +Goal "(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; by (Blast_tac 1); qed "Un_INT_distrib2"; @@ -573,155 +573,155 @@ (** The following are not added to the default simpset because (a) they duplicate the body and (b) there are no similar rules for Int. **) -goal thy "(ALL x:A Un B. P x) = ((ALL x:A. P x) & (ALL x:B. P x))"; +Goal "(ALL x:A Un B. P x) = ((ALL x:A. P x) & (ALL x:B. P x))"; by (Blast_tac 1); qed "ball_Un"; -goal thy "(EX x:A Un B. P x) = ((EX x:A. P x) | (EX x:B. P x))"; +Goal "(EX x:A Un B. P x) = ((EX x:A. P x) | (EX x:B. P x))"; by (Blast_tac 1); qed "bex_Un"; -goal thy "(ALL z: UNION A B. P z) = (ALL x:A. ALL z:B x. P z)"; +Goal "(ALL z: UNION A B. P z) = (ALL x:A. ALL z:B x. P z)"; by (Blast_tac 1); qed "ball_UN"; -goal thy "(EX z: UNION A B. P z) = (EX x:A. EX z:B x. P z)"; +Goal "(EX z: UNION A B. P z) = (EX x:A. EX z:B x. P z)"; by (Blast_tac 1); qed "bex_UN"; section "-"; -goal thy "A-B = A Int Compl B"; +Goal "A-B = A Int Compl B"; by (Blast_tac 1); qed "Diff_eq"; -goal thy "A-A = {}"; +Goal "A-A = {}"; by (Blast_tac 1); qed "Diff_cancel"; Addsimps[Diff_cancel]; -goal thy "!!A B. A Int B = {} ==> A-B = A"; +Goal "!!A B. A Int B = {} ==> A-B = A"; by (blast_tac (claset() addEs [equalityE]) 1); qed "Diff_triv"; -goal thy "{}-A = {}"; +Goal "{}-A = {}"; by (Blast_tac 1); qed "empty_Diff"; Addsimps[empty_Diff]; -goal thy "A-{} = A"; +Goal "A-{} = A"; by (Blast_tac 1); qed "Diff_empty"; Addsimps[Diff_empty]; -goal thy "A-UNIV = {}"; +Goal "A-UNIV = {}"; by (Blast_tac 1); qed "Diff_UNIV"; Addsimps[Diff_UNIV]; -goal thy "!!x. x~:A ==> A - insert x B = A-B"; +Goal "!!x. x~:A ==> A - insert x B = A-B"; by (Blast_tac 1); qed "Diff_insert0"; Addsimps [Diff_insert0]; (*NOT SUITABLE FOR REWRITING since {a} == insert a 0*) -goal thy "A - insert a B = A - B - {a}"; +Goal "A - insert a B = A - B - {a}"; by (Blast_tac 1); qed "Diff_insert"; (*NOT SUITABLE FOR REWRITING since {a} == insert a 0*) -goal thy "A - insert a B = A - {a} - B"; +Goal "A - insert a B = A - {a} - B"; by (Blast_tac 1); qed "Diff_insert2"; -goal thy "insert x A - B = (if x:B then A-B else insert x (A-B))"; +Goal "insert x A - B = (if x:B then A-B else insert x (A-B))"; by (Simp_tac 1); by (Blast_tac 1); qed "insert_Diff_if"; -goal thy "!!x. x:B ==> insert x A - B = A-B"; +Goal "!!x. x:B ==> insert x A - B = A-B"; by (Blast_tac 1); qed "insert_Diff1"; Addsimps [insert_Diff1]; -goal thy "!!a. a:A ==> insert a (A-{a}) = A"; +Goal "!!a. a:A ==> insert a (A-{a}) = A"; by (Blast_tac 1); qed "insert_Diff"; -goal thy "A Int (B-A) = {}"; +Goal "A Int (B-A) = {}"; by (Blast_tac 1); qed "Diff_disjoint"; Addsimps[Diff_disjoint]; -goal thy "!!A. A<=B ==> A Un (B-A) = B"; +Goal "!!A. A<=B ==> A Un (B-A) = B"; by (Blast_tac 1); qed "Diff_partition"; -goal thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)"; +Goal "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)"; by (Blast_tac 1); qed "double_diff"; -goal thy "A Un (B-A) = A Un B"; +Goal "A Un (B-A) = A Un B"; by (Blast_tac 1); qed "Un_Diff_cancel"; -goal thy "(B-A) Un A = B Un A"; +Goal "(B-A) Un A = B Un A"; by (Blast_tac 1); qed "Un_Diff_cancel2"; Addsimps [Un_Diff_cancel, Un_Diff_cancel2]; -goal thy "A - (B Un C) = (A-B) Int (A-C)"; +Goal "A - (B Un C) = (A-B) Int (A-C)"; by (Blast_tac 1); qed "Diff_Un"; -goal thy "A - (B Int C) = (A-B) Un (A-C)"; +Goal "A - (B Int C) = (A-B) Un (A-C)"; by (Blast_tac 1); qed "Diff_Int"; -goal thy "(A Un B) - C = (A - C) Un (B - C)"; +Goal "(A Un B) - C = (A - C) Un (B - C)"; by (Blast_tac 1); qed "Un_Diff"; -goal thy "(A Int B) - C = A Int (B - C)"; +Goal "(A Int B) - C = A Int (B - C)"; by (Blast_tac 1); qed "Int_Diff"; -goal thy "C Int (A-B) = (C Int A) - (C Int B)"; +Goal "C Int (A-B) = (C Int A) - (C Int B)"; by (Blast_tac 1); qed "Diff_Int_distrib"; -goal thy "(A-B) Int C = (A Int C) - (B Int C)"; +Goal "(A-B) Int C = (A Int C) - (B Int C)"; by (Blast_tac 1); qed "Diff_Int_distrib2"; section "Miscellany"; -goal thy "(A = B) = ((A <= (B::'a set)) & (B<=A))"; +Goal "(A = B) = ((A <= (B::'a set)) & (B<=A))"; by (Blast_tac 1); qed "set_eq_subset"; -goal thy "A <= B = (! t. t:A --> t:B)"; +Goal "A <= B = (! t. t:A --> t:B)"; by (Blast_tac 1); qed "subset_iff"; -goalw thy [psubset_def] "((A::'a set) <= B) = ((A < B) | (A=B))"; +Goalw [psubset_def] "((A::'a set) <= B) = ((A < B) | (A=B))"; by (Blast_tac 1); qed "subset_iff_psubset_eq"; -goal thy "(!x. x ~: A) = (A={})"; +Goal "(!x. x ~: A) = (A={})"; by (Blast_tac 1); qed "all_not_in_conv"; AddIffs [all_not_in_conv]; -goalw thy [Pow_def] "Pow {} = {{}}"; +Goalw [Pow_def] "Pow {} = {{}}"; by Auto_tac; qed "Pow_empty"; Addsimps [Pow_empty]; -goal thy "Pow (insert a A) = Pow A Un (insert a `` Pow A)"; +Goal "Pow (insert a A) = Pow A Un (insert a `` Pow A)"; by Safe_tac; by (etac swap 1); by (res_inst_tac [("x", "x-{a}")] image_eqI 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/BT.ML --- a/src/HOL/ex/BT.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/BT.ML Mon Jun 22 17:26:46 1998 +0200 @@ -10,48 +10,48 @@ (** BT simplification **) -goal BT.thy "n_leaves(reflect(t)) = n_leaves(t)"; +Goal "n_leaves(reflect(t)) = n_leaves(t)"; by (bt.induct_tac "t" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_commute]))); qed "n_leaves_reflect"; -goal BT.thy "n_nodes(reflect(t)) = n_nodes(t)"; +Goal "n_nodes(reflect(t)) = n_nodes(t)"; by (bt.induct_tac "t" 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_commute]))); qed "n_nodes_reflect"; (*The famous relationship between the numbers of leaves and nodes*) -goal BT.thy "n_leaves(t) = Suc(n_nodes(t))"; +Goal "n_leaves(t) = Suc(n_nodes(t))"; by (bt.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "n_leaves_nodes"; -goal BT.thy "reflect(reflect(t))=t"; +Goal "reflect(reflect(t))=t"; by (bt.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "reflect_reflect_ident"; -goal BT.thy "bt_map f (reflect t) = reflect (bt_map f t)"; +Goal "bt_map f (reflect t) = reflect (bt_map f t)"; by (bt.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "bt_map_reflect"; -goal BT.thy "inorder (bt_map f t) = map f (inorder t)"; +Goal "inorder (bt_map f t) = map f (inorder t)"; by (bt.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "inorder_bt_map"; -goal BT.thy "preorder (reflect t) = rev (postorder t)"; +Goal "preorder (reflect t) = rev (postorder t)"; by (bt.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "preorder_reflect"; -goal BT.thy "inorder (reflect t) = rev (inorder t)"; +Goal "inorder (reflect t) = rev (inorder t)"; by (bt.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "inorder_reflect"; -goal BT.thy "postorder (reflect t) = rev (preorder t)"; +Goal "postorder (reflect t) = rev (preorder t)"; by (bt.induct_tac "t" 1); by (ALLGOALS Asm_simp_tac); qed "postorder_reflect"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/Fib.ML --- a/src/HOL/ex/Fib.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/Fib.ML Mon Jun 22 17:26:46 1998 +0200 @@ -25,7 +25,7 @@ val fib_Suc3 = read_instantiate [("x", "(Suc ?n)")] fib_Suc_Suc; (*Concrete Mathematics, page 280*) -goal thy "fib (Suc (n + k)) = fib(Suc k) * fib(Suc n) + fib k * fib n"; +Goal "fib (Suc (n + k)) = fib(Suc k) * fib(Suc n) + fib k * fib n"; by (res_inst_tac [("u","n")] fib.induct 1); (*Simplify the LHS just enough to apply the induction hypotheses*) by (asm_full_simp_tac @@ -37,7 +37,7 @@ qed "fib_add"; -goal thy "fib (Suc n) ~= 0"; +Goal "fib (Suc n) ~= 0"; by (res_inst_tac [("u","n")] fib.induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [fib_Suc_Suc]))); qed "fib_Suc_neq_0"; @@ -45,14 +45,14 @@ (* Also add 0 < fib (Suc n) *) Addsimps [fib_Suc_neq_0, [neq0_conv, fib_Suc_neq_0] MRS iffD1]; -goal thy "!!n. 0 0 < fib n"; +Goal "!!n. 0 0 < fib n"; by (rtac (not0_implies_Suc RS exE) 1); by Auto_tac; qed "fib_gr_0"; (*Concrete Mathematics, page 278: Cassini's identity*) -goal thy "fib (Suc (Suc n)) * fib n = \ +Goal "fib (Suc (Suc n)) * fib n = \ \ (if n mod 2 = 0 then (fib(Suc n) * fib(Suc n)) - 1 \ \ else Suc (fib(Suc n) * fib(Suc n)))"; by (res_inst_tac [("u","n")] fib.induct 1); @@ -73,7 +73,7 @@ (** Towards Law 6.111 of Concrete Mathematics **) -goal thy "gcd(fib n, fib (Suc n)) = 1"; +Goal "gcd(fib n, fib (Suc n)) = 1"; by (res_inst_tac [("u","n")] fib.induct 1); by (asm_simp_tac (simpset() addsimps [fib_Suc3, gcd_commute, gcd_add2]) 3); by (ALLGOALS (simp_tac (simpset() addsimps [fib_Suc_Suc]))); @@ -82,7 +82,7 @@ val gcd_fib_commute = read_instantiate_sg (sign_of thy) [("m", "fib m")] gcd_commute; -goal thy "gcd(fib m, fib (n+m)) = gcd(fib m, fib n)"; +Goal "gcd(fib m, fib (n+m)) = gcd(fib m, fib n)"; by (simp_tac (simpset() addsimps [gcd_fib_commute]) 1); by (case_tac "m=0" 1); by (Asm_simp_tac 1); @@ -93,12 +93,12 @@ by (asm_simp_tac (simpset() addsimps [gcd_fib_Suc_eq_1, gcd_mult_cancel]) 1); qed "gcd_fib_add"; -goal thy "!!m. m <= n ==> gcd(fib m, fib (n-m)) = gcd(fib m, fib n)"; +Goal "!!m. m <= n ==> gcd(fib m, fib (n-m)) = gcd(fib m, fib n)"; by (rtac (gcd_fib_add RS sym RS trans) 1); by (Asm_simp_tac 1); qed "gcd_fib_diff"; -goal thy "!!m. 0 gcd (fib m, fib (n mod m)) = gcd (fib m, fib n)"; +Goal "!!m. 0 gcd (fib m, fib (n mod m)) = gcd (fib m, fib n)"; by (res_inst_tac [("n","n")] less_induct 1); by (stac mod_if 1); by (Asm_simp_tac 1); @@ -107,7 +107,7 @@ qed "gcd_fib_mod"; (*Law 6.111*) -goal thy "fib(gcd(m,n)) = gcd(fib m, fib n)"; +Goal "fib(gcd(m,n)) = gcd(fib m, fib n)"; by (res_inst_tac [("m","m"),("n","n")] gcd_induct 1); by (Asm_simp_tac 1); by (asm_full_simp_tac (simpset() addsimps [gcd_non_0]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/InSort.ML --- a/src/HOL/ex/InSort.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/InSort.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,18 +6,18 @@ Correctness proof of insertion sort. *) -goal thy "!y. mset(ins f x xs) y = mset (x#xs) y"; +Goal "!y. mset(ins f x xs) y = mset (x#xs) y"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "mset_ins"; Addsimps [mset_ins]; -goal thy "!x. mset(insort f xs) x = mset xs x"; +Goal "!x. mset(insort f xs) x = mset xs x"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "insort_permutes"; -goal thy "set(ins f x xs) = insert x (set xs)"; +Goal "set(ins f x xs) = insert x (set xs)"; by (asm_simp_tac (simpset() addsimps [set_via_mset]) 1); by (Fast_tac 1); qed "set_ins"; @@ -32,7 +32,7 @@ qed "sorted_ins"; Addsimps [sorted_ins]; -goal InSort.thy "!!f. [| total(f); transf(f) |] ==> sorted f (insort f xs)"; +Goal "!!f. [| total(f); transf(f) |] ==> sorted f (insort f xs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "sorted_insort"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/MT.ML --- a/src/HOL/ex/MT.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/MT.ML Mon Jun 22 17:26:46 1998 +0200 @@ -140,13 +140,13 @@ (* Monotonicity of eval_fun *) -goalw MT.thy [mono_def, eval_fun_def] "mono(eval_fun)"; +Goalw [mono_def, eval_fun_def] "mono(eval_fun)"; by infsys_mono_tac; qed "eval_fun_mono"; (* Introduction rules *) -goalw MT.thy [eval_def, eval_rel_def] "ve |- e_const(c) ---> v_const(c)"; +Goalw [eval_def, eval_rel_def] "ve |- e_const(c) ---> v_const(c)"; by (rtac lfp_intro2 1); by (rtac eval_fun_mono 1); by (rewtac eval_fun_def); @@ -263,13 +263,13 @@ (* Elaborations *) (* ############################################################ *) -goalw MT.thy [mono_def, elab_fun_def] "mono(elab_fun)"; +Goalw [mono_def, elab_fun_def] "mono(elab_fun)"; by infsys_mono_tac; qed "elab_fun_mono"; (* Introduction rules *) -goalw MT.thy [elab_def, elab_rel_def] +Goalw [elab_def, elab_rel_def] "!!te. c isof ty ==> te |- e_const(c) ===> ty"; by (rtac lfp_intro2 1); by (rtac elab_fun_mono 1); @@ -277,7 +277,7 @@ by (blast_tac (claset() addSIs [exI]) 1); qed "elab_const"; -goalw MT.thy [elab_def, elab_rel_def] +Goalw [elab_def, elab_rel_def] "!!te. x:te_dom(te) ==> te |- e_var(x) ===> te_app te x"; by (rtac lfp_intro2 1); by (rtac elab_fun_mono 1); @@ -285,7 +285,7 @@ by (blast_tac (claset() addSIs [exI]) 1); qed "elab_var"; -goalw MT.thy [elab_def, elab_rel_def] +Goalw [elab_def, elab_rel_def] "!!te. te + {x |=> ty1} |- e ===> ty2 ==> te |- fn x => e ===> ty1->ty2"; by (rtac lfp_intro2 1); by (rtac elab_fun_mono 1); @@ -293,7 +293,7 @@ by (blast_tac (claset() addSIs [exI]) 1); qed "elab_fn"; -goalw MT.thy [elab_def, elab_rel_def] +Goalw [elab_def, elab_rel_def] "!!te. te + {f |=> ty1->ty2} + {x |=> ty1} |- e ===> ty2 ==> \ \ te |- fix f(x) = e ===> ty1->ty2"; by (rtac lfp_intro2 1); @@ -302,7 +302,7 @@ by (blast_tac (claset() addSIs [exI]) 1); qed "elab_fix"; -goalw MT.thy [elab_def, elab_rel_def] +Goalw [elab_def, elab_rel_def] "!!te. [| te |- e1 ===> ty1->ty2; te |- e2 ===> ty1 |] ==> \ \ te |- e1 @ e2 ===> ty2"; by (rtac lfp_intro2 1); @@ -493,7 +493,7 @@ (* Monotonicity of hasty_fun *) -goalw MT.thy [mono_def,MT.hasty_fun_def] "mono(hasty_fun)"; +Goalw [mono_def,MT.hasty_fun_def] "mono(hasty_fun)"; by infsys_mono_tac; by (Blast_tac 1); qed "mono_hasty_fun"; @@ -506,7 +506,7 @@ (* First strong indtroduction (co-induction) rule for hasty_rel *) val prems = - goalw MT.thy [hasty_rel_def] "c isof t ==> (v_const(c),t) : hasty_rel"; +Goalw [hasty_rel_def] "c isof t ==> (v_const(c),t) : hasty_rel"; by (cut_facts_tac prems 1); by (rtac gfp_coind2 1); by (rewtac MT.hasty_fun_def); @@ -574,11 +574,11 @@ (* Introduction rules for hasty *) -goalw MT.thy [hasty_def] "!!c. c isof t ==> v_const(c) hasty t"; +Goalw [hasty_def] "!!c. c isof t ==> v_const(c) hasty t"; by (etac hasty_rel_const_coind 1); qed "hasty_const"; -goalw MT.thy [hasty_def,hasty_env_def] +Goalw [hasty_def,hasty_env_def] "!!t. te |- fn ev => e ===> t & ve hastyenv te ==> v_clos(<|ev,e,ve|>) hasty t"; by (rtac hasty_rel_clos_coind 1); by (ALLGOALS (blast_tac (claset() delrules [equalityI]))); @@ -586,13 +586,13 @@ (* Elimination on constants for hasty *) -goalw MT.thy [hasty_def] +Goalw [hasty_def] "!!v. v hasty t ==> (!c.(v = v_const(c) --> c isof t))"; by (rtac hasty_rel_elim 1); by (ALLGOALS (blast_tac (v_ext_cs HOL_cs))); qed "hasty_elim_const_lem"; -goal MT.thy "!!c. v_const(c) hasty t ==> c isof t"; +Goal "!!c. v_const(c) hasty t ==> c isof t"; by (dtac hasty_elim_const_lem 1); by (Blast_tac 1); qed "hasty_elim_const"; @@ -608,7 +608,7 @@ by (ALLGOALS (blast_tac (v_ext_cs HOL_cs))); qed "hasty_elim_clos_lem"; -goal MT.thy +Goal "!!t. v_clos(<|ev,e,ve|>) hasty t ==> \ \ ? te. te |- fn ev => e ===> t & ve hastyenv te "; by (dtac hasty_elim_clos_lem 1); @@ -619,7 +619,7 @@ (* The pointwise extension of hasty to environments *) (* ############################################################ *) -goal MT.thy +Goal "!!ve. [| ve hastyenv te; v hasty t |] ==> \ \ ve + {ev |-> v} hastyenv te + {ev |=> t}"; by (rewtac hasty_env_def); @@ -636,27 +636,27 @@ (* The Consistency theorem *) (* ############################################################ *) -goal MT.thy +Goal "!!t. [| ve hastyenv te ; te |- e_const(c) ===> t |] ==> v_const(c) hasty t"; by (dtac elab_const_elim 1); by (etac hasty_const 1); qed "consistency_const"; -goalw MT.thy [hasty_env_def] +Goalw [hasty_env_def] "!!t. [| ev : ve_dom(ve); ve hastyenv te ; te |- e_var(ev) ===> t |] ==> \ \ ve_app ve ev hasty t"; by (dtac elab_var_elim 1); by (Blast_tac 1); qed "consistency_var"; -goal MT.thy +Goal "!!t. [| ve hastyenv te ; te |- fn ev => e ===> t |] ==> \ \ v_clos(<| ev, e, ve |>) hasty t"; by (rtac hasty_clos 1); by (Blast_tac 1); qed "consistency_fn"; -goalw MT.thy [hasty_env_def,hasty_def] +Goalw [hasty_env_def,hasty_def] "!!t. [| cl = <| ev1, e, ve + { ev2 |-> v_clos(cl) } |>; \ \ ve hastyenv te ; \ \ te |- fix ev2 ev1 = e ===> t \ @@ -683,7 +683,7 @@ by (Blast_tac 1); qed "consistency_fix"; -goal MT.thy +Goal "!!t. [| ! t te. ve hastyenv te --> te |- e1 ===> t --> v_const(c1) hasty t;\ \ ! t te. ve hastyenv te --> te |- e2 ===> t --> v_const(c2) hasty t; \ \ ve hastyenv te ; te |- e1 @ e2 ===> t \ @@ -699,7 +699,7 @@ by (Blast_tac 1); qed "consistency_app1"; -goal MT.thy +Goal "!!t. [| ! t te. \ \ ve hastyenv te --> \ \ te |- e1 ===> t --> v_clos(<|evm, em, vem|>) hasty t; \ diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/NatSum.ML --- a/src/HOL/ex/NatSum.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/NatSum.ML Mon Jun 22 17:26:46 1998 +0200 @@ -12,14 +12,14 @@ Addsimps [add_mult_distrib, add_mult_distrib2]; (*The sum of the first n positive integers equals n(n+1)/2.*) -goal NatSum.thy "2*sum (%i. i) (Suc n) = n*Suc(n)"; +Goal "2*sum (%i. i) (Suc n) = n*Suc(n)"; by (Simp_tac 1); by (induct_tac "n" 1); by (Simp_tac 1); by (Asm_simp_tac 1); qed "sum_of_naturals"; -goal NatSum.thy +Goal "Suc(Suc(Suc(Suc 2)))*sum (%i. i*i) (Suc n) = n*Suc(n)*Suc(2*n)"; by (Simp_tac 1); by (induct_tac "n" 1); @@ -27,7 +27,7 @@ by (Asm_simp_tac 1); qed "sum_of_squares"; -goal NatSum.thy +Goal "Suc(Suc 2)*sum (%i. i*i*i) (Suc n) = n*n*Suc(n)*Suc(n)"; by (Simp_tac 1); by (induct_tac "n" 1); @@ -36,7 +36,7 @@ qed "sum_of_cubes"; (*The sum of the first n odd numbers equals n squared.*) -goal NatSum.thy "sum (%i. Suc(i+i)) n = n*n"; +Goal "sum (%i. Suc(i+i)) n = n*n"; by (induct_tac "n" 1); by (Simp_tac 1); by (Asm_simp_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/Primes.ML --- a/src/HOL/ex/Primes.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/Primes.ML Mon Jun 22 17:26:46 1998 +0200 @@ -40,25 +40,25 @@ qed "gcd_induct"; -goal thy "gcd(m,0) = m"; +Goal "gcd(m,0) = m"; by (rtac (gcd_eq RS trans) 1); by (Simp_tac 1); qed "gcd_0"; Addsimps [gcd_0]; -goal thy "!!m. 0 gcd(m,n) = gcd (n, m mod n)"; +Goal "!!m. 0 gcd(m,n) = gcd (n, m mod n)"; by (rtac (gcd_eq RS trans) 1); by (Asm_simp_tac 1); by (Blast_tac 1); qed "gcd_non_0"; -goal thy "gcd(m,1) = 1"; +Goal "gcd(m,1) = 1"; by (simp_tac (simpset() addsimps [gcd_non_0]) 1); qed "gcd_1"; Addsimps [gcd_1]; (*gcd(m,n) divides m and n. The conjunctions don't seem provable separately*) -goal thy "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)"; +Goal "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)"; by (res_inst_tac [("m","m"),("n","n")] gcd_induct 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [gcd_non_0, mod_less_divisor]))); @@ -71,7 +71,7 @@ (*Maximality: for all m,n,f naturals, if f divides m and f divides n then f divides gcd(m,n)*) -goal thy "!!k. (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)"; +Goal "!!k. (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)"; by (res_inst_tac [("m","m"),("n","n")] gcd_induct 1); by (ALLGOALS (asm_full_simp_tac (simpset() addsimps[gcd_non_0, dvd_mod, @@ -79,34 +79,34 @@ qed_spec_mp "gcd_greatest"; (*Function gcd yields the Greatest Common Divisor*) -goalw thy [is_gcd_def] "is_gcd (gcd(m,n)) m n"; +Goalw [is_gcd_def] "is_gcd (gcd(m,n)) m n"; by (asm_simp_tac (simpset() addsimps [gcd_greatest, gcd_dvd_both]) 1); qed "is_gcd"; (*uniqueness of GCDs*) -goalw thy [is_gcd_def] "!!m n. [| is_gcd m a b; is_gcd n a b |] ==> m=n"; +Goalw [is_gcd_def] "!!m n. [| is_gcd m a b; is_gcd n a b |] ==> m=n"; by (blast_tac (claset() addIs [dvd_anti_sym]) 1); qed "is_gcd_unique"; (*USED??*) -goalw thy [is_gcd_def] +Goalw [is_gcd_def] "!!m n. [| is_gcd m a b; k dvd a; k dvd b |] ==> k dvd m"; by (Blast_tac 1); qed "is_gcd_dvd"; (** Commutativity **) -goalw thy [is_gcd_def] "is_gcd k m n = is_gcd k n m"; +Goalw [is_gcd_def] "is_gcd k m n = is_gcd k n m"; by (Blast_tac 1); qed "is_gcd_commute"; -goal thy "gcd(m,n) = gcd(n,m)"; +Goal "gcd(m,n) = gcd(n,m)"; by (rtac is_gcd_unique 1); by (rtac is_gcd 2); by (asm_simp_tac (simpset() addsimps [is_gcd, is_gcd_commute]) 1); qed "gcd_commute"; -goal thy "gcd(gcd(k,m),n) = gcd(k,gcd(m,n))"; +Goal "gcd(gcd(k,m),n) = gcd(k,gcd(m,n))"; by (rtac is_gcd_unique 1); by (rtac is_gcd 2); by (rewtac is_gcd_def); @@ -114,12 +114,12 @@ addIs [gcd_greatest, dvd_trans]) 1); qed "gcd_assoc"; -goal thy "gcd(0,m) = m"; +Goal "gcd(0,m) = m"; by (stac gcd_commute 1); by (rtac gcd_0 1); qed "gcd_0_left"; -goal thy "gcd(1,m) = 1"; +Goal "gcd(1,m) = 1"; by (stac gcd_commute 1); by (rtac gcd_1 1); qed "gcd_1_left"; @@ -129,7 +129,7 @@ (** Multiplication laws **) (*Davenport, page 27*) -goal thy "k * gcd(m,n) = gcd(k*m, k*n)"; +Goal "k * gcd(m,n) = gcd(k*m, k*n)"; by (res_inst_tac [("m","m"),("n","n")] gcd_induct 1); by (Asm_full_simp_tac 1); by (case_tac "k=0" 1); @@ -138,39 +138,39 @@ (simpset() addsimps [mod_geq, gcd_non_0, mod_mult_distrib2]) 1); qed "gcd_mult_distrib2"; -goal thy "gcd(m,m) = m"; +Goal "gcd(m,m) = m"; by (cut_inst_tac [("k","m"),("m","1"),("n","1")] gcd_mult_distrib2 1); by (Asm_full_simp_tac 1); qed "gcd_self"; Addsimps [gcd_self]; -goal thy "gcd(k, k*n) = k"; +Goal "gcd(k, k*n) = k"; by (cut_inst_tac [("k","k"),("m","1"),("n","n")] gcd_mult_distrib2 1); by (Asm_full_simp_tac 1); qed "gcd_mult"; Addsimps [gcd_mult]; -goal thy "!!k. [| gcd(k,n)=1; k dvd (m*n) |] ==> k dvd m"; +Goal "!!k. [| gcd(k,n)=1; k dvd (m*n) |] ==> k dvd m"; by (subgoal_tac "m = gcd(m*k, m*n)" 1); by (etac ssubst 1 THEN rtac gcd_greatest 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [gcd_mult_distrib2 RS sym]))); qed "relprime_dvd_mult"; -goalw thy [prime_def] "!!p. [| p: prime; ~ p dvd n |] ==> gcd (p, n) = 1"; +Goalw [prime_def] "!!p. [| p: prime; ~ p dvd n |] ==> gcd (p, n) = 1"; by (cut_inst_tac [("m","p"),("n","n")] gcd_dvd_both 1); by (fast_tac (claset() addss (simpset())) 1); qed "prime_imp_relprime"; (*This theorem leads immediately to a proof of the uniqueness of factorization. If p divides a product of primes then it is one of those primes.*) -goal thy "!!p. [| p: prime; p dvd (m*n) |] ==> p dvd m | p dvd n"; +Goal "!!p. [| p: prime; p dvd (m*n) |] ==> p dvd m | p dvd n"; by (blast_tac (claset() addIs [relprime_dvd_mult, prime_imp_relprime]) 1); qed "prime_dvd_mult"; (** Addition laws **) -goal thy "gcd(m, m+n) = gcd(m,n)"; +Goal "gcd(m, m+n) = gcd(m,n)"; by (res_inst_tac [("n1", "m+n")] (gcd_commute RS ssubst) 1); by (rtac (gcd_eq RS trans) 1); by Auto_tac; @@ -180,19 +180,19 @@ by Safe_tac; qed "gcd_add"; -goal thy "gcd(m, n+m) = gcd(m,n)"; +Goal "gcd(m, n+m) = gcd(m,n)"; by (asm_simp_tac (simpset() addsimps [add_commute, gcd_add]) 1); qed "gcd_add2"; (** More multiplication laws **) -goal thy "gcd(m,n) dvd gcd(k*m, n)"; +Goal "gcd(m,n) dvd gcd(k*m, n)"; by (blast_tac (claset() addIs [gcd_greatest, dvd_trans, gcd_dvd1, gcd_dvd2]) 1); qed "gcd_dvd_gcd_mult"; -goal thy "!!n. gcd(k,n) = 1 ==> gcd(k*m, n) = gcd(m,n)"; +Goal "!!n. gcd(k,n) = 1 ==> gcd(k*m, n) = gcd(m,n)"; by (rtac dvd_anti_sym 1); by (rtac gcd_dvd_gcd_mult 2); by (rtac ([relprime_dvd_mult, gcd_dvd2] MRS gcd_greatest) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/Primrec.ML --- a/src/HOL/ex/Primrec.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/Primrec.ML Mon Jun 22 17:26:46 1998 +0200 @@ -17,27 +17,27 @@ (** Useful special cases of evaluation ***) -goalw thy [SC_def] "SC (x#l) = Suc x"; +Goalw [SC_def] "SC (x#l) = Suc x"; by (Asm_simp_tac 1); qed "SC"; -goalw thy [CONST_def] "CONST k l = k"; +Goalw [CONST_def] "CONST k l = k"; by (Asm_simp_tac 1); qed "CONST"; -goalw thy [PROJ_def] "PROJ(0) (x#l) = x"; +Goalw [PROJ_def] "PROJ(0) (x#l) = x"; by (Asm_simp_tac 1); qed "PROJ_0"; -goalw thy [COMP_def] "COMP g [f] l = g [f l]"; +Goalw [COMP_def] "COMP g [f] l = g [f l]"; by (Asm_simp_tac 1); qed "COMP_1"; -goalw thy [PREC_def] "PREC f g (0#l) = f l"; +Goalw [PREC_def] "PREC f g (0#l) = f l"; by (Asm_simp_tac 1); qed "PREC_0"; -goalw thy [PREC_def] "PREC f g (Suc x # l) = g (PREC f g (x#l) # x # l)"; +Goalw [PREC_def] "PREC f g (Suc x # l) = g (PREC f g (x#l) # x # l)"; by (Asm_simp_tac 1); qed "PREC_Suc"; @@ -47,7 +47,7 @@ Addsimps ack.rules; (*PROPERTY A 4*) -goal thy "j < ack(i,j)"; +Goal "j < ack(i,j)"; by (res_inst_tac [("u","i"),("v","j")] ack.induct 1); by (ALLGOALS Asm_simp_tac); by (ALLGOALS trans_tac); @@ -56,7 +56,7 @@ AddIffs [less_ack2]; (*PROPERTY A 5-, the single-step lemma*) -goal thy "ack(i,j) < ack(i, Suc(j))"; +Goal "ack(i,j) < ack(i, Suc(j))"; by (res_inst_tac [("u","i"),("v","j")] ack.induct 1); by (ALLGOALS Asm_simp_tac); qed "ack_less_ack_Suc2"; @@ -64,20 +64,20 @@ AddIffs [ack_less_ack_Suc2]; (*PROPERTY A 5, monotonicity for < *) -goal thy "j ack(i,j) < ack(i,k)"; +Goal "j ack(i,j) < ack(i,k)"; by (res_inst_tac [("u","i"),("v","k")] ack.induct 1); by (ALLGOALS Asm_simp_tac); by (blast_tac (claset() addSEs [less_SucE] addIs [less_trans]) 1); qed_spec_mp "ack_less_mono2"; (*PROPERTY A 5', monotonicity for<=*) -goal thy "!!i j k. j<=k ==> ack(i,j)<=ack(i,k)"; +Goal "!!i j k. j<=k ==> ack(i,j)<=ack(i,k)"; by (full_simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); by (blast_tac (claset() addIs [ack_less_mono2]) 1); qed "ack_le_mono2"; (*PROPERTY A 6*) -goal thy "ack(i, Suc(j)) <= ack(Suc(i), j)"; +Goal "ack(i, Suc(j)) <= ack(Suc(i), j)"; by (induct_tac "j" 1); by (ALLGOALS Asm_simp_tac); by (blast_tac (claset() addIs [ack_le_mono2, less_ack2 RS Suc_leI, @@ -87,14 +87,14 @@ AddIffs [ack2_le_ack1]; (*PROPERTY A 7-, the single-step lemma*) -goal thy "ack(i,j) < ack(Suc(i),j)"; +Goal "ack(i,j) < ack(Suc(i),j)"; by (blast_tac (claset() addIs [ack_less_mono2, less_le_trans]) 1); qed "ack_less_ack_Suc1"; AddIffs [ack_less_ack_Suc1]; (*PROPERTY A 4'? Extra lemma needed for CONST case, constant functions*) -goal thy "i < ack(i,j)"; +Goal "i < ack(i,j)"; by (induct_tac "i" 1); by (ALLGOALS Asm_simp_tac); by (blast_tac (claset() addIs [Suc_leI, le_less_trans]) 1); @@ -102,14 +102,14 @@ AddIffs [less_ack1]; (*PROPERTY A 8*) -goal thy "ack(1,j) = Suc(Suc(j))"; +Goal "ack(1,j) = Suc(Suc(j))"; by (induct_tac "j" 1); by (ALLGOALS Asm_simp_tac); qed "ack_1"; Addsimps [ack_1]; (*PROPERTY A 9*) -goal thy "ack(Suc(1),j) = Suc(Suc(Suc(j+j)))"; +Goal "ack(Suc(1),j) = Suc(Suc(Suc(j+j)))"; by (induct_tac "j" 1); by (ALLGOALS Asm_simp_tac); qed "ack_2"; @@ -117,7 +117,7 @@ (*PROPERTY A 7, monotonicity for < [not clear why ack_1 is now needed first!]*) -goal thy "ack(i,k) < ack(Suc(i+i'),k)"; +Goal "ack(i,k) < ack(Suc(i+i'),k)"; by (res_inst_tac [("u","i"),("v","k")] ack.induct 1); by (ALLGOALS Asm_full_simp_tac); by (blast_tac (claset() addIs [less_trans, ack_less_mono2]) 2); @@ -127,19 +127,19 @@ by (blast_tac (claset() addIs [Suc_leI RS le_less_trans, ack_less_mono2]) 1); val lemma = result(); -goal thy "!!i j k. i ack(i,k) < ack(j,k)"; +Goal "!!i j k. i ack(i,k) < ack(j,k)"; by (etac less_natE 1); by (blast_tac (claset() addSIs [lemma]) 1); qed "ack_less_mono1"; (*PROPERTY A 7', monotonicity for<=*) -goal thy "!!i j k. i<=j ==> ack(i,k)<=ack(j,k)"; +Goal "!!i j k. i<=j ==> ack(i,k)<=ack(j,k)"; by (full_simp_tac (simpset() addsimps [le_eq_less_or_eq]) 1); by (blast_tac (claset() addIs [ack_less_mono1]) 1); qed "ack_le_mono1"; (*PROPERTY A 10*) -goal thy "ack(i1, ack(i2,j)) < ack(Suc(Suc(i1+i2)), j)"; +Goal "ack(i1, ack(i2,j)) < ack(Suc(Suc(i1+i2)), j)"; by (rtac (ack2_le_ack1 RSN (2,less_le_trans)) 1); by (Asm_simp_tac 1); by (rtac (le_add1 RS ack_le_mono1 RS le_less_trans) 1); @@ -148,7 +148,7 @@ qed "ack_nest_bound"; (*PROPERTY A 11*) -goal thy "ack(i1,j) + ack(i2,j) < ack(Suc(Suc(Suc(Suc(i1+i2)))), j)"; +Goal "ack(i1,j) + ack(i2,j) < ack(Suc(Suc(Suc(Suc(i1+i2)))), j)"; by (res_inst_tac [("j", "ack(Suc(1), ack(i1 + i2, j))")] less_trans 1); by (Asm_simp_tac 1); by (rtac (ack_nest_bound RS less_le_trans) 2); @@ -160,7 +160,7 @@ (*PROPERTY A 12. Article uses existential quantifier but the ALF proof used k+4. Quantified version must be nested EX k'. ALL i,j... *) -goal thy "!!i j k. i < ack(k,j) ==> i+j < ack(Suc(Suc(Suc(Suc(k)))), j)"; +Goal "!!i j k. i < ack(k,j) ==> i+j < ack(Suc(Suc(Suc(Suc(k)))), j)"; by (res_inst_tac [("j", "ack(k,j) + ack(0,j)")] less_trans 1); by (rtac (ack_add_bound RS less_le_trans) 2); by (Asm_simp_tac 2); @@ -172,16 +172,16 @@ (*** MAIN RESULT ***) -goalw thy [SC_def] "SC l < ack(1, list_add l)"; +Goalw [SC_def] "SC l < ack(1, list_add l)"; by (induct_tac "l" 1); by (ALLGOALS (simp_tac (simpset() addsimps [le_add1, le_imp_less_Suc]))); qed "SC_case"; -goal thy "CONST k l < ack(k, list_add l)"; +Goal "CONST k l < ack(k, list_add l)"; by (Simp_tac 1); qed "CONST_case"; -goalw thy [PROJ_def] "ALL i. PROJ i l < ack(0, list_add l)"; +Goalw [PROJ_def] "ALL i. PROJ i l < ack(0, list_add l)"; by (Simp_tac 1); by (induct_tac "l" 1); by (ALLGOALS Asm_simp_tac); @@ -195,7 +195,7 @@ (** COMP case **) -goal thy +Goal "!!fs. fs : lists(PRIMREC Int {f. EX kf. ALL l. f l < ack(kf, list_add l)}) \ \ ==> EX k. ALL l. list_add (map(%f. f l) fs) < ack(k, list_add l)"; by (etac lists.induct 1); @@ -205,7 +205,7 @@ by (blast_tac (claset() addIs [add_less_mono, ack_add_bound, less_trans]) 1); qed "COMP_map_lemma"; -goalw thy [COMP_def] +Goalw [COMP_def] "!!g. [| ALL l. g l < ack(kg, list_add l); \ \ fs: lists(PRIMREC Int {f. EX kf. ALL l. f l < ack(kf, list_add l)}) \ \ |] ==> EX k. ALL l. COMP g fs l < ack(k, list_add l)"; @@ -220,7 +220,7 @@ (** PREC case **) -goalw thy [PREC_def] +Goalw [PREC_def] "!!f g. [| ALL l. f l + list_add l < ack(kf, list_add l); \ \ ALL l. g l + list_add l < ack(kg, list_add l) \ \ |] ==> PREC f g l + list_add l < ack(Suc(kf+kg), list_add l)"; @@ -244,7 +244,7 @@ by (etac ack_less_mono2 1); qed "PREC_case_lemma"; -goal thy +Goal "!!f g. [| ALL l. f l < ack(kf, list_add l); \ \ ALL l. g l < ack(kg, list_add l) \ \ |] ==> EX k. ALL l. PREC f g l< ack(k, list_add l)"; @@ -254,14 +254,14 @@ by (REPEAT (blast_tac (claset() addIs [ack_add_bound2]) 1)); qed "PREC_case"; -goal thy "!!f. f:PRIMREC ==> EX k. ALL l. f l < ack(k, list_add l)"; +Goal "!!f. f:PRIMREC ==> EX k. ALL l. f l < ack(k, list_add l)"; by (etac PRIMREC.induct 1); by (ALLGOALS (blast_tac (claset() addIs [SC_case, CONST_case, PROJ_case, COMP_case, PREC_case]))); qed "ack_bounds_PRIMREC"; -goal thy "(%l. case l of [] => 0 | x#l' => ack(x,x)) ~: PRIMREC"; +Goal "(%l. case l of [] => 0 | x#l' => ack(x,x)) ~: PRIMREC"; by (rtac notI 1); by (etac (ack_bounds_PRIMREC RS exE) 1); by (rtac less_irrefl 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/Puzzle.ML --- a/src/HOL/ex/Puzzle.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/Puzzle.ML Mon Jun 22 17:26:46 1998 +0200 @@ -13,7 +13,7 @@ by (EVERY1 [rtac (nat_induct RS allI), resolve_tac prems, resolve_tac prems]); qed "nat_exh"; -goal Puzzle.thy "! n. k=f(n) --> n <= f(n)"; +Goal "! n. k=f(n) --> n <= f(n)"; by (res_inst_tac [("n","k")] less_induct 1); by (rtac nat_exh 1); by (Simp_tac 1); @@ -27,11 +27,11 @@ addIs [Puzzle.f_ax, le_less_trans]) 1); val lemma = result() RS spec RS mp; -goal Puzzle.thy "n <= f(n)"; +Goal "n <= f(n)"; by (fast_tac (claset() addIs [lemma]) 1); qed "lemma1"; -goal Puzzle.thy "f(n) < f(Suc(n))"; +Goal "f(n) < f(Suc(n))"; by (deepen_tac (claset() addIs [Puzzle.f_ax, le_less_trans, lemma1]) 0 1); qed "lemma2"; @@ -53,7 +53,7 @@ by (fast_tac (claset() addIs ([mono_lemma,less_imp_le,lemma2]@prems)) 1); qed "f_mono"; -goal Puzzle.thy "f(n) = n"; +Goal "f(n) = n"; by (rtac le_anti_sym 1); by (rtac lemma1 2); by (fast_tac (claset() addIs [Puzzle.f_ax,leI] addDs [leD,f_mono,Suc_leI]) 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/Qsort.ML --- a/src/HOL/ex/Qsort.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/Qsort.ML Mon Jun 22 17:26:46 1998 +0200 @@ -24,12 +24,12 @@ Addsimps ([Qsort.qsort_Nil,Qsort.qsort_Cons]@conj_comms); -goal Qsort.thy "!x. mset (qsort le xs) x = mset xs x"; +Goal "!x. mset (qsort le xs) x = mset xs x"; by (res_inst_tac[("xs","xs"),("p","le")]Qsort.qsort_ind 1); by (ALLGOALS Asm_simp_tac); qed "qsort_permutes"; -goal Qsort.thy "set(qsort le xs) = set xs"; +Goal "set(qsort le xs) = set xs"; by (simp_tac (simpset() addsimps [set_via_mset,qsort_permutes]) 1); qed "set_qsort"; Addsimps [set_qsort]; @@ -41,7 +41,7 @@ qed"Ball_set_filter"; Addsimps [Ball_set_filter]; -goal Qsort.thy +Goal "sorted le (xs@ys) = (sorted le xs & sorted le ys & \ \ (!x:set xs. !y:set ys. le x y))"; by (list.induct_tac "xs" 1); @@ -49,7 +49,7 @@ qed "sorted_append"; Addsimps [sorted_append]; -goal Qsort.thy +Goal "!!le. [| total(le); transf(le) |] ==> sorted le (qsort le xs)"; by (res_inst_tac[("xs","xs"),("p","le")]Qsort.qsort_ind 1); by (ALLGOALS Asm_simp_tac); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/Recdef.ML --- a/src/HOL/ex/Recdef.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/Recdef.ML Mon Jun 22 17:26:46 1998 +0200 @@ -11,7 +11,7 @@ Addsimps qsort.rules; -goal thy "(x mem qsort (ord,l)) = (x mem l)"; +Goal "(x mem qsort (ord,l)) = (x mem l)"; by (res_inst_tac [("u","ord"),("v","l")] qsort.induct 1); by (ALLGOALS Asm_simp_tac); by (Blast_tac 1); @@ -22,13 +22,13 @@ Addsimps g.rules; -goal thy "g x < Suc x"; +Goal "g x < Suc x"; by (res_inst_tac [("u","x")] g.induct 1); by Auto_tac; by (trans_tac 1); qed "g_terminates"; -goal thy "g x = 0"; +Goal "g x = 0"; by (res_inst_tac [("u","x")] g.induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [g_terminates]))); qed "g_zero"; diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/Sorting.ML --- a/src/HOL/ex/Sorting.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/Sorting.ML Mon Jun 22 17:26:46 1998 +0200 @@ -6,12 +6,12 @@ Some general lemmas *) -goal Sorting.thy "!x. mset (xs@ys) x = mset xs x + mset ys x"; +Goal "!x. mset (xs@ys) x = mset xs x + mset ys x"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "mset_append"; -goal Sorting.thy "!x. mset [x:xs. ~p(x)] x + mset [x:xs. p(x)] x = \ +Goal "!x. mset [x:xs. ~p(x)] x + mset [x:xs. p(x)] x = \ \ mset xs x"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); @@ -19,7 +19,7 @@ Addsimps [mset_append, mset_compl_add]; -goal Sorting.thy "set xs = {x. mset xs x ~= 0}"; +Goal "set xs = {x. mset xs x ~= 0}"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); by (Fast_tac 1); diff -r fb28eaa07e01 -r 3ea049f7979d src/HOL/ex/String.ML --- a/src/HOL/ex/String.ML Mon Jun 22 17:13:09 1998 +0200 +++ b/src/HOL/ex/String.ML Mon Jun 22 17:26:46 1998 +0200 @@ -1,20 +1,20 @@ -goal String.thy "hd(''ABCD'') = CHR ''A''"; +Goal "hd(''ABCD'') = CHR ''A''"; by (Simp_tac 1); result(); -goal String.thy "hd(''ABCD'') ~= CHR ''B''"; +Goal "hd(''ABCD'') ~= CHR ''B''"; by (Simp_tac 1); result(); -goal String.thy "''ABCD'' ~= ''ABCX''"; +Goal "''ABCD'' ~= ''ABCX''"; by (Simp_tac 1); result(); -goal String.thy "''ABCD'' = ''ABCD''"; +Goal "''ABCD'' = ''ABCD''"; by (Simp_tac 1); result(); -goal String.thy +Goal "''ABCDEFGHIJKLMNOPQRSTUVWXYZ'' ~= ''ABCDEFGHIJKLMNOPQRSTUVWXY''"; by (Simp_tac 1); result();