# HG changeset patch # User clasohm # Date 794228545 -3600 # Node ID ff1574a81019c69a0fba59b9be6adf7d7ca367ff # Parent 196ca0973a6d679150ac86426aba621fb0347d7b new version of HOL with curried function application diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Arith.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,373 @@ +(* Title: HOL/Arith.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Proofs about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +open Arith; + +(*** Basic rewrite rules for the arithmetic operators ***) + +val [pred_0, pred_Suc] = nat_recs pred_def; +val [add_0,add_Suc] = nat_recs add_def; +val [mult_0,mult_Suc] = nat_recs mult_def; + +(** Difference **) + +val diff_0 = diff_def RS def_nat_rec_0; + +qed_goalw "diff_0_eq_0" Arith.thy [diff_def, pred_def] + "0 - n = 0" + (fn _ => [nat_ind_tac "n" 1, ALLGOALS(asm_simp_tac nat_ss)]); + +(*Must simplify BEFORE the induction!! (Else we get a critical pair) + Suc(m) - Suc(n) rewrites to pred(Suc(m) - n) *) +qed_goalw "diff_Suc_Suc" Arith.thy [diff_def, pred_def] + "Suc(m) - Suc(n) = m - n" + (fn _ => + [simp_tac nat_ss 1, nat_ind_tac "n" 1, ALLGOALS(asm_simp_tac nat_ss)]); + +(*** Simplification over add, mult, diff ***) + +val arith_simps = + [pred_0, pred_Suc, add_0, add_Suc, mult_0, mult_Suc, + diff_0, diff_0_eq_0, diff_Suc_Suc]; + +val arith_ss = nat_ss addsimps arith_simps; + +(**** Inductive properties of the operators ****) + +(*** Addition ***) + +qed_goal "add_0_right" Arith.thy "m + 0 = m" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +qed_goal "add_Suc_right" Arith.thy "m + Suc(n) = Suc(m+n)" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +val arith_ss = arith_ss addsimps [add_0_right,add_Suc_right]; + +(*Associative law for addition*) +qed_goal "add_assoc" Arith.thy "(m + n) + k = m + ((n + k)::nat)" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*Commutative law for addition*) +qed_goal "add_commute" Arith.thy "m + n = n + (m::nat)" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +qed_goal "add_left_commute" Arith.thy "x+(y+z)=y+((x+z)::nat)" + (fn _ => [rtac (add_commute RS trans) 1, rtac (add_assoc RS trans) 1, + rtac (add_commute RS arg_cong) 1]); + +(*Addition is an AC-operator*) +val add_ac = [add_assoc, add_commute, add_left_commute]; + +goal Arith.thy "!!k::nat. (k + m = k + n) = (m=n)"; +by (nat_ind_tac "k" 1); +by (simp_tac arith_ss 1); +by (asm_simp_tac arith_ss 1); +qed "add_left_cancel"; + +goal Arith.thy "!!k::nat. (m + k = n + k) = (m=n)"; +by (nat_ind_tac "k" 1); +by (simp_tac arith_ss 1); +by (asm_simp_tac arith_ss 1); +qed "add_right_cancel"; + +goal Arith.thy "!!k::nat. (k + m <= k + n) = (m<=n)"; +by (nat_ind_tac "k" 1); +by (simp_tac arith_ss 1); +by (asm_simp_tac (arith_ss addsimps [Suc_le_mono]) 1); +qed "add_left_cancel_le"; + +goal Arith.thy "!!k::nat. (k + m < k + n) = (m [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*right Sucessor law for multiplication*) +qed_goal "mult_Suc_right" Arith.thy "m * Suc(n) = m + (m * n)" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac (arith_ss addsimps add_ac))]); + +val arith_ss = arith_ss addsimps [mult_0_right,mult_Suc_right]; + +(*Commutative law for multiplication*) +qed_goal "mult_commute" Arith.thy "m * n = n * (m::nat)" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS (asm_simp_tac arith_ss)]); + +(*addition distributes over multiplication*) +qed_goal "add_mult_distrib" Arith.thy "(m + n)*k = (m*k) + ((n*k)::nat)" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac (arith_ss addsimps add_ac))]); + +qed_goal "add_mult_distrib2" Arith.thy "k*(m + n) = (k*m) + ((k*n)::nat)" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac (arith_ss addsimps add_ac))]); + +val arith_ss = arith_ss addsimps [add_mult_distrib,add_mult_distrib2]; + +(*Associative law for multiplication*) +qed_goal "mult_assoc" Arith.thy "(m * n) * k = m * ((n * k)::nat)" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +qed_goal "mult_left_commute" Arith.thy "x*(y*z) = y*((x*z)::nat)" + (fn _ => [rtac trans 1, rtac mult_commute 1, rtac trans 1, + rtac mult_assoc 1, rtac (mult_commute RS arg_cong) 1]); + +val mult_ac = [mult_assoc,mult_commute,mult_left_commute]; + +(*** Difference ***) + +qed_goal "diff_self_eq_0" Arith.thy "m - m = 0" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *) +val [prem] = goal Arith.thy "[| ~ m n+(m-n) = (m::nat)"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS(asm_simp_tac arith_ss)); +qed "add_diff_inverse"; + + +(*** Remainder ***) + +goal Arith.thy "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 arith_ss)); +qed "diff_less_Suc"; + +goal Arith.thy "!!m::nat. m - n <= m"; +by (res_inst_tac [("m","m"), ("n","n")] diff_induct 1); +by (ALLGOALS (asm_simp_tac arith_ss)); +by (etac le_trans 1); +by (simp_tac (HOL_ss addsimps [le_eq_less_or_eq, lessI]) 1); +qed "diff_le_self"; + +goal Arith.thy "!!n::nat. (n+m) - n = m"; +by (nat_ind_tac "n" 1); +by (ALLGOALS (asm_simp_tac arith_ss)); +qed "diff_add_inverse"; + +goal Arith.thy "!!n::nat. n - (n+m) = 0"; +by (nat_ind_tac "n" 1); +by (ALLGOALS (asm_simp_tac arith_ss)); +qed "diff_add_0"; + +(*In ordinary notation: if 0 m - n < m"; +by (subgoal_tac "0 ~ m m - n < m" 1); +by (fast_tac HOL_cs 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS(asm_simp_tac(arith_ss addsimps [diff_less_Suc]))); +qed "div_termination"; + +val wf_less_trans = wf_pred_nat RS wf_trancl RSN (2, def_wfrec RS trans); + +goalw Nat.thy [less_def] " : pred_nat^+ = (m m mod n = m"; +by (rtac (mod_def RS wf_less_trans) 1); +by(asm_simp_tac HOL_ss 1); +qed "mod_less"; + +goal Arith.thy "!!m. [| 0 m mod n = (m-n) mod n"; +by (rtac (mod_def RS wf_less_trans) 1); +by(asm_simp_tac (nat_ss addsimps [div_termination, cut_apply, less_eq]) 1); +qed "mod_geq"; + + +(*** Quotient ***) + +goal Arith.thy "!!m. m m div n = 0"; +by (rtac (div_def RS wf_less_trans) 1); +by(asm_simp_tac nat_ss 1); +qed "div_less"; + +goal Arith.thy "!!M. [| 0 m div n = Suc((m-n) div n)"; +by (rtac (div_def RS wf_less_trans) 1); +by(asm_simp_tac (nat_ss addsimps [div_termination, cut_apply, less_eq]) 1); +qed "div_geq"; + +(*Main Result about quotient and remainder.*) +goal Arith.thy "!!m. 0 (m div n)*n + m mod n = m"; +by (res_inst_tac [("n","m")] less_induct 1); +by (rename_tac "k" 1); (*Variable name used in line below*) +by (case_tac "k m-n = 0"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS (asm_simp_tac arith_ss)); +qed "less_imp_diff_is_0"; + +val prems = goal Arith.thy "m-n = 0 --> n-m = 0 --> m=n"; +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (REPEAT(simp_tac arith_ss 1 THEN TRY(atac 1))); +qed "diffs0_imp_equal_lemma"; + +(* [| m-n = 0; n-m = 0 |] ==> m=n *) +bind_thm ("diffs0_imp_equal", (diffs0_imp_equal_lemma RS mp RS mp)); + +val [prem] = goal Arith.thy "m 0 Suc(m)-n = Suc(m-n)"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); +by (ALLGOALS(asm_simp_tac arith_ss)); +qed "Suc_diff_n"; + +goal Arith.thy "Suc(m)-n = if (m (!n. P(Suc(n))--> P(n)) --> P(k-i)"; +by (res_inst_tac [("m","k"),("n","i")] diff_induct 1); +by (ALLGOALS (strip_tac THEN' simp_tac arith_ss THEN' TRY o fast_tac HOL_cs)); +qed "zero_induct_lemma"; + +val prems = goal Arith.thy "[| P(k); !!n. P(Suc(n)) ==> P(n) |] ==> P(0)"; +by (rtac (diff_self_eq_0 RS subst) 1); +by (rtac (zero_induct_lemma RS mp RS mp) 1); +by (REPEAT (ares_tac ([impI,allI]@prems) 1)); +qed "zero_induct"; + +(*13 July 1992: loaded in 105.7s*) + +(**** Additional theorems about "less than" ****) + +goal Arith.thy "!!m. m (? k. n=Suc(m+k))"; +by (nat_ind_tac "n" 1); +by (ALLGOALS(simp_tac arith_ss)); +by (REPEAT_FIRST (ares_tac [conjI, impI])); +by (res_inst_tac [("x","0")] exI 2); +by (simp_tac arith_ss 2); +by (safe_tac HOL_cs); +by (res_inst_tac [("x","Suc(k)")] exI 1); +by (simp_tac arith_ss 1); +val less_eq_Suc_add_lemma = result(); + +(*"m ? k. n = Suc(m+k)"*) +bind_thm ("less_eq_Suc_add", less_eq_Suc_add_lemma RS mp); + + +goal Arith.thy "n <= ((m + n)::nat)"; +by (nat_ind_tac "m" 1); +by (ALLGOALS(simp_tac arith_ss)); +by (etac le_trans 1); +by (rtac (lessI RS less_imp_le) 1); +qed "le_add2"; + +goal Arith.thy "n <= ((n + m)::nat)"; +by (simp_tac (arith_ss addsimps add_ac) 1); +by (rtac le_add2 1); +qed "le_add1"; + +bind_thm ("less_add_Suc1", (lessI RS (le_add1 RS le_less_trans))); +bind_thm ("less_add_Suc2", (lessI RS (le_add2 RS le_less_trans))); + +(*"i <= j ==> i <= j+m"*) +bind_thm ("trans_le_add1", le_add1 RSN (2,le_trans)); + +(*"i <= j ==> i <= m+j"*) +bind_thm ("trans_le_add2", le_add2 RSN (2,le_trans)); + +(*"i < j ==> i < j+m"*) +bind_thm ("trans_less_add1", le_add1 RSN (2,less_le_trans)); + +(*"i < j ==> i < m+j"*) +bind_thm ("trans_less_add2", le_add2 RSN (2,less_le_trans)); + +goal Arith.thy "!!k::nat. m <= n ==> m <= n+k"; +by (eresolve_tac [le_trans] 1); +by (resolve_tac [le_add1] 1); +qed "le_imp_add_le"; + +goal Arith.thy "!!k::nat. m < n ==> m < n+k"; +by (eresolve_tac [less_le_trans] 1); +by (resolve_tac [le_add1] 1); +qed "less_imp_add_less"; + +goal Arith.thy "m+k<=n --> m<=(n::nat)"; +by (nat_ind_tac "k" 1); +by (ALLGOALS (asm_simp_tac arith_ss)); +by (fast_tac (HOL_cs addDs [Suc_leD]) 1); +val add_leD1_lemma = result(); +bind_thm ("add_leD1", add_leD1_lemma RS mp);; + +goal Arith.thy "!!k l::nat. [| k m i + k < j + k"; +by (nat_ind_tac "k" 1); +by (ALLGOALS (asm_simp_tac arith_ss)); +qed "add_less_mono1"; + +(*strict, in both arguments*) +goal Arith.thy "!!i j k::nat. [|i < j; k < l|] ==> i + k < j + l"; +by (rtac (add_less_mono1 RS less_trans) 1); +by (REPEAT (etac asm_rl 1)); +by (nat_ind_tac "j" 1); +by (ALLGOALS(asm_simp_tac arith_ss)); +qed "add_less_mono"; + +(*A [clumsy] way of lifting < monotonicity to <= monotonicity *) +val [lt_mono,le] = goal Arith.thy + "[| !!i j::nat. i f(i) < f(j); \ +\ i <= j \ +\ |] ==> f(i) <= (f(j)::nat)"; +by (cut_facts_tac [le] 1); +by (asm_full_simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1); +by (fast_tac (HOL_cs addSIs [lt_mono]) 1); +qed "less_mono_imp_le_mono"; + +(*non-strict, in 1st argument*) +goal Arith.thy "!!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 (eresolve_tac [add_less_mono1] 1); +by (assume_tac 1); +qed "add_le_mono1"; + +(*non-strict, in both arguments*) +goal Arith.thy "!!k l::nat. [|i<=j; k<=l |] ==> i + k <= j + l"; +by (etac (add_le_mono1 RS le_trans) 1); +by (simp_tac (HOL_ss addsimps [add_commute]) 1); +(*j moves to the end because it is free while k, l are bound*) +by (eresolve_tac [add_le_mono1] 1); +qed "add_le_mono"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Arith.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,31 @@ +(* Title: HOL/Arith.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Arithmetic operators and their definitions +*) + +Arith = Nat + + +instance + nat :: {plus, minus, times} + +consts + pred :: "nat => nat" + div, mod :: "[nat, nat] => nat" (infixl 70) + +defs + pred_def "pred(m) == nat_rec m 0 (%n r.n)" + add_def "m+n == nat_rec m n (%u v. Suc(v))" + diff_def "m-n == nat_rec n m (%u v. pred(v))" + mult_def "m*n == nat_rec m 0 (%u v. n + v)" + mod_def "m mod n == wfrec (trancl pred_nat) m (%j f.(if (jn. + Also, nat_rec(m, 0, %z w.z) is pred(m). *) + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Finite.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Finite.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,84 @@ +(* Title: HOL/Finite.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Finite powerset operator +*) + +open Finite; + +goalw Finite.thy Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)"; +br lfp_mono 1; +by (REPEAT (ares_tac basic_monos 1)); +qed "Fin_mono"; + +goalw Finite.thy Fin.defs "Fin(A) <= Pow(A)"; +by (fast_tac (set_cs addSIs [lfp_lowerbound]) 1); +qed "Fin_subset_Pow"; + +(* A : Fin(B) ==> A <= B *) +val FinD = Fin_subset_Pow RS subsetD RS PowD; + +(*Discharging ~ x:y entails extra work*) +val major::prems = goal Finite.thy + "[| F:Fin(A); P({}); \ +\ !!F x. [| x:A; F:Fin(A); x~:F; P(F) |] ==> P(insert x F) \ +\ |] ==> P(F)"; +by (rtac (major RS Fin.induct) 1); +by (excluded_middle_tac "a:b" 2); +by (etac (insert_absorb RS ssubst) 3 THEN assume_tac 3); (*backtracking!*) +by (REPEAT (ares_tac prems 1)); +qed "Fin_induct"; + +(** Simplification for Fin **) + +val Fin_ss = set_ss addsimps Fin.intrs; + +(*The union of two finite sets is finite*) +val major::prems = goal Finite.thy + "[| F: Fin(A); G: Fin(A) |] ==> F Un G : Fin(A)"; +by (rtac (major RS Fin_induct) 1); +by (ALLGOALS (asm_simp_tac (Fin_ss addsimps (prems@[Un_insert_left])))); +qed "Fin_UnI"; + +(*Every subset of a finite set is finite*) +val [subs,fin] = goal Finite.thy "[| A<=B; B: Fin(M) |] ==> A: Fin(M)"; +by (EVERY1 [subgoal_tac "ALL C. C<=B --> C: Fin(M)", + rtac mp, etac spec, + rtac subs]); +by (rtac (fin RS Fin_induct) 1); +by (simp_tac (Fin_ss addsimps [subset_Un_eq]) 1); +by (safe_tac (set_cs addSDs [subset_insert_iff RS iffD1])); +by (eres_inst_tac [("t","C")] (insert_Diff RS subst) 2); +by (ALLGOALS (asm_simp_tac Fin_ss)); +qed "Fin_subset"; + +(*The image of a finite set is finite*) +val major::_ = goal Finite.thy + "F: Fin(A) ==> h``F : Fin(h``A)"; +by (rtac (major RS Fin_induct) 1); +by (simp_tac Fin_ss 1); +by (asm_simp_tac (set_ss addsimps [image_eqI RS Fin.insertI, image_insert]) 1); +qed "Fin_imageI"; + +val major::prems = goal Finite.thy + "[| c: Fin(A); b: Fin(A); \ +\ P(b); \ +\ !!(x::'a) y. [| x:A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) \ +\ |] ==> c<=b --> P(b-c)"; +by (rtac (major RS Fin_induct) 1); +by (rtac (Diff_insert RS ssubst) 2); +by (ALLGOALS (asm_simp_tac + (Fin_ss addsimps (prems@[Diff_subset RS Fin_subset])))); +qed "Fin_empty_induct_lemma"; + +val prems = goal Finite.thy + "[| b: Fin(A); \ +\ P(b); \ +\ !!x y. [| x:A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) \ +\ |] ==> P({})"; +by (rtac (Diff_cancel RS subst) 1); +by (rtac (Fin_empty_induct_lemma RS mp) 1); +by (REPEAT (ares_tac (subset_refl::prems) 1)); +qed "Fin_empty_induct"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Finite.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Finite.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,17 @@ +(* Title: HOL/Finite.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Finite powerset operator +*) + +Finite = Lfp + +consts Fin :: "'a set => 'a set set" + +inductive "Fin(A)" + intrs + emptyI "{} : Fin(A)" + insertI "[| a: A; b: Fin(A) |] ==> insert a b : Fin(A)" + +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Fun.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Fun.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,200 @@ +(* Title: HOL/Fun + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Lemmas about functions. +*) + +goal Fun.thy "(f = g) = (!x. f(x)=g(x))"; +by (rtac iffI 1); +by(asm_simp_tac HOL_ss 1); +by(rtac ext 1 THEN asm_simp_tac HOL_ss 1); +qed "expand_fun_eq"; + +val prems = goal Fun.thy + "[| f(x)=u; !!x. P(x) ==> g(f(x)) = x; P(x) |] ==> x=g(u)"; +by (rtac (arg_cong RS box_equals) 1); +by (REPEAT (resolve_tac (prems@[refl]) 1)); +qed "apply_inverse"; + + +(*** Range of a function ***) + +(*Frequently b does not have the syntactic form of f(x).*) +val [prem] = goalw Fun.thy [range_def] "b=f(x) ==> b : range(f)"; +by (EVERY1 [rtac CollectI, rtac exI, rtac prem]); +qed "range_eqI"; + +val rangeI = refl RS range_eqI; + +val [major,minor] = goalw Fun.thy [range_def] + "[| b : range(%x.f(x)); !!x. b=f(x) ==> P |] ==> P"; +by (rtac (major RS CollectD RS exE) 1); +by (etac minor 1); +qed "rangeE"; + +(*** Image of a set under a function ***) + +val prems = goalw Fun.thy [image_def] "[| b=f(x); x:A |] ==> b : f``A"; +by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1)); +qed "image_eqI"; + +val imageI = refl RS image_eqI; + +(*The eta-expansion gives variable-name preservation.*) +val major::prems = goalw Fun.thy [image_def] + "[| b : (%x.f(x))``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P"; +by (rtac (major RS CollectD RS bexE) 1); +by (REPEAT (ares_tac prems 1)); +qed "imageE"; + +goalw Fun.thy [o_def] "(f o g)``r = f``(g``r)"; +by (rtac set_ext 1); +by (fast_tac (HOL_cs addIs [imageI] addSEs [imageE]) 1); +qed "image_compose"; + +goal Fun.thy "f``(A Un B) = f``A Un f``B"; +by (rtac set_ext 1); +by (fast_tac (HOL_cs addIs [imageI,UnCI] addSEs [imageE,UnE]) 1); +qed "image_Un"; + +(*** inj(f): f is a one-to-one function ***) + +val prems = goalw Fun.thy [inj_def] + "[| !! x y. f(x) = f(y) ==> x=y |] ==> inj(f)"; +by (fast_tac (HOL_cs addIs prems) 1); +qed "injI"; + +val [major] = goal Fun.thy "(!!x. g(f(x)) = x) ==> inj(f)"; +by (rtac injI 1); +by (etac (arg_cong RS box_equals) 1); +by (rtac major 1); +by (rtac major 1); +qed "inj_inverseI"; + +val [major,minor] = goalw Fun.thy [inj_def] + "[| inj(f); f(x) = f(y) |] ==> x=y"; +by (rtac (major RS spec RS spec RS mp) 1); +by (rtac minor 1); +qed "injD"; + +(*Useful with the simplifier*) +val [major] = goal Fun.thy "inj(f) ==> (f(x) = f(y)) = (x=y)"; +by (rtac iffI 1); +by (etac (major RS injD) 1); +by (etac arg_cong 1); +qed "inj_eq"; + +val [major] = goal Fun.thy "inj(f) ==> (@x.f(x)=f(y)) = y"; +by (rtac (major RS injD) 1); +by (rtac selectI 1); +by (rtac refl 1); +qed "inj_select"; + +(*A one-to-one function has an inverse (given using select).*) +val [major] = goalw Fun.thy [Inv_def] "inj(f) ==> Inv f (f x) = x"; +by (EVERY1 [rtac (major RS inj_select)]); +qed "Inv_f_f"; + +(* Useful??? *) +val [oneone,minor] = goal Fun.thy + "[| inj(f); !!y. y: range(f) ==> P(Inv f y) |] ==> P(x)"; +by (res_inst_tac [("t", "x")] (oneone RS (Inv_f_f RS subst)) 1); +by (rtac (rangeI RS minor) 1); +qed "inj_transfer"; + + +(*** inj_onto f A: f is one-to-one over A ***) + +val prems = goalw Fun.thy [inj_onto_def] + "(!! x y. [| f(x) = f(y); x:A; y:A |] ==> x=y) ==> inj_onto f A"; +by (fast_tac (HOL_cs addIs prems addSIs [ballI]) 1); +qed "inj_ontoI"; + +val [major] = goal Fun.thy + "(!!x. x:A ==> g(f(x)) = x) ==> inj_onto f A"; +by (rtac inj_ontoI 1); +by (etac (apply_inverse RS trans) 1); +by (REPEAT (eresolve_tac [asm_rl,major] 1)); +qed "inj_onto_inverseI"; + +val major::prems = goalw Fun.thy [inj_onto_def] + "[| inj_onto f A; f(x)=f(y); x:A; y:A |] ==> x=y"; +by (rtac (major RS bspec RS bspec RS mp) 1); +by (REPEAT (resolve_tac prems 1)); +qed "inj_ontoD"; + +goal Fun.thy "!!x y.[| inj_onto f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)"; +by (fast_tac (HOL_cs addSEs [inj_ontoD]) 1); +qed "inj_onto_iff"; + +val major::prems = goal Fun.thy + "[| inj_onto f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)"; +by (rtac contrapos 1); +by (etac (major RS inj_ontoD) 2); +by (REPEAT (resolve_tac prems 1)); +qed "inj_onto_contraD"; + + +(*** Lemmas about inj ***) + +val prems = goalw Fun.thy [o_def] + "[| inj(f); inj_onto g (range f) |] ==> inj(g o f)"; +by (cut_facts_tac prems 1); +by (fast_tac (HOL_cs addIs [injI,rangeI] + addEs [injD,inj_ontoD]) 1); +qed "comp_inj"; + +val [prem] = goal Fun.thy "inj(f) ==> inj_onto f A"; +by (fast_tac (HOL_cs addIs [prem RS injD, inj_ontoI]) 1); +qed "inj_imp"; + +val [prem] = goalw Fun.thy [Inv_def] "y : range(f) ==> f(Inv f y) = y"; +by (EVERY1 [rtac (prem RS rangeE), rtac selectI, etac sym]); +qed "f_Inv_f"; + +val prems = goal Fun.thy + "[| Inv f x=Inv f y; x: range(f); y: range(f) |] ==> x=y"; +by (rtac (arg_cong RS box_equals) 1); +by (REPEAT (resolve_tac (prems @ [f_Inv_f]) 1)); +qed "Inv_injective"; + +val prems = goal Fun.thy + "[| inj(f); A<=range(f) |] ==> inj_onto (Inv f) A"; +by (cut_facts_tac prems 1); +by (fast_tac (HOL_cs addIs [inj_ontoI] + addEs [Inv_injective,injD,subsetD]) 1); +qed "inj_onto_Inv"; + + +(*** Set reasoning tools ***) + +val set_cs = HOL_cs + addSIs [ballI, PowI, subsetI, InterI, INT_I, INT1_I, CollectI, + ComplI, IntI, DiffI, UnCI, insertCI] + addIs [bexI, UnionI, UN_I, UN1_I, imageI, rangeI] + addSEs [bexE, make_elim PowD, UnionE, UN_E, UN1_E, DiffE, + CollectE, ComplE, IntE, UnE, insertE, imageE, rangeE, emptyE] + addEs [ballE, InterD, InterE, INT_D, INT_E, make_elim INT1_D, + subsetD, subsetCE]; + +fun cfast_tac prems = cut_facts_tac prems THEN' fast_tac set_cs; + + +fun prover s = prove_goal Fun.thy s (fn _=>[fast_tac set_cs 1]); + +val mem_simps = map prover + [ "(a : A Un B) = (a:A | a:B)", + "(a : A Int B) = (a:A & a:B)", + "(a : Compl(B)) = (~a:B)", + "(a : A-B) = (a:A & ~a:B)", + "(a : {b}) = (a=b)", + "(a : {x.P(x)}) = P(a)" ]; + +val mksimps_pairs = ("Ball",[bspec]) :: mksimps_pairs; + +val set_ss = + HOL_ss addsimps mem_simps + addcongs [ball_cong,bex_cong] + setmksimps (mksimps mksimps_pairs); diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Fun.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Fun.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,9 @@ +(* Title: HOL/Fun.thy + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Lemmas about functions. +*) + +Fun = Set diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Gfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Gfp.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,145 @@ +(* Title: HOL/gfp + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For gfp.thy. The Knaster-Tarski Theorem for greatest fixed points. +*) + +open Gfp; + +(*** Proof of Knaster-Tarski Theorem using gfp ***) + +(* gfp(f) is the least upper bound of {u. u <= f(u)} *) + +val prems = goalw Gfp.thy [gfp_def] "[| X <= f(X) |] ==> X <= gfp(f)"; +by (rtac (CollectI RS Union_upper) 1); +by (resolve_tac prems 1); +qed "gfp_upperbound"; + +val prems = goalw Gfp.thy [gfp_def] + "[| !!u. u <= f(u) ==> u<=X |] ==> gfp(f) <= X"; +by (REPEAT (ares_tac ([Union_least]@prems) 1)); +by (etac CollectD 1); +qed "gfp_least"; + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) <= f(gfp(f))"; +by (EVERY1 [rtac gfp_least, rtac subset_trans, atac, + rtac (mono RS monoD), rtac gfp_upperbound, atac]); +qed "gfp_lemma2"; + +val [mono] = goal Gfp.thy "mono(f) ==> f(gfp(f)) <= gfp(f)"; +by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD), + rtac gfp_lemma2, rtac mono]); +qed "gfp_lemma3"; + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))"; +by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1)); +qed "gfp_Tarski"; + +(*** Coinduction rules for greatest fixed points ***) + +(*weak version*) +val prems = goal Gfp.thy + "[| a: X; X <= f(X) |] ==> a : gfp(f)"; +by (rtac (gfp_upperbound RS subsetD) 1); +by (REPEAT (ares_tac prems 1)); +qed "weak_coinduct"; + +val [prem,mono] = goal Gfp.thy + "[| X <= f(X Un gfp(f)); mono(f) |] ==> \ +\ X Un gfp(f) <= f(X Un gfp(f))"; +by (rtac (prem RS Un_least) 1); +by (rtac (mono RS gfp_lemma2 RS subset_trans) 1); +by (rtac (Un_upper2 RS subset_trans) 1); +by (rtac (mono RS mono_Un) 1); +qed "coinduct_lemma"; + +(*strong version, thanks to Coen & Frost*) +goal Gfp.thy + "!!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)); +qed "coinduct"; + +val [mono,prem] = goal Gfp.thy + "[| mono(f); a: gfp(f) |] ==> a: f(X Un gfp(f))"; +br (mono RS mono_Un RS subsetD) 1; +br (mono RS gfp_lemma2 RS subsetD RS UnI2) 1; +by (rtac prem 1); +qed "gfp_fun_UnI2"; + +(*** Even Stronger version of coinduct [by Martin Coen] + - instead of the condition X <= f(X) + consider X <= (f(X) Un f(f(X)) ...) Un gfp(X) ***) + +val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un X Un B)"; +by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1)); +qed "coinduct3_mono_lemma"; + +val [prem,mono] = goal Gfp.thy + "[| X <= f(lfp(%x.f(x) Un X Un gfp(f))); mono(f) |] ==> \ +\ lfp(%x.f(x) Un X Un gfp(f)) <= f(lfp(%x.f(x) Un X Un gfp(f)))"; +by (rtac subset_trans 1); +by (rtac (mono RS coinduct3_mono_lemma RS lfp_lemma3) 1); +by (rtac (Un_least RS Un_least) 1); +by (rtac subset_refl 1); +by (rtac prem 1); +by (rtac (mono RS gfp_Tarski RS equalityD1 RS subset_trans) 1); +by (rtac (mono RS monoD) 1); +by (rtac (mono RS coinduct3_mono_lemma RS lfp_Tarski RS ssubst) 1); +by (rtac Un_upper2 1); +qed "coinduct3_lemma"; + +val prems = goal Gfp.thy + "[| mono(f); a:X; X <= f(lfp(%x.f(x) Un X Un gfp(f))) |] ==> a : gfp(f)"; +by (rtac (coinduct3_lemma RSN (2,weak_coinduct)) 1); +by (resolve_tac (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1); +by (rtac (UnI2 RS UnI1) 1); +by (REPEAT (resolve_tac prems 1)); +qed "coinduct3"; + + +(** Definition forms of gfp_Tarski and coinduct, to control unfolding **) + +val [rew,mono] = goal Gfp.thy "[| A==gfp(f); mono(f) |] ==> A = f(A)"; +by (rewtac rew); +by (rtac (mono RS gfp_Tarski) 1); +qed "def_gfp_Tarski"; + +val rew::prems = goal Gfp.thy + "[| A==gfp(f); mono(f); a:X; X <= f(X Un A) |] ==> a: A"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct]) 1)); +qed "def_coinduct"; + +(*The version used in the induction/coinduction package*) +val prems = goal Gfp.thy + "[| A == gfp(%w. Collect(P(w))); mono(%w. Collect(P(w))); \ +\ a: X; !!z. z: X ==> P (X Un A) z |] ==> \ +\ a : A"; +by (rtac def_coinduct 1); +by (REPEAT (ares_tac (prems @ [subsetI,CollectI]) 1)); +qed "def_Collect_coinduct"; + +val rew::prems = goal Gfp.thy + "[| A==gfp(f); mono(f); a:X; X <= f(lfp(%x.f(x) Un X Un A)) |] ==> a: A"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1)); +qed "def_coinduct3"; + +(*Monotonicity of gfp!*) +val prems = goal Gfp.thy + "[| mono(f); !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)"; +by (rtac gfp_upperbound 1); +by (rtac subset_trans 1); +by (rtac gfp_lemma2 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +val gfp_mono = result(); + +(*Monotonicity of gfp!*) +val [prem] = goal Gfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)"; +br (gfp_upperbound RS gfp_least) 1; +be (prem RSN (2,subset_trans)) 1; +qed "gfp_mono"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Gfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Gfp.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,14 @@ +(* Title: HOL/gfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Greatest fixed points (requires Lfp too!) +*) + +Gfp = Lfp + +consts gfp :: "['a set=>'a set] => 'a set" +defs + (*greatest fixed point*) + gfp_def "gfp(f) == Union({u. u <= f(u)})" +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/HOL.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HOL.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,266 @@ +(* Title: HOL/hol.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +For hol.thy +Derived rules from Appendix of Mike Gordons HOL Report, Cambridge TR 68 +*) + +open HOL; + + +(** Equality **) + +qed_goal "sym" HOL.thy "s=t ==> t=s" + (fn prems => [cut_facts_tac prems 1, etac subst 1, rtac refl 1]); + +(*calling "standard" reduces maxidx to 0*) +bind_thm ("ssubst", (sym RS subst)); + +qed_goal "trans" HOL.thy "[| r=s; s=t |] ==> r=t" + (fn prems => + [rtac subst 1, resolve_tac prems 1, resolve_tac prems 1]); + +(*Useful with eresolve_tac for proving equalties from known equalities. + a = b + | | + c = d *) +qed_goal "box_equals" HOL.thy + "[| a=b; a=c; b=d |] ==> c=d" + (fn prems=> + [ (rtac trans 1), + (rtac trans 1), + (rtac sym 1), + (REPEAT (resolve_tac prems 1)) ]); + +(** Congruence rules for meta-application **) + +(*similar to AP_THM in Gordon's HOL*) +qed_goal "fun_cong" HOL.thy "(f::'a=>'b) = g ==> f(x)=g(x)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*) +qed_goal "arg_cong" HOL.thy "x=y ==> f(x)=f(y)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +qed_goal "cong" HOL.thy + "[| f = g; (x::'a) = y |] ==> f(x) = g(y)" + (fn [prem1,prem2] => + [rtac (prem1 RS subst) 1, rtac (prem2 RS subst) 1, rtac refl 1]); + +(** Equality of booleans -- iff **) + +qed_goal "iffI" HOL.thy + "[| P ==> Q; Q ==> P |] ==> P=Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI, iff RS mp RS mp]) 1)) ]); + +qed_goal "iffD2" HOL.thy "[| P=Q; Q |] ==> P" + (fn prems => + [rtac ssubst 1, resolve_tac prems 1, resolve_tac prems 1]); + +val iffD1 = sym RS iffD2; + +qed_goal "iffE" HOL.thy + "[| P=Q; [| P --> Q; Q --> P |] ==> R |] ==> R" + (fn [p1,p2] => [REPEAT(ares_tac([p1 RS iffD2, p1 RS iffD1, p2, impI])1)]); + +(** True **) + +qed_goalw "TrueI" HOL.thy [True_def] "True" + (fn _ => [rtac refl 1]); + +qed_goal "eqTrueI " HOL.thy "P ==> P=True" + (fn prems => [REPEAT(resolve_tac ([iffI,TrueI]@prems) 1)]); + +qed_goal "eqTrueE" HOL.thy "P=True ==> P" + (fn prems => [REPEAT(resolve_tac (prems@[TrueI,iffD2]) 1)]); + +(** Universal quantifier **) + +qed_goalw "allI" HOL.thy [All_def] "(!!x::'a. P(x)) ==> !x. P(x)" + (fn prems => [resolve_tac (prems RL [eqTrueI RS ext]) 1]); + +qed_goalw "spec" HOL.thy [All_def] "! x::'a.P(x) ==> P(x)" + (fn prems => [rtac eqTrueE 1, resolve_tac (prems RL [fun_cong]) 1]); + +qed_goal "allE" HOL.thy "[| !x.P(x); P(x) ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (resolve_tac (prems @ [major RS spec]) 1)) ]); + +qed_goal "all_dupE" HOL.thy + "[| ! x.P(x); [| P(x); ! x.P(x) |] ==> R |] ==> R" + (fn prems => + [ (REPEAT (resolve_tac (prems @ (prems RL [spec])) 1)) ]); + + +(** False ** Depends upon spec; it is impossible to do propositional logic + before quantifiers! **) + +qed_goalw "FalseE" HOL.thy [False_def] "False ==> P" + (fn [major] => [rtac (major RS spec) 1]); + +qed_goal "False_neq_True" HOL.thy "False=True ==> P" + (fn [prem] => [rtac (prem RS eqTrueE RS FalseE) 1]); + + +(** Negation **) + +qed_goalw "notI" HOL.thy [not_def] "(P ==> False) ==> ~P" + (fn prems=> [rtac impI 1, eresolve_tac prems 1]); + +qed_goalw "notE" HOL.thy [not_def] "[| ~P; P |] ==> R" + (fn prems => [rtac (prems MRS mp RS FalseE) 1]); + +(** Implication **) + +qed_goal "impE" HOL.thy "[| P-->Q; P; Q ==> R |] ==> R" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +(* Reduces Q to P-->Q, allowing substitution in P. *) +qed_goal "rev_mp" HOL.thy "[| P; P --> Q |] ==> Q" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +qed_goal "contrapos" HOL.thy "[| ~Q; P==>Q |] ==> ~P" + (fn [major,minor]=> + [ (rtac (major RS notE RS notI) 1), + (etac minor 1) ]); + +(* ~(?t = ?s) ==> ~(?s = ?t) *) +val [not_sym] = compose(sym,2,contrapos); + + +(** Existential quantifier **) + +qed_goalw "exI" HOL.thy [Ex_def] "P(x) ==> ? x::'a.P(x)" + (fn prems => [rtac selectI 1, resolve_tac prems 1]); + +qed_goalw "exE" HOL.thy [Ex_def] + "[| ? x::'a.P(x); !!x. P(x) ==> Q |] ==> Q" + (fn prems => [REPEAT(resolve_tac prems 1)]); + + +(** Conjunction **) + +qed_goalw "conjI" HOL.thy [and_def] "[| P; Q |] ==> P&Q" + (fn prems => + [REPEAT (resolve_tac (prems@[allI,impI]) 1 ORELSE etac (mp RS mp) 1)]); + +qed_goalw "conjunct1" HOL.thy [and_def] "[| P & Q |] ==> P" + (fn prems => + [resolve_tac (prems RL [spec] RL [mp]) 1, REPEAT(ares_tac [impI] 1)]); + +qed_goalw "conjunct2" HOL.thy [and_def] "[| P & Q |] ==> Q" + (fn prems => + [resolve_tac (prems RL [spec] RL [mp]) 1, REPEAT(ares_tac [impI] 1)]); + +qed_goal "conjE" HOL.thy "[| P&Q; [| P; Q |] ==> R |] ==> R" + (fn prems => + [cut_facts_tac prems 1, resolve_tac prems 1, + etac conjunct1 1, etac conjunct2 1]); + +(** Disjunction *) + +qed_goalw "disjI1" HOL.thy [or_def] "P ==> P|Q" + (fn [prem] => [REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]); + +qed_goalw "disjI2" HOL.thy [or_def] "Q ==> P|Q" + (fn [prem] => [REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]); + +qed_goalw "disjE" HOL.thy [or_def] "[| P | Q; P ==> R; Q ==> R |] ==> R" + (fn [a1,a2,a3] => + [rtac (mp RS mp) 1, rtac spec 1, rtac a1 1, + rtac (a2 RS impI) 1, assume_tac 1, rtac (a3 RS impI) 1, assume_tac 1]); + +(** CCONTR -- classical logic **) + +qed_goalw "classical" HOL.thy [not_def] "(~P ==> P) ==> P" + (fn [prem] => + [rtac (True_or_False RS (disjE RS eqTrueE)) 1, assume_tac 1, + rtac (impI RS prem RS eqTrueI) 1, + etac subst 1, assume_tac 1]); + +val ccontr = FalseE RS classical; + +(*Double negation law*) +qed_goal "notnotD" HOL.thy "~~P ==> P" + (fn [major]=> + [ (rtac classical 1), (eresolve_tac [major RS notE] 1) ]); + + +(** Unique existence **) + +qed_goalw "ex1I" HOL.thy [Ex1_def] + "[| P(a); !!x. P(x) ==> x=a |] ==> ?! x. P(x)" + (fn prems => + [REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)]); + +qed_goalw "ex1E" HOL.thy [Ex1_def] + "[| ?! x.P(x); !!x. [| P(x); ! y. P(y) --> y=x |] ==> R |] ==> R" + (fn major::prems => + [rtac (major RS exE) 1, REPEAT (etac conjE 1 ORELSE ares_tac prems 1)]); + + +(** Select: Hilbert's Epsilon-operator **) + +(*Easier to apply than selectI: conclusion has only one occurrence of P*) +qed_goal "selectI2" HOL.thy + "[| P(a); !!x. P(x) ==> Q(x) |] ==> Q(@x.P(x))" + (fn prems => [ resolve_tac prems 1, + rtac selectI 1, + resolve_tac prems 1 ]); + +qed_goal "select_equality" HOL.thy + "[| P(a); !!x. P(x) ==> x=a |] ==> (@x.P(x)) = a" + (fn prems => [ rtac selectI2 1, + REPEAT (ares_tac prems 1) ]); + + +(** Classical intro rules for disjunction and existential quantifiers *) + +qed_goal "disjCI" HOL.thy "(~Q ==> P) ==> P|Q" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[disjI1,notI]) 1)), + (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ]); + +qed_goal "excluded_middle" HOL.thy "~P | P" + (fn _ => [ (REPEAT (ares_tac [disjCI] 1)) ]); + +(*For disjunctive case analysis*) +fun excluded_middle_tac sP = + res_inst_tac [("Q",sP)] (excluded_middle RS disjE); + +(*Classical implies (-->) elimination. *) +qed_goal "impCE" HOL.thy "[| P-->Q; ~P ==> R; Q ==> R |] ==> R" + (fn major::prems=> + [ rtac (excluded_middle RS disjE) 1, + REPEAT (DEPTH_SOLVE_1 (ares_tac (prems @ [major RS mp]) 1))]); + +(*Classical <-> elimination. *) +qed_goal "iffCE" HOL.thy + "[| P=Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R" + (fn major::prems => + [ (rtac (major RS iffE) 1), + (REPEAT (DEPTH_SOLVE_1 + (eresolve_tac ([asm_rl,impCE,notE]@prems) 1))) ]); + +qed_goal "exCI" HOL.thy "(! x. ~P(x) ==> P(a)) ==> ? x.P(x)" + (fn prems=> + [ (rtac ccontr 1), + (REPEAT (ares_tac (prems@[exI,allI,notI,notE]) 1)) ]); + + +(* case distinction *) + +qed_goal "case_split_thm" HOL.thy "[| P ==> Q; ~P ==> Q |] ==> Q" + (fn [p1,p2] => [cut_facts_tac [excluded_middle] 1, etac disjE 1, + etac p2 1, etac p1 1]); + +fun case_tac a = res_inst_tac [("P",a)] case_split_thm; + +(** Standard abbreviations **) + +fun stac th = rtac(th RS ssubst); +fun sstac ths = EVERY' (map stac ths); +fun strip_tac i = REPEAT(resolve_tac [impI,allI] i); diff -r 196ca0973a6d -r ff1574a81019 src/HOL/HOL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HOL.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,165 @@ +(* Title: HOL/HOL.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +Higher-Order Logic +*) + +HOL = CPure + + +classes + term < logic + +axclass + plus < term + +axclass + minus < term + +axclass + times < term + +default + term + +types + bool + +arities + fun :: (term, term) term + bool :: term + + +consts + + (* Constants *) + + Trueprop :: "bool => prop" ("(_)" 5) + not :: "bool => bool" ("~ _" [40] 40) + True, False :: "bool" + if :: "[bool, 'a, 'a] => 'a" + Inv :: "('a => 'b) => ('b => 'a)" + + (* Binders *) + + Eps :: "('a => bool) => 'a" (binder "@" 10) + All :: "('a => bool) => bool" (binder "! " 10) + Ex :: "('a => bool) => bool" (binder "? " 10) + Ex1 :: "('a => bool) => bool" (binder "?! " 10) + Let :: "['a, 'a => 'b] => 'b" + + (* Infixes *) + + o :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixr 50) + "=" :: "['a, 'a] => bool" (infixl 50) +(*"~=" :: "['a, 'a] => bool" (infixl 50)*) + "&" :: "[bool, bool] => bool" (infixr 35) + "|" :: "[bool, bool] => bool" (infixr 30) + "-->" :: "[bool, bool] => bool" (infixr 25) + + (* Overloaded Constants *) + + "+" :: "['a::plus, 'a] => 'a" (infixl 65) + "-" :: "['a::minus, 'a] => 'a" (infixl 65) + "*" :: "['a::times, 'a] => 'a" (infixl 70) + + +types + letbinds letbind + case_syn cases_syn + +syntax + + "~=" :: "['a, 'a] => bool" (infixl 50) + + (* Alternative Quantifiers *) + + "*All" :: "[idts, bool] => bool" ("(3ALL _./ _)" 10) + "*Ex" :: "[idts, bool] => bool" ("(3EX _./ _)" 10) + "*Ex1" :: "[idts, bool] => bool" ("(3EX! _./ _)" 10) + + (* Let expressions *) + + "_bind" :: "[idt, 'a] => letbind" ("(2_ =/ _)" 10) + "" :: "letbind => letbinds" ("_") + "_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") + "_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) + + (* Case expressions *) + + "@case" :: "['a, cases_syn] => 'b" ("(case _ of/ _)" 10) + "@case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10) + "" :: "case_syn => cases_syn" ("_") + "@case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _") + +translations + "x ~= y" == "~ (x = y)" + "ALL xs. P" => "! xs. P" + "EX xs. P" => "? xs. P" + "EX! xs. P" => "?! xs. P" + "_Let (_binds b bs) e" == "_Let b (_Let bs e)" + "let x = a in e" == "Let a (%x. e)" + + +rules + + eq_reflection "(x=y) ==> (x==y)" + + (* Basic Rules *) + + refl "t = (t::'a)" + subst "[| s = t; P(s) |] ==> P(t::'a)" + ext "(!!x::'a. (f(x)::'b) = g(x)) ==> (%x.f(x)) = (%x.g(x))" + selectI "P(x::'a) ==> P(@x.P(x))" + + impI "(P ==> Q) ==> P-->Q" + mp "[| P-->Q; P |] ==> Q" + +defs + + True_def "True == ((%x::bool.x)=(%x.x))" + All_def "All(P) == (P = (%x.True))" + Ex_def "Ex(P) == P(@x.P(x))" + False_def "False == (!P.P)" + not_def "~ P == P-->False" + and_def "P & Q == !R. (P-->Q-->R) --> R" + or_def "P | Q == !R. (P-->R) --> (Q-->R) --> R" + Ex1_def "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" + +rules + (* Axioms *) + + iff "(P-->Q) --> (Q-->P) --> (P=Q)" + True_or_False "(P=True) | (P=False)" + +defs + (* Misc Definitions *) + + Let_def "Let s f == f(s)" + Inv_def "Inv(f::'a=>'b) == (% y. @x. f(x)=y)" + 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)" + +end + + +ML + +(** Choice between the HOL and Isabelle style of quantifiers **) + +val HOL_quantifiers = ref true; + +fun alt_ast_tr' (name, alt_name) = + let + fun ast_tr' (*name*) args = + if ! HOL_quantifiers then raise Match + else Syntax.mk_appl (Syntax.Constant alt_name) args; + in + (name, ast_tr') + end; + + +val print_ast_translation = + map alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")]; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Inductive.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Inductive.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,94 @@ +(* Title: HOL/inductive.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +(Co)Inductive Definitions for HOL + +Inductive definitions use least fixedpoints with standard products and sums +Coinductive definitions use greatest fixedpoints with Quine products and sums + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations +*) + +local open Ind_Syntax +in + +fun gen_fp_oper a (X,T,t) = + let val setT = mk_setT T + in Const(a, (setT-->setT)-->setT) $ absfree(X, setT, t) end; + +structure Lfp_items = + struct + val oper = gen_fp_oper "lfp" + val Tarski = def_lfp_Tarski + val induct = def_induct + end; + +structure Gfp_items = + struct + val oper = gen_fp_oper "gfp" + val Tarski = def_gfp_Tarski + val induct = def_Collect_coinduct + end; + +end; + + +functor Ind_section_Fun (Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end) + : sig include INTR_ELIM INDRULE end = +struct +structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and + Fp=Lfp_items); + +structure Indrule = Indrule_Fun + (structure Inductive=Inductive and Intr_elim=Intr_elim); + +open Intr_elim Indrule +end; + + +structure Ind = Add_inductive_def_Fun (Lfp_items); + + +signature INDUCTIVE_STRING = + sig + val thy_name : string (*name of the new theory*) + val srec_tms : string list (*recursion terms*) + val sintrs : string list (*desired introduction rules*) + end; + + +(*For upwards compatibility: can be called directly from ML*) +functor Inductive_Fun + (Inductive: sig include INDUCTIVE_STRING INDUCTIVE_ARG end) + : sig include INTR_ELIM INDRULE end = +Ind_section_Fun + (open Inductive Ind_Syntax + val sign = sign_of thy; + val rec_tms = map (readtm sign termTVar) srec_tms + and intr_tms = map (readtm sign propT) sintrs; + val thy = thy |> Ind.add_fp_def_i(rec_tms, intr_tms) + |> add_thyname thy_name); + + + +signature COINDRULE = + sig + val coinduct : thm + end; + + +functor CoInd_section_Fun + (Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end) + : sig include INTR_ELIM COINDRULE end = +struct +structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and Fp=Gfp_items); + +open Intr_elim +val coinduct = raw_induct +end; + + +structure CoInd = Add_inductive_def_Fun(Gfp_items); diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Inductive.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Inductive.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,1 @@ +Inductive = Gfp + Prod diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Lfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Lfp.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,74 @@ +(* Title: HOL/lfp.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For lfp.thy. The Knaster-Tarski Theorem +*) + +open Lfp; + +(*** Proof of Knaster-Tarski Theorem ***) + +(* lfp(f) is the greatest lower bound of {u. f(u) <= u} *) + +val prems = goalw Lfp.thy [lfp_def] "[| f(A) <= A |] ==> lfp(f) <= A"; +by (rtac (CollectI RS Inter_lower) 1); +by (resolve_tac prems 1); +qed "lfp_lowerbound"; + +val prems = goalw Lfp.thy [lfp_def] + "[| !!u. f(u) <= u ==> A<=u |] ==> A <= lfp(f)"; +by (REPEAT (ares_tac ([Inter_greatest]@prems) 1)); +by (etac CollectD 1); +qed "lfp_greatest"; + +val [mono] = goal Lfp.thy "mono(f) ==> f(lfp(f)) <= lfp(f)"; +by (EVERY1 [rtac lfp_greatest, rtac subset_trans, + rtac (mono RS monoD), rtac lfp_lowerbound, atac, atac]); +qed "lfp_lemma2"; + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) <= f(lfp(f))"; +by (EVERY1 [rtac lfp_lowerbound, rtac (mono RS monoD), + rtac lfp_lemma2, rtac mono]); +qed "lfp_lemma3"; + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))"; +by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1)); +qed "lfp_Tarski"; + + +(*** General induction rule for least fixed points ***) + +val [lfp,mono,indhyp] = goal Lfp.thy + "[| a: lfp(f); mono(f); \ +\ !!x. [| x: f(lfp(f) Int {x.P(x)}) |] ==> P(x) \ +\ |] ==> P(a)"; +by (res_inst_tac [("a","a")] (Int_lower2 RS subsetD RS CollectD) 1); +by (rtac (lfp RSN (2, lfp_lowerbound RS subsetD)) 1); +by (EVERY1 [rtac Int_greatest, rtac subset_trans, + rtac (Int_lower1 RS (mono RS monoD)), + rtac (mono RS lfp_lemma2), + rtac (CollectI RS subsetI), rtac indhyp, atac]); +qed "induct"; + +(** Definition forms of lfp_Tarski and induct, to control unfolding **) + +val [rew,mono] = goal Lfp.thy "[| h==lfp(f); mono(f) |] ==> h = f(h)"; +by (rewtac rew); +by (rtac (mono RS lfp_Tarski) 1); +qed "def_lfp_Tarski"; + +val rew::prems = goal Lfp.thy + "[| A == lfp(f); mono(f); a:A; \ +\ !!x. [| x: f(A Int {x.P(x)}) |] ==> P(x) \ +\ |] ==> P(a)"; +by (EVERY1 [rtac induct, (*backtracking to force correct induction*) + REPEAT1 o (ares_tac (map (rewrite_rule [rew]) prems))]); +qed "def_induct"; + +(*Monotonicity of lfp!*) +val [prem] = goal Lfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> lfp(f) <= lfp(g)"; +br (lfp_lowerbound RS lfp_greatest) 1; +be (prem RS subset_trans) 1; +qed "lfp_mono"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Lfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Lfp.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,14 @@ +(* Title: HOL/lfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The Knaster-Tarski Theorem +*) + +Lfp = mono + +consts lfp :: "['a set=>'a set] => 'a set" +defs + (*least fixed point*) + lfp_def "lfp(f) == Inter({u. f(u) <= u})" +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/List.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/List.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,148 @@ +(* Title: HOL/List + ID: $Id$ + Author: Tobias Nipkow + Copyright 1994 TU Muenchen + +List lemmas +*) + +open List; + +val [Nil_not_Cons,Cons_not_Nil] = list.distinct; + +bind_thm("Cons_neq_Nil", Cons_not_Nil RS notE); +bind_thm("Nil_neq_Cons", sym RS Cons_neq_Nil); + +bind_thm("Cons_inject", (hd list.inject) RS iffD1 RS conjE); + +val list_ss = HOL_ss addsimps list.simps; + +goal List.thy "!x. xs ~= x#xs"; +by (list.induct_tac "xs" 1); +by (ALLGOALS (asm_simp_tac list_ss)); +qed "not_Cons_self"; + +goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)"; +by (list.induct_tac "xs" 1); +by(simp_tac list_ss 1); +by(asm_simp_tac list_ss 1); +by(REPEAT(resolve_tac [exI,refl,conjI] 1)); +qed "neq_Nil_conv"; + +val list_ss = arith_ss addsimps list.simps @ + [null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons, + mem_Nil, mem_Cons, + append_Nil, append_Cons, + map_Nil, map_Cons, + flat_Nil, flat_Cons, + list_all_Nil, list_all_Cons, + filter_Nil, filter_Cons]; + + +(** @ - append **) + +goal List.thy "(xs@ys)@zs = xs@(ys@zs)"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac list_ss)); +qed "append_assoc"; + +goal List.thy "xs @ [] = xs"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac list_ss)); +qed "append_Nil2"; + +goal List.thy "(xs@ys = []) = (xs=[] & ys=[])"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac list_ss)); +qed "append_is_Nil"; + +goal List.thy "(xs @ ys = xs @ zs) = (ys=zs)"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac list_ss)); +qed "same_append_eq"; + + +(** mem **) + +goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); +qed "mem_append"; + +goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); +qed "mem_filter"; + +(** list_all **) + +goal List.thy "(Alls x:xs.True) = True"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac list_ss)); +qed "list_all_True"; + +goal List.thy "list_all p (xs@ys) = (list_all p xs & list_all p ys)"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac list_ss)); +qed "list_all_conj"; + +goal List.thy "(Alls x:xs.P(x)) = (!x. x mem xs --> P(x))"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); +by(fast_tac HOL_cs 1); +qed "list_all_mem_conv"; + + +(** list_case **) + +goal List.thy + "P(list_case a f xs) = ((xs=[] --> P(a)) & \ +\ (!y ys. xs=y#ys --> P(f y ys)))"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac list_ss)); +by(fast_tac HOL_cs 1); +qed "expand_list_case"; + +goal List.thy "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)"; +by(list.induct_tac "xs" 1); +by(fast_tac HOL_cs 1); +by(fast_tac HOL_cs 1); +bind_thm("list_eq_cases", + impI RSN (2,allI RSN (2,allI RSN (2,impI RS (conjI RS (result() RS mp)))))); + +(** flat **) + +goal List.thy "flat(xs@ys) = flat(xs)@flat(ys)"; +by (list.induct_tac "xs" 1); +by(ALLGOALS(asm_simp_tac (list_ss addsimps [append_assoc]))); +qed"flat_append"; + +(** nth **) + +val [nth_0,nth_Suc] = nat_recs nth_def; +store_thm("nth_0",nth_0); +store_thm("nth_Suc",nth_Suc); + +(** Additional mapping lemmas **) + +goal List.thy "map (%x.x) xs = xs"; +by (list.induct_tac "xs" 1); +by (ALLGOALS (asm_simp_tac list_ss)); +qed "map_ident"; + +goal List.thy "map f (xs@ys) = map f xs @ map f ys"; +by (list.induct_tac "xs" 1); +by (ALLGOALS (asm_simp_tac list_ss)); +qed "map_append"; + +goalw List.thy [o_def] "map (f o g) xs = map f (map g xs)"; +by (list.induct_tac "xs" 1); +by (ALLGOALS (asm_simp_tac list_ss)); +qed "map_compose"; + +val list_ss = list_ss addsimps + [not_Cons_self, append_assoc, append_Nil2, append_is_Nil, same_append_eq, + mem_append, mem_filter, + map_ident, map_append, map_compose, + flat_append, list_all_True, list_all_conj, nth_0, nth_Suc]; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/List.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/List.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,83 @@ +(* Title: HOL/List.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1994 TU Muenchen + +Definition of type 'a list as a datatype. This allows primrec to work. + +*) + +List = Arith + + +datatype 'a list = "[]" ("[]") | "#"('a,'a list) (infixr 65) + +consts + + null :: "'a list => bool" + hd :: "'a list => 'a" + tl,ttl :: "'a list => 'a list" + mem :: "['a, 'a list] => bool" (infixl 55) + list_all :: "('a => bool) => ('a list => bool)" + map :: "('a=>'b) => ('a list => 'b list)" + "@" :: "['a list, 'a list] => 'a list" (infixr 65) + filter :: "['a => bool, 'a list] => 'a list" + foldl :: "[['b,'a] => 'b, 'b, 'a list] => 'b" + length :: "'a list => nat" + flat :: "'a list list => 'a list" + nth :: "[nat, 'a list] => 'a" + +syntax + (* list Enumeration *) + "@list" :: "args => 'a list" ("[(_)]") + + (* Special syntax for list_all and filter *) + "@Alls" :: "[idt, 'a list, bool] => bool" ("(2Alls _:_./ _)" 10) + "@filter" :: "[idt, 'a list, bool] => 'a list" ("(1[_:_ ./ _])") + +translations + "[x, xs]" == "x#[xs]" + "[x]" == "x#[]" + + "[x:xs . P]" == "filter (%x.P) xs" + "Alls x:xs.P" == "list_all (%x.P) xs" + +primrec null list + null_Nil "null([]) = True" + null_Cons "null(x#xs) = False" +primrec hd list + hd_Nil "hd([]) = (@x.False)" + hd_Cons "hd(x#xs) = x" +primrec tl list + tl_Nil "tl([]) = (@x.False)" + tl_Cons "tl(x#xs) = xs" +primrec ttl list + (* a "total" version of tl: *) + ttl_Nil "ttl([]) = []" + ttl_Cons "ttl(x#xs) = xs" +primrec "op mem" list + mem_Nil "x mem [] = False" + mem_Cons "x mem (y#ys) = if (y=x) True (x mem ys)" +primrec list_all list + list_all_Nil "list_all P [] = True" + list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)" +primrec map list + map_Nil "map f [] = []" + map_Cons "map f (x#xs) = f(x)#map f xs" +primrec "op @" list + append_Nil "[] @ ys = ys" + append_Cons "(x#xs)@ys = x#(xs@ys)" +primrec filter list + filter_Nil "filter P [] = []" + filter_Cons "filter P (x#xs) = if (P x) (x#filter P xs) (filter P xs)" +primrec foldl list + foldl_Nil "foldl f a [] = a" + foldl_Cons "foldl f a (x#xs) = foldl f (f a x) xs" +primrec length list + length_Nil "length([]) = 0" + length_Cons "length(x#xs) = Suc(length(xs))" +primrec flat list + flat_Nil "flat([]) = []" + flat_Cons "flat(x#xs) = x @ flat(xs)" +defs + nth_def "nth(n) == nat_rec n hd (%m r xs. r(tl(xs)))" +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Makefile Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,114 @@ +######################################################################### +# # +# Makefile for Isabelle (CHOL) # +# # +######################################################################### + +#To make the system, cd to this directory and type +# make -f Makefile +#To make the system and test it on standard examples, type +# make -f Makefile test + +#Environment variable ISABELLECOMP specifies the compiler. +#Environment variable ISABELLEBIN specifies the destination directory. +#For Poly/ML, ISABELLEBIN must begin with a / + +#Makes pure Isabelle (Pure) if this file is ABSENT -- but not +#if it is out of date, since this Makefile does not know its dependencies! + +BIN = $(ISABELLEBIN) +COMP = $(ISABELLECOMP) +THYS = HOL.thy Ord.thy Set.thy Fun.thy subset.thy \ + equalities.thy Prod.thy Trancl.thy Sum.thy WF.thy \ + mono.thy Lfp.thy Gfp.thy Nat.thy Inductive.thy \ + Finite.thy Arith.thy Sexp.thy Univ.thy List.thy + +FILES = ROOT.ML add_ind_def.ML datatype.ML hologic.ML\ + ind_syntax.ML indrule.ML intr_elim.ML simpdata.ML\ + subtype.ML thy_syntax.ML ../Pure/section_utils.ML\ + ../Provers/classical.ML ../Provers/simplifier.ML \ + ../Provers/splitter.ML ../Provers/ind.ML $(THYS) $(THYS:.thy=.ML) + +$(BIN)/CHOL: $(BIN)/Pure $(FILES) + if [ -d $${ISABELLEBIN:?}/Pure ];\ + then echo Bad value for ISABELLEBIN: \ + $(BIN) is the Isabelle source directory; \ + exit 1; \ + fi;\ + case "$(COMP)" in \ + poly*) echo 'make_database"$(BIN)/CHOL"; quit();' \ + | $(COMP) $(BIN)/Pure;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/CHOL ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/CHOL" banner;' | $(BIN)/Pure ;;\ + *) echo Bad value for ISABELLECOMP: \ + $(COMP) is not poly or sml; exit 1;;\ + esac + +$(BIN)/Pure: + cd ../Pure; $(MAKE) + +#### Testing of CHOL + +#A macro referring to the object-logic (depends on ML compiler) +LOGIC:sh=case $ISABELLECOMP in \ + poly*) echo "$ISABELLECOMP $ISABELLEBIN/CHOL" ;;\ + sml*) echo "$ISABELLEBIN/CHOL" ;;\ + *) echo "echo Bad value for ISABELLECOMP: \ + $ISABELLEBIN is not poly or sml; exit 1" ;;\ + esac + +##IMP-semantics example +IMP_THYS = IMP/Com.thy IMP/Denotation.thy IMP/Equiv.thy IMP/Properties.thy +IMP_FILES = IMP/ROOT.ML $(IMP_THYS) $(IMP_THYS:.thy=.ML) + +IMP: $(BIN)/CHOL $(IMP_FILES) + echo 'use"IMP/ROOT.ML";quit();' | $(LOGIC) + +##The integers in CHOL +INTEG_THYS = Integ/Relation.thy Integ/Equiv.thy Integ/Integ.thy + +INTEG_FILES = Integ/ROOT.ML $(INTEG_THYS) $(INTEG_THYS:.thy=.ML) + +Integ: $(BIN)/CHOL $(INTEG_FILES) + echo 'use"Integ/ROOT.ML";quit();' | $(LOGIC) + +##I/O Automata +IOA_THYS = IOA/example/Action.thy IOA/example/Channels.thy\ + IOA/example/Correctness.thy IOA/example/Impl.thy \ + IOA/example/Lemmas.thy IOA/example/Multiset.thy \ + IOA/example/Receiver.thy IOA/example/Sender.thy \ + IOA/meta_theory/Asig.thy IOA/meta_theory/IOA.thy \ + IOA/meta_theory/Option.thy IOA/meta_theory/Solve.thy + +IOA_FILES = IOA/ROOT.ML IOA/example/Packet.thy IOA/example/Spec.thy\ + $(IOA_THYS) $(IOA_THYS:.thy=.ML) + +IOA: $(BIN)/CHOL $(IOA_FILES) + echo 'use"IOA/ROOT.ML";quit();' | $(LOGIC) + +##Properties of substitutions +SUBST_THYS = Subst/AList.thy Subst/Setplus.thy\ + Subst/Subst.thy Subst/Unifier.thy\ + Subst/UTerm.thy Subst/UTLemmas.thy + +SUBST_FILES = Subst/ROOT.ML $(SUBST_THYS) $(SUBST_THYS:.thy=.ML) + +Subst: $(BIN)/CHOL $(SUBST_FILES) + echo 'use"Subst/ROOT.ML";quit();' | $(LOGIC) + +##Miscellaneous examples +EX_THYS = ex/LexProd.thy ex/MT.thy ex/Acc.thy \ + ex/PropLog.thy ex/Puzzle.thy ex/Qsort.thy ex/LList.thy \ + ex/Rec.thy ex/Simult.thy ex/Term.thy ex/String.thy + +EX_FILES = ex/ROOT.ML ex/cla.ML ex/meson.ML ex/mesontest.ML ex/rel.ML \ + ex/set.ML $(EX_THYS) $(EX_THYS:.thy=.ML) + +ex: $(BIN)/CHOL $(EX_FILES) + echo 'use"ex/ROOT.ML";quit();' | $(LOGIC) + +#Full test. +test: $(BIN)/CHOL IMP Integ IOA Subst ex + echo 'Test examples ran successfully' > test + +.PRECIOUS: $(BIN)/Pure $(BIN)/CHOL diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nat.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,436 @@ +(* Title: HOL/nat + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For nat.thy. Type nat is defined as a set (Nat) over the type ind. +*) + +open Nat; + +goal Nat.thy "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 Nat.thy "Zero_Rep: Nat"; +by (rtac (Nat_unfold RS ssubst) 1); +by (rtac (singletonI RS UnI1) 1); +qed "Zero_RepI"; + +val prems = goal Nat.thy "i: Nat ==> Suc_Rep(i) : Nat"; +by (rtac (Nat_unfold RS ssubst) 1); +by (rtac (imageI RS UnI2) 1); +by (resolve_tac prems 1); +qed "Suc_RepI"; + +(*** Induction ***) + +val major::prems = goal Nat.thy + "[| i: Nat; P(Zero_Rep); \ +\ !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |] ==> P(i)"; +by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1); +by (fast_tac (set_cs addIs prems) 1); +qed "Nat_induct"; + +val prems = goalw Nat.thy [Zero_def,Suc_def] + "[| P(0); \ +\ !!k. P(k) ==> P(Suc(k)) |] ==> P(n)"; +by (rtac (Rep_Nat_inverse RS subst) 1); (*types force good instantiation*) +by (rtac (Rep_Nat RS Nat_induct) 1); +by (REPEAT (ares_tac prems 1 + ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1)); +qed "nat_induct"; + +(*Perform induction on n. *) +fun nat_ind_tac a i = + EVERY [res_inst_tac [("n",a)] nat_induct i, + rename_last_tac a ["1"] (i+1)]; + +(*A special form of induction for reasoning about m P (Suc x) (Suc y) \ +\ |] ==> P m n"; +by (res_inst_tac [("x","m")] spec 1); +by (nat_ind_tac "n" 1); +by (rtac allI 2); +by (nat_ind_tac "x" 2); +by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1)); +qed "diff_induct"; + +(*Case analysis on the natural numbers*) +val prems = goal Nat.thy + "[| n=0 ==> P; !!x. n = Suc(x) ==> P |] ==> P"; +by (subgoal_tac "n=0 | (EX x. n = Suc(x))" 1); +by (fast_tac (HOL_cs addSEs prems) 1); +by (nat_ind_tac "n" 1); +by (rtac (refl RS disjI1) 1); +by (fast_tac HOL_cs 1); +qed "natE"; + +(*** Isomorphisms: Abs_Nat and Rep_Nat ***) + +(*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 Nat.thy "inj(Rep_Nat)"; +by (rtac inj_inverseI 1); +by (rtac Rep_Nat_inverse 1); +qed "inj_Rep_Nat"; + +goal Nat.thy "inj_onto Abs_Nat Nat"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Nat_inverse 1); +qed "inj_onto_Abs_Nat"; + +(*** Distinctness of constructors ***) + +goalw Nat.thy [Zero_def,Suc_def] "Suc(m) ~= 0"; +by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1); +by (rtac Suc_Rep_not_Zero_Rep 1); +by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1)); +qed "Suc_not_Zero"; + +bind_thm ("Zero_not_Suc", (Suc_not_Zero RS not_sym)); + +bind_thm ("Suc_neq_Zero", (Suc_not_Zero RS notE)); +val Zero_neq_Suc = sym RS Suc_neq_Zero; + +(** Injectiveness of Suc **) + +goalw Nat.thy [Suc_def] "inj(Suc)"; +by (rtac injI 1); +by (dtac (inj_onto_Abs_Nat RS inj_ontoD) 1); +by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1)); +by (dtac (inj_Suc_Rep RS injD) 1); +by (etac (inj_Rep_Nat RS injD) 1); +qed "inj_Suc"; + +val Suc_inject = inj_Suc RS injD;; + +goal Nat.thy "(Suc(m)=Suc(n)) = (m=n)"; +by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]); +qed "Suc_Suc_eq"; + +goal Nat.thy "n ~= Suc(n)"; +by (nat_ind_tac "n" 1); +by (ALLGOALS(asm_simp_tac (HOL_ss addsimps [Zero_not_Suc,Suc_Suc_eq]))); +qed "n_not_Suc_n"; + +val Suc_n_not_n = n_not_Suc_n RS not_sym; + +(*** nat_case -- the selection operator for nat ***) + +goalw Nat.thy [nat_case_def] "nat_case a f 0 = a"; +by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1); +qed "nat_case_0"; + +goalw Nat.thy [nat_case_def] "nat_case a f (Suc k) = f(k)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim Suc_inject, Suc_neq_Zero]) 1); +qed "nat_case_Suc"; + +(** Introduction rules for 'pred_nat' **) + +goalw Nat.thy [pred_nat_def] " : pred_nat"; +by (fast_tac set_cs 1); +qed "pred_natI"; + +val major::prems = goalw Nat.thy [pred_nat_def] + "[| p : pred_nat; !!x n. [| p = |] ==> R \ +\ |] ==> R"; +by (rtac (major RS CollectE) 1); +by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1)); +qed "pred_natE"; + +goalw Nat.thy [wf_def] "wf(pred_nat)"; +by (strip_tac 1); +by (nat_ind_tac "x" 1); +by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, + make_elim Suc_inject]) 2); +by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, Zero_neq_Suc]) 1); +qed "wf_pred_nat"; + + +(*** nat_rec -- by wf recursion on pred_nat ***) + +bind_thm ("nat_rec_unfold", (wf_pred_nat RS (nat_rec_def RS def_wfrec))); + +(** conversion rules **) + +goal Nat.thy "nat_rec 0 c h = c"; +by (rtac (nat_rec_unfold RS trans) 1); +by (simp_tac (HOL_ss addsimps [nat_case_0]) 1); +qed "nat_rec_0"; + +goal Nat.thy "nat_rec (Suc n) c h = h n (nat_rec n c h)"; +by (rtac (nat_rec_unfold RS trans) 1); +by (simp_tac (HOL_ss addsimps [nat_case_Suc, pred_natI, cut_apply]) 1); +qed "nat_rec_Suc"; + +(*These 2 rules ease the use of primitive recursion. NOTE USE OF == *) +val [rew] = goal Nat.thy + "[| !!n. f(n) == nat_rec n c h |] ==> f(0) = c"; +by (rewtac rew); +by (rtac nat_rec_0 1); +qed "def_nat_rec_0"; + +val [rew] = goal Nat.thy + "[| !!n. f(n) == nat_rec n c h |] ==> f(Suc(n)) = h n (f n)"; +by (rewtac rew); +by (rtac nat_rec_Suc 1); +qed "def_nat_rec_Suc"; + +fun nat_recs def = + [standard (def RS def_nat_rec_0), + standard (def RS def_nat_rec_Suc)]; + + +(*** Basic properties of "less than" ***) + +(** Introduction properties **) + +val prems = goalw Nat.thy [less_def] "[| i i<(k::nat)"; +by (rtac (trans_trancl RS transD) 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +qed "less_trans"; + +goalw Nat.thy [less_def] "n < Suc(n)"; +by (rtac (pred_natI RS r_into_trancl) 1); +qed "lessI"; + +(* i i ~ m<(n::nat)"; +by(fast_tac (HOL_cs addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1); +qed "less_not_sym"; + +(* [| n R *) +bind_thm ("less_asym", (less_not_sym RS notE)); + +goalw Nat.thy [less_def] "~ n<(n::nat)"; +by (rtac notI 1); +by (etac (wf_pred_nat RS wf_trancl RS wf_anti_refl) 1); +qed "less_not_refl"; + +(* n R *) +bind_thm ("less_anti_refl", (less_not_refl RS notE)); + +goal Nat.thy "!!m. n m ~= (n::nat)"; +by(fast_tac (HOL_cs addEs [less_anti_refl]) 1); +qed "less_not_refl2"; + + +val major::prems = goalw Nat.thy [less_def] + "[| i P; !!j. [| i P \ +\ |] ==> P"; +by (rtac (major RS tranclE) 1); +by (fast_tac (HOL_cs addSEs (prems@[pred_natE, Pair_inject])) 1); +by (fast_tac (HOL_cs addSEs (prems@[pred_natE, Pair_inject])) 1); +qed "lessE"; + +goal Nat.thy "~ n<0"; +by (rtac notI 1); +by (etac lessE 1); +by (etac Zero_neq_Suc 1); +by (etac Zero_neq_Suc 1); +qed "not_less0"; + +(* n<0 ==> R *) +bind_thm ("less_zeroE", (not_less0 RS notE)); + +val [major,less,eq] = goal Nat.thy + "[| m < Suc(n); m P; m=n ==> P |] ==> P"; +by (rtac (major RS lessE) 1); +by (rtac eq 1); +by (fast_tac (HOL_cs addSDs [Suc_inject]) 1); +by (rtac less 1); +by (fast_tac (HOL_cs addSDs [Suc_inject]) 1); +qed "less_SucE"; + +goal Nat.thy "(m < Suc(n)) = (m < n | m = n)"; +by (fast_tac (HOL_cs addSIs [lessI] + addEs [less_trans, less_SucE]) 1); +qed "less_Suc_eq"; + + +(** Inductive (?) properties **) + +val [prem] = goal Nat.thy "Suc(m) < n ==> m P \ +\ |] ==> P"; +by (rtac (major RS lessE) 1); +by (etac (lessI RS minor) 1); +by (etac (Suc_lessD RS minor) 1); +by (assume_tac 1); +qed "Suc_lessE"; + +val [major] = goal Nat.thy "Suc(m) < Suc(n) ==> m Suc(m) < Suc(n)"; +by (subgoal_tac "m Suc(m) < Suc(n)" 1); +by (fast_tac (HOL_cs addIs prems) 1); +by (nat_ind_tac "n" 1); +by (rtac impI 1); +by (etac less_zeroE 1); +by (fast_tac (HOL_cs addSIs [lessI] + addSDs [Suc_inject] + addEs [less_trans, lessE]) 1); +qed "Suc_mono"; + +goal Nat.thy "(Suc(m) < Suc(n)) = (m P(m) |] ==> P(n) |] ==> P(n)"; +by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); +by (eresolve_tac prems 1); +qed "less_induct"; + + +(*** Properties of <= ***) + +goalw Nat.thy [le_def] "0 <= n"; +by (rtac not_less0 1); +qed "le0"; + +val nat_simps = [not_less0, less_not_refl, zero_less_Suc, lessI, + Suc_less_eq, less_Suc_eq, le0, not_Suc_n_less_n, + Suc_not_Zero, Zero_not_Suc, Suc_Suc_eq, + n_not_Suc_n, Suc_n_not_n, + nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; + +val nat_ss0 = sum_ss addsimps nat_simps; + +(*Prevents simplification of f and g: much faster*) +qed_goal "nat_case_weak_cong" Nat.thy + "m=n ==> nat_case a f m = nat_case a f n" + (fn [prem] => [rtac (prem RS arg_cong) 1]); + +qed_goal "nat_rec_weak_cong" Nat.thy + "m=n ==> nat_rec m a f = nat_rec n a f" + (fn [prem] => [rtac (prem RS arg_cong) 1]); + +val prems = goalw Nat.thy [le_def] "~(n m<=(n::nat)"; +by (resolve_tac prems 1); +qed "leI"; + +val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n<(m::nat))"; +by (resolve_tac prems 1); +qed "leD"; + +val leE = make_elim leD; + +goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)"; +by (fast_tac HOL_cs 1); +qed "not_leE"; + +goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n"; +by(simp_tac nat_ss0 1); +by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); +qed "lessD"; + +goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n"; +by(asm_full_simp_tac nat_ss0 1); +by(fast_tac HOL_cs 1); +qed "Suc_leD"; + +goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)"; +by (fast_tac (HOL_cs addEs [less_asym]) 1); +qed "less_imp_le"; + +goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)"; +by (cut_facts_tac [less_linear] 1); +by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); +qed "le_imp_less_or_eq"; + +goalw Nat.thy [le_def] "!!m. m m <=(n::nat)"; +by (cut_facts_tac [less_linear] 1); +by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1); +by (flexflex_tac); +qed "less_or_eq_imp_le"; + +goal Nat.thy "(x <= (y::nat)) = (x < y | x=y)"; +by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1)); +qed "le_eq_less_or_eq"; + +goal Nat.thy "n <= (n::nat)"; +by(simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1); +qed "le_refl"; + +val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)"; +by (dtac le_imp_less_or_eq 1); +by (fast_tac (HOL_cs addIs [less_trans]) 1); +qed "le_less_trans"; + +goal Nat.thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)"; +by (dtac le_imp_less_or_eq 1); +by (fast_tac (HOL_cs addIs [less_trans]) 1); +qed "less_le_trans"; + +goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)"; +by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq, + rtac less_or_eq_imp_le, fast_tac (HOL_cs addIs [less_trans])]); +qed "le_trans"; + +val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)"; +by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq, + fast_tac (HOL_cs addEs [less_anti_refl,less_asym])]); +qed "le_anti_sym"; + +goal Nat.thy "(Suc(n) <= Suc(m)) = (n <= m)"; +by (simp_tac (nat_ss0 addsimps [le_eq_less_or_eq]) 1); +qed "Suc_le_mono"; + +val nat_ss = nat_ss0 addsimps [le_refl]; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nat.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,70 @@ +(* Title: HOL/Nat.thy + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Definition of types ind and nat. + +Type nat is defined as a set Nat over type ind. +*) + +Nat = WF + + +(** type ind **) + +types + ind + +arities + ind :: term + +consts + Zero_Rep :: "ind" + Suc_Rep :: "ind => ind" + +rules + (*the axiom of infinity in 2 parts*) + inj_Suc_Rep "inj(Suc_Rep)" + Suc_Rep_not_Zero_Rep "Suc_Rep(x) ~= Zero_Rep" + + + +(** type nat **) + +(* type definition *) + +subtype (Nat) + nat = "lfp(%X. {Zero_Rep} Un (Suc_Rep``X))" (lfp_def) + +instance + nat :: ord + + +(* abstract constants and syntax *) + +consts + "0" :: "nat" ("0") + Suc :: "nat => nat" + nat_case :: "['a, nat => 'a, nat] => 'a" + pred_nat :: "(nat * nat) set" + nat_rec :: "[nat, 'a, [nat, 'a] => 'a] => 'a" + +translations + "case p of 0 => a | Suc(y) => b" == "nat_case a (%y.b) p" + +defs + Zero_def "0 == Abs_Nat(Zero_Rep)" + Suc_def "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))" + + (*nat operations and recursion*) + nat_case_def "nat_case a f n == @z. (n=0 --> z=a) \ +\ & (!x. n=Suc(x) --> z=f(x))" + pred_nat_def "pred_nat == {p. ? n. p = }" + + less_def "m:trancl(pred_nat)" + + le_def "m<=(n::nat) == ~(n f(A) <= f(B) |] ==> mono(f)"; +by (REPEAT (ares_tac [allI, impI, prem] 1)); +qed "monoI"; + +val [major,minor] = goalw Ord.thy [mono_def] + "[| mono(f); A <= B |] ==> f(A) <= f(B)"; +by (rtac (major RS spec RS spec RS mp) 1); +by (rtac minor 1); +qed "monoD"; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Ord.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Ord.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,25 @@ +(* Title: HOL/Ord.thy + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +The type class for ordered types (* FIXME improve comment *) +*) + +Ord = HOL + + +axclass + ord < term + +consts + "<", "<=" :: "['a::ord, 'a] => bool" (infixl 50) + mono :: "['a::ord => 'b::ord] => bool" (*monotonicity*) + min, max :: "['a::ord, 'a] => 'a" + +defs + mono_def "mono(f) == (!A B. A <= B --> f(A) <= f(B))" + min_def "min a b == if (a <= b) a b" + max_def "max a b == if (a <= b) b a" + +end + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Prod.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Prod.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,237 @@ +(* Title: HOL/prod + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For prod.thy. Ordered Pairs, the Cartesian product type, the unit type +*) + +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"; +by (EVERY1 [rtac CollectI, rtac exI, rtac exI, rtac refl]); +qed "ProdI"; + +val [major] = goalw Prod.thy [Pair_Rep_def] + "Pair_Rep a b = Pair_Rep a' b' ==> a=a' & b=b'"; +by (EVERY1 [rtac (major RS fun_cong RS fun_cong RS subst), + rtac conjI, rtac refl, rtac refl]); +qed "Pair_Rep_inject"; + +goal Prod.thy "inj_onto Abs_Prod Prod"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Prod_inverse 1); +qed "inj_onto_Abs_Prod"; + +val prems = goalw Prod.thy [Pair_def] + "[| = ; [| a=a'; b=b' |] ==> R |] ==> R"; +by (rtac (inj_onto_Abs_Prod RS inj_ontoD RS Pair_Rep_inject RS conjE) 1); +by (REPEAT (ares_tac (prems@[ProdI]) 1)); +qed "Pair_inject"; + +goal Prod.thy "( = ) = (a=a' & b=b')"; +by (fast_tac (set_cs addIs [Pair_inject]) 1); +qed "Pair_eq"; + +goalw Prod.thy [fst_def] "fst() = a"; +by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); +qed "fst_conv"; + +goalw Prod.thy [snd_def] "snd() = b"; +by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); +qed "snd_conv"; + +goalw Prod.thy [Pair_def] "? x y. p = "; +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]); +qed "PairE_lemma"; + +val [prem] = goal Prod.thy "[| !!x y. p = ==> Q |] ==> Q"; +by (rtac (PairE_lemma RS exE) 1); +by (REPEAT (eresolve_tac [prem,exE] 1)); +qed "PairE"; + +goalw Prod.thy [split_def] "split c = c a b"; +by (sstac [fst_conv, snd_conv] 1); +by (rtac refl 1); +qed "split"; + +val prod_ss = set_ss addsimps [fst_conv, snd_conv, split, Pair_eq]; + +goal Prod.thy "(s=t) = (fst(s)=fst(t) & snd(s)=snd(t))"; +by (res_inst_tac[("p","s")] PairE 1); +by (res_inst_tac[("p","t")] PairE 1); +by (asm_simp_tac prod_ss 1); +qed "Pair_fst_snd_eq"; + +(*Prevents simplification of c: much faster*) +qed_goal "split_weak_cong" Prod.thy + "p=q ==> split c p = split c q" + (fn [prem] => [rtac (prem RS arg_cong) 1]); + +(* Do not add as rewrite rule: invalidates some proofs in IMP *) +goal Prod.thy "p = "; +by (res_inst_tac [("p","p")] PairE 1); +by (asm_simp_tac prod_ss 1); +qed "surjective_pairing"; + +goal Prod.thy "p = split (%x y.) p"; +by (res_inst_tac [("p","p")] PairE 1); +by (asm_simp_tac prod_ss 1); +qed "surjective_pairing2"; + +(*For use with split_tac and the simplifier*) +goal Prod.thy "R(split c p) = (! x y. p = --> R(c x y))"; +by (stac surjective_pairing 1); +by (stac split 1); +by (fast_tac (HOL_cs addSEs [Pair_inject]) 1); +qed "expand_split"; + +(** split used as a logical connective or set former **) + +(*These rules are for use with fast_tac. + Could instead call simp_tac/asm_full_simp_tac using split as rewrite.*) + +goal Prod.thy "!!a b c. c a b ==> split c "; +by (asm_simp_tac prod_ss 1); +qed "splitI"; + +val prems = goalw Prod.thy [split_def] + "[| split c p; !!x y. [| p = ; c x y |] ==> Q |] ==> Q"; +by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); +qed "splitE"; + +goal Prod.thy "!!R a b. split R ==> R a b"; +by (etac (split RS iffD1) 1); +qed "splitD"; + +goal Prod.thy "!!a b c. z: c a b ==> z: split c "; +by (asm_simp_tac prod_ss 1); +qed "mem_splitI"; + +val prems = goalw Prod.thy [split_def] + "[| z: split c p; !!x y. [| p = ; z: c x y |] ==> Q |] ==> Q"; +by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); +qed "mem_splitE"; + +(*** prod_fun -- action of the product functor upon functions ***) + +goalw Prod.thy [prod_fun_def] "prod_fun f g = "; +by (rtac split 1); +qed "prod_fun"; + +goal Prod.thy + "prod_fun (f1 o f2) (g1 o g2) = ((prod_fun f1 g1) o (prod_fun f2 g2))"; +by (rtac ext 1); +by (res_inst_tac [("p","x")] PairE 1); +by (asm_simp_tac (prod_ss addsimps [prod_fun,o_def]) 1); +qed "prod_fun_compose"; + +goal Prod.thy "prod_fun (%x.x) (%y.y) = (%z.z)"; +by (rtac ext 1); +by (res_inst_tac [("p","z")] PairE 1); +by (asm_simp_tac (prod_ss addsimps [prod_fun]) 1); +qed "prod_fun_ident"; + +val prems = goal Prod.thy ":r ==> : (prod_fun f g)``r"; +by (rtac image_eqI 1); +by (rtac (prod_fun RS sym) 1); +by (resolve_tac prems 1); +qed "prod_fun_imageI"; + +val major::prems = goal Prod.thy + "[| c: (prod_fun f g)``r; !!x y. [| c=; :r |] ==> P \ +\ |] ==> P"; +by (rtac (major RS imageE) 1); +by (res_inst_tac [("p","x")] PairE 1); +by (resolve_tac prems 1); +by (fast_tac HOL_cs 2); +by (fast_tac (HOL_cs addIs [prod_fun]) 1); +qed "prod_fun_imageE"; + +(*** Disjoint union of a family of sets - Sigma ***) + +qed_goalw "SigmaI" Prod.thy [Sigma_def] + "[| a:A; b:B(a) |] ==> : Sigma A B" + (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]); + +(*The general elimination rule*) +qed_goalw "SigmaE" Prod.thy [Sigma_def] + "[| c: Sigma A B; \ +\ !!x y.[| x:A; y:B(x); c= |] ==> P \ +\ |] ==> P" + (fn major::prems=> + [ (cut_facts_tac [major] 1), + (REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]); + +(** Elimination of :A*B -- introduces no eigenvariables **) +qed_goal "SigmaD1" Prod.thy " : Sigma A B ==> a : A" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +qed_goal "SigmaD2" Prod.thy " : Sigma A B ==> b : B(a)" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +qed_goal "SigmaE2" Prod.thy + "[| : Sigma A B; \ +\ [| a:A; b:B(a) |] ==> P \ +\ |] ==> P" + (fn [major,minor]=> + [ (rtac minor 1), + (rtac (major RS SigmaD1) 1), + (rtac (major RS SigmaD2) 1) ]); + +(*** Domain of a relation ***) + +val prems = goalw Prod.thy [image_def] " : r ==> a : fst``r"; +by (rtac CollectI 1); +by (rtac bexI 1); +by (rtac (fst_conv RS sym) 1); +by (resolve_tac prems 1); +qed "fst_imageI"; + +val major::prems = goal Prod.thy + "[| a : fst``r; !!y.[| : r |] ==> P |] ==> P"; +by (rtac (major RS imageE) 1); +by (resolve_tac prems 1); +by (etac ssubst 1); +by (rtac (surjective_pairing RS subst) 1); +by (assume_tac 1); +qed "fst_imageE"; + +(*** Range of a relation ***) + +val prems = goalw Prod.thy [image_def] " : r ==> b : snd``r"; +by (rtac CollectI 1); +by (rtac bexI 1); +by (rtac (snd_conv RS sym) 1); +by (resolve_tac prems 1); +qed "snd_imageI"; + +val major::prems = goal Prod.thy + "[| a : snd``r; !!y.[| : r |] ==> P |] ==> P"; +by (rtac (major RS imageE) 1); +by (resolve_tac prems 1); +by (etac ssubst 1); +by (rtac (surjective_pairing RS subst) 1); +by (assume_tac 1); +qed "snd_imageE"; + +(** Exhaustion rule for unit -- a degenerate form of induction **) + +goalw Prod.thy [Unity_def] + "u = Unity"; +by (stac (rewrite_rule [Unit_def] Rep_Unit RS CollectD RS sym) 1); +by (rtac (Rep_Unit_inverse RS sym) 1); +qed "unit_eq"; + +val prod_cs = set_cs addSIs [SigmaI, mem_splitI] + addIs [fst_imageI, snd_imageI, prod_fun_imageI] + addSEs [SigmaE2, SigmaE, mem_splitE, + fst_imageE, snd_imageE, prod_fun_imageE, + Pair_inject]; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Prod.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Prod.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,66 @@ +(* Title: HOL/Prod.thy + ID: Prod.thy,v 1.5 1994/08/19 09:04:27 lcp Exp + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ordered Pairs and the Cartesian product type. +The unit type. +*) + +Prod = Fun + + +(** Products **) + +(* type definition *) + +consts + Pair_Rep :: "['a, 'b] => ['a, 'b] => bool" + +defs + Pair_Rep_def "Pair_Rep == (%a b. %x y. x=a & y=b)" + +subtype (Prod) + ('a, 'b) "*" (infixr 20) + = "{f. ? a b. f = Pair_Rep (a::'a) (b::'b)}" + + +(* abstract constants and syntax *) + +consts + fst :: "'a * 'b => 'a" + snd :: "'a * 'b => 'b" + split :: "[['a, 'b] => 'c, 'a * 'b] => 'c" + prod_fun :: "['a => 'b, 'c => 'd, 'a * 'c] => 'b * 'd" + Pair :: "['a, 'b] => 'a * 'b" + Sigma :: "['a set, 'a => 'b set] => ('a * 'b) set" + +syntax + "@Tuple" :: "args => 'a * 'b" ("(1<_>)") + +translations + "" == ">" + "" == "Pair x y" + "" => "x" + +defs + Pair_def "Pair a b == Abs_Prod(Pair_Rep a b)" + fst_def "fst(p) == @a. ? b. p = " + snd_def "snd(p) == @b. ? a. p = " + split_def "split c p == c (fst p) (snd p)" + prod_fun_def "prod_fun f g == split(%x y.)" + Sigma_def "Sigma A B == UN x:A. UN y:B(x). {}" + + + +(** Unit **) + +subtype (Unit) + unit = "{p. p = True}" + +consts + Unity :: "unit" ("<>") + +defs + Unity_def "Unity == Abs_Unit(True)" + +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/README Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,23 @@ + CHOL: Higher-Order Logic with curried functions + +This directory contains the Standard ML sources of the Isabelle system for +Higher-Order Logic with curried functions. Important files include + +ROOT.ML -- loads all source files. Enter an ML image containing Pure +Isabelle and type: use "ROOT.ML"; + +Makefile -- compiles the files under Poly/ML or SML of New Jersey + +ex -- subdirectory containing examples. To execute them, enter an ML image +containing CHOL and type: use "ex/ROOT.ML"; + +Subst -- subdirectory defining a theory of substitution and unification. + +Useful references on Higher-Order Logic: + + P. B. Andrews, + An Introduction to Mathematical Logic and Type Theory + (Academic Press, 1986). + + J. Lambek and P. J. Scott, + Introduction to Higher Order Categorical Logic (CUP, 1986) diff -r 196ca0973a6d -r ff1574a81019 src/HOL/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ROOT.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,87 @@ +(* Title: CHOL/ROOT.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +Adds Classical Higher-order Logic to a database containing Pure Isabelle. +Should be executed in the subdirectory HOL. +*) + +val banner = "Higher-Order Logic with curried functions"; +writeln banner; + +print_depth 1; +set eta_contract; + +(* Add user sections *) +use "../Pure/section_utils.ML"; +use "thy_syntax.ML"; + +use_thy "HOL"; +use "../Provers/hypsubst.ML"; +use "../Provers/classical.ML"; +use "../Provers/simplifier.ML"; +use "../Provers/splitter.ML"; + +(** Applying HypsubstFun to generate hyp_subst_tac **) + +structure Hypsubst_Data = + struct + (*Take apart an equality judgement; otherwise raise Match!*) + fun dest_eq (Const("Trueprop",_) $ (Const("op =",_) $ t $ u)) = (t,u); + val imp_intr = impI + val rev_mp = rev_mp + val subst = subst + val sym = sym + end; + +structure Hypsubst = HypsubstFun(Hypsubst_Data); +open Hypsubst; + +(*** Applying ClassicalFun to create a classical prover ***) +structure Classical_Data = + struct + val sizef = size_of_thm + val mp = mp + val not_elim = notE + val classical = classical + val hyp_subst_tacs=[hyp_subst_tac] + end; + +structure Classical = ClassicalFun(Classical_Data); +open Classical; + +(*Propositional rules*) +val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI] + addSEs [conjE,disjE,impCE,FalseE,iffE]; + +(*Quantifier rules*) +val HOL_cs = prop_cs addSIs [allI] addIs [exI,ex1I] + addSEs [exE,ex1E] addEs [allE]; + +use "simpdata.ML"; +use_thy "Ord"; +use_thy "subset"; +use_thy "equalities"; +use "hologic.ML"; +use "subtype.ML"; +use_thy "Prod"; +use_thy "Sum"; +use_thy "Gfp"; +use_thy "Nat"; + +use "datatype.ML"; +use "ind_syntax.ML"; +use "add_ind_def.ML"; +use "intr_elim.ML"; +use "indrule.ML"; +use_thy "Inductive"; + +use_thy "Finite"; +use_thy "Sexp"; +use_thy "List"; + +init_pps (); +print_depth 8; + +val CHOL_build_completed = (); (*indicate successful build*) diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Set.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,447 @@ +(* Title: HOL/set + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For set.thy. Set theory for higher-order logic. A set is simply a predicate. +*) + +open Set; + +val [prem] = goal Set.thy "[| P(a) |] ==> a : {x.P(x)}"; +by (rtac (mem_Collect_eq RS ssubst) 1); +by (rtac prem 1); +qed "CollectI"; + +val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)"; +by (resolve_tac (prems RL [mem_Collect_eq RS subst]) 1); +qed "CollectD"; + +val [prem] = goal Set.thy "[| !!x. (x:A) = (x:B) |] ==> A = B"; +by (rtac (prem RS ext RS arg_cong RS box_equals) 1); +by (rtac Collect_mem_eq 1); +by (rtac Collect_mem_eq 1); +qed "set_ext"; + +val [prem] = goal Set.thy "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}"; +by (rtac (prem RS ext RS arg_cong) 1); +qed "Collect_cong"; + +val CollectE = make_elim CollectD; + +(*** Bounded quantifiers ***) + +val prems = goalw Set.thy [Ball_def] + "[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)"; +by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); +qed "ballI"; + +val [major,minor] = goalw Set.thy [Ball_def] + "[| ! x:A. P(x); x:A |] ==> P(x)"; +by (rtac (minor RS (major RS spec RS mp)) 1); +qed "bspec"; + +val major::prems = goalw Set.thy [Ball_def] + "[| ! x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q"; +by (rtac (major RS spec RS impCE) 1); +by (REPEAT (eresolve_tac prems 1)); +qed "ballE"; + +(*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*) +fun ball_tac i = etac ballE i THEN contr_tac (i+1); + +val prems = goalw Set.thy [Bex_def] + "[| P(x); x:A |] ==> ? x:A. P(x)"; +by (REPEAT (ares_tac (prems @ [exI,conjI]) 1)); +qed "bexI"; + +qed_goal "bexCI" Set.thy + "[| ! x:A. ~P(x) ==> P(a); a:A |] ==> ? x:A.P(x)" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); + +val major::prems = goalw Set.thy [Bex_def] + "[| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q"; +by (rtac (major RS exE) 1); +by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)); +qed "bexE"; + +(*Trival rewrite rule; (! x:A.P)=P holds only if A is nonempty!*) +val prems = goal Set.thy + "(! x:A. True) = True"; +by (REPEAT (ares_tac [TrueI,ballI,iffI] 1)); +qed "ball_rew"; + +(** Congruence rules **) + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ +\ (! x:A. P(x)) = (! x:B. Q(x))"; +by (resolve_tac (prems RL [ssubst]) 1); +by (REPEAT (ares_tac [ballI,iffI] 1 + ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1)); +qed "ball_cong"; + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ +\ (? x:A. P(x)) = (? x:B. Q(x))"; +by (resolve_tac (prems RL [ssubst]) 1); +by (REPEAT (etac bexE 1 + ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1)); +qed "bex_cong"; + +(*** Subsets ***) + +val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B"; +by (REPEAT (ares_tac (prems @ [ballI]) 1)); +qed "subsetI"; + +(*Rule in Modus Ponens style*) +val major::prems = goalw Set.thy [subset_def] "[| A <= B; c:A |] ==> c:B"; +by (rtac (major RS bspec) 1); +by (resolve_tac prems 1); +qed "subsetD"; + +(*The same, with reversed premises for use with etac -- cf rev_mp*) +qed_goal "rev_subsetD" Set.thy "[| c:A; A <= B |] ==> c:B" + (fn prems=> [ (REPEAT (resolve_tac (prems@[subsetD]) 1)) ]); + +(*Classical elimination rule*) +val major::prems = goalw Set.thy [subset_def] + "[| A <= B; c~:A ==> P; c:B ==> P |] ==> P"; +by (rtac (major RS ballE) 1); +by (REPEAT (eresolve_tac prems 1)); +qed "subsetCE"; + +(*Takes assumptions A<=B; c:A and creates the assumption c:B *) +fun set_mp_tac i = etac subsetCE i THEN mp_tac i; + +qed_goal "subset_refl" Set.thy "A <= (A::'a set)" + (fn _=> [ (REPEAT (ares_tac [subsetI] 1)) ]); + +val prems = goal Set.thy "[| A<=B; B<=C |] ==> A<=(C::'a set)"; +by (cut_facts_tac prems 1); +by (REPEAT (ares_tac [subsetI] 1 ORELSE set_mp_tac 1)); +qed "subset_trans"; + + +(*** Equality ***) + +(*Anti-symmetry of the subset relation*) +val prems = goal Set.thy "[| A <= B; B <= A |] ==> A = (B::'a set)"; +by (rtac (iffI RS set_ext) 1); +by (REPEAT (ares_tac (prems RL [subsetD]) 1)); +qed "subset_antisym"; +val equalityI = subset_antisym; + +(* Equality rules from ZF set theory -- are they appropriate here? *) +val prems = goal Set.thy "A = B ==> A<=(B::'a set)"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +qed "equalityD1"; + +val prems = goal Set.thy "A = B ==> B<=(A::'a set)"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +qed "equalityD2"; + +val prems = goal Set.thy + "[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P"; +by (resolve_tac prems 1); +by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1)); +qed "equalityE"; + +val major::prems = goal Set.thy + "[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P"; +by (rtac (major RS equalityE) 1); +by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)); +qed "equalityCE"; + +(*Lemma for creating induction formulae -- for "pattern matching" on p + To make the induction hypotheses usable, apply "spec" or "bspec" to + put universal quantifiers over the free variables in p. *) +val prems = goal Set.thy + "[| p:A; !!z. z:A ==> p=z --> R |] ==> R"; +by (rtac mp 1); +by (REPEAT (resolve_tac (refl::prems) 1)); +qed "setup_induction"; + + +(*** Set complement -- Compl ***) + +val prems = goalw Set.thy [Compl_def] + "[| c:A ==> False |] ==> c : Compl(A)"; +by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); +qed "ComplI"; + +(*This form, with negated conclusion, works well with the Classical prover. + Negated assumptions behave like formulae on the right side of the notional + turnstile...*) +val major::prems = goalw Set.thy [Compl_def] + "[| c : Compl(A) |] ==> c~:A"; +by (rtac (major RS CollectD) 1); +qed "ComplD"; + +val ComplE = make_elim ComplD; + + +(*** Binary union -- Un ***) + +val prems = goalw Set.thy [Un_def] "c:A ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI1]) 1)); +qed "UnI1"; + +val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1)); +qed "UnI2"; + +(*Classical introduction rule: no commitment to A vs B*) +qed_goal "UnCI" Set.thy "(c~:B ==> c:A) ==> c : A Un B" + (fn prems=> + [ (rtac classical 1), + (REPEAT (ares_tac (prems@[UnI1,notI]) 1)), + (REPEAT (ares_tac (prems@[UnI2,notE]) 1)) ]); + +val major::prems = goalw Set.thy [Un_def] + "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"; +by (rtac (major RS CollectD RS disjE) 1); +by (REPEAT (eresolve_tac prems 1)); +qed "UnE"; + + +(*** Binary intersection -- Int ***) + +val prems = goalw Set.thy [Int_def] + "[| c:A; c:B |] ==> c : A Int B"; +by (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)); +qed "IntI"; + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A"; +by (rtac (major RS CollectD RS conjunct1) 1); +qed "IntD1"; + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B"; +by (rtac (major RS CollectD RS conjunct2) 1); +qed "IntD2"; + +val [major,minor] = goal Set.thy + "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"; +by (rtac minor 1); +by (rtac (major RS IntD1) 1); +by (rtac (major RS IntD2) 1); +qed "IntE"; + + +(*** Set difference ***) + +qed_goalw "DiffI" Set.thy [set_diff_def] + "[| c : A; c ~: B |] ==> c : A - B" + (fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)) ]); + +qed_goalw "DiffD1" Set.thy [set_diff_def] + "c : A - B ==> c : A" + (fn [major]=> [ (rtac (major RS CollectD RS conjunct1) 1) ]); + +qed_goalw "DiffD2" Set.thy [set_diff_def] + "[| c : A - B; c : B |] ==> P" + (fn [major,minor]=> + [rtac (minor RS (major RS CollectD RS conjunct2 RS notE)) 1]); + +qed_goal "DiffE" Set.thy + "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P" + (fn prems=> + [ (resolve_tac prems 1), + (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]); + +qed_goal "Diff_iff" Set.thy "(c : A-B) = (c:A & c~:B)" + (fn _ => [ (fast_tac (HOL_cs addSIs [DiffI] addSEs [DiffE]) 1) ]); + + +(*** The empty set -- {} ***) + +qed_goalw "emptyE" Set.thy [empty_def] "a:{} ==> P" + (fn [prem] => [rtac (prem RS CollectD RS FalseE) 1]); + +qed_goal "empty_subsetI" Set.thy "{} <= A" + (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]); + +qed_goal "equals0I" Set.thy "[| !!y. y:A ==> False |] ==> A={}" + (fn prems=> + [ (REPEAT (ares_tac (prems@[empty_subsetI,subsetI,equalityI]) 1 + ORELSE eresolve_tac (prems RL [FalseE]) 1)) ]); + +qed_goal "equals0D" Set.thy "[| A={}; a:A |] ==> P" + (fn [major,minor]=> + [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]); + + +(*** Augmenting a set -- insert ***) + +qed_goalw "insertI1" Set.thy [insert_def] "a : insert a B" + (fn _ => [rtac (CollectI RS UnI1) 1, rtac refl 1]); + +qed_goalw "insertI2" Set.thy [insert_def] "a : B ==> a : insert b B" + (fn [prem]=> [ (rtac (prem RS UnI2) 1) ]); + +qed_goalw "insertE" Set.thy [insert_def] + "[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS UnE) 1), + (REPEAT (eresolve_tac (prems @ [CollectE]) 1)) ]); + +qed_goal "insert_iff" Set.thy "a : insert b A = (a=b | a:A)" + (fn _ => [fast_tac (HOL_cs addIs [insertI1,insertI2] addSEs [insertE]) 1]); + +(*Classical introduction rule*) +qed_goal "insertCI" Set.thy "(a~:B ==> a=b) ==> a: insert b B" + (fn [prem]=> + [ (rtac (disjCI RS (insert_iff RS iffD2)) 1), + (etac prem 1) ]); + +(*** Singletons, using insert ***) + +qed_goal "singletonI" Set.thy "a : {a}" + (fn _=> [ (rtac insertI1 1) ]); + +qed_goal "singletonE" Set.thy "[| a: {b}; a=b ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS insertE) 1), + (REPEAT (eresolve_tac (prems @ [emptyE]) 1)) ]); + +goalw Set.thy [insert_def] "!!a. b : {a} ==> b=a"; +by(fast_tac (HOL_cs addSEs [emptyE,CollectE,UnE]) 1); +qed "singletonD"; + +val singletonE = make_elim singletonD; + +val [major] = goal Set.thy "{a}={b} ==> a=b"; +by (rtac (major RS equalityD1 RS subsetD RS singletonD) 1); +by (rtac singletonI 1); +qed "singleton_inject"; + +(*** Unions of families -- UNION x:A. B(x) is Union(B``A) ***) + +(*The order of the premises presupposes that A is rigid; b may be flexible*) +val prems = goalw Set.thy [UNION_def] + "[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; +by (REPEAT (resolve_tac (prems @ [bexI,CollectI]) 1)); +qed "UN_I"; + +val major::prems = goalw Set.thy [UNION_def] + "[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R"; +by (rtac (major RS CollectD RS bexE) 1); +by (REPEAT (ares_tac prems 1)); +qed "UN_E"; + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ +\ (UN x:A. C(x)) = (UN x:B. D(x))"; +by (REPEAT (etac UN_E 1 + ORELSE ares_tac ([UN_I,equalityI,subsetI] @ + (prems RL [equalityD1,equalityD2] RL [subsetD])) 1)); +qed "UN_cong"; + + +(*** Intersections of families -- INTER x:A. B(x) is Inter(B``A) *) + +val prems = goalw Set.thy [INTER_def] + "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"; +by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); +qed "INT_I"; + +val major::prems = goalw Set.thy [INTER_def] + "[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; +by (rtac (major RS CollectD RS bspec) 1); +by (resolve_tac prems 1); +qed "INT_D"; + +(*"Classical" elimination -- by the Excluded Middle on a:A *) +val major::prems = goalw Set.thy [INTER_def] + "[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R"; +by (rtac (major RS CollectD RS ballE) 1); +by (REPEAT (eresolve_tac prems 1)); +qed "INT_E"; + +val prems = goal Set.thy + "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ +\ (INT x:A. C(x)) = (INT x:B. D(x))"; +by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI])); +by (REPEAT (dtac INT_D 1 + ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1)); +qed "INT_cong"; + + +(*** Unions over a type; UNION1(B) = Union(range(B)) ***) + +(*The order of the premises presupposes that A is rigid; b may be flexible*) +val prems = goalw Set.thy [UNION1_def] + "b: B(x) ==> b: (UN x. B(x))"; +by (REPEAT (resolve_tac (prems @ [TrueI, CollectI RS UN_I]) 1)); +qed "UN1_I"; + +val major::prems = goalw Set.thy [UNION1_def] + "[| b : (UN x. B(x)); !!x. b: B(x) ==> R |] ==> R"; +by (rtac (major RS UN_E) 1); +by (REPEAT (ares_tac prems 1)); +qed "UN1_E"; + + +(*** Intersections over a type; INTER1(B) = Inter(range(B)) *) + +val prems = goalw Set.thy [INTER1_def] + "(!!x. b: B(x)) ==> b : (INT x. B(x))"; +by (REPEAT (ares_tac (INT_I::prems) 1)); +qed "INT1_I"; + +val [major] = goalw Set.thy [INTER1_def] + "b : (INT x. B(x)) ==> b: B(a)"; +by (rtac (TrueI RS (CollectI RS (major RS INT_D))) 1); +qed "INT1_D"; + +(*** Unions ***) + +(*The order of the premises presupposes that C is rigid; A may be flexible*) +val prems = goalw Set.thy [Union_def] + "[| X:C; A:X |] ==> A : Union(C)"; +by (REPEAT (resolve_tac (prems @ [UN_I]) 1)); +qed "UnionI"; + +val major::prems = goalw Set.thy [Union_def] + "[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R"; +by (rtac (major RS UN_E) 1); +by (REPEAT (ares_tac prems 1)); +qed "UnionE"; + +(*** Inter ***) + +val prems = goalw Set.thy [Inter_def] + "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; +by (REPEAT (ares_tac ([INT_I] @ prems) 1)); +qed "InterI"; + +(*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". *) +val major::prems = goalw Set.thy [Inter_def] + "[| A : Inter(C); X:C |] ==> A:X"; +by (rtac (major RS INT_D) 1); +by (resolve_tac prems 1); +qed "InterD"; + +(*"Classical" elimination rule -- does not require proving X:C *) +val major::prems = goalw Set.thy [Inter_def] + "[| A : Inter(C); A:X ==> R; X~:C ==> R |] ==> R"; +by (rtac (major RS INT_E) 1); +by (REPEAT (eresolve_tac prems 1)); +qed "InterE"; + +(*** Powerset ***) + +qed_goalw "PowI" Set.thy [Pow_def] "!!A B. A <= B ==> A : Pow(B)" + (fn _ => [ (etac CollectI 1) ]); + +qed_goalw "PowD" Set.thy [Pow_def] "!!A B. A : Pow(B) ==> A<=B" + (fn _=> [ (etac CollectD 1) ]); + +val Pow_bottom = empty_subsetI RS PowI; (* {}: Pow(B) *) +val Pow_top = subset_refl RS PowI; (* A : Pow(A) *) diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Set.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,145 @@ +(* Title: HOL/Set.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge +*) + +Set = Ord + + +types + 'a set + +arities + set :: (term) term + +instance + set :: (term) {ord, minus} + +consts + "{}" :: "'a set" ("{}") + insert :: "['a, 'a set] => 'a set" + Collect :: "('a => bool) => 'a set" (*comprehension*) + Compl :: "('a set) => 'a set" (*complement*) + Int :: "['a set, 'a set] => 'a set" (infixl 70) + Un :: "['a set, 'a set] => 'a set" (infixl 65) + UNION, INTER :: "['a set, 'a => 'b set] => 'b set" (*general*) + UNION1 :: "['a => 'b set] => 'b set" (binder "UN " 10) + INTER1 :: "['a => 'b set] => 'b set" (binder "INT " 10) + Union, Inter :: "(('a set)set) => 'a set" (*of a set*) + Pow :: "'a set => 'a set set" (*powerset*) + range :: "('a => 'b) => 'b set" (*of function*) + Ball, Bex :: "['a set, 'a => bool] => bool" (*bounded quantifiers*) + inj, surj :: "('a => 'b) => bool" (*inj/surjective*) + inj_onto :: "['a => 'b, 'a set] => bool" + "``" :: "['a => 'b, 'a set] => ('b set)" (infixl 90) + ":" :: "['a, 'a set] => bool" (infixl 50) (*membership*) + + +syntax + + "~:" :: "['a, 'a set] => bool" (infixl 50) + + "@Finset" :: "args => 'a set" ("{(_)}") + + "@Coll" :: "[idt, bool] => 'a set" ("(1{_./ _})") + "@SetCompr" :: "['a, idts, bool] => 'a set" ("(1{_ |/_./ _})") + + (* Big Intersection / Union *) + + "@INTER" :: "[idt, 'a set, 'b set] => 'b set" ("(3INT _:_./ _)" 10) + "@UNION" :: "[idt, 'a set, 'b set] => 'b set" ("(3UN _:_./ _)" 10) + + (* Bounded Quantifiers *) + + "@Ball" :: "[idt, 'a set, bool] => bool" ("(3! _:_./ _)" 10) + "@Bex" :: "[idt, 'a set, bool] => bool" ("(3? _:_./ _)" 10) + "*Ball" :: "[idt, 'a set, bool] => bool" ("(3ALL _:_./ _)" 10) + "*Bex" :: "[idt, 'a set, bool] => bool" ("(3EX _:_./ _)" 10) + +translations + "x ~: y" == "~ (x : y)" + "{x, xs}" == "insert x {xs}" + "{x}" == "insert x {}" + "{x. P}" == "Collect (%x. P)" + "INT x:A. B" == "INTER A (%x. B)" + "UN x:A. B" == "UNION A (%x. B)" + "! x:A. P" == "Ball A (%x. P)" + "? x:A. P" == "Bex A (%x. P)" + "ALL x:A. P" => "Ball A (%x. P)" + "EX x:A. P" => "Bex A (%x. P)" + + +rules + + (* Isomorphisms between Predicates and Sets *) + + mem_Collect_eq "(a : {x.P(x)}) = P(a)" + Collect_mem_eq "{x.x:A} = A" + + +defs + Ball_def "Ball A P == ! x. x:A --> P(x)" + Bex_def "Bex A P == ? x. x:A & P(x)" + subset_def "A <= B == ! x:A. x:B" + Compl_def "Compl(A) == {x. ~x:A}" + Un_def "A Un B == {x.x:A | x:B}" + Int_def "A Int B == {x.x:A & x:B}" + set_diff_def "A - B == {x. x:A & ~x:B}" + INTER_def "INTER A B == {y. ! x:A. y: B(x)}" + UNION_def "UNION A B == {y. ? x:A. y: B(x)}" + INTER1_def "INTER1(B) == INTER {x.True} B" + UNION1_def "UNION1(B) == UNION {x.True} B" + Inter_def "Inter(S) == (INT x:S. x)" + Union_def "Union(S) == (UN x:S. x)" + Pow_def "Pow(A) == {B. B <= A}" + empty_def "{} == {x. False}" + insert_def "insert a B == {x.x=a} Un B" + range_def "range(f) == {y. ? x. y=f(x)}" + image_def "f``A == {y. ? x:A. y=f(x)}" + inj_def "inj(f) == ! x y. f(x)=f(y) --> x=y" + inj_onto_def "inj_onto f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" + surj_def "surj(f) == ! y. ? x. y=f(x)" + +end + +ML + +local + +(* Translates between { e | x1..xn. P} and {u. ? x1..xn. u=e & P} *) +(* {y. ? x1..xn. y = e & P} is only translated if [0..n] subset bvs(e) *) + +val ex_tr = snd(mk_binder_tr("? ","Ex")); + +fun nvars(Const("_idts",_) $ _ $ idts) = nvars(idts)+1 + | nvars(_) = 1; + +fun setcompr_tr[e,idts,b] = + let val eq = Syntax.const("op =") $ Bound(nvars(idts)) $ e + val P = Syntax.const("op &") $ eq $ b + val exP = ex_tr [idts,P] + in Syntax.const("Collect") $ Abs("",dummyT,exP) end; + +val ex_tr' = snd(mk_binder_tr' ("Ex","DUMMY")); + +fun setcompr_tr'[Abs(_,_,P)] = + let fun ok(Const("Ex",_)$Abs(_,_,P),n) = ok(P,n+1) + | ok(Const("op &",_) $ (Const("op =",_) $ Bound(m) $ e) $ _, n) = + if n>0 andalso m=n andalso + ((0 upto (n-1)) subset add_loose_bnos(e,0,[])) + then () else raise Match + + fun tr'(_ $ abs) = + let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr'[abs] + in Syntax.const("@SetCompr") $ e $ idts $ Q end + in ok(P,0); tr'(P) end; + +in + +val parse_translation = [("@SetCompr", setcompr_tr)]; +val print_translation = [("Collect", setcompr_tr')]; +val print_ast_translation = + map HOL.alt_ast_tr' [("@Ball", "*Ball"), ("@Bex", "*Bex")]; + +end; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Sexp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Sexp.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,135 @@ +(* Title: HOL/Sexp + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +S-expressions, general binary trees for defining recursive data structures +*) + +open Sexp; + +(** sexp_case **) + +val sexp_free_cs = + set_cs addSDs [Leaf_inject, Numb_inject, Scons_inject] + addSEs [Leaf_neq_Scons, Leaf_neq_Numb, + Numb_neq_Scons, Numb_neq_Leaf, + Scons_neq_Leaf, Scons_neq_Numb]; + +goalw Sexp.thy [sexp_case_def] "sexp_case c d e (Leaf a) = c(a)"; +by (resolve_tac [select_equality] 1); +by (ALLGOALS (fast_tac sexp_free_cs)); +qed "sexp_case_Leaf"; + +goalw Sexp.thy [sexp_case_def] "sexp_case c d e (Numb k) = d(k)"; +by (resolve_tac [select_equality] 1); +by (ALLGOALS (fast_tac sexp_free_cs)); +qed "sexp_case_Numb"; + +goalw Sexp.thy [sexp_case_def] "sexp_case c d e (M$N) = e M N"; +by (resolve_tac [select_equality] 1); +by (ALLGOALS (fast_tac sexp_free_cs)); +qed "sexp_case_Scons"; + + +(** Introduction rules for sexp constructors **) + +val [prem] = goalw Sexp.thy [In0_def] + "M: sexp ==> In0(M) : sexp"; +by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1); +qed "sexp_In0I"; + +val [prem] = goalw Sexp.thy [In1_def] + "M: sexp ==> In1(M) : sexp"; +by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1); +qed "sexp_In1I"; + +val sexp_cs = set_cs addIs sexp.intrs@[SigmaI, uprodI]; + +goal Sexp.thy "range(Leaf) <= sexp"; +by (fast_tac sexp_cs 1); +qed "range_Leaf_subset_sexp"; + +val [major] = goal Sexp.thy "M$N : sexp ==> M: sexp & N: sexp"; +by (rtac (major RS setup_induction) 1); +by (etac sexp.induct 1); +by (ALLGOALS + (fast_tac (set_cs addSEs [Scons_neq_Leaf,Scons_neq_Numb,Scons_inject]))); +qed "Scons_D"; + +(** Introduction rules for 'pred_sexp' **) + +goalw Sexp.thy [pred_sexp_def] "pred_sexp <= Sigma sexp (%u.sexp)"; +by (fast_tac sexp_cs 1); +qed "pred_sexp_subset_Sigma"; + +(* : pred_sexp^+ ==> a : sexp *) +val trancl_pred_sexpD1 = + pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD RS SigmaD1 +and trancl_pred_sexpD2 = + pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD RS SigmaD2; + +val prems = goalw Sexp.thy [pred_sexp_def] + "[| M: sexp; N: sexp |] ==> : pred_sexp"; +by (fast_tac (set_cs addIs prems) 1); +qed "pred_sexpI1"; + +val prems = goalw Sexp.thy [pred_sexp_def] + "[| M: sexp; N: sexp |] ==> : pred_sexp"; +by (fast_tac (set_cs addIs prems) 1); +qed "pred_sexpI2"; + +(*Combinations involving transitivity and the rules above*) +val pred_sexp_t1 = pred_sexpI1 RS r_into_trancl +and pred_sexp_t2 = pred_sexpI2 RS r_into_trancl; + +val pred_sexp_trans1 = pred_sexp_t1 RSN (2, trans_trancl RS transD) +and pred_sexp_trans2 = pred_sexp_t2 RSN (2, trans_trancl RS transD); + +(*Proves goals of the form :pred_sexp^+ provided M,N:sexp*) +val pred_sexp_simps = + sexp.intrs @ + [pred_sexp_t1, pred_sexp_t2, + pred_sexp_trans1, pred_sexp_trans2, cut_apply]; +val pred_sexp_ss = HOL_ss addsimps pred_sexp_simps; + +val major::prems = goalw Sexp.thy [pred_sexp_def] + "[| p : pred_sexp; \ +\ !!M N. [| p = ; M: sexp; N: sexp |] ==> R; \ +\ !!M N. [| p = ; M: sexp; N: sexp |] ==> R \ +\ |] ==> R"; +by (cut_facts_tac [major] 1); +by (REPEAT (eresolve_tac ([asm_rl,emptyE,insertE,UN_E]@prems) 1)); +qed "pred_sexpE"; + +goal Sexp.thy "wf(pred_sexp)"; +by (rtac (pred_sexp_subset_Sigma RS wfI) 1); +by (etac sexp.induct 1); +by (fast_tac (HOL_cs addSEs [mp, pred_sexpE, Pair_inject, Scons_inject]) 3); +by (fast_tac (HOL_cs addSEs [mp, pred_sexpE, Pair_inject, Numb_neq_Scons]) 2); +by (fast_tac (HOL_cs addSEs [mp, pred_sexpE, Pair_inject, Leaf_neq_Scons]) 1); +qed "wf_pred_sexp"; + +(*** sexp_rec -- by wf recursion on pred_sexp ***) + +(** conversion rules **) + +val sexp_rec_unfold = wf_pred_sexp RS (sexp_rec_def RS def_wfrec); + + +goal Sexp.thy "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)"; +by (stac sexp_rec_unfold 1); +by (rtac sexp_case_Numb 1); +qed "sexp_rec_Numb"; + +goal Sexp.thy "!!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(HOL_ss addsimps + [sexp_case_Scons,pred_sexpI1,pred_sexpI2,cut_apply])1); +qed "sexp_rec_Scons"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Sexp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Sexp.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,40 @@ +(* Title: HOL/Sexp + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +S-expressions, general binary trees for defining recursive data structures +*) + +Sexp = Univ + +consts + sexp :: "'a item set" + + sexp_case :: "['a=>'b, nat=>'b, ['a item, 'a item]=>'b, \ +\ 'a item] => 'b" + + sexp_rec :: "['a item, 'a=>'b, nat=>'b, \ +\ ['a item, 'a item, 'b, 'b]=>'b] => 'b" + + pred_sexp :: "('a item * 'a item)set" + +inductive "sexp" + intrs + LeafI "Leaf(a): sexp" + NumbI "Numb(a): sexp" + SconsI "[| M: sexp; N: sexp |] ==> M$N : sexp" + +defs + + sexp_case_def + "sexp_case c d e M == @ z. (? x. M=Leaf(x) & z=c(x)) \ +\ | (? k. M=Numb(k) & z=d(k)) \ +\ | (? N1 N2. M = N1 $ N2 & z=e N1 N2)" + + pred_sexp_def + "pred_sexp == UN M: sexp. UN N: sexp. {, }" + + sexp_rec_def + "sexp_rec M c d e == wfrec pred_sexp M \ +\ (%M g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2)) M)" +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Sum.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Sum.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,204 @@ +(* Title: HOL/Sum.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For Sum.thy. The disjoint sum of two types +*) + +open Sum; + +(** 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"; +by (EVERY1 [rtac CollectI, rtac disjI1, rtac exI, rtac refl]); +qed "Inl_RepI"; + +goalw Sum.thy [Sum_def] "Inr_Rep(b) : Sum"; +by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]); +qed "Inr_RepI"; + +goal Sum.thy "inj_onto Abs_Sum Sum"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Sum_inverse 1); +qed "inj_onto_Abs_Sum"; + +(** Distinctness of Inl and Inr **) + +goalw Sum.thy [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)"; +by (rtac (inj_onto_Abs_Sum RS inj_onto_contraD) 1); +by (rtac Inl_Rep_not_Inr_Rep 1); +by (rtac Inl_RepI 1); +by (rtac Inr_RepI 1); +qed "Inl_not_Inr"; + +bind_thm ("Inl_neq_Inr", (Inl_not_Inr RS notE)); +val Inr_neq_Inl = sym RS Inl_neq_Inr; + +goal Sum.thy "(Inl(a)=Inr(b)) = False"; +by (simp_tac (HOL_ss addsimps [Inl_not_Inr]) 1); +qed "Inl_Inr_eq"; + +goal Sum.thy "(Inr(b)=Inl(a)) = False"; +by (simp_tac (HOL_ss addsimps [Inl_not_Inr RS not_sym]) 1); +qed "Inr_Inl_eq"; + + +(** Injectiveness of Inl and Inr **) + +val [major] = goalw Sum.thy [Inl_Rep_def] "Inl_Rep(a) = Inl_Rep(c) ==> a=c"; +by (rtac (major RS fun_cong RS fun_cong RS fun_cong RS iffE) 1); +by (fast_tac HOL_cs 1); +qed "Inl_Rep_inject"; + +val [major] = goalw Sum.thy [Inr_Rep_def] "Inr_Rep(b) = Inr_Rep(d) ==> b=d"; +by (rtac (major RS fun_cong RS fun_cong RS fun_cong RS iffE) 1); +by (fast_tac HOL_cs 1); +qed "Inr_Rep_inject"; + +goalw Sum.thy [Inl_def] "inj(Inl)"; +by (rtac injI 1); +by (etac (inj_onto_Abs_Sum RS inj_ontoD RS Inl_Rep_inject) 1); +by (rtac Inl_RepI 1); +by (rtac Inl_RepI 1); +qed "inj_Inl"; +val Inl_inject = inj_Inl RS injD; + +goalw Sum.thy [Inr_def] "inj(Inr)"; +by (rtac injI 1); +by (etac (inj_onto_Abs_Sum RS inj_ontoD RS Inr_Rep_inject) 1); +by (rtac Inr_RepI 1); +by (rtac Inr_RepI 1); +qed "inj_Inr"; +val Inr_inject = inj_Inr RS injD; + +goal Sum.thy "(Inl(x)=Inl(y)) = (x=y)"; +by (fast_tac (HOL_cs addSEs [Inl_inject]) 1); +qed "Inl_eq"; + +goal Sum.thy "(Inr(x)=Inr(y)) = (x=y)"; +by (fast_tac (HOL_cs addSEs [Inr_inject]) 1); +qed "Inr_eq"; + +(*** Rules for the disjoint sum of two SETS ***) + +(** Introduction rules for the injections **) + +goalw Sum.thy [sum_def] "!!a A B. a : A ==> Inl(a) : A plus B"; +by (REPEAT (ares_tac [UnI1,imageI] 1)); +qed "InlI"; + +goalw Sum.thy [sum_def] "!!b A B. b : B ==> Inr(b) : A plus B"; +by (REPEAT (ares_tac [UnI2,imageI] 1)); +qed "InrI"; + +(** Elimination rules **) + +val major::prems = goalw Sum.thy [sum_def] + "[| u: A plus B; \ +\ !!x. [| x:A; u=Inl(x) |] ==> P; \ +\ !!y. [| y:B; u=Inr(y) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (REPEAT (rtac refl 1 + ORELSE eresolve_tac (prems@[imageE,ssubst]) 1)); +qed "plusE"; + + +val sum_cs = set_cs addSIs [InlI, InrI] + addSEs [plusE, Inl_neq_Inr, Inr_neq_Inl] + addSDs [Inl_inject, Inr_inject]; + + +(** sum_case -- the selection operator for sums **) + +goalw Sum.thy [sum_case_def] "sum_case f g (Inl x) = f(x)"; +by (fast_tac (sum_cs addIs [select_equality]) 1); +qed "sum_case_Inl"; + +goalw Sum.thy [sum_case_def] "sum_case f g (Inr x) = g(x)"; +by (fast_tac (sum_cs addIs [select_equality]) 1); +qed "sum_case_Inr"; + +(** Exhaustion rule for sums -- a degenerate form of induction **) + +val prems = goalw Sum.thy [Inl_def,Inr_def] + "[| !!x::'a. s = Inl(x) ==> P; !!y::'b. s = Inr(y) ==> P \ +\ |] ==> P"; +by (rtac (rewrite_rule [Sum_def] Rep_Sum RS CollectE) 1); +by (REPEAT (eresolve_tac [disjE,exE] 1 + ORELSE EVERY1 [resolve_tac prems, + etac subst, + 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)"; +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) = \ +\ ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))"; +by (rtac sumE 1); +by (etac ssubst 1); +by (stac sum_case_Inl 1); +by (fast_tac (set_cs addSEs [make_elim Inl_inject, Inl_neq_Inr]) 1); +by (etac ssubst 1); +by (stac sum_case_Inr 1); +by (fast_tac (set_cs addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1); +qed "expand_sum_case"; + +val sum_ss = prod_ss addsimps [Inl_eq, Inr_eq, Inl_Inr_eq, Inr_Inl_eq, + sum_case_Inl, sum_case_Inr]; + +(*Prevents simplification of f and g: much faster*) +qed_goal "sum_case_weak_cong" Sum.thy + "s=t ==> sum_case f g s = sum_case f g t" + (fn [prem] => [rtac (prem RS arg_cong) 1]); + + + + +(** Rules for the Part primitive **) + +goalw Sum.thy [Part_def] + "!!a b A h. [| a : A; a=h(b) |] ==> a : Part A h"; +by (fast_tac set_cs 1); +qed "Part_eqI"; + +val PartI = refl RSN (2,Part_eqI); + +val major::prems = goalw Sum.thy [Part_def] + "[| a : Part A h; !!z. [| a : A; a=h(z) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS IntE) 1); +by (etac CollectE 1); +by (etac exE 1); +by (REPEAT (ares_tac prems 1)); +qed "PartE"; + +goalw Sum.thy [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"; +by (fast_tac (set_cs addSIs [PartI] addSEs [PartE]) 1); +qed "Part_mono"; + +goalw Sum.thy [Part_def] "!!a. a : Part A h ==> a : A"; +by (etac IntD1 1); +qed "PartD1"; + +goal Sum.thy "Part A (%x.x) = A"; +by (fast_tac (set_cs addIs [PartI,equalityI] addSEs [PartE]) 1); +qed "Part_id"; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Sum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Sum.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,51 @@ +(* Title: HOL/Sum.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The disjoint sum of two types. +*) + +Sum = Prod + + +(* type definition *) + +consts + Inl_Rep :: "['a, 'a, 'b, bool] => bool" + Inr_Rep :: "['b, 'a, 'b, bool] => bool" + +defs + Inl_Rep_def "Inl_Rep == (%a. %x y p. x=a & p)" + Inr_Rep_def "Inr_Rep == (%b. %x y p. y=b & ~p)" + +subtype (Sum) + ('a, 'b) "+" (infixr 10) + = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}" + + +(* abstract constants and syntax *) + +consts + Inl :: "'a => 'a + 'b" + Inr :: "'b => 'a + 'b" + sum_case :: "['a => 'c, 'b => 'c, 'a + 'b] => 'c" + + (*disjoint sum for sets; the operator + is overloaded with wrong type!*) + "plus" :: "['a set, 'b set] => ('a + 'b) set" (infixr 65) + Part :: "['a set, 'b => 'a] => 'a set" + +translations + "case p of Inl(x) => a | Inr(y) => b" == "sum_case (%x.a) (%y.b) p" + +defs + Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))" + Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))" + sum_case_def "sum_case f g p == @z. (!x. p=Inl(x) --> z=f(x)) \ +\ & (!y. p=Inr(y) --> z=g(y))" + + sum_def "A plus B == (Inl``A) Un (Inr``B)" + + (*for selecting out the components of a mutually recursive definition*) + Part_def "Part A h == A Int {x. ? z. x = h(z)}" + +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Trancl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Trancl.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,237 @@ +(* Title: HOL/trancl + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For trancl.thy. Theorems about the transitive closure of a relation +*) + +open Trancl; + +(** Natural deduction for trans(r) **) + +val prems = goalw Trancl.thy [trans_def] + "(!! x y z. [| :r; :r |] ==> :r) ==> trans(r)"; +by (REPEAT (ares_tac (prems@[allI,impI]) 1)); +qed "transI"; + +val major::prems = goalw Trancl.thy [trans_def] + "[| trans(r); :r; :r |] ==> :r"; +by (cut_facts_tac [major] 1); +by (fast_tac (HOL_cs addIs prems) 1); +qed "transD"; + +(** Identity relation **) + +goalw Trancl.thy [id_def] " : id"; +by (rtac CollectI 1); +by (rtac exI 1); +by (rtac refl 1); +qed "idI"; + +val major::prems = goalw Trancl.thy [id_def] + "[| p: id; !!x.[| p = |] ==> P \ +\ |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac exE 1); +by (eresolve_tac prems 1); +qed "idE"; + +goalw Trancl.thy [id_def] ":id = (a=b)"; +by(fast_tac prod_cs 1); +qed "pair_in_id_conv"; + +(** Composition of two relations **) + +val prems = goalw Trancl.thy [comp_def] + "[| :s; :r |] ==> : r O s"; +by (fast_tac (set_cs addIs prems) 1); +qed "compI"; + +(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*) +val prems = goalw Trancl.thy [comp_def] + "[| xz : r O s; \ +\ !!x y z. [| xz = ; :s; :r |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac prems 1); +by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1)); +qed "compE"; + +val prems = goal Trancl.thy + "[| : r O s; \ +\ !!y. [| :s; :r |] ==> P \ +\ |] ==> P"; +by (rtac compE 1); +by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Pair_inject,ssubst] 1)); +qed "compEpair"; + +val comp_cs = prod_cs addIs [compI, idI] addSEs [compE, idE]; + +goal Trancl.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +by (fast_tac comp_cs 1); +qed "comp_mono"; + +goal Trancl.thy + "!!r s. [| s <= Sigma A (%x.B); r <= Sigma B (%x.C) |] ==> \ +\ (r O s) <= Sigma A (%x.C)"; +by (fast_tac comp_cs 1); +qed "comp_subset_Sigma"; + + +(** The relation rtrancl **) + +goal Trancl.thy "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"; + +val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski); + +(*Reflexivity of rtrancl*) +goal Trancl.thy " : r^*"; +by (stac rtrancl_unfold 1); +by (fast_tac comp_cs 1); +qed "rtrancl_refl"; + +(*Closure under composition with r*) +val prems = goal Trancl.thy + "[| : r^*; : r |] ==> : r^*"; +by (stac rtrancl_unfold 1); +by (fast_tac (comp_cs addIs prems) 1); +qed "rtrancl_into_rtrancl"; + +(*rtrancl of r contains r*) +val [prem] = goal Trancl.thy "[| : r |] ==> : r^*"; +by (rtac (rtrancl_refl RS rtrancl_into_rtrancl) 1); +by (rtac prem 1); +qed "r_into_rtrancl"; + +(*monotonicity of rtrancl*) +goalw Trancl.thy [rtrancl_def] "!!r s. r <= s ==> r^* <= s^*"; +by(REPEAT(ares_tac [lfp_mono,Un_mono,comp_mono,subset_refl] 1)); +qed "rtrancl_mono"; + +(** standard induction rule **) + +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ !!x. P(); \ +\ !!x y z.[| P(); : r^*; : r |] ==> P() |] \ +\ ==> P()"; +by (rtac ([rtrancl_def, rtrancl_fun_mono, major] MRS def_induct) 1); +by (fast_tac (comp_cs addIs prems) 1); +qed "rtrancl_full_induct"; + +(*nice induction rule*) +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ P(a); \ +\ !!y z.[| : r^*; : r; P(y) |] ==> P(z) |] \ +\ ==> P(b)"; +(*by induction on this formula*) +by (subgoal_tac "! y. = --> P(y)" 1); +(*now solve first subgoal: this formula is sufficient*) +by (fast_tac HOL_cs 1); +(*now do the induction*) +by (resolve_tac [major RS rtrancl_full_induct] 1); +by (fast_tac (comp_cs addIs prems) 1); +by (fast_tac (comp_cs addIs prems) 1); +qed "rtrancl_induct"; + +(*transitivity of transitive closure!! -- by induction.*) +goal Trancl.thy "trans(r^*)"; +by (rtac transI 1); +by (res_inst_tac [("b","z")] rtrancl_induct 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); +qed "trans_rtrancl"; + +(*elimination of rtrancl -- by induction on a special formula*) +val major::prems = goal Trancl.thy + "[| : r^*; (a = b) ==> P; \ +\ !!y.[| : r^*; : r |] ==> P \ +\ |] ==> P"; +by (subgoal_tac "(a::'a) = b | (? y. : r^* & : r)" 1); +by (rtac (major RS rtrancl_induct) 2); +by (fast_tac (set_cs addIs prems) 2); +by (fast_tac (set_cs addIs prems) 2); +by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1)); +qed "rtranclE"; + + +(**** The relation trancl ****) + +(** Conversions between trancl and rtrancl **) + +val [major] = goalw Trancl.thy [trancl_def] + " : r^+ ==> : r^*"; +by (resolve_tac [major RS compEpair] 1); +by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1)); +qed "trancl_into_rtrancl"; + +(*r^+ contains r*) +val [prem] = goalw Trancl.thy [trancl_def] + "[| : r |] ==> : r^+"; +by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); +qed "r_into_trancl"; + +(*intro rule by definition: from rtrancl and r*) +val prems = goalw Trancl.thy [trancl_def] + "[| : r^*; : r |] ==> : r^+"; +by (REPEAT (resolve_tac ([compI]@prems) 1)); +qed "rtrancl_into_trancl1"; + +(*intro rule from r and rtrancl*) +val prems = goal Trancl.thy + "[| : r; : r^* |] ==> : r^+"; +by (resolve_tac (prems RL [rtranclE]) 1); +by (etac subst 1); +by (resolve_tac (prems RL [r_into_trancl]) 1); +by (rtac (trans_rtrancl RS transD RS rtrancl_into_trancl1) 1); +by (REPEAT (ares_tac (prems@[r_into_rtrancl]) 1)); +qed "rtrancl_into_trancl2"; + +(*elimination of r^+ -- NOT an induction rule*) +val major::prems = goal Trancl.thy + "[| : r^+; \ +\ : r ==> P; \ +\ !!y.[| : r^+; : r |] ==> P \ +\ |] ==> P"; +by (subgoal_tac " : r | (? y. : r^+ & : r)" 1); +by (REPEAT (eresolve_tac ([asm_rl,disjE,exE,conjE]@prems) 1)); +by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); +by (etac rtranclE 1); +by (fast_tac comp_cs 1); +by (fast_tac (comp_cs addSIs [rtrancl_into_trancl1]) 1); +qed "tranclE"; + +(*Transitivity of r^+. + Proved by unfolding since it uses transitivity of rtrancl. *) +goalw Trancl.thy [trancl_def] "trans(r^+)"; +by (rtac transI 1); +by (REPEAT (etac compEpair 1)); +by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1); +by (REPEAT (assume_tac 1)); +qed "trans_trancl"; + +val prems = goal Trancl.thy + "[| : r; : r^+ |] ==> : r^+"; +by (rtac (r_into_trancl RS (trans_trancl RS transD)) 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +qed "trancl_into_trancl2"; + + +val major::prems = goal Trancl.thy + "[| : r^*; r <= Sigma A (%x.A) |] ==> a=b | a:A"; +by (cut_facts_tac prems 1); +by (rtac (major RS rtrancl_induct) 1); +by (rtac (refl RS disjI1) 1); +by (fast_tac (comp_cs addSEs [SigmaE2]) 1); +qed "trancl_subset_Sigma_lemma"; + +goalw Trancl.thy [trancl_def] + "!!r. r <= Sigma A (%x.A) ==> trancl(r) <= Sigma A (%x.A)"; +by (fast_tac (comp_cs addSDs [trancl_subset_Sigma_lemma]) 1); +qed "trancl_subset_Sigma"; + +val prod_ss = prod_ss addsimps [pair_in_id_conv]; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Trancl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Trancl.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,26 @@ +(* Title: HOL/trancl.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Transitive closure of a relation + +rtrancl is refl/transitive closure; trancl is transitive closure +*) + +Trancl = Lfp + Prod + +consts + trans :: "('a * 'a)set => bool" (*transitivity predicate*) + id :: "('a * 'a)set" + rtrancl :: "('a * 'a)set => ('a * 'a)set" ("(_^*)" [100] 100) + trancl :: "('a * 'a)set => ('a * 'a)set" ("(_^+)" [100] 100) + O :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60) +defs +trans_def "trans(r) == (!x y z. :r --> :r --> :r)" +comp_def (*composition of relations*) + "r O s == {xz. ? x y z. xz = & :s & :r}" +id_def (*the identity relation*) + "id == {p. ? x. p = }" +rtrancl_def "r^* == lfp(%s. id Un (r O s))" +trancl_def "r^+ == r O rtrancl(r)" +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Univ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Univ.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,615 @@ +(* Title: HOL/univ + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For univ.thy +*) + +open Univ; + +(** LEAST -- the least number operator **) + + +val [prem1,prem2] = goalw Univ.thy [Least_def] + "[| P(k); !!x. x ~P(x) |] ==> (LEAST x.P(x)) = k"; +by (rtac select_equality 1); +by (fast_tac (HOL_cs addSIs [prem1,prem2]) 1); +by (cut_facts_tac [less_linear] 1); +by (fast_tac (HOL_cs addSIs [prem1] addSDs [prem2]) 1); +qed "Least_equality"; + +val [prem] = goal Univ.thy "P(k) ==> P(LEAST x.P(x))"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("n","k")] less_induct 1); +by (rtac impI 1); +by (rtac classical 1); +by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1); +by (assume_tac 1); +by (assume_tac 2); +by (fast_tac HOL_cs 1); +qed "LeastI"; + +(*Proof is almost identical to the one above!*) +val [prem] = goal Univ.thy "P(k) ==> (LEAST x.P(x)) <= k"; +by (rtac (prem RS rev_mp) 1); +by (res_inst_tac [("n","k")] less_induct 1); +by (rtac impI 1); +by (rtac classical 1); +by (res_inst_tac [("s","n")] (Least_equality RS ssubst) 1); +by (assume_tac 1); +by (rtac le_refl 2); +by (fast_tac (HOL_cs addIs [less_imp_le,le_trans]) 1); +qed "Least_le"; + +val [prem] = goal Univ.thy "k < (LEAST x.P(x)) ==> ~P(k)"; +by (rtac notI 1); +by (etac (rewrite_rule [le_def] Least_le RS notE) 1); +by (rtac prem 1); +qed "not_less_Least"; + + +(** apfst -- can be used in similar type definitions **) + +goalw Univ.thy [apfst_def] "apfst f = "; +by (rtac split 1); +qed "apfst"; + +val [major,minor] = goal Univ.thy + "[| q = apfst f p; !!x y. [| p = ; q = |] ==> R \ +\ |] ==> R"; +by (rtac PairE 1); +by (rtac minor 1); +by (assume_tac 1); +by (rtac (major RS trans) 1); +by (etac ssubst 1); +by (rtac apfst 1); +qed "apfstE"; + +(** Push -- an injection, analogous to Cons on lists **) + +val [major] = goalw Univ.thy [Push_def] "Push i f =Push j g ==> i=j"; +by (rtac (major RS fun_cong RS box_equals RS Suc_inject) 1); +by (rtac nat_case_0 1); +by (rtac nat_case_0 1); +qed "Push_inject1"; + +val [major] = goalw Univ.thy [Push_def] "Push i f =Push j g ==> f=g"; +by (rtac (major RS fun_cong RS ext RS box_equals) 1); +by (rtac (nat_case_Suc RS ext) 1); +by (rtac (nat_case_Suc RS ext) 1); +qed "Push_inject2"; + +val [major,minor] = goal Univ.thy + "[| Push i f =Push j g; [| i=j; f=g |] ==> P \ +\ |] ==> P"; +by (rtac ((major RS Push_inject2) RS ((major RS Push_inject1) RS minor)) 1); +qed "Push_inject"; + +val [major] = goalw Univ.thy [Push_def] "Push k f =(%z.0) ==> P"; +by (rtac (major RS fun_cong RS box_equals RS Suc_neq_Zero) 1); +by (rtac nat_case_0 1); +by (rtac refl 1); +qed "Push_neq_K0"; + +(*** Isomorphisms ***) + +goal Univ.thy "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_onto Abs_Node Node"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Node_inverse 1); +qed "inj_onto_Abs_Node"; + +val Abs_Node_inject = inj_onto_Abs_Node RS inj_ontoD; + + +(*** Introduction rules for Node ***) + +goalw Univ.thy [Node_def] "<%k. 0,a> : Node"; +by (fast_tac set_cs 1); +qed "Node_K0_I"; + +goalw Univ.thy [Node_def,Push_def] + "!!p. p: Node ==> apfst (Push i) p : Node"; +by (fast_tac (set_cs addSIs [apfst, nat_case_Suc RS trans]) 1); +qed "Node_Push_I"; + + +(*** Distinctness of constructors ***) + +(** Scons vs Atom **) + +goalw Univ.thy [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); +by (REPEAT (eresolve_tac [imageE, Abs_Node_inject RS apfstE, + Pair_inject, sym RS Push_neq_K0] 1 + ORELSE resolve_tac [Node_K0_I, Rep_Node RS Node_Push_I] 1)); +qed "Scons_not_Atom"; +bind_thm ("Atom_not_Scons", (Scons_not_Atom RS not_sym)); + +bind_thm ("Scons_neq_Atom", (Scons_not_Atom RS notE)); +val Atom_neq_Scons = sym RS Scons_neq_Atom; + +(*** Injectiveness ***) + +(** Atomic nodes **) + +goalw Univ.thy [Atom_def] "inj(Atom)"; +by (rtac injI 1); +by (etac (singleton_inject RS Abs_Node_inject RS Pair_inject) 1); +by (REPEAT (ares_tac [Node_K0_I] 1)); +qed "inj_Atom"; +val Atom_inject = inj_Atom RS injD; + +goalw Univ.thy [Leaf_def,o_def] "inj(Leaf)"; +by (rtac injI 1); +by (etac (Atom_inject RS Inl_inject) 1); +qed "inj_Leaf"; + +val Leaf_inject = inj_Leaf RS injD; + +goalw Univ.thy [Numb_def,o_def] "inj(Numb)"; +by (rtac injI 1); +by (etac (Atom_inject RS Inr_inject) 1); +qed "inj_Numb"; + +val Numb_inject = inj_Numb RS injD; + +(** Injectiveness of Push_Node **) + +val [major,minor] = goalw Univ.thy [Push_Node_def] + "[| Push_Node i m =Push_Node j n; [| i=j; m=n |] ==> P \ +\ |] ==> P"; +by (rtac (major RS Abs_Node_inject RS apfstE) 1); +by (REPEAT (resolve_tac [Rep_Node RS Node_Push_I] 1)); +by (etac (sym RS apfstE) 1); +by (rtac minor 1); +by (etac Pair_inject 1); +by (etac (Push_inject1 RS sym) 1); +by (rtac (inj_Rep_Node RS injD) 1); +by (etac trans 1); +by (safe_tac (HOL_cs addSEs [Pair_inject,Push_inject,sym])); +qed "Push_Node_inject"; + + +(** Injectiveness of Scons **) + +val [major] = goalw Univ.thy [Scons_def] "M$N <= M'$N' ==> M<=M'"; +by (cut_facts_tac [major] 1); +by (fast_tac (set_cs addSDs [Suc_inject] + addSEs [Push_Node_inject, Zero_neq_Suc]) 1); +qed "Scons_inject_lemma1"; + +val [major] = goalw Univ.thy [Scons_def] "M$N <= M'$N' ==> N<=N'"; +by (cut_facts_tac [major] 1); +by (fast_tac (set_cs addSDs [Suc_inject] + addSEs [Push_Node_inject, Suc_neq_Zero]) 1); +qed "Scons_inject_lemma2"; + +val [major] = goal Univ.thy "M$N = M'$N' ==> M=M'"; +by (rtac (major RS equalityE) 1); +by (REPEAT (ares_tac [equalityI, Scons_inject_lemma1] 1)); +qed "Scons_inject1"; + +val [major] = goal Univ.thy "M$N = M'$N' ==> N=N'"; +by (rtac (major RS equalityE) 1); +by (REPEAT (ares_tac [equalityI, Scons_inject_lemma2] 1)); +qed "Scons_inject2"; + +val [major,minor] = goal Univ.thy + "[| M$N = M'$N'; [| M=M'; N=N' |] ==> P \ +\ |] ==> P"; +by (rtac ((major RS Scons_inject2) RS ((major RS Scons_inject1) RS minor)) 1); +qed "Scons_inject"; + +(*rewrite rules*) +goal Univ.thy "(Atom(a)=Atom(b)) = (a=b)"; +by (fast_tac (HOL_cs addSEs [Atom_inject]) 1); +qed "Atom_Atom_eq"; + +goal Univ.thy "(M$N = M'$N') = (M=M' & N=N')"; +by (fast_tac (HOL_cs addSEs [Scons_inject]) 1); +qed "Scons_Scons_eq"; + +(*** Distinctness involving Leaf and Numb ***) + +(** Scons vs Leaf **) + +goalw Univ.thy [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)); + +bind_thm ("Scons_neq_Leaf", (Scons_not_Leaf RS notE)); +val Leaf_neq_Scons = sym RS Scons_neq_Leaf; + +(** Scons vs Numb **) + +goalw Univ.thy [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)); + +bind_thm ("Scons_neq_Numb", (Scons_not_Numb RS notE)); +val Numb_neq_Scons = sym RS Scons_neq_Numb; + +(** Leaf vs Numb **) + +goalw Univ.thy [Leaf_def,Numb_def] "Leaf(a) ~= Numb(k)"; +by (simp_tac (HOL_ss addsimps [Atom_Atom_eq,Inl_not_Inr]) 1); +qed "Leaf_not_Numb"; +bind_thm ("Numb_not_Leaf", (Leaf_not_Numb RS not_sym)); + +bind_thm ("Leaf_neq_Numb", (Leaf_not_Numb RS notE)); +val Numb_neq_Leaf = sym RS Leaf_neq_Numb; + + +(*** ndepth -- the depth of a node ***) + +val univ_simps = [apfst,Scons_not_Atom,Atom_not_Scons,Scons_Scons_eq]; +val univ_ss = nat_ss addsimps univ_simps; + + +goalw Univ.thy [ndepth_def] "ndepth (Abs_Node(<%k.0, x>)) = 0"; +by (sstac [Node_K0_I RS Abs_Node_inverse, split] 1); +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) --> nat_case (Suc i) f k ~= 0"; +by (nat_ind_tac "k" 1); +by (ALLGOALS (simp_tac nat_ss)); +by (rtac impI 1); +by (etac not_less_Least 1); +qed "ndepth_Push_lemma"; + +goalw Univ.thy [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); +by (safe_tac set_cs); +be ssubst 1; (*instantiates type variables!*) +by (simp_tac univ_ss 1); +by (rtac Least_equality 1); +by (rewtac Push_def); +by (rtac (nat_case_Suc RS trans) 1); +by (etac LeastI 1); +by (etac (ndepth_Push_lemma RS mp) 1); +qed "ndepth_Push_Node"; + + +(*** ntrunc applied to the various node sets ***) + +goalw Univ.thy [ntrunc_def] "ntrunc 0 M = {}"; +by (safe_tac (set_cs addSIs [equalityI] addSEs [less_zeroE])); +qed "ntrunc_0"; + +goalw Univ.thy [Atom_def,ntrunc_def] "ntrunc (Suc k) (Atom a) = Atom(a)"; +by (safe_tac (set_cs addSIs [equalityI])); +by (stac ndepth_K0 1); +by (rtac zero_less_Suc 1); +qed "ntrunc_Atom"; + +goalw Univ.thy [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)"; +by (rtac ntrunc_Atom 1); +qed "ntrunc_Numb"; + +goalw Univ.thy [Scons_def,ntrunc_def] + "ntrunc (Suc k) (M$N) = ntrunc k M $ ntrunc k N"; +by (safe_tac (set_cs addSIs [equalityI,imageI])); +by (REPEAT (stac ndepth_Push_Node 3 THEN etac Suc_mono 3)); +by (REPEAT (rtac Suc_less_SucD 1 THEN + rtac (ndepth_Push_Node RS subst) 1 THEN + assume_tac 1)); +qed "ntrunc_Scons"; + +(** Injection nodes **) + +goalw Univ.thy [In0_def] "ntrunc (Suc 0) (In0 M) = {}"; +by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_0]) 1); +by (rewtac Scons_def); +by (safe_tac (set_cs addSIs [equalityI])); +qed "ntrunc_one_In0"; + +goalw Univ.thy [In0_def] + "ntrunc (Suc (Suc k)) (In0 M) = In0 (ntrunc (Suc k) M)"; +by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_Numb]) 1); +qed "ntrunc_In0"; + +goalw Univ.thy [In1_def] "ntrunc (Suc 0) (In1 M) = {}"; +by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_0]) 1); +by (rewtac Scons_def); +by (safe_tac (set_cs addSIs [equalityI])); +qed "ntrunc_one_In1"; + +goalw Univ.thy [In1_def] + "ntrunc (Suc (Suc k)) (In1 M) = In1 (ntrunc (Suc k) M)"; +by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_Numb]) 1); +qed "ntrunc_In1"; + + +(*** Cartesian Product ***) + +goalw Univ.thy [uprod_def] "!!M N. [| M:A; N:B |] ==> (M$N) : A<*>B"; +by (REPEAT (ares_tac [singletonI,UN_I] 1)); +qed "uprodI"; + +(*The general elimination rule*) +val major::prems = goalw Univ.thy [uprod_def] + "[| c : A<*>B; \ +\ !!x y. [| x:A; y:B; c=x$y |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac [major] 1); +by (REPEAT (eresolve_tac [asm_rl,singletonE,UN_E] 1 + ORELSE resolve_tac prems 1)); +qed "uprodE"; + +(*Elimination of a pair -- introduces no eigenvariables*) +val prems = goal Univ.thy + "[| (M$N) : A<*>B; [| M:A; N:B |] ==> P \ +\ |] ==> P"; +by (rtac uprodE 1); +by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Scons_inject,ssubst] 1)); +qed "uprodE2"; + + +(*** Disjoint Sum ***) + +goalw Univ.thy [usum_def] "!!M. M:A ==> In0(M) : A<+>B"; +by (fast_tac set_cs 1); +qed "usum_In0I"; + +goalw Univ.thy [usum_def] "!!N. N:B ==> In1(N) : A<+>B"; +by (fast_tac set_cs 1); +qed "usum_In1I"; + +val major::prems = goalw Univ.thy [usum_def] + "[| u : A<+>B; \ +\ !!x. [| x:A; u=In0(x) |] ==> P; \ +\ !!y. [| y:B; u=In1(y) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (REPEAT (rtac refl 1 + ORELSE eresolve_tac (prems@[imageE,ssubst]) 1)); +qed "usumE"; + + +(** Injection **) + +goalw Univ.thy [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"; + +bind_thm ("In1_not_In0", (In0_not_In1 RS not_sym)); +bind_thm ("In0_neq_In1", (In0_not_In1 RS notE)); +val In1_neq_In0 = sym RS In0_neq_In1; + +val [major] = goalw Univ.thy [In0_def] "In0(M) = In0(N) ==> M=N"; +by (rtac (major RS Scons_inject2) 1); +qed "In0_inject"; + +val [major] = goalw Univ.thy [In1_def] "In1(M) = In1(N) ==> M=N"; +by (rtac (major RS Scons_inject2) 1); +qed "In1_inject"; + + +(*** proving equality of sets and functions using ntrunc ***) + +goalw Univ.thy [ntrunc_def] "ntrunc k M <= M"; +by (fast_tac set_cs 1); +qed "ntrunc_subsetI"; + +val [major] = goalw Univ.thy [ntrunc_def] + "(!!k. ntrunc k M <= N) ==> M<=N"; +by (fast_tac (set_cs addIs [less_add_Suc1, less_add_Suc2, + major RS subsetD]) 1); +qed "ntrunc_subsetD"; + +(*A generalized form of the take-lemma*) +val [major] = goal Univ.thy "(!!k. ntrunc k M = ntrunc k N) ==> M=N"; +by (rtac equalityI 1); +by (ALLGOALS (rtac ntrunc_subsetD)); +by (ALLGOALS (rtac (ntrunc_subsetI RSN (2, subset_trans)))); +by (rtac (major RS equalityD1) 1); +by (rtac (major RS equalityD2) 1); +qed "ntrunc_equality"; + +val [major] = goalw Univ.thy [o_def] + "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"; +by (rtac (ntrunc_equality RS ext) 1); +by (rtac (major RS fun_cong) 1); +qed "ntrunc_o_equality"; + +(*** Monotonicity ***) + +goalw Univ.thy [uprod_def] "!!A B. [| A<=A'; B<=B' |] ==> A<*>B <= A'<*>B'"; +by (fast_tac set_cs 1); +qed "uprod_mono"; + +goalw Univ.thy [usum_def] "!!A B. [| A<=A'; B<=B' |] ==> A<+>B <= A'<+>B'"; +by (fast_tac set_cs 1); +qed "usum_mono"; + +goalw Univ.thy [Scons_def] "!!M N. [| M<=M'; N<=N' |] ==> M$N <= M'$N'"; +by (fast_tac set_cs 1); +qed "Scons_mono"; + +goalw Univ.thy [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)"; +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"; +by (fast_tac (set_cs addIs [select_equality] addEs [Scons_inject]) 1); +qed "Split"; + +goalw Univ.thy [Case_def] "Case c d (In0 M) = c(M)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim In0_inject, In0_neq_In1]) 1); +qed "Case_In0"; + +goalw Univ.thy [Case_def] "Case c d (In1 N) = d(N)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim In1_inject, In1_neq_In0]) 1); +qed "Case_In1"; + +(**** UN x. B(x) rules ****) + +goalw Univ.thy [ntrunc_def] "ntrunc k (UN x.f(x)) = (UN x. ntrunc k (f x))"; +by (fast_tac (set_cs addIs [equalityI]) 1); +qed "ntrunc_UN1"; + +goalw Univ.thy [Scons_def] "(UN x.f(x)) $ M = (UN x. f(x) $ M)"; +by (fast_tac (set_cs addIs [equalityI]) 1); +qed "Scons_UN1_x"; + +goalw Univ.thy [Scons_def] "M $ (UN x.f(x)) = (UN x. M $ f(x))"; +by (fast_tac (set_cs addIs [equalityI]) 1); +qed "Scons_UN1_y"; + +goalw Univ.thy [In0_def] "In0(UN x.f(x)) = (UN x. In0(f(x)))"; +br Scons_UN1_y 1; +qed "In0_UN1"; + +goalw Univ.thy [In1_def] "In1(UN x.f(x)) = (UN x. In1(f(x)))"; +br Scons_UN1_y 1; +qed "In1_UN1"; + + +(*** Equality : the diagonal relation ***) + +goalw Univ.thy [diag_def] "!!a A. [| a=b; a:A |] ==> : diag(A)"; +by (fast_tac set_cs 1); +qed "diag_eqI"; + +val diagI = refl RS diag_eqI |> standard; + +(*The general elimination rule*) +val major::prems = goalw Univ.thy [diag_def] + "[| c : diag(A); \ +\ !!x y. [| x:A; c = |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UN_E) 1); +by (REPEAT (eresolve_tac [asm_rl,singletonE] 1 ORELSE resolve_tac prems 1)); +qed "diagE"; + +(*** Equality for Cartesian Product ***) + +goalw Univ.thy [dprod_def] + "!!r s. [| :r; :s |] ==> : r<**>s"; +by (fast_tac prod_cs 1); +qed "dprodI"; + +(*The general elimination rule*) +val major::prems = goalw Univ.thy [dprod_def] + "[| c : r<**>s; \ +\ !!x y x' y'. [| : r; : s; c = |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac [major] 1); +by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, mem_splitE, singletonE])); +by (REPEAT (ares_tac prems 1 ORELSE hyp_subst_tac 1)); +qed "dprodE"; + + +(*** Equality for Disjoint Sum ***) + +goalw Univ.thy [dsum_def] "!!r. :r ==> : r<++>s"; +by (fast_tac prod_cs 1); +qed "dsum_In0I"; + +goalw Univ.thy [dsum_def] "!!r. :s ==> : r<++>s"; +by (fast_tac prod_cs 1); +qed "dsum_In1I"; + +val major::prems = goalw Univ.thy [dsum_def] + "[| w : r<++>s; \ +\ !!x x'. [| : r; w = |] ==> P; \ +\ !!y y'. [| : s; w = |] ==> P \ +\ |] ==> P"; +by (cut_facts_tac [major] 1); +by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, UnE, mem_splitE, singletonE])); +by (DEPTH_SOLVE (ares_tac prems 1 ORELSE hyp_subst_tac 1)); +qed "dsumE"; + + +val univ_cs = + prod_cs addSIs [diagI, uprodI, dprodI] + addIs [usum_In0I, usum_In1I, dsum_In0I, dsum_In1I] + addSEs [diagE, uprodE, dprodE, usumE, dsumE]; + + +(*** Monotonicity ***) + +goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<**>s <= r'<**>s'"; +by (fast_tac univ_cs 1); +qed "dprod_mono"; + +goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<++>s <= r'<++>s'"; +by (fast_tac univ_cs 1); +qed "dsum_mono"; + + +(*** Bounding theorems ***) + +goal Univ.thy "diag(A) <= Sigma A (%x.A)"; +by (fast_tac univ_cs 1); +qed "diag_subset_Sigma"; + +goal Univ.thy "(Sigma A (%x.B) <**> Sigma C (%x.D)) <= Sigma (A<*>C) (%z. B<*>D)"; +by (fast_tac univ_cs 1); +qed "dprod_Sigma"; + +val dprod_subset_Sigma = [dprod_mono, dprod_Sigma] MRS subset_trans |>standard; + +(*Dependent version*) +goal Univ.thy + "(Sigma A B <**> Sigma C D) <= Sigma (A<*>C) (Split(%x y. B(x)<*>D(y)))"; +by (safe_tac univ_cs); +by (stac Split 1); +by (fast_tac univ_cs 1); +qed "dprod_subset_Sigma2"; + +goal Univ.thy "(Sigma A (%x.B) <++> Sigma C (%x.D)) <= Sigma (A<+>C) (%z. B<+>D)"; +by (fast_tac univ_cs 1); +qed "dsum_Sigma"; + +val dsum_subset_Sigma = [dsum_mono, dsum_Sigma] MRS subset_trans |> standard; + + +(*** Domain ***) + +goal Univ.thy "fst `` diag(A) = A"; +by (fast_tac (prod_cs addIs [equalityI, diagI] addSEs [diagE]) 1); +qed "fst_image_diag"; + +goal Univ.thy "fst `` (r<**>s) = (fst``r) <*> (fst``s)"; +by (fast_tac (prod_cs addIs [equalityI, uprodI, dprodI] + addSEs [uprodE, dprodE]) 1); +qed "fst_image_dprod"; + +goal Univ.thy "fst `` (r<++>s) = (fst``r) <+> (fst``s)"; +by (fast_tac (prod_cs addIs [equalityI, usum_In0I, usum_In1I, + dsum_In0I, dsum_In1I] + addSEs [usumE, dsumE]) 1); +qed "fst_image_dsum"; + +val fst_image_simps = [fst_image_diag, fst_image_dprod, fst_image_dsum]; +val fst_image_ss = univ_ss addsimps fst_image_simps; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/Univ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Univ.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,103 @@ +(* Title: HOL/Univ.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Move LEAST to Nat.thy??? Could it be defined for all types 'a::ord? + +Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat) + +Defines "Cartesian Product" and "Disjoint Sum" as set operations. +Could <*> be generalized to a general summation (Sigma)? +*) + +Univ = Arith + Sum + + +(** lists, trees will be sets of nodes **) + +subtype (Node) + 'a node = "{p. EX f x k. p = nat, x::'a+nat> & f(k)=0}" + +types + 'a item = "'a node set" + +consts + Least :: "(nat=>bool) => nat" (binder "LEAST " 10) + + apfst :: "['a=>'c, 'a*'b] => 'c*'b" + Push :: "[nat, nat=>nat] => (nat=>nat)" + + Push_Node :: "[nat, 'a node] => 'a node" + ndepth :: "'a node => nat" + + Atom :: "('a+nat) => 'a item" + Leaf :: "'a => 'a item" + Numb :: "nat => 'a item" + "$" :: "['a item, 'a item]=> 'a item" (infixr 60) + In0,In1 :: "'a item => 'a item" + + ntrunc :: "[nat, 'a item] => 'a item" + + "<*>" :: "['a item set, 'a item set]=> 'a item set" (infixr 80) + "<+>" :: "['a item set, 'a item set]=> 'a item set" (infixr 70) + + Split :: "[['a item, 'a item]=>'b, 'a item] => 'b" + Case :: "[['a item]=>'b, ['a item]=>'b, 'a item] => 'b" + + diag :: "'a set => ('a * 'a)set" + "<**>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \ +\ => ('a item * 'a item)set" (infixr 80) + "<++>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \ +\ => ('a item * 'a item)set" (infixr 70) + +defs + + (*least number operator*) + Least_def "Least(P) == @k. P(k) & (ALL j. j ~P(j))" + + Push_Node_def "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))" + + (*crude "lists" of nats -- needed for the constructions*) + apfst_def "apfst == (%f. split(%x y. ))" + Push_def "Push == (%b h. nat_case (Suc b) h)" + + (** operations on S-expressions -- sets of nodes **) + + (*S-expression constructors*) + Atom_def "Atom == (%x. {Abs_Node(<%k.0, x>)})" + Scons_def "M$N == (Push_Node(0) `` M) Un (Push_Node(Suc(0)) `` N)" + + (*Leaf nodes, with arbitrary or nat labels*) + Leaf_def "Leaf == Atom o Inl" + Numb_def "Numb == Atom o Inr" + + (*Injections of the "disjoint sum"*) + In0_def "In0(M) == Numb(0) $ M" + In1_def "In1(M) == Numb(Suc(0)) $ M" + + (*the set of nodes with depth less than k*) + ndepth_def "ndepth(n) == split (%f x. LEAST k. f(k)=0) (Rep_Node n)" + ntrunc_def "ntrunc k N == {n. n:N & ndepth(n)B == UN x:A. UN y:B. { (x$y) }" + usum_def "A<+>B == In0``A Un In1``B" + + (*the corresponding eliminators*) + Split_def "Split c M == @u. ? x y. M = x$y & u = c x y" + + Case_def "Case c d M == @u. (? x . M = In0(x) & u = c(x)) \ +\ | (? y . M = In1(y) & u = d(y))" + + + (** diagonal sets and equality for the "universe" **) + + diag_def "diag(A) == UN x:A. {}" + + dprod_def "r<**>s == UN u:r. split (%x x'. \ +\ UN v:s. split (%y y'. {}) v) u" + + dsum_def "r<++>s == (UN u:r. split (%x x'. {}) u) Un \ +\ (UN v:s. split (%y y'. {}) v)" + +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/WF.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/WF.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,198 @@ +(* Title: HOL/wf.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +For wf.thy. Well-founded Recursion +*) + +open WF; + +val H_cong = read_instantiate [("f","H::[?'a, ?'a=>?'b]=>?'b")] + (standard(refl RS cong RS cong)); +val H_cong1 = refl RS H_cong; + +(*Restriction to domain A. If r is well-founded over A then wf(r)*) +val [prem1,prem2] = goalw WF.thy [wf_def] + "[| r <= Sigma A (%u.A); \ +\ !!x P. [| ! x. (! y. : r --> P(y)) --> P(x); x:A |] ==> P(x) |] \ +\ ==> wf(r)"; +by (strip_tac 1); +by (rtac allE 1); +by (assume_tac 1); +by (best_tac (HOL_cs addSEs [prem1 RS subsetD RS SigmaE2] addIs [prem2]) 1); +qed "wfI"; + +val major::prems = goalw WF.thy [wf_def] + "[| wf(r); \ +\ !!x.[| ! y. : r --> P(y) |] ==> P(x) \ +\ |] ==> P(a)"; +by (rtac (major RS spec RS mp RS spec) 1); +by (fast_tac (HOL_cs addEs prems) 1); +qed "wf_induct"; + +(*Perform induction on i, then prove the wf(r) subgoal using prems. *) +fun wf_ind_tac a prems i = + EVERY [res_inst_tac [("a",a)] wf_induct i, + rename_last_tac a ["1"] (i+1), + ares_tac prems i]; + +val prems = goal WF.thy "[| wf(r); :r; :r |] ==> P"; +by (subgoal_tac "! x. :r --> :r --> P" 1); +by (fast_tac (HOL_cs addIs prems) 1); +by (wf_ind_tac "a" prems 1); +by (fast_tac set_cs 1); +qed "wf_asym"; + +val prems = goal WF.thy "[| wf(r); : r |] ==> P"; +by (rtac wf_asym 1); +by (REPEAT (resolve_tac prems 1)); +qed "wf_anti_refl"; + +(*transitive closure of a WF relation is WF!*) +val [prem] = goal WF.thy "wf(r) ==> wf(r^+)"; +by (rewtac wf_def); +by (strip_tac 1); +(*must retain the universal formula for later use!*) +by (rtac allE 1 THEN assume_tac 1); +by (etac mp 1); +by (res_inst_tac [("a","x")] (prem RS wf_induct) 1); +by (rtac (impI RS allI) 1); +by (etac tranclE 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "wf_trancl"; + + +(** cut **) + +(*This rewrite rule works upon formulae; thus it requires explicit use of + H_cong to expose the equality*) +goalw WF.thy [cut_def] + "(cut f r x = cut g r x) = (!y. :r --> f(y)=g(y))"; +by(simp_tac (HOL_ss addsimps [expand_fun_eq] + setloop (split_tac [expand_if])) 1); +qed "cut_cut_eq"; + +goalw WF.thy [cut_def] "!!x. :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] + "!!f. [| is_recfun r a H f; ~:r |] ==> f(b) = (@z.True)"; +by (etac ssubst 1); +by(asm_simp_tac HOL_ss 1); +qed "is_recfun_undef"; + +(*eresolve_tac transD solves :r using transitivity AT MOST ONCE + mp amd allE instantiate induction hypotheses*) +fun indhyp_tac hyps = + ares_tac (TrueI::hyps) ORELSE' + (cut_facts_tac hyps THEN' + DEPTH_SOLVE_1 o (ares_tac [TrueI] ORELSE' + eresolve_tac [transD, mp, allE])); + +(*** NOTE! some simplifications need a different finish_tac!! ***) +fun indhyp_tac hyps = + resolve_tac (TrueI::refl::hyps) ORELSE' + (cut_facts_tac hyps THEN' + DEPTH_SOLVE_1 o (ares_tac [TrueI] ORELSE' + eresolve_tac [transD, mp, allE])); +val wf_super_ss = HOL_ss setsolver indhyp_tac; + +val prems = goalw WF.thy [is_recfun_def,cut_def] + "[| wf(r); trans(r); is_recfun r a H f; is_recfun r b H g |] ==> \ + \ :r --> :r --> f(x)=g(x)"; +by (cut_facts_tac prems 1); +by (etac wf_induct 1); +by (REPEAT (rtac impI 1 ORELSE etac ssubst 1)); +by (asm_simp_tac (wf_super_ss addcongs [if_cong]) 1); +qed "is_recfun_equal_lemma"; +bind_thm ("is_recfun_equal", (is_recfun_equal_lemma RS mp RS mp)); + + +val prems as [wfr,transr,recfa,recgb,_] = goalw WF.thy [cut_def] + "[| wf(r); trans(r); \ +\ is_recfun r a H f; is_recfun r b H g; :r |] ==> \ +\ cut f r b = g"; +val gundef = recgb RS is_recfun_undef +and fisg = recgb RS (recfa RS (transr RS (wfr RS is_recfun_equal))); +by (cut_facts_tac prems 1); +by (rtac ext 1); +by (asm_simp_tac (wf_super_ss addsimps [gundef,fisg] + setloop (split_tac [expand_if])) 1); +qed "is_recfun_cut"; + +(*** Main Existence Lemma -- Basic Properties of the_recfun ***) + +val prems = goalw WF.thy [the_recfun_def] + "is_recfun r a H f ==> is_recfun r a H (the_recfun r a H)"; +by (res_inst_tac [("P", "is_recfun r a H")] selectI 1); +by (resolve_tac prems 1); +qed "is_the_recfun"; + +val prems = goal WF.thy + "[| wf(r); trans(r) |] ==> is_recfun r a H (the_recfun r a H)"; +by (cut_facts_tac prems 1); +by (wf_ind_tac "a" prems 1); +by (res_inst_tac [("f", "cut (%y. wftrec r y H) r a1")] is_the_recfun 1); +by (rewrite_goals_tac [is_recfun_def, wftrec_def]); +by (rtac (cut_cut_eq RS ssubst) 1); +(*Applying the substitution: must keep the quantified assumption!!*) +by (EVERY1 [strip_tac, rtac H_cong1, rtac allE, atac, + etac (mp RS ssubst), atac]); +by (fold_tac [is_recfun_def]); +by (asm_simp_tac (wf_super_ss addsimps[cut_apply,is_recfun_cut,cut_cut_eq]) 1); +qed "unfold_the_recfun"; + + +(*Beware incompleteness of unification!*) +val prems = goal WF.thy + "[| wf(r); trans(r); :r; :r |] \ +\ ==> the_recfun r a H c = the_recfun r b H c"; +by (DEPTH_SOLVE (ares_tac (prems@[is_recfun_equal,unfold_the_recfun]) 1)); +qed "the_recfun_equal"; + +val prems = goal WF.thy + "[| wf(r); trans(r); :r |] \ +\ ==> cut (the_recfun r a H) r b = the_recfun r b H"; +by (REPEAT (ares_tac (prems@[is_recfun_cut,unfold_the_recfun]) 1)); +qed "the_recfun_cut"; + +(*** Unfolding wftrec ***) + +goalw WF.thy [wftrec_def] + "!!r. [| wf(r); trans(r) |] ==> \ +\ wftrec r a H = H a (cut (%x.wftrec r x H) r a)"; +by (EVERY1 [stac (rewrite_rule [is_recfun_def] unfold_the_recfun), + REPEAT o atac, rtac H_cong1]); +by (asm_simp_tac (HOL_ss addsimps [cut_cut_eq,the_recfun_cut]) 1); +qed "wftrec"; + +(*Unused but perhaps interesting*) +val prems = goal WF.thy + "[| wf(r); trans(r); !!f x. H x (cut f r x) = H x f |] ==> \ +\ wftrec r a H = H a (%x.wftrec r x H)"; +by (rtac (wftrec RS trans) 1); +by (REPEAT (resolve_tac prems 1)); +qed "wftrec2"; + +(** Removal of the premise trans(r) **) + +goalw WF.thy [wfrec_def] + "!!r. wf(r) ==> wfrec r a H = H a (cut (%x.wfrec r x H) r a)"; +by (etac (wf_trancl RS wftrec RS ssubst) 1); +by (rtac trans_trancl 1); +by (rtac (refl RS H_cong) 1); (*expose the equality of cuts*) +by (simp_tac (HOL_ss addsimps [cut_cut_eq, cut_apply, r_into_trancl]) 1); +qed "wfrec"; + +(*This form avoids giant explosions in proofs. NOTE USE OF == *) +val rew::prems = goal WF.thy + "[| !!x. f(x)==wfrec r x H; wf(r) |] ==> f(a) = H a (cut (%x.f(x)) r a)"; +by (rewtac rew); +by (REPEAT (resolve_tac (prems@[wfrec]) 1)); +qed "def_wfrec"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/WF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/WF.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,30 @@ +(* Title: HOL/wf.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1992 University of Cambridge + +Well-founded Recursion +*) + +WF = Trancl + +consts + wf :: "('a * 'a)set => bool" + cut :: "['a => 'b, ('a * 'a)set, 'a] => 'a => 'b" + wftrec,wfrec :: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'b" + is_recfun :: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b, 'a=>'b] => bool" + the_recfun :: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'a=>'b" + +defs + wf_def "wf(r) == (!P. (!x. (!y. :r --> P(y)) --> P(x)) --> (!x.P(x)))" + + cut_def "cut f r x == (%y. if (:r) (f y) (@z.True))" + + is_recfun_def "is_recfun r a H f == (f = cut (%x.(H x (cut f r x))) r a)" + + the_recfun_def "the_recfun r a H == (@f.is_recfun r a H f)" + + wftrec_def "wftrec r a H == H a (the_recfun r a H)" + + (*version not requiring transitivity*) + wfrec_def "wfrec r a H == wftrec (trancl r) a (%x f.(H x (cut f r x)))" +end diff -r 196ca0973a6d -r ff1574a81019 src/HOL/add_ind_def.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/add_ind_def.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,244 @@ +(* Title: HOL/add_ind_def.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Fixedpoint definition module -- for Inductive/Coinductive Definitions + +Features: +* least or greatest fixedpoints +* user-specified product and sum constructions +* mutually recursive definitions +* definitions involving arbitrary monotone operators +* automatically proves introduction and elimination rules + +The recursive sets must *already* be declared as constants in parent theory! + + Introduction rules have the form + [| ti:M(Sj), ..., P(x), ... |] ==> t: Sk |] + where M is some monotone operator (usually the identity) + P(x) is any (non-conjunctive) side condition on the free variables + ti, t are any terms + Sj, Sk are two of the sets being defined in mutual recursion + +Sums are used only for mutual recursion; +Products are used only to derive "streamlined" induction rules for relations + +Nestings of disjoint sum types: + (a+(b+c)) for 3, ((a+b)+(c+d)) for 4, ((a+b)+(c+(d+e))) for 5, + ((a+(b+c))+(d+(e+f))) for 6 +*) + +signature FP = (** Description of a fixed point operator **) + sig + val oper : string * typ * term -> term (*fixed point operator*) + val Tarski : thm (*Tarski's fixed point theorem*) + val induct : thm (*induction/coinduction rule*) + end; + + +signature ADD_INDUCTIVE_DEF = + sig + val add_fp_def_i : term list * term list -> theory -> theory + end; + + + +(*Declares functions to add fixedpoint/constructor defs to a theory*) +functor Add_inductive_def_Fun (Fp: FP) : ADD_INDUCTIVE_DEF = +struct +open Logic Ind_Syntax; + +(*internal version*) +fun add_fp_def_i (rec_tms, intr_tms) thy = + let + val sign = sign_of thy; + + (*recT and rec_params should agree for all mutually recursive components*) + val rec_hds = map head_of rec_tms; + + val _ = assert_all is_Const rec_hds + (fn t => "Recursive set not previously declared as constant: " ^ + Sign.string_of_term sign t); + + (*Now we know they are all Consts, so get their names, type and params*) + val rec_names = map (#1 o dest_Const) rec_hds + and (Const(_,recT),rec_params) = strip_comb (hd rec_tms); + + val _ = assert_all Syntax.is_identifier rec_names + (fn a => "Name of recursive set not an identifier: " ^ a); + + local (*Checking the introduction rules*) + val intr_sets = map (#2 o rule_concl_msg sign) intr_tms; + fun intr_ok set = + case head_of set of Const(a,_) => a mem rec_names | _ => false; + in + val _ = assert_all intr_ok intr_sets + (fn t => "Conclusion of rule does not name a recursive set: " ^ + Sign.string_of_term sign t); + end; + + val _ = assert_all is_Free rec_params + (fn t => "Param in recursion term not a free variable: " ^ + Sign.string_of_term sign t); + + (*** Construct the lfp definition ***) + val mk_variant = variant (foldr add_term_names (intr_tms,[])); + + val z = mk_variant"z" and X = mk_variant"X" and w = mk_variant"w"; + + (*Probably INCORRECT for mutual recursion!*) + val domTs = summands(dest_setT (body_type recT)); + val dom_sumT = fold_bal mk_sum domTs; + val dom_set = mk_setT dom_sumT; + + val freez = Free(z, dom_sumT) + and freeX = Free(X, dom_set); + (*type of w may be any of the domTs*) + + fun dest_tprop (Const("Trueprop",_) $ P) = P + | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ + Sign.string_of_term sign Q); + + (*Makes a disjunct from an introduction rule*) + fun lfp_part intr = (*quantify over rule's free vars except parameters*) + let val prems = map dest_tprop (strip_imp_prems intr) + val _ = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds + val exfrees = term_frees intr \\ rec_params + val zeq = eq_const dom_sumT $ freez $ (#1 (rule_concl intr)) + in foldr mk_exists (exfrees, fold_bal (app conj) (zeq::prems)) end; + + (*The Part(A,h) terms -- compose injections to make h*) + fun mk_Part (Bound 0, _) = freeX (*no mutual rec, no Part needed*) + | mk_Part (h, domT) = + let val goodh = mend_sum_types (h, dom_sumT) + and Part_const = + Const("Part", [dom_set, domT-->dom_sumT]---> dom_set) + in Part_const $ freeX $ Abs(w,domT,goodh) end; + + (*Access to balanced disjoint sums via injections*) + val parts = map mk_Part + (accesses_bal (ap Inl, ap Inr, Bound 0) (length domTs) ~~ + domTs); + + (*replace each set by the corresponding Part(A,h)*) + val part_intrs = map (subst_free (rec_tms ~~ parts) o lfp_part) intr_tms; + + val lfp_rhs = Fp.oper(X, dom_sumT, + mk_Collect(z, dom_sumT, + fold_bal (app disj) part_intrs)) + + val _ = seq (fn rec_hd => deny (rec_hd occs lfp_rhs) + "Illegal occurrence of recursion operator") + rec_hds; + + (*** Make the new theory ***) + + (*A key definition: + If no mutual recursion then it equals the one recursive set. + If mutual recursion then it differs from all the recursive sets. *) + val big_rec_name = space_implode "_" rec_names; + + (*Big_rec... is the union of the mutually recursive sets*) + val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); + + (*The individual sets must already be declared*) + val axpairs = map mk_defpair + ((big_rec_tm, lfp_rhs) :: + (case parts of + [_] => [] (*no mutual recursion*) + | _ => rec_tms ~~ (*define the sets as Parts*) + map (subst_atomic [(freeX, big_rec_tm)]) parts)); + + val _ = seq (writeln o Sign.string_of_term sign o #2) axpairs + + in thy |> add_defs_i axpairs end + + +(****************************************************************OMITTED + +(*Expects the recursive sets to have been defined already. + con_ty_lists specifies the constructors in the form (name,prems,mixfix) *) +fun add_constructs_def (rec_names, con_ty_lists) thy = +* let +* val _ = writeln" Defining the constructor functions..."; +* val case_name = "f"; (*name for case variables*) + +* (** Define the constructors **) + +* (*The empty tuple is 0*) +* fun mk_tuple [] = Const("0",iT) +* | mk_tuple args = foldr1 mk_Pair args; + +* fun mk_inject n k u = access_bal(ap Inl, ap Inr, u) n k; + +* val npart = length rec_names; (*total # of mutually recursive parts*) + +* (*Make constructor definition; kpart is # of this mutually recursive part*) +* fun mk_con_defs (kpart, con_ty_list) = +* let val ncon = length con_ty_list (*number of constructors*) + fun mk_def (((id,T,syn), name, args, prems), kcon) = + (*kcon is index of constructor*) + mk_defpair (list_comb (Const(name,T), args), + mk_inject npart kpart + (mk_inject ncon kcon (mk_tuple args))) +* in map mk_def (con_ty_list ~~ (1 upto ncon)) end; + +* (** Define the case operator **) + +* (*Combine split terms using case; yields the case operator for one part*) +* fun call_case case_list = +* let fun call_f (free,args) = + ap_split T free (map (#2 o dest_Free) args) +* in fold_bal (app sum_case) (map call_f case_list) end; + +* (** Generating function variables for the case definition + Non-identifiers (e.g. infixes) get a name of the form f_op_nnn. **) + +* (*Treatment of a single constructor*) +* fun add_case (((id,T,syn), name, args, prems), (opno,cases)) = + if Syntax.is_identifier id + then (opno, + (Free(case_name ^ "_" ^ id, T), args) :: cases) + else (opno+1, + (Free(case_name ^ "_op_" ^ string_of_int opno, T), args) :: + cases) + +* (*Treatment of a list of constructors, for one part*) +* fun add_case_list (con_ty_list, (opno,case_lists)) = + let val (opno',case_list) = foldr add_case (con_ty_list, (opno,[])) + in (opno', case_list :: case_lists) end; + +* (*Treatment of all parts*) +* val (_, case_lists) = foldr add_case_list (con_ty_lists, (1,[])); + +* val big_case_typ = flat (map (map (#2 o #1)) con_ty_lists) ---> (iT-->iT); + +* val big_rec_name = space_implode "_" rec_names; + +* val big_case_name = big_rec_name ^ "_case"; + +* (*The list of all the function variables*) +* val big_case_args = flat (map (map #1) case_lists); + +* val big_case_tm = + list_comb (Const(big_case_name, big_case_typ), big_case_args); + +* val big_case_def = mk_defpair + (big_case_tm, fold_bal (app sum_case) (map call_case case_lists)); + +* (** Build the new theory **) + +* val const_decs = + (big_case_name, big_case_typ, NoSyn) :: map #1 (flat con_ty_lists); + +* val axpairs = + big_case_def :: flat (map mk_con_defs ((1 upto npart) ~~ con_ty_lists)) + +* in thy |> add_consts_i const_decs |> add_defs_i axpairs end; +****************************************************************) +end; + + + + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/datatype.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/datatype.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,488 @@ +(* Title: HOL/datatype.ML + ID: $Id$ + Author: Max Breitling, Carsten Clasohm, Tobias Nipkow, Norbert Voelker + Copyright 1995 TU Muenchen +*) + + +(*used for constructor parameters*) +datatype dt_type = dtVar of string | + dtTyp of dt_type list * string | + dtRek of dt_type list * string; + +structure Datatype = +struct +local + +val mysort = sort; +open ThyParse HOLogic; +exception Impossible; +exception RecError of string; + +val is_dtRek = (fn dtRek _ => true | _ => false); +fun opt_parens s = if s = "" then "" else enclose "(" ")" s; + +(* ----------------------------------------------------------------------- *) +(* Derivation of the primrec combinator application from the equations *) + +(* substitute fname(ls,xk,rs) by yk(ls,rs) in t for (xk,yk) in pairs *) + +fun subst_apps (_,_) [] t = t + | subst_apps (fname,rpos) pairs t = + let + fun subst (Abs(a,T,t)) = Abs(a,T,subst t) + | subst (funct $ body) = + let val (f,b) = strip_comb (funct$body) + in + if is_Const f andalso fst(dest_Const f) = fname + then + let val (ls,rest) = (take(rpos,b), drop(rpos,b)); + val (xk,rs) = (hd rest,tl rest) + handle LIST _ => raise RecError "not enough arguments \ + \ in recursive application on rhs" + in + (case assoc (pairs,xk) of + None => raise RecError + ("illegal occurence of " ^ fname ^ " on rhs") + | Some(U) => list_comb(U,map subst (ls @ rs))) + end + else list_comb(f, map subst b) + end + | subst(t) = t + in subst t end; + +(* abstract rhs *) + +fun abst_rec (fname,rpos,tc,ls,cargs,rs,rhs) = + let val rargs = (map fst o + (filter (fn (a,T) => is_dtRek T))) (cargs ~~ tc); + val subs = map (fn (s,T) => (s,dummyT)) + (rev(rename_wrt_term rhs rargs)); + val subst_rhs = subst_apps (fname,rpos) + (map Free rargs ~~ map Free subs) rhs; + in + list_abs_free (cargs @ subs @ ls @ rs, subst_rhs) + end; + +(* parsing the prim rec equations *) + +fun dest_eq ( Const("Trueprop",_) $ (Const ("op =",_) $ lhs $ rhs)) + = (lhs, rhs) + | dest_eq _ = raise RecError "not a proper equation"; + +fun dest_rec eq = + let val (lhs,rhs) = dest_eq eq; + val (name,args) = strip_comb lhs; + val (ls',rest) = take_prefix is_Free args; + val (middle,rs') = take_suffix is_Free rest; + val rpos = length ls'; + val (c,cargs') = strip_comb (hd middle) + handle LIST "hd" => raise RecError "constructor missing"; + val (ls,cargs,rs) = (map dest_Free ls', map dest_Free cargs' + , map dest_Free rs') + handle TERM ("dest_Free",_) => + raise RecError "constructor has illegal argument in pattern"; + in + if length middle > 1 then + raise RecError "more than one non-variable in pattern" + else if not(null(findrep (map fst (ls @ rs @ cargs)))) then + raise RecError "repeated variable name in pattern" + else (fst(dest_Const name) handle TERM _ => + raise RecError "function is not declared as constant in theory" + ,rpos,ls,fst( dest_Const c),cargs,rs,rhs) + end; + +(* check function specified for all constructors and sort function terms *) + +fun check_and_sort (n,its) = + if length its = n + then map snd (mysort (fn ((i : int,_),(j,_)) => i + error("Primrec definition error: " ^ s ^ ":\n" + ^ " " ^ Sign.string_of_term (sign_of thy) eq1); + val tcs = map (fn (_,c,T,_,_) => (c,T)) cs'; + val cs = map fst tcs; + fun trans_recs' _ [] = [] + | trans_recs' cis (eq::eqs) = + let val (name,rpos,ls,c,cargs,rs,rhs) = dest_rec eq; + val tc = assoc(tcs,c); + val i = (1 + find (c,cs)) handle LIST "find" => 0; + in + if name <> name1 then + raise RecError "function names inconsistent" + else if rpos <> rpos1 then + raise RecError "position of rec. argument inconsistent" + else if i = 0 then + raise RecError "illegal argument in pattern" + else if i mem cis then + raise RecError "constructor already occured as pattern " + else (i,abst_rec (name,rpos,the tc,ls,cargs,rs,rhs)) + :: trans_recs' (i::cis) eqs + end + handle RecError s => + error("Primrec definition error\n" ^ s ^ "\n" + ^ " " ^ Sign.string_of_term (sign_of thy) eq); + in ( name1, ls1 + , check_and_sort (length cs, trans_recs' [] (eq1::eqs))) + end ; + +in + fun add_datatype (typevars, tname, cons_list') thy = + let + fun typid(dtRek(_,id)) = id + | typid(dtVar s) = implode (tl (explode s)) + | typid(dtTyp(_,id)) = id; + + fun index_vnames(vn::vns,tab) = + (case assoc(tab,vn) of + None => if vn mem vns + then (vn^"1") :: index_vnames(vns,(vn,2)::tab) + else vn :: index_vnames(vns,tab) + | Some(i) => (vn^(string_of_int i)) :: + index_vnames(vns,(vn,i+1)::tab)) + | index_vnames([],tab) = []; + + fun mk_var_names types = index_vnames(map typid types,[]); + + (*search for free type variables and convert recursive *) + fun analyse_types (cons, types, syn) = + let fun analyse(t as dtVar v) = + if t mem typevars then t + else error ("Free type variable " ^ v ^ " on rhs.") + | analyse(dtTyp(typl,s)) = + if tname <> s then dtTyp(analyses typl, s) + else if typevars = typl then dtRek(typl, s) + else error (s ^ " used in different ways") + | analyse(dtRek _) = raise Impossible + and analyses ts = map analyse ts; + in (cons, Syntax.const_name cons syn, analyses types, + mk_var_names types, syn) + end; + + (*test if all elements are recursive, i.e. if the type is empty*) + + fun non_empty (cs : ('a * 'b * dt_type list * 'c *'d) list) = + not(forall (exists is_dtRek o #3) cs) orelse + error("Empty datatype not allowed!"); + + val cons_list = map analyse_types cons_list'; + val dummy = non_empty cons_list; + val num_of_cons = length cons_list; + + (* Auxiliary functions to construct argument and equation lists *) + + (*generate 'var_n, ..., var_m'*) + fun Args(var, delim, n, m) = + space_implode delim (map (fn n => var^string_of_int(n)) (n upto m)); + + fun C_exp name vns = name ^ opt_parens(space_implode ") (" vns); + + (*Arg_eqs([x1,...,xn],[y1,...,yn]) = "x1 = y1 & ... & xn = yn" *) + fun arg_eqs vns vns' = + let fun mkeq(x,x') = x ^ "=" ^ x' + in space_implode " & " (map mkeq (vns~~vns')) end; + + (*Pretty printers for type lists; + pp_typlist1: parentheses, pp_typlist2: brackets*) + fun pp_typ (dtVar s) = s + | pp_typ (dtTyp (typvars, id)) = + if null typvars then id else (pp_typlist1 typvars) ^ id + | pp_typ (dtRek (typvars, id)) = (pp_typlist1 typvars) ^ id + and + pp_typlist' ts = commas (map pp_typ ts) + and + pp_typlist1 ts = if null ts then "" else parens (pp_typlist' ts); + + fun pp_typlist2 ts = if null ts then "" else brackets (pp_typlist' ts); + + (* Generate syntax translation for case rules *) + fun calc_xrules c_nr y_nr ((_, name, _, vns, _) :: cs) = + let val arity = length vns; + val body = "z" ^ string_of_int(c_nr); + val args1 = if arity=0 then "" + else parens (Args ("y", ") (", y_nr, y_nr+arity-1)); + val args2 = if arity=0 then "" + else "% " ^ Args ("y", " ", y_nr, y_nr+arity-1) + ^ ". "; + val (rest1,rest2) = + if null cs then ("","") + else let val (h1, h2) = calc_xrules (c_nr+1) (y_nr+arity) cs + in (" | " ^ h1, " " ^ h2) end; + in (name ^ args1 ^ " => " ^ body ^ rest1, + "(" ^ args2 ^ body ^ rest2 ^ ")") + end + | calc_xrules _ _ [] = raise Impossible; + + val xrules = + let val (first_part, scnd_part) = calc_xrules 1 1 cons_list + in [("logic", "case x of " ^ first_part) <-> + ("logic", tname ^ "_case (" ^ scnd_part ^ ") x" )] + end; + + (*type declarations for constructors*) + fun const_type (id, _, typlist, _, syn) = + (id, + (if null typlist then "" else pp_typlist2 typlist ^ " => ") ^ + pp_typlist1 typevars ^ tname, syn); + + + fun assumpt (dtRek _ :: ts, v :: vs ,found) = + let val h = if found then ";P(" ^ v ^ ")" else "[| P(" ^ v ^ ")" + in h ^ (assumpt (ts, vs, true)) end + | assumpt (t :: ts, v :: vs, found) = assumpt (ts, vs, found) + | assumpt ([], [], found) = if found then "|] ==>" else "" + | assumpt _ = raise Impossible; + + fun t_inducting ((_, name, types, vns, _) :: cs) = + let + val h = if null types then " P(" ^ name ^ ")" + else " !!" ^ (space_implode " " vns) ^ "." ^ + (assumpt (types, vns, false)) ^ + "P(" ^ C_exp name vns ^ ")"; + val rest = t_inducting cs; + in if rest = "" then h else h ^ "; " ^ rest end + | t_inducting [] = ""; + + fun t_induct cl typ_name = + "[|" ^ t_inducting cl ^ "|] ==> P(" ^ typ_name ^ ")"; + + fun gen_typlist typevar f ((_, _, ts, _, _) :: cs) = + let val h = if (length ts) > 0 + then pp_typlist2(f ts) ^ "=>" + else "" + in h ^ typevar ^ "," ^ (gen_typlist typevar f cs) end + | gen_typlist _ _ [] = ""; + + +(* -------------------------------------------------------------------- *) +(* The case constant and rules *) + + val t_case = tname ^ "_case"; + + fun case_rule n (id, name, _, vns, _) = + let val args = opt_parens(space_implode ") (" vns) + in (t_case ^ "_" ^ id, + t_case ^ "(" ^ Args("f", ") (", 1, num_of_cons) + ^ ") (" ^ name ^ args ^ ") = f"^string_of_int(n) ^ args) + end + + fun case_rules n (c :: cs) = case_rule n c :: case_rules(n+1) cs + | case_rules _ [] = []; + + val datatype_arity = length typevars; + + val types = [(tname, datatype_arity, NoSyn)]; + + val arities = + let val term_list = replicate datatype_arity termS; + in [(tname, term_list, termS)] + end; + + val datatype_name = pp_typlist1 typevars ^ tname; + + val new_tvar_name = variant (map (fn dtVar s => s) typevars) "'z"; + + val case_const = + (t_case, + "[" ^ gen_typlist new_tvar_name I cons_list + ^ pp_typlist1 typevars ^ tname ^ "] =>" ^ new_tvar_name, + NoSyn); + + val rules_case = case_rules 1 cons_list; + +(* -------------------------------------------------------------------- *) +(* The prim-rec combinator *) + + val t_rec = tname ^ "_rec" + +(* adding type variables for dtRek types to end of list of dt_types *) + + fun add_reks ts = + ts @ map (fn _ => dtVar new_tvar_name) (filter is_dtRek ts); + +(* positions of the dtRek types in a list of dt_types, starting from 1 *) + fun rek_vars ts vns = map snd (filter (is_dtRek o fst) (ts ~~ vns)) + + fun rec_rule n (id,name,ts,vns,_) = + let val args = space_implode ") (" vns + val fargs = Args("f",") (",1,num_of_cons) + fun rarg vn = ") (" ^ t_rec ^ parens(fargs ^ ") (" ^ vn) + val rargs = implode (map rarg (rek_vars ts vns)) + in + ( t_rec ^ "_" ^ id + , t_rec ^ parens(fargs ^ ") (" ^ name ^ (opt_parens args)) ^ " = f" + ^ string_of_int(n) ^ opt_parens (args ^ rargs)) + end + + fun rec_rules n (c::cs) = rec_rule n c :: rec_rules (n+1) cs + | rec_rules _ [] = []; + + val rec_const = + (t_rec, + "[" ^ (gen_typlist new_tvar_name add_reks cons_list) + ^ (pp_typlist1 typevars) ^ tname ^ "] =>" ^ new_tvar_name, + NoSyn); + + val rules_rec = rec_rules 1 cons_list + +(* -------------------------------------------------------------------- *) + val consts = + map const_type cons_list + @ (if num_of_cons < dtK then [] + else [(tname ^ "_ord", datatype_name ^ "=>nat", NoSyn)]) + @ [case_const,rec_const]; + + + fun Ci_ing ((id, name, _, vns, _) :: cs) = + if null vns then Ci_ing cs + else let val vns' = variantlist(vns,vns) + in ("inject_" ^ id, + "(" ^ (C_exp name vns) ^ "=" ^ (C_exp name vns') + ^ ") = (" ^ (arg_eqs vns vns') ^ ")") :: (Ci_ing cs) + end + | Ci_ing [] = []; + + fun Ci_negOne (id1,name1,_,vns1,_) (id2,name2,_,vns2,_) = + let val vns2' = variantlist(vns2,vns1) + val ax = C_exp name1 vns1 ^ "~=" ^ C_exp name2 vns2' + in (id1 ^ "_not_" ^ id2, ax) end; + + fun Ci_neg1 [] = [] + | Ci_neg1 (c1::cs) = (map (Ci_negOne c1) cs) @ Ci_neg1 cs; + + fun suc_expr n = + if n=0 then "0" else "Suc(" ^ suc_expr(n-1) ^ ")"; + + fun Ci_neg2() = + let val ord_t = tname ^ "_ord"; + val cis = cons_list ~~ (0 upto (num_of_cons - 1)) + fun Ci_neg2equals ((id, name, _, vns, _), n) = + let val ax = ord_t ^ "(" ^ (C_exp name vns) ^ ") = " ^ (suc_expr n) + in (ord_t ^ "_" ^ id, ax) end + in (ord_t ^ "_distinct", ord_t^"(x) ~= "^ord_t^"(y) ==> x ~= y") :: + (map Ci_neg2equals cis) + end; + + val rules_distinct = if num_of_cons < dtK then Ci_neg1 cons_list + else Ci_neg2(); + + val rules_inject = Ci_ing cons_list; + + val rule_induct = (tname ^ "_induct", t_induct cons_list tname); + + val rules = rule_induct :: + (rules_inject @ rules_distinct @ rules_case @ rules_rec); + + fun add_primrec eqns thy = + let val rec_comb = Const(t_rec,dummyT) + val teqns = map (fn neq => snd(read_axm (sign_of thy) neq)) eqns + val (fname,ls,fns) = trans_recs thy cons_list teqns + val rhs = + list_abs_free + (ls @ [(tname,dummyT)] + ,list_comb(rec_comb + , fns @ map Bound (0 ::(length ls downto 1)))); + val sg = sign_of thy; + val defpair = mk_defpair (Const(fname,dummyT),rhs) + val defpairT as (_, _ $ Const(_,T) $ _ ) = inferT_axm sg defpair; + val varT = Type.varifyT T; + val ftyp = the (Sign.const_type sg fname); + in + if Type.typ_instance (#tsig(Sign.rep_sg sg), ftyp, varT) + then add_defs_i [defpairT] thy + else error("Primrec definition error: \ntype of " ^ fname + ^ " is not instance of type deduced from equations") + end; + + in + (thy + |> add_types types + |> add_arities arities + |> add_consts consts + |> add_trrules xrules + |> add_axioms rules,add_primrec) + end +end +end + +(* +Informal description of functions used in datatype.ML for the Isabelle/HOL +implementation of prim. rec. function definitions. (N. Voelker, Feb. 1995) + +* subst_apps (fname,rpos) pairs t: + substitute the term + fname(ls,xk,rs) + by + yk(ls,rs) + in t for (xk,yk) in pairs, where rpos = length ls. + Applied with : + fname = function name + rpos = position of recursive argument + pairs = list of pairs (xk,yk), where + xk are the rec. arguments of the constructor in the pattern, + yk is a variable with name derived from xk + t = rhs of equation + +* abst_rec (fname,rpos,tc,ls,cargs,rs,rhs) + - filter recursive arguments from constructor arguments cargs, + - perform substitutions on rhs, + - derive list subs of new variable names yk for use in subst_apps, + - abstract rhs with respect to cargs, subs, ls and rs. + +* dest_eq t + destruct a term denoting an equation into lhs and rhs. + +* dest_req eq + destruct an equation of the form + name (vl1..vlrpos, Ci(vi1..vin), vr1..vrn) = rhs + into + - function name (name) + - position of the first non-variable parameter (rpos) + - the list of first rpos parameters (ls = [vl1..vlrpos]) + - the constructor (fst( dest_Const c) = Ci) + - the arguments of the constructor (cargs = [vi1..vin]) + - the rest of the variables in the pattern (rs = [vr1..vrn]) + - the right hand side of the equation (rhs). + +* check_and_sort (n,its) + check that n = length its holds, and sort elements of its by + first component. + +* trans_recs thy cs' (eq1::eqs) + destruct eq1 into name1, rpos1, ls1, etc.. + get constructor list with and without type (tcs resp. cs) from cs', + for every equation: + destruct it into (name,rpos,ls,c,cargs,rs,rhs) + get typed constructor tc from c and tcs + determine the index i of the constructor + check function name and position of rec. argument by comparison + with first equation + check for repeated variable names in pattern + derive function term f_i which is used as argument of the rec. combinator + sort the terms f_i according to i and return them together + with the function name and the parameter of the definition (ls). + +* Application: + + The rec. combinator is applied to the function terms resulting from + trans_rec. This results in a function which takes the recursive arg. + as first parameter and then the arguments corresponding to ls. The + order of parameters is corrected by setting the rhs equal to + + list_abs_free + (ls @ [(tname,dummyT)] + ,list_comb(rec_comb + , fns @ map Bound (0 ::(length ls downto 1)))); + + Note the de-Bruijn indices counting the number of lambdas between the + variable and its binding. +*) diff -r 196ca0973a6d -r ff1574a81019 src/HOL/equalities.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/equalities.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,333 @@ +(* Title: HOL/equalities + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Equalities involving union, intersection, inclusion, etc. +*) + +writeln"File HOL/equalities"; + +val eq_cs = set_cs addSIs [equalityI]; + +(** The membership relation, : **) + +goal Set.thy "x ~: {}"; +by(fast_tac set_cs 1); +qed "in_empty"; + +goal Set.thy "x : insert y A = (x=y | x:A)"; +by(fast_tac set_cs 1); +qed "in_insert"; + +(** insert **) + +goal Set.thy "!!a. a:A ==> insert a A = A"; +by (fast_tac eq_cs 1); +qed "insert_absorb"; + +goal Set.thy "(insert x A <= B) = (x:B & A <= B)"; +by (fast_tac set_cs 1); +qed "insert_subset"; + +(** Image **) + +goal Set.thy "f``{} = {}"; +by (fast_tac eq_cs 1); +qed "image_empty"; + +goal Set.thy "f``insert a B = insert (f a) (f``B)"; +by (fast_tac eq_cs 1); +qed "image_insert"; + +(** Binary Intersection **) + +goal Set.thy "A Int A = A"; +by (fast_tac eq_cs 1); +qed "Int_absorb"; + +goal Set.thy "A Int B = B Int A"; +by (fast_tac eq_cs 1); +qed "Int_commute"; + +goal Set.thy "(A Int B) Int C = A Int (B Int C)"; +by (fast_tac eq_cs 1); +qed "Int_assoc"; + +goal Set.thy "{} Int B = {}"; +by (fast_tac eq_cs 1); +qed "Int_empty_left"; + +goal Set.thy "A Int {} = {}"; +by (fast_tac eq_cs 1); +qed "Int_empty_right"; + +goal Set.thy "A Int (B Un C) = (A Int B) Un (A Int C)"; +by (fast_tac eq_cs 1); +qed "Int_Un_distrib"; + +goal Set.thy "(A<=B) = (A Int B = A)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +qed "subset_Int_eq"; + +(** Binary Union **) + +goal Set.thy "A Un A = A"; +by (fast_tac eq_cs 1); +qed "Un_absorb"; + +goal Set.thy "A Un B = B Un A"; +by (fast_tac eq_cs 1); +qed "Un_commute"; + +goal Set.thy "(A Un B) Un C = A Un (B Un C)"; +by (fast_tac eq_cs 1); +qed "Un_assoc"; + +goal Set.thy "{} Un B = B"; +by(fast_tac eq_cs 1); +qed "Un_empty_left"; + +goal Set.thy "A Un {} = A"; +by(fast_tac eq_cs 1); +qed "Un_empty_right"; + +goal Set.thy "insert a B Un C = insert a (B Un C)"; +by(fast_tac eq_cs 1); +qed "Un_insert_left"; + +goal Set.thy "(A Int B) Un C = (A Un C) Int (B Un C)"; +by (fast_tac eq_cs 1); +qed "Un_Int_distrib"; + +goal Set.thy + "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; +by (fast_tac eq_cs 1); +qed "Un_Int_crazy"; + +goal Set.thy "(A<=B) = (A Un B = B)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +qed "subset_Un_eq"; + +goal Set.thy "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)"; +by (fast_tac eq_cs 1); +qed "subset_insert_iff"; + +goal Set.thy "(A Un B = {}) = (A = {} & B = {})"; +by (fast_tac (eq_cs addEs [equalityCE]) 1); +qed "Un_empty"; + +(** Simple properties of Compl -- complement of a set **) + +goal Set.thy "A Int Compl(A) = {}"; +by (fast_tac eq_cs 1); +qed "Compl_disjoint"; + +goal Set.thy "A Un Compl(A) = {x.True}"; +by (fast_tac eq_cs 1); +qed "Compl_partition"; + +goal Set.thy "Compl(Compl(A)) = A"; +by (fast_tac eq_cs 1); +qed "double_complement"; + +goal Set.thy "Compl(A Un B) = Compl(A) Int Compl(B)"; +by (fast_tac eq_cs 1); +qed "Compl_Un"; + +goal Set.thy "Compl(A Int B) = Compl(A) Un Compl(B)"; +by (fast_tac eq_cs 1); +qed "Compl_Int"; + +goal Set.thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"; +by (fast_tac eq_cs 1); +qed "Compl_UN"; + +goal Set.thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"; +by (fast_tac eq_cs 1); +qed "Compl_INT"; + +(*Halmos, Naive Set Theory, page 16.*) + +goal Set.thy "((A Int B) Un C = A Int (B Un C)) = (C<=A)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +qed "Un_Int_assoc_eq"; + + +(** Big Union and Intersection **) + +goal Set.thy "Union({}) = {}"; +by (fast_tac eq_cs 1); +qed "Union_empty"; + +goal Set.thy "Union(insert a B) = a Un Union(B)"; +by (fast_tac eq_cs 1); +qed "Union_insert"; + +goal Set.thy "Union(A Un B) = Union(A) Un Union(B)"; +by (fast_tac eq_cs 1); +qed "Union_Un_distrib"; + +goal Set.thy "Union(A Int B) <= Union(A) Int Union(B)"; +by (fast_tac set_cs 1); +qed "Union_Int_subset"; + +val prems = goal Set.thy + "(Union(C) Int A = {}) = (! B:C. B Int A = {})"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +qed "Union_disjoint"; + +goal Set.thy "Inter(A Un B) = Inter(A) Int Inter(B)"; +by (best_tac eq_cs 1); +qed "Inter_Un_distrib"; + +(** Unions and Intersections of Families **) + +(*Basic identities*) + +goal Set.thy "Union(range(f)) = (UN x.f(x))"; +by (fast_tac eq_cs 1); +qed "Union_range_eq"; + +goal Set.thy "Inter(range(f)) = (INT x.f(x))"; +by (fast_tac eq_cs 1); +qed "Inter_range_eq"; + +goal Set.thy "Union(B``A) = (UN x:A. B(x))"; +by (fast_tac eq_cs 1); +qed "Union_image_eq"; + +goal Set.thy "Inter(B``A) = (INT x:A. B(x))"; +by (fast_tac eq_cs 1); +qed "Inter_image_eq"; + +goal Set.thy "!!A. a: A ==> (UN y:A. c) = c"; +by (fast_tac eq_cs 1); +qed "UN_constant"; + +goal Set.thy "!!A. a: A ==> (INT y:A. c) = c"; +by (fast_tac eq_cs 1); +qed "INT_constant"; + +goal Set.thy "(UN x.B) = B"; +by (fast_tac eq_cs 1); +qed "UN1_constant"; + +goal Set.thy "(INT x.B) = B"; +by (fast_tac eq_cs 1); +qed "INT1_constant"; + +goal Set.thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})"; +by (fast_tac eq_cs 1); +qed "UN_eq"; + +(*Look: it has an EXISTENTIAL quantifier*) +goal Set.thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})"; +by (fast_tac eq_cs 1); +qed "INT_eq"; + +(*Distributive laws...*) + +goal Set.thy "A Int Union(B) = (UN C:B. A Int C)"; +by (fast_tac eq_cs 1); +qed "Int_Union"; + +(* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: + Union of a family of unions **) +goal Set.thy "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)"; +by (fast_tac eq_cs 1); +qed "Un_Union_image"; + +(*Equivalent version*) +goal Set.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; +by (fast_tac eq_cs 1); +qed "UN_Un_distrib"; + +goal Set.thy "A Un Inter(B) = (INT C:B. A Un C)"; +by (fast_tac eq_cs 1); +qed "Un_Inter"; + +goal Set.thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)"; +by (best_tac eq_cs 1); +qed "Int_Inter_image"; + +(*Equivalent version*) +goal Set.thy "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; +by (fast_tac eq_cs 1); +qed "INT_Int_distrib"; + +(*Halmos, Naive Set Theory, page 35.*) +goal Set.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; +by (fast_tac eq_cs 1); +qed "Int_UN_distrib"; + +goal Set.thy "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; +by (fast_tac eq_cs 1); +qed "Un_INT_distrib"; + +goal Set.thy + "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; +by (fast_tac eq_cs 1); +qed "Int_UN_distrib2"; + +goal Set.thy + "(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; +by (fast_tac eq_cs 1); +qed "Un_INT_distrib2"; + +(** Simple properties of Diff -- set difference **) + +goal Set.thy "A-A = {}"; +by (fast_tac eq_cs 1); +qed "Diff_cancel"; + +goal Set.thy "{}-A = {}"; +by (fast_tac eq_cs 1); +qed "empty_Diff"; + +goal Set.thy "A-{} = A"; +by (fast_tac eq_cs 1); +qed "Diff_empty"; + +(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*) +goal Set.thy "A - insert a B = A - B - {a}"; +by (fast_tac eq_cs 1); +qed "Diff_insert"; + +(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*) +goal Set.thy "A - insert a B = A - {a} - B"; +by (fast_tac eq_cs 1); +qed "Diff_insert2"; + +val prems = goal Set.thy "a:A ==> insert a (A-{a}) = A"; +by (fast_tac (eq_cs addSIs prems) 1); +qed "insert_Diff"; + +goal Set.thy "A Int (B-A) = {}"; +by (fast_tac eq_cs 1); +qed "Diff_disjoint"; + +goal Set.thy "!!A. A<=B ==> A Un (B-A) = B"; +by (fast_tac eq_cs 1); +qed "Diff_partition"; + +goal Set.thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)"; +by (fast_tac eq_cs 1); +qed "double_diff"; + +goal Set.thy "A - (B Un C) = (A-B) Int (A-C)"; +by (fast_tac eq_cs 1); +qed "Diff_Un"; + +goal Set.thy "A - (B Int C) = (A-B) Un (A-C)"; +by (fast_tac eq_cs 1); +qed "Diff_Int"; + +val set_ss = set_ss addsimps + [in_empty,in_insert,insert_subset, + Int_absorb,Int_empty_left,Int_empty_right, + Un_absorb,Un_empty_left,Un_empty_right,Un_empty, + UN1_constant,image_empty, + Compl_disjoint,double_complement, + Union_empty,Union_insert,empty_subsetI,subset_refl, + Diff_cancel,empty_Diff,Diff_empty,Diff_disjoint]; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/equalities.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/equalities.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,9 @@ +(* Title: HOL/equalities + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Equalities involving union, intersection, inclusion, etc. +*) + +equalities = subset diff -r 196ca0973a6d -r ff1574a81019 src/HOL/hologic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/hologic.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,86 @@ +(* Title: HOL/hologic.ML + ID: $Id$ + Author: Lawrence C Paulson and Markus Wenzel + +Abstract syntax operations for HOL. +*) + +signature HOLOGIC = +sig + val termC: class + val termS: sort + val termTVar: typ + val boolT: typ + val mk_setT: typ -> typ + val dest_setT: typ -> typ + val mk_Trueprop: term -> term + val dest_Trueprop: term -> term + val conj: term + val disj: term + val imp: term + val eq_const: typ -> term + val all_const: typ -> term + val exists_const: typ -> term + val Collect_const: typ -> term + val mk_eq: term * term -> term + val mk_all: string * typ * term -> term + val mk_exists: string * typ * term -> term + val mk_Collect: string * typ * term -> term + val mk_mem: term * term -> term +end; + +structure HOLogic: HOLOGIC = +struct + +(* classes *) + +val termC: class = "term"; +val termS: sort = [termC]; + + +(* types *) + +val termTVar = TVar (("'a", 0), termS); + +val boolT = Type ("bool", []); + +fun mk_setT T = Type ("set", [T]); + +fun dest_setT (Type ("set", [T])) = T + | dest_setT T = raise_type "dest_setT: set type expected" [T] []; + + +(* terms *) + +val Trueprop = Const ("Trueprop", boolT --> propT); + +fun mk_Trueprop P = Trueprop $ P; + +fun dest_Trueprop (Const ("Trueprop", _) $ P) = P + | dest_Trueprop t = raise_term "dest_Trueprop" [t]; + + +val conj = Const ("op &", [boolT, boolT] ---> boolT) +and disj = Const ("op |", [boolT, boolT] ---> boolT) +and imp = Const ("op -->", [boolT, boolT] ---> boolT); + +fun eq_const T = Const ("op =", [T, T] ---> boolT); +fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u; + +fun all_const T = Const ("All", [T --> boolT] ---> boolT); +fun mk_all (x, T, P) = all_const T $ absfree (x, T, P); + +fun exists_const T = Const ("Ex", [T --> boolT] ---> boolT); +fun mk_exists (x, T, P) = exists_const T $ absfree (x, T, P); + +fun Collect_const T = Const ("Collect", [T --> boolT] ---> mk_setT T); +fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T, t); + +fun mk_mem (x, A) = + let val setT = fastype_of A in + Const ("op :", [dest_setT setT, setT] ---> boolT) $ x $ A + end; + + +end; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/ind_syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ind_syntax.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,124 @@ +(* Title: HOL/ind_syntax.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Abstract Syntax functions for Inductive Definitions +See also hologic.ML and ../Pure/section-utils.ML +*) + +(*The structure protects these items from redeclaration (somewhat!). The + datatype definitions in theory files refer to these items by name! +*) +structure Ind_Syntax = +struct + +(** Abstract syntax definitions for HOL **) + +open HOLogic; + +fun Int_const T = + let val sT = mk_setT T + in Const("op Int", [sT,sT]--->sT) end; + +fun mk_exists (Free(x,T),P) = exists_const T $ (absfree (x,T,P)); + +fun mk_all (Free(x,T),P) = all_const T $ (absfree (x,T,P)); + +(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *) +fun mk_all_imp (A,P) = + let val T = dest_setT (fastype_of A) + in all_const T $ Abs("v", T, imp $ (mk_mem (Bound 0, A)) $ (P $ Bound 0)) + end; + +(** Cartesian product type **) + +val unitT = Type("unit",[]); + +fun mk_prod (T1,T2) = Type("*", [T1,T2]); + +(*Maps the type T1*...*Tn to [T1,...,Tn], if nested to the right*) +fun factors (Type("*", [T1,T2])) = T1 :: factors T2 + | factors T = [T]; + +(*Make a correctly typed ordered pair*) +fun mk_Pair (t1,t2) = + let val T1 = fastype_of t1 + and T2 = fastype_of t2 + in Const("Pair", [T1, T2] ---> mk_prod(T1,T2)) $ t1 $ t2 end; + +fun split_const(Ta,Tb,Tc) = + Const("split", [[Ta,Tb]--->Tc, mk_prod(Ta,Tb)] ---> Tc); + +(*Given u expecting arguments of types [T1,...,Tn], create term of + type T1*...*Tn => Tc using split. Here * associates to the LEFT*) +fun ap_split_l Tc u [ ] = Abs("null", unitT, u) + | ap_split_l Tc u [_] = u + | ap_split_l Tc u (Ta::Tb::Ts) = ap_split_l Tc (split_const(Ta,Tb,Tc) $ u) + (mk_prod(Ta,Tb) :: Ts); + +(*Given u expecting arguments of types [T1,...,Tn], create term of + type T1*...*Tn => i using split. Here * associates to the RIGHT*) +fun ap_split Tc u [ ] = Abs("null", unitT, u) + | ap_split Tc u [_] = u + | ap_split Tc u [Ta,Tb] = split_const(Ta,Tb,Tc) $ u + | ap_split Tc u (Ta::Ts) = + split_const(Ta, foldr1 mk_prod Ts, Tc) $ + (Abs("v", Ta, ap_split Tc (u $ Bound(length Ts - 2)) Ts)); + +(** Disjoint sum type **) + +fun mk_sum (T1,T2) = Type("+", [T1,T2]); +val Inl = Const("Inl", dummyT) +and Inr = Const("Inr", dummyT); (*correct types added later!*) +(*val elim = Const("case", [iT-->iT, iT-->iT, iT]--->iT)*) + +fun summands (Type("+", [T1,T2])) = summands T1 @ summands T2 + | summands T = [T]; + +(*Given the destination type, fills in correct types of an Inl/Inr nest*) +fun mend_sum_types (h,T) = + (case (h,T) of + (Const("Inl",_) $ h1, Type("+", [T1,T2])) => + Const("Inl", T1 --> T) $ (mend_sum_types (h1, T1)) + | (Const("Inr",_) $ h2, Type("+", [T1,T2])) => + Const("Inr", T2 --> T) $ (mend_sum_types (h2, T2)) + | _ => h); + + + +(*simple error-checking in the premises of an inductive definition*) +fun chk_prem rec_hd (Const("op &",_) $ _ $ _) = + error"Premises may not be conjuctive" + | chk_prem rec_hd (Const("op :",_) $ t $ X) = + deny (Logic.occs(rec_hd,t)) "Recursion term on left of member symbol" + | chk_prem rec_hd t = + deny (Logic.occs(rec_hd,t)) "Recursion term in side formula"; + +(*Return the conclusion of a rule, of the form t:X*) +fun rule_concl rl = + let val Const("Trueprop",_) $ (Const("op :",_) $ t $ X) = + Logic.strip_imp_concl rl + in (t,X) end; + +(*As above, but return error message if bad*) +fun rule_concl_msg sign rl = rule_concl rl + handle Bind => error ("Ill-formed conclusion of introduction rule: " ^ + Sign.string_of_term sign rl); + +(*For simplifying the elimination rule*) +val sumprod_free_SEs = + Pair_inject :: + map make_elim [(*Inl_neq_Inr, Inr_neq_Inl, Inl_inject, Inr_inject*)]; + +(*For deriving cases rules. + read_instantiate replaces a propositional variable by a formula variable*) +val equals_CollectD = + read_instantiate [("W","?Q")] + (make_elim (equalityD1 RS subsetD RS CollectD)); + +(*Delete needless equality assumptions*) +val refl_thin = prove_goal HOL.thy "!!P. [| a=a; P |] ==> P" + (fn _ => [assume_tac 1]); + +end; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/indrule.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/indrule.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,184 @@ +(* Title: HOL/indrule.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Induction rule module -- for Inductive/Coinductive Definitions + +Proves a strong induction rule and a mutual induction rule +*) + +signature INDRULE = + sig + val induct : thm (*main induction rule*) + val mutual_induct : thm (*mutual induction rule*) + end; + + +functor Indrule_Fun + (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end and + Intr_elim: INTR_ELIM) : INDRULE = +struct +open Logic Ind_Syntax Inductive Intr_elim; + +val sign = sign_of thy; + +val (Const(_,recT),rec_params) = strip_comb (hd rec_tms); + +val elem_type = dest_setT (body_type recT); +val domTs = summands(elem_type); +val big_rec_name = space_implode "_" rec_names; +val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); + +val _ = writeln " Proving the induction rules..."; + +(*** Prove the main induction rule ***) + +val pred_name = "P"; (*name for predicate variables*) + +val big_rec_def::part_rec_defs = Intr_elim.defs; + +(*Used to express induction rules: adds induction hypotheses. + ind_alist = [(rec_tm1,pred1),...] -- associates predicates with rec ops + prem is a premise of an intr rule*) +fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ + (Const("op :",_)$t$X), iprems) = + (case gen_assoc (op aconv) (ind_alist, X) of + Some pred => prem :: mk_Trueprop (pred $ t) :: iprems + | None => (*possibly membership in M(rec_tm), for M monotone*) + let fun mk_sb (rec_tm,pred) = + (case binder_types (fastype_of pred) of + [T] => (rec_tm, + Int_const T $ rec_tm $ (Collect_const T $ pred)) + | _ => error + "Bug: add_induct_prem called with non-unary predicate") + in subst_free (map mk_sb ind_alist) prem :: iprems end) + | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; + +(*Make a premise of the induction rule.*) +fun induct_prem ind_alist intr = + let val quantfrees = map dest_Free (term_frees intr \\ rec_params) + val iprems = foldr (add_induct_prem ind_alist) + (strip_imp_prems intr,[]) + val (t,X) = rule_concl intr + val (Some pred) = gen_assoc (op aconv) (ind_alist, X) + val concl = mk_Trueprop (pred $ t) + in list_all_free (quantfrees, list_implies (iprems,concl)) end + handle Bind => error"Recursion term not found in conclusion"; + +(*Avoids backtracking by delivering the correct premise to each goal*) +fun ind_tac [] 0 = all_tac + | ind_tac(prem::prems) i = + DEPTH_SOLVE_1 (ares_tac [Part_eqI, prem, refl] i) THEN + ind_tac prems (i-1); + +val pred = Free(pred_name, elem_type --> boolT); + +val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) intr_tms; + +val quant_induct = + prove_goalw_cterm part_rec_defs + (cterm_of sign (list_implies (ind_prems, + mk_Trueprop (mk_all_imp(big_rec_tm,pred))))) + (fn prems => + [rtac (impI RS allI) 1, + etac raw_induct 1, + REPEAT (FIRSTGOAL (eresolve_tac [IntE, CollectE, exE, conjE, disjE] + ORELSE' hyp_subst_tac)), + REPEAT (FIRSTGOAL (eresolve_tac [PartE, CollectE])), + ind_tac (rev prems) (length prems)]) + handle e => print_sign_exn sign e; + +(*** Prove the simultaneous induction rule ***) + +(*Make distinct predicates for each inductive set. + Splits cartesian products in domT, IF nested to the right! *) + +(*Given a recursive set and its domain, return the "split" predicate + and a conclusion for the simultaneous induction rule*) +fun mk_predpair (rec_tm,domT) = + let val rec_name = (#1 o dest_Const o head_of) rec_tm + val T = factors domT ---> boolT + val pfree = Free(pred_name ^ "_" ^ rec_name, T) + val frees = mk_frees "za" (binder_types T) + val qconcl = + foldr mk_all (frees, + imp $ (mk_mem (foldr1 mk_Pair frees, rec_tm)) + $ (list_comb (pfree,frees))) + in (ap_split boolT pfree (binder_types T), + qconcl) + end; + +val (preds,qconcls) = split_list (map mk_predpair (rec_tms~~domTs)); + +(*Used to form simultaneous induction lemma*) +fun mk_rec_imp (rec_tm,pred) = + imp $ (mk_mem (Bound 0, rec_tm)) $ (pred $ Bound 0); + +(*To instantiate the main induction rule*) +val induct_concl = + mk_Trueprop(mk_all_imp(big_rec_tm, + Abs("z", elem_type, + fold_bal (app conj) + (map mk_rec_imp (rec_tms~~preds))))) +and mutual_induct_concl = mk_Trueprop(fold_bal (app conj) qconcls); + +val lemma = (*makes the link between the two induction rules*) + prove_goalw_cterm part_rec_defs + (cterm_of sign (mk_implies (induct_concl,mutual_induct_concl))) + (fn prems => + [cut_facts_tac prems 1, + REPEAT (eresolve_tac [asm_rl, conjE, PartE, mp] 1 + ORELSE resolve_tac [allI, impI, conjI, Part_eqI, refl] 1 + ORELSE dresolve_tac [spec, mp, splitD] 1)]) + handle e => print_sign_exn sign e; + +(*Mutual induction follows by freeness of Inl/Inr.*) + +(*Removes Collects caused by M-operators in the intro rules*) +val cmonos = [subset_refl RS Int_Collect_mono] RL monos RLN (2,[rev_subsetD]); + +(*Avoids backtracking by delivering the correct premise to each goal*) +fun mutual_ind_tac [] 0 = all_tac + | mutual_ind_tac(prem::prems) i = + DETERM + (SELECT_GOAL + ((*unpackage and use "prem" in the corresponding place*) + REPEAT (FIRSTGOAL + (etac conjE ORELSE' eq_mp_tac ORELSE' + ares_tac [impI, conjI])) + (*prem is not allowed in the REPEAT, lest it loop!*) + THEN TRYALL (rtac prem) + THEN REPEAT + (FIRSTGOAL (ares_tac [impI] ORELSE' + eresolve_tac (mp::cmonos))) + (*prove remaining goals by contradiction*) + THEN rewrite_goals_tac (con_defs@part_rec_defs) + THEN DEPTH_SOLVE (eresolve_tac (PartE :: sumprod_free_SEs) 1)) + i) + THEN mutual_ind_tac prems (i-1); + +val mutual_induct_split = + prove_goalw_cterm [] + (cterm_of sign + (list_implies (map (induct_prem (rec_tms~~preds)) intr_tms, + mutual_induct_concl))) + (fn prems => + [rtac (quant_induct RS lemma) 1, + mutual_ind_tac (rev prems) (length prems)]) + handle e => print_sign_exn sign e; + +(*Attempts to remove all occurrences of split*) +val split_tac = + REPEAT (SOMEGOAL (FIRST' [rtac splitI, + dtac splitD, + etac splitE, + bound_hyp_subst_tac])) + THEN prune_params_tac; + +(*strip quantifier*) +val induct = standard (quant_induct RS spec RSN (2,rev_mp)); + +val mutual_induct = rule_by_tactic split_tac mutual_induct_split; + +end; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/intr_elim.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/intr_elim.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,141 @@ +(* Title: HOL/intr_elim.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Introduction/elimination rule module -- for Inductive/Coinductive Definitions +*) + +signature INDUCTIVE_ARG = (** Description of a (co)inductive def **) + sig + val thy : theory (*new theory with inductive defs*) + val monos : thm list (*monotonicity of each M operator*) + val con_defs : thm list (*definitions of the constructors*) + end; + +(*internal items*) +signature INDUCTIVE_I = + sig + val rec_tms : term list (*the recursive sets*) + val intr_tms : term list (*terms for the introduction rules*) + end; + +signature INTR_ELIM = + sig + val thy : theory (*copy of input theory*) + val defs : thm list (*definitions made in thy*) + val mono : thm (*monotonicity for the lfp definition*) + val unfold : thm (*fixed-point equation*) + val intrs : thm list (*introduction rules*) + val elim : thm (*case analysis theorem*) + val raw_induct : thm (*raw induction rule from Fp.induct*) + val mk_cases : thm list -> string -> thm (*generates case theorems*) + val rec_names : string list (*names of recursive sets*) + end; + +(*prove intr/elim rules for a fixedpoint definition*) +functor Intr_elim_Fun + (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end + and Fp: FP) : INTR_ELIM = +struct +open Logic Inductive Ind_Syntax; + +val rec_names = map (#1 o dest_Const o head_of) rec_tms; +val big_rec_name = space_implode "_" rec_names; + +val _ = deny (big_rec_name mem map ! (stamps_of_thy thy)) + ("Definition " ^ big_rec_name ^ + " would clash with the theory of the same name!"); + +(*fetch fp definitions from the theory*) +val big_rec_def::part_rec_defs = + map (get_def thy) + (case rec_names of [_] => rec_names | _ => big_rec_name::rec_names); + + +val sign = sign_of thy; + +(********) +val _ = writeln " Proving monotonicity..."; + +val Const("==",_) $ _ $ (Const(_,fpT) $ fp_abs) = + big_rec_def |> rep_thm |> #prop |> unvarify; + +(*For the type of the argument of mono*) +val [monoT] = binder_types fpT; + +val mono = + prove_goalw_cterm [] + (cterm_of sign (mk_Trueprop (Const("mono", monoT-->boolT) $ fp_abs))) + (fn _ => + [rtac monoI 1, + REPEAT (ares_tac (basic_monos @ monos) 1)]); + +val unfold = standard (mono RS (big_rec_def RS Fp.Tarski)); + +(********) +val _ = writeln " Proving the introduction rules..."; + +fun intro_tacsf disjIn prems = + [(*insert prems and underlying sets*) + cut_facts_tac prems 1, + rtac (unfold RS ssubst) 1, + REPEAT (resolve_tac [Part_eqI,CollectI] 1), + (*Now 1-2 subgoals: the disjunction, perhaps equality.*) + rtac disjIn 1, + (*Not ares_tac, since refl must be tried before any equality assumptions; + backtracking may occur if the premises have extra variables!*) + DEPTH_SOLVE_1 (resolve_tac [refl,exI,conjI] 1 ORELSE assume_tac 1)]; + +(*combines disjI1 and disjI2 to access the corresponding nested disjunct...*) +val mk_disj_rls = + let fun f rl = rl RS disjI1 + and g rl = rl RS disjI2 + in accesses_bal(f, g, asm_rl) end; + +val intrs = map (uncurry (prove_goalw_cterm part_rec_defs)) + (map (cterm_of sign) intr_tms ~~ + map intro_tacsf (mk_disj_rls(length intr_tms))); + +(********) +val _ = writeln " Proving the elimination rule..."; + +(*Includes rules for Suc and Pair since they are common constructions*) +val elim_rls = [asm_rl, FalseE, Suc_neq_Zero, Zero_neq_Suc, + make_elim Suc_inject, + refl_thin, conjE, exE, disjE]; + +(*Breaks down logical connectives in the monotonic function*) +val basic_elim_tac = + REPEAT (SOMEGOAL (eresolve_tac (elim_rls@sumprod_free_SEs) + ORELSE' bound_hyp_subst_tac)) + THEN prune_params_tac; + +val elim = rule_by_tactic basic_elim_tac (unfold RS equals_CollectD); + +(*Applies freeness of the given constructors, which *must* be unfolded by + the given defs. Cannot simply use the local con_defs because con_defs=[] + for inference systems. +fun con_elim_tac defs = + rewrite_goals_tac defs THEN basic_elim_tac THEN fold_tac defs; + *) +fun con_elim_tac simps = + let val elim_tac = REPEAT o (eresolve_tac (elim_rls@sumprod_free_SEs)) + in ALLGOALS(EVERY'[elim_tac, + asm_full_simp_tac (nat_ss addsimps simps), + elim_tac, + REPEAT o bound_hyp_subst_tac]) + THEN prune_params_tac + end; + + +(*String s should have the form t:Si where Si is an inductive set*) +fun mk_cases defs s = + rule_by_tactic (con_elim_tac defs) + (assume_read thy s RS elim); + +val defs = big_rec_def::part_rec_defs; + +val raw_induct = standard ([big_rec_def, mono] MRS Fp.induct); +end; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/mono.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/mono.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,123 @@ +(* Title: HOL/mono + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Monotonicity of various operations +*) + +goal Set.thy "!!A B. A<=B ==> f``A <= f``B"; +by (fast_tac set_cs 1); +qed "image_mono"; + +goal Set.thy "!!A B. A<=B ==> Pow(A) <= Pow(B)"; +by (fast_tac set_cs 1); +qed "Pow_mono"; + +goal Set.thy "!!A B. A<=B ==> Union(A) <= Union(B)"; +by (fast_tac set_cs 1); +qed "Union_mono"; + +goal Set.thy "!!A B. B<=A ==> Inter(A) <= Inter(B)"; +by (fast_tac set_cs 1); +qed "Inter_anti_mono"; + +val prems = goal Set.thy + "[| A<=B; !!x. x:A ==> f(x)<=g(x) |] ==> \ +\ (UN x:A. f(x)) <= (UN x:B. g(x))"; +by (fast_tac (set_cs addIs (prems RL [subsetD])) 1); +qed "UN_mono"; + +val [prem] = goal Set.thy + "[| !!x. f(x)<=g(x) |] ==> (UN x. f(x)) <= (UN x. g(x))"; +by (fast_tac (set_cs addIs [prem RS subsetD]) 1); +qed "UN1_mono"; + +val prems = goal Set.thy + "[| B<=A; !!x. x:A ==> f(x)<=g(x) |] ==> \ +\ (INT x:A. f(x)) <= (INT x:A. g(x))"; +by (fast_tac (set_cs addIs (prems RL [subsetD])) 1); +qed "INT_anti_mono"; + +(*The inclusion is POSITIVE! *) +val [prem] = goal Set.thy + "[| !!x. f(x)<=g(x) |] ==> (INT x. f(x)) <= (INT x. g(x))"; +by (fast_tac (set_cs addIs [prem RS subsetD]) 1); +qed "INT1_mono"; + +goal Set.thy "!!A B. [| A<=C; B<=D |] ==> A Un B <= C Un D"; +by (fast_tac set_cs 1); +qed "Un_mono"; + +goal Set.thy "!!A B. [| A<=C; B<=D |] ==> A Int B <= C Int D"; +by (fast_tac set_cs 1); +qed "Int_mono"; + +goal Set.thy "!!A::'a set. [| A<=C; D<=B |] ==> A-B <= C-D"; +by (fast_tac set_cs 1); +qed "Diff_mono"; + +goal Set.thy "!!A B. A<=B ==> Compl(B) <= Compl(A)"; +by (fast_tac set_cs 1); +qed "Compl_anti_mono"; + +val prems = goal Prod.thy + "[| A<=C; !!x. x:A ==> B<=D |] ==> Sigma A (%x.B) <= Sigma C (%x.D)"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addIs (prems RL [subsetD]) + addSIs [SigmaI] + addSEs [SigmaE]) 1); +qed "Sigma_mono"; + + +(** Monotonicity of implications. For inductive definitions **) + +goal Set.thy "!!A B x. A<=B ==> x:A --> x:B"; +by (rtac impI 1); +by (etac subsetD 1); +by (assume_tac 1); +qed "in_mono"; + +goal HOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)"; +by (fast_tac HOL_cs 1); +qed "conj_mono"; + +goal HOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)"; +by (fast_tac HOL_cs 1); +qed "disj_mono"; + +goal HOL.thy "!!P1 P2 Q1 Q2.[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)"; +by (fast_tac HOL_cs 1); +qed "imp_mono"; + +goal HOL.thy "P-->P"; +by (rtac impI 1); +by (assume_tac 1); +qed "imp_refl"; + +val [PQimp] = goal HOL.thy + "[| !!x. P(x) --> Q(x) |] ==> (EX x.P(x)) --> (EX x.Q(x))"; +by (fast_tac (HOL_cs addIs [PQimp RS mp]) 1); +qed "ex_mono"; + +val [PQimp] = goal HOL.thy + "[| !!x. P(x) --> Q(x) |] ==> (ALL x.P(x)) --> (ALL x.Q(x))"; +by (fast_tac (HOL_cs addIs [PQimp RS mp]) 1); +qed "all_mono"; + +val [PQimp] = goal Set.thy + "[| !!x. P(x) --> Q(x) |] ==> Collect(P) <= Collect(Q)"; +by (fast_tac (set_cs addIs [PQimp RS mp]) 1); +qed "Collect_mono"; + +(*Used in indrule.ML*) +val [subs,PQimp] = goal Set.thy + "[| A<=B; !!x. x:A ==> P(x) --> Q(x) \ +\ |] ==> A Int Collect(P) <= B Int Collect(Q)"; +by (fast_tac (set_cs addIs [subs RS subsetD, PQimp RS mp]) 1); +qed "Int_Collect_mono"; + +(*Used in intr_elim.ML and in individual datatype definitions*) +val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono, + ex_mono, Collect_mono, Part_mono, in_mono]; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/mono.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/mono.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,8 @@ +(* Title: HOL/mono + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +*) + +mono = subset diff -r 196ca0973a6d -r ff1574a81019 src/HOL/simpdata.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/simpdata.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,163 @@ +(* Title: HOL/simpdata.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Instantiation of the generic simplifier +*) + +open Simplifier; + +local + +fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]); + +val P_imp_P_iff_True = prover "P --> (P = True)" RS mp; +val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; + +val not_P_imp_P_iff_F = prover "~P --> (P = False)" RS mp; +val not_P_imp_P_eq_False = not_P_imp_P_iff_F RS eq_reflection; + +fun atomize pairs = + let fun atoms th = + (case concl_of th of + Const("Trueprop",_) $ p => + (case head_of p of + Const(a,_) => + (case assoc(pairs,a) of + Some(rls) => flat (map atoms ([th] RL rls)) + | None => [th]) + | _ => [th]) + | _ => [th]) + in atoms end; + +fun mk_meta_eq r = case concl_of r of + Const("==",_)$_$_ => r + | _$(Const("op =",_)$_$_) => r RS eq_reflection + | _$(Const("not",_)$_) => r RS not_P_imp_P_eq_False + | _ => r RS P_imp_P_eq_True; +(* last 2 lines requires all formulae to be of the from Trueprop(.) *) + +fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th; + +val imp_cong = impI RSN + (2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))" + (fn _=> [fast_tac HOL_cs 1]) RS mp RS mp); + +val o_apply = prove_goalw HOL.thy [o_def] "(f o g)(x) = f(g(x))" + (fn _ => [rtac refl 1]); + +val simp_thms = map prover + [ "(x=x) = True", + "(~True) = False", "(~False) = True", "(~ ~ P) = P", + "(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))", + "(True=P) = P", "(P=True) = P", + "(True --> P) = P", "(False --> P) = True", + "(P --> True) = True", "(P --> P) = True", + "(P --> False) = (~P)", "(P --> ~P) = (~P)", + "(P & True) = P", "(True & P) = P", + "(P & False) = False", "(False & P) = False", "(P & P) = P", + "(P | True) = True", "(True | P) = True", + "(P | False) = P", "(False | P) = P", "(P | P) = P", + "(!x.P) = P", "(? x.P) = P", "? x. x=t", "(? x. x=t & P(x)) = P(t)", + "(P|Q --> R) = ((P-->R)&(Q-->R))" ]; + +in + +val meta_eq_to_obj_eq = prove_goal HOL.thy "x==y ==> x=y" + (fn [prem] => [rewtac prem, rtac refl 1]); + +val eq_sym_conv = prover "(x=y) = (y=x)"; + +val conj_assoc = prover "((P&Q)&R) = (P&(Q&R))"; + +val if_True = prove_goalw HOL.thy [if_def] "if True x y = x" + (fn _=>[fast_tac (HOL_cs addIs [select_equality]) 1]); + +val if_False = prove_goalw HOL.thy [if_def] "if False x y = y" + (fn _=>[fast_tac (HOL_cs addIs [select_equality]) 1]); + +val if_P = prove_goal HOL.thy "P ==> if P x y = x" + (fn [prem] => [ stac (prem RS eqTrueI) 1, rtac if_True 1 ]); + +val if_not_P = prove_goal HOL.thy "~P ==> if P x y = y" + (fn [prem] => [ stac (prem RS not_P_imp_P_iff_F) 1, rtac if_False 1 ]); + +val expand_if = prove_goal HOL.thy + "P(if Q x y) = ((Q --> P(x)) & (~Q --> P(y)))" + (fn _=> [ (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1), + rtac (if_P RS ssubst) 2, + rtac (if_not_P RS ssubst) 1, + REPEAT(fast_tac HOL_cs 1) ]); + +val if_bool_eq = prove_goal HOL.thy "if P Q R = ((P-->Q) & (~P-->R))" + (fn _ => [rtac expand_if 1]); + +infix addcongs; +fun ss addcongs congs = ss addeqcongs (congs RL [eq_reflection]); + +val mksimps_pairs = + [("op -->", [mp]), ("op &", [conjunct1,conjunct2]), + ("All", [spec]), ("True", []), ("False", []), + ("if", [if_bool_eq RS iffD1])]; + +fun mksimps pairs = map mk_meta_eq o atomize pairs o gen_all; + +val HOL_ss = empty_ss + setmksimps (mksimps mksimps_pairs) + setsolver (fn prems => resolve_tac (TrueI::refl::prems) ORELSE' atac + ORELSE' etac FalseE) + setsubgoaler asm_simp_tac + addsimps ([if_True, if_False, o_apply, conj_assoc] @ simp_thms) + addcongs [imp_cong]; + +fun split_tac splits = + mk_case_split_tac (meta_eq_to_obj_eq RS iffD2) (map mk_meta_eq splits); + +(* eliminiation of existential quantifiers in assumptions *) + +val ex_all_equiv = + let val lemma1 = prove_goal HOL.thy + "(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)" + (fn prems => [resolve_tac prems 1, etac exI 1]); + val lemma2 = prove_goalw HOL.thy [Ex_def] + "(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)" + (fn prems => [REPEAT(resolve_tac prems 1)]) + in equal_intr lemma1 lemma2 end; + +(* '&' congruence rule: not included by default! + May slow rewrite proofs down by as much as 50% *) + +val conj_cong = impI RSN + (2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))" + (fn _=> [fast_tac HOL_cs 1]) RS mp RS mp); + +(** 'if' congruence rules: neither included by default! *) + +(*Simplifies x assuming c and y assuming ~c*) +val if_cong = prove_goal HOL.thy + "[| b=c; c ==> x=u; ~c ==> y=v |] ==> if b x y = if c u v" + (fn rew::prems => + [stac rew 1, stac expand_if 1, stac expand_if 1, + fast_tac (HOL_cs addDs prems) 1]); + +(*Prevents simplification of x and y: much faster*) +val if_weak_cong = prove_goal HOL.thy + "b=c ==> if b x y = if c x y" + (fn [prem] => [rtac (prem RS arg_cong) 1]); + +(*Prevents simplification of t: much faster*) +val let_weak_cong = prove_goal HOL.thy + "a = b ==> (let x=a in t(x)) = (let x=b in t(x))" + (fn [prem] => [rtac (prem RS arg_cong) 1]); + +end; + +fun prove nm thm = qed_goal nm HOL.thy thm (fn _ => [fast_tac HOL_cs 1]); + +prove "conj_commute" "(P&Q) = (Q&P)"; +prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))"; +val conj_comms = [conj_commute, conj_left_commute]; + +prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)"; +prove "conj_disj_distribR" "((P|Q)&R) = (P&R | Q&R)"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/subset.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/subset.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,135 @@ +(* Title: HOL/subset + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Derived rules involving subsets +Union and Intersection as lattice operations +*) + +(*** insert ***) + +qed_goal "subset_insertI" Set.thy "B <= insert a B" + (fn _=> [ (rtac subsetI 1), (etac insertI2 1) ]); + +(*** Big Union -- least upper bound of a set ***) + +val prems = goal Set.thy + "B:A ==> B <= Union(A)"; +by (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1)); +qed "Union_upper"; + +val [prem] = goal Set.thy + "[| !!X. X:A ==> X<=C |] ==> Union(A) <= C"; +br subsetI 1; +by (REPEAT (eresolve_tac [asm_rl, UnionE, prem RS subsetD] 1)); +qed "Union_least"; + +(** General union **) + +val prems = goal Set.thy + "a:A ==> B(a) <= (UN x:A. B(x))"; +by (REPEAT (ares_tac (prems@[UN_I RS subsetI]) 1)); +qed "UN_upper"; + +val [prem] = goal Set.thy + "[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A. B(x)) <= C"; +br subsetI 1; +by (REPEAT (eresolve_tac [asm_rl, UN_E, prem RS subsetD] 1)); +qed "UN_least"; + +goal Set.thy "B(a) <= (UN x. B(x))"; +by (REPEAT (ares_tac [UN1_I RS subsetI] 1)); +qed "UN1_upper"; + +val [prem] = goal Set.thy "[| !!x. B(x)<=C |] ==> (UN x. B(x)) <= C"; +br subsetI 1; +by (REPEAT (eresolve_tac [asm_rl, UN1_E, prem RS subsetD] 1)); +qed "UN1_least"; + + +(*** Big Intersection -- greatest lower bound of a set ***) + +val prems = goal Set.thy "B:A ==> Inter(A) <= B"; +br subsetI 1; +by (REPEAT (resolve_tac prems 1 ORELSE etac InterD 1)); +qed "Inter_lower"; + +val [prem] = goal Set.thy + "[| !!X. X:A ==> C<=X |] ==> C <= Inter(A)"; +br (InterI RS subsetI) 1; +by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1)); +qed "Inter_greatest"; + +val prems = goal Set.thy "a:A ==> (INT x:A. B(x)) <= B(a)"; +br subsetI 1; +by (REPEAT (resolve_tac prems 1 ORELSE etac INT_D 1)); +qed "INT_lower"; + +val [prem] = goal Set.thy + "[| !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A. B(x))"; +br (INT_I RS subsetI) 1; +by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1)); +qed "INT_greatest"; + +goal Set.thy "(INT x. B(x)) <= B(a)"; +br subsetI 1; +by (REPEAT (resolve_tac prems 1 ORELSE etac INT1_D 1)); +qed "INT1_lower"; + +val [prem] = goal Set.thy + "[| !!x. C<=B(x) |] ==> C <= (INT x. B(x))"; +br (INT1_I RS subsetI) 1; +by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1)); +qed "INT1_greatest"; + +(*** Finite Union -- the least upper bound of 2 sets ***) + +goal Set.thy "A <= A Un B"; +by (REPEAT (ares_tac [subsetI,UnI1] 1)); +qed "Un_upper1"; + +goal Set.thy "B <= A Un B"; +by (REPEAT (ares_tac [subsetI,UnI2] 1)); +qed "Un_upper2"; + +val prems = goal Set.thy "[| A<=C; B<=C |] ==> A Un B <= C"; +by (cut_facts_tac prems 1); +by (DEPTH_SOLVE (ares_tac [subsetI] 1 + ORELSE eresolve_tac [UnE,subsetD] 1)); +qed "Un_least"; + +(*** Finite Intersection -- the greatest lower bound of 2 sets *) + +goal Set.thy "A Int B <= A"; +by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)); +qed "Int_lower1"; + +goal Set.thy "A Int B <= B"; +by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)); +qed "Int_lower2"; + +val prems = goal Set.thy "[| C<=A; C<=B |] ==> C <= A Int B"; +by (cut_facts_tac prems 1); +by (REPEAT (ares_tac [subsetI,IntI] 1 + ORELSE etac subsetD 1)); +qed "Int_greatest"; + +(*** Set difference ***) + +qed_goal "Diff_subset" Set.thy "A-B <= (A::'a set)" + (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac DiffE 1)) ]); + +(*** Monotonicity ***) + +val [prem] = goal Set.thy "mono(f) ==> f(A) Un f(B) <= f(A Un B)"; +by (rtac Un_least 1); +by (rtac (Un_upper1 RS (prem RS monoD)) 1); +by (rtac (Un_upper2 RS (prem RS monoD)) 1); +qed "mono_Un"; + +val [prem] = goal Set.thy "mono(f) ==> f(A Int B) <= f(A) Int f(B)"; +by (rtac Int_greatest 1); +by (rtac (Int_lower1 RS (prem RS monoD)) 1); +by (rtac (Int_lower2 RS (prem RS monoD)) 1); +qed "mono_Int"; diff -r 196ca0973a6d -r ff1574a81019 src/HOL/subset.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/subset.thy Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,7 @@ +(* Title: HOL/subset.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge +*) + +subset = Fun diff -r 196ca0973a6d -r ff1574a81019 src/HOL/subtype.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/subtype.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,141 @@ +(* Title: HOL/subtype.ML + ID: $Id$ + Author: Markus Wenzel, TU Muenchen + +Internal interface for subtype definitions. +*) + +signature SUBTYPE = +sig + val prove_nonempty: cterm -> thm list -> tactic option -> thm + val add_subtype: string -> string * string list * mixfix -> + string -> string list -> thm list -> tactic option -> theory -> theory + val add_subtype_i: string -> string * string list * mixfix -> + term -> string list -> thm list -> tactic option -> theory -> theory +end; + +structure Subtype: SUBTYPE = +struct + +open Syntax Logic HOLogic; + + +(* prove non-emptyness of a set *) (*exception ERROR*) + +val is_def = is_equals o #prop o rep_thm; + +fun prove_nonempty cset thms usr_tac = + let + val {T = setT, t = set, maxidx, sign} = rep_cterm cset; + val T = dest_setT setT; + val goal = + cterm_of sign (mk_Trueprop (mk_mem (Var (("x", maxidx + 1), T), set))); + val tac = + TRY (rewrite_goals_tac (filter is_def thms)) THEN + TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN + if_none usr_tac (TRY (ALLGOALS (fast_tac set_cs))); + in + prove_goalw_cterm [] goal (K [tac]) + end + handle ERROR => + error ("Failed to prove non-emptyness of " ^ quote (string_of_cterm cset)); + + +(* ext_subtype *) + +fun ext_subtype prep_term name (t, vs, mx) raw_set axms thms usr_tac thy = + let + val _ = require_thy thy "Set" "subtype definitions"; + val sign = sign_of thy; + + (*rhs*) + val cset = prep_term sign raw_set; + val {T = setT, t = set, ...} = rep_cterm cset; + val rhs_tfrees = term_tfrees set; + val oldT = dest_setT setT handle TYPE _ => + error ("Not a set type: " ^ quote (Sign.string_of_typ sign setT)); + + (*lhs*) + val lhs_tfrees = + map (fn v => (v, if_none (assoc (rhs_tfrees, v)) termS)) vs; + + val tname = type_name t mx; + val tlen = length vs; + val newT = Type (tname, map TFree lhs_tfrees); + + val Rep_name = "Rep_" ^ name; + val Abs_name = "Abs_" ^ name; + val setC = Const (name, setT); + val RepC = Const (Rep_name, newT --> oldT); + val AbsC = Const (Abs_name, oldT --> newT); + val x_new = Free ("x", newT); + val y_old = Free ("y", oldT); + + (*axioms*) + val rep_type = mk_Trueprop (mk_mem (RepC $ x_new, setC)); + val rep_type_inv = mk_Trueprop (mk_eq (AbsC $ (RepC $ x_new), x_new)); + val abs_type_inv = mk_implies (mk_Trueprop (mk_mem (y_old, setC)), + mk_Trueprop (mk_eq (RepC $ (AbsC $ y_old), y_old))); + + + (* errors *) + + val show_names = commas_quote o map fst; + + val illegal_vars = + if null (term_vars set) andalso null (term_tvars set) then [] + else ["Illegal schematic variable(s) on rhs"]; + + val dup_lhs_tfrees = + (case duplicates lhs_tfrees of [] => [] + | dups => ["Duplicate type variables on lhs: " ^ show_names dups]); + + val extra_rhs_tfrees = + (case gen_rems (op =) (rhs_tfrees, lhs_tfrees) of [] => [] + | extras => ["Extra type variables on rhs: " ^ show_names extras]); + + val illegal_frees = + (case term_frees set of [] => [] + | xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]); + + val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees; + in + if null errs then () + else error (cat_lines errs); + + prove_nonempty cset (map (get_axiom thy) axms @ thms) usr_tac; + + thy + |> add_types [(t, tlen, mx)] + |> add_arities + [(tname, replicate tlen logicS, logicS), + (tname, replicate tlen termS, termS)] + |> add_consts_i + [(name, setT, NoSyn), + (Rep_name, newT --> oldT, NoSyn), + (Abs_name, oldT --> newT, NoSyn)] + |> add_defs_i + [(name ^ "_def", mk_equals (setC, set))] + |> add_axioms_i + [(Rep_name, rep_type), + (Rep_name ^ "_inverse", rep_type_inv), + (Abs_name ^ "_inverse", abs_type_inv)] + end + handle ERROR => + error ("The error(s) above occurred in subtype definition " ^ quote name); + + +(* external interfaces *) + +fun cert_term sg tm = + cterm_of sg tm handle TERM (msg, _) => error msg; + +fun read_term sg str = + read_cterm sg (str, termTVar); + +val add_subtype = ext_subtype read_term; +val add_subtype_i = ext_subtype cert_term; + + +end; + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/thy_syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/thy_syntax.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,187 @@ +(* Title: HOL/thy_syntax.ML + ID: $Id$ + Author: Markus Wenzel and Lawrence C Paulson and Carsten Clasohm + +Additional theory file sections for HOL. + +TODO: + move datatype / primrec stuff to pre_datatype.ML (?) +*) + +(*the kind of distinctiveness axioms depends on number of constructors*) +val dtK = 5; (* FIXME rename?, move? *) + +structure ThySynData: THY_SYN_DATA = +struct + +open ThyParse; + + +(** subtype **) + +fun mk_subtype_decl (((((opt_name, vs), t), mx), rhs), wt) = + let + val name' = if_none opt_name t; + val name = strip_quotes name'; + in + (cat_lines [name', mk_triple (t, mk_list vs, mx), rhs, wt], + [name ^ "_def", "Rep_" ^ name, "Rep_" ^ name ^ "_inverse", + "Abs_" ^ name ^ "_inverse"]) + end; + +val subtype_decl = + optional ("(" $$-- name --$$ ")" >> Some) None -- + type_args -- name -- opt_infix --$$ "=" -- string -- opt_witness + >> mk_subtype_decl; + + + +(** (co)inductive **) + +(*co is either "" or "Co"*) +fun inductive_decl co = + let + fun mk_intr_name (s, _) = (*the "op" cancels any infix status*) + if Syntax.is_identifier s then "op " ^ s else "_"; + fun mk_params (((recs, ipairs), monos), con_defs) = + let val big_rec_name = space_implode "_" (map (scan_to_id o trim) recs) + and srec_tms = mk_list recs + and sintrs = mk_big_list (map snd ipairs) + val stri_name = big_rec_name ^ "_Intrnl" + in + (";\n\n\ + \structure " ^ stri_name ^ " =\n\ + \ let open Ind_Syntax in\n\ + \ struct\n\ + \ val _ = writeln \"" ^ co ^ + "Inductive definition " ^ big_rec_name ^ "\"\n\ + \ val rec_tms\t= map (readtm (sign_of thy) termTVar) " + ^ srec_tms ^ "\n\ + \ and intr_tms\t= map (readtm (sign_of thy) propT)\n" + ^ sintrs ^ "\n\ + \ end\n\ + \ end;\n\n\ + \val thy = thy |> " ^ co ^ "Ind.add_fp_def_i \n (" ^ + stri_name ^ ".rec_tms, " ^ + stri_name ^ ".intr_tms)" + , + "structure " ^ big_rec_name ^ " =\n\ + \ struct\n\ + \ structure Result = " ^ co ^ "Ind_section_Fun\n\ + \ (open " ^ stri_name ^ "\n\ + \ val thy\t\t= thy\n\ + \ val monos\t\t= " ^ monos ^ "\n\ + \ val con_defs\t\t= " ^ con_defs ^ ");\n\n\ + \ val " ^ mk_list (map mk_intr_name ipairs) ^ " = Result.intrs;\n\ + \ open Result\n\ + \ end\n" + ) + end + val ipairs = "intrs" $$-- repeat1 (ident -- !! string) + fun optstring s = optional (s $$-- string) "\"[]\"" >> trim + in + repeat1 string -- ipairs -- optstring "monos" -- optstring "con_defs" + >> mk_params + end; + + + +(** datatype **) + +local + (* FIXME err -> add_datatype *) + fun mk_cons cs = + (case duplicates (map (fst o fst) cs) of + [] => map (fn ((s, ts), syn) => mk_triple (s, mk_list ts, syn)) cs + | dups => error ("Duplicate constructors: " ^ commas_quote dups)); + + (*generate names of distinctiveness axioms*) + fun mk_distinct_rules cs tname = + let + val uqcs = map (fn ((s, _), _) => strip_quotes s) cs; + (*combine all constructor names with all others w/o duplicates*) + fun neg_one c = map (fn c2 => quote (c ^ "_not_" ^ c2)); + fun neg1 [] = [] + | neg1 (c1 :: cs) = neg_one c1 cs @ neg1 cs; + in + if length uqcs < dtK then neg1 uqcs + else quote (tname ^ "_ord_distinct") :: + map (fn c => quote (tname ^ "_ord_" ^ c)) uqcs + end; + + fun mk_rules tname cons pre = " map (get_axiom thy) " ^ + mk_list (map (fn ((s, _), _) => quote (tname ^ pre ^ strip_quotes s)) cons); + + (*generate string for calling add_datatype*) + fun mk_params ((ts, tname), cons) = + ("val (thy, " ^ tname ^ "_add_primrec) = Datatype.add_datatype\n" + ^ mk_triple (mk_list ts, quote tname, mk_list (mk_cons cons)) ^ " thy\n\ + \val thy = thy", + "structure " ^ tname ^ " =\n\ + \struct\n\ + \ val inject = map (get_axiom thy) " ^ + mk_list (map (fn ((s, _), _) => quote ("inject_" ^ strip_quotes s)) + (filter_out (null o snd o fst) cons)) ^ ";\n\ + \ val distinct = " ^ + (if length cons < dtK then "let val distinct' = " else "") ^ + "map (get_axiom thy) " ^ mk_list (mk_distinct_rules cons tname) ^ + (if length cons < dtK then + " in distinct' @ (map (fn t => sym COMP (t RS contrapos))\ + \ distinct') end" + else "") ^ ";\n\ + \ val induct = get_axiom thy \"" ^ tname ^ "_induct\";\n\ + \ val cases =" ^ mk_rules tname cons "_case_" ^ ";\n\ + \ val recs =" ^ mk_rules tname cons "_rec_" ^ ";\n\ + \ val simps = inject @ distinct @ cases @ recs;\n\ + \ fun induct_tac a = res_inst_tac [(" ^ quote tname ^ ", a)] induct;\n\ + \end;\n"); + + (*parsers*) + val tvars = type_args >> map (cat "dtVar"); + val typ = + tvars -- (ident>>quote) >> (cat "dtTyp" o mk_pair o apfst mk_list) || + type_var >> cat "dtVar"; + val opt_typs = optional ("(" $$-- list1 typ --$$ ")") []; + val constructor = name -- opt_typs -- opt_mixfix; +in + val datatype_decl = + (* FIXME tvars -> type_args *) + tvars -- ident --$$ "=" -- enum1 "|" constructor >> mk_params; +end; + + + +(** primrec **) + +fun mk_primrec_decl ((fname, tname), axms) = + let + fun mk_prove (name, eqn) = + "val " ^ name ^ " = store_thm (" ^ quote name ^ ", prove_goalw thy [get_def thy " + ^ fname ^ "] " ^ eqn ^ "\n\ + \ (fn _ => [simp_tac (HOL_ss addsimps " ^ tname ^ ".recs) 1]));"; + val axs = mk_list (map (fn (n, a) => mk_pair (quote n, a)) axms); + in ("|> " ^ tname ^ "_add_primrec " ^ axs, cat_lines (map mk_prove axms)) end; + +val primrec_decl = + name -- long_id -- repeat1 (ident -- string) >> mk_primrec_decl; + + + +(** sections **) + +val user_keywords = ["intrs", "monos", "con_defs", "|"]; + +val user_sections = + [axm_section "subtype" "|> Subtype.add_subtype" subtype_decl, + ("inductive", inductive_decl ""), + ("coinductive", inductive_decl "Co"), + ("datatype", datatype_decl), + ("primrec", primrec_decl)]; + + +end; + + +structure ThySyn = ThySynFun(ThySynData); +init_thy_reader (); + diff -r 196ca0973a6d -r ff1574a81019 src/HOL/typedef.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/typedef.ML Fri Mar 03 12:02:25 1995 +0100 @@ -0,0 +1,141 @@ +(* Title: HOL/subtype.ML + ID: $Id$ + Author: Markus Wenzel, TU Muenchen + +Internal interface for subtype definitions. +*) + +signature SUBTYPE = +sig + val prove_nonempty: cterm -> thm list -> tactic option -> thm + val add_subtype: string -> string * string list * mixfix -> + string -> string list -> thm list -> tactic option -> theory -> theory + val add_subtype_i: string -> string * string list * mixfix -> + term -> string list -> thm list -> tactic option -> theory -> theory +end; + +structure Subtype: SUBTYPE = +struct + +open Syntax Logic HOLogic; + + +(* prove non-emptyness of a set *) (*exception ERROR*) + +val is_def = is_equals o #prop o rep_thm; + +fun prove_nonempty cset thms usr_tac = + let + val {T = setT, t = set, maxidx, sign} = rep_cterm cset; + val T = dest_setT setT; + val goal = + cterm_of sign (mk_Trueprop (mk_mem (Var (("x", maxidx + 1), T), set))); + val tac = + TRY (rewrite_goals_tac (filter is_def thms)) THEN + TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN + if_none usr_tac (TRY (ALLGOALS (fast_tac set_cs))); + in + prove_goalw_cterm [] goal (K [tac]) + end + handle ERROR => + error ("Failed to prove non-emptyness of " ^ quote (string_of_cterm cset)); + + +(* ext_subtype *) + +fun ext_subtype prep_term name (t, vs, mx) raw_set axms thms usr_tac thy = + let + val _ = require_thy thy "Set" "subtype definitions"; + val sign = sign_of thy; + + (*rhs*) + val cset = prep_term sign raw_set; + val {T = setT, t = set, ...} = rep_cterm cset; + val rhs_tfrees = term_tfrees set; + val oldT = dest_setT setT handle TYPE _ => + error ("Not a set type: " ^ quote (Sign.string_of_typ sign setT)); + + (*lhs*) + val lhs_tfrees = + map (fn v => (v, if_none (assoc (rhs_tfrees, v)) termS)) vs; + + val tname = type_name t mx; + val tlen = length vs; + val newT = Type (tname, map TFree lhs_tfrees); + + val Rep_name = "Rep_" ^ name; + val Abs_name = "Abs_" ^ name; + val setC = Const (name, setT); + val RepC = Const (Rep_name, newT --> oldT); + val AbsC = Const (Abs_name, oldT --> newT); + val x_new = Free ("x", newT); + val y_old = Free ("y", oldT); + + (*axioms*) + val rep_type = mk_Trueprop (mk_mem (RepC $ x_new, setC)); + val rep_type_inv = mk_Trueprop (mk_eq (AbsC $ (RepC $ x_new), x_new)); + val abs_type_inv = mk_implies (mk_Trueprop (mk_mem (y_old, setC)), + mk_Trueprop (mk_eq (RepC $ (AbsC $ y_old), y_old))); + + + (* errors *) + + val show_names = commas_quote o map fst; + + val illegal_vars = + if null (term_vars set) andalso null (term_tvars set) then [] + else ["Illegal schematic variable(s) on rhs"]; + + val dup_lhs_tfrees = + (case duplicates lhs_tfrees of [] => [] + | dups => ["Duplicate type variables on lhs: " ^ show_names dups]); + + val extra_rhs_tfrees = + (case gen_rems (op =) (rhs_tfrees, lhs_tfrees) of [] => [] + | extras => ["Extra type variables on rhs: " ^ show_names extras]); + + val illegal_frees = + (case term_frees set of [] => [] + | xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]); + + val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees; + in + if null errs then () + else error (cat_lines errs); + + prove_nonempty cset (map (get_axiom thy) axms @ thms) usr_tac; + + thy + |> add_types [(t, tlen, mx)] + |> add_arities + [(tname, replicate tlen logicS, logicS), + (tname, replicate tlen termS, termS)] + |> add_consts_i + [(name, setT, NoSyn), + (Rep_name, newT --> oldT, NoSyn), + (Abs_name, oldT --> newT, NoSyn)] + |> add_defs_i + [(name ^ "_def", mk_equals (setC, set))] + |> add_axioms_i + [(Rep_name, rep_type), + (Rep_name ^ "_inverse", rep_type_inv), + (Abs_name ^ "_inverse", abs_type_inv)] + end + handle ERROR => + error ("The error(s) above occurred in subtype definition " ^ quote name); + + +(* external interfaces *) + +fun cert_term sg tm = + cterm_of sg tm handle TERM (msg, _) => error msg; + +fun read_term sg str = + read_cterm sg (str, termTVar); + +val add_subtype = ext_subtype read_term; +val add_subtype_i = ext_subtype cert_term; + + +end; +