# HG changeset patch # User clasohm # Date 748174867 -7200 # Node ID 7949f97df77a3f85ef54318fccfff3c225a75a9b Initial revision diff -r 000000000000 -r 7949f97df77a Arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Arith.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,229 @@ +(* Title: HOL/arith.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For HOL/arith.thy. + +Proofs about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +open Arith; + +(*** Basic rewrite and congruence rules for the arithmetic operators ***) + +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; + +val diff_0_eq_0 = prove_goalw Arith.thy [diff_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) *) +val diff_Suc_Suc = prove_goalw Arith.thy [diff_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 = [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 ***) + +val add_0_right = prove_goal Arith.thy "m + 0 = m" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +val add_Suc_right = prove_goal 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*) +val add_assoc = prove_goal 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. Must simplify after first induction! + Orientation of rewrites is delicate*) +val add_commute = prove_goal Arith.thy "m + n = n + m::nat" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*** Multiplication ***) + +(*right annihilation in product*) +val mult_0_right = prove_goal Arith.thy "m * 0 = 0" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*right Sucessor law for multiplication*) +val mult_Suc_right = prove_goal Arith.thy + "m * Suc(n) = m + (m * n)" + (fn _ => + [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_assoc RS sym])), + (*The final goal requires the commutative law for addition*) + stac add_commute 1, rtac refl 1]); + +val arith_ss = arith_ss addsimps [mult_0_right,mult_Suc_right]; + +(*Commutative law for multiplication*) +val mult_commute = prove_goal Arith.thy "m * n = n * m::nat" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS (asm_simp_tac arith_ss)]); + +(*addition distributes over multiplication*) +val add_mult_distrib = prove_goal Arith.thy "(m + n)*k = (m*k) + (n*k)::nat" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_assoc]))]); + +(*Associative law for multiplication*) +val mult_assoc = prove_goal Arith.thy "(m * n) * k = m * (n * k)::nat" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_mult_distrib]))]); + + +(*** Difference ***) + +val diff_self_eq_0 = prove_goal 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)); +val add_diff_inverse = result(); + + +(*** 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)); +val diff_less_Suc = result(); + +(*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]))); +val div_termination = result(); + +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); +val mod_less = result(); + +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); +val mod_geq = result(); + + +(*** 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); +val div_less = result(); + +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); +val div_geq = result(); + +(*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 (res_inst_tac [("Q","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)); +val less_imp_diff_is_0 = result(); + +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))); +val diffs0_imp_equal_lemma = result(); + +(* [| m-n = 0; n-m = 0 |] ==> m=n *) +val diffs0_imp_equal = standard (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)); +val Suc_diff_n = result(); + +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' fast_tac HOL_cs)); +val zero_induct_lemma = result(); + +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)); +val zero_induct = result(); + +(*13 July 1992: loaded in 105.7s*) + +(**** Additional theorems about "less than" ****) + +goal Arith.thy "n <= (m + n)::nat"; +by (nat_ind_tac "m" 1); +by (ALLGOALS(simp_tac (arith_ss addsimps [le_refl]))); +by (etac le_trans 1); +by (rtac (lessI RS less_imp_le) 1); +val le_add2 = result(); + +goal Arith.thy "m <= (m + n)::nat"; +by (stac add_commute 1); +by (rtac le_add2 1); +val le_add1 = result(); + +val less_add_Suc1 = standard (lessI RS (le_add1 RS le_less_trans)); +val less_add_Suc2 = standard (lessI RS (le_add2 RS le_less_trans)); diff -r 000000000000 -r 7949f97df77a Arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Arith.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,26 @@ +(* 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 + +arities nat::plus + nat::minus + nat::times +consts + div,mod :: "[nat,nat]=>nat" (infixl 70) +rules + add_def "m+n == nat_rec(m, n, %u v.Suc(v))" + diff_def "m-n == nat_rec(n, m, %u v. nat_rec(v, 0, %x y.x))" + 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 000000000000 -r 7949f97df77a Gfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Gfp.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,136 @@ +(* 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] "[| A <= f(A) |] ==> A <= gfp(f)"; +by (rtac (CollectI RS Union_upper) 1); +by (resolve_tac prems 1); +val gfp_upperbound = result(); + +val prems = goalw Gfp.thy [gfp_def] + "[| !!u. u <= f(u) ==> u<=A |] ==> gfp(f) <= A"; +by (REPEAT (ares_tac ([Union_least]@prems) 1)); +by (etac CollectD 1); +val gfp_least = result(); + +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]); +val gfp_lemma2 = result(); + +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]); +val gfp_lemma3 = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))"; +by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1)); +val gfp_Tarski = result(); + +(*** Coinduction rules for greatest fixed points ***) + +(*weak version*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) |] ==> a : gfp(f)"; +by (rtac (gfp_upperbound RS subsetD) 1); +by (REPEAT (ares_tac prems 1)); +val coinduct = result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(A) Un gfp(f); mono(f) |] ==> \ +\ A Un gfp(f) <= f(A Un gfp(f))"; +by (rtac subset_trans 1); +by (rtac (mono RS mono_Un) 2); +by (rtac (mono RS gfp_Tarski RS subst) 1); +by (rtac (prem RS Un_least) 1); +by (rtac Un_upper2 1); +val coinduct2_lemma = result(); + +(*strong version, thanks to Martin Coen*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) Un gfp(f); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct2_lemma RSN (2,coinduct)) 1); +by (REPEAT (resolve_tac (prems@[UnI1]) 1)); +val coinduct2 = result(); + +(*** Even Stronger version of coinduct [by Martin Coen] + - instead of the condition A <= f(A) + consider A <= (f(A) Un f(f(A)) ...) Un gfp(A) ***) + +val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un A Un B)"; +by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1)); +val coinduct3_mono_lemma= result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> \ +\ lfp(%x.f(x) Un A Un gfp(f)) <= f(lfp(%x.f(x) Un A 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); +val coinduct3_lemma = result(); + +val prems = goal Gfp.thy + "[| a:A; A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct3_lemma RSN (2,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)); +val coinduct3 = result(); + + +(** Definition forms of gfp_Tarski and coinduct, to control unfolding **) + +val [rew,mono] = goal Gfp.thy "[| h==gfp(f); mono(f) |] ==> h = f(h)"; +by (rewtac rew); +by (rtac (mono RS gfp_Tarski) 1); +val def_gfp_Tarski = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (prems @ [coinduct]) 1)); +val def_coinduct = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) Un h; mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct2]) 1)); +val def_coinduct2 = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(lfp(%x.f(x) Un A Un h)); mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1)); +val def_coinduct3 = result(); + +(*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; +val gfp_mono = result(); diff -r 000000000000 -r 7949f97df77a Gfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Gfp.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: HOL/gfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Greatest fixed points +*) + +Gfp = Lfp + +consts gfp :: "['a set=>'a set] => 'a set" +rules + (*greatest fixed point*) + gfp_def "gfp(f) == Union({u. u <= f(u)})" +end diff -r 000000000000 -r 7949f97df77a HOL.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/HOL.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,334 @@ +(* 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; + +signature HOL_LEMMAS = + sig + val allE: thm + val all_dupE: thm + val allI: thm + val arg_cong: thm + val fun_cong: thm + val box_equals: thm + val ccontr: thm + val classical: thm + val cong: thm + val conjunct1: thm + val conjunct2: thm + val conjE: thm + val conjI: thm + val contrapos: thm + val disjCI: thm + val disjE: thm + val disjI1: thm + val disjI2: thm + val eqTrueI: thm + val eqTrueE: thm + val ex1E: thm + val ex1I: thm + val exCI: thm + val exI: thm + val exE: thm + val excluded_middle: thm + val FalseE: thm + val False_neq_True: thm + val iffCE : thm + val iffD1: thm + val iffD2: thm + val iffE: thm + val iffI: thm + val impCE: thm + val impE: thm + val not_sym: thm + val notE: thm + val notI: thm + val notnotD : thm + val rev_mp: thm + val select_equality: thm + val spec: thm + val sstac: thm list -> int -> tactic + val ssubst: thm + val stac: thm -> int -> tactic + val strip_tac: int -> tactic + val swap: thm + val sym: thm + val trans: thm + val TrueI: thm + end; + +structure HOL_Lemmas : HOL_LEMMAS = + +struct + +(** Equality **) + +val sym = prove_goal 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*) +val ssubst = standard (sym RS subst); + +val trans = prove_goal 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 *) +val box_equals = prove_goal 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*) +val fun_cong = prove_goal 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*) +val arg_cong = prove_goal HOL.thy "x=y ==> f(x)=f(y)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +val cong = prove_goal 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 **) + +val iffI = prove_goal HOL.thy + "[| P ==> Q; Q ==> P |] ==> P=Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI, iff RS mp RS mp]) 1)) ]); + +val iffD2 = prove_goal HOL.thy "[| P=Q; Q |] ==> P" + (fn prems => + [rtac ssubst 1, resolve_tac prems 1, resolve_tac prems 1]); + +val iffD1 = sym RS iffD2; + +val iffE = prove_goal 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 **) + +val TrueI = refl RS (True_def RS iffD2); + +val eqTrueI = prove_goal HOL.thy "P ==> P=True" + (fn prems => [REPEAT(resolve_tac ([iffI,TrueI]@prems) 1)]); + +val eqTrueE = prove_goal HOL.thy "P=True ==> P" + (fn prems => [REPEAT(resolve_tac (prems@[TrueI,iffD2]) 1)]); + +(** Universal quantifier **) + +val allI = prove_goal HOL.thy "(!!x::'a. P(x)) ==> !x. P(x)" + (fn [asm] => [rtac (All_def RS ssubst) 1, rtac (asm RS (eqTrueI RS ext)) 1]); + +val spec = prove_goal HOL.thy "! x::'a.P(x) ==> P(x)" + (fn prems => + [ rtac eqTrueE 1, + resolve_tac (prems RL [All_def RS subst] RL [fun_cong]) 1 ]); + +val allE = prove_goal HOL.thy "[| !x.P(x); P(x) ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (resolve_tac (prems @ [major RS spec]) 1)) ]); + +val all_dupE = prove_goal 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! **) + +val FalseE = prove_goal HOL.thy "False ==> P" + (fn prems => [rtac spec 1, rtac (False_def RS subst) 1, resolve_tac prems 1]); + +val False_neq_True = prove_goal HOL.thy "False=True ==> P" + (fn [prem] => [rtac (prem RS eqTrueE RS FalseE) 1]); + + +(** Negation **) + +val notI = prove_goal HOL.thy "(P ==> False) ==> ~P" + (fn prems=> [rtac (not_def RS ssubst) 1, rtac impI 1, eresolve_tac prems 1]); + +val notE = prove_goal HOL.thy "[| ~P; P |] ==> R" + (fn prems => + [rtac (mp RS FalseE) 1, + resolve_tac prems 2, rtac (not_def RS subst) 1, + resolve_tac prems 1]); + +(** Implication **) + +val impE = prove_goal 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. *) +val rev_mp = prove_goal HOL.thy "[| P; P --> Q |] ==> Q" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +val contrapos = prove_goal 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 **) + +val exI = prove_goal HOL.thy "P(x) ==> ? x::'a.P(x)" + (fn prems => + [rtac (selectI RS (Ex_def RS ssubst)) 1, + resolve_tac prems 1]); + +val exE = prove_goal HOL.thy "[| ? x::'a.P(x); !!x. P(x) ==> Q |] ==> Q" + (fn prems => + [resolve_tac prems 1, res_inst_tac [("P","%C.C(P)")] subst 1, + rtac Ex_def 1, resolve_tac prems 1]); + + +(** Conjunction **) + +val conjI = prove_goal HOL.thy "[| P; Q |] ==> P&Q" + (fn prems => + [ (rtac (and_def RS ssubst) 1), + (REPEAT (resolve_tac (prems@[allI,impI]) 1 ORELSE etac (mp RS mp) 1)) ]); + +val conjunct1 = prove_goal HOL.thy "[| P & Q |] ==> P" + (fn prems => + [ (resolve_tac (prems RL [and_def RS subst] RL [spec] RL [mp]) 1), + (REPEAT(ares_tac [impI] 1)) ]); + +val conjunct2 = prove_goal HOL.thy "[| P & Q |] ==> Q" + (fn prems => + [ (resolve_tac (prems RL [and_def RS subst] RL [spec] RL [mp]) 1), + (REPEAT(ares_tac [impI] 1)) ]); + +val conjE = prove_goal 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 *) + +val disjI1 = prove_goal HOL.thy "P ==> P|Q" + (fn [prem] => + [rtac (or_def RS ssubst) 1, + REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]); + +val disjI2 = prove_goal HOL.thy "Q ==> P|Q" + (fn [prem] => + [rtac (or_def RS ssubst) 1, + REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]); + +val disjE = prove_goal HOL.thy "[| P | Q; P ==> R; Q ==> R |] ==> R" + (fn [a1,a2,a3] => + [rtac (mp RS mp) 1, rtac spec 1, rtac (or_def RS subst) 1, rtac a1 1, + rtac (a2 RS impI) 1, atac 1, rtac (a3 RS impI) 1, atac 1]); + +(** CCONTR -- classical logic **) + +val ccontr = prove_goal HOL.thy "(~P ==> False) ==> P" + (fn prems => + [rtac (True_or_False RS (disjE RS eqTrueE)) 1, atac 1, + rtac spec 1, rtac (False_def RS subst) 1, resolve_tac prems 1, + rtac ssubst 1, atac 1, rtac (not_def RS ssubst) 1, + REPEAT (ares_tac [impI] 1) ]); + +val classical = prove_goal HOL.thy "(~P ==> P) ==> P" + (fn prems => + [rtac ccontr 1, + REPEAT (ares_tac (prems@[notE]) 1)]); + + +(*Double negation law*) +val notnotD = prove_goal HOL.thy "~~P ==> P" + (fn [major]=> + [ (rtac classical 1), (eresolve_tac [major RS notE] 1) ]); + + +(** Unique existence **) + +val ex1I = prove_goal HOL.thy + "[| P(a); !!x. P(x) ==> x=a |] ==> ?! x. P(x)" + (fn prems => + [ (rtac (Ex1_def RS ssubst) 1), + (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); + +val ex1E = prove_goal HOL.thy + "[| ?! x.P(x); !!x. [| P(x); ! y. P(y) --> y=x |] ==> R |] ==> R" + (fn major::prems => + [ (resolve_tac ([major] RL [Ex1_def RS subst] RL [exE]) 1), + (REPEAT (etac conjE 1 ORELSE ares_tac prems 1)) ]); + + +(** Select: Hilbert's Epsilon-operator **) + +val select_equality = prove_goal HOL.thy + "[| P(a); !!x. P(x) ==> x=a |] ==> (@x.P(x)) = a" + (fn prems => [ resolve_tac prems 1, + rtac selectI 1, + resolve_tac prems 1 ]); + +(** Classical intro rules for disjunction and existential quantifiers *) + +val disjCI = prove_goal 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)) ]); + +val excluded_middle = prove_goal HOL.thy "~P | P" + (fn _ => [ (REPEAT (ares_tac [disjCI] 1)) ]); + +(*Classical implies (-->) elimination. *) +val impCE = prove_goal 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. *) +val iffCE = prove_goal 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))) ]); + +val exCI = prove_goal HOL.thy "(! x. ~P(x) ==> P(a)) ==> ? x.P(x)" + (fn prems=> + [ (rtac ccontr 1), + (REPEAT (ares_tac (prems@[exI,allI,notI,notE]) 1)) ]); + +(*Required by the "classical" module*) +val swap = prove_goal HOL.thy "~P ==> (~Q ==> P) ==> Q" + (fn major::prems=> + [ rtac ccontr 1, rtac (major RS notE) 1, REPEAT (ares_tac prems 1)]); + +(** 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); + +end; + +open HOL_Lemmas; + diff -r 000000000000 -r 7949f97df77a HOL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/HOL.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,142 @@ +(* Title: HOL/hol.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +Higher-Order Logic +*) + +HOL = Pure + + +classes + term < logic + plus < term + minus < term + times < term + +default term + +types + bool 0 + +arities + fun :: (term, term) term + bool :: term + + +consts + + (* Constants *) + + Trueprop :: "bool => prop" ("(_)" [0] 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" + "@let" :: "[idt,'a,'b] => 'b" ("(let _ = (2_)/ in (2_))" 10) + + (* Alternative Quantifiers *) + + "*All" :: "[idts, bool] => bool" ("(3ALL _./ _)" [0,0] 10) + "*Ex" :: "[idts, bool] => bool" ("(3EX _./ _)" [0,0] 10) + "*Ex1" :: "[idts, bool] => bool" ("(3EX! _./ _)" [0,0] 10) + + (* Infixes *) + + o :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixr 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) + + (* Rewriting Gadget *) + + NORM :: "'a => 'a" + + +translations + "ALL xs. P" => "! xs. P" + "EX xs. P" => "? xs. P" + "EX! xs. P" => "?! xs. P" + + "let x = s in t" == "Let(s,%x.t)" + +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" + + (* Definitions *) + + True_def "True = ((%x.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 "not = (%P. P-->False)" + and_def "op & = (%P Q. !R. (P-->Q-->R) --> R)" + or_def "op | = (%P Q. !R. (P-->R) --> (Q-->R) --> R)" + Ex1_def "Ex1 = (%P. ? x. P(x) & (! y. P(y) --> y=x))" + Let_def "Let(s,f) = f(s)" + + (* Axioms *) + + iff "(P-->Q) --> (Q-->P) --> (P=Q)" + True_or_False "(P=True) | (P=False)" + + (* Misc Definitions *) + + Inv_def "Inv = (%(f::'a=>'b) y. @x. f(x)=y)" + o_def "op o = (%(f::'b=>'c) g (x::'a). f(g(x)))" + + if_def "if = (%P x y.@z::'a. (P=True --> z=x) & (P=False --> z=y))" + + (* Rewriting -- special constant to flag normalized terms *) + + NORM_def "NORM(x) = x" + +end + + +ML + +(** Choice between the HOL and Isabelle style of quantifiers **) + +val HOL_quantifiers = ref true; + +fun mk_alt_ast_tr' (name, alt_name) = + let + fun ast_tr' (*name*) args = + if ! HOL_quantifiers then raise Match + else Ast.mk_appl (Ast.Constant alt_name) args; + in + (name, ast_tr') + end; + + +val print_ast_translation = + map mk_alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")]; + diff -r 000000000000 -r 7949f97df77a LList.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/LList.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,899 @@ +(* Title: HOL/llist + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For llist.thy. + +SHOULD LListD_Fun_CONS_I, etc., be equations (for rewriting) +*) + +open LList; + +(** Simplification **) + +val llist_simps = [case_Inl, case_Inr]; +val llist_ss = univ_ss addsimps llist_simps + setloop (split_tac [expand_split,expand_case]); + +(** the llist functional **) + +val LList_unfold = rewrite_rule [List_Fun_def] + (List_Fun_mono RS (LList_def RS def_gfp_Tarski)); + +(*This justifies using LList in other recursive type definitions*) +goalw LList.thy [LList_def] "!!A B. A<=B ==> LList(A) <= LList(B)"; +by (rtac gfp_mono 1); +by (etac List_Fun_mono2 1); +val LList_mono = result(); + +(*Elimination is case analysis, not induction.*) +val [major,prem1,prem2] = goalw LList.thy [NIL_def,CONS_def] + "[| L : LList(A); \ +\ L=NIL ==> P; \ +\ !!M N. [| M:A; N: LList(A); L=CONS(M,N) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS (LList_unfold RS equalityD1 RS subsetD RS usumE)) 1); +by (etac uprodE 2); +by (rtac prem2 2); +by (rtac prem1 1); +by (REPEAT (ares_tac [refl] 1 + ORELSE eresolve_tac [singletonE,ssubst] 1)); +val LListE = result(); + + +(*** Type checking by co-induction, using List_Fun ***) + +val prems = goalw LList.thy [LList_def] + "[| M: X; X <= List_Fun(A,X) |] ==> M: LList(A)"; +by (REPEAT (resolve_tac (prems@[coinduct]) 1)); +val LList_coinduct = result(); + +(*stronger version*) +val prems = goalw LList.thy [LList_def] + "[| M : X; X <= List_Fun(A, X) Un LList(A) |] ==> M : LList(A)"; +by (REPEAT (resolve_tac (prems@[coinduct2,List_Fun_mono]) 1)); +val LList_coinduct2 = result(); + +(** Rules to prove the 2nd premise of LList_coinduct **) + +goalw LList.thy [List_Fun_def,NIL_def] "NIL: List_Fun(A,X)"; +by (resolve_tac [singletonI RS usum_In0I] 1); +val List_Fun_NIL_I = result(); + +goalw LList.thy [List_Fun_def,CONS_def] + "!!M N. [| M: A; N: X |] ==> CONS(M,N) : List_Fun(A,X)"; +by (REPEAT (ares_tac [uprodI RS usum_In1I] 1)); +val List_Fun_CONS_I = result(); + +(*** LList_corec satisfies the desired recurion equation ***) + +(*A continuity result?*) +goalw LList.thy [CONS_def] "CONS(M, UN x.f(x)) = (UN x. CONS(M, f(x)))"; +by(simp_tac (univ_ss addsimps [In1_UN1, Scons_UN1_y]) 1); +val CONS_UN1 = result(); + +goal Prod.thy "split(p, %x y.UN z.f(x,y,z)) = (UN z. split(p, %x y.f(x,y,z)))"; +by(simp_tac (pair_ss setloop (split_tac [expand_split])) 1); +val split_UN1 = result(); + +goal Sum.thy "case(s, f, %y. UN z.g(y,z)) = (UN z. case(s, f, %y. g(y,z)))"; +by(simp_tac (sum_ss setloop (split_tac [expand_case])) 1); +val case2_UN1 = result(); + +val prems = goalw LList.thy [CONS_def] + "[| M<=M'; N<=N' |] ==> CONS(M,N) <= CONS(M',N')"; +by (REPEAT (resolve_tac ([In1_mono,Scons_mono]@prems) 1)); +val CONS_mono = result(); + +val corec_fun_simps = [LList_corec_fun_def RS def_nat_rec_0, + LList_corec_fun_def RS def_nat_rec_Suc]; +val corec_fun_ss = llist_ss addsimps corec_fun_simps; + +(** The directions of the equality are proved separately **) + +goalw LList.thy [LList_corec_def] + "LList_corec(a,f) <= case(f(a), %u.NIL, \ +\ %v. split(v, %z w. CONS(z, LList_corec(w,f))))"; +by (rtac UN1_least 1); +by (nat_ind_tac "k" 1); +by(ALLGOALS(simp_tac corec_fun_ss)); +by (REPEAT (resolve_tac [allI, impI, subset_refl RS CONS_mono, UN1_upper] 1)); +val LList_corec_subset1 = result(); + +goalw LList.thy [LList_corec_def] + "case(f(a), %u.NIL, %v. split(v, %z w. CONS(z, LList_corec(w,f)))) <= \ +\ LList_corec(a,f)"; +by (simp_tac (corec_fun_ss addsimps [CONS_UN1]) 1); +by (safe_tac set_cs); +by (ALLGOALS (res_inst_tac [("x","Suc(?k)")] UN1_I THEN' + asm_simp_tac corec_fun_ss)); +val LList_corec_subset2 = result(); + +(*the recursion equation for LList_corec -- NOT SUITABLE FOR REWRITING!*) +goal LList.thy + "LList_corec(a,f) = case(f(a), %u. NIL, \ +\ %v. split(v, %z w. CONS(z, LList_corec(w,f))))"; +by (REPEAT (resolve_tac [equalityI, LList_corec_subset1, + LList_corec_subset2] 1)); +val LList_corec = result(); + +(*definitional version of same*) +val [rew] = goal LList.thy + "[| !!x. h(x) == LList_corec(x,f) |] ==> \ +\ h(a) = case(f(a), %u.NIL, %v. split(v, %z w. CONS(z, h(w))))"; +by (rewtac rew); +by (rtac LList_corec 1); +val def_LList_corec = result(); + +(*A typical use of co-induction to show membership in the gfp. + Bisimulation is range(%x. LList_corec(x,f)) *) +goal LList.thy "LList_corec(a,f) : LList({u.True})"; +by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] LList_coinduct 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac LList_corec 1); +by(simp_tac (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I, + CollectI, range_eqI]) 1); +(* 6.7 vs 3.4 !!! +by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I, + CollectI, rangeI]) 1); +*) +val LList_corec_type = result(); + +(*Lemma for the proof of llist_corec*) +goal LList.thy + "LList_corec(a, %z. case(f(z),Inl,%x. split(x,%v w. Inr()))) : \ +\ LList(range(Leaf))"; +by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] LList_coinduct 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac LList_corec 1); +(*nested "case"; requires an explicit split*) +by (res_inst_tac [("s", "f(xa)")] sumE 1); +by(asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_NIL_I])) 1); +by(asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_CONS_I, range_eqI]) + setloop (split_tac [expand_split])) 1); +(* FIXME: can the selection of the case split be automated? +by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_CONS_I, rangeI]) 1);*) +val LList_corec_type2 = result(); + +(**** LList equality as a gfp; the bisimulation principle ****) + +goalw LList.thy [LListD_Fun_def] "mono(LListD_Fun(r))"; +by (REPEAT (ares_tac [monoI, subset_refl, dsum_mono, dprod_mono] 1)); +val LListD_fun_mono = result(); + +val LListD_unfold = rewrite_rule [LListD_Fun_def] + (LListD_fun_mono RS (LListD_def RS def_gfp_Tarski)); + +goal LList.thy "!M N. : LListD(diag(A)) --> ntrunc(k,M) = ntrunc(k,N)"; +by (res_inst_tac [("n", "k")] less_induct 1); +by (safe_tac set_cs); +by (etac (LListD_unfold RS equalityD1 RS subsetD RS dsumE) 1); +by (safe_tac (set_cs addSEs [Pair_inject, dprodE, diagE])); +by (res_inst_tac [("n", "n")] natE 1); +by(asm_simp_tac (univ_ss addsimps [ntrunc_0]) 1); +by (res_inst_tac [("n", "xb")] natE 1); +by(asm_simp_tac (univ_ss addsimps [ntrunc_one_In1]) 1); +by(asm_simp_tac (univ_ss addsimps [ntrunc_In1, ntrunc_Scons]) 1); +val LListD_implies_ntrunc_equality = result(); + +goalw LList.thy [LList_def,List_Fun_def] "fst``LListD(diag(A)) <= LList(A)"; +by (rtac gfp_upperbound 1); +by (res_inst_tac [("P", "%x. fst``x <= ?B")] (LListD_unfold RS ssubst) 1); +by(simp_tac fst_image_ss 1); +val fst_image_LListD = result(); + +(*This inclusion justifies the use of coinduction to show M=N*) +goal LList.thy "LListD(diag(A)) <= diag(LList(A))"; +by (rtac subsetI 1); +by (res_inst_tac [("p","x")] PairE 1); +by (safe_tac HOL_cs); +by (res_inst_tac [("s","xa")] subst 1); +by (rtac (LListD_implies_ntrunc_equality RS spec RS spec RS mp RS + ntrunc_equality) 1); +by (assume_tac 1); +by (rtac diagI 1); +by (etac (fst_imageI RS (fst_image_LListD RS subsetD)) 1); +val LListD_subset_diag = result(); + +(*This converse inclusion helps to strengthen LList_equalityI*) +goalw LList.thy [LListD_def] "diag(LList(A)) <= LListD(diag(A))"; +by (rtac gfp_upperbound 1); +by (rtac subsetI 1); +by (etac diagE 1); +by (etac ssubst 1); +by (etac (LList_unfold RS equalityD1 RS subsetD RS usumE) 1); +by (rewtac LListD_Fun_def); +by (ALLGOALS (fast_tac (set_cs addIs [diagI,dsum_In0I,dsum_In1I,dprodI] + addSEs [uprodE]))); +val diag_subset_LListD = result(); + +goal LList.thy "LListD(diag(A)) = diag(LList(A))"; +by (REPEAT (resolve_tac [equalityI, LListD_subset_diag, + diag_subset_LListD] 1)); +val LListD_eq_diag = result(); + +(** To show two LLists are equal, exhibit a bisimulation! **) +(* Replace "A" by some particular set, like {x.True}??? *) +val prems = goal LList.thy + "[| : r; r <= LListD_Fun(diag(A), r) |] ==> M=N"; +by (rtac (rewrite_rule [LListD_def] + (LListD_subset_diag RS subsetD RS diagE)) 1); +by (REPEAT (resolve_tac (prems@[coinduct]) 1)); +by (safe_tac (set_cs addSEs [Pair_inject])); +val LList_equalityI = result(); + +(*Stronger notion of bisimulation -- also admits true equality*) +val prems = goal LList.thy + "[| : r; r <= LListD_Fun(diag(A), r) Un diag(LList(A)) |] ==> M=N"; +by (rtac (rewrite_rule [LListD_def] + (LListD_subset_diag RS subsetD RS diagE)) 1); +by (rtac coinduct2 1); +by (stac (rewrite_rule [LListD_def] LListD_eq_diag) 2); +by (REPEAT (resolve_tac (prems@[LListD_fun_mono]) 1)); +by (safe_tac (set_cs addSEs [Pair_inject])); +val LList_equalityI2 = result(); + +(** Rules to prove the 2nd premise of LList_equalityI **) + +goalw LList.thy [LListD_Fun_def,NIL_def] " : LListD_Fun(r,s)"; +by (rtac (singletonI RS diagI RS dsum_In0I) 1); +val LListD_Fun_NIL_I = result(); + +val prems = goalw LList.thy [LListD_Fun_def,CONS_def] + "[| x:A; :s |] ==> : LListD_Fun(diag(A),s)"; +by (rtac (dprodI RS dsum_In1I) 1); +by (REPEAT (resolve_tac (diagI::prems) 1)); +val LListD_Fun_CONS_I = result(); + + +(*** Finality of LList(A): Uniqueness of functions defined by corecursion ***) + +(*abstract proof using a bisimulation*) +val [prem1,prem2] = goal LList.thy + "[| !!x. h1(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h1(w)))); \ +\ !!x. h2(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h2(w)))) |] \ +\ ==> h1=h2"; +by (rtac ext 1); +(*next step avoids an unknown (and flexflex pair) in simplification*) +by (res_inst_tac [("A", "{u.True}"), + ("r", "range(%u. )")] LList_equalityI 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac prem1 1); +by (stac prem2 1); +by(simp_tac (llist_ss addsimps [LListD_Fun_NIL_I, range_eqI, + CollectI RS LListD_Fun_CONS_I]) 1); +(* 9.5 vs 9.2/4.1/4.3 +by (ASM_SIMP_TAC (llist_ss addsimps [LListD_Fun_NIL_I, rangeI, + CollectI RS LListD_Fun_CONS_I]) 1);*) +val LList_corec_unique = result(); + +val [prem] = goal LList.thy + "[| !!x. h(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h(w)))) |] \ +\ ==> h = (%x.LList_corec(x,f))"; +by (rtac (LList_corec RS (prem RS LList_corec_unique)) 1); +val equals_LList_corec = result(); + + +(** Obsolete LList_corec_unique proof: complete induction, not coinduction **) + +goalw LList.thy [CONS_def] "ntrunc(Suc(0), CONS(M,N)) = {}"; +by (rtac ntrunc_one_In1 1); +val ntrunc_one_CONS = result(); + +goalw LList.thy [CONS_def] + "ntrunc(Suc(Suc(k)), CONS(M,N)) = CONS (ntrunc(k,M), ntrunc(k,N))"; +by(simp_tac (HOL_ss addsimps [ntrunc_Scons,ntrunc_In1]) 1); +val ntrunc_CONS = result(); + +val [prem1,prem2] = goal LList.thy + "[| !!x. h1(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h1(w)))); \ +\ !!x. h2(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h2(w)))) |] \ +\ ==> h1=h2"; +by (rtac (ntrunc_equality RS ext) 1); +by (res_inst_tac [("x", "x")] spec 1); +by (res_inst_tac [("n", "k")] less_induct 1); +by (rtac allI 1); +by (stac prem1 1); +by (stac prem2 1); +by(simp_tac (sum_ss setloop (split_tac [expand_split,expand_case])) 1); +by (strip_tac 1); +by (res_inst_tac [("n", "n")] natE 1); +by (res_inst_tac [("n", "xc")] natE 2); +by(ALLGOALS(asm_simp_tac(nat_ss addsimps + [ntrunc_0,ntrunc_one_CONS,ntrunc_CONS]))); +val LList_corec_unique = result(); + + +(*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***) + +goal LList.thy "mono(CONS(M))"; +by (REPEAT (ares_tac [monoI, subset_refl, CONS_mono] 1)); +val Lconst_fun_mono = result(); + +(* Lconst(M) = CONS(M,Lconst(M)) *) +val Lconst = standard (Lconst_fun_mono RS (Lconst_def RS def_lfp_Tarski)); + +(*A typical use of co-induction to show membership in the gfp. + The containing set is simply the singleton {Lconst(M)}. *) +goal LList.thy "!!M A. M:A ==> Lconst(M): LList(A)"; +by (rtac (singletonI RS LList_coinduct) 1); +by (safe_tac set_cs); +by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1); +by (REPEAT (ares_tac [List_Fun_CONS_I, singletonI] 1)); +val Lconst_type = result(); + +goal LList.thy "Lconst(M) = LList_corec(M, %x.Inr())"; +by (rtac (equals_LList_corec RS fun_cong) 1); +by(simp_tac sum_ss 1); +by (rtac Lconst 1); +val Lconst_eq_LList_corec = result(); + +(*Thus we could have used gfp in the definition of Lconst*) +goal LList.thy "gfp(%N. CONS(M, N)) = LList_corec(M, %x.Inr())"; +by (rtac (equals_LList_corec RS fun_cong) 1); +by(simp_tac sum_ss 1); +by (rtac (Lconst_fun_mono RS gfp_Tarski) 1); +val gfp_Lconst_eq_LList_corec = result(); + + +(** Introduction rules for LList constructors **) + +(* c : {Numb(0)} <+> A <*> LList(A) ==> c : LList(A) *) +val LListI = LList_unfold RS equalityD2 RS subsetD; + +(*This justifies the type definition: LList(A) is nonempty.*) +goalw LList.thy [NIL_def] "NIL: LList(A)"; +by (rtac (singletonI RS usum_In0I RS LListI) 1); +val NIL_LListI = result(); + +val prems = goalw LList.thy [CONS_def] + "[| M: A; N: LList(A) |] ==> CONS(M,N) : LList(A)"; +by (rtac (uprodI RS usum_In1I RS LListI) 1); +by (REPEAT (resolve_tac prems 1)); +val CONS_LListI = result(); + +(*** Isomorphisms ***) + +goal LList.thy "inj(Rep_LList)"; +by (rtac inj_inverseI 1); +by (rtac Rep_LList_inverse 1); +val inj_Rep_LList = result(); + +goal LList.thy "inj_onto(Abs_LList,LList(range(Leaf)))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_LList_inverse 1); +val inj_onto_Abs_LList = result(); + +(** Distinctness of constructors **) + +goalw LList.thy [LNil_def,LCons_def] "~ LCons(x,xs) = LNil"; +by (rtac (CONS_not_NIL RS (inj_onto_Abs_LList RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, NIL_LListI, CONS_LListI, Rep_LList] 1)); +val LCons_not_LNil = result(); + +val LNil_not_LCons = standard (LCons_not_LNil RS not_sym); + +val LCons_neq_LNil = standard (LCons_not_LNil RS notE); +val LNil_neq_LCons = sym RS LCons_neq_LNil; + +(** llist constructors **) + +goalw LList.thy [LNil_def] + "Rep_LList(LNil) = NIL"; +by (rtac (NIL_LListI RS Abs_LList_inverse) 1); +val Rep_LList_LNil = result(); + +goalw LList.thy [LCons_def] + "Rep_LList(LCons(x,l)) = CONS(Leaf(x),Rep_LList(l))"; +by (REPEAT (resolve_tac [CONS_LListI RS Abs_LList_inverse, + rangeI, Rep_LList] 1)); +val Rep_LList_LCons = result(); + +(** Injectiveness of CONS and LCons **) + +goalw LList.thy [CONS_def] "(CONS(M,N)=CONS(M',N')) = (M=M' & N=N')"; +by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1); +val CONS_CONS_eq = result(); + +val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE); + + +(*For reasoning about abstract llist constructors*) +val LList_cs = set_cs addIs [Rep_LList, NIL_LListI, CONS_LListI] + addSEs [CONS_neq_NIL,NIL_neq_CONS,CONS_inject] + addSDs [inj_onto_Abs_LList RS inj_ontoD, + inj_Rep_LList RS injD, Leaf_inject]; + +goalw LList.thy [LCons_def] "(LCons(x,xs)=LCons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac LList_cs 1); +val LCons_LCons_eq = result(); +val LCons_inject = standard (LCons_LCons_eq RS iffD1 RS conjE); + +val [major] = goal LList.thy "CONS(M,N): LList(A) ==> M: A & N: LList(A)"; +by (rtac (major RS LListE) 1); +by (etac CONS_neq_NIL 1); +by (fast_tac LList_cs 1); +val CONS_D = result(); + + +(****** Reasoning about LList(A) ******) + +val List_case_simps = [List_case_NIL, List_case_CONS]; +val List_case_ss = llist_ss addsimps List_case_simps; + +(*A special case of list_equality for functions over lazy lists*) +val [MList,gMList,NILcase,CONScase] = goal LList.thy + "[| M: LList(A); g(NIL): LList(A); \ +\ f(NIL)=g(NIL); \ +\ !!x l. [| x:A; l: LList(A) |] ==> \ +\ : \ +\ LListD_Fun(diag(A), (%u.)``LList(A)) Un \ +\ diag(LList(A)) \ +\ |] ==> f(M) = g(M)"; +by (rtac LList_equalityI2 1); +br (MList RS imageI) 1; +by (rtac subsetI 1); +by (etac imageE 1); +by (etac ssubst 1); +by (etac LListE 1); +by (etac ssubst 1); +by (stac NILcase 1); +br (gMList RS diagI RS UnI2) 1; +by (etac ssubst 1); +by (REPEAT (ares_tac [CONScase] 1)); +val LList_fun_equalityI = result(); + + +(*** The functional "Lmap" ***) + +goal LList.thy "Lmap(f,NIL) = NIL"; +by (rtac (Lmap_def RS def_LList_corec RS trans) 1); +by(simp_tac List_case_ss 1); +val Lmap_NIL = result(); + +goal LList.thy "Lmap(f, CONS(M,N)) = CONS(f(M), Lmap(f,N))"; +by (rtac (Lmap_def RS def_LList_corec RS trans) 1); +by(simp_tac List_case_ss 1); +val Lmap_CONS = result(); + +(*Another type-checking proof by coinduction*) +val [major,minor] = goal LList.thy + "[| M: LList(A); !!x. x:A ==> f(x):B |] ==> Lmap(f,M): LList(B)"; +by (rtac (major RS imageI RS LList_coinduct) 1); +by (safe_tac set_cs); +by (etac LListE 1); +by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); +by (REPEAT (ares_tac [List_Fun_NIL_I, List_Fun_CONS_I, minor, imageI] 1)); +val Lmap_type = result(); + +(*This type checking rule synthesises a sufficiently large set for f*) +val [major] = goal LList.thy "M: LList(A) ==> Lmap(f,M): LList(f``A)"; +by (rtac (major RS Lmap_type) 1); +by (etac imageI 1); +val Lmap_type2 = result(); + +(** Two easy results about Lmap **) + +val [prem] = goal LList.thy + "M: LList(A) ==> Lmap(f o g, M) = Lmap(f, Lmap(g, M))"; +by (rtac (prem RS imageI RS LList_equalityI) 1); +by (stac o_def 1); +by (safe_tac set_cs); +by (etac LListE 1); +by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); +by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, + rangeI RS LListD_Fun_CONS_I] 1)); +val Lmap_compose = result(); + +val [prem] = goal LList.thy "M: LList(A) ==> Lmap(%x.x, M) = M"; +by (rtac (prem RS imageI RS LList_equalityI) 1); +by (safe_tac set_cs); +by (etac LListE 1); +by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); +by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, + rangeI RS LListD_Fun_CONS_I] 1)); +val Lmap_ident = result(); + + +(*** Lappend -- its two arguments cause some complications! ***) + +goalw LList.thy [Lappend_def] "Lappend(NIL,NIL) = NIL"; +by (rtac (LList_corec RS trans) 1); +(* takes 2.4(3.4 w NORM) vs 0.9 w/o NORM terms *) +by(simp_tac List_case_ss 1); +(*by (SIMP_TAC List_case_ss 1);*) +val Lappend_NIL_NIL = result(); + +goalw LList.thy [Lappend_def] + "Lappend(NIL,CONS(N,N')) = CONS(N, Lappend(NIL,N'))"; +by (rtac (LList_corec RS trans) 1); +(* takes 5(7 w NORM) vs 2.1 w/o NORM terms *) +by(simp_tac List_case_ss 1); +(*by (SIMP_TAC List_case_ss 1);*) +val Lappend_NIL_CONS = result(); + +goalw LList.thy [Lappend_def] + "Lappend(CONS(M,M'), N) = CONS(M, Lappend(M',N))"; +by (rtac (LList_corec RS trans) 1); +(* takes 4.9(6.7) vs 2.2 w/o NORM terms *) +by(simp_tac List_case_ss 1); +(*by (SIMP_TAC List_case_ss 1);*) +val Lappend_CONS = result(); + +val Lappend_ss = List_case_ss addsimps + [NIL_LListI, Lappend_NIL_NIL, Lappend_NIL_CONS, + Lappend_CONS, image_eqI, LListD_Fun_CONS_I]; + +goal LList.thy "!!M. M: LList(A) ==> Lappend(NIL,M) = M"; +by (etac LList_fun_equalityI 1); +by (ALLGOALS (asm_simp_tac Lappend_ss)); +val Lappend_NIL = result(); + +goal LList.thy "!!M. M: LList(A) ==> Lappend(M,NIL) = M"; +by (etac LList_fun_equalityI 1); +by (ALLGOALS (asm_simp_tac Lappend_ss)); +val Lappend_NIL2 = result(); + +(** Alternative type-checking proofs for Lappend **) + +(*weak co-induction: bisimulation and case analysis on both variables*) +goal LList.thy + "!!M N. [| M: LList(A); N: LList(A) |] ==> Lappend(M,N): LList(A)"; +by (res_inst_tac + [("X", "UN u:LList(A). UN v: LList(A). {Lappend(u,v)}")] LList_coinduct 1); +by (fast_tac set_cs 1); +by (safe_tac set_cs); +by (eres_inst_tac [("L", "u")] LListE 1); +by (eres_inst_tac [("L", "v")] LListE 1); +(* 7/12 vs 7.8/13.3/8.2/13.4 *) +by (ALLGOALS + (asm_simp_tac Lappend_ss THEN' + fast_tac (set_cs addSIs [NIL_LListI,List_Fun_NIL_I,List_Fun_CONS_I]) )); +(* +by (REPEAT + (ASM_SIMP_TAC Lappend_ss 1 THEN + fast_tac (set_cs addSIs [NIL_LListI,List_Fun_NIL_I,List_Fun_CONS_I])1)); +*) +val Lappend_type = result(); + +(*strong co-induction: bisimulation and case analysis on one variable*) +goal LList.thy + "!!M N. [| M: LList(A); N: LList(A) |] ==> Lappend(M,N): LList(A)"; +by (res_inst_tac [("X", "(%u.Lappend(u,N))``LList(A)")] LList_coinduct2 1); +fe imageI; +br subsetI 1; +be imageE 1; +by (eres_inst_tac [("L", "u")] LListE 1); +by (asm_simp_tac (Lappend_ss addsimps [Lappend_NIL]) 1); +by (asm_simp_tac Lappend_ss 1); +by (fast_tac (set_cs addSIs [List_Fun_CONS_I]) 1); +val Lappend_type = result(); + +(**** Lazy lists as the type 'a llist -- strongly typed versions of above ****) + +(** llist_case: case analysis for 'a llist **) + +val Rep_LList_simps = + [List_case_NIL, List_case_CONS, + Abs_LList_inverse, Rep_LList_inverse, NIL_LListI, CONS_LListI, + Rep_LList, rangeI, inj_Leaf, Inv_f_f]; +val Rep_LList_ss = llist_ss addsimps Rep_LList_simps; + +goalw LList.thy [llist_case_def,LNil_def] "llist_case(LNil, c, d) = c"; +by(simp_tac Rep_LList_ss 1); +val llist_case_LNil = result(); + +goalw LList.thy [llist_case_def,LCons_def] + "llist_case(LCons(M,N), c, d) = d(M,N)"; +by(simp_tac Rep_LList_ss 1); +val llist_case_LCons = result(); + +(*Elimination is case analysis, not induction.*) +val [prem1,prem2] = goalw LList.thy [NIL_def,CONS_def] + "[| l=LNil ==> P; !!x l'. l=LCons(x,l') ==> P \ +\ |] ==> P"; +by (rtac (Rep_LList RS LListE) 1); +by (rtac (inj_Rep_LList RS injD RS prem1) 1); +by (stac Rep_LList_LNil 1); +by (assume_tac 1); +by (etac rangeE 1); +by (rtac (inj_Rep_LList RS injD RS prem2) 1); +by(asm_simp_tac (HOL_ss addsimps [Rep_LList_LCons]) 1); +by (etac (Abs_LList_inverse RS ssubst) 1); +by (rtac refl 1); +val llistE = result(); + +(** llist_corec: corecursion for 'a llist **) + +goalw LList.thy [llist_corec_def,LNil_def,LCons_def] + "llist_corec(a,f) = case(f(a), %u. LNil, \ +\ %v. split(v, %z w. LCons(z, llist_corec(w,f))))"; +by (stac LList_corec 1); +by(res_inst_tac [("s","f(a)")] sumE 1); +by(asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1); +by(res_inst_tac [("p","y")] PairE 1); +by(asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1); +(*FIXME: correct case splits usd to be found automatically: +by(ASM_SIMP_TAC(llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);*) +val llist_corec = result(); + +(*definitional version of same*) +val [rew] = goal LList.thy + "[| !!x. h(x) == llist_corec(x,f) |] ==> \ +\ h(a) = case(f(a), %u.LNil, %v. split(v, %z w. LCons(z, h(w))))"; +by (rewtac rew); +by (rtac llist_corec 1); +val def_llist_corec = result(); + +(**** Proofs about type 'a llist functions ****) + +(*** Deriving llist_equalityI -- llist equality is a bisimulation ***) + +val prems = goalw LList.thy [LListD_Fun_def] + "r <= Sigma(LList(A), %x.LList(A)) ==> \ +\ LListD_Fun(diag(A),r) <= Sigma(LList(A), %x.LList(A))"; +by (stac LList_unfold 1); +by (cut_facts_tac prems 1); +by (fast_tac univ_cs 1); +val LListD_Fun_subset_Sigma_LList = result(); + +goal LList.thy + "prod_fun(Rep_LList,Rep_LList) `` r <= \ +\ Sigma(LList(range(Leaf)), %x.LList(range(Leaf)))"; +by (fast_tac (set_cs addSEs [prod_fun_imageE] addIs [SigmaI, Rep_LList]) 1); +val subset_Sigma_LList = result(); + +val [prem] = goal LList.thy + "r <= Sigma(LList(range(Leaf)), %x.LList(range(Leaf))) ==> \ +\ prod_fun(Rep_LList o Abs_LList, Rep_LList o Abs_LList) `` r <= r"; +by (safe_tac (set_cs addSEs [prod_fun_imageE])); +by (rtac (prem RS subsetD RS SigmaE2) 1); +by (assume_tac 1); +by(asm_simp_tac (HOL_ss addsimps [o_def,prod_fun,Abs_LList_inverse]) 1); +val prod_fun_lemma = result(); + +(** To show two llists are equal, exhibit a bisimulation! **) +val [prem1,prem2] = goalw LList.thy [llistD_Fun_def] + "[| : r; r <= llistD_Fun(r) |] ==> l1=l2"; +by (rtac (inj_Rep_LList RS injD) 1); +by (res_inst_tac [("r", "prod_fun(Rep_LList,Rep_LList)``r")] + LList_equalityI 1); +by (rtac (prem1 RS prod_fun_imageI) 1); +by (rtac (prem2 RS image_mono RS subset_trans) 1); +by (rtac (image_compose RS subst) 1); +by (rtac (prod_fun_compose RS subst) 1); +by (rtac (subset_Sigma_LList RS LListD_Fun_subset_Sigma_LList RS + prod_fun_lemma) 1); +val llist_equalityI = result(); + + +(*Stronger notion of bisimulation -- also admits true equality*) +val [prem1,prem2] = goalw LList.thy [llistD_Fun_def] + "[| : r; r <= llistD_Fun(r) Un range(%x.) |] ==> l1=l2"; +by (rtac (inj_Rep_LList RS injD) 1); +by (res_inst_tac [("r", "prod_fun(Rep_LList,Rep_LList)``r"), + ("A", "range(Leaf)")] + LList_equalityI2 1); +by (rtac (prem1 RS prod_fun_imageI) 1); +by (rtac (prem2 RS image_mono RS subset_trans) 1); +by (rtac (image_Un RS ssubst) 1); +by (rtac Un_least 1); +by (rtac (image_compose RS subst) 1); +by (rtac (prod_fun_compose RS subst) 1); +by (rtac (subset_Sigma_LList RS LListD_Fun_subset_Sigma_LList RS + prod_fun_lemma RS subset_trans) 1); +by (rtac Un_upper1 1); +by (fast_tac (set_cs addSEs [prod_fun_imageE, Pair_inject] + addIs [diagI,Rep_LList]) 1); +val llist_equalityI2 = result(); + +(** Rules to prove the 2nd premise of llist_equalityI **) +goalw LList.thy [llistD_Fun_def,LNil_def] " : llistD_Fun(r)"; +by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1); +val llistD_Fun_LNil_I = result(); + +val [prem] = goalw LList.thy [llistD_Fun_def,LCons_def] + ":r ==> : llistD_Fun(r)"; +by (rtac (rangeI RS LListD_Fun_CONS_I RS prod_fun_imageI) 1); +by (rtac (prem RS prod_fun_imageI) 1); +val llistD_Fun_LCons_I = result(); + + +(*A special case of list_equality for functions over lazy lists*) +val [prem1,prem2] = goal LList.thy + "[| f(LNil)=g(LNil); \ +\ !!x l. : \ +\ llistD_Fun(range(%u. )) Un range(%v. ) \ +\ |] ==> f(l) = g(l :: 'a llist) :: 'b llist"; +by (res_inst_tac [("r", "range(%u. )")] llist_equalityI2 1); +by (rtac rangeI 1); +by (rtac subsetI 1); +by (etac rangeE 1); +by (etac ssubst 1); +by (res_inst_tac [("l", "u")] llistE 1); +by (etac ssubst 1); +by (stac prem1 1); +by (fast_tac set_cs 1); +by (etac ssubst 1); +by (rtac prem2 1); +val llist_fun_equalityI = result(); + +(*simpset for llist bisimulations*) +val llistD_simps = [llist_case_LNil, llist_case_LCons, range_eqI, + llistD_Fun_LNil_I, llistD_Fun_LCons_I]; +val llistD_ss = llist_ss addsimps llistD_simps; + + +(*** The functional "lmap" ***) + +goal LList.thy "lmap(f,LNil) = LNil"; +by (rtac (lmap_def RS def_llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +val lmap_LNil = result(); + +goal LList.thy "lmap(f, LCons(M,N)) = LCons(f(M), lmap(f,N))"; +by (rtac (lmap_def RS def_llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +val lmap_LCons = result(); + + +(** Two easy results about lmap **) + +goal LList.thy "lmap(f o g, l) = lmap(f, lmap(g, l))"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons]))); +val lmap_compose = result(); + +goal LList.thy "lmap(%x.x, l) = l"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons]))); +val lmap_ident = result(); + + +(*** iterates -- llist_fun_equalityI cannot be used! ***) + +goal LList.thy "iterates(f,x) = LCons(x, iterates(f,f(x)))"; +by (rtac (iterates_def RS def_llist_corec RS trans) 1); +by(simp_tac sum_ss 1); +val iterates = result(); + +goal LList.thy "lmap(f, iterates(f,x)) = iterates(f,f(x))"; +by (res_inst_tac [("r", "range(%u.)")] + llist_equalityI 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (res_inst_tac [("x1", "f(u)")] (iterates RS ssubst) 1); +by (res_inst_tac [("x1", "u")] (iterates RS ssubst) 1); +by (simp_tac (llistD_ss addsimps [lmap_LCons]) 1); +val lmap_iterates = result(); + +goal LList.thy "iterates(f,x) = LCons(x, lmap(f, iterates(f,x)))"; +br (lmap_iterates RS ssubst) 1; +br iterates 1; +val iterates_lmap = result(); + +(*** A rather complex proof about iterates -- cf Andy Pitts ***) + +(** Two lemmas about natrec(n,x,%m.g), which is essentially (g^n)(x) **) + +goal LList.thy + "nat_rec(n, LCons(b, l), %m. lmap(f)) = \ +\ LCons(nat_rec(n, b, %m. f), nat_rec(n, l, %m. lmap(f)))"; +by (nat_ind_tac "n" 1); +by(ALLGOALS (asm_simp_tac (nat_ss addsimps [lmap_LCons]))); +val fun_power_lmap = result(); + +goal Nat.thy "nat_rec(n, g(x), %m. g) = nat_rec(Suc(n), x, %m. g)"; +by (nat_ind_tac "n" 1); +by(ALLGOALS (asm_simp_tac nat_ss)); +val fun_power_Suc = result(); + +val Pair_cong = read_instantiate_sg (sign_of Prod.thy) + [("f","Pair")] (standard(refl RS cong RS cong)); + +(*The bisimulation consists of {} + for all u and all n::nat.*) +val [prem] = goal LList.thy + "(!!x. h(x) = LCons(x, lmap(f,h(x)))) ==> h = iterates(f)"; +br ext 1; +by (res_inst_tac [("r", + "UN u. range(%n. )")] + llist_equalityI 1); +by (REPEAT (resolve_tac [UN1_I, range_eqI, Pair_cong, nat_rec_0 RS sym] 1)); +by (safe_tac set_cs); +by (stac iterates 1); +by (stac prem 1); +by (stac fun_power_lmap 1); +by (stac fun_power_lmap 1); +br llistD_Fun_LCons_I 1; +by (rtac (lmap_iterates RS subst) 1); +by (stac fun_power_Suc 1); +by (stac fun_power_Suc 1); +br UN1_I 1; +br rangeI 1; +val iterates_equality = result(); + + +(*** lappend -- its two arguments cause some complications! ***) + +goalw LList.thy [lappend_def] "lappend(LNil,LNil) = LNil"; +by (rtac (llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +val lappend_LNil_LNil = result(); + +goalw LList.thy [lappend_def] + "lappend(LNil,LCons(l,l')) = LCons(l, lappend(LNil,l'))"; +by (rtac (llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +(* 3.3(5.7) vs 1.3 !by (SIMP_TAC llistD_ss 1);*) +val lappend_LNil_LCons = result(); + +goalw LList.thy [lappend_def] + "lappend(LCons(l,l'), N) = LCons(l, lappend(l',N))"; +by (rtac (llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +(* 5(5.5) vs 1.3 !by (SIMP_TAC llistD_ss 1);*) +val lappend_LCons = result(); + +goal LList.thy "lappend(LNil,l) = l"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac (llistD_ss addsimps [lappend_LNil_LNil, + lappend_LNil_LCons]))); +val lappend_LNil = result(); + +goal LList.thy "lappend(l,LNil) = l"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac(llistD_ss addsimps [lappend_LNil_LNil,lappend_LCons]))); +val lappend_LNil2 = result(); + +(*The infinite first argument blocks the second*) +goal LList.thy "lappend(iterates(f,x), N) = iterates(f,x)"; +by (res_inst_tac [("r", "range(%u.)")] + llist_equalityI 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac iterates 1); +by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1); +val lappend_iterates = result(); + +(** Two proofs that lmap distributes over lappend **) + +(*Long proof requiring case analysis on both both arguments*) +goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))"; +by (res_inst_tac + [("r", + "UN n. range(%l.)")] + llist_equalityI 1); +by (rtac UN1_I 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (res_inst_tac [("l", "l")] llistE 1); +by (res_inst_tac [("l", "n")] llistE 1); +by (ALLGOALS (asm_simp_tac (llistD_ss addsimps + [lappend_LNil_LNil,lappend_LCons,lappend_LNil_LCons, + lmap_LNil,lmap_LCons]))); +by (REPEAT_SOME (ares_tac [llistD_Fun_LCons_I, UN1_I, rangeI])); +by (rtac range_eqI 1); +by (rtac (refl RS Pair_cong) 1); +by (stac lmap_LNil 1); +by (rtac refl 1); +val lmap_lappend_distrib = result(); + +(*Shorter proof of the theorem above using llist_equalityI2*) +goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (simp_tac (llistD_ss addsimps [lappend_LNil, lmap_LNil])1); +by (simp_tac (llistD_ss addsimps [lappend_LCons, lmap_LCons]) 1); +val lmap_lappend_distrib = result(); + +(*Without llist_equalityI2, three case analyses might be needed*) +goal LList.thy "lappend(lappend(l1,l2) ,l3) = lappend(l1, lappend(l2,l3))"; +by (res_inst_tac [("l","l1")] llist_fun_equalityI 1); +by (simp_tac (llistD_ss addsimps [lappend_LNil])1); +by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1); +val lappend_assoc = result(); diff -r 000000000000 -r 7949f97df77a LList.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/LList.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,126 @@ +(* Title: HOL/llist.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Definition of type 'a llist by a greatest fixed point + +Shares NIL, CONS, List_Fun, List_case with list.thy + +Still needs filter and flatten functions -- hard because they need +bounds on the amount of lookahead required. + +Could try (but would it work for the gfp analogue of term?) + LListD_Fun_def "LListD_Fun(A) == (%Z.diag({Numb(0)}) <++> diag(A) <**> Z)" + +A nice but complex example would be [ML for the Working Programmer, page 176] + from(1) = enumerate (Lmap (Lmap(pack), makeqq(from(1),from(1)))) + +Previous definition of llistD_Fun was explicit: + llistD_Fun_def + "llistD_Fun(r) == \ +\ {} Un \ +\ (UN x. (%z.split(z, %l1 l2.))``r)" +*) + +LList = Gfp + List + +types llist 1 +arities llist :: (term)term +consts + LList :: "'a node set set => 'a node set set" + LListD_Fun :: + "[('a node set * 'a node set)set, ('a node set * 'a node set)set] => \ +\ ('a node set * 'a node set)set" + LListD :: + "('a node set * 'a node set)set => ('a node set * 'a node set)set" + llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set" + + Rep_LList :: "'a llist => 'a node set" + Abs_LList :: "'a node set => 'a llist" + LNil :: "'a llist" + LCons :: "['a, 'a llist] => 'a llist" + + llist_case :: "['a llist, 'b, ['a, 'a llist]=>'b] => 'b" + + LList_corec_fun :: "[nat, 'a=>unit+('b node set * 'a), 'a] => 'b node set" + LList_corec :: "['a, 'a => unit + ('b node set * 'a)] => 'b node set" + llist_corec :: "['a, 'a => unit + ('b * 'a)] => 'b llist" + + Lmap :: "('a node set => 'b node set) => ('a node set => 'b node set)" + lmap :: "('a=>'b) => ('a llist => 'b llist)" + + iterates :: "['a => 'a, 'a] => 'a llist" + + Lconst :: "'a node set => 'a node set" + Lappend :: "['a node set, 'a node set] => 'a node set" + lappend :: "['a llist, 'a llist] => 'a llist" + + +rules + LListD_Fun_def "LListD_Fun(r) == (%Z.diag({Numb(0)}) <++> r <**> Z)" + LList_def "LList(A) == gfp(List_Fun(A))" + LListD_def "LListD(r) == gfp(LListD_Fun(r))" + (*faking a type definition...*) + Rep_LList "Rep_LList(xs): LList(range(Leaf))" + Rep_LList_inverse "Abs_LList(Rep_LList(xs)) = xs" + Abs_LList_inverse "M: LList(range(Leaf)) ==> Rep_LList(Abs_LList(M)) = M" + (*defining the abstract constructors*) + LNil_def "LNil == Abs_LList(NIL)" + LCons_def "LCons(x,xs) == Abs_LList(CONS(Leaf(x), Rep_LList(xs)))" + + llist_case_def + "llist_case(l,c,d) == \ +\ List_case(Rep_LList(l), c, %x y. d(Inv(Leaf,x), Abs_LList(y)))" + + LList_corec_fun_def + "LList_corec_fun(k,f) == \ +\ nat_rec(k, %x. {}, \ +\ %j r x. case(f(x), %u.NIL, \ +\ %v. split(v, %z w. CONS(z, r(w)))))" + + LList_corec_def + "LList_corec(a,f) == UN k. LList_corec_fun(k,f,a)" + + llist_corec_def + "llist_corec(a,f) == \ +\ Abs_LList(LList_corec(a, %z.case(f(z), %x.Inl(x), \ +\ %y.split(y, %v w. Inr()))))" + + llistD_Fun_def + "llistD_Fun(r) == \ +\ prod_fun(Abs_LList,Abs_LList) `` \ +\ LListD_Fun(diag(range(Leaf)), \ +\ prod_fun(Rep_LList,Rep_LList) `` r)" + + Lconst_def "Lconst(M) == lfp(%N. CONS(M, N))" + + Lmap_def + "Lmap(f,M) == \ +\ LList_corec(M, %M. List_case(M, Inl(Unity), %x M'. Inr()))" + + lmap_def + "lmap(f,l) == \ +\ llist_corec(l, %l. llist_case(l, Inl(Unity), %y z. Inr()))" + + iterates_def "iterates(f,a) == llist_corec(a, %x. Inr())" + +(*Append generates its result by applying f, where + f() = Inl(Unity) + f() = Inr() + f() = Inr() +*) + + Lappend_def + "Lappend(M,N) == LList_corec(, %p. split(p, \ +\ %M N. List_case(M, List_case(N, Inl(Unity), \ +\ %N1 N2. Inr(>)), \ +\ %M1 M2. Inr(>))))" + + lappend_def + "lappend(l,n) == llist_corec(, %p. split(p, \ +\ %l n. llist_case(l, llist_case(n, Inl(Unity), \ +\ %n1 n2. Inr(>)), \ +\ %l1 l2. Inr(>))))" + + +end diff -r 000000000000 -r 7949f97df77a Lfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Lfp.ML Thu Sep 16 12:21:07 1993 +0200 @@ -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); +val lfp_lowerbound = result(); + +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); +val lfp_greatest = result(); + +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]); +val lfp_lemma2 = result(); + +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]); +val lfp_lemma3 = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))"; +by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1)); +val lfp_Tarski = result(); + + +(*** 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]); +val induct = result(); + +(** 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); +val def_lfp_Tarski = result(); + +val rew::prems = goal Lfp.thy + "[| A == lfp(f); a:A; mono(f); \ +\ !!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))]); +val def_induct = result(); + +(*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; +val lfp_mono = result(); diff -r 000000000000 -r 7949f97df77a Lfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Lfp.thy Thu Sep 16 12:21:07 1993 +0200 @@ -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 = Sum + +consts lfp :: "['a set=>'a set] => 'a set" +rules + (*least fixed point*) + lfp_def "lfp(f) == Inter({u. f(u) <= u})" +end diff -r 000000000000 -r 7949f97df77a List.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/List.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,362 @@ +(* Title: HOL/list + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For list.thy. +*) + +open List; + +(** the list functional **) + +goalw List.thy [List_Fun_def] "mono(List_Fun(A))"; +by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono] 1)); +val List_Fun_mono = result(); + +goalw List.thy [List_Fun_def] + "!!A B. A<=B ==> List_Fun(A,Z) <= List_Fun(B,Z)"; +by (REPEAT (ares_tac [subset_refl, usum_mono, uprod_mono] 1)); +val List_Fun_mono2 = result(); + +(*This justifies using List in other recursive type definitions*) +goalw List.thy [List_def] "!!A B. A<=B ==> List(A) <= List(B)"; +by (rtac lfp_mono 1); +by (etac List_Fun_mono2 1); +val List_mono = result(); + +(** Type checking rules -- List creates well-founded sets **) + +val prems = goalw List.thy [List_def,List_Fun_def] "List(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_NumbI,Sexp_In0I,Sexp_In1I,Sexp_SconsI]) 1); +val List_Sexp = result(); + +(* A <= Sexp ==> List(A) <= Sexp *) +val List_subset_Sexp = standard + (List_mono RS (List_Sexp RSN (2,subset_trans))); + +(** Induction **) + +(*Induction for the set List(A) *) +val major::prems = goalw List.thy [NIL_def,CONS_def] + "[| M: List(A); P(NIL); \ +\ !!M N. [| M: A; N: List(A); P(N) |] ==> P(CONS(M,N)) |] \ +\ ==> P(M)"; +by (rtac (major RS (List_def RS def_induct)) 1); +by (rtac List_Fun_mono 1); +by (rewtac List_Fun_def); +by (fast_tac (set_cs addIs prems addEs [usumE,uprodE]) 1); +val List_induct = result(); + +(*Induction for the type 'a list *) +val prems = goalw List.thy [Nil_def,Cons_def] + "[| P(Nil); \ +\ !!x xs. P(xs) ==> P(Cons(x,xs)) |] ==> P(l)"; +by (rtac (Rep_List_inverse RS subst) 1); (*types force good instantiation*) +by (rtac (Rep_List RS List_induct) 1); +by (REPEAT (ares_tac prems 1 + ORELSE eresolve_tac [rangeE, ssubst, Abs_List_inverse RS subst] 1)); +val list_induct = result(); + +(*Perform induction on xs. *) +fun list_ind_tac a M = + EVERY [res_inst_tac [("l",a)] list_induct M, + rename_last_tac a ["1"] (M+1)]; + +(** Introduction rules for List constructors **) + +val List_unfold = rewrite_rule [List_Fun_def] + (List_Fun_mono RS (List_def RS def_lfp_Tarski)); + +(* c : {Numb(0)} <+> A <*> List(A) ==> c : List(A) *) +val ListI = List_unfold RS equalityD2 RS subsetD; + +(* NIL is a List -- this also justifies the type definition*) +goalw List.thy [NIL_def] "NIL: List(A)"; +by (rtac (singletonI RS usum_In0I RS ListI) 1); +val NIL_I = result(); + +goalw List.thy [CONS_def] + "!!a A M. [| a: A; M: List(A) |] ==> CONS(a,M) : List(A)"; +by (REPEAT (ares_tac [uprodI RS usum_In1I RS ListI] 1)); +val CONS_I = result(); + +(*** Isomorphisms ***) + +goal List.thy "inj(Rep_List)"; +by (rtac inj_inverseI 1); +by (rtac Rep_List_inverse 1); +val inj_Rep_List = result(); + +goal List.thy "inj_onto(Abs_List,List(range(Leaf)))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_List_inverse 1); +val inj_onto_Abs_List = result(); + +(** Distinctness of constructors **) + +goalw List.thy [NIL_def,CONS_def] "~ CONS(M,N) = NIL"; +by (rtac In1_not_In0 1); +val CONS_not_NIL = result(); +val NIL_not_CONS = standard (CONS_not_NIL RS not_sym); + +val CONS_neq_NIL = standard (CONS_not_NIL RS notE); +val NIL_neq_CONS = sym RS CONS_neq_NIL; + +goalw List.thy [Nil_def,Cons_def] "~ Cons(x,xs) = Nil"; +by (rtac (CONS_not_NIL RS (inj_onto_Abs_List RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, NIL_I, CONS_I, Rep_List] 1)); +val Cons_not_Nil = result(); + +val Nil_not_Cons = standard (Cons_not_Nil RS not_sym); + +val Cons_neq_Nil = standard (Cons_not_Nil RS notE); +val Nil_neq_Cons = sym RS Cons_neq_Nil; + +(** Injectiveness of CONS and Cons **) + +goalw List.thy [CONS_def] "(CONS(K,M)=CONS(L,N)) = (K=L & M=N)"; +by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1); +val CONS_CONS_eq = result(); + +val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE); + +(*For reasoning about abstract list constructors*) +val List_cs = set_cs addIs [Rep_List, NIL_I, CONS_I] + addSEs [CONS_neq_NIL,NIL_neq_CONS,CONS_inject] + addSDs [inj_onto_Abs_List RS inj_ontoD, + inj_Rep_List RS injD, Leaf_inject]; + +goalw List.thy [Cons_def] "(Cons(x,xs)=Cons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac List_cs 1); +val Cons_Cons_eq = result(); +val Cons_inject = standard (Cons_Cons_eq RS iffD1 RS conjE); + +val [major] = goal List.thy "CONS(M,N): List(A) ==> M: A & N: List(A)"; +by (rtac (major RS setup_induction) 1); +by (etac List_induct 1); +by (ALLGOALS (fast_tac List_cs)); +val CONS_D = result(); + +val prems = goalw List.thy [CONS_def,In1_def] + "CONS(M,N): Sexp ==> M: Sexp & N: Sexp"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSDs [Scons_D]) 1); +val Sexp_CONS_D = result(); + + +(*Basic ss with constructors and their freeness*) +val list_free_simps = [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq, + CONS_not_NIL, NIL_not_CONS, CONS_CONS_eq, + NIL_I, CONS_I]; +val list_free_ss = HOL_ss addsimps list_free_simps; + +goal List.thy "!!N. N: List(A) ==> !M. ~ N = CONS(M,N)"; +by (etac List_induct 1); +by (ALLGOALS (asm_simp_tac list_free_ss)); +val not_CONS_self = result(); + +goal List.thy "!x. ~ l=Cons(x,l)"; +by (list_ind_tac "l" 1); +by (ALLGOALS (asm_simp_tac list_free_ss)); +val not_Cons_self = result(); + + +(** Conversion rules for List_case: case analysis operator **) + +goalw List.thy [List_case_def,NIL_def] "List_case(NIL,c,h) = c"; +by (rtac Case_In0 1); +val List_case_NIL = result(); + +goalw List.thy [List_case_def,CONS_def] "List_case(CONS(M,N), c, h) = h(M,N)"; +by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1); +val List_case_CONS = result(); + +(*** List_rec -- by wf recursion on pred_Sexp ***) + +(* The trancl(pred_sexp) is essential because pred_Sexp_CONS_I1,2 would not + hold if pred_Sexp^+ were changed to pred_Sexp. *) + +val List_rec_unfold = wf_pred_Sexp RS wf_trancl RS (List_rec_def RS def_wfrec); + +(** pred_Sexp lemmas **) + +goalw List.thy [CONS_def,In1_def] + "!!M. [| M: Sexp; N: Sexp |] ==> : pred_Sexp^+"; +by (asm_simp_tac pred_Sexp_ss 1); +val pred_Sexp_CONS_I1 = result(); + +goalw List.thy [CONS_def,In1_def] + "!!M. [| M: Sexp; N: Sexp |] ==> : pred_Sexp^+"; +by (asm_simp_tac pred_Sexp_ss 1); +val pred_Sexp_CONS_I2 = result(); + +val [prem] = goal List.thy + " : pred_Sexp^+ ==> \ +\ : pred_Sexp^+ & : pred_Sexp^+"; +by (rtac (prem RS (pred_Sexp_subset_Sigma RS trancl_subset_Sigma RS + subsetD RS SigmaE2)) 1); +by (etac (Sexp_CONS_D RS conjE) 1); +by (REPEAT (ares_tac [conjI, pred_Sexp_CONS_I1, pred_Sexp_CONS_I2, + prem RSN (2, trans_trancl RS transD)] 1)); +val pred_Sexp_CONS_D = result(); + +(** Conversion rules for List_rec **) + +goal List.thy "List_rec(NIL,c,h) = c"; +by (rtac (List_rec_unfold RS trans) 1); +by (rtac List_case_NIL 1); +val List_rec_NIL = result(); + +goal List.thy "!!M. [| M: Sexp; N: Sexp |] ==> \ +\ List_rec(CONS(M,N), c, h) = h(M, N, List_rec(N,c,h))"; +by (rtac (List_rec_unfold RS trans) 1); +by (rtac (List_case_CONS RS trans) 1); +by(asm_simp_tac(HOL_ss addsimps [CONS_I, pred_Sexp_CONS_I2, cut_apply])1); +val List_rec_CONS = result(); + +(*** list_rec -- by List_rec ***) + +val Rep_List_in_Sexp = + Rep_List RS (range_Leaf_subset_Sexp RS List_subset_Sexp RS subsetD); + +local + val list_rec_simps = list_free_simps @ + [List_rec_NIL, List_rec_CONS, + Abs_List_inverse, Rep_List_inverse, + Rep_List, rangeI, inj_Leaf, Inv_f_f, + Sexp_LeafI, Rep_List_in_Sexp] +in + val list_rec_Nil = prove_goalw List.thy [list_rec_def, Nil_def] + "list_rec(Nil,c,h) = c" + (fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]); + + val list_rec_Cons = prove_goalw List.thy [list_rec_def, Cons_def] + "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))" + (fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]); +end; + +val list_simps = [List_rec_NIL, List_rec_CONS, + list_rec_Nil, list_rec_Cons]; +val list_ss = list_free_ss addsimps list_simps; + + +(*Type checking. Useful?*) +val major::A_subset_Sexp::prems = goal List.thy + "[| M: List(A); \ +\ A<=Sexp; \ +\ c: C(NIL); \ +\ !!x y r. [| x: A; y: List(A); r: C(y) |] ==> h(x,y,r): C(CONS(x,y)) \ +\ |] ==> List_rec(M,c,h) : C(M :: 'a node set)"; +val Sexp_ListA_I = A_subset_Sexp RS List_subset_Sexp RS subsetD; +val Sexp_A_I = A_subset_Sexp RS subsetD; +by (rtac (major RS List_induct) 1); +by (ALLGOALS(asm_simp_tac (list_ss addsimps ([Sexp_A_I,Sexp_ListA_I]@prems)))); +val List_rec_type = result(); + +(** Generalized map functionals **) + +goalw List.thy [Rep_map_def] "Rep_map(f,Nil) = NIL"; +by (rtac list_rec_Nil 1); +val Rep_map_Nil = result(); + +goalw List.thy [Rep_map_def] + "Rep_map(f, Cons(x,xs)) = CONS(f(x), Rep_map(f,xs))"; +by (rtac list_rec_Cons 1); +val Rep_map_Cons = result(); + +goalw List.thy [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map(f,xs): List(A)"; +by (rtac list_induct 1); +by(ALLGOALS(asm_simp_tac list_ss)); +val Rep_map_type = result(); + +goalw List.thy [Abs_map_def] "Abs_map(g,NIL) = Nil"; +by (rtac List_rec_NIL 1); +val Abs_map_NIL = result(); + +val prems = goalw List.thy [Abs_map_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ Abs_map(g, CONS(M,N)) = Cons(g(M), Abs_map(g,N))"; +by (REPEAT (resolve_tac (List_rec_CONS::prems) 1)); +val Abs_map_CONS = result(); + +(** null, hd, tl, list_case **) + +goalw List.thy [null_def] "null([]) = True"; +by (rtac list_rec_Nil 1); +val null_Nil = result(); + +goalw List.thy [null_def] "null(Cons(x,xs)) = False"; +by (rtac list_rec_Cons 1); +val null_Cons = result(); + + +goalw List.thy [hd_def] "hd(Cons(x,xs)) = x"; +by (rtac list_rec_Cons 1); +val hd_Cons = result(); + + +goalw List.thy [tl_def] "tl(Cons(x,xs)) = xs"; +by (rtac list_rec_Cons 1); +val tl_Cons = result(); + + +goalw List.thy [list_case_def] "list_case([],a,f) = a"; +by (rtac list_rec_Nil 1); +val list_case_Nil = result(); + +goalw List.thy [list_case_def] "list_case(Cons(x,xs),a,f) = f(x,xs)"; +by (rtac list_rec_Cons 1); +val list_case_Cons = result(); + + +(** The functional "map" **) + +goalw List.thy [map_def] "map(f,Nil) = Nil"; +by (rtac list_rec_Nil 1); +val map_Nil = result(); + +goalw List.thy [map_def] "map(f, Cons(x,xs)) = Cons(f(x), map(f,xs))"; +by (rtac list_rec_Cons 1); +val map_Cons = result(); + +val map_simps = [Abs_map_NIL, Abs_map_CONS, + Rep_map_Nil, Rep_map_Cons, + map_Nil, map_Cons]; +val map_ss = list_free_ss addsimps map_simps; + +val [major,A_subset_Sexp,minor] = goal List.thy + "[| M: List(A); A<=Sexp; !!z. z: A ==> f(g(z)) = z |] \ +\ ==> Rep_map(f, Abs_map(g,M)) = M"; +by (rtac (major RS List_induct) 1); +by (ALLGOALS (asm_simp_tac(map_ss addsimps [Sexp_A_I,Sexp_ListA_I,minor]))); +val Abs_map_inverse = result(); + +(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*) + + +(** The functional "list_all" -- creates predicates over lists **) + +goalw List.thy [list_all_def] "list_all(P,Nil) = True"; +by (rtac list_rec_Nil 1); +val list_all_Nil = result(); + +goalw List.thy [list_all_def] + "list_all(P, Cons(x,xs)) = (P(x) & list_all(P,xs))"; +by (rtac list_rec_Cons 1); +val list_all_Cons = result(); + +(** Additional mapping lemmas **) + +goal List.thy "map(%x.x, xs) = xs"; +by (list_ind_tac "xs" 1); +by (ALLGOALS (asm_simp_tac map_ss)); +val map_ident = result(); + +goal List.thy "!!f. (!!x. f(x): Sexp) ==> \ +\ Abs_map(g, Rep_map(f,xs)) = map(%t. g(f(t)), xs)"; +by (list_ind_tac "xs" 1); +by(ALLGOALS(asm_simp_tac(map_ss addsimps + [Rep_map_type,List_Sexp RS subsetD]))); +val Abs_Rep_map = result(); diff -r 000000000000 -r 7949f97df77a List.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/List.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,104 @@ +(* Title: HOL/list + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Definition of type 'a list by a least fixed point + +We use List(A) == lfp(%Z. {NUMB(0)} <+> A <*> Z) +and not List == lfp(%Z. {NUMB(0)} <+> range(Leaf) <*> Z) +so that List can serve as a "functor" for defining other recursive types +*) + +List = Sexp + + +types + list 1 + +arities + list :: (term) term + + +consts + + List_Fun :: "['a node set set, 'a node set set] => 'a node set set" + List :: "'a node set set => 'a node set set" + Rep_List :: "'a list => 'a node set" + Abs_List :: "'a node set => 'a list" + NIL :: "'a node set" + CONS :: "['a node set, 'a node set] => 'a node set" + Nil :: "'a list" + Cons :: "['a, 'a list] => 'a list" + List_case :: "['a node set, 'b, ['a node set, 'a node set]=>'b] => 'b" + List_rec :: "['a node set, 'b, ['a node set, 'a node set, 'b]=>'b] => 'b" + list_rec :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b" + Rep_map :: "('b => 'a node set) => ('b list => 'a node set)" + Abs_map :: "('a node set => 'b) => 'a node set => 'b list" + null :: "'a list => bool" + hd :: "'a list => 'a" + tl :: "'a list => 'a list" + list_all :: "('a => bool) => ('a list => bool)" + map :: "('a=>'b) => ('a list => 'b list)" + list_case :: "['a list, 'b, ['a, 'a list]=>'b] => 'b" + + (* List Enumeration *) + + "[]" :: "'a list" ("[]") + "@List" :: "args => 'a list" ("[(_)]") + + +translations + "[x, xs]" == "Cons(x, [xs])" + "[x]" == "Cons(x, [])" + "[]" == "Nil" + + +rules + + List_Fun_def "List_Fun(A) == (%Z. {Numb(0)} <+> A <*> Z)" + List_def "List(A) == lfp(List_Fun(A))" + + (* Faking a Type Definition ... *) + + Rep_List "Rep_List(xs): List(range(Leaf))" + Rep_List_inverse "Abs_List(Rep_List(xs)) = xs" + Abs_List_inverse "M: List(range(Leaf)) ==> Rep_List(Abs_List(M)) = M" + + (* Defining the Concrete Constructors *) + + NIL_def "NIL == In0(Numb(0))" + CONS_def "CONS(M, N) == In1(M . N)" + + (* Defining the Abstract Constructors *) + + Nil_def "Nil == Abs_List(NIL)" + Cons_def "Cons(x, xs) == Abs_List(CONS(Leaf(x), Rep_List(xs)))" + + List_case_def "List_case(M, c, d) == Case(M, %x.c, %u. Split(u, %x y.d(x, y)))" + + (* List Recursion -- the trancl is Essential; see list.ML *) + + List_rec_def + "List_rec(M, c, d) == wfrec(trancl(pred_Sexp), M, \ +\ %z g. List_case(z, c, %x y. d(x, y, g(y))))" + + list_rec_def + "list_rec(l, c, d) == \ +\ List_rec(Rep_List(l), c, %x y r. d(Inv(Leaf, x), Abs_List(y), r))" + + (* Generalized Map Functionals *) + + Rep_map_def + "Rep_map(f, xs) == list_rec(xs, NIL, %x l r. CONS(f(x), r))" + Abs_map_def + "Abs_map(g, M) == List_rec(M, Nil, %N L r. Cons(g(N), r))" + + null_def "null(xs) == list_rec(xs, True, %x xs r.False)" + hd_def "hd(xs) == list_rec(xs, @x.True, %x xs r.x)" + tl_def "tl(xs) == list_rec(xs, @xs.True, %x xs r.xs)" + list_all_def "list_all(P, xs) == list_rec(xs, True, %x l r. P(x) & r)" + map_def "map(f, xs) == list_rec(xs, Nil, %x l r. Cons(f(x), r))" + list_case_def "list_case(xs, a, f) == list_rec(xs, a, %x xs r.f(x, xs))" + +end + diff -r 000000000000 -r 7949f97df77a Makefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Makefile Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,49 @@ +######################################################################### +# # +# Makefile for Isabelle (HOL) # +# # +######################################################################### + +#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) +FILES = ROOT.ML hol.thy hol.ML simpdata.ML ord.thy ord.ML \ + fun.ML set.thy set.ML subset.ML equalities.ML \ + prod.thy prod.ML sum.thy sum.ML wf.thy wf.ML \ + mono.ML lfp.thy lfp.ML gfp.thy gfp.ML nat.thy nat.ML sexp.thy \ + sexp.ML univ.thy univ.ML llist.thy llist.ML list.thy list.ML \ + ../Provers/classical.ML ../Provers/simplifier.ML \ + ../Provers/splitter.ML ../Provers/ind.ML + +$(BIN)/HOL: $(BIN)/Pure $(FILES) + case "$(COMP)" in \ + poly*) echo 'make_database"$(BIN)/HOL"; quit();' \ + | $(COMP) $(BIN)/Pure;\ + echo 'open PolyML; use"ROOT";' | $(COMP) $(BIN)/HOL ;;\ + sml*) echo 'use"ROOT.ML"; xML"$(BIN)/HOL" banner;' | $(BIN)/Pure ;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +$(BIN)/Pure: + cd ../Pure; $(MAKE) + +test: ex/ROOT.ML $(BIN)/HOL + case "$(COMP)" in \ + poly*) echo 'use"ex/ROOT.ML"; use"Subst/ROOT.ML"; quit();' \ + | $(COMP) $(BIN)/HOL ;;\ + sml*) echo 'use"ex/ROOT.ML";use"Subst/ROOT.ML";' | $(BIN)/HOL;;\ + *) echo Bad value for ISABELLECOMP;;\ + esac + +.PRECIOUS: $(BIN)/Pure $(BIN)/HOL diff -r 000000000000 -r 7949f97df77a Nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nat.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,409 @@ +(* 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)); +val Nat_fun_mono = result(); + +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); +val Zero_RepI = result(); + +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); +val Suc_RepI = result(); + +(*** 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 (major RS (Nat_def RS def_induct)) 1); +by (rtac Nat_fun_mono 1); +by (fast_tac (set_cs addIs prems) 1); +val Nat_induct = result(); + +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)); +val nat_induct = result(); + +(*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)); +val diff_induct = result(); + +(*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); +val natE = result(); + +(*** 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); +val inj_Rep_Nat = result(); + +goal Nat.thy "inj_onto(Abs_Nat,Nat)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Nat_inverse 1); +val inj_onto_Abs_Nat = result(); + +(*** 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)); +val Suc_not_Zero = result(); + +val Zero_not_Suc = standard (Suc_not_Zero RS not_sym); + +val Suc_neq_Zero = standard (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); +val inj_Suc = result(); + +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]); +val Suc_Suc_eq = result(); + +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]))); +val n_not_Suc_n = result(); + +(*** nat_case -- the selection operator for nat ***) + +goalw Nat.thy [nat_case_def] "nat_case(0, a, f) = a"; +by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1); +val nat_case_0 = result(); + +goalw Nat.thy [nat_case_def] "nat_case(Suc(k), a, f) = f(k)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim Suc_inject, Suc_neq_Zero]) 1); +val nat_case_Suc = result(); + +(** Introduction rules for 'pred_nat' **) + +goalw Nat.thy [pred_nat_def] " : pred_nat"; +by (fast_tac set_cs 1); +val pred_natI = result(); + +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)); +val pred_natE = result(); + +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); +val wf_pred_nat = result(); + + +(*** nat_rec -- by wf recursion on pred_nat ***) + +val nat_rec_unfold = standard (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 (rtac nat_case_0 1); +val nat_rec_0 = result(); + +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 (rtac (nat_case_Suc RS trans) 1); +by(simp_tac (HOL_ss addsimps [pred_natI,cut_apply]) 1); +val nat_rec_Suc = result(); + +(*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); +val def_nat_rec_0 = result(); + +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); +val def_nat_rec_Suc = result(); + +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 i ~ m R *) +val less_anti_sym = standard (less_not_sym RS mp RS notE); + + +goalw Nat.thy [less_def] "~ n R *) +val less_anti_refl = standard (less_not_refl RS notE); + + +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); +val lessE = result(); + +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); +val not_less0 = result(); + +(* n<0 ==> R *) +val less_zeroE = standard (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); +val less_SucE = result(); + +goal Nat.thy "(m < Suc(n)) = (m < n | m = n)"; +by (fast_tac (HOL_cs addSIs [lessI] + addEs [less_trans, less_SucE]) 1); +val less_Suc_eq = result(); + + +(** 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); +val Suc_lessE = result(); + +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); +val Suc_mono = result(); + +goal Nat.thy "(Suc(m) < Suc(n)) = (m P"; +by (rtac (major RS Suc_lessD RS less_anti_refl) 1); +val not_Suc_n_less_n = result(); + +(*"Less than" is a linear ordering*) +goal Nat.thy "m P(m) |] ==> P(n) |] ==> P(n)"; +by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); +by (eresolve_tac prems 1); +val less_induct = result(); + + +(*** Properties of <= ***) + +goalw Nat.thy [le_def] "0 <= n"; +by (rtac not_less0 1); +val le0 = result(); + +val nat_simps = [not_less0, less_not_refl, zero_less_Suc, lessI, + Suc_less_eq, less_Suc_eq, le0, + Suc_not_Zero, Zero_not_Suc, Suc_Suc_eq, + nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; + +val nat_ss = pair_ss addsimps nat_simps; + +val prems = goalw Nat.thy [le_def] "~(n m<=n::nat"; +by (resolve_tac prems 1); +val leI = result(); + +val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n n Suc(m) <= n"; +by(simp_tac (HOL_ss addsimps [less_Suc_eq]) 1); +by (fast_tac (HOL_cs addEs [less_anti_refl,less_anti_sym]) 1); +val lessD = result(); + +goalw Nat.thy [le_def] "!!m. m < n ==> m <= n::nat"; +by (fast_tac (HOL_cs addEs [less_anti_sym]) 1); +val less_imp_le = result(); + +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_anti_sym]) 1); +val le_imp_less_or_eq = result(); + +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_anti_sym]) 1); +by (flexflex_tac); +val less_or_eq_imp_le = result(); + +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)); +val le_eq_less_or_eq = result(); + +goal Nat.thy "n <= n::nat"; +by(simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1); +val le_refl = result(); + +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); +val le_less_trans = result(); + +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])]); +val le_trans = result(); + +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_anti_sym])]); +val le_anti_sym = result(); diff -r 000000000000 -r 7949f97df77a Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nat.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,51 @@ +(* Title: HOL/nat + 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 + +types ind,nat 0 +arities ind,nat :: term + nat :: ord + +consts + Zero_Rep :: "ind" + Suc_Rep :: "ind => ind" + Nat :: "ind set" + Rep_Nat :: "nat => ind" + Abs_Nat :: "ind => nat" + Suc :: "nat => nat" + nat_case :: "[nat, 'a, nat=>'a] =>'a" + pred_nat :: "(nat*nat) set" + nat_rec :: "[nat, 'a, [nat, 'a]=>'a] => 'a" + "0" :: "nat" ("0") + +rules + (*the axiom of infinity in 2 parts*) + inj_Suc_Rep "inj(Suc_Rep)" + Suc_Rep_not_Zero_Rep "~(Suc_Rep(x) = Zero_Rep)" + Nat_def "Nat == lfp(%X. {Zero_Rep} Un (Suc_Rep``X))" + (*faking a type definition...*) + Rep_Nat "Rep_Nat(n): Nat" + Rep_Nat_inverse "Abs_Nat(Rep_Nat(n)) = n" + Abs_Nat_inverse "i: Nat ==> Rep_Nat(Abs_Nat(i)) = i" + (*defining the abstract constants*) + 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 == (%n a f. @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)); +val monoI = result(); + +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); +val monoD = result(); + diff -r 000000000000 -r 7949f97df77a Ord.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Ord.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,15 @@ +Ord = HOL + +classes + ord < term +consts + "<", "<=" :: "['a::ord, 'a] => bool" (infixl 50) + mono :: "['a::ord => 'b::ord] => bool" (*monotonicity*) + min,max :: "['a::ord,'a] => 'a" + +rules + +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 000000000000 -r 7949f97df77a Prod.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Prod.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,204 @@ +(* 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]); +val ProdI = result(); + +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]); +val Pair_Rep_inject = result(); + +goal Prod.thy "inj_onto(Abs_Prod,Prod)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Prod_inverse 1); +val inj_onto_Abs_Prod = result(); + +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)); +val Pair_inject = result(); + +goal Prod.thy "( = ) = (a=a' & b=b')"; +by (fast_tac (set_cs addIs [Pair_inject]) 1); +val Pair_eq = result(); + +goalw Prod.thy [fst_def] "fst() = a"; +by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); +val fst_conv = result(); + +goalw Prod.thy [snd_def] "snd() = b"; +by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); +val snd_conv = result(); + +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]); +val PairE_lemma = result(); + +val [prem] = goal Prod.thy "[| !!x y. p = ==> Q |] ==> Q"; +by (rtac (PairE_lemma RS exE) 1); +by (REPEAT (eresolve_tac [prem,exE] 1)); +val PairE = result(); + +goalw Prod.thy [split_def] "split(, c) = c(a,b)"; +by (sstac [fst_conv, snd_conv] 1); +by (rtac refl 1); +val split = result(); + +(*FIXME: split's congruence rule should only simplifies the pair*) +val pair_ss = set_ss addsimps [fst_conv, snd_conv, split]; + +goal Prod.thy "p = "; +by (res_inst_tac [("p","p")] PairE 1); +by(asm_simp_tac pair_ss 1); +val surjective_pairing = result(); + +goal Prod.thy "p = split(p, %x y.)"; +by (res_inst_tac [("p","p")] PairE 1); +by(asm_simp_tac pair_ss 1); +val surjective_pairing2 = result(); + +(** split used as a logical connective, with result type bool **) + +val prems = goal Prod.thy "c(a,b) ==> split(, c)"; +by (stac split 1); +by (resolve_tac prems 1); +val splitI = result(); + +val prems = goalw Prod.thy [split_def] + "[| split(p,c); !!x y. [| p = ; c(x,y) |] ==> Q |] ==> Q"; +by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); +val splitE = result(); + +goal Prod.thy "R(split(p,c)) = (! 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); +val expand_split = result(); + +(*** prod_fun -- action of the product functor upon functions ***) + +goalw Prod.thy [prod_fun_def] "prod_fun(f,g,) = "; +by (rtac split 1); +val prod_fun = result(); + +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 (pair_ss addsimps [prod_fun,o_def]) 1); +val prod_fun_compose = result(); + +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 (pair_ss addsimps [prod_fun]) 1); +val prod_fun_ident = result(); + +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); +val prod_fun_imageI = result(); + +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); +val prod_fun_imageE = result(); + +(*** Disjoint union of a family of sets - Sigma ***) + +val SigmaI = prove_goalw 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*) +val SigmaE = prove_goalw 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 **) +val SigmaD1 = prove_goal Prod.thy " : Sigma(A,B) ==> a : A" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +val SigmaD2 = prove_goal Prod.thy " : Sigma(A,B) ==> b : B(a)" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +val SigmaE2 = prove_goal 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); +val fst_imageI = result(); + +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); +val fst_imageE = result(); + +(*** 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); +val snd_imageI = result(); + +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); +val snd_imageE = result(); + +(** 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); +val unit_eq = result(); diff -r 000000000000 -r 7949f97df77a Prod.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Prod.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,66 @@ +(* Title: HOL/prod + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ordered Pairs and the Cartesian product type +The unit type + +The type definition admits the following unused axiom: + Abs_Unit_inverse "f: Unit ==> Rep_Unit(Abs_Unit(f)) = f" +*) + +Prod = Set + +types + "*" 2 (infixr 20) + unit 0 +arities + "*" :: (term,term)term + unit :: term +consts + Pair_Rep :: "['a,'b] => ['a,'b] => bool" + Prod :: "('a => 'b => bool)set" + Rep_Prod :: "'a * 'b => ('a => 'b => bool)" + Abs_Prod :: "('a => 'b => bool) => 'a * 'b" + fst :: "'a * 'b => 'a" + snd :: "'a * 'b => 'b" + split :: "['a * 'b, ['a,'b]=>'c] => 'c" + prod_fun :: "['a=>'b, 'c=>'d, 'a*'c] => 'b*'d" + Pair :: "['a,'b] => 'a * 'b" + "@Tuple" :: "args => 'a*'b" ("(1<_>)") + Sigma :: "['a set, 'a => 'b set] => ('a*'b)set" + + Unit :: "bool set" + Rep_Unit :: "unit => bool" + Abs_Unit :: "bool => unit" + Unity :: "unit" ("<>") + +translations + + "" == ">" + "" == "Pair(x,y)" + "" => "x" + +rules + + Pair_Rep_def "Pair_Rep == (%a b. %x y. x=a & y=b)" + Prod_def "Prod == {f. ? a b. f = Pair_Rep(a,b)}" + (*faking a type definition...*) + Rep_Prod "Rep_Prod(p): Prod" + Rep_Prod_inverse "Abs_Prod(Rep_Prod(p)) = p" + Abs_Prod_inverse "f: Prod ==> Rep_Prod(Abs_Prod(f)) = f" + (*defining the abstract constants*) + 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(p,c) == c(fst(p),snd(p))" + prod_fun_def "prod_fun(f,g) == (%z.split(z, %x y.))" + Sigma_def "Sigma(A,B) == UN x:A. UN y:B(x). {}" + + Unit_def "Unit == {p. p=True}" + (*faking a type definition...*) + Rep_Unit "Rep_Unit(u): Unit" + Rep_Unit_inverse "Abs_Unit(Rep_Unit(u)) = u" + (*defining the abstract constants*) + Unity_def "Unity == Abs_Unit(True)" +end diff -r 000000000000 -r 7949f97df77a README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/README Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,21 @@ + HOL: Higher-Order Logic + +This directory contains the Standard ML sources of the Isabelle system for +Higher-Order Logic. 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 HOL and type: use "ex/ROOT.ML"; + +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 000000000000 -r 7949f97df77a ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ROOT.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,93 @@ +(* Title: HOL/ROOT + 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"; +writeln banner; + +print_depth 1; +eta_contract := true; +use_thy "hol"; +use "../Provers/hypsubst.ML"; +use "../Provers/classical.ML"; +use "../Provers/simplifier.ML"; +use "../Provers/splitter.ML"; +use "../Provers/ind.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 + + (*etac rev_cut_eq moves an equality to be the last premise. *) + val rev_cut_eq = prove_goal HOL.thy "[| a=b; a=b ==> R |] ==> R" + (fn prems => [ REPEAT(resolve_tac prems 1) ]); + + 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 swap = swap + 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]; + +val HOL_dup_cs = prop_cs addSIs [allI] addIs [exCI,ex1I] + addSEs [exE,ex1E] addEs [all_dupE]; + +structure HOL_Induction = InductionFun(struct val spec=spec end); +open HOL_Induction; + +use "simpdata.ML"; +use_thy "ord"; +use_thy "set"; +use "fun.ML"; +use "subset.ML"; +use "equalities.ML"; +use_thy "prod"; +use_thy "sum"; +use "mono.ML"; +use_thy "lfp"; +use_thy "gfp"; +use_thy "trancl"; +use_thy "wf"; +use_thy "nat"; +use_thy "arith"; +use_thy "univ"; +use_thy "sexp"; +use_thy "list"; +use_thy "llist"; + +use "../Pure/install_pp.ML"; +print_depth 8; + +val HOL_build_completed = (); (*indicate successful build*) diff -r 000000000000 -r 7949f97df77a Set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Set.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,436 @@ +(* 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); +val CollectI = result(); + +val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)"; +by (resolve_tac (prems RL [mem_Collect_eq RS subst]) 1); +val CollectD = result(); + +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); +val set_ext = result(); + +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); +val Collect_cong = result(); + +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)); +val ballI = result(); + +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); +val bspec = result(); + +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)); +val ballE = result(); + +(*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)); +val bexI = result(); + +val bexCI = prove_goal 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)); +val bexE = result(); + +(*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)); +val ball_rew = result(); + +(** 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)); +val ball_cong = result(); + +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)); +val bex_cong = result(); + +(*** Subsets ***) + +val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B"; +by (REPEAT (ares_tac (prems @ [ballI]) 1)); +val subsetI = result(); + +(*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); +val subsetD = result(); + +(*The same, with reversed premises for use with etac -- cf rev_mp*) +val rev_subsetD = prove_goal 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)); +val subsetCE = result(); + +(*Takes assumptions A<=B; c:A and creates the assumption c:B *) +fun set_mp_tac i = etac subsetCE i THEN mp_tac i; + +val subset_refl = prove_goal 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)); +val subset_trans = result(); + + +(*** 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)); +val subset_antisym = result(); +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); +val equalityD1 = result(); + +val prems = goal Set.thy "A = B ==> B<=A::'a set"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +val equalityD2 = result(); + +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)); +val equalityE = result(); + +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)); +val equalityCE = result(); + +(*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)); +val setup_induction = result(); + + +(*** Set complement -- Compl ***) + +val prems = goalw Set.thy [Compl_def] + "[| c:A ==> False |] ==> c : Compl(A)"; +by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); +val ComplI = result(); + +(*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); +val ComplD = result(); + +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)); +val UnI1 = result(); + +val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1)); +val UnI2 = result(); + +(*Classical introduction rule: no commitment to A vs B*) +val UnCI = prove_goal 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)); +val UnE = result(); + + +(*** 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)); +val IntI = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A"; +by (rtac (major RS CollectD RS conjunct1) 1); +val IntD1 = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B"; +by (rtac (major RS CollectD RS conjunct2) 1); +val IntD2 = result(); + +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); +val IntE = result(); + + +(*** Set difference ***) + +val DiffI = prove_goalw Set.thy [set_diff_def] + "[| c : A; ~ c : B |] ==> c : A - B" + (fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)) ]); + +val DiffD1 = prove_goalw Set.thy [set_diff_def] + "c : A - B ==> c : A" + (fn [major]=> [ (rtac (major RS CollectD RS conjunct1) 1) ]); + +val DiffD2 = prove_goalw 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]); + +val DiffE = prove_goal 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)) ]); + +val Diff_iff = prove_goal Set.thy "(c : A-B) = (c:A & ~ c:B)" + (fn _ => [ (fast_tac (HOL_cs addSIs [DiffI] addSEs [DiffE]) 1) ]); + + +(*** The empty set -- {} ***) + +val emptyE = prove_goalw Set.thy [empty_def] "a:{} ==> P" + (fn [prem] => [rtac (prem RS CollectD RS FalseE) 1]); + +val empty_subsetI = prove_goal Set.thy "{} <= A" + (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]); + +val equals0I = prove_goal 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)) ]); + +val equals0D = prove_goal Set.thy "[| A={}; a:A |] ==> P" + (fn [major,minor]=> + [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]); + + +(*** Augmenting a set -- insert ***) + +val insertI1 = prove_goalw Set.thy [insert_def] "a : insert(a,B)" + (fn _ => [rtac (CollectI RS UnI1) 1, rtac refl 1]); + +val insertI2 = prove_goalw Set.thy [insert_def] "a : B ==> a : insert(b,B)" + (fn [prem]=> [ (rtac (prem RS UnI2) 1) ]); + +val insertE = prove_goalw 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)) ]); + +val insert_iff = prove_goal Set.thy "a : insert(b,A) = (a=b | a:A)" + (fn _ => [fast_tac (HOL_cs addIs [insertI1,insertI2] addSEs [insertE]) 1]); + +(*Classical introduction rule*) +val insertCI = prove_goal 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 ***) + +val singletonI = prove_goal Set.thy "a : {a}" + (fn _=> [ (rtac insertI1 1) ]); + +val singletonE = prove_goal 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); +val singletonD = result(); + +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); +val singleton_inject = result(); + +(*** 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)); +val UN_I = result(); + +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)); +val UN_E = result(); + +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)); +val UN_cong = result(); + + +(*** 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)); +val INT_I = result(); + +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); +val INT_D = result(); + +(*"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)); +val INT_E = result(); + +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)); +val INT_cong = result(); + + +(*** 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)); +val UN1_I = result(); + +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)); +val UN1_E = result(); + + +(*** 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)); +val INT1_I = result(); + +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); +val INT1_D = result(); + +(*** 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)); +val UnionI = result(); + +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)); +val UnionE = result(); + +(*** Inter ***) + +val prems = goalw Set.thy [Inter_def] + "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; +by (REPEAT (ares_tac ([INT_I] @ prems) 1)); +val InterI = result(); + +(*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); +val InterD = result(); + +(*"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)); +val InterE = result(); diff -r 000000000000 -r 7949f97df77a Set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Set.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,111 @@ +(* Title: HOL/set.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge +*) + +Set = Ord + + +types + set 1 + +arities + set :: (term) term + set :: (term) ord + set :: (term) minus + + +consts + + (* Constants *) + + 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*) + range :: "('a => 'b) => 'b set" (*of function*) + Ball, Bex :: "['a set, 'a => bool] => bool" (*bounded quantifiers*) + inj, surj :: "('a => 'b) => bool" (*injective/surjective*) + inj_onto :: "['a => 'b, 'a set] => bool" + "``" :: "['a => 'b, 'a set] => ('b set)" (infixl 90) + ":" :: "['a, 'a set] => bool" (infixl 50) (*membership*) + + (* Finite Sets *) + + insert :: "['a, 'a set] => 'a set" + "{}" :: "'a set" ("{}") + "@Finset" :: "args => 'a set" ("{(_)}") + + + (** Binding Constants **) + + "@Coll" :: "[idt, bool] => 'a set" ("(1{_./ _})") (*collection*) + + (* 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. 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)" + + "{x, xs}" == "insert(x, {xs})" + "{x}" == "insert(x, {})" + + +rules + + (* Isomorphisms between Predicates and Sets *) + + mem_Collect_eq "(a : {x.P(x)}) = P(a)" + Collect_mem_eq "{x.x:A} = A" + + (* Definitions *) + + 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)" + 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 + +val print_ast_translation = + map HOL.mk_alt_ast_tr' [("@Ball", "*Ball"), ("@Bex", "*Bex")]; + diff -r 000000000000 -r 7949f97df77a Sexp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Sexp.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,166 @@ +(* Title: HOL/sexp + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For sexp.thy. S-expressions. +*) + +open Sexp; + +(** the sexp functional **) + +goal Sexp.thy "mono(%Z. range(Leaf) Un range(Numb) Un Z<*>Z)"; +by (REPEAT (ares_tac [monoI, subset_refl, Un_mono, uprod_mono] 1)); +val Sexp_fun_mono = result(); + +val Sexp_unfold = Sexp_fun_mono RS (Sexp_def RS def_lfp_Tarski); + +(** Induction **) + +val major::prems = goal Sexp.thy + "[| ii: Sexp; !!a. P(Leaf(a)); !!k. P(Numb(k)); \ +\ !!i j. [| i: Sexp; j: Sexp; P(i); P(j) |] ==> P(i.j) \ +\ |] ==> P(ii)"; +by (rtac (major RS (Sexp_def RS def_induct)) 1); +by (rtac Sexp_fun_mono 1); +by (fast_tac (set_cs addIs prems addSEs [uprodE]) 1); +val Sexp_induct = result(); + +(** Sexp_case **) + +goalw Sexp.thy [Sexp_case_def] "Sexp_case(Leaf(a),c,d,e) = c(a)"; +by (fast_tac (HOL_cs addIs [select_equality] + addSDs [Leaf_inject] + addSEs [Leaf_neq_Scons, Leaf_neq_Numb]) 1); +val Sexp_case_Leaf = result(); + +goalw Sexp.thy [Sexp_case_def] "Sexp_case(Numb(k),c,d,e) = d(k)"; +by (fast_tac (HOL_cs addIs [select_equality] + addSDs [Numb_inject] + addSEs [Numb_neq_Scons, Numb_neq_Leaf]) 1); +val Sexp_case_Numb = result(); + +goalw Sexp.thy [Sexp_case_def] "Sexp_case(M.N, c, d, e) = e(M,N)"; +by (fast_tac (HOL_cs addIs [select_equality] + addSDs [Scons_inject] + addSEs [Scons_neq_Leaf, Scons_neq_Numb]) 1); +val Sexp_case_Scons = result(); + + +(** Introduction rules for Sexp constructors **) + +val SexpI = Sexp_unfold RS equalityD2 RS subsetD; + +goal Sexp.thy "Leaf(a): Sexp"; +by (fast_tac (set_cs addIs [SexpI]) 1); +val Sexp_LeafI = result(); + +goal Sexp.thy "Numb(a): Sexp"; +by (fast_tac (set_cs addIs [SexpI]) 1); +val Sexp_NumbI = result(); + +val prems = goal Sexp.thy + "[| M: Sexp; N: Sexp |] ==> M.N : Sexp"; +by (fast_tac (set_cs addIs ([uprodI,SexpI]@prems)) 1); +val Sexp_SconsI = result(); + +val [prem] = goalw Sexp.thy [In0_def] + "M: Sexp ==> In0(M) : Sexp"; +by (rtac (prem RS (Sexp_NumbI RS Sexp_SconsI)) 1); +val Sexp_In0I = result(); + +val [prem] = goalw Sexp.thy [In1_def] + "M: Sexp ==> In1(M) : Sexp"; +by (rtac (prem RS (Sexp_NumbI RS Sexp_SconsI)) 1); +val Sexp_In1I = result(); + +goal Sexp.thy "range(Leaf) <= Sexp"; +by (fast_tac (set_cs addIs [SexpI]) 1); +val range_Leaf_subset_Sexp = result(); + +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]))); +val Scons_D = result(); + +(** Introduction rules for 'pred_Sexp' **) + +val sexp_cs = set_cs addIs [SigmaI, uprodI, SexpI]; + +goalw Sexp.thy [pred_Sexp_def] "pred_Sexp <= Sigma(Sexp, %u.Sexp)"; +by (fast_tac sexp_cs 1); +val pred_Sexp_subset_Sigma = result(); + +(* : 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); +val pred_SexpI1 = result(); + +val prems = goalw Sexp.thy [pred_Sexp_def] + "[| M: Sexp; N: Sexp |] ==> : pred_Sexp"; +by (fast_tac (set_cs addIs prems) 1); +val pred_SexpI2 = result(); + +(*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_LeafI, Sexp_NumbI, Sexp_SconsI, + 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)); +val pred_SexpE = result(); + +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); +val wf_pred_Sexp = result(); + +(*** 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); +val Sexp_rec_Leaf = result(); + +goal Sexp.thy "Sexp_rec(Numb(k), c, d, h) = d(k)"; +by (stac Sexp_rec_unfold 1); +by (rtac Sexp_case_Numb 1); +val Sexp_rec_Numb = result(); + +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); +val Sexp_rec_Scons = result(); diff -r 000000000000 -r 7949f97df77a Sexp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Sexp.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,36 @@ +(* 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 node set set" + + Sexp_case :: "['a node set, 'a=>'b, nat=>'b, \ +\ ['a node set,'a node set]=>'b] => 'b" + + Sexp_rec :: "['a node set, 'a=>'b, nat=>'b, \ +\ ['a node set,'a node set,'b,'b]=>'b] => 'b" + + pred_Sexp :: "('a node set * 'a node set)set" + +rules + Sexp_def "Sexp == lfp(%Z. range(Leaf) Un range(Numb) Un Z<*>Z)" + + Sexp_case_def + "Sexp_case(M,c,d,e) == @ 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(M, c, d, %N1 N2. e(N1, N2, g(N1), g(N2))))" +end + diff -r 000000000000 -r 7949f97df77a Subst/AList.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/AList.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,52 @@ +(* Title: Substitutions/alist.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For alist.thy. +*) + +open AList; + +(*********) + +val al_defs = [alist_rec_def,assoc_def]; + +val alist_ss = pair_ss addsimps [Nil_not_Cons,Cons_not_Nil,Cons_Cons_eq, + list_rec_Nil,list_rec_Cons]; + +val al_rews = + let fun mk_thm s = prove_goalw AList.thy al_defs s + (fn _ => [simp_tac alist_ss 1]) + in map mk_thm + ["alist_rec(Nil,c,d) = c", + "alist_rec(Cons(,al),c,d) = d(a,b,al,alist_rec(al,c,d))", + "assoc(v,d,Nil) = d", + "assoc(v,d,Cons(,al)) = if(v=a,b,assoc(v,d,al))"] end; + +(*********) + +val prems = goal AList.thy + "[| P(Nil); \ +\ !!x y xs. P(xs) ==> P(Cons(,xs)) |] ==> P(l)"; +by (list_ind_tac "l" 1); +by (resolve_tac prems 1); +by (rtac PairE 1); +by (etac ssubst 1); +by (resolve_tac prems 1); +by (assume_tac 1); +val alist_induct = result(); + +(*Perform induction on xs. *) +fun alist_ind_tac a M = + EVERY [res_inst_tac [("l",a)] alist_induct M, + rename_last_tac a ["1"] (M+1)]; + +(* +val prems = goal AList.thy + "[| P(Nil); !! x y xs. P(xs) --> P(Cons(,xs)) |] ==> P(a)"; +by (alist_ind_tac "a" 1); +by (ALLGOALS (cut_facts_tac prems THEN' fast_tac HOL_cs)); +val alist_induct2 = result(); + +add_inds alist_induct2; +*) diff -r 000000000000 -r 7949f97df77a Subst/AList.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/AList.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,22 @@ +(* Title: Substitutions/alist.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Association lists. +*) + +AList = List + + +consts + + alist_rec :: "[('a*'b) list,'c,['a,'b,('a*'b) list,'c]=>'c]=>'c" + assoc :: "['a,'b,('a*'b) list] => 'b" + +rules + + alist_rec_def + "alist_rec(al,b,c) == list_rec(al,b,%p xs g.split(p,%x y.c(x,y,xs,g)))" + + assoc_def "assoc(v,d,al) == alist_rec(al,d,%x y xs g.if(v=x,y,g))" + +end diff -r 000000000000 -r 7949f97df77a Subst/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/ROOT.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,36 @@ +(* Title: HOL/Subst + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Substitution and Unification in Higher-Order Logic. + +Implements Manna & Waldinger's formalization, with Paulson's simplifications: + +Z Manna & R Waldinger, Deductive Synthesis of the Unification Algorithm. +SCP 1 (1981), 5-48 + +L C Paulson, Verifying the Unification Algorithm in LCF. SCP 5 (1985), 143-170 + +setplus - minor additions to HOL's set theory +alist - association lists +uterm - inductive data type of terms +utermlemmas - definition of occurs and vars_of for terms +subst - substitutions +unifier - specification of unification and conditions for + correctness and termination + +To load, go to the parent directory and type use"Substitutions/ROOT.ML"; +*) + +HOL_build_completed; (*Cause examples to fail if HOL did*) + +writeln"Root file for Substitutions and Unification"; +use_thy "Subst/setplus"; + +use_thy "Subst/alist"; +use_thy "Subst/uterm"; +use_thy "Subst/utermlemmas"; + +use_thy "Subst/subst"; +use_thy "Subst/unifier"; +writeln"END: Root file for Substitutions and Unification"; diff -r 000000000000 -r 7949f97df77a Subst/Setplus.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/Setplus.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,64 @@ +(* Title: Substitutions/setplus.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For setplus.thy. +Properties of subsets and empty sets. +*) + +open Setplus; + +(*********) + +(*** Rules for subsets ***) + +goal Set.thy "A <= B = (! t.t:A --> t:B)"; +by (fast_tac set_cs 1); +val subset_iff = result(); + +goalw Setplus.thy [ssubset_def] "A < B = ((A <= B) & ~(A=B))"; +by (rtac refl 1); +val ssubset_iff = result(); + +goal Setplus.thy "(A::'a set <= B) = ((A < B) | (A=B))"; +by (simp_tac (set_ss addsimps [ssubset_iff]) 1); +by (fast_tac set_cs 1); +val subseteq_iff_subset_eq = result(); + +(*Rule in Modus Ponens style*) +goal Setplus.thy "A < B --> c:A --> c:B"; +by (simp_tac (set_ss addsimps [ssubset_iff]) 1); +by (fast_tac set_cs 1); +val ssubsetD = result(); + +(*********) + +goalw Setplus.thy [empty_def] "~ a : {}"; +by (fast_tac set_cs 1); +val not_in_empty = result(); + +goalw Setplus.thy [empty_def] "(A = {}) = (ALL a.~ a:A)"; +by (fast_tac (set_cs addIs [set_ext]) 1); +val empty_iff = result(); + + +(*********) + +goal Set.thy "(~A=B) = ((? x.x:A & ~x:B) | (? x.~x:A & x:B))"; +by (fast_tac (set_cs addIs [set_ext]) 1); +val not_equal_iff = result(); + +(*********) + +val setplus_rews = [ssubset_iff,not_in_empty,empty_iff]; + +(*********) + +(*Case analysis for rewriting; P also gets rewritten*) +val [prem1,prem2] = goal HOL.thy "[| P-->Q; ~P-->Q |] ==> Q"; +by (rtac (excluded_middle RS disjE) 1); +by (etac (prem2 RS mp) 1); +by (etac (prem1 RS mp) 1); +val imp_excluded_middle = result(); + +fun imp_excluded_middle_tac s = res_inst_tac [("P",s)] imp_excluded_middle; diff -r 000000000000 -r 7949f97df77a Subst/Setplus.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/Setplus.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: Substitutions/setplus.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Minor additions to HOL's set theory +*) + +Setplus = Set + + +rules + + ssubset_def "A < B == A <= B & ~ A=B" + +end diff -r 000000000000 -r 7949f97df77a Subst/Subst.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/Subst.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,185 @@ +(* Title: Substitutions/subst.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For subst.thy. +*) + +open Subst; + +(***********) + +val subst_defs = [subst_def,comp_def,sdom_def]; + +val raw_subst_ss = utlemmas_ss addsimps al_rews; + +local fun mk_thm s = prove_goalw Subst.thy subst_defs s + (fn _ => [simp_tac raw_subst_ss 1]) +in val subst_rews = map mk_thm +["Const(c) <| al = Const(c)", + "Comb(t,u) <| al = Comb(t <| al, u <| al)", + "Nil <> bl = bl", + "Cons(,al) <> bl = Cons(, al <> bl)", + "sdom(Nil) = {}", + "sdom(Cons(,al)) = if(Var(a)=b,sdom(al) Int Compl({a}),sdom(al) Un {a})" +]; + (* This rewrite isn't always desired *) + val Var_subst = mk_thm "Var(x) <| al = assoc(x,Var(x),al)"; +end; + +val subst_ss = raw_subst_ss addsimps subst_rews; + +(**** Substitutions ****) + +goal Subst.thy "t <| Nil = t"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); +val subst_Nil = result(); + +goal Subst.thy "t <: u --> t <| s <: u <| s"; +by (uterm_ind_tac "u" 1); +by (ALLGOALS (asm_simp_tac subst_ss)); +val subst_mono = result() RS mp; + +goal Subst.thy "~ (Var(v) <: t) --> t <| Cons(,s) = t <| s"; +by (imp_excluded_middle_tac "t = Var(v)" 1); +by (res_inst_tac [("P", + "%x.~x=Var(v) --> ~(Var(v) <: x) --> x <| Cons(,s)=x<|s")] + uterm_induct 2); +by (ALLGOALS (simp_tac (subst_ss addsimps [Var_subst]))); +by (fast_tac HOL_cs 1); +val Var_not_occs = result() RS mp; + +goal Subst.thy + "(t <|r = t <|s) = (! v.v : vars_of(t) --> Var(v) <|r = Var(v) <|s)"; +by (uterm_ind_tac "t" 1); +by (REPEAT (etac rev_mp 3)); +by (ALLGOALS (asm_simp_tac subst_ss)); +by (ALLGOALS (fast_tac HOL_cs)); +val agreement = result(); + +goal Subst.thy "~ v: vars_of(t) --> t <| Cons(,s) = t <| s"; +by(simp_tac(subst_ss addsimps [agreement,Var_subst] + setloop (split_tac [expand_if])) 1); +val repl_invariance = result() RS mp; + +val asms = goal Subst.thy + "v : vars_of(t) --> w : vars_of(t <| Cons(,s))"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); +val Var_in_subst = result() RS mp; + +(**** Equality between Substitutions ****) + +goalw Subst.thy [subst_eq_def] "r =s= s = (! t.t <| r = t <| s)"; +by (simp_tac subst_ss 1); +val subst_eq_iff = result(); + +local fun mk_thm s = prove_goal Subst.thy s + (fn prems => [cut_facts_tac prems 1, + REPEAT (etac rev_mp 1), + simp_tac (subst_ss addsimps [subst_eq_iff]) 1]) +in + val subst_refl = mk_thm "r = s ==> r =s= s"; + val subst_sym = mk_thm "r =s= s ==> s =s= r"; + val subst_trans = mk_thm "[| q =s= r; r =s= s |] ==> q =s= s"; +end; + +val eq::prems = goalw Subst.thy [subst_eq_def] + "[| r =s= s; P(t <| r,u <| r) |] ==> P(t <| s,u <| s)"; +by (resolve_tac [eq RS spec RS subst] 1); +by (resolve_tac (prems RL [eq RS spec RS subst]) 1); +val subst_subst2 = result(); + +val ssubst_subst2 = subst_sym RS subst_subst2; + +(**** Composition of Substitutions ****) + +goal Subst.thy "s <> Nil = s"; +by (alist_ind_tac "s" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [subst_Nil]))); +val comp_Nil = result(); + +goal Subst.thy "(t <| r <> s) = (t <| r <| s)"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); +by (alist_ind_tac "r" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst,subst_Nil] + setloop (split_tac [expand_if])))); +val subst_comp = result(); + +goal Subst.thy "q <> r <> s =s= q <> (r <> s)"; +by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); +val comp_assoc = result(); + +goal Subst.thy "Cons(,s) =s= s"; +by (rtac (allI RS (subst_eq_iff RS iffD2)) 1); +by (uterm_ind_tac "t" 1); +by (REPEAT (etac rev_mp 3)); +by (ALLGOALS (simp_tac (subst_ss addsimps[Var_subst] + setloop (split_tac [expand_if])))); +val Cons_trivial = result(); + +val [prem] = goal Subst.thy "q <> r =s= s ==> t <| q <| r = t <| s"; +by (simp_tac (subst_ss addsimps [prem RS (subst_eq_iff RS iffD1), + subst_comp RS sym]) 1); +val comp_subst_subst = result(); + +(**** Domain and range of Substitutions ****) + +goal Subst.thy "(v : sdom(s)) = (~ Var(v) <| s = Var(v))"; +by (alist_ind_tac "s" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst] + setloop (split_tac[expand_if])))); +by (fast_tac HOL_cs 1); +val sdom_iff = result(); + +goalw Subst.thy [srange_def] + "v : srange(s) = (? w.w : sdom(s) & v : vars_of(Var(w) <| s))"; +by (fast_tac set_cs 1); +val srange_iff = result(); + +goal Subst.thy "(t <| s = t) = (sdom(s) Int vars_of(t) = {})"; +by (uterm_ind_tac "t" 1); +by (REPEAT (etac rev_mp 3)); +by (ALLGOALS (simp_tac (subst_ss addsimps [sdom_iff,Var_subst]))); +by (ALLGOALS (fast_tac set_cs)); +val invariance = result(); + +goal Subst.thy "v : sdom(s) --> ~v : srange(s) --> ~v : vars_of(t <| s)"; +by (uterm_ind_tac "t" 1); +by (imp_excluded_middle_tac "x : sdom(s)" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [sdom_iff,srange_iff]))); +by (ALLGOALS (fast_tac set_cs)); +val Var_elim = result() RS mp RS mp; + +val asms = goal Subst.thy + "[| v : sdom(s); v : vars_of(t <| s) |] ==> v : srange(s)"; +by (REPEAT (ares_tac (asms @ [Var_elim RS swap RS classical]) 1)); +val Var_elim2 = result(); + +goal Subst.thy "v : vars_of(t <| s) --> v : srange(s) | v : vars_of(t)"; +by (uterm_ind_tac "t" 1); +by (REPEAT_SOME (etac rev_mp )); +by (ALLGOALS (simp_tac (subst_ss addsimps [sdom_iff,srange_iff]))); +by (REPEAT (step_tac (set_cs addIs [vars_var_iff RS iffD1 RS sym]) 1)); +by (etac notE 1); +by (etac subst 1); +by (ALLGOALS (fast_tac set_cs)); +val Var_intro = result() RS mp; + +goal Subst.thy + "v : srange(s) --> (? w.w : sdom(s) & v : vars_of(Var(w) <| s))"; +by (simp_tac (subst_ss addsimps [srange_iff]) 1); +val srangeE = make_elim (result() RS mp); + +val asms = goal Subst.thy + "sdom(s) Int srange(s) = {} = (! t.sdom(s) Int vars_of(t <| s) = {})"; +by (simp_tac subst_ss 1); +by (fast_tac (set_cs addIs [Var_elim2] addEs [srangeE]) 1); +val dom_range_disjoint = result(); + +val asms = goal Subst.thy "~ u <| s = u --> (? x.x : sdom(s))"; +by (simp_tac (subst_ss addsimps [invariance]) 1); +by (fast_tac set_cs 1); +val subst_not_empty = result() RS mp; diff -r 000000000000 -r 7949f97df77a Subst/Subst.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/Subst.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,37 @@ +(* Title: Substitutions/subst.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Substitutions on uterms +*) + +Subst = Setplus + AList + UTLemmas + + +consts + + "=s=" :: "[('a*('a uterm)) list,('a*('a uterm)) list] => bool" (infixr 52) + + "<|" :: "['a uterm,('a*('a uterm)) list] => 'a uterm" (infixl 55) + "<>" :: "[('a*('a uterm)) list, ('a*('a uterm)) list] => \ +\ ('a*('a uterm)) list" (infixl 56) + sdom :: "('a*('a uterm)) list => 'a set" + srange :: "('a*('a uterm)) list => 'a set" + +rules + + subst_eq_def "r =s= s == ALL t.t <| r = t <| s" + + subst_def + "t <| al == uterm_rec(t, %x.assoc(x,Var(x),al), \ +\ %x.Const(x), \ +\ %u v q r.Comb(q,r))" + + comp_def "al <> bl == alist_rec(al,bl,%x y xs g.Cons(,g))" + + sdom_def + "sdom(al) == alist_rec(al, {}, \ +\ %x y xs g.if(Var(x)=y, g Int Compl({x}), g Un {x}))" + srange_def + "srange(al) == Union({y. EX x:sdom(al).y=vars_of(Var(x) <| al)})" + +end diff -r 000000000000 -r 7949f97df77a Subst/UTLemmas.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/UTLemmas.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,77 @@ +(* Title: Substitutions/utermlemmas.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For utermlemmas.thy. +*) + +open UTLemmas; + +(***********) + +val utlemmas_defs = [vars_of_def,occs_def]; + +local fun mk_thm s = prove_goalw UTLemmas.thy utlemmas_defs s + (fn _ => [simp_tac uterm_ss 1]) +in val utlemmas_rews = map mk_thm + ["vars_of(Const(c)) = {}", + "vars_of(Var(x)) = {x}", + "vars_of(Comb(t,u)) = vars_of(t) Un vars_of(u)", + "t <: Const(c) = False", + "t <: Var(x) = False", + "t <: Comb(u,v) = (t=u | t=v | t <: u | t <: v)"]; +end; + +val utlemmas_ss = pair_ss addsimps (setplus_rews @ uterm_rews @ utlemmas_rews); + +(**** occs irrefl ****) + +goal UTLemmas.thy "t <: u & u <: v --> t <: v"; +by (uterm_ind_tac "v" 1); +by (ALLGOALS (simp_tac utlemmas_ss)); +by (fast_tac HOL_cs 1); +val occs_trans = conjI RS (result() RS mp); + +goal UTLemmas.thy "~ t <: v --> ~ t <: u | ~ u <: v"; +by (fast_tac (HOL_cs addIs [occs_trans]) 1); +val contr_occs_trans = result() RS mp; + +goal UTLemmas.thy "t <: Comb(t,u)"; +by (simp_tac utlemmas_ss 1); +val occs_Comb1 = result(); + +goal UTLemmas.thy "u <: Comb(t,u)"; +by (simp_tac utlemmas_ss 1); +val occs_Comb2 = result(); + +goal HOL.thy "(~(P|Q)) = (~P & ~Q)"; +by (fast_tac HOL_cs 1); +val demorgan_disj = result(); + +goal UTLemmas.thy "~ t <: t"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (simp_tac (utlemmas_ss addsimps [demorgan_disj]))); +by (REPEAT (resolve_tac [impI,conjI] 1 ORELSE + (etac contrapos 1 THEN etac subst 1 THEN + resolve_tac [occs_Comb1,occs_Comb2] 1) ORELSE + (etac (contr_occs_trans RS disjE) 1 THEN assume_tac 2) ORELSE + eresolve_tac ([occs_Comb1,occs_Comb2] RLN(2,[notE])) 1)); +val occs_irrefl = result(); + +goal UTLemmas.thy "t <: u --> ~t=u"; +by (fast_tac (HOL_cs addEs [occs_irrefl RS notE]) 1); +val occs_irrefl2 = result() RS mp; + + +(**** vars_of lemmas ****) + +goal UTLemmas.thy "(v : vars_of(Var(w))) = (w=v)"; +by (simp_tac utlemmas_ss 1); +by (fast_tac HOL_cs 1); +val vars_var_iff = result(); + +goal UTLemmas.thy "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (simp_tac utlemmas_ss)); +by (fast_tac HOL_cs 1); +val vars_iff_occseq = result(); diff -r 000000000000 -r 7949f97df77a Subst/UTLemmas.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/UTLemmas.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,20 @@ +(* Title: Substitutions/utermlemmas.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Additional definitions for uterms that are not part of the basic inductive definition. +*) + +UTLemmas = UTerm + Setplus + + +consts + + vars_of :: "'a uterm=>'a set" + "<:" :: "['a uterm,'a uterm]=>bool" (infixl 54) + +rules (*Definitions*) + + vars_of_def "vars_of(t) == uterm_rec(t,%x.{x},%x.{},%u v q r.q Un r)" + occs_def "s <: t == uterm_rec(t,%x.False,%x.False,%u v q r.s=u | s=v | q | r)" + +end diff -r 000000000000 -r 7949f97df77a Subst/UTerm.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/UTerm.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,320 @@ +(* Title: Substitutions/uterm.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For uterm.thy. +*) + +open UTerm; + +(** the uterm functional **) + +goal UTerm.thy "mono(%Z. A <+> A <+> Z <*> Z)"; +by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono] 1)); +val UTerm_fun_mono = result(); + +val UTerm_unfold = UTerm_fun_mono RS (UTerm_def RS def_lfp_Tarski); + +(*This justifies using UTerm in other recursive type definitions*) +val prems = goalw UTerm.thy [UTerm_def] "[| A<=B |] ==> UTerm(A) <= UTerm(B)"; +by (REPEAT (ares_tac (prems@[monoI, subset_refl, lfp_mono, + usum_mono, uprod_mono]) 1)); +val UTerm_mono = result(); + +(** Type checking rules -- UTerm creates well-founded sets **) + +val prems = goalw UTerm.thy [UTerm_def] "UTerm(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_In0I,Sexp_In1I,Sexp_SconsI]) 1); +val UTerm_Sexp = result(); + +(* A <= Sexp ==> UTerm(A) <= Sexp *) +val UTerm_subset_Sexp = standard + (UTerm_mono RS (UTerm_Sexp RSN (2,subset_trans))); + +(** Induction **) + +(*Induction for the set UTerm(A) *) +val major::prems = goalw UTerm.thy [VAR_def,CONST_def,COMB_def] + "[| M: UTerm(A); !!M.M : A ==> P(VAR(M)); !!M.M : A ==> P(CONST(M)); \ +\ !!M N. [| M: UTerm(A); N: UTerm(A); P(M); P(N) |] ==> P(COMB(M,N)) |] \ +\ ==> P(M)"; +by (rtac (major RS (UTerm_def RS def_induct)) 1); +by (rtac UTerm_fun_mono 1); +by (fast_tac (set_cs addIs prems addEs [usumE,uprodE]) 1); +val UTerm_induct = result(); + +(*Induction for the type 'a uterm *) +val prems = goalw UTerm.thy [Var_def,Const_def,Comb_def] + "[| !!x.P(Var(x)); !!x.P(Const(x)); \ +\ !!u v. [| P(u); P(v) |] ==> P(Comb(u,v)) |] ==> P(t)"; +by (rtac (Rep_UTerm_inverse RS subst) 1); (*types force good instantiation*) +by (rtac (Rep_UTerm RS UTerm_induct) 1); +by (REPEAT (ares_tac prems 1 + ORELSE eresolve_tac [rangeE, ssubst, Abs_UTerm_inverse RS subst] 1)); +val uterm_induct = result(); + +(*Perform induction on xs. *) +fun uterm_ind_tac a M = + EVERY [res_inst_tac [("t",a)] uterm_induct M, + rename_last_tac a ["1"] (M+1)]; + +(** Introduction rules for UTerm constructors **) + +(* c : A <+> A <+> UTerm(A) <*> UTerm(A) ==> c : UTerm(A) *) +val UTermI = UTerm_unfold RS equalityD2 RS subsetD; + +(* Nil is a UTerm -- this also justifies the type definition*) +val prems = goalw UTerm.thy [VAR_def] "v:A ==> VAR(v) : UTerm(A)"; +by (REPEAT (resolve_tac ([singletonI, UTermI, usum_In0I]@prems) 1)); +val VAR_I = result(); + +val prems = goalw UTerm.thy [CONST_def] "c:A ==> CONST(c) : UTerm(A)"; +by (REPEAT (resolve_tac ([singletonI, UTermI, usum_In0I, usum_In1I]@prems) 1)); +val CONST_I = result(); + +val prems = goalw UTerm.thy [COMB_def] + "[| M:UTerm(A); N:UTerm(A) |] ==> COMB(M,N) : UTerm(A)"; +by (REPEAT (resolve_tac (prems@[UTermI, uprodI, usum_In1I]) 1)); +val COMB_I = result(); + +(*** Isomorphisms ***) + +goal UTerm.thy "inj(Rep_UTerm)"; +by (rtac inj_inverseI 1); +by (rtac Rep_UTerm_inverse 1); +val inj_Rep_UTerm = result(); + +goal UTerm.thy "inj_onto(Abs_UTerm,UTerm(range(Leaf)))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_UTerm_inverse 1); +val inj_onto_Abs_UTerm = result(); + +(** Distinctness of constructors **) + +goalw UTerm.thy [CONST_def,COMB_def] "~ CONST(c) = COMB(u,v)"; +by (rtac notI 1); +by (etac (In1_inject RS (In0_not_In1 RS notE)) 1); +val CONST_not_COMB = result(); +val COMB_not_CONST = standard (CONST_not_COMB RS not_sym); +val CONST_neq_COMB = standard (CONST_not_COMB RS notE); +val COMB_neq_CONST = sym RS CONST_neq_COMB; + +goalw UTerm.thy [COMB_def,VAR_def] "~ COMB(u,v) = VAR(x)"; +by (rtac In1_not_In0 1); +val COMB_not_VAR = result(); +val VAR_not_COMB = standard (COMB_not_VAR RS not_sym); +val COMB_neq_VAR = standard (COMB_not_VAR RS notE); +val VAR_neq_COMB = sym RS COMB_neq_VAR; + +goalw UTerm.thy [VAR_def,CONST_def] "~ VAR(x) = CONST(c)"; +by (rtac In0_not_In1 1); +val VAR_not_CONST = result(); +val CONST_not_VAR = standard (VAR_not_CONST RS not_sym); +val VAR_neq_CONST = standard (VAR_not_CONST RS notE); +val CONST_neq_VAR = sym RS VAR_neq_CONST; + + +goalw UTerm.thy [Const_def,Comb_def] "~ Const(c) = Comb(u,v)"; +by (rtac (CONST_not_COMB RS (inj_onto_Abs_UTerm RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, VAR_I, CONST_I, COMB_I, Rep_UTerm] 1)); +val Const_not_Comb = result(); +val Comb_not_Const = standard (Const_not_Comb RS not_sym); +val Const_neq_Comb = standard (Const_not_Comb RS notE); +val Comb_neq_Const = sym RS Const_neq_Comb; + +goalw UTerm.thy [Comb_def,Var_def] "~ Comb(u,v) = Var(x)"; +by (rtac (COMB_not_VAR RS (inj_onto_Abs_UTerm RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, VAR_I, CONST_I, COMB_I, Rep_UTerm] 1)); +val Comb_not_Var = result(); +val Var_not_Comb = standard (Comb_not_Var RS not_sym); +val Comb_neq_Var = standard (Comb_not_Var RS notE); +val Var_neq_Comb = sym RS Comb_neq_Var; + +goalw UTerm.thy [Var_def,Const_def] "~ Var(x) = Const(c)"; +by (rtac (VAR_not_CONST RS (inj_onto_Abs_UTerm RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, VAR_I, CONST_I, COMB_I, Rep_UTerm] 1)); +val Var_not_Const = result(); +val Const_not_Var = standard (Var_not_Const RS not_sym); +val Var_neq_Const = standard (Var_not_Const RS notE); +val Const_neq_Var = sym RS Var_neq_Const; + + +(** Injectiveness of CONST and Const **) + +val inject_cs = HOL_cs addSEs [Scons_inject] + addSDs [In0_inject,In1_inject]; + +goalw UTerm.thy [VAR_def] "(VAR(M)=VAR(N)) = (M=N)"; +by (fast_tac inject_cs 1); +val VAR_VAR_eq = result(); + +goalw UTerm.thy [CONST_def] "(CONST(M)=CONST(N)) = (M=N)"; +by (fast_tac inject_cs 1); +val CONST_CONST_eq = result(); + +goalw UTerm.thy [COMB_def] "(COMB(K,L)=COMB(M,N)) = (K=M & L=N)"; +by (fast_tac inject_cs 1); +val COMB_COMB_eq = result(); + +val VAR_inject = standard (VAR_VAR_eq RS iffD1); +val CONST_inject = standard (CONST_CONST_eq RS iffD1); +val COMB_inject = standard (COMB_COMB_eq RS iffD1 RS conjE); + + +(*For reasoning about abstract uterm constructors*) +val UTerm_cs = set_cs addIs [Rep_UTerm, VAR_I, CONST_I, COMB_I] + addSEs [CONST_neq_COMB,COMB_neq_VAR,VAR_neq_CONST, + COMB_neq_CONST,VAR_neq_COMB,CONST_neq_VAR, + COMB_inject] + addSDs [VAR_inject,CONST_inject, + inj_onto_Abs_UTerm RS inj_ontoD, + inj_Rep_UTerm RS injD, Leaf_inject]; + +goalw UTerm.thy [Var_def] "(Var(x)=Var(y)) = (x=y)"; +by (fast_tac UTerm_cs 1); +val Var_Var_eq = result(); +val Var_inject = standard (Var_Var_eq RS iffD1); + +goalw UTerm.thy [Const_def] "(Const(x)=Const(y)) = (x=y)"; +by (fast_tac UTerm_cs 1); +val Const_Const_eq = result(); +val Const_inject = standard (Const_Const_eq RS iffD1); + +goalw UTerm.thy [Comb_def] "(Comb(u,v)=Comb(x,y)) = (u=x & v=y)"; +by (fast_tac UTerm_cs 1); +val Comb_Comb_eq = result(); +val Comb_inject = standard (Comb_Comb_eq RS iffD1 RS conjE); + +val [major] = goal UTerm.thy "VAR(M): UTerm(A) ==> M : A"; +by (rtac (major RS setup_induction) 1); +by (etac UTerm_induct 1); +by (ALLGOALS (fast_tac UTerm_cs)); +val VAR_D = result(); + +val [major] = goal UTerm.thy "CONST(M): UTerm(A) ==> M : A"; +by (rtac (major RS setup_induction) 1); +by (etac UTerm_induct 1); +by (ALLGOALS (fast_tac UTerm_cs)); +val CONST_D = result(); + +val [major] = goal UTerm.thy + "COMB(M,N): UTerm(A) ==> M: UTerm(A) & N: UTerm(A)"; +by (rtac (major RS setup_induction) 1); +by (etac UTerm_induct 1); +by (ALLGOALS (fast_tac UTerm_cs)); +val COMB_D = result(); + +(*Basic ss with constructors and their freeness*) +val uterm_free_simps = [Const_not_Comb,Comb_not_Var,Var_not_Const, + Comb_not_Const,Var_not_Comb,Const_not_Var, + Var_Var_eq,Const_Const_eq,Comb_Comb_eq, + CONST_not_COMB,COMB_not_VAR,VAR_not_CONST, + COMB_not_CONST,VAR_not_COMB,CONST_not_VAR, + VAR_VAR_eq,CONST_CONST_eq,COMB_COMB_eq, + VAR_I, CONST_I, COMB_I]; +val uterm_free_ss = HOL_ss addsimps uterm_free_simps; + +goal UTerm.thy "!u. ~(t=Comb(t,u))"; +by (uterm_ind_tac "t" 1); +by (rtac (Var_not_Comb RS allI) 1); +by (rtac (Const_not_Comb RS allI) 1); +by (asm_simp_tac uterm_free_ss 1); +val t_not_Comb_t = result(); + +goal UTerm.thy "!t. ~(u=Comb(t,u))"; +by (uterm_ind_tac "u" 1); +by (rtac (Var_not_Comb RS allI) 1); +by (rtac (Const_not_Comb RS allI) 1); +by (asm_simp_tac uterm_free_ss 1); +val u_not_Comb_u = result(); + + +(*** UTerm_rec -- by wf recursion on pred_Sexp ***) + +val UTerm_rec_unfold = + wf_pred_Sexp RS wf_trancl RS (UTerm_rec_def RS def_wfrec); + +(** conversion rules **) + +goalw UTerm.thy [VAR_def] "UTerm_rec(VAR(x),b,c,d) = b(x)"; +by (rtac (UTerm_rec_unfold RS trans) 1); +by (rtac Case_In0 1); +val UTerm_rec_VAR = result(); + +goalw UTerm.thy [CONST_def] "UTerm_rec(CONST(x),b,c,d) = c(x)"; +by (rtac (UTerm_rec_unfold RS trans) 1); +by (simp_tac (HOL_ss addsimps [Case_In0,Case_In1]) 1); +val UTerm_rec_CONST = result(); + +val prems = goalw UTerm.thy [COMB_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ UTerm_rec(COMB(M,N), b, c, d) = \ +\ d(M, N, UTerm_rec(M,b,c,d), UTerm_rec(N,b,c,d))"; +by (rtac (UTerm_rec_unfold RS trans) 1); +by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1); +by (simp_tac (pred_Sexp_ss addsimps (In1_def::prems)) 1); +val UTerm_rec_COMB = result(); + +(*** uterm_rec -- by UTerm_rec ***) + +val Rep_UTerm_in_Sexp = + Rep_UTerm RS (range_Leaf_subset_Sexp RS UTerm_subset_Sexp RS subsetD); + +val uterm_rec_simps = [UTerm_rec_VAR, UTerm_rec_CONST, UTerm_rec_COMB, + Abs_UTerm_inverse, Rep_UTerm_inverse, VAR_I, CONST_I, COMB_I, + Rep_UTerm, rangeI, inj_Leaf, Inv_f_f, Rep_UTerm_in_Sexp]; +val uterm_rec_ss = HOL_ss addsimps uterm_rec_simps; + +goalw UTerm.thy [uterm_rec_def, Var_def] "uterm_rec(Var(x),b,c,d) = b(x)"; +by (simp_tac uterm_rec_ss 1); +val uterm_rec_Var = result(); + +goalw UTerm.thy [uterm_rec_def, Const_def] "uterm_rec(Const(x),b,c,d) = c(x)"; +by (simp_tac uterm_rec_ss 1); +val uterm_rec_Const = result(); + +goalw UTerm.thy [uterm_rec_def, Comb_def] + "uterm_rec(Comb(u,v),b,c,d) = d(u,v,uterm_rec(u,b,c,d),uterm_rec(v,b,c,d))"; +by (simp_tac uterm_rec_ss 1); +val uterm_rec_Comb = result(); + +val uterm_simps = [UTerm_rec_VAR, UTerm_rec_CONST, UTerm_rec_COMB, + uterm_rec_Var, uterm_rec_Const, uterm_rec_Comb]; +val uterm_ss = uterm_free_ss addsimps uterm_simps; + +(*Type checking. Useful?*) +val major::A_subset_Sexp::prems = goal UTerm.thy + "[| M: UTerm(A); \ +\ A<=Sexp; \ +\ !!x.x:A ==> b(x): C(VAR(x)); \ +\ !!x.x:A ==> c(x): C(CONST(x)); \ +\ !!x y q r. [| x: UTerm(A); y: UTerm(A); q: C(x); r: C(y) |] ==> \ +\ d(x,y,q,r): C(COMB(x,y)) \ +\ |] ==> UTerm_rec(M,b,c,d) : C(M)"; +val Sexp_UTermA_I = A_subset_Sexp RS UTerm_subset_Sexp RS subsetD; +val Sexp_A_I = A_subset_Sexp RS subsetD; +by (rtac (major RS UTerm_induct) 1); +by (ALLGOALS + (asm_simp_tac (uterm_ss addsimps ([Sexp_A_I,Sexp_UTermA_I] @ prems)))); +val UTerm_rec_type = result(); + + +(**********) + +val uterm_rews = [uterm_rec_Var,uterm_rec_Const,uterm_rec_Comb, + t_not_Comb_t,u_not_Comb_u, + Const_not_Comb,Comb_not_Var,Var_not_Const, + Comb_not_Const,Var_not_Comb,Const_not_Var, + Var_Var_eq,Const_Const_eq,Comb_Comb_eq]; + +(* +val prems = goal Subst.thy + "[| !!x.P(Var(x)); !!x.P(Const(x)); \ +\ !!u v. P(u) --> P(v) --> P(Comb(u,v)) |] ==> P(a)"; +by (uterm_ind_tac "a" 1); +by (ALLGOALS (cut_facts_tac prems THEN' fast_tac HOL_cs)); +val uterm_induct2 = result(); + +add_inds uterm_induct2; +*) diff -r 000000000000 -r 7949f97df77a Subst/UTerm.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/UTerm.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,58 @@ +(* Title: Substitutions/uterm.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Simple term structure for unifiation. +Binary trees with leaves that are constants or variables. +*) + +UTerm = Sexp + + +types uterm 1 + +arities + uterm :: (term)term + +consts + UTerm :: "'a node set set => 'a node set set" + Rep_UTerm :: "'a uterm => 'a node set" + Abs_UTerm :: "'a node set => 'a uterm" + VAR :: "'a node set => 'a node set" + CONST :: "'a node set => 'a node set" + COMB :: "['a node set, 'a node set] => 'a node set" + Var :: "'a => 'a uterm" + Const :: "'a => 'a uterm" + Comb :: "['a uterm, 'a uterm] => 'a uterm" + UTerm_rec :: "['a node set, 'a node set => 'b, 'a node set => 'b, \ +\ ['a node set , 'a node set, 'b, 'b]=>'b] => 'b" + uterm_rec :: "['a uterm, 'a => 'b, 'a => 'b, \ +\ ['a uterm, 'a uterm,'b,'b]=>'b] => 'b" + +rules + UTerm_def "UTerm(A) == lfp(%Z. A <+> A <+> Z <*> Z)" + (*faking a type definition...*) + Rep_UTerm "Rep_UTerm(xs): UTerm(range(Leaf))" + Rep_UTerm_inverse "Abs_UTerm(Rep_UTerm(xs)) = xs" + Abs_UTerm_inverse "M: UTerm(range(Leaf)) ==> Rep_UTerm(Abs_UTerm(M)) = M" + (*defining the concrete constructors*) + VAR_def "VAR(v) == In0(v)" + CONST_def "CONST(v) == In1(In0(v))" + COMB_def "COMB(t,u) == In1(In1(t . u))" + (*defining the abstract constructors*) + Var_def "Var(v) == Abs_UTerm(VAR(Leaf(v)))" + Const_def "Const(c) == Abs_UTerm(CONST(Leaf(c)))" + Comb_def "Comb(t,u) == Abs_UTerm(COMB(Rep_UTerm(t),Rep_UTerm(u)))" + + (*uterm recursion*) + UTerm_rec_def + "UTerm_rec(M,b,c,d) == wfrec(trancl(pred_Sexp), M, \ +\ %z g. Case(z, %x. b(x), \ +\ %w. Case(w, %x. c(x), \ +\ %v. Split(v, %x y. d(x,y,g(x),g(y))))))" + + uterm_rec_def + "uterm_rec(t,b,c,d) == \ +\ UTerm_rec(Rep_UTerm(t), %x.b(Inv(Leaf,x)), %x.c(Inv(Leaf,x)), \ +\ %x y q r.d(Abs_UTerm(x),Abs_UTerm(y),q,r))" + +end diff -r 000000000000 -r 7949f97df77a Subst/Unifier.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/Unifier.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,299 @@ +(* Title: Substitutions/unifier.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For unifier.thy. +Properties of unifiers. +Cases for partial correctness of algorithm and conditions for termination. +*) + +open Unifier; + +val unify_defs = + [Idem_def,Unifier_def,MoreGeneral_def,MGUnifier_def,MGIUnifier_def]; + +(**** Unifiers ****) + +goalw Unifier.thy [Unifier_def] "Unifier(s,t,u) = (t <| s = u <| s)"; +by (rtac refl 1); +val Unifier_iff = result(); + +goal Unifier.thy + "Unifier(s,Comb(t,u),Comb(v,w)) --> Unifier(s,t,v) & Unifier(s,u,w)"; +by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); +val Unifier_Comb = result() RS mp RS conjE; + +goal Unifier.thy + "~v : vars_of(t) --> ~v : vars_of(u) -->Unifier(s,t,u) --> \ +\ Unifier(Cons(,s),t,u)"; +by (simp_tac (subst_ss addsimps [Unifier_iff,repl_invariance]) 1); +val Cons_Unifier = result() RS mp RS mp RS mp; + +(**** Most General Unifiers ****) + +goalw Unifier.thy [MoreGeneral_def] "r >> s = (EX q. s =s= r <> q)"; +by (rtac refl 1); +val MoreGen_iff = result(); + +goal Unifier.thy "Nil >> s"; +by (simp_tac (subst_ss addsimps [MoreGen_iff]) 1); +by (fast_tac (set_cs addIs [refl RS subst_refl]) 1); +val MoreGen_Nil = result(); + +goalw Unifier.thy unify_defs + "MGUnifier(s,t,u) = (ALL r.Unifier(r,t,u) = s >> r)"; +by (REPEAT (ares_tac [iffI,allI] 1 ORELSE + eresolve_tac [conjE,allE,mp,exE,ssubst_subst2] 1)); +by (asm_simp_tac (subst_ss addsimps [subst_comp]) 1); +by (fast_tac (set_cs addIs [comp_Nil RS sym RS subst_refl]) 1); +val MGU_iff = result(); + +val [prem] = goal Unifier.thy + "~ Var(v) <: t ==> MGUnifier(Cons(,Nil),Var(v),t)"; +by (simp_tac (subst_ss addsimps [MGU_iff,MoreGen_iff,Unifier_iff]) 1); +by (REPEAT_SOME (step_tac set_cs)); +by (etac subst 1); +by (etac ssubst_subst2 2); +by (rtac (Cons_trivial RS subst_sym) 1); +by (simp_tac (subst_ss addsimps [prem RS Var_not_occs,Var_subst]) 1); +val MGUnifier_Var = result(); + +(**** Most General Idempotent Unifiers ****) + +goal Unifier.thy "r <> r =s= r --> s =s= r <> q --> r <> s =s= s"; +by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); +val MGIU_iff_lemma = result() RS mp RS mp; + +goalw Unifier.thy unify_defs + "MGIUnifier(s,t,u) = \ +\ (Idem(s) & Unifier(s,t,u) & (ALL r.Unifier(r,t,u) --> s<>r=s=r))"; +by (fast_tac (set_cs addEs [subst_sym,MGIU_iff_lemma]) 1); +val MGIU_iff = result(); + +(**** Idempotence ****) + +goalw Unifier.thy unify_defs "Idem(s) = (s <> s =s= s)"; +by (rtac refl 1); +val raw_Idem_iff = result(); + +goal Unifier.thy "Idem(s) = (sdom(s) Int srange(s) = {})"; +by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp, + invariance,dom_range_disjoint])1); +val Idem_iff = result(); + +goal Unifier.thy "Idem(Nil)"; +by (simp_tac (subst_ss addsimps [raw_Idem_iff,refl RS subst_refl]) 1); +val Idem_Nil = result(); + +goal Unifier.thy "~ (Var(v) <: t) --> Idem(Cons(,Nil))"; +by (simp_tac (subst_ss addsimps [Var_subst,vars_iff_occseq,Idem_iff,srange_iff] + setloop (split_tac [expand_if])) 1); +by (fast_tac set_cs 1); +val Var_Idem = result() RS mp; + +val [prem] = goalw Unifier.thy [Idem_def] + "Idem(r) ==> Unifier(s,t <| r,u <| r) --> Unifier(r <> s,t <| r,u <| r)"; +by (simp_tac (subst_ss addsimps + [Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1); +val Unifier_Idem_subst = result() RS mp; + +val [prem] = goal Unifier.thy + "r <> s =s= s ==> Unifier(s,t,u) --> Unifier(s,t <| r,u <| r)"; +by (simp_tac (subst_ss addsimps + [Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1); +val Unifier_comp_subst = result() RS mp; + +(*** The domain of a MGIU is a subset of the variables in the terms ***) +(*** NB this and one for range are only needed for termination ***) + +val [prem] = goal Unifier.thy + "~ vars_of(Var(x) <| r) = vars_of(Var(x) <| s) ==> ~r =s= s"; +by (rtac (prem RS contrapos) 1); +by (fast_tac (set_cs addEs [subst_subst2]) 1); +val lemma_lemma = result(); + +val prems = goal Unifier.thy + "x : sdom(s) --> ~x : srange(s) --> \ +\ ~vars_of(Var(x) <| s<>Cons(,s)) = \ +\ vars_of(Var(x) <| Cons(,s))"; +by (simp_tac (subst_ss addsimps [not_equal_iff]) 1); +by (REPEAT (resolve_tac [impI,disjI2] 1)); +by(res_inst_tac [("x","x")] exI 1); +br conjI 1; +by (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,repl_invariance]) 1); +by (asm_simp_tac (subst_ss addsimps [Var_subst]) 1); +val MGIU_sdom_lemma = result() RS mp RS mp RS lemma_lemma RS notE;; + +goal Unifier.thy "MGIUnifier(s,t,u) --> sdom(s) <= vars_of(t) Un vars_of(u)"; +by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1); +by (asm_simp_tac (subst_ss addsimps [MGIU_iff,Idem_iff,subset_iff]) 1); +by (safe_tac set_cs); +by (eresolve_tac ([spec] RL [impE]) 1); +by (rtac Cons_Unifier 1); +by (ALLGOALS (fast_tac (set_cs addIs [Cons_Unifier,MGIU_sdom_lemma]))); +val MGIU_sdom = result() RS mp; + +(*** The range of a MGIU is a subset of the variables in the terms ***) + +val prems = goal HOL.thy "P = Q ==> (~P) = (~Q)"; +by (simp_tac (set_ss addsimps prems) 1); +val not_cong = result(); + +val prems = goal Unifier.thy + "~w=x --> x : vars_of(Var(w) <| s) --> w : sdom(s) --> ~w : srange(s) --> \ +\ ~vars_of(Var(w) <| s<>Cons(,s)) = \ +\ vars_of(Var(w) <| Cons(,s))"; +by (simp_tac (subst_ss addsimps [not_equal_iff]) 1); +by (REPEAT (resolve_tac [impI,disjI1] 1)); +by(res_inst_tac [("x","w")] exI 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp, + vars_var_iff RS not_cong RS iffD2 RS repl_invariance]) )); +by (fast_tac (set_cs addIs [Var_in_subst]) 1); +val MGIU_srange_lemma =result() RS mp RS mp RS mp RS mp RS lemma_lemma RS notE; + +goal Unifier.thy "MGIUnifier(s,t,u) --> srange(s) <= vars_of(t) Un vars_of(u)"; +by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1); +by (asm_simp_tac (subst_ss addsimps [MGIU_iff,srange_iff,subset_iff]) 1); +by (simp_tac (subst_ss addsimps [Idem_iff]) 1); +by (safe_tac set_cs); +by (eresolve_tac ([spec] RL [impE]) 1); +by (rtac Cons_Unifier 1); +by (imp_excluded_middle_tac "w=ta" 4); +by (fast_tac (set_cs addEs [MGIU_srange_lemma]) 5); +by (ALLGOALS (fast_tac (set_cs addIs [Var_elim2]))); +val MGIU_srange = result() RS mp; + +(*************** Correctness of a simple unification algorithm ***************) +(* *) +(* fun unify Const(m) Const(n) = if m=n then Nil else Fail *) +(* | unify Const(m) _ = Fail *) +(* | unify Var(v) t = if Var(v)<:t then Fail else Cons(,Nil)*) +(* | unify Comb(t,u) Const(n) = Fail *) +(* | unify Comb(t,u) Var(v) = if Var(v) <: Comb(t,u) then Fail *) +(* else Cons(,Nil) *) +(* | unify Comb(t,u) Comb(v,w) = let s = unify t v *) +(* in if s=Fail then Fail *) +(* else unify (u<|s) (w<|s); *) + +(**** Cases for the partial correctness of the algorithm ****) + +goalw Unifier.thy unify_defs "MGIUnifier(s,t,u) = MGIUnifier(s,u,t)"; +by (safe_tac (HOL_cs addSEs ([sym]@([spec] RL [mp])))); +val Unify_comm = result(); + +goal Unifier.thy "MGIUnifier(Nil,Const(n),Const(n))"; +by (simp_tac (subst_ss addsimps + [MGIU_iff,MGU_iff,Unifier_iff,subst_eq_iff,Idem_Nil]) 1); +val Unify1 = result(); + +goal Unifier.thy "~m=n --> (ALL l.~Unifier(l,Const(m),Const(n)))"; +by (simp_tac (subst_ss addsimps[Unifier_iff]) 1); +val Unify2 = result() RS mp; + +val [prem] = goalw Unifier.thy [MGIUnifier_def] + "~Var(v) <: t ==> MGIUnifier(Cons(,Nil),Var(v),t)"; +by (fast_tac (HOL_cs addSIs [prem RS MGUnifier_Var,prem RS Var_Idem]) 1); +val Unify3 = result(); + +val [prem] = goal Unifier.thy "Var(v) <: t ==> (ALL l.~Unifier(l,Var(v),t))"; +by (simp_tac (subst_ss addsimps + [Unifier_iff,prem RS subst_mono RS occs_irrefl2]) 1); +val Unify4 = result(); + +goal Unifier.thy "ALL l.~Unifier(l,Const(m),Comb(t,u))"; +by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); +val Unify5 = result(); + +goal Unifier.thy + "(ALL l.~Unifier(l,t,v)) --> (ALL l.~Unifier(l,Comb(t,u),Comb(v,w)))"; +by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); +val Unify6 = result() RS mp; + +goal Unifier.thy "MGIUnifier(s,t,v) --> (ALL l.~Unifier(l,u <| s,w <| s)) --> \ +\ (ALL l.~Unifier(l,Comb(t,u),Comb(v,w)))"; +by (simp_tac (subst_ss addsimps [MGIU_iff]) 1); +by (fast_tac (set_cs addIs [Unifier_comp_subst] addSEs [Unifier_Comb]) 1); +val Unify7 = result() RS mp RS mp; + +val [p1,p2,p3] = goal Unifier.thy + "[| Idem(r); Unifier(s,t <| r,u <| r); \ +\ (! q.Unifier(q,t <| r,u <| r) --> s <> q =s= q) |] ==> \ +\ Idem(r <> s)"; +by (cut_facts_tac [p1, + p2 RS (p1 RS Unifier_Idem_subst RS (p3 RS spec RS mp))] 1); +by (REPEAT_SOME (etac rev_mp)); +by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp]) 1); +val Unify8_lemma1 = result(); + +val [p1,p2,p3,p4] = goal Unifier.thy + "[| Unifier(q,t,v); Unifier(q,u,w); (! q.Unifier(q,t,v) --> r <> q =s= q); \ +\ (! q.Unifier(q,u <| r,w <| r) --> s <> q =s= q) |] ==> \ +\ r <> s <> q =s= q"; +val pp = p1 RS (p3 RS spec RS mp); +by (cut_facts_tac [pp, + p2 RS (pp RS Unifier_comp_subst) RS (p4 RS spec RS mp)] 1); +by (REPEAT_SOME (etac rev_mp)); +by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); +val Unify8_lemma2 = result(); + +goal Unifier.thy "MGIUnifier(r,t,v) --> MGIUnifier(s,u <| r,w <| r) --> \ +\ MGIUnifier(r <> s,Comb(t,u),Comb(v,w))"; +by (simp_tac (subst_ss addsimps [MGIU_iff,subst_comp,comp_assoc]) 1); +by (safe_tac HOL_cs); +by (REPEAT (etac rev_mp 2)); +by (simp_tac (subst_ss addsimps + [Unifier_iff,MGIU_iff,subst_comp,comp_assoc]) 2); +by (ALLGOALS (fast_tac (set_cs addEs + [Unifier_Comb,Unify8_lemma1,Unify8_lemma2]))); +val Unify8 = result(); + + +(********************** Termination of the algorithm *************************) +(* *) +(*UWFD is a well-founded relation that orders the 2 recursive calls in unify *) +(* NB well-foundedness of UWFD isn't proved *) + + +goalw Unifier.thy [UWFD_def] "UWFD(t,t',Comb(t,u),Comb(t',u'))"; +by (simp_tac subst_ss 1); +by (fast_tac set_cs 1); +val UnifyWFD1 = result(); + +val [prem] = goal Unifier.thy + "MGIUnifier(s,t,t') ==> vars_of(u <| s) Un vars_of(u' <| s) <= \ +\ vars_of(Comb(t,u)) Un vars_of(Comb(t',u'))"; +by (subgoal_tac "vars_of(u <| s) Un vars_of(u' <| s) <= \ +\ srange(s) Un vars_of(u) Un srange(s) Un vars_of(u')" 1); +by (etac subset_trans 1); +by (ALLGOALS (simp_tac (subst_ss addsimps [Var_intro,subset_iff]))); +by (ALLGOALS (fast_tac (set_cs addDs + [Var_intro,prem RS MGIU_srange RS subsetD]))); +val UWFD2_lemma1 = result(); + +val [major,minor] = goal Unifier.thy + "[| MGIUnifier(s,t,t'); ~ u <| s = u |] ==> \ +\ ~ vars_of(u <| s) Un vars_of(u' <| s) = \ +\ (vars_of(t) Un vars_of(u)) Un (vars_of(t') Un vars_of(u'))"; +by (cut_facts_tac + [major RS (MGIU_iff RS iffD1) RS conjunct1 RS (Idem_iff RS iffD1)] 1); +by (rtac (minor RS subst_not_empty RS exE) 1); +by (rtac (make_elim ((major RS MGIU_sdom) RS subsetD)) 1 THEN assume_tac 1); +by (rtac (disjI2 RS (not_equal_iff RS iffD2)) 1); +by (REPEAT (etac rev_mp 1)); +by (asm_simp_tac subst_ss 1); +by (fast_tac (set_cs addIs [Var_elim2]) 1); +val UWFD2_lemma2 = result(); + +val [prem] = goalw Unifier.thy [UWFD_def] + "MGIUnifier(s,t,t') ==> UWFD(u <| s,u' <| s,Comb(t,u),Comb(t',u'))"; +by (cut_facts_tac + [prem RS UWFD2_lemma1 RS (subseteq_iff_subset_eq RS iffD1)] 1); +by (imp_excluded_middle_tac "u <| s = u" 1); +by (simp_tac (set_ss addsimps [occs_Comb2] ) 1); +by (rtac impI 1 THEN etac subst 1 THEN assume_tac 1); +by (rtac impI 1); +by (rtac (conjI RS (ssubset_iff RS iffD2) RS disjI1) 1); +by (asm_simp_tac (set_ss addsimps [subseteq_iff_subset_eq]) 1); +by (asm_simp_tac subst_ss 1); +by (fast_tac (set_cs addDs [prem RS UWFD2_lemma2]) 1); +val UnifyWFD2 = result(); diff -r 000000000000 -r 7949f97df77a Subst/Unifier.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/Unifier.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,33 @@ +(* Title: Subst/unifier.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Definition of most general idempotent unifier +*) + +Unifier = Subst + + +consts + + Idem :: "('a*('a uterm))list=> bool" + Unifier :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool" + ">>" :: "[('a*('a uterm))list,('a*('a uterm))list] => bool" (infixr 52) + MGUnifier :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool" + MGIUnifier :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool" + UWFD :: "['a uterm,'a uterm,'a uterm,'a uterm] => bool" + +rules (*Definitions*) + + Idem_def "Idem(s) == s <> s =s= s" + Unifier_def "Unifier(s,t,u) == t <| s = u <| s" + MoreGeneral_def "r >> s == ? q.s =s= r <> q" + MGUnifier_def "MGUnifier(s,t,u) == Unifier(s,t,u) & \ +\ (! r.Unifier(r,t,u) --> s >> r)" + MGIUnifier_def "MGIUnifier(s,t,u) == MGUnifier(s,t,u) & Idem(s)" + + UWFD_def + "UWFD(x,y,x',y') == \ +\ (vars_of(x) Un vars_of(y) < vars_of(x') Un vars_of(y')) | \ +\ (vars_of(x) Un vars_of(y) = vars_of(x') Un vars_of(y') & x <: x')" + +end diff -r 000000000000 -r 7949f97df77a Subst/alist.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/alist.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,52 @@ +(* Title: Substitutions/alist.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For alist.thy. +*) + +open AList; + +(*********) + +val al_defs = [alist_rec_def,assoc_def]; + +val alist_ss = pair_ss addsimps [Nil_not_Cons,Cons_not_Nil,Cons_Cons_eq, + list_rec_Nil,list_rec_Cons]; + +val al_rews = + let fun mk_thm s = prove_goalw AList.thy al_defs s + (fn _ => [simp_tac alist_ss 1]) + in map mk_thm + ["alist_rec(Nil,c,d) = c", + "alist_rec(Cons(,al),c,d) = d(a,b,al,alist_rec(al,c,d))", + "assoc(v,d,Nil) = d", + "assoc(v,d,Cons(,al)) = if(v=a,b,assoc(v,d,al))"] end; + +(*********) + +val prems = goal AList.thy + "[| P(Nil); \ +\ !!x y xs. P(xs) ==> P(Cons(,xs)) |] ==> P(l)"; +by (list_ind_tac "l" 1); +by (resolve_tac prems 1); +by (rtac PairE 1); +by (etac ssubst 1); +by (resolve_tac prems 1); +by (assume_tac 1); +val alist_induct = result(); + +(*Perform induction on xs. *) +fun alist_ind_tac a M = + EVERY [res_inst_tac [("l",a)] alist_induct M, + rename_last_tac a ["1"] (M+1)]; + +(* +val prems = goal AList.thy + "[| P(Nil); !! x y xs. P(xs) --> P(Cons(,xs)) |] ==> P(a)"; +by (alist_ind_tac "a" 1); +by (ALLGOALS (cut_facts_tac prems THEN' fast_tac HOL_cs)); +val alist_induct2 = result(); + +add_inds alist_induct2; +*) diff -r 000000000000 -r 7949f97df77a Subst/alist.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/alist.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,22 @@ +(* Title: Substitutions/alist.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Association lists. +*) + +AList = List + + +consts + + alist_rec :: "[('a*'b) list,'c,['a,'b,('a*'b) list,'c]=>'c]=>'c" + assoc :: "['a,'b,('a*'b) list] => 'b" + +rules + + alist_rec_def + "alist_rec(al,b,c) == list_rec(al,b,%p xs g.split(p,%x y.c(x,y,xs,g)))" + + assoc_def "assoc(v,d,al) == alist_rec(al,d,%x y xs g.if(v=x,y,g))" + +end diff -r 000000000000 -r 7949f97df77a Subst/setplus.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/setplus.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,64 @@ +(* Title: Substitutions/setplus.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For setplus.thy. +Properties of subsets and empty sets. +*) + +open Setplus; + +(*********) + +(*** Rules for subsets ***) + +goal Set.thy "A <= B = (! t.t:A --> t:B)"; +by (fast_tac set_cs 1); +val subset_iff = result(); + +goalw Setplus.thy [ssubset_def] "A < B = ((A <= B) & ~(A=B))"; +by (rtac refl 1); +val ssubset_iff = result(); + +goal Setplus.thy "(A::'a set <= B) = ((A < B) | (A=B))"; +by (simp_tac (set_ss addsimps [ssubset_iff]) 1); +by (fast_tac set_cs 1); +val subseteq_iff_subset_eq = result(); + +(*Rule in Modus Ponens style*) +goal Setplus.thy "A < B --> c:A --> c:B"; +by (simp_tac (set_ss addsimps [ssubset_iff]) 1); +by (fast_tac set_cs 1); +val ssubsetD = result(); + +(*********) + +goalw Setplus.thy [empty_def] "~ a : {}"; +by (fast_tac set_cs 1); +val not_in_empty = result(); + +goalw Setplus.thy [empty_def] "(A = {}) = (ALL a.~ a:A)"; +by (fast_tac (set_cs addIs [set_ext]) 1); +val empty_iff = result(); + + +(*********) + +goal Set.thy "(~A=B) = ((? x.x:A & ~x:B) | (? x.~x:A & x:B))"; +by (fast_tac (set_cs addIs [set_ext]) 1); +val not_equal_iff = result(); + +(*********) + +val setplus_rews = [ssubset_iff,not_in_empty,empty_iff]; + +(*********) + +(*Case analysis for rewriting; P also gets rewritten*) +val [prem1,prem2] = goal HOL.thy "[| P-->Q; ~P-->Q |] ==> Q"; +by (rtac (excluded_middle RS disjE) 1); +by (etac (prem2 RS mp) 1); +by (etac (prem1 RS mp) 1); +val imp_excluded_middle = result(); + +fun imp_excluded_middle_tac s = res_inst_tac [("P",s)] imp_excluded_middle; diff -r 000000000000 -r 7949f97df77a Subst/setplus.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/setplus.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: Substitutions/setplus.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Minor additions to HOL's set theory +*) + +Setplus = Set + + +rules + + ssubset_def "A < B == A <= B & ~ A=B" + +end diff -r 000000000000 -r 7949f97df77a Subst/subst.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/subst.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,185 @@ +(* Title: Substitutions/subst.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For subst.thy. +*) + +open Subst; + +(***********) + +val subst_defs = [subst_def,comp_def,sdom_def]; + +val raw_subst_ss = utlemmas_ss addsimps al_rews; + +local fun mk_thm s = prove_goalw Subst.thy subst_defs s + (fn _ => [simp_tac raw_subst_ss 1]) +in val subst_rews = map mk_thm +["Const(c) <| al = Const(c)", + "Comb(t,u) <| al = Comb(t <| al, u <| al)", + "Nil <> bl = bl", + "Cons(,al) <> bl = Cons(, al <> bl)", + "sdom(Nil) = {}", + "sdom(Cons(,al)) = if(Var(a)=b,sdom(al) Int Compl({a}),sdom(al) Un {a})" +]; + (* This rewrite isn't always desired *) + val Var_subst = mk_thm "Var(x) <| al = assoc(x,Var(x),al)"; +end; + +val subst_ss = raw_subst_ss addsimps subst_rews; + +(**** Substitutions ****) + +goal Subst.thy "t <| Nil = t"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); +val subst_Nil = result(); + +goal Subst.thy "t <: u --> t <| s <: u <| s"; +by (uterm_ind_tac "u" 1); +by (ALLGOALS (asm_simp_tac subst_ss)); +val subst_mono = result() RS mp; + +goal Subst.thy "~ (Var(v) <: t) --> t <| Cons(,s) = t <| s"; +by (imp_excluded_middle_tac "t = Var(v)" 1); +by (res_inst_tac [("P", + "%x.~x=Var(v) --> ~(Var(v) <: x) --> x <| Cons(,s)=x<|s")] + uterm_induct 2); +by (ALLGOALS (simp_tac (subst_ss addsimps [Var_subst]))); +by (fast_tac HOL_cs 1); +val Var_not_occs = result() RS mp; + +goal Subst.thy + "(t <|r = t <|s) = (! v.v : vars_of(t) --> Var(v) <|r = Var(v) <|s)"; +by (uterm_ind_tac "t" 1); +by (REPEAT (etac rev_mp 3)); +by (ALLGOALS (asm_simp_tac subst_ss)); +by (ALLGOALS (fast_tac HOL_cs)); +val agreement = result(); + +goal Subst.thy "~ v: vars_of(t) --> t <| Cons(,s) = t <| s"; +by(simp_tac(subst_ss addsimps [agreement,Var_subst] + setloop (split_tac [expand_if])) 1); +val repl_invariance = result() RS mp; + +val asms = goal Subst.thy + "v : vars_of(t) --> w : vars_of(t <| Cons(,s))"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); +val Var_in_subst = result() RS mp; + +(**** Equality between Substitutions ****) + +goalw Subst.thy [subst_eq_def] "r =s= s = (! t.t <| r = t <| s)"; +by (simp_tac subst_ss 1); +val subst_eq_iff = result(); + +local fun mk_thm s = prove_goal Subst.thy s + (fn prems => [cut_facts_tac prems 1, + REPEAT (etac rev_mp 1), + simp_tac (subst_ss addsimps [subst_eq_iff]) 1]) +in + val subst_refl = mk_thm "r = s ==> r =s= s"; + val subst_sym = mk_thm "r =s= s ==> s =s= r"; + val subst_trans = mk_thm "[| q =s= r; r =s= s |] ==> q =s= s"; +end; + +val eq::prems = goalw Subst.thy [subst_eq_def] + "[| r =s= s; P(t <| r,u <| r) |] ==> P(t <| s,u <| s)"; +by (resolve_tac [eq RS spec RS subst] 1); +by (resolve_tac (prems RL [eq RS spec RS subst]) 1); +val subst_subst2 = result(); + +val ssubst_subst2 = subst_sym RS subst_subst2; + +(**** Composition of Substitutions ****) + +goal Subst.thy "s <> Nil = s"; +by (alist_ind_tac "s" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [subst_Nil]))); +val comp_Nil = result(); + +goal Subst.thy "(t <| r <> s) = (t <| r <| s)"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); +by (alist_ind_tac "r" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst,subst_Nil] + setloop (split_tac [expand_if])))); +val subst_comp = result(); + +goal Subst.thy "q <> r <> s =s= q <> (r <> s)"; +by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); +val comp_assoc = result(); + +goal Subst.thy "Cons(,s) =s= s"; +by (rtac (allI RS (subst_eq_iff RS iffD2)) 1); +by (uterm_ind_tac "t" 1); +by (REPEAT (etac rev_mp 3)); +by (ALLGOALS (simp_tac (subst_ss addsimps[Var_subst] + setloop (split_tac [expand_if])))); +val Cons_trivial = result(); + +val [prem] = goal Subst.thy "q <> r =s= s ==> t <| q <| r = t <| s"; +by (simp_tac (subst_ss addsimps [prem RS (subst_eq_iff RS iffD1), + subst_comp RS sym]) 1); +val comp_subst_subst = result(); + +(**** Domain and range of Substitutions ****) + +goal Subst.thy "(v : sdom(s)) = (~ Var(v) <| s = Var(v))"; +by (alist_ind_tac "s" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst] + setloop (split_tac[expand_if])))); +by (fast_tac HOL_cs 1); +val sdom_iff = result(); + +goalw Subst.thy [srange_def] + "v : srange(s) = (? w.w : sdom(s) & v : vars_of(Var(w) <| s))"; +by (fast_tac set_cs 1); +val srange_iff = result(); + +goal Subst.thy "(t <| s = t) = (sdom(s) Int vars_of(t) = {})"; +by (uterm_ind_tac "t" 1); +by (REPEAT (etac rev_mp 3)); +by (ALLGOALS (simp_tac (subst_ss addsimps [sdom_iff,Var_subst]))); +by (ALLGOALS (fast_tac set_cs)); +val invariance = result(); + +goal Subst.thy "v : sdom(s) --> ~v : srange(s) --> ~v : vars_of(t <| s)"; +by (uterm_ind_tac "t" 1); +by (imp_excluded_middle_tac "x : sdom(s)" 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [sdom_iff,srange_iff]))); +by (ALLGOALS (fast_tac set_cs)); +val Var_elim = result() RS mp RS mp; + +val asms = goal Subst.thy + "[| v : sdom(s); v : vars_of(t <| s) |] ==> v : srange(s)"; +by (REPEAT (ares_tac (asms @ [Var_elim RS swap RS classical]) 1)); +val Var_elim2 = result(); + +goal Subst.thy "v : vars_of(t <| s) --> v : srange(s) | v : vars_of(t)"; +by (uterm_ind_tac "t" 1); +by (REPEAT_SOME (etac rev_mp )); +by (ALLGOALS (simp_tac (subst_ss addsimps [sdom_iff,srange_iff]))); +by (REPEAT (step_tac (set_cs addIs [vars_var_iff RS iffD1 RS sym]) 1)); +by (etac notE 1); +by (etac subst 1); +by (ALLGOALS (fast_tac set_cs)); +val Var_intro = result() RS mp; + +goal Subst.thy + "v : srange(s) --> (? w.w : sdom(s) & v : vars_of(Var(w) <| s))"; +by (simp_tac (subst_ss addsimps [srange_iff]) 1); +val srangeE = make_elim (result() RS mp); + +val asms = goal Subst.thy + "sdom(s) Int srange(s) = {} = (! t.sdom(s) Int vars_of(t <| s) = {})"; +by (simp_tac subst_ss 1); +by (fast_tac (set_cs addIs [Var_elim2] addEs [srangeE]) 1); +val dom_range_disjoint = result(); + +val asms = goal Subst.thy "~ u <| s = u --> (? x.x : sdom(s))"; +by (simp_tac (subst_ss addsimps [invariance]) 1); +by (fast_tac set_cs 1); +val subst_not_empty = result() RS mp; diff -r 000000000000 -r 7949f97df77a Subst/subst.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/subst.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,37 @@ +(* Title: Substitutions/subst.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Substitutions on uterms +*) + +Subst = Setplus + AList + UTLemmas + + +consts + + "=s=" :: "[('a*('a uterm)) list,('a*('a uterm)) list] => bool" (infixr 52) + + "<|" :: "['a uterm,('a*('a uterm)) list] => 'a uterm" (infixl 55) + "<>" :: "[('a*('a uterm)) list, ('a*('a uterm)) list] => \ +\ ('a*('a uterm)) list" (infixl 56) + sdom :: "('a*('a uterm)) list => 'a set" + srange :: "('a*('a uterm)) list => 'a set" + +rules + + subst_eq_def "r =s= s == ALL t.t <| r = t <| s" + + subst_def + "t <| al == uterm_rec(t, %x.assoc(x,Var(x),al), \ +\ %x.Const(x), \ +\ %u v q r.Comb(q,r))" + + comp_def "al <> bl == alist_rec(al,bl,%x y xs g.Cons(,g))" + + sdom_def + "sdom(al) == alist_rec(al, {}, \ +\ %x y xs g.if(Var(x)=y, g Int Compl({x}), g Un {x}))" + srange_def + "srange(al) == Union({y. EX x:sdom(al).y=vars_of(Var(x) <| al)})" + +end diff -r 000000000000 -r 7949f97df77a Subst/unifier.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/unifier.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,299 @@ +(* Title: Substitutions/unifier.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For unifier.thy. +Properties of unifiers. +Cases for partial correctness of algorithm and conditions for termination. +*) + +open Unifier; + +val unify_defs = + [Idem_def,Unifier_def,MoreGeneral_def,MGUnifier_def,MGIUnifier_def]; + +(**** Unifiers ****) + +goalw Unifier.thy [Unifier_def] "Unifier(s,t,u) = (t <| s = u <| s)"; +by (rtac refl 1); +val Unifier_iff = result(); + +goal Unifier.thy + "Unifier(s,Comb(t,u),Comb(v,w)) --> Unifier(s,t,v) & Unifier(s,u,w)"; +by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); +val Unifier_Comb = result() RS mp RS conjE; + +goal Unifier.thy + "~v : vars_of(t) --> ~v : vars_of(u) -->Unifier(s,t,u) --> \ +\ Unifier(Cons(,s),t,u)"; +by (simp_tac (subst_ss addsimps [Unifier_iff,repl_invariance]) 1); +val Cons_Unifier = result() RS mp RS mp RS mp; + +(**** Most General Unifiers ****) + +goalw Unifier.thy [MoreGeneral_def] "r >> s = (EX q. s =s= r <> q)"; +by (rtac refl 1); +val MoreGen_iff = result(); + +goal Unifier.thy "Nil >> s"; +by (simp_tac (subst_ss addsimps [MoreGen_iff]) 1); +by (fast_tac (set_cs addIs [refl RS subst_refl]) 1); +val MoreGen_Nil = result(); + +goalw Unifier.thy unify_defs + "MGUnifier(s,t,u) = (ALL r.Unifier(r,t,u) = s >> r)"; +by (REPEAT (ares_tac [iffI,allI] 1 ORELSE + eresolve_tac [conjE,allE,mp,exE,ssubst_subst2] 1)); +by (asm_simp_tac (subst_ss addsimps [subst_comp]) 1); +by (fast_tac (set_cs addIs [comp_Nil RS sym RS subst_refl]) 1); +val MGU_iff = result(); + +val [prem] = goal Unifier.thy + "~ Var(v) <: t ==> MGUnifier(Cons(,Nil),Var(v),t)"; +by (simp_tac (subst_ss addsimps [MGU_iff,MoreGen_iff,Unifier_iff]) 1); +by (REPEAT_SOME (step_tac set_cs)); +by (etac subst 1); +by (etac ssubst_subst2 2); +by (rtac (Cons_trivial RS subst_sym) 1); +by (simp_tac (subst_ss addsimps [prem RS Var_not_occs,Var_subst]) 1); +val MGUnifier_Var = result(); + +(**** Most General Idempotent Unifiers ****) + +goal Unifier.thy "r <> r =s= r --> s =s= r <> q --> r <> s =s= s"; +by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); +val MGIU_iff_lemma = result() RS mp RS mp; + +goalw Unifier.thy unify_defs + "MGIUnifier(s,t,u) = \ +\ (Idem(s) & Unifier(s,t,u) & (ALL r.Unifier(r,t,u) --> s<>r=s=r))"; +by (fast_tac (set_cs addEs [subst_sym,MGIU_iff_lemma]) 1); +val MGIU_iff = result(); + +(**** Idempotence ****) + +goalw Unifier.thy unify_defs "Idem(s) = (s <> s =s= s)"; +by (rtac refl 1); +val raw_Idem_iff = result(); + +goal Unifier.thy "Idem(s) = (sdom(s) Int srange(s) = {})"; +by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp, + invariance,dom_range_disjoint])1); +val Idem_iff = result(); + +goal Unifier.thy "Idem(Nil)"; +by (simp_tac (subst_ss addsimps [raw_Idem_iff,refl RS subst_refl]) 1); +val Idem_Nil = result(); + +goal Unifier.thy "~ (Var(v) <: t) --> Idem(Cons(,Nil))"; +by (simp_tac (subst_ss addsimps [Var_subst,vars_iff_occseq,Idem_iff,srange_iff] + setloop (split_tac [expand_if])) 1); +by (fast_tac set_cs 1); +val Var_Idem = result() RS mp; + +val [prem] = goalw Unifier.thy [Idem_def] + "Idem(r) ==> Unifier(s,t <| r,u <| r) --> Unifier(r <> s,t <| r,u <| r)"; +by (simp_tac (subst_ss addsimps + [Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1); +val Unifier_Idem_subst = result() RS mp; + +val [prem] = goal Unifier.thy + "r <> s =s= s ==> Unifier(s,t,u) --> Unifier(s,t <| r,u <| r)"; +by (simp_tac (subst_ss addsimps + [Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1); +val Unifier_comp_subst = result() RS mp; + +(*** The domain of a MGIU is a subset of the variables in the terms ***) +(*** NB this and one for range are only needed for termination ***) + +val [prem] = goal Unifier.thy + "~ vars_of(Var(x) <| r) = vars_of(Var(x) <| s) ==> ~r =s= s"; +by (rtac (prem RS contrapos) 1); +by (fast_tac (set_cs addEs [subst_subst2]) 1); +val lemma_lemma = result(); + +val prems = goal Unifier.thy + "x : sdom(s) --> ~x : srange(s) --> \ +\ ~vars_of(Var(x) <| s<>Cons(,s)) = \ +\ vars_of(Var(x) <| Cons(,s))"; +by (simp_tac (subst_ss addsimps [not_equal_iff]) 1); +by (REPEAT (resolve_tac [impI,disjI2] 1)); +by(res_inst_tac [("x","x")] exI 1); +br conjI 1; +by (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,repl_invariance]) 1); +by (asm_simp_tac (subst_ss addsimps [Var_subst]) 1); +val MGIU_sdom_lemma = result() RS mp RS mp RS lemma_lemma RS notE;; + +goal Unifier.thy "MGIUnifier(s,t,u) --> sdom(s) <= vars_of(t) Un vars_of(u)"; +by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1); +by (asm_simp_tac (subst_ss addsimps [MGIU_iff,Idem_iff,subset_iff]) 1); +by (safe_tac set_cs); +by (eresolve_tac ([spec] RL [impE]) 1); +by (rtac Cons_Unifier 1); +by (ALLGOALS (fast_tac (set_cs addIs [Cons_Unifier,MGIU_sdom_lemma]))); +val MGIU_sdom = result() RS mp; + +(*** The range of a MGIU is a subset of the variables in the terms ***) + +val prems = goal HOL.thy "P = Q ==> (~P) = (~Q)"; +by (simp_tac (set_ss addsimps prems) 1); +val not_cong = result(); + +val prems = goal Unifier.thy + "~w=x --> x : vars_of(Var(w) <| s) --> w : sdom(s) --> ~w : srange(s) --> \ +\ ~vars_of(Var(w) <| s<>Cons(,s)) = \ +\ vars_of(Var(w) <| Cons(,s))"; +by (simp_tac (subst_ss addsimps [not_equal_iff]) 1); +by (REPEAT (resolve_tac [impI,disjI1] 1)); +by(res_inst_tac [("x","w")] exI 1); +by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp, + vars_var_iff RS not_cong RS iffD2 RS repl_invariance]) )); +by (fast_tac (set_cs addIs [Var_in_subst]) 1); +val MGIU_srange_lemma =result() RS mp RS mp RS mp RS mp RS lemma_lemma RS notE; + +goal Unifier.thy "MGIUnifier(s,t,u) --> srange(s) <= vars_of(t) Un vars_of(u)"; +by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1); +by (asm_simp_tac (subst_ss addsimps [MGIU_iff,srange_iff,subset_iff]) 1); +by (simp_tac (subst_ss addsimps [Idem_iff]) 1); +by (safe_tac set_cs); +by (eresolve_tac ([spec] RL [impE]) 1); +by (rtac Cons_Unifier 1); +by (imp_excluded_middle_tac "w=ta" 4); +by (fast_tac (set_cs addEs [MGIU_srange_lemma]) 5); +by (ALLGOALS (fast_tac (set_cs addIs [Var_elim2]))); +val MGIU_srange = result() RS mp; + +(*************** Correctness of a simple unification algorithm ***************) +(* *) +(* fun unify Const(m) Const(n) = if m=n then Nil else Fail *) +(* | unify Const(m) _ = Fail *) +(* | unify Var(v) t = if Var(v)<:t then Fail else Cons(,Nil)*) +(* | unify Comb(t,u) Const(n) = Fail *) +(* | unify Comb(t,u) Var(v) = if Var(v) <: Comb(t,u) then Fail *) +(* else Cons(,Nil) *) +(* | unify Comb(t,u) Comb(v,w) = let s = unify t v *) +(* in if s=Fail then Fail *) +(* else unify (u<|s) (w<|s); *) + +(**** Cases for the partial correctness of the algorithm ****) + +goalw Unifier.thy unify_defs "MGIUnifier(s,t,u) = MGIUnifier(s,u,t)"; +by (safe_tac (HOL_cs addSEs ([sym]@([spec] RL [mp])))); +val Unify_comm = result(); + +goal Unifier.thy "MGIUnifier(Nil,Const(n),Const(n))"; +by (simp_tac (subst_ss addsimps + [MGIU_iff,MGU_iff,Unifier_iff,subst_eq_iff,Idem_Nil]) 1); +val Unify1 = result(); + +goal Unifier.thy "~m=n --> (ALL l.~Unifier(l,Const(m),Const(n)))"; +by (simp_tac (subst_ss addsimps[Unifier_iff]) 1); +val Unify2 = result() RS mp; + +val [prem] = goalw Unifier.thy [MGIUnifier_def] + "~Var(v) <: t ==> MGIUnifier(Cons(,Nil),Var(v),t)"; +by (fast_tac (HOL_cs addSIs [prem RS MGUnifier_Var,prem RS Var_Idem]) 1); +val Unify3 = result(); + +val [prem] = goal Unifier.thy "Var(v) <: t ==> (ALL l.~Unifier(l,Var(v),t))"; +by (simp_tac (subst_ss addsimps + [Unifier_iff,prem RS subst_mono RS occs_irrefl2]) 1); +val Unify4 = result(); + +goal Unifier.thy "ALL l.~Unifier(l,Const(m),Comb(t,u))"; +by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); +val Unify5 = result(); + +goal Unifier.thy + "(ALL l.~Unifier(l,t,v)) --> (ALL l.~Unifier(l,Comb(t,u),Comb(v,w)))"; +by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); +val Unify6 = result() RS mp; + +goal Unifier.thy "MGIUnifier(s,t,v) --> (ALL l.~Unifier(l,u <| s,w <| s)) --> \ +\ (ALL l.~Unifier(l,Comb(t,u),Comb(v,w)))"; +by (simp_tac (subst_ss addsimps [MGIU_iff]) 1); +by (fast_tac (set_cs addIs [Unifier_comp_subst] addSEs [Unifier_Comb]) 1); +val Unify7 = result() RS mp RS mp; + +val [p1,p2,p3] = goal Unifier.thy + "[| Idem(r); Unifier(s,t <| r,u <| r); \ +\ (! q.Unifier(q,t <| r,u <| r) --> s <> q =s= q) |] ==> \ +\ Idem(r <> s)"; +by (cut_facts_tac [p1, + p2 RS (p1 RS Unifier_Idem_subst RS (p3 RS spec RS mp))] 1); +by (REPEAT_SOME (etac rev_mp)); +by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp]) 1); +val Unify8_lemma1 = result(); + +val [p1,p2,p3,p4] = goal Unifier.thy + "[| Unifier(q,t,v); Unifier(q,u,w); (! q.Unifier(q,t,v) --> r <> q =s= q); \ +\ (! q.Unifier(q,u <| r,w <| r) --> s <> q =s= q) |] ==> \ +\ r <> s <> q =s= q"; +val pp = p1 RS (p3 RS spec RS mp); +by (cut_facts_tac [pp, + p2 RS (pp RS Unifier_comp_subst) RS (p4 RS spec RS mp)] 1); +by (REPEAT_SOME (etac rev_mp)); +by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); +val Unify8_lemma2 = result(); + +goal Unifier.thy "MGIUnifier(r,t,v) --> MGIUnifier(s,u <| r,w <| r) --> \ +\ MGIUnifier(r <> s,Comb(t,u),Comb(v,w))"; +by (simp_tac (subst_ss addsimps [MGIU_iff,subst_comp,comp_assoc]) 1); +by (safe_tac HOL_cs); +by (REPEAT (etac rev_mp 2)); +by (simp_tac (subst_ss addsimps + [Unifier_iff,MGIU_iff,subst_comp,comp_assoc]) 2); +by (ALLGOALS (fast_tac (set_cs addEs + [Unifier_Comb,Unify8_lemma1,Unify8_lemma2]))); +val Unify8 = result(); + + +(********************** Termination of the algorithm *************************) +(* *) +(*UWFD is a well-founded relation that orders the 2 recursive calls in unify *) +(* NB well-foundedness of UWFD isn't proved *) + + +goalw Unifier.thy [UWFD_def] "UWFD(t,t',Comb(t,u),Comb(t',u'))"; +by (simp_tac subst_ss 1); +by (fast_tac set_cs 1); +val UnifyWFD1 = result(); + +val [prem] = goal Unifier.thy + "MGIUnifier(s,t,t') ==> vars_of(u <| s) Un vars_of(u' <| s) <= \ +\ vars_of(Comb(t,u)) Un vars_of(Comb(t',u'))"; +by (subgoal_tac "vars_of(u <| s) Un vars_of(u' <| s) <= \ +\ srange(s) Un vars_of(u) Un srange(s) Un vars_of(u')" 1); +by (etac subset_trans 1); +by (ALLGOALS (simp_tac (subst_ss addsimps [Var_intro,subset_iff]))); +by (ALLGOALS (fast_tac (set_cs addDs + [Var_intro,prem RS MGIU_srange RS subsetD]))); +val UWFD2_lemma1 = result(); + +val [major,minor] = goal Unifier.thy + "[| MGIUnifier(s,t,t'); ~ u <| s = u |] ==> \ +\ ~ vars_of(u <| s) Un vars_of(u' <| s) = \ +\ (vars_of(t) Un vars_of(u)) Un (vars_of(t') Un vars_of(u'))"; +by (cut_facts_tac + [major RS (MGIU_iff RS iffD1) RS conjunct1 RS (Idem_iff RS iffD1)] 1); +by (rtac (minor RS subst_not_empty RS exE) 1); +by (rtac (make_elim ((major RS MGIU_sdom) RS subsetD)) 1 THEN assume_tac 1); +by (rtac (disjI2 RS (not_equal_iff RS iffD2)) 1); +by (REPEAT (etac rev_mp 1)); +by (asm_simp_tac subst_ss 1); +by (fast_tac (set_cs addIs [Var_elim2]) 1); +val UWFD2_lemma2 = result(); + +val [prem] = goalw Unifier.thy [UWFD_def] + "MGIUnifier(s,t,t') ==> UWFD(u <| s,u' <| s,Comb(t,u),Comb(t',u'))"; +by (cut_facts_tac + [prem RS UWFD2_lemma1 RS (subseteq_iff_subset_eq RS iffD1)] 1); +by (imp_excluded_middle_tac "u <| s = u" 1); +by (simp_tac (set_ss addsimps [occs_Comb2] ) 1); +by (rtac impI 1 THEN etac subst 1 THEN assume_tac 1); +by (rtac impI 1); +by (rtac (conjI RS (ssubset_iff RS iffD2) RS disjI1) 1); +by (asm_simp_tac (set_ss addsimps [subseteq_iff_subset_eq]) 1); +by (asm_simp_tac subst_ss 1); +by (fast_tac (set_cs addDs [prem RS UWFD2_lemma2]) 1); +val UnifyWFD2 = result(); diff -r 000000000000 -r 7949f97df77a Subst/unifier.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/unifier.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,33 @@ +(* Title: Subst/unifier.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Definition of most general idempotent unifier +*) + +Unifier = Subst + + +consts + + Idem :: "('a*('a uterm))list=> bool" + Unifier :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool" + ">>" :: "[('a*('a uterm))list,('a*('a uterm))list] => bool" (infixr 52) + MGUnifier :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool" + MGIUnifier :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool" + UWFD :: "['a uterm,'a uterm,'a uterm,'a uterm] => bool" + +rules (*Definitions*) + + Idem_def "Idem(s) == s <> s =s= s" + Unifier_def "Unifier(s,t,u) == t <| s = u <| s" + MoreGeneral_def "r >> s == ? q.s =s= r <> q" + MGUnifier_def "MGUnifier(s,t,u) == Unifier(s,t,u) & \ +\ (! r.Unifier(r,t,u) --> s >> r)" + MGIUnifier_def "MGIUnifier(s,t,u) == MGUnifier(s,t,u) & Idem(s)" + + UWFD_def + "UWFD(x,y,x',y') == \ +\ (vars_of(x) Un vars_of(y) < vars_of(x') Un vars_of(y')) | \ +\ (vars_of(x) Un vars_of(y) = vars_of(x') Un vars_of(y') & x <: x')" + +end diff -r 000000000000 -r 7949f97df77a Subst/uterm.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/uterm.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,320 @@ +(* Title: Substitutions/uterm.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For uterm.thy. +*) + +open UTerm; + +(** the uterm functional **) + +goal UTerm.thy "mono(%Z. A <+> A <+> Z <*> Z)"; +by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono] 1)); +val UTerm_fun_mono = result(); + +val UTerm_unfold = UTerm_fun_mono RS (UTerm_def RS def_lfp_Tarski); + +(*This justifies using UTerm in other recursive type definitions*) +val prems = goalw UTerm.thy [UTerm_def] "[| A<=B |] ==> UTerm(A) <= UTerm(B)"; +by (REPEAT (ares_tac (prems@[monoI, subset_refl, lfp_mono, + usum_mono, uprod_mono]) 1)); +val UTerm_mono = result(); + +(** Type checking rules -- UTerm creates well-founded sets **) + +val prems = goalw UTerm.thy [UTerm_def] "UTerm(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_In0I,Sexp_In1I,Sexp_SconsI]) 1); +val UTerm_Sexp = result(); + +(* A <= Sexp ==> UTerm(A) <= Sexp *) +val UTerm_subset_Sexp = standard + (UTerm_mono RS (UTerm_Sexp RSN (2,subset_trans))); + +(** Induction **) + +(*Induction for the set UTerm(A) *) +val major::prems = goalw UTerm.thy [VAR_def,CONST_def,COMB_def] + "[| M: UTerm(A); !!M.M : A ==> P(VAR(M)); !!M.M : A ==> P(CONST(M)); \ +\ !!M N. [| M: UTerm(A); N: UTerm(A); P(M); P(N) |] ==> P(COMB(M,N)) |] \ +\ ==> P(M)"; +by (rtac (major RS (UTerm_def RS def_induct)) 1); +by (rtac UTerm_fun_mono 1); +by (fast_tac (set_cs addIs prems addEs [usumE,uprodE]) 1); +val UTerm_induct = result(); + +(*Induction for the type 'a uterm *) +val prems = goalw UTerm.thy [Var_def,Const_def,Comb_def] + "[| !!x.P(Var(x)); !!x.P(Const(x)); \ +\ !!u v. [| P(u); P(v) |] ==> P(Comb(u,v)) |] ==> P(t)"; +by (rtac (Rep_UTerm_inverse RS subst) 1); (*types force good instantiation*) +by (rtac (Rep_UTerm RS UTerm_induct) 1); +by (REPEAT (ares_tac prems 1 + ORELSE eresolve_tac [rangeE, ssubst, Abs_UTerm_inverse RS subst] 1)); +val uterm_induct = result(); + +(*Perform induction on xs. *) +fun uterm_ind_tac a M = + EVERY [res_inst_tac [("t",a)] uterm_induct M, + rename_last_tac a ["1"] (M+1)]; + +(** Introduction rules for UTerm constructors **) + +(* c : A <+> A <+> UTerm(A) <*> UTerm(A) ==> c : UTerm(A) *) +val UTermI = UTerm_unfold RS equalityD2 RS subsetD; + +(* Nil is a UTerm -- this also justifies the type definition*) +val prems = goalw UTerm.thy [VAR_def] "v:A ==> VAR(v) : UTerm(A)"; +by (REPEAT (resolve_tac ([singletonI, UTermI, usum_In0I]@prems) 1)); +val VAR_I = result(); + +val prems = goalw UTerm.thy [CONST_def] "c:A ==> CONST(c) : UTerm(A)"; +by (REPEAT (resolve_tac ([singletonI, UTermI, usum_In0I, usum_In1I]@prems) 1)); +val CONST_I = result(); + +val prems = goalw UTerm.thy [COMB_def] + "[| M:UTerm(A); N:UTerm(A) |] ==> COMB(M,N) : UTerm(A)"; +by (REPEAT (resolve_tac (prems@[UTermI, uprodI, usum_In1I]) 1)); +val COMB_I = result(); + +(*** Isomorphisms ***) + +goal UTerm.thy "inj(Rep_UTerm)"; +by (rtac inj_inverseI 1); +by (rtac Rep_UTerm_inverse 1); +val inj_Rep_UTerm = result(); + +goal UTerm.thy "inj_onto(Abs_UTerm,UTerm(range(Leaf)))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_UTerm_inverse 1); +val inj_onto_Abs_UTerm = result(); + +(** Distinctness of constructors **) + +goalw UTerm.thy [CONST_def,COMB_def] "~ CONST(c) = COMB(u,v)"; +by (rtac notI 1); +by (etac (In1_inject RS (In0_not_In1 RS notE)) 1); +val CONST_not_COMB = result(); +val COMB_not_CONST = standard (CONST_not_COMB RS not_sym); +val CONST_neq_COMB = standard (CONST_not_COMB RS notE); +val COMB_neq_CONST = sym RS CONST_neq_COMB; + +goalw UTerm.thy [COMB_def,VAR_def] "~ COMB(u,v) = VAR(x)"; +by (rtac In1_not_In0 1); +val COMB_not_VAR = result(); +val VAR_not_COMB = standard (COMB_not_VAR RS not_sym); +val COMB_neq_VAR = standard (COMB_not_VAR RS notE); +val VAR_neq_COMB = sym RS COMB_neq_VAR; + +goalw UTerm.thy [VAR_def,CONST_def] "~ VAR(x) = CONST(c)"; +by (rtac In0_not_In1 1); +val VAR_not_CONST = result(); +val CONST_not_VAR = standard (VAR_not_CONST RS not_sym); +val VAR_neq_CONST = standard (VAR_not_CONST RS notE); +val CONST_neq_VAR = sym RS VAR_neq_CONST; + + +goalw UTerm.thy [Const_def,Comb_def] "~ Const(c) = Comb(u,v)"; +by (rtac (CONST_not_COMB RS (inj_onto_Abs_UTerm RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, VAR_I, CONST_I, COMB_I, Rep_UTerm] 1)); +val Const_not_Comb = result(); +val Comb_not_Const = standard (Const_not_Comb RS not_sym); +val Const_neq_Comb = standard (Const_not_Comb RS notE); +val Comb_neq_Const = sym RS Const_neq_Comb; + +goalw UTerm.thy [Comb_def,Var_def] "~ Comb(u,v) = Var(x)"; +by (rtac (COMB_not_VAR RS (inj_onto_Abs_UTerm RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, VAR_I, CONST_I, COMB_I, Rep_UTerm] 1)); +val Comb_not_Var = result(); +val Var_not_Comb = standard (Comb_not_Var RS not_sym); +val Comb_neq_Var = standard (Comb_not_Var RS notE); +val Var_neq_Comb = sym RS Comb_neq_Var; + +goalw UTerm.thy [Var_def,Const_def] "~ Var(x) = Const(c)"; +by (rtac (VAR_not_CONST RS (inj_onto_Abs_UTerm RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, VAR_I, CONST_I, COMB_I, Rep_UTerm] 1)); +val Var_not_Const = result(); +val Const_not_Var = standard (Var_not_Const RS not_sym); +val Var_neq_Const = standard (Var_not_Const RS notE); +val Const_neq_Var = sym RS Var_neq_Const; + + +(** Injectiveness of CONST and Const **) + +val inject_cs = HOL_cs addSEs [Scons_inject] + addSDs [In0_inject,In1_inject]; + +goalw UTerm.thy [VAR_def] "(VAR(M)=VAR(N)) = (M=N)"; +by (fast_tac inject_cs 1); +val VAR_VAR_eq = result(); + +goalw UTerm.thy [CONST_def] "(CONST(M)=CONST(N)) = (M=N)"; +by (fast_tac inject_cs 1); +val CONST_CONST_eq = result(); + +goalw UTerm.thy [COMB_def] "(COMB(K,L)=COMB(M,N)) = (K=M & L=N)"; +by (fast_tac inject_cs 1); +val COMB_COMB_eq = result(); + +val VAR_inject = standard (VAR_VAR_eq RS iffD1); +val CONST_inject = standard (CONST_CONST_eq RS iffD1); +val COMB_inject = standard (COMB_COMB_eq RS iffD1 RS conjE); + + +(*For reasoning about abstract uterm constructors*) +val UTerm_cs = set_cs addIs [Rep_UTerm, VAR_I, CONST_I, COMB_I] + addSEs [CONST_neq_COMB,COMB_neq_VAR,VAR_neq_CONST, + COMB_neq_CONST,VAR_neq_COMB,CONST_neq_VAR, + COMB_inject] + addSDs [VAR_inject,CONST_inject, + inj_onto_Abs_UTerm RS inj_ontoD, + inj_Rep_UTerm RS injD, Leaf_inject]; + +goalw UTerm.thy [Var_def] "(Var(x)=Var(y)) = (x=y)"; +by (fast_tac UTerm_cs 1); +val Var_Var_eq = result(); +val Var_inject = standard (Var_Var_eq RS iffD1); + +goalw UTerm.thy [Const_def] "(Const(x)=Const(y)) = (x=y)"; +by (fast_tac UTerm_cs 1); +val Const_Const_eq = result(); +val Const_inject = standard (Const_Const_eq RS iffD1); + +goalw UTerm.thy [Comb_def] "(Comb(u,v)=Comb(x,y)) = (u=x & v=y)"; +by (fast_tac UTerm_cs 1); +val Comb_Comb_eq = result(); +val Comb_inject = standard (Comb_Comb_eq RS iffD1 RS conjE); + +val [major] = goal UTerm.thy "VAR(M): UTerm(A) ==> M : A"; +by (rtac (major RS setup_induction) 1); +by (etac UTerm_induct 1); +by (ALLGOALS (fast_tac UTerm_cs)); +val VAR_D = result(); + +val [major] = goal UTerm.thy "CONST(M): UTerm(A) ==> M : A"; +by (rtac (major RS setup_induction) 1); +by (etac UTerm_induct 1); +by (ALLGOALS (fast_tac UTerm_cs)); +val CONST_D = result(); + +val [major] = goal UTerm.thy + "COMB(M,N): UTerm(A) ==> M: UTerm(A) & N: UTerm(A)"; +by (rtac (major RS setup_induction) 1); +by (etac UTerm_induct 1); +by (ALLGOALS (fast_tac UTerm_cs)); +val COMB_D = result(); + +(*Basic ss with constructors and their freeness*) +val uterm_free_simps = [Const_not_Comb,Comb_not_Var,Var_not_Const, + Comb_not_Const,Var_not_Comb,Const_not_Var, + Var_Var_eq,Const_Const_eq,Comb_Comb_eq, + CONST_not_COMB,COMB_not_VAR,VAR_not_CONST, + COMB_not_CONST,VAR_not_COMB,CONST_not_VAR, + VAR_VAR_eq,CONST_CONST_eq,COMB_COMB_eq, + VAR_I, CONST_I, COMB_I]; +val uterm_free_ss = HOL_ss addsimps uterm_free_simps; + +goal UTerm.thy "!u. ~(t=Comb(t,u))"; +by (uterm_ind_tac "t" 1); +by (rtac (Var_not_Comb RS allI) 1); +by (rtac (Const_not_Comb RS allI) 1); +by (asm_simp_tac uterm_free_ss 1); +val t_not_Comb_t = result(); + +goal UTerm.thy "!t. ~(u=Comb(t,u))"; +by (uterm_ind_tac "u" 1); +by (rtac (Var_not_Comb RS allI) 1); +by (rtac (Const_not_Comb RS allI) 1); +by (asm_simp_tac uterm_free_ss 1); +val u_not_Comb_u = result(); + + +(*** UTerm_rec -- by wf recursion on pred_Sexp ***) + +val UTerm_rec_unfold = + wf_pred_Sexp RS wf_trancl RS (UTerm_rec_def RS def_wfrec); + +(** conversion rules **) + +goalw UTerm.thy [VAR_def] "UTerm_rec(VAR(x),b,c,d) = b(x)"; +by (rtac (UTerm_rec_unfold RS trans) 1); +by (rtac Case_In0 1); +val UTerm_rec_VAR = result(); + +goalw UTerm.thy [CONST_def] "UTerm_rec(CONST(x),b,c,d) = c(x)"; +by (rtac (UTerm_rec_unfold RS trans) 1); +by (simp_tac (HOL_ss addsimps [Case_In0,Case_In1]) 1); +val UTerm_rec_CONST = result(); + +val prems = goalw UTerm.thy [COMB_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ UTerm_rec(COMB(M,N), b, c, d) = \ +\ d(M, N, UTerm_rec(M,b,c,d), UTerm_rec(N,b,c,d))"; +by (rtac (UTerm_rec_unfold RS trans) 1); +by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1); +by (simp_tac (pred_Sexp_ss addsimps (In1_def::prems)) 1); +val UTerm_rec_COMB = result(); + +(*** uterm_rec -- by UTerm_rec ***) + +val Rep_UTerm_in_Sexp = + Rep_UTerm RS (range_Leaf_subset_Sexp RS UTerm_subset_Sexp RS subsetD); + +val uterm_rec_simps = [UTerm_rec_VAR, UTerm_rec_CONST, UTerm_rec_COMB, + Abs_UTerm_inverse, Rep_UTerm_inverse, VAR_I, CONST_I, COMB_I, + Rep_UTerm, rangeI, inj_Leaf, Inv_f_f, Rep_UTerm_in_Sexp]; +val uterm_rec_ss = HOL_ss addsimps uterm_rec_simps; + +goalw UTerm.thy [uterm_rec_def, Var_def] "uterm_rec(Var(x),b,c,d) = b(x)"; +by (simp_tac uterm_rec_ss 1); +val uterm_rec_Var = result(); + +goalw UTerm.thy [uterm_rec_def, Const_def] "uterm_rec(Const(x),b,c,d) = c(x)"; +by (simp_tac uterm_rec_ss 1); +val uterm_rec_Const = result(); + +goalw UTerm.thy [uterm_rec_def, Comb_def] + "uterm_rec(Comb(u,v),b,c,d) = d(u,v,uterm_rec(u,b,c,d),uterm_rec(v,b,c,d))"; +by (simp_tac uterm_rec_ss 1); +val uterm_rec_Comb = result(); + +val uterm_simps = [UTerm_rec_VAR, UTerm_rec_CONST, UTerm_rec_COMB, + uterm_rec_Var, uterm_rec_Const, uterm_rec_Comb]; +val uterm_ss = uterm_free_ss addsimps uterm_simps; + +(*Type checking. Useful?*) +val major::A_subset_Sexp::prems = goal UTerm.thy + "[| M: UTerm(A); \ +\ A<=Sexp; \ +\ !!x.x:A ==> b(x): C(VAR(x)); \ +\ !!x.x:A ==> c(x): C(CONST(x)); \ +\ !!x y q r. [| x: UTerm(A); y: UTerm(A); q: C(x); r: C(y) |] ==> \ +\ d(x,y,q,r): C(COMB(x,y)) \ +\ |] ==> UTerm_rec(M,b,c,d) : C(M)"; +val Sexp_UTermA_I = A_subset_Sexp RS UTerm_subset_Sexp RS subsetD; +val Sexp_A_I = A_subset_Sexp RS subsetD; +by (rtac (major RS UTerm_induct) 1); +by (ALLGOALS + (asm_simp_tac (uterm_ss addsimps ([Sexp_A_I,Sexp_UTermA_I] @ prems)))); +val UTerm_rec_type = result(); + + +(**********) + +val uterm_rews = [uterm_rec_Var,uterm_rec_Const,uterm_rec_Comb, + t_not_Comb_t,u_not_Comb_u, + Const_not_Comb,Comb_not_Var,Var_not_Const, + Comb_not_Const,Var_not_Comb,Const_not_Var, + Var_Var_eq,Const_Const_eq,Comb_Comb_eq]; + +(* +val prems = goal Subst.thy + "[| !!x.P(Var(x)); !!x.P(Const(x)); \ +\ !!u v. P(u) --> P(v) --> P(Comb(u,v)) |] ==> P(a)"; +by (uterm_ind_tac "a" 1); +by (ALLGOALS (cut_facts_tac prems THEN' fast_tac HOL_cs)); +val uterm_induct2 = result(); + +add_inds uterm_induct2; +*) diff -r 000000000000 -r 7949f97df77a Subst/uterm.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/uterm.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,58 @@ +(* Title: Substitutions/uterm.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Simple term structure for unifiation. +Binary trees with leaves that are constants or variables. +*) + +UTerm = Sexp + + +types uterm 1 + +arities + uterm :: (term)term + +consts + UTerm :: "'a node set set => 'a node set set" + Rep_UTerm :: "'a uterm => 'a node set" + Abs_UTerm :: "'a node set => 'a uterm" + VAR :: "'a node set => 'a node set" + CONST :: "'a node set => 'a node set" + COMB :: "['a node set, 'a node set] => 'a node set" + Var :: "'a => 'a uterm" + Const :: "'a => 'a uterm" + Comb :: "['a uterm, 'a uterm] => 'a uterm" + UTerm_rec :: "['a node set, 'a node set => 'b, 'a node set => 'b, \ +\ ['a node set , 'a node set, 'b, 'b]=>'b] => 'b" + uterm_rec :: "['a uterm, 'a => 'b, 'a => 'b, \ +\ ['a uterm, 'a uterm,'b,'b]=>'b] => 'b" + +rules + UTerm_def "UTerm(A) == lfp(%Z. A <+> A <+> Z <*> Z)" + (*faking a type definition...*) + Rep_UTerm "Rep_UTerm(xs): UTerm(range(Leaf))" + Rep_UTerm_inverse "Abs_UTerm(Rep_UTerm(xs)) = xs" + Abs_UTerm_inverse "M: UTerm(range(Leaf)) ==> Rep_UTerm(Abs_UTerm(M)) = M" + (*defining the concrete constructors*) + VAR_def "VAR(v) == In0(v)" + CONST_def "CONST(v) == In1(In0(v))" + COMB_def "COMB(t,u) == In1(In1(t . u))" + (*defining the abstract constructors*) + Var_def "Var(v) == Abs_UTerm(VAR(Leaf(v)))" + Const_def "Const(c) == Abs_UTerm(CONST(Leaf(c)))" + Comb_def "Comb(t,u) == Abs_UTerm(COMB(Rep_UTerm(t),Rep_UTerm(u)))" + + (*uterm recursion*) + UTerm_rec_def + "UTerm_rec(M,b,c,d) == wfrec(trancl(pred_Sexp), M, \ +\ %z g. Case(z, %x. b(x), \ +\ %w. Case(w, %x. c(x), \ +\ %v. Split(v, %x y. d(x,y,g(x),g(y))))))" + + uterm_rec_def + "uterm_rec(t,b,c,d) == \ +\ UTerm_rec(Rep_UTerm(t), %x.b(Inv(Leaf,x)), %x.c(Inv(Leaf,x)), \ +\ %x y q r.d(Abs_UTerm(x),Abs_UTerm(y),q,r))" + +end diff -r 000000000000 -r 7949f97df77a Subst/utlemmas.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/utlemmas.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,77 @@ +(* Title: Substitutions/utermlemmas.ML + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For utermlemmas.thy. +*) + +open UTLemmas; + +(***********) + +val utlemmas_defs = [vars_of_def,occs_def]; + +local fun mk_thm s = prove_goalw UTLemmas.thy utlemmas_defs s + (fn _ => [simp_tac uterm_ss 1]) +in val utlemmas_rews = map mk_thm + ["vars_of(Const(c)) = {}", + "vars_of(Var(x)) = {x}", + "vars_of(Comb(t,u)) = vars_of(t) Un vars_of(u)", + "t <: Const(c) = False", + "t <: Var(x) = False", + "t <: Comb(u,v) = (t=u | t=v | t <: u | t <: v)"]; +end; + +val utlemmas_ss = pair_ss addsimps (setplus_rews @ uterm_rews @ utlemmas_rews); + +(**** occs irrefl ****) + +goal UTLemmas.thy "t <: u & u <: v --> t <: v"; +by (uterm_ind_tac "v" 1); +by (ALLGOALS (simp_tac utlemmas_ss)); +by (fast_tac HOL_cs 1); +val occs_trans = conjI RS (result() RS mp); + +goal UTLemmas.thy "~ t <: v --> ~ t <: u | ~ u <: v"; +by (fast_tac (HOL_cs addIs [occs_trans]) 1); +val contr_occs_trans = result() RS mp; + +goal UTLemmas.thy "t <: Comb(t,u)"; +by (simp_tac utlemmas_ss 1); +val occs_Comb1 = result(); + +goal UTLemmas.thy "u <: Comb(t,u)"; +by (simp_tac utlemmas_ss 1); +val occs_Comb2 = result(); + +goal HOL.thy "(~(P|Q)) = (~P & ~Q)"; +by (fast_tac HOL_cs 1); +val demorgan_disj = result(); + +goal UTLemmas.thy "~ t <: t"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (simp_tac (utlemmas_ss addsimps [demorgan_disj]))); +by (REPEAT (resolve_tac [impI,conjI] 1 ORELSE + (etac contrapos 1 THEN etac subst 1 THEN + resolve_tac [occs_Comb1,occs_Comb2] 1) ORELSE + (etac (contr_occs_trans RS disjE) 1 THEN assume_tac 2) ORELSE + eresolve_tac ([occs_Comb1,occs_Comb2] RLN(2,[notE])) 1)); +val occs_irrefl = result(); + +goal UTLemmas.thy "t <: u --> ~t=u"; +by (fast_tac (HOL_cs addEs [occs_irrefl RS notE]) 1); +val occs_irrefl2 = result() RS mp; + + +(**** vars_of lemmas ****) + +goal UTLemmas.thy "(v : vars_of(Var(w))) = (w=v)"; +by (simp_tac utlemmas_ss 1); +by (fast_tac HOL_cs 1); +val vars_var_iff = result(); + +goal UTLemmas.thy "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)"; +by (uterm_ind_tac "t" 1); +by (ALLGOALS (simp_tac utlemmas_ss)); +by (fast_tac HOL_cs 1); +val vars_iff_occseq = result(); diff -r 000000000000 -r 7949f97df77a Subst/utlemmas.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Subst/utlemmas.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,20 @@ +(* Title: Substitutions/utermlemmas.thy + Author: Martin Coen, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Additional definitions for uterms that are not part of the basic inductive definition. +*) + +UTLemmas = UTerm + Setplus + + +consts + + vars_of :: "'a uterm=>'a set" + "<:" :: "['a uterm,'a uterm]=>bool" (infixl 54) + +rules (*Definitions*) + + vars_of_def "vars_of(t) == uterm_rec(t,%x.{x},%x.{},%u v q r.q Un r)" + occs_def "s <: t == uterm_rec(t,%x.False,%x.False,%u v q r.s=u | s=v | q | r)" + +end diff -r 000000000000 -r 7949f97df77a Sum.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Sum.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,130 @@ +(* Title: HOL/sum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For sum.ML. 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]); +val Inl_RepI = result(); + +goalw Sum.thy [Sum_def] "Inr_Rep(b) : Sum"; +by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]); +val Inr_RepI = result(); + +goal Sum.thy "inj_onto(Abs_Sum,Sum)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Sum_inverse 1); +val inj_onto_Abs_Sum = result(); + +(** 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]) ]); +val Inl_Rep_not_Inr_Rep = result(); + +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); +val Inl_not_Inr = result(); + +val Inl_neq_Inr = standard (Inl_not_Inr RS notE); +val Inr_neq_Inl = sym RS Inl_neq_Inr; + +(** 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); +val Inl_Rep_inject = result(); + +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); +val Inr_Rep_inject = result(); + +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); +val inj_Inl = result(); +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); +val inj_Inr = result(); +val Inr_inject = inj_Inr RS injD; + +goal Sum.thy "(Inl(x)=Inl(y)) = (x=y)"; +br iffI 1; +be (rewrite_rule [inj_def] Inl_inject) 1; +be ssubst 1; +br refl 1; +val Inl_inj = result(); + +goal Sum.thy "(Inr(x)=Inr(y)) = (x=y)"; +br iffI 1; +be (rewrite_rule [inj_def] Inr_inject) 1; +be ssubst 1; +br refl 1; +val Inr_inj = result(); + +(** case -- the selection operator for sums **) + +goalw Sum.thy [case_def] "case(Inl(x), f, g) = f(x)"; +by (fast_tac (set_cs addIs [select_equality] + addSEs [make_elim Inl_inject, Inl_neq_Inr]) 1); +val case_Inl = result(); + +goalw Sum.thy [case_def] "case(Inr(x), f, g) = g(x)"; +by (fast_tac (set_cs addIs [select_equality] + addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1); +val case_Inr = result(); + +(** 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)])); +val sumE = result(); + +goal Sum.thy "case(s, %x::'a. f(Inl(x)), %y::'b. f(Inr(y))) = f(s)"; +by (EVERY1 [res_inst_tac [("s","s")] sumE, + etac ssubst, rtac case_Inl, + etac ssubst, rtac case_Inr]); +val surjective_sum = result(); + +goal Sum.thy "R(case(s,f,g)) = \ +\ ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))"; +by (rtac sumE 1); +by (etac ssubst 1); +by (stac case_Inl 1); +by (fast_tac (set_cs addSEs [make_elim Inl_inject, Inl_neq_Inr]) 1); +by (etac ssubst 1); +by (stac case_Inr 1); +by (fast_tac (set_cs addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1); +val expand_case = result(); + +(*FIXME: case's congruence rule only should simplifies the first argument*) +val sum_ss = pair_ss addsimps [case_Inl, case_Inr]; diff -r 000000000000 -r 7949f97df77a Sum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Sum.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,34 @@ +(* Title: HOL/sum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The disjoint sum of two types +*) + +Sum = Prod + +types "+" 2 (infixl 10) +arities "+" :: (term,term)term +consts + Inl_Rep :: "['a,'a,'b,bool] => bool" + Inr_Rep :: "['b,'a,'b,bool] => bool" + Sum :: "(['a,'b,bool] => bool)set" + Rep_Sum :: "'a + 'b => (['a,'b,bool] => bool)" + Abs_Sum :: "(['a,'b,bool] => bool) => 'a+'b" + Inl :: "'a => 'a+'b" + Inr :: "'b => 'a+'b" + case :: "['a+'b, 'a=>'c,'b=>'c] =>'c" +rules + Inl_Rep_def "Inl_Rep == (%a. %x y p. x=a & p)" + Inr_Rep_def "Inr_Rep == (%b. %x y p. y=b & ~p)" + Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}" + (*faking a type definition...*) + Rep_Sum "Rep_Sum(s): Sum" + Rep_Sum_inverse "Abs_Sum(Rep_Sum(s)) = s" + Abs_Sum_inverse "f: Sum ==> Rep_Sum(Abs_Sum(f)) = f" + (*defining the abstract constants*) + Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))" + Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))" + case_def "case == (%p f g. @z. (!x. p=Inl(x) --> z=f(x))\ +\ & (!y. p=Inr(y) --> z=g(y)))" +end diff -r 000000000000 -r 7949f97df77a Trancl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Trancl.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,240 @@ +(* 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)); +val transI = result(); + +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); +val transD = result(); + +(** Identity relation **) + +goalw Trancl.thy [id_def] " : id"; +by (rtac CollectI 1); +by (rtac exI 1); +by (rtac refl 1); +val idI = result(); + +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); +val idE = result(); + +(** Composition of two relations **) + +val prems = goalw Trancl.thy [comp_def] + "[| :s; :r |] ==> : r O s"; +by (fast_tac (set_cs addIs prems) 1); +val compI = result(); + +(*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)); +val compE = result(); + +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)); +val compEpair = result(); + +val comp_cs = set_cs addIs [compI,idI] + addSEs [compE,idE,Pair_inject]; + +val prems = goal Trancl.thy + "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +by (cut_facts_tac prems 1); +by (fast_tac comp_cs 1); +val comp_mono = result(); + +val prems = goal Trancl.thy + "[| s <= Sigma(A,%x.B); r <= Sigma(B,%x.C) |] ==> \ +\ (r O s) <= Sigma(A,%x.C)"; +by (cut_facts_tac prems 1); +by (fast_tac (comp_cs addIs [SigmaI] addSEs [SigmaE2]) 1); +val comp_subset_Sigma = result(); + + +(** 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)); +val rtrancl_fun_mono = result(); + +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); +val rtrancl_refl = result(); + +(*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); +val rtrancl_into_rtrancl = result(); + +(*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); +val r_into_rtrancl = result(); + +(*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)); +val rtrancl_mono = result(); + +(** standard induction rule **) + +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ !!x. P(); \ +\ !!x y z.[| P(); : r^*; : r |] ==> P() |] \ +\ ==> P()"; +by (rtac (major RS (rtrancl_def RS def_induct)) 1); +by (rtac rtrancl_fun_mono 1); +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_full_induct = result(); + +(*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); +val rtrancl_induct = result(); + +(*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)); +val trans_rtrancl = result(); + +(*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)); +val rtranclE = result(); + + +(**** 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)); +val trancl_into_rtrancl = result(); + +(*r^+ contains r*) +val [prem] = goalw Trancl.thy [trancl_def] + "[| : r |] ==> : r^+"; +by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); +val r_into_trancl = result(); + +(*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)); +val rtrancl_into_trancl1 = result(); + +(*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)); +val rtrancl_into_trancl2 = result(); + +(*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); +val tranclE = result(); + +(*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)); +val trans_trancl = result(); + +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); +val trancl_into_trancl2 = result(); + + +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); +val trancl_subset_Sigma_lemma = result(); + +val prems = goalw Trancl.thy [trancl_def] + "r <= Sigma(A,%x.A) ==> trancl(r) <= Sigma(A,%x.A)"; +by (cut_facts_tac prems 1); +by (fast_tac (comp_cs addIs [SigmaI] + addSDs [trancl_subset_Sigma_lemma] + addSEs [SigmaE2]) 1); +val trancl_subset_Sigma = result(); + diff -r 000000000000 -r 7949f97df77a Trancl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Trancl.thy Thu Sep 16 12:21:07 1993 +0200 @@ -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 + +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) +rules +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 000000000000 -r 7949f97df77a Univ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Univ.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,635 @@ +(* 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); +val Least_equality = result(); + +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); +val LeastI = result(); + +(*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); +val Least_le = result(); + +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); +val not_less_Least = result(); + + +(** apfst -- can be used in similar type definitions **) + +goalw Univ.thy [apfst_def] "apfst(f,) = "; +by (rtac split 1); +val apfst = result(); + +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); +val apfstE = result(); + +(** 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); +val Push_inject1 = result(); + +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); +val Push_inject2 = result(); + +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); +val Push_inject = result(); + +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); +val Push_neq_K0 = result(); + +(*** Isomorphisms ***) + +goal Univ.thy "inj(Rep_Node)"; +by (rtac inj_inverseI 1); (*cannot combine by RS: multiple unifiers*) +by (rtac Rep_Node_inverse 1); +val inj_Rep_Node = result(); + +goal Univ.thy "inj_onto(Abs_Node,Node)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Node_inverse 1); +val inj_onto_Abs_Node = result(); + +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); +val Node_K0_I = result(); + +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); +val Node_Push_I = result(); + + +(*** 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)); +val Scons_not_Atom = result(); +val Atom_not_Scons = standard (Scons_not_Atom RS not_sym); + +val Scons_neq_Atom = standard (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)); +val inj_Atom = result(); +val Atom_inject = inj_Atom RS injD; + +goalw Univ.thy [Leaf_def] "inj(Leaf)"; +by (stac o_def 1); +by (rtac injI 1); +by (etac (Atom_inject RS Inl_inject) 1); +val inj_Leaf = result(); + +val Leaf_inject = inj_Leaf RS injD; + +goalw Univ.thy [Numb_def] "inj(Numb)"; +by (stac o_def 1); +by (rtac injI 1); +by (etac (Atom_inject RS Inr_inject) 1); +val inj_Numb = result(); + +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])); +val Push_Node_inject = result(); + + +(** 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); +val Scons_inject_lemma1 = result(); + +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); +val Scons_inject_lemma2 = result(); + +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)); +val Scons_inject1 = result(); + +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)); +val Scons_inject2 = result(); + +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); +val Scons_inject = result(); + +(*rewrite rules*) +goal Univ.thy "(Atom(a)=Atom(b)) = (a=b)"; +by (fast_tac (HOL_cs addSEs [Atom_inject]) 1); +val Atom_Atom_eq = result(); + +goal Univ.thy "(M.N = M'.N') = (M=M' & N=N')"; +by (fast_tac (HOL_cs addSEs [Scons_inject]) 1); +val Scons_Scons_eq = result(); + +(*** Distinctness involving Leaf and Numb ***) + +(** Scons vs Leaf **) + +goalw Univ.thy [Leaf_def] "~ ((M.N) = Leaf(a))"; +by (stac o_def 1); +by (rtac Scons_not_Atom 1); +val Scons_not_Leaf = result(); +val Leaf_not_Scons = standard (Scons_not_Leaf RS not_sym); + +val Scons_neq_Leaf = standard (Scons_not_Leaf RS notE); +val Leaf_neq_Scons = sym RS Scons_neq_Leaf; + +(** Scons vs Numb **) + +goalw Univ.thy [Numb_def] "~ ((M.N) = Numb(k))"; +by (stac o_def 1); +by (rtac Scons_not_Atom 1); +val Scons_not_Numb = result(); +val Numb_not_Scons = standard (Scons_not_Numb RS not_sym); + +val Scons_neq_Numb = standard (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); +val Leaf_not_Numb = result(); +val Numb_not_Leaf = standard (Leaf_not_Numb RS not_sym); + +val Leaf_neq_Numb = standard (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); +val ndepth_K0 = result(); + +goal Univ.thy "k < Suc(LEAST x. f(x)=0) --> ~ nat_case(k, Suc(i), f) = 0"; +by (nat_ind_tac "k" 1); +by (ALLGOALS (simp_tac nat_ss)); +by (rtac impI 1); +by (etac not_less_Least 1); +val ndepth_Push_lemma = result(); + +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); +val ndepth_Push_Node = result(); + + +(*** 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])); +val ntrunc_0 = result(); + +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); +val ntrunc_Atom = result(); + +goalw Univ.thy [Leaf_def] "ntrunc(Suc(k), Leaf(a)) = Leaf(a)"; +by (stac o_def 1); +by (rtac ntrunc_Atom 1); +val ntrunc_Leaf = result(); + +goalw Univ.thy [Numb_def] "ntrunc(Suc(k), Numb(i)) = Numb(i)"; +by (stac o_def 1); +by (rtac ntrunc_Atom 1); +val ntrunc_Numb = result(); + +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)); +val ntrunc_Scons = result(); + +(** 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])); +val ntrunc_one_In0 = result(); + +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); +val ntrunc_In0 = result(); + +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])); +val ntrunc_one_In1 = result(); + +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); +val ntrunc_In1 = result(); + + +(*** 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)); +val uprodI = result(); + +(*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)); +val uprodE = result(); + +(*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)); +val uprodE2 = result(); + + +(*** Disjoint Sum ***) + +goalw Univ.thy [usum_def] "!!M. M:A ==> In0(M) : A<+>B"; +by (fast_tac set_cs 1); +val usum_In0I = result(); + +goalw Univ.thy [usum_def] "!!N. N:B ==> In1(N) : A<+>B"; +by (fast_tac set_cs 1); +val usum_In1I = result(); + +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)); +val usumE = result(); + + +(** 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); +val In0_not_In1 = result(); + +val In1_not_In0 = standard (In0_not_In1 RS not_sym); +val In0_neq_In1 = standard (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); +val In0_inject = result(); + +val [major] = goalw Univ.thy [In1_def] "In1(M) = In1(N) ==> M=N"; +by (rtac (major RS Scons_inject2) 1); +val In1_inject = result(); + + +(*** proving equality of sets and functions using ntrunc ***) + +goalw Univ.thy [ntrunc_def] "ntrunc(k,M) <= M"; +by (fast_tac set_cs 1); +val ntrunc_subsetI = result(); + +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); +val ntrunc_subsetD = result(); + +(*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); +val ntrunc_equality = result(); + +val [major] = goal Univ.thy + "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"; +by (rtac (ntrunc_equality RS ext) 1); +by (resolve_tac ([major RS fun_cong] RL [o_def RS subst]) 1); +val ntrunc_o_equality = result(); + +(*** Monotonicity ***) + +goalw Univ.thy [uprod_def] "!!A B. [| A<=A'; B<=B' |] ==> A<*>B <= A'<*>B'"; +by (fast_tac set_cs 1); +val uprod_mono = result(); + +goalw Univ.thy [usum_def] "!!A B. [| A<=A'; B<=B' |] ==> A<+>B <= A'<+>B'"; +by (fast_tac set_cs 1); +val usum_mono = result(); + +goalw Univ.thy [Scons_def] "!!M N. [| M<=M'; N<=N' |] ==> M.N <= M'.N'"; +by (fast_tac set_cs 1); +val Scons_mono = result(); + +goalw Univ.thy [In0_def] "!!M N. M<=N ==> In0(M) <= In0(N)"; +by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); +val In0_mono = result(); + +goalw Univ.thy [In1_def] "!!M N. M<=N ==> In1(M) <= In1(N)"; +by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); +val In1_mono = result(); + + +(*** Split and Case ***) + +goalw Univ.thy [Split_def] "Split(M.N, c) = c(M,N)"; +by (fast_tac (set_cs addIs [select_equality] addEs [Scons_inject]) 1); +val Split = result(); + +goalw Univ.thy [Case_def] "Case(In0(M), c, d) = c(M)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim In0_inject, In0_neq_In1]) 1); +val Case_In0 = result(); + +goalw Univ.thy [Case_def] "Case(In1(N), c, d) = d(N)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim In1_inject, In1_neq_In0]) 1); +val Case_In1 = result(); + +(**** 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); +val ntrunc_UN1 = result(); + +goalw Univ.thy [Scons_def] "(UN x.f(x)) . M = (UN x. f(x) . M)"; +by (fast_tac (set_cs addIs [equalityI]) 1); +val Scons_UN1_x = result(); + +goalw Univ.thy [Scons_def] "M . (UN x.f(x)) = (UN x. M . f(x))"; +by (fast_tac (set_cs addIs [equalityI]) 1); +val Scons_UN1_y = result(); + +goalw Univ.thy [In0_def] "In0(UN x.f(x)) = (UN x. In0(f(x)))"; +br Scons_UN1_y 1; +val In0_UN1 = result(); + +goalw Univ.thy [In1_def] "In1(UN x.f(x)) = (UN x. In1(f(x)))"; +br Scons_UN1_y 1; +val In1_UN1 = result(); + + +(*** Equality : the diagonal relation ***) + +goalw Univ.thy [diag_def] "!!a A. a:A ==> : diag(A)"; +by (REPEAT (ares_tac [singletonI,UN_I] 1)); +val diagI = result(); + +(*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)); +val diagE = result(); + +(*** Equality for Cartesian Product ***) + +goal Univ.thy + "split(, %x x'. split(, %y y'. {})) = {}"; +by (simp_tac univ_ss 1); +val dprod_lemma = result(); + +goalw Univ.thy [dprod_def] + "!!r s. [| :r; :s |] ==> : r<**>s"; +by (REPEAT (ares_tac [UN_I] 1)); +by (rtac (singletonI RS (dprod_lemma RS equalityD2 RS subsetD)) 1); +val dprodI = result(); + +(*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 (eresolve_tac [asm_rl,singletonE,UN_E] 1)); +by (res_inst_tac [("p","u")] PairE 1); +by (res_inst_tac [("p","v")] PairE 1); +by (safe_tac HOL_cs); +by (REPEAT (ares_tac prems 1)); +by (safe_tac (set_cs addSDs [dprod_lemma RS equalityD1 RS subsetD])); +val dprodE = result(); + + +(*** Equality for Disjoint Sum ***) + +goalw Univ.thy [dsum_def] "!!r. :r ==> : r<++>s"; +by (fast_tac (set_cs addSIs [split RS equalityD2 RS subsetD]) 1); +val dsum_In0I = result(); + +goalw Univ.thy [dsum_def] "!!r. :s ==> : r<++>s"; +by (fast_tac (set_cs addSIs [split RS equalityD2 RS subsetD]) 1); +val dsum_In1I = result(); + +val major::prems = goalw Univ.thy [dsum_def] + "[| w : r<++>s; \ +\ !!x x'. [| : r; w = |] ==> P; \ +\ !!y y'. [| : s; w = |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (safe_tac set_cs); +by (res_inst_tac [("p","u")] PairE 1); +by (res_inst_tac [("p","v")] PairE 2); +by (safe_tac (set_cs addSEs prems + addSDs [split RS equalityD1 RS subsetD])); +val dsumE = result(); + + +(*** Monotonicity ***) + +goalw Univ.thy [dprod_def] "!!r s. [| r<=r'; s<=s' |] ==> r<**>s <= r'<**>s'"; +by (fast_tac set_cs 1); +val dprod_mono = result(); + +goalw Univ.thy [dsum_def] "!!r s. [| r<=r'; s<=s' |] ==> r<++>s <= r'<++>s'"; +by (fast_tac set_cs 1); +val dsum_mono = result(); + + +(*** Bounding theorems ***) + +goal Univ.thy "diag(A) <= Sigma(A,%x.A)"; +by (fast_tac (set_cs addIs [SigmaI] addSEs [diagE]) 1); +val diag_subset_Sigma = result(); + +val prems = goal Univ.thy + "[| r <= Sigma(A,%x.B); s <= Sigma(C,%x.D) |] ==> \ +\ (r<**>s) <= Sigma(A<*>C, %z. B<*>D)"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSIs [SigmaI,uprodI] + addSEs [dprodE,SigmaE2]) 1); +val dprod_subset_Sigma = result(); + +goal Univ.thy + "!!r s. [| r <= Sigma(A,B); s <= Sigma(C,D) |] ==> \ +\ (r<**>s) <= Sigma(A<*>C, %z. Split(z, %x y. B(x)<*>D(y)))"; +by (safe_tac (set_cs addSIs [SigmaI,uprodI] addSEs [dprodE])); +by (stac Split 3); +by (ALLGOALS (fast_tac (set_cs addSIs [uprodI] addSEs [SigmaE2]))); +val dprod_subset_Sigma2 = result(); + +goal Univ.thy + "!!r s. [| r <= Sigma(A,%x.B); s <= Sigma(C,%x.D) |] ==> \ +\ (r<++>s) <= Sigma(A<+>C, %z. B<+>D)"; +by (fast_tac (set_cs addSIs [SigmaI,usum_In0I,usum_In1I] + addSEs [dsumE,SigmaE2]) 1); +val dsum_subset_Sigma = result(); + + +(*** Domain ***) + +goal Univ.thy "fst `` diag(A) = A"; +by (fast_tac (set_cs addIs [equalityI, fst_imageI, diagI] + addSEs [fst_imageE, Pair_inject, diagE]) 1); +val fst_image_diag = result(); + +goal Univ.thy "fst `` (r<**>s) = (fst``r) <*> (fst``s)"; +by (fast_tac (set_cs addIs [equalityI, fst_imageI, uprodI, dprodI] + addSEs [fst_imageE, Pair_inject, uprodE, dprodE]) 1); +val fst_image_dprod = result(); + +goal Univ.thy "fst `` (r<++>s) = (fst``r) <+> (fst``s)"; +by (fast_tac (set_cs addIs [equalityI, fst_imageI, usum_In0I, usum_In1I, + dsum_In0I, dsum_In1I] + addSEs [fst_imageE, Pair_inject, usumE, dsumE]) 1); +val fst_image_dsum = result(); + +val fst_image_simps = [fst_image_diag, fst_image_dprod, fst_image_dsum]; +val fst_image_ss = univ_ss addsimps fst_image_simps; + +val univ_cs = + set_cs addSIs [SigmaI,uprodI,dprodI] + addIs [usum_In0I,usum_In1I,dsum_In0I,dsum_In1I] + addSEs [diagE,uprodE,dprodE,usumE,dsumE,SigmaE2,Pair_inject]; diff -r 000000000000 -r 7949f97df77a Univ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Univ.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,105 @@ +(* 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 + +types node 1 +arities node :: (term)term +consts + Least :: "(nat=>bool) => nat" (binder "LEAST " 10) + + apfst :: "['a=>'c, 'a*'b] => 'c*'b" + Push :: "[nat, nat=>nat] => (nat=>nat)" + + Node :: "((nat=>nat) * ('a+nat)) set" + Rep_Node :: "'a node => (nat=>nat) * ('a+nat)" + Abs_Node :: "(nat=>nat) * ('a+nat) => 'a node" + Push_Node :: "[nat, 'a node] => 'a node" + ndepth :: "'a node => nat" + + Atom :: "('a+nat) => 'a node set" + Leaf :: "'a => 'a node set" + Numb :: "nat => 'a node set" + "." :: "['a node set, 'a node set]=> 'a node set" (infixr 60) + In0,In1 :: "'a node set => 'a node set" + + ntrunc :: "[nat, 'a node set] => 'a node set" + + "<*>" :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 80) + "<+>" :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 70) + + Split :: "['a node set, ['a node set, 'a node set]=>'b] => 'b" + Case :: "['a node set, ['a node set]=>'b, ['a node set]=>'b] => 'b" + + diag :: "'a set => ('a * 'a)set" + "<**>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \ +\ => ('a node set * 'a node set)set" (infixr 80) + "<++>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \ +\ => ('a node set * 'a node set)set" (infixr 70) + +rules + + (*least number operator*) + Least_def "Least(P) == @k. P(k) & (ALL j. j ~P(j))" + + (** lists, trees will be sets of nodes **) + Node_def "Node == {p. EX f x k. p = & f(k)=0}" + (*faking the type definition 'a node == (nat=>nat) * ('a+nat) *) + Rep_Node "Rep_Node(n): Node" + Rep_Node_inverse "Abs_Node(Rep_Node(n)) = n" + Abs_Node_inverse "p: Node ==> Rep_Node(Abs_Node(p)) = p" + 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 p.split(p, %x y. ))" + Push_def "Push == (%b h n. nat_case(n,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(Rep_Node(n), %f x. LEAST k. f(k)=0)" + 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(M,c) == @u. ? x y. M = x.y & u = c(x,y)" + + Case_def "Case(M,c,d) == @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. UN v:s. \ +\ split(u, %x x'. split(v, %y y'. {}))" + + dsum_def "r<++>s == (UN u:r. split(u, %x x'. {})) Un \ +\ (UN v:s. split(v, %y y'. {}))" + +end diff -r 000000000000 -r 7949f97df77a WF.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/WF.ML Thu Sep 16 12:21:07 1993 +0200 @@ -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); +val wfI = result(); + +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); +val wf_induct = result(); + +(*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); +val wf_anti_sym = result(); + +val prems = goal WF.thy "[| wf(r); : r |] ==> P"; +by (rtac wf_anti_sym 1); +by (REPEAT (resolve_tac prems 1)); +val wf_anti_refl = result(); + +(*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); +val wf_trancl = result(); + + +(** 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); +val cut_cut_eq = result(); + +goalw WF.thy [cut_def] "!!x. :r ==> cut(f,r,a)(x) = f(x)"; +by(asm_simp_tac HOL_ss 1); +val cut_apply = result(); + + +(*** 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); +val is_recfun_undef = result(); + +(*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); +val is_recfun_equal_lemma = result(); +val is_recfun_equal = standard (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); +val is_recfun_cut = result(); + +(*** 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); +val is_the_recfun = result(); + +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); +val unfold_the_recfun = result(); + + +(*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)); +val the_recfun_equal = result(); + +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)); +val the_recfun_cut = result(); + +(*** 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); +val wftrec = result(); + +(*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)); +val wftrec2 = result(); + +(** 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); +val wfrec = result(); + +(*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)); +val def_wfrec = result(); diff -r 000000000000 -r 7949f97df77a WF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/WF.thy Thu Sep 16 12:21:07 1993 +0200 @@ -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" + +rules + 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 000000000000 -r 7949f97df77a arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/arith.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,229 @@ +(* Title: HOL/arith.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For HOL/arith.thy. + +Proofs about elementary arithmetic: addition, multiplication, etc. +Tests definitions and simplifier. +*) + +open Arith; + +(*** Basic rewrite and congruence rules for the arithmetic operators ***) + +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; + +val diff_0_eq_0 = prove_goalw Arith.thy [diff_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) *) +val diff_Suc_Suc = prove_goalw Arith.thy [diff_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 = [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 ***) + +val add_0_right = prove_goal Arith.thy "m + 0 = m" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +val add_Suc_right = prove_goal 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*) +val add_assoc = prove_goal 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. Must simplify after first induction! + Orientation of rewrites is delicate*) +val add_commute = prove_goal Arith.thy "m + n = n + m::nat" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*** Multiplication ***) + +(*right annihilation in product*) +val mult_0_right = prove_goal Arith.thy "m * 0 = 0" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]); + +(*right Sucessor law for multiplication*) +val mult_Suc_right = prove_goal Arith.thy + "m * Suc(n) = m + (m * n)" + (fn _ => + [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_assoc RS sym])), + (*The final goal requires the commutative law for addition*) + stac add_commute 1, rtac refl 1]); + +val arith_ss = arith_ss addsimps [mult_0_right,mult_Suc_right]; + +(*Commutative law for multiplication*) +val mult_commute = prove_goal Arith.thy "m * n = n * m::nat" + (fn _ => [nat_ind_tac "m" 1, ALLGOALS (asm_simp_tac arith_ss)]); + +(*addition distributes over multiplication*) +val add_mult_distrib = prove_goal Arith.thy "(m + n)*k = (m*k) + (n*k)::nat" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_assoc]))]); + +(*Associative law for multiplication*) +val mult_assoc = prove_goal Arith.thy "(m * n) * k = m * (n * k)::nat" + (fn _ => [nat_ind_tac "m" 1, + ALLGOALS(asm_simp_tac(arith_ss addsimps [add_mult_distrib]))]); + + +(*** Difference ***) + +val diff_self_eq_0 = prove_goal 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)); +val add_diff_inverse = result(); + + +(*** 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)); +val diff_less_Suc = result(); + +(*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]))); +val div_termination = result(); + +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); +val mod_less = result(); + +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); +val mod_geq = result(); + + +(*** 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); +val div_less = result(); + +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); +val div_geq = result(); + +(*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 (res_inst_tac [("Q","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)); +val less_imp_diff_is_0 = result(); + +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))); +val diffs0_imp_equal_lemma = result(); + +(* [| m-n = 0; n-m = 0 |] ==> m=n *) +val diffs0_imp_equal = standard (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)); +val Suc_diff_n = result(); + +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' fast_tac HOL_cs)); +val zero_induct_lemma = result(); + +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)); +val zero_induct = result(); + +(*13 July 1992: loaded in 105.7s*) + +(**** Additional theorems about "less than" ****) + +goal Arith.thy "n <= (m + n)::nat"; +by (nat_ind_tac "m" 1); +by (ALLGOALS(simp_tac (arith_ss addsimps [le_refl]))); +by (etac le_trans 1); +by (rtac (lessI RS less_imp_le) 1); +val le_add2 = result(); + +goal Arith.thy "m <= (m + n)::nat"; +by (stac add_commute 1); +by (rtac le_add2 1); +val le_add1 = result(); + +val less_add_Suc1 = standard (lessI RS (le_add1 RS le_less_trans)); +val less_add_Suc2 = standard (lessI RS (le_add2 RS le_less_trans)); diff -r 000000000000 -r 7949f97df77a arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/arith.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,26 @@ +(* 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 + +arities nat::plus + nat::minus + nat::times +consts + div,mod :: "[nat,nat]=>nat" (infixl 70) +rules + add_def "m+n == nat_rec(m, n, %u v.Suc(v))" + diff_def "m-n == nat_rec(n, m, %u v. nat_rec(v, 0, %x y.x))" + 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 000000000000 -r 7949f97df77a equalities.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/equalities.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,271 @@ +(* Title: HOL/equalities + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Equalities involving union, intersection, inclusion, etc. +*) + +writeln"File HOL/equalities"; + +val eq_cs = set_cs addSIs [equalityI]; + +(** insert **) + +goal Set.thy "!!a. a:A ==> insert(a,A) = A"; +by (fast_tac eq_cs 1); +val insert_absorb = result(); + +(** Image **) + +goal Set.thy "f``{} = {}"; +by (fast_tac eq_cs 1); +val image_empty = result(); + +goal Set.thy "f``insert(a,B) = insert(f(a), f``B)"; +by (fast_tac eq_cs 1); +val image_insert = result(); + +(** Binary Intersection **) + +goal Set.thy "A Int A = A"; +by (fast_tac eq_cs 1); +val Int_absorb = result(); + +goal Set.thy "A Int B = B Int A"; +by (fast_tac eq_cs 1); +val Int_commute = result(); + +goal Set.thy "(A Int B) Int C = A Int (B Int C)"; +by (fast_tac eq_cs 1); +val Int_assoc = result(); + +goal Set.thy "{} Int B = {}"; +by (fast_tac eq_cs 1); +val Int_empty_left = result(); + +goal Set.thy "A Int {} = {}"; +by (fast_tac eq_cs 1); +val Int_empty_right = result(); + +goal Set.thy "(A Un B) Int C = (A Int C) Un (B Int C)"; +by (fast_tac eq_cs 1); +val Int_Un_distrib = result(); + +goal Set.thy "(A<=B) = (A Int B = A)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val subset_Int_eq = result(); + +(** Binary Union **) + +goal Set.thy "A Un A = A"; +by (fast_tac eq_cs 1); +val Un_absorb = result(); + +goal Set.thy "A Un B = B Un A"; +by (fast_tac eq_cs 1); +val Un_commute = result(); + +goal Set.thy "(A Un B) Un C = A Un (B Un C)"; +by (fast_tac eq_cs 1); +val Un_assoc = result(); + +goal Set.thy "{} Un B = B"; +by(fast_tac eq_cs 1); +val Un_empty_left = result(); + +goal Set.thy "A Un {} = A"; +by(fast_tac eq_cs 1); +val Un_empty_right = result(); + +goal Set.thy "insert(a,B) Un C = insert(a,B Un C)"; +by(fast_tac eq_cs 1); +val Un_insert_left = result(); + +goal Set.thy "(A Int B) Un C = (A Un C) Int (B Un C)"; +by (fast_tac eq_cs 1); +val Un_Int_distrib = result(); + +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); +val Un_Int_crazy = result(); + +goal Set.thy "(A<=B) = (A Un B = B)"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val subset_Un_eq = result(); + +goal Set.thy "(A <= insert(b,C)) = (A <= C | b:A & A-{b} <= C)"; +by (fast_tac eq_cs 1); +val subset_insert_iff = result(); + +(** Simple properties of Compl -- complement of a set **) + +goal Set.thy "A Int Compl(A) = {}"; +by (fast_tac eq_cs 1); +val Compl_disjoint = result(); + +goal Set.thy "A Un Compl(A) = {x.True}"; +by (fast_tac eq_cs 1); +val Compl_partition = result(); + +goal Set.thy "Compl(Compl(A)) = A"; +by (fast_tac eq_cs 1); +val double_complement = result(); + +goal Set.thy "Compl(A Un B) = Compl(A) Int Compl(B)"; +by (fast_tac eq_cs 1); +val Compl_Un = result(); + +goal Set.thy "Compl(A Int B) = Compl(A) Un Compl(B)"; +by (fast_tac eq_cs 1); +val Compl_Int = result(); + +goal Set.thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"; +by (fast_tac eq_cs 1); +val Compl_UN = result(); + +goal Set.thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"; +by (fast_tac eq_cs 1); +val Compl_INT = result(); + +(*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); +val Un_Int_assoc_eq = result(); + + +(** Big Union and Intersection **) + +goal Set.thy "Union({}) = {}"; +by (fast_tac eq_cs 1); +val Union_empty = result(); + +goal Set.thy "Union(insert(a,B)) = a Un Union(B)"; +by (fast_tac eq_cs 1); +val Union_insert = result(); + +goal Set.thy "Union(A Un B) = Union(A) Un Union(B)"; +by (fast_tac eq_cs 1); +val Union_Un_distrib = result(); + +goal Set.thy "Union(A Un B) = Union(A) Un Union(B)"; +by (fast_tac eq_cs 1); +val Union_Un_distrib = result(); + +val prems = goal Set.thy + "(Union(C) Int A = {}) = (! B:C. B Int A = {})"; +by (fast_tac (eq_cs addSEs [equalityE]) 1); +val Union_disjoint = result(); + +goal Set.thy "Inter(A Un B) = Inter(A) Int Inter(B)"; +by (best_tac eq_cs 1); +val Inter_Un_distrib = result(); + +(** Unions and Intersections of Families **) + +goal Set.thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})"; +by (fast_tac eq_cs 1); +val UN_eq = result(); + +(*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); +val INT_eq = result(); + +goal Set.thy "A Int Union(B) = (UN C:B. A Int C)"; +by (fast_tac eq_cs 1); +val Int_Union_image = result(); + +(* Devlin, page 12: 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); +val Un_Union_image = result(); + +goal Set.thy "A Un Inter(B) = (INT C:B. A Un C)"; +by (fast_tac eq_cs 1); +val Un_Inter_image = result(); + +goal Set.thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)"; +by (best_tac eq_cs 1); +val Int_Inter_image = result(); + +(** Other identities about Unions and Intersections **) + +goal Set.thy "Union(range(f)) = (UN x.f(x))"; +by (fast_tac eq_cs 1); +val Union_range_eq = result(); + +goal Set.thy "Inter(range(f)) = (INT x.f(x))"; +by (fast_tac eq_cs 1); +val Inter_range_eq = result(); + +goal Set.thy "Union(B``A) = (UN x:A. B(x))"; +by (fast_tac eq_cs 1); +val Union_image_eq = result(); + +goal Set.thy "Inter(B``A) = (INT x:A. B(x))"; +by (fast_tac eq_cs 1); +val Inter_image_eq = result(); + +goal Set.thy "(UN x.B) = B"; +by (fast_tac eq_cs 1); +val constant_UN = result(); + +(** Simple properties of Diff -- set difference **) + +goal Set.thy "A-A = {}"; +by (fast_tac eq_cs 1); +val Diff_cancel = result(); + +goal Set.thy "{}-A = {}"; +by (fast_tac eq_cs 1); +val empty_Diff = result(); + +goal Set.thy "A-{} = A"; +by (fast_tac eq_cs 1); +val Diff_empty = result(); + +(*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); +val Diff_insert = result(); + +(*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); +val Diff_insert2 = result(); + +val prems = goal Set.thy "a:A ==> insert(a,A-{a}) = A"; +by (fast_tac (eq_cs addSIs prems) 1); +val insert_Diff = result(); + +goal Set.thy "A Int (B-A) = {}"; +by (fast_tac eq_cs 1); +val Diff_disjoint = result(); + +goal Set.thy "!!A. A<=B ==> A Un (B-A) = B"; +by (fast_tac eq_cs 1); +val Diff_partition = result(); +(* +goal Set.thy "!!A. [| A<=B; B<= C |] ==> (B - (C-A)) = A"; +by (cut_facts_tac prems 1); +by (fast_tac eq_cs 1); +val double_complement = result(); +*) +goal Set.thy "A - (B Un C) = (A-B) Int (A-C)"; +by (fast_tac eq_cs 1); +val Diff_Un = result(); + +goal Set.thy "A - (B Int C) = (A-B) Un (A-C)"; +by (fast_tac eq_cs 1); +val Diff_Int = result(); + +val set_ss = set_ss addsimps + [Int_absorb,Int_empty_left,Int_empty_right, + Un_absorb,Un_empty_left,Un_empty_right, + constant_UN,image_empty, + Compl_disjoint,double_complement, + Union_empty,Union_insert,empty_subsetI,subset_refl, + Diff_cancel,empty_Diff,Diff_empty,Diff_disjoint]; diff -r 000000000000 -r 7949f97df77a ex/Finite.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Finite.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,101 @@ +open Finite; + +(** Monotonicity and unfolding of the function **) + +goal Finite.thy "mono(%F. insert({}, UN x:A. insert(x)``F))"; +by (rtac monoI 1); +by (fast_tac set_cs 1); +val Fin_bnd_mono = result(); + +goalw Finite.thy [Fin_def] "!!A B. A<=B ==> Fin(A) <= Fin(B)"; +br lfp_mono 1; +by(fast_tac set_cs 1); +val Fin_mono = result(); + +goalw Finite.thy [Fin_def] "Fin(A) <= {B. B <= A}"; +by (fast_tac (set_cs addSIs [lfp_lowerbound]) 1); +val Fin_subset_Pow = result(); + +(* A : Fin(B) ==> A <= B *) +val FinD = Fin_subset_Pow RS subsetD RS CollectD; + +val Fin_unfold = Fin_bnd_mono RS (Fin_def RS def_lfp_Tarski); + +goal Finite.thy "{} : Fin(A)"; +by (rtac (Fin_unfold RS ssubst) 1); +by (rtac insertI1 1); +val Fin_0I = result(); + +goal Finite.thy "!!a. [| a:A; F:Fin(A) |] ==> insert(a,F) : Fin(A)"; +by (rtac (Fin_unfold RS ssubst) 1); +by(fast_tac set_cs 1); +val Fin_insertI = result(); + +(*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_def RS def_induct)) 1); +by (rtac Fin_bnd_mono 1); +by (safe_tac set_cs); +(*Excluded middle on x:y would be more straightforward; + actual proof is inspired by Dov Gabbay's Restart Rule*) +by (rtac classical 2); +by (REPEAT (ares_tac prems 1)); +by (etac contrapos 1); +by (rtac (insert_absorb RS ssubst) 1); +by (REPEAT (assume_tac 1)); +val Fin_induct = result(); + +(** Simplification for Fin **) + +val Fin_ss = set_ss addsimps [Fin_0I, Fin_insertI]; + +(*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])))); +val Fin_UnI = result(); + +(*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)); +val Fin_subset = result(); + +(*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); +val Fin_imageI = result(); + +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@[insert_subset_iff, + Diff_subset RS Fin_subset])))); +val Fin_empty_induct_lemma = result(); + +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)); +val Fin_empty_induct = result(); diff -r 000000000000 -r 7949f97df77a ex/Finite.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Finite.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,5 @@ +Finite = Lfp + +consts Fin :: "'a set => 'a set set" +rules + Fin_def "Fin(A) == lfp(%F. insert({}, UN x:A. insert(x)``F))" +end diff -r 000000000000 -r 7949f97df77a ex/LexProd.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/LexProd.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,24 @@ +(* Title: HOL/ex/lex-prod.ML + ID: $Id$ + Author: Tobias Nipkow, TU Munich + Copyright 1993 TU Munich + +For lex-prod.thy. +The lexicographic product of two wellfounded relations is again wellfounded. +*) + +val prems = goal Prod.thy "!a b. P() ==> !p.P(p)"; +by (cut_facts_tac prems 1); +by (rtac allI 1); +by (rtac (surjective_pairing RS ssubst) 1); +by (fast_tac HOL_cs 1); +val split_all_pair = result(); + +val [wfa,wfb] = goalw LexProd.thy [wf_def,LexProd.lex_prod_def] + "[| wf(ra); wf(rb) |] ==> wf(ra**rb)"; +by (EVERY1 [rtac allI,rtac impI, rtac split_all_pair]); +by (rtac (wfa RS spec RS mp) 1); +by (EVERY1 [rtac allI,rtac impI]); +by (rtac (wfb RS spec RS mp) 1); +by (fast_tac (set_cs addSEs [Pair_inject]) 1); +val wf_lex_prod = result(); diff -r 000000000000 -r 7949f97df77a ex/LexProd.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/LexProd.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,15 @@ +(* Title: HOL/ex/lex-prod.thy + ID: $Id$ + Author: Tobias Nipkow, TU Munich + Copyright 1993 TU Munich + +The lexicographic product of two relations. +*) + +LexProd = WF + Prod + +consts "**" :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set" (infixl 70) +rules +lex_prod_def "ra**rb == {p. ? a a' b b'. \ +\ p = <,> & ( : ra | a=a' & : rb)}" +end + diff -r 000000000000 -r 7949f97df77a ex/PL.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/PL.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,344 @@ +(* Title: HOL/ex/prop-log.ML + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1993 TU Muenchen & University of Cambridge + +For ex/prop-log.thy. Inductive definition of propositional logic. +Soundness and completeness w.r.t. truth-tables. + +Prove: If H|=p then G|=p where G:Fin(H) +*) + +open PL; + +val rule_defs = [axK_def, axS_def, axDN_def, ruleMP_def]; + + +(** Monotonicity and unfolding of the function **) + +goalw PL.thy rule_defs "mono(%X. H Un axK Un axS Un axDN Un ruleMP(X))"; +by (rtac monoI 1); +by(fast_tac set_cs 1); +val thms_bnd_mono = result(); + +goalw PL.thy [thms_def] "!!G H. G<=H ==> thms(G) <= thms(H)"; +by (REPEAT (ares_tac [subset_refl, Un_mono, lfp_mono] 1)); +val thms_mono = result(); + +(** Introduction rules for the consequence relation **) + +(* thms(H) = H Int Un axK Un axS Un ruleMP(thms(H)) *) +val thms_unfold = thms_bnd_mono RS (thms_def RS def_lfp_Tarski); + +(*Proof by hypothesis*) +val prems = goalw PL.thy [conseq_def] "p:H ==> H |- p"; +by (rtac (thms_unfold RS ssubst) 1); +by (fast_tac (set_cs addSIs prems) 1); +val conseq_H = result(); + +(*Proof by axiom K*) +goalw PL.thy [conseq_def] "H |- p->q->p"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac axK_def); +by (fast_tac set_cs 1); +val conseq_K = result(); + +(*Proof by axiom S*) +goalw PL.thy [conseq_def] "H |- (p->q->r) -> (p->q) -> p -> r"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac axS_def); +by (fast_tac set_cs 1); +val conseq_S = result(); + +(*Proof by axiom DN (double negation) *) +goalw PL.thy [conseq_def] "H |- ((p->false) -> false) -> p"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac axDN_def); +by (fast_tac set_cs 1); +val conseq_DN = result(); + +(*Proof by rule MP (Modus Ponens) *) +val [prempq,premp] = goalw PL.thy [conseq_def] + "[| H |- p->q; H |- p |] ==> H |- q"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac ruleMP_def); +by (fast_tac (set_cs addSIs [premp,prempq]) 1); +val conseq_MP = result(); + +(*Rule is called I for Identity Combinator, not for Introduction*) +goal PL.thy "H |- p->p"; +by (rtac (conseq_S RS conseq_MP RS conseq_MP) 1); +by (rtac conseq_K 2); +by (rtac conseq_K 1); +val conseq_I = result(); + +(** Weakening, left and right **) + +(*This order of premises is convenient with RS*) +val prems = goalw PL.thy [conseq_def] "[| G<=H; G |- p |] ==> H |- p"; +by (rtac (thms_mono RS subsetD) 1); +by (REPEAT (resolve_tac prems 1)); +val weaken_left = result(); + +(* H |- p ==> insert(a,H) |- p *) +val weaken_left_insert = subset_insertI RS weaken_left; + +val weaken_left_Un1 = Un_upper1 RS weaken_left; +val weaken_left_Un2 = Un_upper2 RS weaken_left; + +val prems = goal PL.thy "H |- q ==> H |- p->q"; +by (rtac (conseq_K RS conseq_MP) 1); +by (REPEAT (resolve_tac prems 1)); +val weaken_right = result(); + +(** Rule induction for H|-p **) + +(*Careful unfolding/folding to avoid a big expansion*) +val major::prems = goalw PL.thy [conseq_def] + "[| H |- a; \ +\ !!x. x:H ==> P(x); \ +\ !!x y. P(x->y->x); \ +\ !!x y z. P((x->y->z)->(x->y)->x->z); \ +\ !!x. P(((x->false)->false)->x); \ +\ !!x y. [| H |- x->y; H |- x; P(x->y); P(x) |] ==> P(y) \ +\ |] ==> P(a)"; +by (rtac (major RS (thms_def RS def_induct)) 1); +by (rtac thms_bnd_mono 1); +by (rewrite_tac rule_defs); +by (fast_tac (set_cs addIs prems) 1); +val conseq_induct = result(); + +(*The deduction theorem*) +val [major] = goal PL.thy "insert(p,H) |- q ==> H |- p->q"; +by (rtac (major RS conseq_induct) 1); +by (fast_tac (set_cs addIs [conseq_I, conseq_H RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_K RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_S RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_DN RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_S RS conseq_MP RS conseq_MP]) 1); +val deduction = result(); + + +(*The cut rule*) +val prems = goal PL.thy "[| H|-p; insert(p,H) |- q |] ==> H |- q"; +by (rtac (deduction RS conseq_MP) 1); +by (REPEAT (resolve_tac prems 1)); +val cut = result(); + +val prems = goal PL.thy "H |- false ==> H |- p"; +by (rtac (conseq_DN RS conseq_MP) 1); +by (rtac weaken_right 1); +by (resolve_tac prems 1); +val conseq_falseE = result(); + +(* [| H |- p->false; H |- p; q: pl |] ==> H |- q *) +val conseq_notE = standard (conseq_MP RS conseq_falseE); + +(** The function eval **) + +val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp]; + +goalw PL.thy [eval_def] "tt[false] = False"; +by (simp_tac pl_ss 1); +val eval_false = result(); + +goalw PL.thy [eval_def] "tt[#v] = (v:tt)"; +by (simp_tac pl_ss 1); +val eval_var = result(); + +goalw PL.thy [eval_def] "tt[p->q] = (tt[p]-->tt[q])"; +by (simp_tac pl_ss 1); +val eval_imp = result(); + +val pl_ss = pl_ss addsimps [eval_false, eval_var, eval_imp]; + +(** The function hyps **) + +goalw PL.thy [hyps_def] "hyps(false,tt) = {}"; +by (simp_tac pl_ss 1); +val hyps_false = result(); + +goalw PL.thy [hyps_def] "hyps(#v,tt) = {if(v:tt, #v, #v->false)}"; +by (simp_tac pl_ss 1); +val hyps_var = result(); + +goalw PL.thy [hyps_def] "hyps(p->q,tt) = hyps(p,tt) Un hyps(q,tt)"; +by (simp_tac pl_ss 1); +val hyps_imp = result(); + +val pl_ss = pl_ss addsimps [hyps_false, hyps_var, hyps_imp]; + +val ball_eq = prove_goalw Set.thy [Ball_def] "(!x:A.P(x)) = (!x.x:A --> P(x))" + (fn _ => [rtac refl 1]); + +(*Soundness of the rules wrt truth-table semantics*) +val [major] = goalw PL.thy [sat_def] "H |- p ==> H |= p"; +by (rtac (major RS conseq_induct) 1); +by (fast_tac (set_cs addSDs [eval_imp RS iffD1 RS mp]) 5); +by (ALLGOALS (asm_simp_tac(pl_ss addsimps + [ball_eq,not_def RS fun_cong RS sym]))); +val soundness = result(); + +(** Structural induction on pl + +val major::prems = goalw PL.thy pl_defs + "[| q: pl; \ +\ P(false); \ +\ !!v. v:nat ==> P(#v); \ +\ !!q1 q2. [| q1: pl; q2: pl; P(q1); P(q2) |] ==> P(q1->q2) \ +\ |] ==> P(q)"; +by (rtac (major RS sexp_induct) 1); +by (etac nat_induct 1); +by (REPEAT (ares_tac prems 1)); +val pl_induct = result(); + **) +(*** Towards the completeness proof ***) + +val [premf] = goal PL.thy "H |- p->false ==> H |- p->q"; +by (rtac deduction 1); +by (rtac (premf RS weaken_left_insert RS conseq_notE) 1); +by (rtac conseq_H 1); +by (rtac insertI1 1); +val false_imp = result(); + +val [premp,premq] = goal PL.thy + "[| H |- p; H |- q->false |] ==> H |- (p->q)->false"; +by (rtac deduction 1); +by (rtac (premq RS weaken_left_insert RS conseq_MP) 1); +by (rtac (conseq_H RS conseq_MP) 1); +by (rtac insertI1 1); +by (rtac (premp RS weaken_left_insert) 1); +val imp_false = result(); + +(*This formulation is required for strong induction hypotheses*) +goal PL.thy "hyps(p,tt) |- if(tt[p], p, p->false)"; +by (rtac (expand_if RS iffD2) 1); +by(res_inst_tac[("x","p")]spec 1); +by (rtac pl_ind 1); +by (ALLGOALS (simp_tac (pl_ss addsimps [conseq_I, conseq_H]))); +by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2, + weaken_right, imp_false] + addSEs [false_imp]) 1); +val hyps_conseq_if = result(); + +(*Key lemma for completeness; yields a set of assumptions satisfying p*) +val [sat] = goalw PL.thy [sat_def] "{} |= p ==> hyps(p,tt) |- p"; +by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN + rtac hyps_conseq_if 2); +by (fast_tac set_cs 1); +val sat_conseq_p = result(); + +(*For proving certain theorems in our new propositional logic*) +val conseq_cs = + set_cs addSIs [deduction] addIs [conseq_H, conseq_H RS conseq_MP]; + +(*The excluded middle in the form of an elimination rule*) +goal PL.thy "H |- (p->q) -> ((p->false)->q) -> q"; +by (rtac (deduction RS deduction) 1); +by (rtac (conseq_DN RS conseq_MP) 1); +by (ALLGOALS (best_tac (conseq_cs addSIs prems))); +val conseq_excluded_middle = result(); + +(*Hard to prove directly because it requires cuts*) +val prems = goal PL.thy + "[| insert(p,H) |- q; insert(p->false,H) |- q |] ==> H |- q"; +by (rtac (conseq_excluded_middle RS conseq_MP RS conseq_MP) 1); +by (REPEAT (resolve_tac (prems@[deduction]) 1)); +val conseq_excluded_middle_rule = result(); + +(*** Completeness -- lemmas for reducing the set of assumptions ***) + +(*For the case hyps(p,t)-insert(#v,Y) |- p; + we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) +goal PL.thy "!p.hyps(p, t-{v}) <= insert(#v->false, hyps(p,t)-{#v})"; +by (rtac pl_ind 1); +by (simp_tac pl_ss 1); +by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); +by (fast_tac (set_cs addSEs [sym RS var_neq_imp] addSDs [var_inject]) 1); +by (simp_tac pl_ss 1); +by (fast_tac set_cs 1); +val hyps_Diff = result() RS spec; + +(*For the case hyps(p,t)-insert(#v -> false,Y) |- p; + we also have hyps(p,t)-{#v->false} <= hyps(p, insert(v,t)) *) +goal PL.thy "!p.hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{#v->false})"; +by (rtac pl_ind 1); +by (simp_tac pl_ss 1); +by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); +by (fast_tac (set_cs addSEs [var_neq_imp, imp_inject] addSDs [var_inject]) 1); +by (simp_tac pl_ss 1); +by (fast_tac set_cs 1); +val hyps_insert = result() RS spec; + +(** Two lemmas for use with weaken_left **) + +goal Set.thy "B-C <= insert(a, B-insert(a,C))"; +by (fast_tac set_cs 1); +val insert_Diff_same = result(); + +goal Set.thy "insert(a, B-{c}) - D <= insert(a, B-insert(c,D))"; +by (fast_tac set_cs 1); +val insert_Diff_subset2 = result(); + +(*The set hyps(p,t) is finite, and elements have the form #v or #v->false; + could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*) +goal PL.thy "!p.hyps(p,t) : Fin(UN v:{x.True}. {#v, #v->false})"; +by (rtac pl_ind 1); +by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN' + fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI]))); +val hyps_finite = result() RS spec; + +val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; + +(*Induction on the finite set of assumptions hyps(p,t0). + We may repeatedly subtract assumptions until none are left!*) +val [sat] = goal PL.thy + "{} |= p ==> !t. hyps(p,t) - hyps(p,t0) |- p"; +by (rtac (hyps_finite RS Fin_induct) 1); +by (simp_tac (pl_ss addsimps [sat RS sat_conseq_p]) 1); +by (safe_tac set_cs); +(*Case hyps(p,t)-insert(#v,Y) |- p *) +by (rtac conseq_excluded_middle_rule 1); +by (rtac (insert_Diff_same RS weaken_left) 1); +by (etac spec 1); +by (rtac (insert_Diff_subset2 RS weaken_left) 1); +by (rtac (hyps_Diff RS Diff_weaken_left) 1); +by (etac spec 1); +(*Case hyps(p,t)-insert(#v -> false,Y) |- p *) +by (rtac conseq_excluded_middle_rule 1); +by (rtac (insert_Diff_same RS weaken_left) 2); +by (etac spec 2); +by (rtac (insert_Diff_subset2 RS weaken_left) 1); +by (rtac (hyps_insert RS Diff_weaken_left) 1); +by (etac spec 1); +val completeness_0_lemma = result(); + +(*The base case for completeness*) +val [sat] = goal PL.thy "{} |= p ==> {} |- p"; +by (rtac (Diff_cancel RS subst) 1); +by (rtac (sat RS (completeness_0_lemma RS spec)) 1); +val completeness_0 = result(); + +(*A semantic analogue of the Deduction Theorem*) +val [sat] = goalw PL.thy [sat_def] "insert(p,H) |= q ==> H |= p->q"; +by (simp_tac pl_ss 1); +by (cfast_tac [sat] 1); +val sat_imp = result(); + +val [finite] = goal PL.thy "H: Fin({p.True}) ==> !p. H |= p --> H |- p"; +by (rtac (finite RS Fin_induct) 1); +by (safe_tac (set_cs addSIs [completeness_0])); +by (rtac (weaken_left_insert RS conseq_MP) 1); +by (fast_tac (set_cs addSIs [sat_imp]) 1); +by (fast_tac conseq_cs 1); +val completeness_lemma = result(); + +val completeness = completeness_lemma RS spec RS mp; + +val [finite] = goal PL.thy "H: Fin({p.True}) ==> (H |- p) = (H |= p)"; +by (fast_tac (set_cs addSEs [soundness, finite RS completeness]) 1); +val conseq_iff = result(); + +writeln"Reached end of file."; + + diff -r 000000000000 -r 7949f97df77a ex/PL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/PL.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,55 @@ +(* Title: HOL/ex/prop-log + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Inductive definition of propositional logic. + +*) + +PL = Finite + +types pl 1 +arities pl :: (term)term +consts + false :: "'a pl" + "->" :: "['a pl,'a pl] => 'a pl" (infixr 90) + var :: "'a => 'a pl" ("#_") + pl_rec :: "['a pl,'a => 'b, 'b, ['b,'b] => 'b] => 'b" + axK,axS,axDN:: "'a pl set" + ruleMP,thms :: "'a pl set => 'a pl set" + "|-" :: "['a pl set, 'a pl] => bool" (infixl 50) + "|=" :: "['a pl set, 'a pl] => bool" (infixl 50) + eval :: "['a set, 'a pl] => bool" ("_[_]" [100,0] 100) + hyps :: "['a pl, 'a set] => 'a pl set" +rules + + (** Proof theory for propositional logic **) + + axK_def "axK == {x . ? p q. x = p->q->p}" + axS_def "axS == {x . ? p q r. x = (p->q->r) -> (p->q) -> p->r}" + axDN_def "axDN == {x . ? p. x = ((p->false) -> false) -> p}" + + (*the use of subsets simplifies the proof of monotonicity*) + ruleMP_def "ruleMP(X) == {q. ? p:X. p->q : X}" + + thms_def + "thms(H) == lfp(%X. H Un axK Un axS Un axDN Un ruleMP(X))" + + conseq_def "H |- p == p : thms(H)" + + sat_def "H |= p == (!tt. (!q:H. tt[q]) --> tt[p])" + +pl_rec_var "pl_rec(#v,f,y,z) = f(v)" +pl_rec_false "pl_rec(false,f,y,z) = y" +pl_rec_imp "pl_rec(p->q,f,y,g) = g(pl_rec(p,f,y,g),pl_rec(q,f,y,g))" + +eval_def "tt[p] == pl_rec(p, %v.v:tt, False, op -->)" + +hyps_def + "hyps(p,tt) == pl_rec(p, %a. {if(a:tt, #a, #a->false)}, {}, op Un)" + +var_inject "(#v = #w) ==> v = w" +imp_inject "[| (p -> q) = (p' -> q'); [| p = p'; q = q' |] ==> R |] ==> R" +var_neq_imp "(#v = (p -> q)) ==> R" +pl_ind "[| P(false); !!v. P(#v); !!p q. P(p)-->P(q)-->P(p->q)|] ==> !t.P(t)" +end diff -r 000000000000 -r 7949f97df77a ex/Puzzle.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Puzzle.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,58 @@ +(* Title: HOL/ex/puzzle.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 TU Muenchen + +For puzzle.thy. A question from "Bundeswettbewerb Mathematik" + +Proof due to Herbert Ehler +*) + +(*specialized form of induction needed below*) +val prems = goal Nat.thy "[| P(0); !!n. P(Suc(n)) |] ==> !n.P(n)"; +by (EVERY1 [rtac (nat_induct RS allI), resolve_tac prems, resolve_tac prems]); +val nat_exh = result(); + +goal Puzzle.thy "! n. k=f(n) --> n <= f(n)"; +by (res_inst_tac [("n","k")] less_induct 1); +by (rtac nat_exh 1); +by (simp_tac nat_ss 1); +by (rtac impI 1); +by (rtac classical 1); +by (dtac not_leE 1); +by (subgoal_tac "f(na) <= f(f(na))" 1); +by (best_tac (HOL_cs addIs [lessD,Puzzle.f_ax,le_less_trans,le_trans]) 1); +by (fast_tac (HOL_cs addIs [Puzzle.f_ax]) 1); +val lemma = result() RS spec RS mp; + +goal Puzzle.thy "n <= f(n)"; +by (fast_tac (HOL_cs addIs [lemma]) 1); +val lemma1 = result(); + +goal Puzzle.thy "f(n) < f(Suc(n))"; +by (fast_tac (HOL_cs addIs [Puzzle.f_ax,le_less_trans,lemma1]) 1); +val lemma2 = result(); + +val prems = goal Puzzle.thy "(!!n.f(n) <= f(Suc(n))) ==> m f(m) <= f(n)"; +by (res_inst_tac[("n","n")]nat_induct 1); +by (simp_tac nat_ss 1); +by (simp_tac nat_ss 1); +by (fast_tac (HOL_cs addIs (le_trans::prems)) 1); +val mono_lemma1 = result() RS mp; + +val [p1,p2] = goal Puzzle.thy + "[| !! n. f(n)<=f(Suc(n)); m<=n |] ==> f(m) <= f(n)"; +by (rtac (p2 RS le_imp_less_or_eq RS disjE) 1); +by (etac (p1 RS mono_lemma1) 1); +by (fast_tac (HOL_cs addIs [le_refl]) 1); +val mono_lemma = result(); + +val prems = goal Puzzle.thy "m <= n ==> f(m) <= f(n)"; +by (fast_tac (HOL_cs addIs ([mono_lemma,less_imp_le,lemma2]@prems)) 1); +val f_mono = result(); + +goal Puzzle.thy "f(n) = n"; +by (rtac le_anti_sym 1); +by (rtac lemma1 2); +by (fast_tac (HOL_cs addIs [Puzzle.f_ax,leI] addDs [leD,f_mono,lessD]) 1); +result(); diff -r 000000000000 -r 7949f97df77a ex/Puzzle.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Puzzle.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,12 @@ +(* Title: HOL/ex/puzzle.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 TU Muenchen + +An question from "Bundeswettbewerb Mathematik" +*) + +Puzzle = Nat + +consts f :: "nat => nat" +rules f_ax "f(f(n)) < f(Suc(n))" +end diff -r 000000000000 -r 7949f97df77a ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/ROOT.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,23 @@ +(* Title: HOL/ex/ROOT + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Executes all examples for Higher-Order Logic. +*) + +HOL_build_completed; (*Cause examples to fail if HOL did*) + +writeln"Root file for HOL examples"; +proof_timing := true; +time_use "ex/cla.ML"; +time_use "ex/meson.ML"; +time_use "ex/meson-test.ML"; +time_use_thy "ex/lex-prod"; +time_use_thy "ex/puzzle"; +time_use "ex/set.ML"; +time_use_thy "ex/finite"; +time_use_thy "ex/prop-log"; +time_use_thy "ex/term"; +time_use_thy "ex/simult"; +maketest "END: Root file for HOL examples"; diff -r 000000000000 -r 7949f97df77a ex/Rec.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Rec.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,5 @@ +open Rec; + +goalw thy [mono_def,Domf_def] "mono(Domf(F))"; +by (DEPTH_SOLVE (slow_step_tac set_cs 1)); +val mono_Domf = result(); diff -r 000000000000 -r 7949f97df77a ex/Rec.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Rec.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,9 @@ +Rec = Fixedpt + +consts +fix :: "('a=>'a) => 'a" +Dom :: "(('a=>'b) => ('a=>'b)) => 'a set" +Domf :: "(('a=>'b) => ('a=>'b)) => 'a set => 'a set" +rules +Domf_def "Domf(F,D) == {y . !f g. (!x:D. f(x)=g(x)) --> F(f,y)=F(g,y)}" +Dom_def "Dom(F) == lfp(Domf(F))" +end diff -r 000000000000 -r 7949f97df77a ex/Simult.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Simult.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,319 @@ +(* Title: HOL/ex/simult.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For simult.thy. + +Primitives for simultaneous recursive type definitions + includes worked example of trees & forests + +This is essentially the same data structure that on ex/term.ML, which is +simpler because it uses List as a new type former. The approach in this +file may be superior for other simultaneous recursions. +*) + +open Simult; + +(*** General rules for Part ***) + +val prems = goalw Simult.thy [Part_def] "h(a) : A ==> h(a) : Part(A,h)"; +by (cfast_tac prems 1); +val PartI = result(); + +val major::prems = goalw Simult.thy [Part_def] + "[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac conjE 1); +by (etac exE 1); +by (REPEAT (ares_tac prems 1)); +val PartE = result(); + +goal Simult.thy "!!A B. A<=B ==> Part(A,h) <= Part(B,h)"; +by (fast_tac (set_cs addSIs [PartI] addSEs [PartE]) 1); +val Part_mono = result(); + +val [prem] = goal Simult.thy "A<=B ==> Part(A,h) <= B"; +by (fast_tac (set_cs addSIs [PartI, prem RS subsetD] addSEs [PartE]) 1); +val Part_subset = result(); + + +(*** Monotonicity and unfolding of the function ***) + +goal Simult.thy "mono(%Z. A <*> Part(Z,In1) \ +\ <+> ({Numb(0)} <+> Part(Z,In0) <*> Part(Z,In1)))"; +by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono, + Part_mono] 1)); +val TF_fun_mono = result(); + +val TF_unfold = TF_fun_mono RS (TF_def RS def_lfp_Tarski); + +goalw Simult.thy [TF_def] "!!A B. A<=B ==> TF(A) <= TF(B)"; +by (REPEAT (ares_tac [lfp_mono, subset_refl, usum_mono, uprod_mono] 1)); +val TF_mono = result(); + +goalw Simult.thy [TF_def] "TF(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_NumbI,Sexp_In0I,Sexp_In1I,Sexp_SconsI] + addSEs [PartE]) 1); +val TF_Sexp = result(); + +(* A <= Sexp ==> TF(A) <= Sexp *) +val TF_subset_Sexp = standard + (TF_mono RS (TF_Sexp RSN (2,subset_trans))); + + +(** Elimination -- structural induction on the set TF **) + +val TF_Rep_defs = [TCONS_def,FNIL_def,FCONS_def,NIL_def,CONS_def]; + +val major::prems = goalw Simult.thy TF_Rep_defs + "[| i: TF(A); \ +\ !!M N. [| M: A; N: Part(TF(A),In1); R(N) |] ==> R(TCONS(M,N)); \ +\ R(FNIL); \ +\ !!M N. [| M: Part(TF(A),In0); N: Part(TF(A),In1); R(M); R(N) \ +\ |] ==> R(FCONS(M,N)) \ +\ |] ==> R(i)"; +by (rtac (major RS (TF_def RS def_induct)) 1); +by (rtac TF_fun_mono 1); +by (fast_tac (set_cs addIs (prems@[PartI]) + addEs [usumE, uprodE, PartE]) 1); +val TF_induct = result(); + +(*This lemma replaces a use of subgoal_tac to prove tree_forest_induct*) +val prems = goalw Simult.thy [Part_def] + "! M: TF(A). (M: Part(TF(A),In0) --> P(M)) & (M: Part(TF(A),In1) --> Q(M)) \ +\ ==> (! M: Part(TF(A),In0). P(M)) & (! M: Part(TF(A),In1). Q(M))"; +by (cfast_tac prems 1); +val TF_induct_lemma = result(); + +val uplus_cs = set_cs addSIs [PartI] + addSDs [In0_inject, In1_inject] + addSEs [In0_neq_In1, In1_neq_In0, PartE]; + +(*Could prove ~ TCONS(M,N) : Part(TF(A),In1) etc. *) + +(*Induction on TF with separate predicates P, Q*) +val prems = goalw Simult.thy TF_Rep_defs + "[| !!M N. [| M: A; N: Part(TF(A),In1); Q(N) |] ==> P(TCONS(M,N)); \ +\ Q(FNIL); \ +\ !!M N. [| M: Part(TF(A),In0); N: Part(TF(A),In1); P(M); Q(N) \ +\ |] ==> Q(FCONS(M,N)) \ +\ |] ==> (! M: Part(TF(A),In0). P(M)) & (! N: Part(TF(A),In1). Q(N))"; +by (rtac (ballI RS TF_induct_lemma) 1); +by (etac TF_induct 1); +by (rewrite_goals_tac TF_Rep_defs); +by (ALLGOALS (fast_tac (uplus_cs addIs prems))); +(*29 secs??*) +val Tree_Forest_induct = result(); + +(*Induction for the abstract types 'a tree, 'a forest*) +val prems = goalw Simult.thy [Tcons_def,Fnil_def,Fcons_def] + "[| !!x ts. Q(ts) ==> P(Tcons(x,ts)); \ +\ Q(Fnil); \ +\ !!t ts. [| P(t); Q(ts) |] ==> Q(Fcons(t,ts)) \ +\ |] ==> (! t. P(t)) & (! ts. Q(ts))"; +by (res_inst_tac [("P1","%z.P(Abs_Tree(z))"), + ("Q1","%z.Q(Abs_Forest(z))")] + (Tree_Forest_induct RS conjE) 1); +(*Instantiates ?A1 to range(Leaf). *) +by (fast_tac (set_cs addSEs [Rep_Tree_inverse RS subst, + Rep_Forest_inverse RS subst] + addSIs [Rep_Tree,Rep_Forest]) 4); +(*Cannot use simplifier: the rewrites work in the wrong direction!*) +by (ALLGOALS (fast_tac (set_cs addSEs [Abs_Tree_inverse RS subst, + Abs_Forest_inverse RS subst] + addSIs prems))); +val tree_forest_induct = result(); + + + +(*** Isomorphisms ***) + +goal Simult.thy "inj(Rep_Tree)"; +by (rtac inj_inverseI 1); +by (rtac Rep_Tree_inverse 1); +val inj_Rep_Tree = result(); + +goal Simult.thy "inj_onto(Abs_Tree,Part(TF(range(Leaf)),In0))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Tree_inverse 1); +val inj_onto_Abs_Tree = result(); + +goal Simult.thy "inj(Rep_Forest)"; +by (rtac inj_inverseI 1); +by (rtac Rep_Forest_inverse 1); +val inj_Rep_Forest = result(); + +goal Simult.thy "inj_onto(Abs_Forest,Part(TF(range(Leaf)),In1))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Forest_inverse 1); +val inj_onto_Abs_Forest = result(); + +(** Introduction rules for constructors **) + +(* c : A <*> Part(TF(A),In1) + <+> {Numb(0)} <+> Part(TF(A),In0) <*> Part(TF(A),In1) ==> c : TF(A) *) +val TF_I = TF_unfold RS equalityD2 RS subsetD; + +(*For reasoning about the representation*) +val TF_Rep_cs = uplus_cs addIs [TF_I, uprodI, usum_In0I, usum_In1I] + addSEs [Scons_inject]; + +val prems = goalw Simult.thy TF_Rep_defs + "[| a: A; M: Part(TF(A),In1) |] ==> TCONS(a,M) : Part(TF(A),In0)"; +by (fast_tac (TF_Rep_cs addIs prems) 1); +val TCONS_I = result(); + +(* FNIL is a TF(A) -- this also justifies the type definition*) +goalw Simult.thy TF_Rep_defs "FNIL: Part(TF(A),In1)"; +by (fast_tac TF_Rep_cs 1); +val FNIL_I = result(); + +val prems = goalw Simult.thy TF_Rep_defs + "[| M: Part(TF(A),In0); N: Part(TF(A),In1) |] ==> \ +\ FCONS(M,N) : Part(TF(A),In1)"; +by (fast_tac (TF_Rep_cs addIs prems) 1); +val FCONS_I = result(); + +(** Injectiveness of TCONS and FCONS **) + +goalw Simult.thy TF_Rep_defs "(TCONS(K,M)=TCONS(L,N)) = (K=L & M=N)"; +by (fast_tac TF_Rep_cs 1); +val TCONS_TCONS_eq = result(); +val TCONS_inject = standard (TCONS_TCONS_eq RS iffD1 RS conjE); + +goalw Simult.thy TF_Rep_defs "(FCONS(K,M)=FCONS(L,N)) = (K=L & M=N)"; +by (fast_tac TF_Rep_cs 1); +val FCONS_FCONS_eq = result(); +val FCONS_inject = standard (FCONS_FCONS_eq RS iffD1 RS conjE); + +(** Distinctness of TCONS, FNIL and FCONS **) + +goalw Simult.thy TF_Rep_defs "~ TCONS(M,N) = FNIL"; +by (fast_tac TF_Rep_cs 1); +val TCONS_not_FNIL = result(); +val FNIL_not_TCONS = standard (TCONS_not_FNIL RS not_sym); + +val TCONS_neq_FNIL = standard (TCONS_not_FNIL RS notE); +val FNIL_neq_TCONS = sym RS TCONS_neq_FNIL; + +goalw Simult.thy TF_Rep_defs "~ FCONS(M,N) = FNIL"; +by (fast_tac TF_Rep_cs 1); +val FCONS_not_FNIL = result(); +val FNIL_not_FCONS = standard (FCONS_not_FNIL RS not_sym); + +val FCONS_neq_FNIL = standard (FCONS_not_FNIL RS notE); +val FNIL_neq_FCONS = sym RS FCONS_neq_FNIL; + +goalw Simult.thy TF_Rep_defs "~ TCONS(M,N) = FCONS(K,L)"; +by (fast_tac TF_Rep_cs 1); +val TCONS_not_FCONS = result(); +val FCONS_not_TCONS = standard (TCONS_not_FCONS RS not_sym); + +val TCONS_neq_FCONS = standard (TCONS_not_FCONS RS notE); +val FCONS_neq_TCONS = sym RS TCONS_neq_FCONS; + +(*???? Too many derived rules ???? + Automatically generate symmetric forms? Always expand TF_Rep_defs? *) + +(** Injectiveness of Tcons and Fcons **) + +(*For reasoning about abstract constructors*) +val TF_cs = set_cs addSIs [Rep_Tree, Rep_Forest, TCONS_I, FNIL_I, FCONS_I] + addSEs [TCONS_inject, FCONS_inject, + TCONS_neq_FNIL, FNIL_neq_TCONS, + FCONS_neq_FNIL, FNIL_neq_FCONS, + TCONS_neq_FCONS, FCONS_neq_TCONS] + addSDs [inj_onto_Abs_Tree RS inj_ontoD, + inj_onto_Abs_Forest RS inj_ontoD, + inj_Rep_Tree RS injD, inj_Rep_Forest RS injD, + Leaf_inject]; + +goalw Simult.thy [Tcons_def] "(Tcons(x,xs)=Tcons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac TF_cs 1); +val Tcons_Tcons_eq = result(); +val Tcons_inject = standard (Tcons_Tcons_eq RS iffD1 RS conjE); + +goalw Simult.thy [Fcons_def,Fnil_def] "~ Fcons(x,xs) = Fnil"; +by (fast_tac TF_cs 1); +val Fcons_not_Fnil = result(); + +val Fcons_neq_Fnil = standard (Fcons_not_Fnil RS notE);; +val Fnil_neq_Fcons = sym RS Fcons_neq_Fnil; + + +(** Injectiveness of Fcons **) + +goalw Simult.thy [Fcons_def] "(Fcons(x,xs)=Fcons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac TF_cs 1); +val Fcons_Fcons_eq = result(); +val Fcons_inject = standard (Fcons_Fcons_eq RS iffD1 RS conjE); + + +(*** TF_rec -- by wf recursion on pred_Sexp ***) + +val TF_rec_unfold = + wf_pred_Sexp RS wf_trancl RS (TF_rec_def RS def_wfrec); + +(** conversion rules for TF_rec **) + +val prems = goalw Simult.thy [TCONS_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ TF_rec(TCONS(M,N),b,c,d) = b(M, N, TF_rec(N,b,c,d))"; +by (rtac (TF_rec_unfold RS trans) 1); +by (rtac (Case_In0 RS trans) 1); +by (rtac (Split RS trans) 1); +by (rewtac In0_def); +by (simp_tac (pred_Sexp_ss addsimps prems) 1); +val TF_rec_TCONS = result(); + +goalw Simult.thy [FNIL_def] "TF_rec(FNIL,b,c,d) = c"; +by (rtac (TF_rec_unfold RS trans) 1); +by (rtac (Case_In1 RS trans) 1); +by (rtac List_case_NIL 1); +val TF_rec_FNIL = result(); + +val prems = goalw Simult.thy [FCONS_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ TF_rec(FCONS(M,N),b,c,d) = d(M, N, TF_rec(M,b,c,d), TF_rec(N,b,c,d))"; +by (rtac (TF_rec_unfold RS trans) 1); +by (rtac (Case_In1 RS trans) 1); +by (rtac (List_case_CONS RS trans) 1); +by (rewrite_goals_tac [CONS_def,In1_def]); +by (simp_tac (pred_Sexp_ss addsimps prems) 1); +val TF_rec_FCONS = result(); + + +(*** tree_rec, forest_rec -- by TF_rec ***) + +val Rep_Tree_in_Sexp = + Rep_Tree RS (range_Leaf_subset_Sexp RS TF_subset_Sexp RS + Part_subset RS subsetD); +val Rep_Forest_in_Sexp = + Rep_Forest RS (range_Leaf_subset_Sexp RS TF_subset_Sexp RS + Part_subset RS subsetD); + +val tf_rec_simps = [TF_rec_TCONS, TF_rec_FNIL, TF_rec_FCONS, + TCONS_I, FNIL_I, FCONS_I, Rep_Tree, Rep_Forest, + Rep_Tree_inverse, Rep_Forest_inverse, + Abs_Tree_inverse, Abs_Forest_inverse, + inj_Leaf, Inv_f_f, Sexp_LeafI, range_eqI, + Rep_Tree_in_Sexp, Rep_Forest_in_Sexp]; +val tf_rec_ss = HOL_ss addsimps tf_rec_simps; + +goalw Simult.thy [tree_rec_def, forest_rec_def, Tcons_def] + "tree_rec(Tcons(a,tf),b,c,d) = b(a,tf,forest_rec(tf,b,c,d))"; +by (simp_tac tf_rec_ss 1); +val tree_rec_Tcons = result(); + +goalw Simult.thy [forest_rec_def, Fnil_def] "forest_rec(Fnil,b,c,d) = c"; +by (simp_tac tf_rec_ss 1); +val forest_rec_Fnil = result(); + +goalw Simult.thy [tree_rec_def, forest_rec_def, Fcons_def] + "forest_rec(Fcons(t,tf),b,c,d) = \ +\ d(t,tf,tree_rec(t,b,c,d), forest_rec(tf,b,c,d))"; +by (simp_tac tf_rec_ss 1); +val forest_rec_Cons = result(); diff -r 000000000000 -r 7949f97df77a ex/Simult.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Simult.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,77 @@ +(* Title: HOL/ex/simult + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Primitives for simultaneous recursive type definitions + includes worked example of trees & forests + +This is essentially the same data structure that on ex/term.ML, which is +simpler because it uses List as a new type former. The approach in this +file may be superior for other simultaneous recursions. +*) + +Simult = List + +types tree,forest 1 +arities tree,forest :: (term)term +consts + Part :: "['a set, 'a=>'a] => 'a set" + TF :: "'a node set set => 'a node set set" + FNIL :: "'a node set" + TCONS,FCONS :: "['a node set, 'a node set] => 'a node set" + Rep_Tree :: "'a tree => 'a node set" + Abs_Tree :: "'a node set => 'a tree" + Rep_Forest :: "'a forest => 'a node set" + Abs_Forest :: "'a node set => 'a forest" + Tcons :: "['a, 'a forest] => 'a tree" + Fcons :: "['a tree, 'a forest] => 'a forest" + Fnil :: "'a forest" + TF_rec :: "['a node set, ['a node set , 'a node set, 'b]=>'b, \ +\ 'b, ['a node set , 'a node set, 'b, 'b]=>'b] => 'b" + tree_rec :: "['a tree, ['a, 'a forest, 'b]=>'b, \ +\ 'b, ['a tree, 'a forest, 'b, 'b]=>'b] => 'b" + forest_rec :: "['a forest, ['a, 'a forest, 'b]=>'b, \ +\ 'b, ['a tree, 'a forest, 'b, 'b]=>'b] => 'b" + +rules + (*operator for selecting out the various types*) + Part_def "Part(A,h) == {x. x:A & (? z. x = h(z))}" + + TF_def "TF(A) == lfp(%Z. A <*> Part(Z,In1) \ +\ <+> ({Numb(0)} <+> Part(Z,In0) <*> Part(Z,In1)))" + (*faking a type definition for tree...*) + Rep_Tree "Rep_Tree(n): Part(TF(range(Leaf)),In0)" + Rep_Tree_inverse "Abs_Tree(Rep_Tree(t)) = t" + Abs_Tree_inverse "z: Part(TF(range(Leaf)),In0) ==> Rep_Tree(Abs_Tree(z)) = z" + (*faking a type definition for forest...*) + Rep_Forest "Rep_Forest(n): Part(TF(range(Leaf)),In1)" + Rep_Forest_inverse "Abs_Forest(Rep_Forest(ts)) = ts" + Abs_Forest_inverse + "z: Part(TF(range(Leaf)),In1) ==> Rep_Forest(Abs_Forest(z)) = z" + + (*the concrete constants*) + TCONS_def "TCONS(M,N) == In0(M . N)" + FNIL_def "FNIL == In1(NIL)" + FCONS_def "FCONS(M,N) == In1(CONS(M,N))" + (*the abstract constants*) + Tcons_def "Tcons(a,ts) == Abs_Tree(TCONS(Leaf(a), Rep_Forest(ts)))" + Fnil_def "Fnil == Abs_Forest(FNIL)" + Fcons_def "Fcons(t,ts) == Abs_Forest(FCONS(Rep_Tree(t), Rep_Forest(ts)))" + + (*recursion*) + TF_rec_def + "TF_rec(M,b,c,d) == wfrec(trancl(pred_Sexp), M, \ +\ %Z g. Case(Z, %U. Split(U, %x y. b(x,y,g(y))), \ +\ %V. List_case(V, c, \ +\ %x y. d(x,y,g(x),g(y)))))" + + tree_rec_def + "tree_rec(t,b,c,d) == \ +\ TF_rec(Rep_Tree(t), %x y r. b(Inv(Leaf,x), Abs_Forest(y), r), \ +\ c, %x y rt rf. d(Abs_Tree(x), Abs_Forest(y), rt, rf))" + + forest_rec_def + "forest_rec(tf,b,c,d) == \ +\ TF_rec(Rep_Forest(tf), %x y r. b(Inv(Leaf,x), Abs_Forest(y), r), \ +\ c, %x y rt rf. d(Abs_Tree(x), Abs_Forest(y), rt, rf))" +end diff -r 000000000000 -r 7949f97df77a ex/Term.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Term.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,177 @@ +(* Title: HOL/ex/term + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For term.thy. illustrates List functor + (essentially the same type as in Trees & Forests) +*) + +open Term; + +(*** Monotonicity and unfolding of the function ***) + +goal Term.thy "mono(%Z. A <*> List(Z))"; +by (REPEAT (ares_tac [monoI, subset_refl, List_mono, uprod_mono] 1)); +val Term_fun_mono = result(); + +val Term_unfold = Term_fun_mono RS (Term_def RS def_lfp_Tarski); + +(*This justifies using Term in other recursive type definitions*) +goalw Term.thy [Term_def] "!!A B. A<=B ==> Term(A) <= Term(B)"; +by (REPEAT (ares_tac [lfp_mono, subset_refl, List_mono, uprod_mono] 1)); +val Term_mono = result(); + +(** Type checking rules -- Term creates well-founded sets **) + +val prems = goalw Term.thy [Term_def] "Term(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_SconsI, List_Sexp RS subsetD]) 1); +val Term_Sexp = result(); + +(* A <= Sexp ==> Term(A) <= Sexp *) +val Term_subset_Sexp = standard + (Term_mono RS (Term_Sexp RSN (2,subset_trans))); + + +(** Elimination -- structural induction on the set Term(A) **) + +(*Induction for the set Term(A) *) +val [major,minor] = goal Term.thy + "[| M: Term(A); \ +\ !!x zs. [| x: A; zs: List(Term(A)); zs: List({x.R(x)}) \ +\ |] ==> R(x.zs) \ +\ |] ==> R(M)"; +by (rtac (major RS (Term_def RS def_induct)) 1); +by (rtac Term_fun_mono 1); +by (REPEAT (eresolve_tac ([uprodE, ssubst, minor] @ + ([Int_lower1,Int_lower2] RL [List_mono RS subsetD])) 1)); +val Term_induct = result(); + +(*Induction on Term(A) followed by induction on List *) +val major::prems = goal Term.thy + "[| M: Term(A); \ +\ !!x. [| x: A |] ==> R(x.NIL); \ +\ !!x z zs. [| x: A; z: Term(A); zs: List(Term(A)); R(x.zs) \ +\ |] ==> R(x . CONS(z,zs)) \ +\ |] ==> R(M)"; +by (rtac (major RS Term_induct) 1); +by (etac List_induct 1); +by (REPEAT (ares_tac prems 1)); +val Term_induct2 = result(); + +(*** Structural Induction on the abstract type 'a term ***) + +val list_all_ss = map_ss addsimps [list_all_Nil, list_all_Cons]; + +val Rep_Term_in_Sexp = + Rep_Term RS (range_Leaf_subset_Sexp RS Term_subset_Sexp RS subsetD); + +(*Induction for the abstract type 'a term*) +val prems = goalw Term.thy [App_def,Rep_TList_def,Abs_TList_def] + "[| !!x ts. list_all(R,ts) ==> R(App(x,ts)) \ +\ |] ==> R(t)"; +by (rtac (Rep_Term_inverse RS subst) 1); (*types force good instantiation*) +by (res_inst_tac [("P","Rep_Term(t) : Sexp")] conjunct2 1); +by (rtac (Rep_Term RS Term_induct) 1); +by (REPEAT (ares_tac [conjI, Sexp_SconsI, Term_subset_Sexp RS + List_subset_Sexp,range_Leaf_subset_Sexp] 1 + ORELSE etac rev_subsetD 1)); +by (eres_inst_tac [("A1","Term(?u)"), ("f1","Rep_Term"), ("g1","Abs_Term")] + (Abs_map_inverse RS subst) 1); +by (rtac (range_Leaf_subset_Sexp RS Term_subset_Sexp) 1); +by (etac Abs_Term_inverse 1); +by (etac rangeE 1); +by (hyp_subst_tac 1); +by (resolve_tac prems 1); +by (etac List_induct 1); +by (etac CollectE 2); +by (stac Abs_map_CONS 2); +by (etac conjunct1 2); +by (etac rev_subsetD 2); +by (rtac List_subset_Sexp 2); +by (fast_tac set_cs 2); +by (ALLGOALS (asm_simp_tac list_all_ss)); +val term_induct = result(); + +(*Induction for the abstract type 'a term*) +val prems = goal Term.thy + "[| !!x. R(App(x,Nil)); \ +\ !!x t ts. R(App(x,ts)) ==> R(App(x, Cons(t,ts))) \ +\ |] ==> R(t)"; +by (rtac term_induct 1); (*types force good instantiation*) +by (etac rev_mp 1); +by (rtac list_induct 1); (*types force good instantiation*) +by (ALLGOALS (asm_simp_tac (list_all_ss addsimps prems))); +val term_induct2 = result(); + +(*Perform induction on xs. *) +fun term_ind2_tac a i = + EVERY [res_inst_tac [("t",a)] term_induct2 i, + rename_last_tac a ["1","s"] (i+1)]; + + +(** Introduction rule for Term **) + +(* c : A <*> List(Term(A)) ==> c : Term(A) *) +val TermI = standard (Term_unfold RS equalityD2 RS subsetD); + +(*The constant APP is not declared; it is simply . *) +val prems = goal Term.thy "[| M: A; N : List(Term(A)) |] ==> M.N : Term(A)"; +by (REPEAT (resolve_tac (prems@[TermI, ListI, uprodI]) 1)); +val APP_I = result(); + + +(*** Term_rec -- by wf recursion on pred_Sexp ***) + +val Term_rec_unfold = + wf_pred_Sexp RS wf_trancl RS (Term_rec_def RS def_wfrec); + +(** conversion rules **) + +val [prem] = goal Term.thy + "N: List(Term(A)) ==> \ +\ !M. : pred_Sexp^+ --> \ +\ Abs_map(cut(h, pred_Sexp^+, M), N) = \ +\ Abs_map(h,N)"; +by (rtac (prem RS List_induct) 1); +by (simp_tac list_all_ss 1); +by (strip_tac 1); +by (etac (pred_Sexp_CONS_D RS conjE) 1); +by (asm_simp_tac (list_all_ss addsimps [trancl_pred_SexpD1, cut_apply]) 1); +val Abs_map_lemma = result(); + +val [prem1,prem2,A_subset_Sexp] = goal Term.thy + "[| M: Sexp; N: List(Term(A)); A<=Sexp |] ==> \ +\ Term_rec(M.N, d) = d(M, N, Abs_map(%Z. Term_rec(Z,d), N))"; +by (rtac (Term_rec_unfold RS trans) 1); +by (rtac (Split RS trans) 1); +by (simp_tac (HOL_ss addsimps + [prem2 RS Abs_map_lemma RS spec RS mp, pred_SexpI2 RS r_into_trancl, + prem1, prem2 RS rev_subsetD, List_subset_Sexp, + Term_subset_Sexp, A_subset_Sexp])1); +val Term_rec = result(); + +(*** term_rec -- by Term_rec ***) + +local + val Rep_map_type1 = read_instantiate_sg (sign_of Term.thy) + [("f","Rep_Term")] Rep_map_type; + val Rep_TList = Rep_Term RS Rep_map_type1; + val Rep_Term_rec = range_Leaf_subset_Sexp RSN (2,Rep_TList RSN(2,Term_rec)); + + (*Now avoids conditional rewriting with the premise N: List(Term(A)), + since A will be uninstantiated and will cause rewriting to fail. *) + val term_rec_ss = HOL_ss + addsimps [Rep_TList RS (rangeI RS APP_I RS Abs_Term_inverse), + Rep_Term_in_Sexp, Rep_Term_rec, Rep_Term_inverse, + inj_Leaf, Inv_f_f, + Abs_Rep_map, map_ident, Sexp_LeafI] +in + +val term_rec = prove_goalw Term.thy + [term_rec_def, App_def, Rep_TList_def, Abs_TList_def] + "term_rec(App(f,ts), d) = d(f, ts, map (%t. term_rec(t,d), ts))" + (fn _ => [simp_tac term_rec_ss 1]) + +end; diff -r 000000000000 -r 7949f97df77a ex/Term.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/Term.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,44 @@ +(* Title: HOL/ex/term + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Terms over a given alphabet -- function applications; illustrates List functor + (essentially the same type as in Trees & Forests) + +There is no constructor APP because it is simply cons (.) +*) + +Term = List + +types term 1 +arities term :: (term)term +consts + Term :: "'a node set set => 'a node set set" + Rep_Term :: "'a term => 'a node set" + Abs_Term :: "'a node set => 'a term" + Rep_TList :: "'a term list => 'a node set" + Abs_TList :: "'a node set => 'a term list" + App :: "['a, ('a term)list] => 'a term" + Term_rec :: + "['a node set, ['a node set , 'a node set, 'b list]=>'b] => 'b" + term_rec :: "['a term, ['a ,'a term list, 'b list]=>'b] => 'b" +rules + Term_def "Term(A) == lfp(%Z. A <*> List(Z))" + (*faking a type definition for term...*) + Rep_Term "Rep_Term(n): Term(range(Leaf))" + Rep_Term_inverse "Abs_Term(Rep_Term(t)) = t" + Abs_Term_inverse "M: Term(range(Leaf)) ==> Rep_Term(Abs_Term(M)) = M" + (*defining abstraction/representation functions for term list...*) + Rep_TList_def "Rep_TList == Rep_map(Rep_Term)" + Abs_TList_def "Abs_TList == Abs_map(Abs_Term)" + (*defining the abstract constants*) + App_def "App(a,ts) == Abs_Term(Leaf(a) . Rep_TList(ts))" + (*list recursion*) + Term_rec_def + "Term_rec(M,d) == wfrec(trancl(pred_Sexp), M, \ +\ %M g. Split(M, %x y. d(x,y, Abs_map(g,y))))" + + term_rec_def + "term_rec(t,d) == \ +\ Term_rec(Rep_Term(t), %x y r. d(Inv(Leaf,x), Abs_TList(y), r))" +end diff -r 000000000000 -r 7949f97df77a ex/cla.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/cla.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,442 @@ +(* Title: HOL/ex/cla + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Higher-Order Logic: predicate calculus problems + +Taken from FOL/cla.ML; beware of precedence of = vs <-> +*) + +writeln"File HOL/ex/cla."; + +goal HOL.thy "(P --> Q | R) --> (P-->Q) | (P-->R)"; +by (fast_tac HOL_cs 1); +result(); + +(*If and only if*) + +goal HOL.thy "(P=Q) = (Q=P::bool)"; +by (fast_tac HOL_cs 1); +result(); + +goal HOL.thy "~ (P = (~P))"; +by (fast_tac HOL_cs 1); +result(); + + +(*Sample problems from + F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. + +The hardest problems -- judging by experience with several theorem provers, +including matrix ones -- are 34 and 43. +*) + +writeln"Pelletier's examples"; +(*1*) +goal HOL.thy "(P-->Q) = (~Q --> ~P)"; +by (fast_tac HOL_cs 1); +result(); + +(*2*) +goal HOL.thy "(~ ~ P) = P"; +by (fast_tac HOL_cs 1); +result(); + +(*3*) +goal HOL.thy "~(P-->Q) --> (Q-->P)"; +by (fast_tac HOL_cs 1); +result(); + +(*4*) +goal HOL.thy "(~P-->Q) = (~Q --> P)"; +by (fast_tac HOL_cs 1); +result(); + +(*5*) +goal HOL.thy "((P|Q)-->(P|R)) --> (P|(Q-->R))"; +by (fast_tac HOL_cs 1); +result(); + +(*6*) +goal HOL.thy "P | ~ P"; +by (fast_tac HOL_cs 1); +result(); + +(*7*) +goal HOL.thy "P | ~ ~ ~ P"; +by (fast_tac HOL_cs 1); +result(); + +(*8. Peirce's law*) +goal HOL.thy "((P-->Q) --> P) --> P"; +by (fast_tac HOL_cs 1); +result(); + +(*9*) +goal HOL.thy "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (fast_tac HOL_cs 1); +result(); + +(*10*) +goal HOL.thy "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P=Q)"; +by (fast_tac HOL_cs 1); +result(); + +(*11. Proved in each direction (incorrectly, says Pelletier!!) *) +goal HOL.thy "P=P::bool"; +by (fast_tac HOL_cs 1); +result(); + +(*12. "Dijkstra's law"*) +goal HOL.thy "((P = Q) = R) = (P = (Q = R))"; +by (fast_tac HOL_cs 1); +result(); + +(*13. Distributive law*) +goal HOL.thy "(P | (Q & R)) = ((P | Q) & (P | R))"; +by (fast_tac HOL_cs 1); +result(); + +(*14*) +goal HOL.thy "(P = Q) = ((Q | ~P) & (~Q|P))"; +by (fast_tac HOL_cs 1); +result(); + +(*15*) +goal HOL.thy "(P --> Q) = (~P | Q)"; +by (fast_tac HOL_cs 1); +result(); + +(*16*) +goal HOL.thy "(P-->Q) | (Q-->P)"; +by (fast_tac HOL_cs 1); +result(); + +(*17*) +goal HOL.thy "((P & (Q-->R))-->S) = ((~P | Q | S) & (~P | ~R | S))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Classical Logic: examples with quantifiers"; + +goal HOL.thy "(! x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; +by (fast_tac HOL_cs 1); +result(); + +goal HOL.thy "(? x. P-->Q(x)) = (P --> (? x.Q(x)))"; +by (fast_tac HOL_cs 1); +result(); + +goal HOL.thy "(? x.P(x)-->Q) = ((! x.P(x)) --> Q)"; +by (fast_tac HOL_cs 1); +result(); + +goal HOL.thy "((! x.P(x)) | Q) = (! x. P(x) | Q)"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problems requiring quantifier duplication"; + +(*Needs multiple instantiation of the quantifier.*) +goal HOL.thy "(! x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; +by (best_tac HOL_dup_cs 1); +result(); + +(*Needs double instantiation of the quantifier*) +goal HOL.thy "? x. P(x) --> P(a) & P(b)"; +by (best_tac HOL_dup_cs 1); +result(); + +goal HOL.thy "? z. P(z) --> (! x. P(x))"; +by (best_tac HOL_dup_cs 1); +result(); + +writeln"Hard examples with quantifiers"; + +writeln"Problem 18"; +goal HOL.thy "? y. ! x. P(y)-->P(x)"; +by (best_tac HOL_dup_cs 1); +result(); + +writeln"Problem 19"; +goal HOL.thy "? x. ! y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; +by (best_tac HOL_dup_cs 1); +result(); + +writeln"Problem 20"; +goal HOL.thy "(! x y. ? z. ! w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (? x y. P(x) & Q(y)) --> (? z. R(z))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 21"; +goal HOL.thy "(? x. P-->Q(x)) & (? x. Q(x)-->P) --> (? x. P=Q(x))"; +by (best_tac HOL_dup_cs 1); +result(); + +writeln"Problem 22"; +goal HOL.thy "(! x. P = Q(x)) --> (P = (! x. Q(x)))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 23"; +goal HOL.thy "(! x. P | Q(x)) = (P | (! x. Q(x)))"; +by (best_tac HOL_cs 1); +result(); + +writeln"Problem 24"; +goal HOL.thy "~(? x. S(x)&Q(x)) & (! x. P(x) --> Q(x)|R(x)) & \ +\ ~(? x.P(x)) --> (? x.Q(x)) & (! x. Q(x)|R(x) --> S(x)) \ +\ --> (? x. P(x)&R(x))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 25"; +goal HOL.thy "(? x. P(x)) & \ +\ (! x. L(x) --> ~ (M(x) & R(x))) & \ +\ (! x. P(x) --> (M(x) & L(x))) & \ +\ ((! x. P(x)-->Q(x)) | (? x. P(x)&R(x))) \ +\ --> (? x. Q(x)&P(x))"; +by (best_tac HOL_cs 1); +result(); + +writeln"Problem 26"; +goal HOL.thy "((? x. p(x)) = (? x. q(x))) & \ +\ (! x. ! y. p(x) & q(y) --> (r(x) = s(y))) \ +\ --> ((! x. p(x)-->r(x)) = (! x. q(x)-->s(x)))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 27"; +goal HOL.thy "(? x. P(x) & ~Q(x)) & \ +\ (! x. P(x) --> R(x)) & \ +\ (! x. M(x) & L(x) --> P(x)) & \ +\ ((? x. R(x) & ~ Q(x)) --> (! x. L(x) --> ~ R(x))) \ +\ --> (! x. M(x) --> ~L(x))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 28. AMENDED"; +goal HOL.thy "(! x. P(x) --> (! x. Q(x))) & \ +\ ((! x. Q(x)|R(x)) --> (? x. Q(x)&S(x))) & \ +\ ((? x.S(x)) --> (! x. L(x) --> M(x))) \ +\ --> (! x. P(x) & L(x) --> M(x))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal HOL.thy "(? x. F(x)) & (? y. G(y)) \ +\ --> ( ((! x. F(x)-->H(x)) & (! y. G(y)-->J(y))) = \ +\ (! x y. F(x) & G(y) --> H(x) & J(y)))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 30"; +goal HOL.thy "(! x. P(x) | Q(x) --> ~ R(x)) & \ +\ (! x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (! x. S(x))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 31"; +goal HOL.thy "~(? x.P(x) & (Q(x) | R(x))) & \ +\ (? x. L(x) & P(x)) & \ +\ (! x. ~ R(x) --> M(x)) \ +\ --> (? x. L(x) & M(x))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 32"; +goal HOL.thy "(! x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (! x. S(x) & R(x) --> L(x)) & \ +\ (! x. M(x) --> R(x)) \ +\ --> (! x. P(x) & M(x) --> L(x))"; +by (best_tac HOL_cs 1); +result(); + +writeln"Problem 33"; +goal HOL.thy "(! x. P(a) & (P(x)-->P(b))-->P(c)) = \ +\ (! x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; +by (best_tac HOL_cs 1); +result(); + +writeln"Problem 34 AMENDED (TWICE!!) NOT PROVED AUTOMATICALLY"; +(*Andrews's challenge*) +goal HOL.thy "((? x. ! y. p(x) = p(y)) = \ +\ ((? x. q(x)) = (! y. p(y)))) = \ +\ ((? x. ! y. q(x) = q(y)) = \ +\ ((? x. p(x)) = (! y. q(y))))"; +by (safe_tac HOL_cs); (*24 secs*) +by (TRYALL (fast_tac HOL_cs)); (*161 secs*) +by (TRYALL (best_tac HOL_dup_cs)); (*86 secs -- much faster than FOL!*) +result(); + +writeln"Problem 35"; +goal HOL.thy "? x y. P(x,y) --> (! u v. P(u,v))"; +by (best_tac HOL_dup_cs 1); +result(); + +writeln"Problem 36"; +goal HOL.thy "(! x. ? y. J(x,y)) & \ +\ (! x. ? y. G(x,y)) & \ +\ (! x y. J(x,y) | G(x,y) --> \ +\ (! z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (! x. ? y. H(x,y))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 37"; +goal HOL.thy "(! z. ? w. ! x. ? y. \ +\ (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (? u.Q(u,w)))) & \ +\ (! x z. ~P(x,z) --> (? y. Q(y,z))) & \ +\ ((? x y. Q(x,y)) --> (! x. R(x,x))) \ +\ --> (! x. ? y. R(x,y))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 38"; +goal HOL.thy + "(! x. p(a) & (p(x) --> (? y. p(y) & r(x,y))) --> \ +\ (? z. ? w. p(z) & r(x,w) & r(w,z))) = \ +\ (! x. (~p(a) | p(x) | (? z. ? w. p(z) & r(x,w) & r(w,z))) & \ +\ (~p(a) | ~(? y. p(y) & r(x,y)) | \ +\ (? z. ? w. p(z) & r(x,w) & r(w,z))))"; + +writeln"Problem 39"; +goal HOL.thy "~ (? x. ! y. F(y,x) = (~F(y,y)))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 40. AMENDED"; +goal HOL.thy "(? y. ! x. F(x,y) = F(x,x)) \ +\ --> ~ (! x. ? y. ! z. F(z,y) = (~F(z,x)))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 41"; +goal HOL.thy "(! z. ? y. ! x. f(x,y) = (f(x,z) & ~ f(x,x))) \ +\ --> ~ (? z. ! x. f(x,z))"; +by (best_tac HOL_cs 1); +result(); + +writeln"Problem 42"; +goal HOL.thy "~ (? y. ! x. p(x,y) = (~ (? z. p(x,z) & p(z,x))))"; + +writeln"Problem 43 NOT PROVED AUTOMATICALLY"; +goal HOL.thy + "(! x::'a. ! y::'a. q(x,y) = (! z. p(z,x) = p(z,y)::bool)) \ +\ --> (! x. (! y. q(x,y) = q(y,x)::bool))"; + + +writeln"Problem 44"; +goal HOL.thy "(! x. f(x) --> \ +\ (? y. g(y) & h(x,y) & (? y. g(y) & ~ h(x,y)))) & \ +\ (? x. j(x) & (! y. g(y) --> h(x,y))) \ +\ --> (? x. j(x) & ~f(x))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 45"; +goal HOL.thy + "(! x. f(x) & (! y. g(y) & h(x,y) --> j(x,y)) \ +\ --> (! y. g(y) & h(x,y) --> k(y))) & \ +\ ~ (? y. l(y) & k(y)) & \ +\ (? x. f(x) & (! y. h(x,y) --> l(y)) \ +\ & (! y. g(y) & h(x,y) --> j(x,y))) \ +\ --> (? x. f(x) & ~ (? y. g(y) & h(x,y)))"; +by (best_tac HOL_cs 1); +result(); + + +writeln"Problems (mainly) involving equality or functions"; + +writeln"Problem 48"; +goal HOL.thy "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 49 NOT PROVED AUTOMATICALLY"; +(*Hard because it involves substitution for Vars; + the type constraint ensures that x,y,z have the same type as a,b,u. *) +goal HOL.thy "(? x y::'a. ! z. z=x | z=y) & P(a) & P(b) & (~a=b) \ +\ --> (! u::'a.P(u))"; +by (Classical.safe_tac HOL_cs); +by (res_inst_tac [("x","a")] allE 1); +by (assume_tac 1); +by (res_inst_tac [("x","b")] allE 1); +by (assume_tac 1); +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 50"; +(*What has this to do with equality?*) +goal HOL.thy "(! x. P(a,x) | (! y.P(x,y))) --> (? x. ! y.P(x,y))"; +by (best_tac HOL_dup_cs 1); +result(); + +writeln"Problem 51"; +goal HOL.thy + "(? z w. ! x y. P(x,y) = (x=z & y=w)) --> \ +\ (? z. ! x. ? w. (! y. P(x,y) = (y=w)) = (x=z))"; +by (best_tac HOL_cs 1); +result(); + +writeln"Problem 52"; +(*Almost the same as 51. *) +goal HOL.thy + "(? z w. ! x y. P(x,y) = (x=z & y=w)) --> \ +\ (? w. ! y. ? z. (! x. P(x,y) = (x=z)) = (y=w))"; +by (best_tac HOL_cs 1); +result(); + +writeln"Problem 55"; + +(*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988). + fast_tac DISCOVERS who killed Agatha. *) +goal HOL.thy "lives(agatha) & lives(butler) & lives(charles) & \ +\ (killed(agatha,agatha) | killed(butler,agatha) | killed(charles,agatha)) & \ +\ (!x y. killed(x,y) --> hates(x,y) & ~richer(x,y)) & \ +\ (!x. hates(agatha,x) --> ~hates(charles,x)) & \ +\ (hates(agatha,agatha) & hates(agatha,charles)) & \ +\ (!x. lives(x) & ~richer(x,agatha) --> hates(butler,x)) & \ +\ (!x. hates(agatha,x) --> hates(butler,x)) & \ +\ (!x. ~hates(x,agatha) | ~hates(x,butler) | ~hates(x,charles)) --> \ +\ killed(?who,agatha)"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 56"; +goal HOL.thy + "(! x. (? y. P(y) & x=f(y)) --> P(x)) = (! x. P(x) --> P(f(x)))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 57"; +goal HOL.thy + "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (! x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Problem 58 NOT PROVED AUTOMATICALLY"; +goal HOL.thy "(! x y. f(x)=g(y)) --> (! x y. f(f(x))=f(g(y)))"; +val f_cong = read_instantiate [("f","f")] arg_cong; +by (fast_tac (HOL_cs addIs [f_cong]) 1); +result(); + +writeln"Problem 59"; +goal HOL.thy "(! x. P(x) = (~P(f(x)))) --> (? x. P(x) & ~P(f(x)))"; +by (best_tac HOL_dup_cs 1); +result(); + +writeln"Problem 60"; +goal HOL.thy + "! x. P(x,f(x)) = (? y. (! z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (fast_tac HOL_cs 1); +result(); + +writeln"Reached end of file."; diff -r 000000000000 -r 7949f97df77a ex/finite.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/finite.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,101 @@ +open Finite; + +(** Monotonicity and unfolding of the function **) + +goal Finite.thy "mono(%F. insert({}, UN x:A. insert(x)``F))"; +by (rtac monoI 1); +by (fast_tac set_cs 1); +val Fin_bnd_mono = result(); + +goalw Finite.thy [Fin_def] "!!A B. A<=B ==> Fin(A) <= Fin(B)"; +br lfp_mono 1; +by(fast_tac set_cs 1); +val Fin_mono = result(); + +goalw Finite.thy [Fin_def] "Fin(A) <= {B. B <= A}"; +by (fast_tac (set_cs addSIs [lfp_lowerbound]) 1); +val Fin_subset_Pow = result(); + +(* A : Fin(B) ==> A <= B *) +val FinD = Fin_subset_Pow RS subsetD RS CollectD; + +val Fin_unfold = Fin_bnd_mono RS (Fin_def RS def_lfp_Tarski); + +goal Finite.thy "{} : Fin(A)"; +by (rtac (Fin_unfold RS ssubst) 1); +by (rtac insertI1 1); +val Fin_0I = result(); + +goal Finite.thy "!!a. [| a:A; F:Fin(A) |] ==> insert(a,F) : Fin(A)"; +by (rtac (Fin_unfold RS ssubst) 1); +by(fast_tac set_cs 1); +val Fin_insertI = result(); + +(*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_def RS def_induct)) 1); +by (rtac Fin_bnd_mono 1); +by (safe_tac set_cs); +(*Excluded middle on x:y would be more straightforward; + actual proof is inspired by Dov Gabbay's Restart Rule*) +by (rtac classical 2); +by (REPEAT (ares_tac prems 1)); +by (etac contrapos 1); +by (rtac (insert_absorb RS ssubst) 1); +by (REPEAT (assume_tac 1)); +val Fin_induct = result(); + +(** Simplification for Fin **) + +val Fin_ss = set_ss addsimps [Fin_0I, Fin_insertI]; + +(*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])))); +val Fin_UnI = result(); + +(*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)); +val Fin_subset = result(); + +(*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); +val Fin_imageI = result(); + +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@[insert_subset_iff, + Diff_subset RS Fin_subset])))); +val Fin_empty_induct_lemma = result(); + +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)); +val Fin_empty_induct = result(); diff -r 000000000000 -r 7949f97df77a ex/finite.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/finite.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,5 @@ +Finite = Lfp + +consts Fin :: "'a set => 'a set set" +rules + Fin_def "Fin(A) == lfp(%F. insert({}, UN x:A. insert(x)``F))" +end diff -r 000000000000 -r 7949f97df77a ex/lexprod.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/lexprod.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,24 @@ +(* Title: HOL/ex/lex-prod.ML + ID: $Id$ + Author: Tobias Nipkow, TU Munich + Copyright 1993 TU Munich + +For lex-prod.thy. +The lexicographic product of two wellfounded relations is again wellfounded. +*) + +val prems = goal Prod.thy "!a b. P() ==> !p.P(p)"; +by (cut_facts_tac prems 1); +by (rtac allI 1); +by (rtac (surjective_pairing RS ssubst) 1); +by (fast_tac HOL_cs 1); +val split_all_pair = result(); + +val [wfa,wfb] = goalw LexProd.thy [wf_def,LexProd.lex_prod_def] + "[| wf(ra); wf(rb) |] ==> wf(ra**rb)"; +by (EVERY1 [rtac allI,rtac impI, rtac split_all_pair]); +by (rtac (wfa RS spec RS mp) 1); +by (EVERY1 [rtac allI,rtac impI]); +by (rtac (wfb RS spec RS mp) 1); +by (fast_tac (set_cs addSEs [Pair_inject]) 1); +val wf_lex_prod = result(); diff -r 000000000000 -r 7949f97df77a ex/lexprod.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/lexprod.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,15 @@ +(* Title: HOL/ex/lex-prod.thy + ID: $Id$ + Author: Tobias Nipkow, TU Munich + Copyright 1993 TU Munich + +The lexicographic product of two relations. +*) + +LexProd = WF + Prod + +consts "**" :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set" (infixl 70) +rules +lex_prod_def "ra**rb == {p. ? a a' b b'. \ +\ p = <,> & ( : ra | a=a' & : rb)}" +end + diff -r 000000000000 -r 7949f97df77a ex/meson-test.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/meson-test.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,496 @@ +(* Title: HOL/ex/meson + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Test data for the MESON proof procedure + (Excludes the equality problems 51, 52, 56, 58) + +show_hyps:=false; + +by (rtac ccontr 1); +val [prem] = gethyps 1; +val nnf = make_nnf prem; +val xsko = skolemize nnf; +by (cut_facts_tac [xsko] 1 THEN REPEAT (etac exE 1)); +val [_,sko] = gethyps 1; +val clauses = make_clauses [sko]; +val horns = make_horns clauses; +val go::_ = neg_clauses clauses; + +goal HOL.thy "False"; +by (rtac (make_goal go) 1); +by (prolog_step_tac horns 1); +by (depth_prolog_tac horns); +by (best_prolog_tac size_of_subgoals horns); +*) + +writeln"File HOL/ex/meson-test."; + +(**** Interactive examples ****) + +(*Generate nice names for Skolem functions*) +Logic.auto_rename := true; Logic.set_rename_prefix "a"; + + +writeln"Problem 25"; +goal HOL.thy "(? x. P(x)) & \ +\ (! x. L(x) --> ~ (M(x) & R(x))) & \ +\ (! x. P(x) --> (M(x) & L(x))) & \ +\ ((! x. P(x)-->Q(x)) | (? x. P(x)&R(x))) \ +\ --> (? x. Q(x)&P(x))"; +by (rtac ccontr 1); +val [prem25] = gethyps 1; +val nnf25 = make_nnf prem25; +val xsko25 = skolemize nnf25; +by (cut_facts_tac [xsko25] 1 THEN REPEAT (etac exE 1)); +val [_,sko25] = gethyps 1; +val clauses25 = make_clauses [sko25]; (*7 clauses*) +val horns25 = make_horns clauses25; (*16 Horn clauses*) +val go25::_ = neg_clauses clauses25; + +goal HOL.thy "False"; +by (rtac (make_goal go25) 1); +by (depth_prolog_tac horns25); + + +writeln"Problem 26"; +goal HOL.thy "((? x. p(x)) = (? x. q(x))) & \ +\ (! x. ! y. p(x) & q(y) --> (r(x) = s(y))) \ +\ --> ((! x. p(x)-->r(x)) = (! x. q(x)-->s(x)))"; +by (rtac ccontr 1); +val [prem26] = gethyps 1; +val nnf26 = make_nnf prem26; +val xsko26 = skolemize nnf26; +by (cut_facts_tac [xsko26] 1 THEN REPEAT (etac exE 1)); +val [_,sko26] = gethyps 1; +val clauses26 = make_clauses [sko26]; (*9 clauses*) +val horns26 = make_horns clauses26; (*24 Horn clauses*) +val go26::_ = neg_clauses clauses26; + +goal HOL.thy "False"; +by (rtac (make_goal go26) 1); +by (depth_prolog_tac horns26); (*6 secs*) + + + +writeln"Problem 43 NOW PROVED AUTOMATICALLY!!"; +goal HOL.thy "(! x. ! y. q(x,y) = (! z. p(z,x) = p(z,y)::bool)) \ +\ --> (! x. (! y. q(x,y) = q(y,x)::bool))"; +by (rtac ccontr 1); +val [prem43] = gethyps 1; +val nnf43 = make_nnf prem43; +val xsko43 = skolemize nnf43; +by (cut_facts_tac [xsko43] 1 THEN REPEAT (etac exE 1)); +val [_,sko43] = gethyps 1; +val clauses43 = make_clauses [sko43]; (*6*) +val horns43 = make_horns clauses43; (*16*) +val go43::_ = neg_clauses clauses43; + +goal HOL.thy "False"; +by (rtac (make_goal go43) 1); +by (best_prolog_tac size_of_subgoals horns43); +(*8.7 secs*) + + +(*Restore variable name preservation*) +Logic.auto_rename := false; + + +(**** Batch test data ****) + +(*Sample problems from + F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. + +The hardest problems -- judging by experience with several theorem provers, +including matrix ones -- are 34 and 43. +*) + +writeln"Pelletier's examples"; +(*1*) +goal HOL.thy "(P-->Q) = (~Q --> ~P)"; +by (safe_meson_tac 1); +result(); + +(*2*) +goal HOL.thy "(~ ~ P) = P"; +by (safe_meson_tac 1); +result(); + +(*3*) +goal HOL.thy "~(P-->Q) --> (Q-->P)"; +by (safe_meson_tac 1); +result(); + +(*4*) +goal HOL.thy "(~P-->Q) = (~Q --> P)"; +by (safe_meson_tac 1); +result(); + +(*5*) +goal HOL.thy "((P|Q)-->(P|R)) --> (P|(Q-->R))"; +by (safe_meson_tac 1); +result(); + +(*6*) +goal HOL.thy "P | ~ P"; +by (safe_meson_tac 1); +result(); + +(*7*) +goal HOL.thy "P | ~ ~ ~ P"; +by (safe_meson_tac 1); +result(); + +(*8. Peirce's law*) +goal HOL.thy "((P-->Q) --> P) --> P"; +by (safe_meson_tac 1); +result(); + +(*9*) +goal HOL.thy "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (safe_meson_tac 1); +result(); + +(*10*) +goal HOL.thy "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P=Q)"; +by (safe_meson_tac 1); +result(); + +(*11. Proved in each direction (incorrectly, says Pelletier!!) *) +goal HOL.thy "P=P::bool"; +by (safe_meson_tac 1); +result(); + +(*12. "Dijkstra's law"*) +goal HOL.thy "((P = Q) = R) = (P = (Q = R))"; +by (best_meson_tac size_of_subgoals 1); +result(); + +(*13. Distributive law*) +goal HOL.thy "(P | (Q & R)) = ((P | Q) & (P | R))"; +by (safe_meson_tac 1); +result(); + +(*14*) +goal HOL.thy "(P = Q) = ((Q | ~P) & (~Q|P))"; +by (safe_meson_tac 1); +result(); + +(*15*) +goal HOL.thy "(P --> Q) = (~P | Q)"; +by (safe_meson_tac 1); +result(); + +(*16*) +goal HOL.thy "(P-->Q) | (Q-->P)"; +by (safe_meson_tac 1); +result(); + +(*17*) +goal HOL.thy "((P & (Q-->R))-->S) = ((~P | Q | S) & (~P | ~R | S))"; +by (safe_meson_tac 1); +result(); + +writeln"Classical Logic: examples with quantifiers"; + +goal HOL.thy "(! x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "(? x. P-->Q(x)) = (P --> (? x.Q(x)))"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "(? x.P(x)-->Q) = ((! x.P(x)) --> Q)"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "((! x.P(x)) | Q) = (! x. P(x) | Q)"; +by (safe_meson_tac 1); +result(); + +writeln"Testing the complete tactic"; + +(*Not provable by pc_tac: needs multiple instantiation of !. + Could be proved trivially by a PROLOG interpreter*) +goal HOL.thy "(! x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; +by (safe_meson_tac 1); +result(); + +(*Not provable by pc_tac: needs double instantiation of EXISTS*) +goal HOL.thy "? x. P(x) --> P(a) & P(b)"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "? z. P(z) --> (! x. P(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Hard examples with quantifiers"; + +writeln"Problem 18"; +goal HOL.thy "? y. ! x. P(y)-->P(x)"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 19"; +goal HOL.thy "? x. ! y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 20"; +goal HOL.thy "(! x y. ? z. ! w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (? x y. P(x) & Q(y)) --> (? z. R(z))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 21"; +goal HOL.thy "(? x. P-->Q(x)) & (? x. Q(x)-->P) --> (? x. P=Q(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 22"; +goal HOL.thy "(! x. P = Q(x)) --> (P = (! x. Q(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 23"; +goal HOL.thy "(! x. P | Q(x)) = (P | (! x. Q(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 24"; +goal HOL.thy "~(? x. S(x)&Q(x)) & (! x. P(x) --> Q(x)|R(x)) & \ +\ ~(? x.P(x)) --> (? x.Q(x)) & (! x. Q(x)|R(x) --> S(x)) \ +\ --> (? x. P(x)&R(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 25"; +goal HOL.thy "(? x. P(x)) & \ +\ (! x. L(x) --> ~ (M(x) & R(x))) & \ +\ (! x. P(x) --> (M(x) & L(x))) & \ +\ ((! x. P(x)-->Q(x)) | (? x. P(x)&R(x))) \ +\ --> (? x. Q(x)&P(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 26"; +goal HOL.thy "((? x. p(x)) = (? x. q(x))) & \ +\ (! x. ! y. p(x) & q(y) --> (r(x) = s(y))) \ +\ --> ((! x. p(x)-->r(x)) = (! x. q(x)-->s(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 27"; +goal HOL.thy "(? x. P(x) & ~Q(x)) & \ +\ (! x. P(x) --> R(x)) & \ +\ (! x. M(x) & L(x) --> P(x)) & \ +\ ((? x. R(x) & ~ Q(x)) --> (! x. L(x) --> ~ R(x))) \ +\ --> (! x. M(x) --> ~L(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 28. AMENDED"; +goal HOL.thy "(! x. P(x) --> (! x. Q(x))) & \ +\ ((! x. Q(x)|R(x)) --> (? x. Q(x)&S(x))) & \ +\ ((? x.S(x)) --> (! x. L(x) --> M(x))) \ +\ --> (! x. P(x) & L(x) --> M(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal HOL.thy "(? x. F(x)) & (? y. G(y)) \ +\ --> ( ((! x. F(x)-->H(x)) & (! y. G(y)-->J(y))) = \ +\ (! x y. F(x) & G(y) --> H(x) & J(y)))"; +by (safe_meson_tac 1); (*5 secs*) +result(); + +writeln"Problem 30"; +goal HOL.thy "(! x. P(x) | Q(x) --> ~ R(x)) & \ +\ (! x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (! x. S(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 31"; +goal HOL.thy "~(? x.P(x) & (Q(x) | R(x))) & \ +\ (? x. L(x) & P(x)) & \ +\ (! x. ~ R(x) --> M(x)) \ +\ --> (? x. L(x) & M(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 32"; +goal HOL.thy "(! x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (! x. S(x) & R(x) --> L(x)) & \ +\ (! x. M(x) --> R(x)) \ +\ --> (! x. P(x) & M(x) --> L(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 33"; +goal HOL.thy "(! x. P(a) & (P(x)-->P(b))-->P(c)) = \ +\ (! x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; +by (safe_meson_tac 1); (*5.6 secs*) +result(); + +writeln"Problem 34 AMENDED (TWICE!!)"; +(*Andrews's challenge*) +goal HOL.thy "((? x. ! y. p(x) = p(y)) = \ +\ ((? x. q(x)) = (! y. p(y)))) = \ +\ ((? x. ! y. q(x) = q(y)) = \ +\ ((? x. p(x)) = (! y. q(y))))"; +by (safe_meson_tac 1); (*90 secs*) +result(); + +writeln"Problem 35"; +goal HOL.thy "? x y. P(x,y) --> (! u v. P(u,v))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 36"; +goal HOL.thy "(! x. ? y. J(x,y)) & \ +\ (! x. ? y. G(x,y)) & \ +\ (! x y. J(x,y) | G(x,y) --> \ +\ (! z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (! x. ? y. H(x,y))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 37"; +goal HOL.thy "(! z. ? w. ! x. ? y. \ +\ (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (? u.Q(u,w)))) & \ +\ (! x z. ~P(x,z) --> (? y. Q(y,z))) & \ +\ ((? x y. Q(x,y)) --> (! x. R(x,x))) \ +\ --> (! x. ? y. R(x,y))"; +by (safe_meson_tac 1); (*causes unification tracing messages*) +result(); + +writeln"Problem 38"; +goal HOL.thy + "(! x. p(a) & (p(x) --> (? y. p(y) & r(x,y))) --> \ +\ (? z. ? w. p(z) & r(x,w) & r(w,z))) = \ +\ (! x. (~p(a) | p(x) | (? z. ? w. p(z) & r(x,w) & r(w,z))) & \ +\ (~p(a) | ~(? y. p(y) & r(x,y)) | \ +\ (? z. ? w. p(z) & r(x,w) & r(w,z))))"; +by (safe_meson_tac 1); (*62 secs*) +result(); + +writeln"Problem 39"; +goal HOL.thy "~ (? x. ! y. F(y,x) = (~F(y,y)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 40. AMENDED"; +goal HOL.thy "(? y. ! x. F(x,y) = F(x,x)) \ +\ --> ~ (! x. ? y. ! z. F(z,y) = (~F(z,x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 41"; +goal HOL.thy "(! z. (? y. (! x. f(x,y) = (f(x,z) & ~ f(x,x))))) \ +\ --> ~ (? z. ! x. f(x,z))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 42"; +goal HOL.thy "~ (? y. ! x. p(x,y) = (~ (? z. p(x,z) & p(z,x))))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 43 NOW PROVED AUTOMATICALLY!!"; +goal HOL.thy "(! x. ! y. q(x,y) = (! z. p(z,x) = p(z,y)::bool)) \ +\ --> (! x. (! y. q(x,y) = q(y,x)::bool))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 44"; +goal HOL.thy "(! x. f(x) --> \ +\ (? y. g(y) & h(x,y) & (? y. g(y) & ~ h(x,y)))) & \ +\ (? x. j(x) & (! y. g(y) --> h(x,y))) \ +\ --> (? x. j(x) & ~f(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 45"; +goal HOL.thy "(! x. f(x) & (! y. g(y) & h(x,y) --> j(x,y)) \ +\ --> (! y. g(y) & h(x,y) --> k(y))) & \ +\ ~ (? y. l(y) & k(y)) & \ +\ (? x. f(x) & (! y. h(x,y) --> l(y)) \ +\ & (! y. g(y) & h(x,y) --> j(x,y))) \ +\ --> (? x. f(x) & ~ (? y. g(y) & h(x,y)))"; +by (safe_meson_tac 1); (*11 secs*) +result(); + +writeln"Problem 46"; +goal HOL.thy + "(! x. f(x) & (! y. f(y) & h(y,x) --> g(y)) --> g(x)) & \ +\ ((? x.f(x) & ~g(x)) --> \ +\ (? x. f(x) & ~g(x) & (! y. f(y) & ~g(y) --> j(x,y)))) & \ +\ (! x y. f(x) & f(y) & h(x,y) --> ~j(y,x)) \ +\ --> (! x. f(x) --> g(x))"; +by (safe_meson_tac 1); (*11 secs*) +result(); + +(* Example suggested by Johannes Schumann and credited to Pelletier *) +goal HOL.thy "(!x y z. P(x,y) --> P(y,z) --> P(x,z)) --> \ +\ (!x y z. Q(x,y) --> Q(y,z) --> Q(x,z)) --> \ +\ (!x y.Q(x,y) --> Q(y,x)) --> (!x y. P(x,y) | Q(x,y)) --> \ +\ (!x y.P(x,y)) | (!x y.Q(x,y))"; +by (safe_meson_tac 1); (*32 secs*) +result(); + +writeln"Problem 50"; +(*What has this to do with equality?*) +goal HOL.thy "(! x. P(a,x) | (! y.P(x,y))) --> (? x. ! y.P(x,y))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 55"; + +(*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988). + meson_tac cannot report who killed Agatha. *) +goal HOL.thy "lives(agatha) & lives(butler) & lives(charles) & \ +\ (killed(agatha,agatha) | killed(butler,agatha) | killed(charles,agatha)) & \ +\ (!x y. killed(x,y) --> hates(x,y) & ~richer(x,y)) & \ +\ (!x. hates(agatha,x) --> ~hates(charles,x)) & \ +\ (hates(agatha,agatha) & hates(agatha,charles)) & \ +\ (!x. lives(x) & ~richer(x,agatha) --> hates(butler,x)) & \ +\ (!x. hates(agatha,x) --> hates(butler,x)) & \ +\ (!x. ~hates(x,agatha) | ~hates(x,butler) | ~hates(x,charles)) --> \ +\ (? x. killed(x,agatha))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 57"; +goal HOL.thy + "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (! x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 58"; +(* Challenge found on info-hol *) +goal HOL.thy + "! P Q R x. ? v w. ! y z. P(x) & Q(y) --> (P(v) | R(w)) & (R(z) --> Q(v))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 59"; +goal HOL.thy "(! x. P(x) = (~P(f(x)))) --> (? x. P(x) & ~P(f(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 60"; +goal HOL.thy "! x. P(x,f(x)) = (? y. (! z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (safe_meson_tac 1); +result(); + +writeln"Reached end of file."; + +(*26 August 1992: loaded in 277 secs. New Jersey v 75*) diff -r 000000000000 -r 7949f97df77a ex/meson.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/meson.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,417 @@ +(* Title: HOL/ex/meson + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The MESON resolution proof procedure for HOL + +When making clauses, avoids using the rewriter -- instead uses RS recursively +*) + +writeln"File HOL/ex/meson."; + +(*Prove theorems using fast_tac*) +fun prove_fun s = + prove_goal HOL.thy s + (fn prems => [ cut_facts_tac prems 1, fast_tac HOL_cs 1 ]); + +(**** Negation Normal Form ****) + +(*** de Morgan laws ***) + +val not_conjD = prove_fun "~(P&Q) ==> ~P | ~Q"; +val not_disjD = prove_fun "~(P|Q) ==> ~P & ~Q"; +val not_notD = prove_fun "~~P ==> P"; +val not_allD = prove_fun "~(! x.P(x)) ==> ? x. ~P(x)"; +val not_exD = prove_fun "~(? x.P(x)) ==> ! x. ~P(x)"; + + +(*** Removal of --> and <-> (positive and negative occurrences) ***) + +val imp_to_disjD = prove_fun "P-->Q ==> ~P | Q"; +val not_impD = prove_fun "~(P-->Q) ==> P & ~Q"; + +val iff_to_disjD = prove_fun "P=Q ==> (~P | Q) & (~Q | P)"; + +(*Much more efficient than (P & ~Q) | (Q & ~P) for computing CNF*) +val not_iffD = prove_fun "~(P=Q) ==> (P | Q) & (~P | ~Q)"; + + +(**** Pulling out the existential quantifiers ****) + +(*** Conjunction ***) + +val conj_exD1 = prove_fun "(? x.P(x)) & Q ==> ? x. P(x) & Q"; +val conj_exD2 = prove_fun "P & (? x.Q(x)) ==> ? x. P & Q(x)"; + +(*** Disjunction ***) + +(*DO NOT USE with forall-Skolemization: makes fewer schematic variables!! + With ex-Skolemization, makes fewer Skolem constants*) +val disj_exD = prove_fun "(? x.P(x)) | (? x.Q(x)) ==> ? x. P(x) | Q(x)"; + +val disj_exD1 = prove_fun "(? x.P(x)) | Q ==> ? x. P(x) | Q"; +val disj_exD2 = prove_fun "P | (? x.Q(x)) ==> ? x. P | Q(x)"; + + +(**** Skolemization -- pulling "?" over "!" ****) + +(*"Axiom" of Choice, proved using the description operator*) +val [major] = goal HOL.thy + "! x. ? y. Q(x,y) ==> ? f. ! x. Q(x,f(x))"; +by (cut_facts_tac [major] 1); +by (fast_tac (HOL_cs addEs [selectI]) 1); +val choice = result(); + + +(***** Generating clauses for the Meson Proof Procedure *****) + +(*** Disjunctions ***) + +val disj_assoc = prove_fun "(P|Q)|R ==> P|(Q|R)"; + +val disj_comm = prove_fun "P|Q ==> Q|P"; + +val disj_FalseD1 = prove_fun "False|P ==> P"; +val disj_FalseD2 = prove_fun "P|False ==> P"; + +(*** Generation of contrapositives ***) + +(*Inserts negated disjunct after removing the negation; P is a literal*) +val [major,minor] = goal HOL.thy "~P|Q ==> ((~P==>P) ==> Q)"; +by (rtac (major RS disjE) 1); +by (rtac notE 1); +by (etac minor 2); +by (ALLGOALS assume_tac); +val make_neg_rule = result(); + +(*For Plaisted's "Postive refinement" of the MESON procedure*) +val [major,minor] = goal HOL.thy "~P|Q ==> (P ==> Q)"; +by (rtac (major RS disjE) 1); +by (rtac notE 1); +by (rtac minor 2); +by (ALLGOALS assume_tac); +val make_refined_neg_rule = result(); + +(*P should be a literal*) +val [major,minor] = goal HOL.thy "P|Q ==> ((P==>~P) ==> Q)"; +by (rtac (major RS disjE) 1); +by (rtac notE 1); +by (etac minor 1); +by (ALLGOALS assume_tac); +val make_pos_rule = result(); + +(*** Generation of a goal clause -- put away the final literal ***) + +val [major,minor] = goal HOL.thy "~P ==> ((~P==>P) ==> False)"; +by (rtac notE 1); +by (rtac minor 2); +by (ALLGOALS (rtac major)); +val make_neg_goal = result(); + +val [major,minor] = goal HOL.thy "P ==> ((P==>~P) ==> False)"; +by (rtac notE 1); +by (rtac minor 1); +by (ALLGOALS (rtac major)); +val make_pos_goal = result(); + + +(**** Lemmas for forward proof (like congruence rules) ****) + +(*NOTE: could handle conjunctions (faster?) by + nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *) +val major::prems = goal HOL.thy + "[| P'&Q'; P' ==> P; Q' ==> Q |] ==> P&Q"; +by (rtac (major RS conjE) 1); +by (rtac conjI 1); +by (ALLGOALS (eresolve_tac prems)); +val conj_forward = result(); + +val major::prems = goal HOL.thy + "[| P'|Q'; P' ==> P; Q' ==> Q |] ==> P|Q"; +by (rtac (major RS disjE) 1); +by (ALLGOALS (dresolve_tac prems)); +by (ALLGOALS (eresolve_tac [disjI1,disjI2])); +val disj_forward = result(); + +val major::prems = goal HOL.thy + "[| ! x. P'(x); !!x. P'(x) ==> P(x) |] ==> ! x. P(x)"; +by (rtac allI 1); +by (resolve_tac prems 1); +by (rtac (major RS spec) 1); +val all_forward = result(); + +val major::prems = goal HOL.thy + "[| ? x. P'(x); !!x. P'(x) ==> P(x) |] ==> ? x. P(x)"; +by (rtac (major RS exE) 1); +by (rtac exI 1); +by (eresolve_tac prems 1); +val ex_forward = result(); + + +(**** Operators for forward proof ****) + +(*raises exception if no rules apply -- unlike RL*) +fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls)) + | tryres (th, []) = raise THM("tryres", 0, [th]); + +val prop_of = #prop o rep_thm; + +(*Permits forward proof from rules that discharge assumptions*) +fun forward_res nf state = + case Sequence.pull + (tapply(ALLGOALS (METAHYPS (fn [prem] => rtac (nf prem) 1)), + state)) + of Some(th,_) => th + | None => raise THM("forward_res", 0, [state]); + + +(*Negation Normal Form*) +val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, + not_impD, not_iffD, not_allD, not_exD, not_notD]; +fun make_nnf th = make_nnf (tryres(th, nnf_rls)) + handle THM _ => + forward_res make_nnf + (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward])) + handle THM _ => th; + + +(*Are any of the constants in "bs" present in the term?*) +fun has_consts bs = + let fun has (Const(a,_)) = a mem bs + | has (f$u) = has f orelse has u + | has (Abs(_,_,t)) = has t + | has _ = false + in has end; + +(*Pull existential quantifiers (Skolemization)*) +fun skolemize th = + if not (has_consts ["Ex"] (prop_of th)) then th + else skolemize (tryres(th, [choice, conj_exD1, conj_exD2, + disj_exD, disj_exD1, disj_exD2])) + handle THM _ => + skolemize (forward_res skolemize + (tryres (th, [conj_forward, disj_forward, all_forward]))) + handle THM _ => forward_res skolemize (th RS ex_forward); + + +(**** Clause handling ****) + +fun literals (Const("Trueprop",_) $ P) = literals P + | literals (Const("op |",_) $ P $ Q) = literals P @ literals Q + | literals (Const("not",_) $ P) = [(false,P)] + | literals P = [(true,P)]; + +(*number of literals in a term*) +val nliterals = length o literals; + +(*to delete tautologous clauses*) +fun taut_lits [] = false + | taut_lits ((flg,t)::ts) = (not flg,t) mem ts orelse taut_lits ts; + +val is_taut = taut_lits o literals o prop_of; + + +(*Generation of unique names -- maxidx cannot be relied upon to increase! + Cannot rely on "variant", since variables might coincide when literals + are joined to make a clause... + 19 chooses "U" as the first variable name*) +val name_ref = ref 19; + +(*Replaces universally quantified variables by FREE variables -- because + assumptions may not contain scheme variables. Later, call "generalize". *) +fun freeze_spec th = + let val sth = th RS spec + val newname = (name_ref := !name_ref + 1; + radixstring(26, "A", !name_ref)) + in read_instantiate [("x", newname)] sth end; + +fun resop nf [prem] = resolve_tac (nf prem) 1; + +(*Conjunctive normal form, detecting tautologies early. + Strips universal quantifiers and breaks up conjunctions. *) +fun cnf_aux seen (th,ths) = + if taut_lits (literals(prop_of th) @ seen) then ths + else if not (has_consts ["All","op &"] (prop_of th)) then th::ths + else (*conjunction?*) + cnf_aux seen (th RS conjunct1, + cnf_aux seen (th RS conjunct2, ths)) + handle THM _ => (*universal quant?*) + cnf_aux seen (freeze_spec th, ths) + handle THM _ => (*disjunction?*) + let val tac = + (METAHYPS (resop (cnf_nil seen)) 1) THEN + (STATE (fn st' => + METAHYPS (resop (cnf_nil (literals (concl_of st') @ seen))) 1)) + in Sequence.list_of_s (tapply(tac, th RS disj_forward)) @ ths + end +and cnf_nil seen th = cnf_aux seen (th,[]); + +(*Top-level call to cnf -- it's safe to reset name_ref*) +fun cnf (th,ths) = + (name_ref := 19; cnf (th RS conjunct1, cnf (th RS conjunct2, ths)) + handle THM _ => (*not a conjunction*) cnf_aux [] (th, ths)); + +(**** Removal of duplicate literals ****) + +(*Version for removal of duplicate literals*) +val major::prems = goal HOL.thy + "[| P'|Q'; P' ==> P; [| Q'; P==>False |] ==> Q |] ==> P|Q"; +by (rtac (major RS disjE) 1); +by (rtac disjI1 1); +by (rtac (disjCI RS disj_comm) 2); +by (ALLGOALS (eresolve_tac prems)); +by (etac notE 1); +by (assume_tac 1); +val disj_forward2 = result(); + +(*Forward proof, passing extra assumptions as theorems to the tactic*) +fun forward_res2 nf hyps state = + case Sequence.pull + (tapply(REPEAT + (METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1), + state)) + of Some(th,_) => th + | None => raise THM("forward_res2", 0, [state]); + +(*Remove duplicates in P|Q by assuming ~P in Q + rls (initially []) accumulates assumptions of the form P==>False*) +fun nodups_aux rls th = nodups_aux rls (th RS disj_assoc) + handle THM _ => tryres(th,rls) + handle THM _ => tryres(forward_res2 nodups_aux rls (th RS disj_forward2), + [disj_FalseD1, disj_FalseD2, asm_rl]) + handle THM _ => th; + +(*Remove duplicate literals, if there are any*) +fun nodups th = + if null(findrep(literals(prop_of th))) then th + else nodups_aux [] th; + + +(**** Generation of contrapositives ****) + +(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) +fun assoc_right th = assoc_right (th RS disj_assoc) + handle THM _ => th; + +(*Must check for negative literal first!*) +val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; +val refined_clause_rules = [disj_assoc, make_refined_neg_rule, make_pos_rule]; + +(*Create a goal or support clause, conclusing False*) +fun make_goal th = (*Must check for negative literal first!*) + make_goal (tryres(th, clause_rules)) + handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); + +(*Sort clauses by number of literals*) +fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2); + +(*TAUTOLOGY CHECK SHOULD NOT BE NECESSARY!*) +fun sort_clauses ths = sort fewerlits (filter (not o is_taut) ths); + +(*Convert all suitable free variables to schematic variables*) +fun generalize th = forall_elim_vars 0 (forall_intr_frees th); + +(*make clauses from a list of theorems*) +fun make_clauses ths = + sort_clauses (map (generalize o nodups) (foldr cnf (ths,[]))); + +(*Create a Horn clause*) +fun make_horn crules th = make_horn crules (tryres(th,crules)) + handle THM _ => th; + +(*Generate Horn clauses for all contrapositives of a clause*) +fun add_contras crules (th,hcs) = + let fun rots (0,th) = hcs + | rots (k,th) = zero_var_indexes (make_horn crules th) :: + rots(k-1, assoc_right (th RS disj_comm)) + in case nliterals(prop_of th) of + 1 => th::hcs + | n => rots(n, assoc_right th) + end; + +(*Convert a list of clauses to (contrapositive) Horn clauses*) +fun make_horns ths = foldr (add_contras clause_rules) (ths,[]); + +(*Find an all-negative support clause*) +fun is_negative th = forall (not o #1) (literals (prop_of th)); + +val neg_clauses = filter is_negative; + + +(***** MESON PROOF PROCEDURE *****) + +fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi, + As) = rhyps(phi, A::As) + | rhyps (_, As) = As; + +(** Detecting repeated assumptions in a subgoal **) + +(*The stringtree detects repeated assumptions.*) +fun ins_term (net,t) = Net.insert_term((t,t), net, op aconv); + +(*detects repetitions in a list of terms*) +fun has_reps [] = false + | has_reps [_] = false + | has_reps [t,u] = (t aconv u) + | has_reps ts = (foldl ins_term (Net.empty, ts); false) + handle INSERT => true; + +(*Loop checking: FAIL if trying to prove the same thing twice + -- repeated literals*) +val check_tac = SUBGOAL (fn (prem,_) => + if has_reps (rhyps(prem,[])) then no_tac else all_tac); + +(* net_resolve_tac actually made it slower... *) +fun prolog_step_tac horns i = + (assume_tac i APPEND resolve_tac horns i) THEN + (ALLGOALS check_tac) THEN + (TRYALL eq_assume_tac); + + +(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) +local fun addconcl(prem,sz) = size_of_term (Logic.strip_assums_concl prem) + sz +in +fun size_of_subgoals st = foldr addconcl (prems_of st, 0) +end; + +(*Could simply use nprems_of, which would count remaining subgoals -- no + discrimination as to their size! With BEST_FIRST, fails for problem 41.*) + +fun best_prolog_tac sizef horns = + BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1); + +fun depth_prolog_tac horns = + DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1); + +(*Return all negative clauses, as possible goal clauses*) +fun gocls cls = map make_goal (neg_clauses cls); + + +fun skolemize_tac prems = + cut_facts_tac (map (skolemize o make_nnf) prems) THEN' + REPEAT o (etac exE); + +fun MESON sko_tac = SELECT_GOAL + (EVERY1 [rtac ccontr, + METAHYPS (fn negs => + EVERY1 [skolemize_tac negs, + METAHYPS (sko_tac o make_clauses)])]); + +fun best_meson_tac sizef = + MESON (fn cls => + resolve_tac (gocls cls) 1 + THEN_BEST_FIRST + (has_fewer_prems 1, sizef, + prolog_step_tac (make_horns cls) 1)); + +(*First, breaks the goal into independent units*) +val safe_meson_tac = + SELECT_GOAL (TRY (safe_tac HOL_cs) THEN + TRYALL (best_meson_tac size_of_subgoals)); + +val depth_meson_tac = + MESON (fn cls => EVERY [resolve_tac (gocls cls) 1, + depth_prolog_tac (make_horns cls)]); + +writeln"Reached end of file."; diff -r 000000000000 -r 7949f97df77a ex/mesontest.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/mesontest.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,496 @@ +(* Title: HOL/ex/meson + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Test data for the MESON proof procedure + (Excludes the equality problems 51, 52, 56, 58) + +show_hyps:=false; + +by (rtac ccontr 1); +val [prem] = gethyps 1; +val nnf = make_nnf prem; +val xsko = skolemize nnf; +by (cut_facts_tac [xsko] 1 THEN REPEAT (etac exE 1)); +val [_,sko] = gethyps 1; +val clauses = make_clauses [sko]; +val horns = make_horns clauses; +val go::_ = neg_clauses clauses; + +goal HOL.thy "False"; +by (rtac (make_goal go) 1); +by (prolog_step_tac horns 1); +by (depth_prolog_tac horns); +by (best_prolog_tac size_of_subgoals horns); +*) + +writeln"File HOL/ex/meson-test."; + +(**** Interactive examples ****) + +(*Generate nice names for Skolem functions*) +Logic.auto_rename := true; Logic.set_rename_prefix "a"; + + +writeln"Problem 25"; +goal HOL.thy "(? x. P(x)) & \ +\ (! x. L(x) --> ~ (M(x) & R(x))) & \ +\ (! x. P(x) --> (M(x) & L(x))) & \ +\ ((! x. P(x)-->Q(x)) | (? x. P(x)&R(x))) \ +\ --> (? x. Q(x)&P(x))"; +by (rtac ccontr 1); +val [prem25] = gethyps 1; +val nnf25 = make_nnf prem25; +val xsko25 = skolemize nnf25; +by (cut_facts_tac [xsko25] 1 THEN REPEAT (etac exE 1)); +val [_,sko25] = gethyps 1; +val clauses25 = make_clauses [sko25]; (*7 clauses*) +val horns25 = make_horns clauses25; (*16 Horn clauses*) +val go25::_ = neg_clauses clauses25; + +goal HOL.thy "False"; +by (rtac (make_goal go25) 1); +by (depth_prolog_tac horns25); + + +writeln"Problem 26"; +goal HOL.thy "((? x. p(x)) = (? x. q(x))) & \ +\ (! x. ! y. p(x) & q(y) --> (r(x) = s(y))) \ +\ --> ((! x. p(x)-->r(x)) = (! x. q(x)-->s(x)))"; +by (rtac ccontr 1); +val [prem26] = gethyps 1; +val nnf26 = make_nnf prem26; +val xsko26 = skolemize nnf26; +by (cut_facts_tac [xsko26] 1 THEN REPEAT (etac exE 1)); +val [_,sko26] = gethyps 1; +val clauses26 = make_clauses [sko26]; (*9 clauses*) +val horns26 = make_horns clauses26; (*24 Horn clauses*) +val go26::_ = neg_clauses clauses26; + +goal HOL.thy "False"; +by (rtac (make_goal go26) 1); +by (depth_prolog_tac horns26); (*6 secs*) + + + +writeln"Problem 43 NOW PROVED AUTOMATICALLY!!"; +goal HOL.thy "(! x. ! y. q(x,y) = (! z. p(z,x) = p(z,y)::bool)) \ +\ --> (! x. (! y. q(x,y) = q(y,x)::bool))"; +by (rtac ccontr 1); +val [prem43] = gethyps 1; +val nnf43 = make_nnf prem43; +val xsko43 = skolemize nnf43; +by (cut_facts_tac [xsko43] 1 THEN REPEAT (etac exE 1)); +val [_,sko43] = gethyps 1; +val clauses43 = make_clauses [sko43]; (*6*) +val horns43 = make_horns clauses43; (*16*) +val go43::_ = neg_clauses clauses43; + +goal HOL.thy "False"; +by (rtac (make_goal go43) 1); +by (best_prolog_tac size_of_subgoals horns43); +(*8.7 secs*) + + +(*Restore variable name preservation*) +Logic.auto_rename := false; + + +(**** Batch test data ****) + +(*Sample problems from + F. J. Pelletier, + Seventy-Five Problems for Testing Automatic Theorem Provers, + J. Automated Reasoning 2 (1986), 191-216. + Errata, JAR 4 (1988), 236-236. + +The hardest problems -- judging by experience with several theorem provers, +including matrix ones -- are 34 and 43. +*) + +writeln"Pelletier's examples"; +(*1*) +goal HOL.thy "(P-->Q) = (~Q --> ~P)"; +by (safe_meson_tac 1); +result(); + +(*2*) +goal HOL.thy "(~ ~ P) = P"; +by (safe_meson_tac 1); +result(); + +(*3*) +goal HOL.thy "~(P-->Q) --> (Q-->P)"; +by (safe_meson_tac 1); +result(); + +(*4*) +goal HOL.thy "(~P-->Q) = (~Q --> P)"; +by (safe_meson_tac 1); +result(); + +(*5*) +goal HOL.thy "((P|Q)-->(P|R)) --> (P|(Q-->R))"; +by (safe_meson_tac 1); +result(); + +(*6*) +goal HOL.thy "P | ~ P"; +by (safe_meson_tac 1); +result(); + +(*7*) +goal HOL.thy "P | ~ ~ ~ P"; +by (safe_meson_tac 1); +result(); + +(*8. Peirce's law*) +goal HOL.thy "((P-->Q) --> P) --> P"; +by (safe_meson_tac 1); +result(); + +(*9*) +goal HOL.thy "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"; +by (safe_meson_tac 1); +result(); + +(*10*) +goal HOL.thy "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P=Q)"; +by (safe_meson_tac 1); +result(); + +(*11. Proved in each direction (incorrectly, says Pelletier!!) *) +goal HOL.thy "P=P::bool"; +by (safe_meson_tac 1); +result(); + +(*12. "Dijkstra's law"*) +goal HOL.thy "((P = Q) = R) = (P = (Q = R))"; +by (best_meson_tac size_of_subgoals 1); +result(); + +(*13. Distributive law*) +goal HOL.thy "(P | (Q & R)) = ((P | Q) & (P | R))"; +by (safe_meson_tac 1); +result(); + +(*14*) +goal HOL.thy "(P = Q) = ((Q | ~P) & (~Q|P))"; +by (safe_meson_tac 1); +result(); + +(*15*) +goal HOL.thy "(P --> Q) = (~P | Q)"; +by (safe_meson_tac 1); +result(); + +(*16*) +goal HOL.thy "(P-->Q) | (Q-->P)"; +by (safe_meson_tac 1); +result(); + +(*17*) +goal HOL.thy "((P & (Q-->R))-->S) = ((~P | Q | S) & (~P | ~R | S))"; +by (safe_meson_tac 1); +result(); + +writeln"Classical Logic: examples with quantifiers"; + +goal HOL.thy "(! x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "(? x. P-->Q(x)) = (P --> (? x.Q(x)))"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "(? x.P(x)-->Q) = ((! x.P(x)) --> Q)"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "((! x.P(x)) | Q) = (! x. P(x) | Q)"; +by (safe_meson_tac 1); +result(); + +writeln"Testing the complete tactic"; + +(*Not provable by pc_tac: needs multiple instantiation of !. + Could be proved trivially by a PROLOG interpreter*) +goal HOL.thy "(! x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"; +by (safe_meson_tac 1); +result(); + +(*Not provable by pc_tac: needs double instantiation of EXISTS*) +goal HOL.thy "? x. P(x) --> P(a) & P(b)"; +by (safe_meson_tac 1); +result(); + +goal HOL.thy "? z. P(z) --> (! x. P(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Hard examples with quantifiers"; + +writeln"Problem 18"; +goal HOL.thy "? y. ! x. P(y)-->P(x)"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 19"; +goal HOL.thy "? x. ! y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 20"; +goal HOL.thy "(! x y. ? z. ! w. (P(x)&Q(y)-->R(z)&S(w))) \ +\ --> (? x y. P(x) & Q(y)) --> (? z. R(z))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 21"; +goal HOL.thy "(? x. P-->Q(x)) & (? x. Q(x)-->P) --> (? x. P=Q(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 22"; +goal HOL.thy "(! x. P = Q(x)) --> (P = (! x. Q(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 23"; +goal HOL.thy "(! x. P | Q(x)) = (P | (! x. Q(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 24"; +goal HOL.thy "~(? x. S(x)&Q(x)) & (! x. P(x) --> Q(x)|R(x)) & \ +\ ~(? x.P(x)) --> (? x.Q(x)) & (! x. Q(x)|R(x) --> S(x)) \ +\ --> (? x. P(x)&R(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 25"; +goal HOL.thy "(? x. P(x)) & \ +\ (! x. L(x) --> ~ (M(x) & R(x))) & \ +\ (! x. P(x) --> (M(x) & L(x))) & \ +\ ((! x. P(x)-->Q(x)) | (? x. P(x)&R(x))) \ +\ --> (? x. Q(x)&P(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 26"; +goal HOL.thy "((? x. p(x)) = (? x. q(x))) & \ +\ (! x. ! y. p(x) & q(y) --> (r(x) = s(y))) \ +\ --> ((! x. p(x)-->r(x)) = (! x. q(x)-->s(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 27"; +goal HOL.thy "(? x. P(x) & ~Q(x)) & \ +\ (! x. P(x) --> R(x)) & \ +\ (! x. M(x) & L(x) --> P(x)) & \ +\ ((? x. R(x) & ~ Q(x)) --> (! x. L(x) --> ~ R(x))) \ +\ --> (! x. M(x) --> ~L(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 28. AMENDED"; +goal HOL.thy "(! x. P(x) --> (! x. Q(x))) & \ +\ ((! x. Q(x)|R(x)) --> (? x. Q(x)&S(x))) & \ +\ ((? x.S(x)) --> (! x. L(x) --> M(x))) \ +\ --> (! x. P(x) & L(x) --> M(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 29. Essentially the same as Principia Mathematica *11.71"; +goal HOL.thy "(? x. F(x)) & (? y. G(y)) \ +\ --> ( ((! x. F(x)-->H(x)) & (! y. G(y)-->J(y))) = \ +\ (! x y. F(x) & G(y) --> H(x) & J(y)))"; +by (safe_meson_tac 1); (*5 secs*) +result(); + +writeln"Problem 30"; +goal HOL.thy "(! x. P(x) | Q(x) --> ~ R(x)) & \ +\ (! x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \ +\ --> (! x. S(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 31"; +goal HOL.thy "~(? x.P(x) & (Q(x) | R(x))) & \ +\ (? x. L(x) & P(x)) & \ +\ (! x. ~ R(x) --> M(x)) \ +\ --> (? x. L(x) & M(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 32"; +goal HOL.thy "(! x. P(x) & (Q(x)|R(x))-->S(x)) & \ +\ (! x. S(x) & R(x) --> L(x)) & \ +\ (! x. M(x) --> R(x)) \ +\ --> (! x. P(x) & M(x) --> L(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 33"; +goal HOL.thy "(! x. P(a) & (P(x)-->P(b))-->P(c)) = \ +\ (! x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"; +by (safe_meson_tac 1); (*5.6 secs*) +result(); + +writeln"Problem 34 AMENDED (TWICE!!)"; +(*Andrews's challenge*) +goal HOL.thy "((? x. ! y. p(x) = p(y)) = \ +\ ((? x. q(x)) = (! y. p(y)))) = \ +\ ((? x. ! y. q(x) = q(y)) = \ +\ ((? x. p(x)) = (! y. q(y))))"; +by (safe_meson_tac 1); (*90 secs*) +result(); + +writeln"Problem 35"; +goal HOL.thy "? x y. P(x,y) --> (! u v. P(u,v))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 36"; +goal HOL.thy "(! x. ? y. J(x,y)) & \ +\ (! x. ? y. G(x,y)) & \ +\ (! x y. J(x,y) | G(x,y) --> \ +\ (! z. J(y,z) | G(y,z) --> H(x,z))) \ +\ --> (! x. ? y. H(x,y))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 37"; +goal HOL.thy "(! z. ? w. ! x. ? y. \ +\ (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (? u.Q(u,w)))) & \ +\ (! x z. ~P(x,z) --> (? y. Q(y,z))) & \ +\ ((? x y. Q(x,y)) --> (! x. R(x,x))) \ +\ --> (! x. ? y. R(x,y))"; +by (safe_meson_tac 1); (*causes unification tracing messages*) +result(); + +writeln"Problem 38"; +goal HOL.thy + "(! x. p(a) & (p(x) --> (? y. p(y) & r(x,y))) --> \ +\ (? z. ? w. p(z) & r(x,w) & r(w,z))) = \ +\ (! x. (~p(a) | p(x) | (? z. ? w. p(z) & r(x,w) & r(w,z))) & \ +\ (~p(a) | ~(? y. p(y) & r(x,y)) | \ +\ (? z. ? w. p(z) & r(x,w) & r(w,z))))"; +by (safe_meson_tac 1); (*62 secs*) +result(); + +writeln"Problem 39"; +goal HOL.thy "~ (? x. ! y. F(y,x) = (~F(y,y)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 40. AMENDED"; +goal HOL.thy "(? y. ! x. F(x,y) = F(x,x)) \ +\ --> ~ (! x. ? y. ! z. F(z,y) = (~F(z,x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 41"; +goal HOL.thy "(! z. (? y. (! x. f(x,y) = (f(x,z) & ~ f(x,x))))) \ +\ --> ~ (? z. ! x. f(x,z))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 42"; +goal HOL.thy "~ (? y. ! x. p(x,y) = (~ (? z. p(x,z) & p(z,x))))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 43 NOW PROVED AUTOMATICALLY!!"; +goal HOL.thy "(! x. ! y. q(x,y) = (! z. p(z,x) = p(z,y)::bool)) \ +\ --> (! x. (! y. q(x,y) = q(y,x)::bool))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 44"; +goal HOL.thy "(! x. f(x) --> \ +\ (? y. g(y) & h(x,y) & (? y. g(y) & ~ h(x,y)))) & \ +\ (? x. j(x) & (! y. g(y) --> h(x,y))) \ +\ --> (? x. j(x) & ~f(x))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 45"; +goal HOL.thy "(! x. f(x) & (! y. g(y) & h(x,y) --> j(x,y)) \ +\ --> (! y. g(y) & h(x,y) --> k(y))) & \ +\ ~ (? y. l(y) & k(y)) & \ +\ (? x. f(x) & (! y. h(x,y) --> l(y)) \ +\ & (! y. g(y) & h(x,y) --> j(x,y))) \ +\ --> (? x. f(x) & ~ (? y. g(y) & h(x,y)))"; +by (safe_meson_tac 1); (*11 secs*) +result(); + +writeln"Problem 46"; +goal HOL.thy + "(! x. f(x) & (! y. f(y) & h(y,x) --> g(y)) --> g(x)) & \ +\ ((? x.f(x) & ~g(x)) --> \ +\ (? x. f(x) & ~g(x) & (! y. f(y) & ~g(y) --> j(x,y)))) & \ +\ (! x y. f(x) & f(y) & h(x,y) --> ~j(y,x)) \ +\ --> (! x. f(x) --> g(x))"; +by (safe_meson_tac 1); (*11 secs*) +result(); + +(* Example suggested by Johannes Schumann and credited to Pelletier *) +goal HOL.thy "(!x y z. P(x,y) --> P(y,z) --> P(x,z)) --> \ +\ (!x y z. Q(x,y) --> Q(y,z) --> Q(x,z)) --> \ +\ (!x y.Q(x,y) --> Q(y,x)) --> (!x y. P(x,y) | Q(x,y)) --> \ +\ (!x y.P(x,y)) | (!x y.Q(x,y))"; +by (safe_meson_tac 1); (*32 secs*) +result(); + +writeln"Problem 50"; +(*What has this to do with equality?*) +goal HOL.thy "(! x. P(a,x) | (! y.P(x,y))) --> (? x. ! y.P(x,y))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 55"; + +(*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988). + meson_tac cannot report who killed Agatha. *) +goal HOL.thy "lives(agatha) & lives(butler) & lives(charles) & \ +\ (killed(agatha,agatha) | killed(butler,agatha) | killed(charles,agatha)) & \ +\ (!x y. killed(x,y) --> hates(x,y) & ~richer(x,y)) & \ +\ (!x. hates(agatha,x) --> ~hates(charles,x)) & \ +\ (hates(agatha,agatha) & hates(agatha,charles)) & \ +\ (!x. lives(x) & ~richer(x,agatha) --> hates(butler,x)) & \ +\ (!x. hates(agatha,x) --> hates(butler,x)) & \ +\ (!x. ~hates(x,agatha) | ~hates(x,butler) | ~hates(x,charles)) --> \ +\ (? x. killed(x,agatha))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 57"; +goal HOL.thy + "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \ +\ (! x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 58"; +(* Challenge found on info-hol *) +goal HOL.thy + "! P Q R x. ? v w. ! y z. P(x) & Q(y) --> (P(v) | R(w)) & (R(z) --> Q(v))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 59"; +goal HOL.thy "(! x. P(x) = (~P(f(x)))) --> (? x. P(x) & ~P(f(x)))"; +by (safe_meson_tac 1); +result(); + +writeln"Problem 60"; +goal HOL.thy "! x. P(x,f(x)) = (? y. (! z. P(z,y) --> P(z,f(x))) & P(x,y))"; +by (safe_meson_tac 1); +result(); + +writeln"Reached end of file."; + +(*26 August 1992: loaded in 277 secs. New Jersey v 75*) diff -r 000000000000 -r 7949f97df77a ex/pl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/pl.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,344 @@ +(* Title: HOL/ex/prop-log.ML + ID: $Id$ + Author: Tobias Nipkow & Lawrence C Paulson + Copyright 1993 TU Muenchen & University of Cambridge + +For ex/prop-log.thy. Inductive definition of propositional logic. +Soundness and completeness w.r.t. truth-tables. + +Prove: If H|=p then G|=p where G:Fin(H) +*) + +open PL; + +val rule_defs = [axK_def, axS_def, axDN_def, ruleMP_def]; + + +(** Monotonicity and unfolding of the function **) + +goalw PL.thy rule_defs "mono(%X. H Un axK Un axS Un axDN Un ruleMP(X))"; +by (rtac monoI 1); +by(fast_tac set_cs 1); +val thms_bnd_mono = result(); + +goalw PL.thy [thms_def] "!!G H. G<=H ==> thms(G) <= thms(H)"; +by (REPEAT (ares_tac [subset_refl, Un_mono, lfp_mono] 1)); +val thms_mono = result(); + +(** Introduction rules for the consequence relation **) + +(* thms(H) = H Int Un axK Un axS Un ruleMP(thms(H)) *) +val thms_unfold = thms_bnd_mono RS (thms_def RS def_lfp_Tarski); + +(*Proof by hypothesis*) +val prems = goalw PL.thy [conseq_def] "p:H ==> H |- p"; +by (rtac (thms_unfold RS ssubst) 1); +by (fast_tac (set_cs addSIs prems) 1); +val conseq_H = result(); + +(*Proof by axiom K*) +goalw PL.thy [conseq_def] "H |- p->q->p"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac axK_def); +by (fast_tac set_cs 1); +val conseq_K = result(); + +(*Proof by axiom S*) +goalw PL.thy [conseq_def] "H |- (p->q->r) -> (p->q) -> p -> r"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac axS_def); +by (fast_tac set_cs 1); +val conseq_S = result(); + +(*Proof by axiom DN (double negation) *) +goalw PL.thy [conseq_def] "H |- ((p->false) -> false) -> p"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac axDN_def); +by (fast_tac set_cs 1); +val conseq_DN = result(); + +(*Proof by rule MP (Modus Ponens) *) +val [prempq,premp] = goalw PL.thy [conseq_def] + "[| H |- p->q; H |- p |] ==> H |- q"; +by (rtac (thms_unfold RS ssubst) 1); +by (rewtac ruleMP_def); +by (fast_tac (set_cs addSIs [premp,prempq]) 1); +val conseq_MP = result(); + +(*Rule is called I for Identity Combinator, not for Introduction*) +goal PL.thy "H |- p->p"; +by (rtac (conseq_S RS conseq_MP RS conseq_MP) 1); +by (rtac conseq_K 2); +by (rtac conseq_K 1); +val conseq_I = result(); + +(** Weakening, left and right **) + +(*This order of premises is convenient with RS*) +val prems = goalw PL.thy [conseq_def] "[| G<=H; G |- p |] ==> H |- p"; +by (rtac (thms_mono RS subsetD) 1); +by (REPEAT (resolve_tac prems 1)); +val weaken_left = result(); + +(* H |- p ==> insert(a,H) |- p *) +val weaken_left_insert = subset_insertI RS weaken_left; + +val weaken_left_Un1 = Un_upper1 RS weaken_left; +val weaken_left_Un2 = Un_upper2 RS weaken_left; + +val prems = goal PL.thy "H |- q ==> H |- p->q"; +by (rtac (conseq_K RS conseq_MP) 1); +by (REPEAT (resolve_tac prems 1)); +val weaken_right = result(); + +(** Rule induction for H|-p **) + +(*Careful unfolding/folding to avoid a big expansion*) +val major::prems = goalw PL.thy [conseq_def] + "[| H |- a; \ +\ !!x. x:H ==> P(x); \ +\ !!x y. P(x->y->x); \ +\ !!x y z. P((x->y->z)->(x->y)->x->z); \ +\ !!x. P(((x->false)->false)->x); \ +\ !!x y. [| H |- x->y; H |- x; P(x->y); P(x) |] ==> P(y) \ +\ |] ==> P(a)"; +by (rtac (major RS (thms_def RS def_induct)) 1); +by (rtac thms_bnd_mono 1); +by (rewrite_tac rule_defs); +by (fast_tac (set_cs addIs prems) 1); +val conseq_induct = result(); + +(*The deduction theorem*) +val [major] = goal PL.thy "insert(p,H) |- q ==> H |- p->q"; +by (rtac (major RS conseq_induct) 1); +by (fast_tac (set_cs addIs [conseq_I, conseq_H RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_K RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_S RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_DN RS weaken_right]) 1); +by (fast_tac (set_cs addIs [conseq_S RS conseq_MP RS conseq_MP]) 1); +val deduction = result(); + + +(*The cut rule*) +val prems = goal PL.thy "[| H|-p; insert(p,H) |- q |] ==> H |- q"; +by (rtac (deduction RS conseq_MP) 1); +by (REPEAT (resolve_tac prems 1)); +val cut = result(); + +val prems = goal PL.thy "H |- false ==> H |- p"; +by (rtac (conseq_DN RS conseq_MP) 1); +by (rtac weaken_right 1); +by (resolve_tac prems 1); +val conseq_falseE = result(); + +(* [| H |- p->false; H |- p; q: pl |] ==> H |- q *) +val conseq_notE = standard (conseq_MP RS conseq_falseE); + +(** The function eval **) + +val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp]; + +goalw PL.thy [eval_def] "tt[false] = False"; +by (simp_tac pl_ss 1); +val eval_false = result(); + +goalw PL.thy [eval_def] "tt[#v] = (v:tt)"; +by (simp_tac pl_ss 1); +val eval_var = result(); + +goalw PL.thy [eval_def] "tt[p->q] = (tt[p]-->tt[q])"; +by (simp_tac pl_ss 1); +val eval_imp = result(); + +val pl_ss = pl_ss addsimps [eval_false, eval_var, eval_imp]; + +(** The function hyps **) + +goalw PL.thy [hyps_def] "hyps(false,tt) = {}"; +by (simp_tac pl_ss 1); +val hyps_false = result(); + +goalw PL.thy [hyps_def] "hyps(#v,tt) = {if(v:tt, #v, #v->false)}"; +by (simp_tac pl_ss 1); +val hyps_var = result(); + +goalw PL.thy [hyps_def] "hyps(p->q,tt) = hyps(p,tt) Un hyps(q,tt)"; +by (simp_tac pl_ss 1); +val hyps_imp = result(); + +val pl_ss = pl_ss addsimps [hyps_false, hyps_var, hyps_imp]; + +val ball_eq = prove_goalw Set.thy [Ball_def] "(!x:A.P(x)) = (!x.x:A --> P(x))" + (fn _ => [rtac refl 1]); + +(*Soundness of the rules wrt truth-table semantics*) +val [major] = goalw PL.thy [sat_def] "H |- p ==> H |= p"; +by (rtac (major RS conseq_induct) 1); +by (fast_tac (set_cs addSDs [eval_imp RS iffD1 RS mp]) 5); +by (ALLGOALS (asm_simp_tac(pl_ss addsimps + [ball_eq,not_def RS fun_cong RS sym]))); +val soundness = result(); + +(** Structural induction on pl + +val major::prems = goalw PL.thy pl_defs + "[| q: pl; \ +\ P(false); \ +\ !!v. v:nat ==> P(#v); \ +\ !!q1 q2. [| q1: pl; q2: pl; P(q1); P(q2) |] ==> P(q1->q2) \ +\ |] ==> P(q)"; +by (rtac (major RS sexp_induct) 1); +by (etac nat_induct 1); +by (REPEAT (ares_tac prems 1)); +val pl_induct = result(); + **) +(*** Towards the completeness proof ***) + +val [premf] = goal PL.thy "H |- p->false ==> H |- p->q"; +by (rtac deduction 1); +by (rtac (premf RS weaken_left_insert RS conseq_notE) 1); +by (rtac conseq_H 1); +by (rtac insertI1 1); +val false_imp = result(); + +val [premp,premq] = goal PL.thy + "[| H |- p; H |- q->false |] ==> H |- (p->q)->false"; +by (rtac deduction 1); +by (rtac (premq RS weaken_left_insert RS conseq_MP) 1); +by (rtac (conseq_H RS conseq_MP) 1); +by (rtac insertI1 1); +by (rtac (premp RS weaken_left_insert) 1); +val imp_false = result(); + +(*This formulation is required for strong induction hypotheses*) +goal PL.thy "hyps(p,tt) |- if(tt[p], p, p->false)"; +by (rtac (expand_if RS iffD2) 1); +by(res_inst_tac[("x","p")]spec 1); +by (rtac pl_ind 1); +by (ALLGOALS (simp_tac (pl_ss addsimps [conseq_I, conseq_H]))); +by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2, + weaken_right, imp_false] + addSEs [false_imp]) 1); +val hyps_conseq_if = result(); + +(*Key lemma for completeness; yields a set of assumptions satisfying p*) +val [sat] = goalw PL.thy [sat_def] "{} |= p ==> hyps(p,tt) |- p"; +by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN + rtac hyps_conseq_if 2); +by (fast_tac set_cs 1); +val sat_conseq_p = result(); + +(*For proving certain theorems in our new propositional logic*) +val conseq_cs = + set_cs addSIs [deduction] addIs [conseq_H, conseq_H RS conseq_MP]; + +(*The excluded middle in the form of an elimination rule*) +goal PL.thy "H |- (p->q) -> ((p->false)->q) -> q"; +by (rtac (deduction RS deduction) 1); +by (rtac (conseq_DN RS conseq_MP) 1); +by (ALLGOALS (best_tac (conseq_cs addSIs prems))); +val conseq_excluded_middle = result(); + +(*Hard to prove directly because it requires cuts*) +val prems = goal PL.thy + "[| insert(p,H) |- q; insert(p->false,H) |- q |] ==> H |- q"; +by (rtac (conseq_excluded_middle RS conseq_MP RS conseq_MP) 1); +by (REPEAT (resolve_tac (prems@[deduction]) 1)); +val conseq_excluded_middle_rule = result(); + +(*** Completeness -- lemmas for reducing the set of assumptions ***) + +(*For the case hyps(p,t)-insert(#v,Y) |- p; + we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) +goal PL.thy "!p.hyps(p, t-{v}) <= insert(#v->false, hyps(p,t)-{#v})"; +by (rtac pl_ind 1); +by (simp_tac pl_ss 1); +by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); +by (fast_tac (set_cs addSEs [sym RS var_neq_imp] addSDs [var_inject]) 1); +by (simp_tac pl_ss 1); +by (fast_tac set_cs 1); +val hyps_Diff = result() RS spec; + +(*For the case hyps(p,t)-insert(#v -> false,Y) |- p; + we also have hyps(p,t)-{#v->false} <= hyps(p, insert(v,t)) *) +goal PL.thy "!p.hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{#v->false})"; +by (rtac pl_ind 1); +by (simp_tac pl_ss 1); +by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); +by (fast_tac (set_cs addSEs [var_neq_imp, imp_inject] addSDs [var_inject]) 1); +by (simp_tac pl_ss 1); +by (fast_tac set_cs 1); +val hyps_insert = result() RS spec; + +(** Two lemmas for use with weaken_left **) + +goal Set.thy "B-C <= insert(a, B-insert(a,C))"; +by (fast_tac set_cs 1); +val insert_Diff_same = result(); + +goal Set.thy "insert(a, B-{c}) - D <= insert(a, B-insert(c,D))"; +by (fast_tac set_cs 1); +val insert_Diff_subset2 = result(); + +(*The set hyps(p,t) is finite, and elements have the form #v or #v->false; + could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*) +goal PL.thy "!p.hyps(p,t) : Fin(UN v:{x.True}. {#v, #v->false})"; +by (rtac pl_ind 1); +by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN' + fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI]))); +val hyps_finite = result() RS spec; + +val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; + +(*Induction on the finite set of assumptions hyps(p,t0). + We may repeatedly subtract assumptions until none are left!*) +val [sat] = goal PL.thy + "{} |= p ==> !t. hyps(p,t) - hyps(p,t0) |- p"; +by (rtac (hyps_finite RS Fin_induct) 1); +by (simp_tac (pl_ss addsimps [sat RS sat_conseq_p]) 1); +by (safe_tac set_cs); +(*Case hyps(p,t)-insert(#v,Y) |- p *) +by (rtac conseq_excluded_middle_rule 1); +by (rtac (insert_Diff_same RS weaken_left) 1); +by (etac spec 1); +by (rtac (insert_Diff_subset2 RS weaken_left) 1); +by (rtac (hyps_Diff RS Diff_weaken_left) 1); +by (etac spec 1); +(*Case hyps(p,t)-insert(#v -> false,Y) |- p *) +by (rtac conseq_excluded_middle_rule 1); +by (rtac (insert_Diff_same RS weaken_left) 2); +by (etac spec 2); +by (rtac (insert_Diff_subset2 RS weaken_left) 1); +by (rtac (hyps_insert RS Diff_weaken_left) 1); +by (etac spec 1); +val completeness_0_lemma = result(); + +(*The base case for completeness*) +val [sat] = goal PL.thy "{} |= p ==> {} |- p"; +by (rtac (Diff_cancel RS subst) 1); +by (rtac (sat RS (completeness_0_lemma RS spec)) 1); +val completeness_0 = result(); + +(*A semantic analogue of the Deduction Theorem*) +val [sat] = goalw PL.thy [sat_def] "insert(p,H) |= q ==> H |= p->q"; +by (simp_tac pl_ss 1); +by (cfast_tac [sat] 1); +val sat_imp = result(); + +val [finite] = goal PL.thy "H: Fin({p.True}) ==> !p. H |= p --> H |- p"; +by (rtac (finite RS Fin_induct) 1); +by (safe_tac (set_cs addSIs [completeness_0])); +by (rtac (weaken_left_insert RS conseq_MP) 1); +by (fast_tac (set_cs addSIs [sat_imp]) 1); +by (fast_tac conseq_cs 1); +val completeness_lemma = result(); + +val completeness = completeness_lemma RS spec RS mp; + +val [finite] = goal PL.thy "H: Fin({p.True}) ==> (H |- p) = (H |= p)"; +by (fast_tac (set_cs addSEs [soundness, finite RS completeness]) 1); +val conseq_iff = result(); + +writeln"Reached end of file."; + + diff -r 000000000000 -r 7949f97df77a ex/pl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/pl.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,55 @@ +(* Title: HOL/ex/prop-log + ID: $Id$ + Author: Tobias Nipkow + Copyright 1991 University of Cambridge + +Inductive definition of propositional logic. + +*) + +PL = Finite + +types pl 1 +arities pl :: (term)term +consts + false :: "'a pl" + "->" :: "['a pl,'a pl] => 'a pl" (infixr 90) + var :: "'a => 'a pl" ("#_") + pl_rec :: "['a pl,'a => 'b, 'b, ['b,'b] => 'b] => 'b" + axK,axS,axDN:: "'a pl set" + ruleMP,thms :: "'a pl set => 'a pl set" + "|-" :: "['a pl set, 'a pl] => bool" (infixl 50) + "|=" :: "['a pl set, 'a pl] => bool" (infixl 50) + eval :: "['a set, 'a pl] => bool" ("_[_]" [100,0] 100) + hyps :: "['a pl, 'a set] => 'a pl set" +rules + + (** Proof theory for propositional logic **) + + axK_def "axK == {x . ? p q. x = p->q->p}" + axS_def "axS == {x . ? p q r. x = (p->q->r) -> (p->q) -> p->r}" + axDN_def "axDN == {x . ? p. x = ((p->false) -> false) -> p}" + + (*the use of subsets simplifies the proof of monotonicity*) + ruleMP_def "ruleMP(X) == {q. ? p:X. p->q : X}" + + thms_def + "thms(H) == lfp(%X. H Un axK Un axS Un axDN Un ruleMP(X))" + + conseq_def "H |- p == p : thms(H)" + + sat_def "H |= p == (!tt. (!q:H. tt[q]) --> tt[p])" + +pl_rec_var "pl_rec(#v,f,y,z) = f(v)" +pl_rec_false "pl_rec(false,f,y,z) = y" +pl_rec_imp "pl_rec(p->q,f,y,g) = g(pl_rec(p,f,y,g),pl_rec(q,f,y,g))" + +eval_def "tt[p] == pl_rec(p, %v.v:tt, False, op -->)" + +hyps_def + "hyps(p,tt) == pl_rec(p, %a. {if(a:tt, #a, #a->false)}, {}, op Un)" + +var_inject "(#v = #w) ==> v = w" +imp_inject "[| (p -> q) = (p' -> q'); [| p = p'; q = q' |] ==> R |] ==> R" +var_neq_imp "(#v = (p -> q)) ==> R" +pl_ind "[| P(false); !!v. P(#v); !!p q. P(p)-->P(q)-->P(p->q)|] ==> !t.P(t)" +end diff -r 000000000000 -r 7949f97df77a ex/puzzle.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/puzzle.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,58 @@ +(* Title: HOL/ex/puzzle.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 TU Muenchen + +For puzzle.thy. A question from "Bundeswettbewerb Mathematik" + +Proof due to Herbert Ehler +*) + +(*specialized form of induction needed below*) +val prems = goal Nat.thy "[| P(0); !!n. P(Suc(n)) |] ==> !n.P(n)"; +by (EVERY1 [rtac (nat_induct RS allI), resolve_tac prems, resolve_tac prems]); +val nat_exh = result(); + +goal Puzzle.thy "! n. k=f(n) --> n <= f(n)"; +by (res_inst_tac [("n","k")] less_induct 1); +by (rtac nat_exh 1); +by (simp_tac nat_ss 1); +by (rtac impI 1); +by (rtac classical 1); +by (dtac not_leE 1); +by (subgoal_tac "f(na) <= f(f(na))" 1); +by (best_tac (HOL_cs addIs [lessD,Puzzle.f_ax,le_less_trans,le_trans]) 1); +by (fast_tac (HOL_cs addIs [Puzzle.f_ax]) 1); +val lemma = result() RS spec RS mp; + +goal Puzzle.thy "n <= f(n)"; +by (fast_tac (HOL_cs addIs [lemma]) 1); +val lemma1 = result(); + +goal Puzzle.thy "f(n) < f(Suc(n))"; +by (fast_tac (HOL_cs addIs [Puzzle.f_ax,le_less_trans,lemma1]) 1); +val lemma2 = result(); + +val prems = goal Puzzle.thy "(!!n.f(n) <= f(Suc(n))) ==> m f(m) <= f(n)"; +by (res_inst_tac[("n","n")]nat_induct 1); +by (simp_tac nat_ss 1); +by (simp_tac nat_ss 1); +by (fast_tac (HOL_cs addIs (le_trans::prems)) 1); +val mono_lemma1 = result() RS mp; + +val [p1,p2] = goal Puzzle.thy + "[| !! n. f(n)<=f(Suc(n)); m<=n |] ==> f(m) <= f(n)"; +by (rtac (p2 RS le_imp_less_or_eq RS disjE) 1); +by (etac (p1 RS mono_lemma1) 1); +by (fast_tac (HOL_cs addIs [le_refl]) 1); +val mono_lemma = result(); + +val prems = goal Puzzle.thy "m <= n ==> f(m) <= f(n)"; +by (fast_tac (HOL_cs addIs ([mono_lemma,less_imp_le,lemma2]@prems)) 1); +val f_mono = result(); + +goal Puzzle.thy "f(n) = n"; +by (rtac le_anti_sym 1); +by (rtac lemma1 2); +by (fast_tac (HOL_cs addIs [Puzzle.f_ax,leI] addDs [leD,f_mono,lessD]) 1); +result(); diff -r 000000000000 -r 7949f97df77a ex/puzzle.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/puzzle.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,12 @@ +(* Title: HOL/ex/puzzle.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 TU Muenchen + +An question from "Bundeswettbewerb Mathematik" +*) + +Puzzle = Nat + +consts f :: "nat => nat" +rules f_ax "f(f(n)) < f(Suc(n))" +end diff -r 000000000000 -r 7949f97df77a ex/rec.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/rec.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,5 @@ +open Rec; + +goalw thy [mono_def,Domf_def] "mono(Domf(F))"; +by (DEPTH_SOLVE (slow_step_tac set_cs 1)); +val mono_Domf = result(); diff -r 000000000000 -r 7949f97df77a ex/rec.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/rec.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,9 @@ +Rec = Fixedpt + +consts +fix :: "('a=>'a) => 'a" +Dom :: "(('a=>'b) => ('a=>'b)) => 'a set" +Domf :: "(('a=>'b) => ('a=>'b)) => 'a set => 'a set" +rules +Domf_def "Domf(F,D) == {y . !f g. (!x:D. f(x)=g(x)) --> F(f,y)=F(g,y)}" +Dom_def "Dom(F) == lfp(Domf(F))" +end diff -r 000000000000 -r 7949f97df77a ex/rel.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/rel.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,109 @@ +(* Title: HOL/ex/rel + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Domain, range of a relation or function -- NOT YET WORKING +*) + +structure Rel = +struct +val thy = extend_theory Univ.thy "Rel" +([], [], [], [], + [ + (["domain"], "('a * 'b)set => 'a set"), + (["range2"], "('a * 'b)set => 'b set"), + (["field"], "('a * 'a)set => 'a set") + ], + None) + [ + ("domain_def", "domain(r) == {a. ? b. : r}" ), + ("range2_def", "range2(r) == {b. ? a. : r}" ), + ("field_def", "field(r) == domain(r) Un range2(r)" ) + ]; +end; + +local val ax = get_axiom Rel.thy +in +val domain_def = ax"domain_def"; +val range2_def = ax"range2_def"; +val field_def = ax"field_def"; +end; + + +(*** domain ***) + +val [prem] = goalw Rel.thy [domain_def,Pair_def] ": r ==> a : domain(r)"; +by (fast_tac (set_cs addIs [prem]) 1); +val domainI = result(); + +val major::prems = goalw Rel.thy [domain_def] + "[| a : domain(r); !!y. : r ==> P |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac exE 1); +by (REPEAT (ares_tac prems 1)); +val domainE = result(); + + +(*** range2 ***) + +val [prem] = goalw Rel.thy [range2_def,Pair_def] ": r ==> b : range2(r)"; +by (fast_tac (set_cs addIs [prem]) 1); +val range2I = result(); + +val major::prems = goalw Rel.thy [range2_def] + "[| b : range2(r); !!x. : r ==> P |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac exE 1); +by (REPEAT (ares_tac prems 1)); +val range2E = result(); + + +(*** field ***) + +val [prem] = goalw Rel.thy [field_def] ": r ==> a : field(r)"; +by (rtac (prem RS domainI RS UnI1) 1); +val fieldI1 = result(); + +val [prem] = goalw Rel.thy [field_def] ": r ==> b : field(r)"; +by (rtac (prem RS range2I RS UnI2) 1); +val fieldI2 = result(); + +val [prem] = goalw Rel.thy [field_def] + "(~ :r ==> : r) ==> a : field(r)"; +by (rtac (prem RS domainI RS UnCI) 1); +by (swap_res_tac [range2I] 1); +by (etac notnotD 1); +val fieldCI = result(); + +val major::prems = goalw Rel.thy [field_def] + "[| a : field(r); \ +\ !!x. : r ==> P; \ +\ !!x. : r ==> P |] ==> P"; +by (rtac (major RS UnE) 1); +by (REPEAT (eresolve_tac (prems@[domainE,range2E]) 1)); +val fieldE = result(); + +goalw Rel.thy [field_def] "domain(r) <= field(r)"; +by (rtac Un_upper1 1); +val domain_in_field = result(); + +goalw Rel.thy [field_def] "range2(r) <= field(r)"; +by (rtac Un_upper2 1); +val range2_in_field = result(); + + +????????????????????????????????????????????????????????????????; + +(*If r allows well-founded induction then wf(r)*) +val [prem1,prem2] = goalw Rel.thy [wf_def] + "[| field(r)<=A; \ +\ !!P u. ! x:A. (! y. : r --> P(y)) --> P(x) ==> P(u) |] \ +\ ==> wf(r)"; +by (rtac (prem1 RS wfI) 1); +by (res_inst_tac [ ("B", "A-Z") ] (prem2 RS subsetCE) 1); +by (fast_tac ZF_cs 3); +by (fast_tac ZF_cs 2); +by (fast_tac ZF_cs 1); +val wfI2 = result(); + diff -r 000000000000 -r 7949f97df77a ex/set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/set.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,125 @@ +(* Title: HOL/ex/set.ML + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Cantor's Theorem; the Schroeder-Berstein Theorem. +*) + + +writeln"File HOL/ex/set."; + +(*** Cantor's Theorem: There is no surjection from a set to its powerset. ***) + +goal Set.thy "~ (? f:: 'a=>'a set. ! S. ? x. f(x) = S)"; +(*requires best-first search because it is undirectional*) +by (best_tac (set_cs addSEs [equalityCE]) 1); +val cantor1 = result(); + +(*This form displays the diagonal term*) +goal Set.thy "! f:: 'a=>'a set. ! x. ~ f(x) = ?S(f)"; +by (best_tac (set_cs addSEs [equalityCE]) 1); +uresult(); + +(*This form exploits the set constructs*) +goal Set.thy "~ ?S : range(f :: 'a=>'a set)"; +by (rtac notI 1); +by (etac rangeE 1); +by (etac equalityCE 1); +by (dtac CollectD 1); +by (contr_tac 1); +by (swap_res_tac [CollectI] 1); +by (assume_tac 1); + +choplev 0; +by (best_tac (set_cs addSEs [equalityCE]) 1); + +(*** The Schroder-Berstein Theorem ***) + +val prems = goalw Lfp.thy [image_def] "inj(f) ==> Inv(f)``(f``X) = X"; +by (cut_facts_tac prems 1); +by (rtac equalityI 1); +by (fast_tac (set_cs addEs [Inv_f_f RS ssubst]) 1); +by (fast_tac (set_cs addEs [Inv_f_f RS ssubst]) 1); +val inv_image_comp = result(); + +val prems = goal Set.thy "~ f(a) : (f``X) ==> ~ a:X"; +by (cfast_tac prems 1); +val contra_imageI = result(); + +goal Lfp.thy "(~ a : Compl(X)) = (a:X)"; +by (fast_tac set_cs 1); +val not_Compl = result(); + +(*Lots of backtracking in this proof...*) +val [compl,fg,Xa] = goal Lfp.thy + "[| Compl(f``X) = g``Compl(X); f(a)=g(b); a:X |] ==> b:X"; +by (EVERY1 [rtac (not_Compl RS subst), rtac contra_imageI, + rtac (compl RS subst), rtac (fg RS subst), stac not_Compl, + rtac imageI, rtac Xa]); +val disj_lemma = result(); + +goal Lfp.thy "range(%z. if(z:X, f(z), g(z))) = f``X Un g``Compl(X)"; +by (rtac equalityI 1); +by (rewtac range_def); +by (fast_tac (set_cs addIs [if_P RS sym, if_not_P RS sym]) 2); +by (rtac subsetI 1); +by (etac CollectE 1); +by (etac exE 1); +by (etac ssubst 1); +by (rtac (excluded_middle RS disjE) 1); +by (EVERY' [rtac (if_P RS ssubst), atac, fast_tac set_cs] 2); +by (EVERY' [rtac (if_not_P RS ssubst), atac, fast_tac set_cs] 1); +val range_if_then_else = result(); + +goal Lfp.thy "a : X Un Compl(X)"; +by (fast_tac set_cs 1); +val X_Un_Compl = result(); + +goalw Lfp.thy [surj_def] "surj(f) = (!a. a : range(f))"; +by (fast_tac (set_cs addEs [ssubst]) 1); +val surj_iff_full_range = result(); + +val [compl] = goal Lfp.thy + "Compl(f``X) = g``Compl(X) ==> surj(%z. if(z:X, f(z), g(z)))"; +by (sstac [surj_iff_full_range, range_if_then_else, compl RS sym] 1); +by (rtac (X_Un_Compl RS allI) 1); +val surj_if_then_else = result(); + +val [injf,injg,compl,bij] = goal Lfp.thy + "[| inj_onto(f,X); inj_onto(g,Compl(X)); Compl(f``X) = g``Compl(X); \ +\ bij = (%z. if(z:X, f(z), g(z))) |] ==> \ +\ inj(bij) & surj(bij)"; +val f_eq_gE = make_elim (compl RS disj_lemma); +by (rtac (bij RS ssubst) 1); +by (rtac conjI 1); +by (rtac (compl RS surj_if_then_else) 2); +by (rewtac inj_def); +by (cut_facts_tac [injf,injg] 1); +by (EVERY1 [rtac allI, rtac allI, stac expand_if, rtac conjI, stac expand_if]); +by (fast_tac (set_cs addEs [inj_ontoD, sym RS f_eq_gE]) 1); +by (stac expand_if 1); +by (fast_tac (set_cs addEs [inj_ontoD, f_eq_gE]) 1); +val bij_if_then_else = result(); + +goal Lfp.thy "? X. X = Compl(g``Compl(f:: 'a=>'b``X))"; +by (rtac exI 1); +by (rtac lfp_Tarski 1); +by (REPEAT (ares_tac [monoI, image_mono, Compl_anti_mono] 1)); +val decomposition = result(); + +val [injf,injg] = goal Lfp.thy + "[| inj(f:: 'a=>'b); inj(g:: 'b=>'a) |] ==> \ +\ ? h:: 'a=>'b. inj(h) & surj(h)"; +by (rtac (decomposition RS exE) 1); +by (rtac exI 1); +by (rtac bij_if_then_else 1); +by (EVERY [rtac refl 4, rtac (injf RS inj_imp) 1, + rtac (injg RS inj_onto_Inv) 1]); +by (EVERY1 [etac ssubst, stac double_complement, rtac subsetI, + etac imageE, etac ssubst, rtac rangeI]); +by (EVERY1 [etac ssubst, stac double_complement, + rtac (injg RS inv_image_comp RS sym)]); +val schroeder_bernstein = result(); + +writeln"Reached end of file."; diff -r 000000000000 -r 7949f97df77a ex/simult.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/simult.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,319 @@ +(* Title: HOL/ex/simult.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For simult.thy. + +Primitives for simultaneous recursive type definitions + includes worked example of trees & forests + +This is essentially the same data structure that on ex/term.ML, which is +simpler because it uses List as a new type former. The approach in this +file may be superior for other simultaneous recursions. +*) + +open Simult; + +(*** General rules for Part ***) + +val prems = goalw Simult.thy [Part_def] "h(a) : A ==> h(a) : Part(A,h)"; +by (cfast_tac prems 1); +val PartI = result(); + +val major::prems = goalw Simult.thy [Part_def] + "[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS CollectE) 1); +by (etac conjE 1); +by (etac exE 1); +by (REPEAT (ares_tac prems 1)); +val PartE = result(); + +goal Simult.thy "!!A B. A<=B ==> Part(A,h) <= Part(B,h)"; +by (fast_tac (set_cs addSIs [PartI] addSEs [PartE]) 1); +val Part_mono = result(); + +val [prem] = goal Simult.thy "A<=B ==> Part(A,h) <= B"; +by (fast_tac (set_cs addSIs [PartI, prem RS subsetD] addSEs [PartE]) 1); +val Part_subset = result(); + + +(*** Monotonicity and unfolding of the function ***) + +goal Simult.thy "mono(%Z. A <*> Part(Z,In1) \ +\ <+> ({Numb(0)} <+> Part(Z,In0) <*> Part(Z,In1)))"; +by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono, + Part_mono] 1)); +val TF_fun_mono = result(); + +val TF_unfold = TF_fun_mono RS (TF_def RS def_lfp_Tarski); + +goalw Simult.thy [TF_def] "!!A B. A<=B ==> TF(A) <= TF(B)"; +by (REPEAT (ares_tac [lfp_mono, subset_refl, usum_mono, uprod_mono] 1)); +val TF_mono = result(); + +goalw Simult.thy [TF_def] "TF(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_NumbI,Sexp_In0I,Sexp_In1I,Sexp_SconsI] + addSEs [PartE]) 1); +val TF_Sexp = result(); + +(* A <= Sexp ==> TF(A) <= Sexp *) +val TF_subset_Sexp = standard + (TF_mono RS (TF_Sexp RSN (2,subset_trans))); + + +(** Elimination -- structural induction on the set TF **) + +val TF_Rep_defs = [TCONS_def,FNIL_def,FCONS_def,NIL_def,CONS_def]; + +val major::prems = goalw Simult.thy TF_Rep_defs + "[| i: TF(A); \ +\ !!M N. [| M: A; N: Part(TF(A),In1); R(N) |] ==> R(TCONS(M,N)); \ +\ R(FNIL); \ +\ !!M N. [| M: Part(TF(A),In0); N: Part(TF(A),In1); R(M); R(N) \ +\ |] ==> R(FCONS(M,N)) \ +\ |] ==> R(i)"; +by (rtac (major RS (TF_def RS def_induct)) 1); +by (rtac TF_fun_mono 1); +by (fast_tac (set_cs addIs (prems@[PartI]) + addEs [usumE, uprodE, PartE]) 1); +val TF_induct = result(); + +(*This lemma replaces a use of subgoal_tac to prove tree_forest_induct*) +val prems = goalw Simult.thy [Part_def] + "! M: TF(A). (M: Part(TF(A),In0) --> P(M)) & (M: Part(TF(A),In1) --> Q(M)) \ +\ ==> (! M: Part(TF(A),In0). P(M)) & (! M: Part(TF(A),In1). Q(M))"; +by (cfast_tac prems 1); +val TF_induct_lemma = result(); + +val uplus_cs = set_cs addSIs [PartI] + addSDs [In0_inject, In1_inject] + addSEs [In0_neq_In1, In1_neq_In0, PartE]; + +(*Could prove ~ TCONS(M,N) : Part(TF(A),In1) etc. *) + +(*Induction on TF with separate predicates P, Q*) +val prems = goalw Simult.thy TF_Rep_defs + "[| !!M N. [| M: A; N: Part(TF(A),In1); Q(N) |] ==> P(TCONS(M,N)); \ +\ Q(FNIL); \ +\ !!M N. [| M: Part(TF(A),In0); N: Part(TF(A),In1); P(M); Q(N) \ +\ |] ==> Q(FCONS(M,N)) \ +\ |] ==> (! M: Part(TF(A),In0). P(M)) & (! N: Part(TF(A),In1). Q(N))"; +by (rtac (ballI RS TF_induct_lemma) 1); +by (etac TF_induct 1); +by (rewrite_goals_tac TF_Rep_defs); +by (ALLGOALS (fast_tac (uplus_cs addIs prems))); +(*29 secs??*) +val Tree_Forest_induct = result(); + +(*Induction for the abstract types 'a tree, 'a forest*) +val prems = goalw Simult.thy [Tcons_def,Fnil_def,Fcons_def] + "[| !!x ts. Q(ts) ==> P(Tcons(x,ts)); \ +\ Q(Fnil); \ +\ !!t ts. [| P(t); Q(ts) |] ==> Q(Fcons(t,ts)) \ +\ |] ==> (! t. P(t)) & (! ts. Q(ts))"; +by (res_inst_tac [("P1","%z.P(Abs_Tree(z))"), + ("Q1","%z.Q(Abs_Forest(z))")] + (Tree_Forest_induct RS conjE) 1); +(*Instantiates ?A1 to range(Leaf). *) +by (fast_tac (set_cs addSEs [Rep_Tree_inverse RS subst, + Rep_Forest_inverse RS subst] + addSIs [Rep_Tree,Rep_Forest]) 4); +(*Cannot use simplifier: the rewrites work in the wrong direction!*) +by (ALLGOALS (fast_tac (set_cs addSEs [Abs_Tree_inverse RS subst, + Abs_Forest_inverse RS subst] + addSIs prems))); +val tree_forest_induct = result(); + + + +(*** Isomorphisms ***) + +goal Simult.thy "inj(Rep_Tree)"; +by (rtac inj_inverseI 1); +by (rtac Rep_Tree_inverse 1); +val inj_Rep_Tree = result(); + +goal Simult.thy "inj_onto(Abs_Tree,Part(TF(range(Leaf)),In0))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Tree_inverse 1); +val inj_onto_Abs_Tree = result(); + +goal Simult.thy "inj(Rep_Forest)"; +by (rtac inj_inverseI 1); +by (rtac Rep_Forest_inverse 1); +val inj_Rep_Forest = result(); + +goal Simult.thy "inj_onto(Abs_Forest,Part(TF(range(Leaf)),In1))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Forest_inverse 1); +val inj_onto_Abs_Forest = result(); + +(** Introduction rules for constructors **) + +(* c : A <*> Part(TF(A),In1) + <+> {Numb(0)} <+> Part(TF(A),In0) <*> Part(TF(A),In1) ==> c : TF(A) *) +val TF_I = TF_unfold RS equalityD2 RS subsetD; + +(*For reasoning about the representation*) +val TF_Rep_cs = uplus_cs addIs [TF_I, uprodI, usum_In0I, usum_In1I] + addSEs [Scons_inject]; + +val prems = goalw Simult.thy TF_Rep_defs + "[| a: A; M: Part(TF(A),In1) |] ==> TCONS(a,M) : Part(TF(A),In0)"; +by (fast_tac (TF_Rep_cs addIs prems) 1); +val TCONS_I = result(); + +(* FNIL is a TF(A) -- this also justifies the type definition*) +goalw Simult.thy TF_Rep_defs "FNIL: Part(TF(A),In1)"; +by (fast_tac TF_Rep_cs 1); +val FNIL_I = result(); + +val prems = goalw Simult.thy TF_Rep_defs + "[| M: Part(TF(A),In0); N: Part(TF(A),In1) |] ==> \ +\ FCONS(M,N) : Part(TF(A),In1)"; +by (fast_tac (TF_Rep_cs addIs prems) 1); +val FCONS_I = result(); + +(** Injectiveness of TCONS and FCONS **) + +goalw Simult.thy TF_Rep_defs "(TCONS(K,M)=TCONS(L,N)) = (K=L & M=N)"; +by (fast_tac TF_Rep_cs 1); +val TCONS_TCONS_eq = result(); +val TCONS_inject = standard (TCONS_TCONS_eq RS iffD1 RS conjE); + +goalw Simult.thy TF_Rep_defs "(FCONS(K,M)=FCONS(L,N)) = (K=L & M=N)"; +by (fast_tac TF_Rep_cs 1); +val FCONS_FCONS_eq = result(); +val FCONS_inject = standard (FCONS_FCONS_eq RS iffD1 RS conjE); + +(** Distinctness of TCONS, FNIL and FCONS **) + +goalw Simult.thy TF_Rep_defs "~ TCONS(M,N) = FNIL"; +by (fast_tac TF_Rep_cs 1); +val TCONS_not_FNIL = result(); +val FNIL_not_TCONS = standard (TCONS_not_FNIL RS not_sym); + +val TCONS_neq_FNIL = standard (TCONS_not_FNIL RS notE); +val FNIL_neq_TCONS = sym RS TCONS_neq_FNIL; + +goalw Simult.thy TF_Rep_defs "~ FCONS(M,N) = FNIL"; +by (fast_tac TF_Rep_cs 1); +val FCONS_not_FNIL = result(); +val FNIL_not_FCONS = standard (FCONS_not_FNIL RS not_sym); + +val FCONS_neq_FNIL = standard (FCONS_not_FNIL RS notE); +val FNIL_neq_FCONS = sym RS FCONS_neq_FNIL; + +goalw Simult.thy TF_Rep_defs "~ TCONS(M,N) = FCONS(K,L)"; +by (fast_tac TF_Rep_cs 1); +val TCONS_not_FCONS = result(); +val FCONS_not_TCONS = standard (TCONS_not_FCONS RS not_sym); + +val TCONS_neq_FCONS = standard (TCONS_not_FCONS RS notE); +val FCONS_neq_TCONS = sym RS TCONS_neq_FCONS; + +(*???? Too many derived rules ???? + Automatically generate symmetric forms? Always expand TF_Rep_defs? *) + +(** Injectiveness of Tcons and Fcons **) + +(*For reasoning about abstract constructors*) +val TF_cs = set_cs addSIs [Rep_Tree, Rep_Forest, TCONS_I, FNIL_I, FCONS_I] + addSEs [TCONS_inject, FCONS_inject, + TCONS_neq_FNIL, FNIL_neq_TCONS, + FCONS_neq_FNIL, FNIL_neq_FCONS, + TCONS_neq_FCONS, FCONS_neq_TCONS] + addSDs [inj_onto_Abs_Tree RS inj_ontoD, + inj_onto_Abs_Forest RS inj_ontoD, + inj_Rep_Tree RS injD, inj_Rep_Forest RS injD, + Leaf_inject]; + +goalw Simult.thy [Tcons_def] "(Tcons(x,xs)=Tcons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac TF_cs 1); +val Tcons_Tcons_eq = result(); +val Tcons_inject = standard (Tcons_Tcons_eq RS iffD1 RS conjE); + +goalw Simult.thy [Fcons_def,Fnil_def] "~ Fcons(x,xs) = Fnil"; +by (fast_tac TF_cs 1); +val Fcons_not_Fnil = result(); + +val Fcons_neq_Fnil = standard (Fcons_not_Fnil RS notE);; +val Fnil_neq_Fcons = sym RS Fcons_neq_Fnil; + + +(** Injectiveness of Fcons **) + +goalw Simult.thy [Fcons_def] "(Fcons(x,xs)=Fcons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac TF_cs 1); +val Fcons_Fcons_eq = result(); +val Fcons_inject = standard (Fcons_Fcons_eq RS iffD1 RS conjE); + + +(*** TF_rec -- by wf recursion on pred_Sexp ***) + +val TF_rec_unfold = + wf_pred_Sexp RS wf_trancl RS (TF_rec_def RS def_wfrec); + +(** conversion rules for TF_rec **) + +val prems = goalw Simult.thy [TCONS_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ TF_rec(TCONS(M,N),b,c,d) = b(M, N, TF_rec(N,b,c,d))"; +by (rtac (TF_rec_unfold RS trans) 1); +by (rtac (Case_In0 RS trans) 1); +by (rtac (Split RS trans) 1); +by (rewtac In0_def); +by (simp_tac (pred_Sexp_ss addsimps prems) 1); +val TF_rec_TCONS = result(); + +goalw Simult.thy [FNIL_def] "TF_rec(FNIL,b,c,d) = c"; +by (rtac (TF_rec_unfold RS trans) 1); +by (rtac (Case_In1 RS trans) 1); +by (rtac List_case_NIL 1); +val TF_rec_FNIL = result(); + +val prems = goalw Simult.thy [FCONS_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ TF_rec(FCONS(M,N),b,c,d) = d(M, N, TF_rec(M,b,c,d), TF_rec(N,b,c,d))"; +by (rtac (TF_rec_unfold RS trans) 1); +by (rtac (Case_In1 RS trans) 1); +by (rtac (List_case_CONS RS trans) 1); +by (rewrite_goals_tac [CONS_def,In1_def]); +by (simp_tac (pred_Sexp_ss addsimps prems) 1); +val TF_rec_FCONS = result(); + + +(*** tree_rec, forest_rec -- by TF_rec ***) + +val Rep_Tree_in_Sexp = + Rep_Tree RS (range_Leaf_subset_Sexp RS TF_subset_Sexp RS + Part_subset RS subsetD); +val Rep_Forest_in_Sexp = + Rep_Forest RS (range_Leaf_subset_Sexp RS TF_subset_Sexp RS + Part_subset RS subsetD); + +val tf_rec_simps = [TF_rec_TCONS, TF_rec_FNIL, TF_rec_FCONS, + TCONS_I, FNIL_I, FCONS_I, Rep_Tree, Rep_Forest, + Rep_Tree_inverse, Rep_Forest_inverse, + Abs_Tree_inverse, Abs_Forest_inverse, + inj_Leaf, Inv_f_f, Sexp_LeafI, range_eqI, + Rep_Tree_in_Sexp, Rep_Forest_in_Sexp]; +val tf_rec_ss = HOL_ss addsimps tf_rec_simps; + +goalw Simult.thy [tree_rec_def, forest_rec_def, Tcons_def] + "tree_rec(Tcons(a,tf),b,c,d) = b(a,tf,forest_rec(tf,b,c,d))"; +by (simp_tac tf_rec_ss 1); +val tree_rec_Tcons = result(); + +goalw Simult.thy [forest_rec_def, Fnil_def] "forest_rec(Fnil,b,c,d) = c"; +by (simp_tac tf_rec_ss 1); +val forest_rec_Fnil = result(); + +goalw Simult.thy [tree_rec_def, forest_rec_def, Fcons_def] + "forest_rec(Fcons(t,tf),b,c,d) = \ +\ d(t,tf,tree_rec(t,b,c,d), forest_rec(tf,b,c,d))"; +by (simp_tac tf_rec_ss 1); +val forest_rec_Cons = result(); diff -r 000000000000 -r 7949f97df77a ex/simult.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/simult.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,77 @@ +(* Title: HOL/ex/simult + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Primitives for simultaneous recursive type definitions + includes worked example of trees & forests + +This is essentially the same data structure that on ex/term.ML, which is +simpler because it uses List as a new type former. The approach in this +file may be superior for other simultaneous recursions. +*) + +Simult = List + +types tree,forest 1 +arities tree,forest :: (term)term +consts + Part :: "['a set, 'a=>'a] => 'a set" + TF :: "'a node set set => 'a node set set" + FNIL :: "'a node set" + TCONS,FCONS :: "['a node set, 'a node set] => 'a node set" + Rep_Tree :: "'a tree => 'a node set" + Abs_Tree :: "'a node set => 'a tree" + Rep_Forest :: "'a forest => 'a node set" + Abs_Forest :: "'a node set => 'a forest" + Tcons :: "['a, 'a forest] => 'a tree" + Fcons :: "['a tree, 'a forest] => 'a forest" + Fnil :: "'a forest" + TF_rec :: "['a node set, ['a node set , 'a node set, 'b]=>'b, \ +\ 'b, ['a node set , 'a node set, 'b, 'b]=>'b] => 'b" + tree_rec :: "['a tree, ['a, 'a forest, 'b]=>'b, \ +\ 'b, ['a tree, 'a forest, 'b, 'b]=>'b] => 'b" + forest_rec :: "['a forest, ['a, 'a forest, 'b]=>'b, \ +\ 'b, ['a tree, 'a forest, 'b, 'b]=>'b] => 'b" + +rules + (*operator for selecting out the various types*) + Part_def "Part(A,h) == {x. x:A & (? z. x = h(z))}" + + TF_def "TF(A) == lfp(%Z. A <*> Part(Z,In1) \ +\ <+> ({Numb(0)} <+> Part(Z,In0) <*> Part(Z,In1)))" + (*faking a type definition for tree...*) + Rep_Tree "Rep_Tree(n): Part(TF(range(Leaf)),In0)" + Rep_Tree_inverse "Abs_Tree(Rep_Tree(t)) = t" + Abs_Tree_inverse "z: Part(TF(range(Leaf)),In0) ==> Rep_Tree(Abs_Tree(z)) = z" + (*faking a type definition for forest...*) + Rep_Forest "Rep_Forest(n): Part(TF(range(Leaf)),In1)" + Rep_Forest_inverse "Abs_Forest(Rep_Forest(ts)) = ts" + Abs_Forest_inverse + "z: Part(TF(range(Leaf)),In1) ==> Rep_Forest(Abs_Forest(z)) = z" + + (*the concrete constants*) + TCONS_def "TCONS(M,N) == In0(M . N)" + FNIL_def "FNIL == In1(NIL)" + FCONS_def "FCONS(M,N) == In1(CONS(M,N))" + (*the abstract constants*) + Tcons_def "Tcons(a,ts) == Abs_Tree(TCONS(Leaf(a), Rep_Forest(ts)))" + Fnil_def "Fnil == Abs_Forest(FNIL)" + Fcons_def "Fcons(t,ts) == Abs_Forest(FCONS(Rep_Tree(t), Rep_Forest(ts)))" + + (*recursion*) + TF_rec_def + "TF_rec(M,b,c,d) == wfrec(trancl(pred_Sexp), M, \ +\ %Z g. Case(Z, %U. Split(U, %x y. b(x,y,g(y))), \ +\ %V. List_case(V, c, \ +\ %x y. d(x,y,g(x),g(y)))))" + + tree_rec_def + "tree_rec(t,b,c,d) == \ +\ TF_rec(Rep_Tree(t), %x y r. b(Inv(Leaf,x), Abs_Forest(y), r), \ +\ c, %x y rt rf. d(Abs_Tree(x), Abs_Forest(y), rt, rf))" + + forest_rec_def + "forest_rec(tf,b,c,d) == \ +\ TF_rec(Rep_Forest(tf), %x y r. b(Inv(Leaf,x), Abs_Forest(y), r), \ +\ c, %x y rt rf. d(Abs_Tree(x), Abs_Forest(y), rt, rf))" +end diff -r 000000000000 -r 7949f97df77a ex/term.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/term.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,177 @@ +(* Title: HOL/ex/term + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For term.thy. illustrates List functor + (essentially the same type as in Trees & Forests) +*) + +open Term; + +(*** Monotonicity and unfolding of the function ***) + +goal Term.thy "mono(%Z. A <*> List(Z))"; +by (REPEAT (ares_tac [monoI, subset_refl, List_mono, uprod_mono] 1)); +val Term_fun_mono = result(); + +val Term_unfold = Term_fun_mono RS (Term_def RS def_lfp_Tarski); + +(*This justifies using Term in other recursive type definitions*) +goalw Term.thy [Term_def] "!!A B. A<=B ==> Term(A) <= Term(B)"; +by (REPEAT (ares_tac [lfp_mono, subset_refl, List_mono, uprod_mono] 1)); +val Term_mono = result(); + +(** Type checking rules -- Term creates well-founded sets **) + +val prems = goalw Term.thy [Term_def] "Term(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_SconsI, List_Sexp RS subsetD]) 1); +val Term_Sexp = result(); + +(* A <= Sexp ==> Term(A) <= Sexp *) +val Term_subset_Sexp = standard + (Term_mono RS (Term_Sexp RSN (2,subset_trans))); + + +(** Elimination -- structural induction on the set Term(A) **) + +(*Induction for the set Term(A) *) +val [major,minor] = goal Term.thy + "[| M: Term(A); \ +\ !!x zs. [| x: A; zs: List(Term(A)); zs: List({x.R(x)}) \ +\ |] ==> R(x.zs) \ +\ |] ==> R(M)"; +by (rtac (major RS (Term_def RS def_induct)) 1); +by (rtac Term_fun_mono 1); +by (REPEAT (eresolve_tac ([uprodE, ssubst, minor] @ + ([Int_lower1,Int_lower2] RL [List_mono RS subsetD])) 1)); +val Term_induct = result(); + +(*Induction on Term(A) followed by induction on List *) +val major::prems = goal Term.thy + "[| M: Term(A); \ +\ !!x. [| x: A |] ==> R(x.NIL); \ +\ !!x z zs. [| x: A; z: Term(A); zs: List(Term(A)); R(x.zs) \ +\ |] ==> R(x . CONS(z,zs)) \ +\ |] ==> R(M)"; +by (rtac (major RS Term_induct) 1); +by (etac List_induct 1); +by (REPEAT (ares_tac prems 1)); +val Term_induct2 = result(); + +(*** Structural Induction on the abstract type 'a term ***) + +val list_all_ss = map_ss addsimps [list_all_Nil, list_all_Cons]; + +val Rep_Term_in_Sexp = + Rep_Term RS (range_Leaf_subset_Sexp RS Term_subset_Sexp RS subsetD); + +(*Induction for the abstract type 'a term*) +val prems = goalw Term.thy [App_def,Rep_TList_def,Abs_TList_def] + "[| !!x ts. list_all(R,ts) ==> R(App(x,ts)) \ +\ |] ==> R(t)"; +by (rtac (Rep_Term_inverse RS subst) 1); (*types force good instantiation*) +by (res_inst_tac [("P","Rep_Term(t) : Sexp")] conjunct2 1); +by (rtac (Rep_Term RS Term_induct) 1); +by (REPEAT (ares_tac [conjI, Sexp_SconsI, Term_subset_Sexp RS + List_subset_Sexp,range_Leaf_subset_Sexp] 1 + ORELSE etac rev_subsetD 1)); +by (eres_inst_tac [("A1","Term(?u)"), ("f1","Rep_Term"), ("g1","Abs_Term")] + (Abs_map_inverse RS subst) 1); +by (rtac (range_Leaf_subset_Sexp RS Term_subset_Sexp) 1); +by (etac Abs_Term_inverse 1); +by (etac rangeE 1); +by (hyp_subst_tac 1); +by (resolve_tac prems 1); +by (etac List_induct 1); +by (etac CollectE 2); +by (stac Abs_map_CONS 2); +by (etac conjunct1 2); +by (etac rev_subsetD 2); +by (rtac List_subset_Sexp 2); +by (fast_tac set_cs 2); +by (ALLGOALS (asm_simp_tac list_all_ss)); +val term_induct = result(); + +(*Induction for the abstract type 'a term*) +val prems = goal Term.thy + "[| !!x. R(App(x,Nil)); \ +\ !!x t ts. R(App(x,ts)) ==> R(App(x, Cons(t,ts))) \ +\ |] ==> R(t)"; +by (rtac term_induct 1); (*types force good instantiation*) +by (etac rev_mp 1); +by (rtac list_induct 1); (*types force good instantiation*) +by (ALLGOALS (asm_simp_tac (list_all_ss addsimps prems))); +val term_induct2 = result(); + +(*Perform induction on xs. *) +fun term_ind2_tac a i = + EVERY [res_inst_tac [("t",a)] term_induct2 i, + rename_last_tac a ["1","s"] (i+1)]; + + +(** Introduction rule for Term **) + +(* c : A <*> List(Term(A)) ==> c : Term(A) *) +val TermI = standard (Term_unfold RS equalityD2 RS subsetD); + +(*The constant APP is not declared; it is simply . *) +val prems = goal Term.thy "[| M: A; N : List(Term(A)) |] ==> M.N : Term(A)"; +by (REPEAT (resolve_tac (prems@[TermI, ListI, uprodI]) 1)); +val APP_I = result(); + + +(*** Term_rec -- by wf recursion on pred_Sexp ***) + +val Term_rec_unfold = + wf_pred_Sexp RS wf_trancl RS (Term_rec_def RS def_wfrec); + +(** conversion rules **) + +val [prem] = goal Term.thy + "N: List(Term(A)) ==> \ +\ !M. : pred_Sexp^+ --> \ +\ Abs_map(cut(h, pred_Sexp^+, M), N) = \ +\ Abs_map(h,N)"; +by (rtac (prem RS List_induct) 1); +by (simp_tac list_all_ss 1); +by (strip_tac 1); +by (etac (pred_Sexp_CONS_D RS conjE) 1); +by (asm_simp_tac (list_all_ss addsimps [trancl_pred_SexpD1, cut_apply]) 1); +val Abs_map_lemma = result(); + +val [prem1,prem2,A_subset_Sexp] = goal Term.thy + "[| M: Sexp; N: List(Term(A)); A<=Sexp |] ==> \ +\ Term_rec(M.N, d) = d(M, N, Abs_map(%Z. Term_rec(Z,d), N))"; +by (rtac (Term_rec_unfold RS trans) 1); +by (rtac (Split RS trans) 1); +by (simp_tac (HOL_ss addsimps + [prem2 RS Abs_map_lemma RS spec RS mp, pred_SexpI2 RS r_into_trancl, + prem1, prem2 RS rev_subsetD, List_subset_Sexp, + Term_subset_Sexp, A_subset_Sexp])1); +val Term_rec = result(); + +(*** term_rec -- by Term_rec ***) + +local + val Rep_map_type1 = read_instantiate_sg (sign_of Term.thy) + [("f","Rep_Term")] Rep_map_type; + val Rep_TList = Rep_Term RS Rep_map_type1; + val Rep_Term_rec = range_Leaf_subset_Sexp RSN (2,Rep_TList RSN(2,Term_rec)); + + (*Now avoids conditional rewriting with the premise N: List(Term(A)), + since A will be uninstantiated and will cause rewriting to fail. *) + val term_rec_ss = HOL_ss + addsimps [Rep_TList RS (rangeI RS APP_I RS Abs_Term_inverse), + Rep_Term_in_Sexp, Rep_Term_rec, Rep_Term_inverse, + inj_Leaf, Inv_f_f, + Abs_Rep_map, map_ident, Sexp_LeafI] +in + +val term_rec = prove_goalw Term.thy + [term_rec_def, App_def, Rep_TList_def, Abs_TList_def] + "term_rec(App(f,ts), d) = d(f, ts, map (%t. term_rec(t,d), ts))" + (fn _ => [simp_tac term_rec_ss 1]) + +end; diff -r 000000000000 -r 7949f97df77a ex/term.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/term.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,44 @@ +(* Title: HOL/ex/term + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Terms over a given alphabet -- function applications; illustrates List functor + (essentially the same type as in Trees & Forests) + +There is no constructor APP because it is simply cons (.) +*) + +Term = List + +types term 1 +arities term :: (term)term +consts + Term :: "'a node set set => 'a node set set" + Rep_Term :: "'a term => 'a node set" + Abs_Term :: "'a node set => 'a term" + Rep_TList :: "'a term list => 'a node set" + Abs_TList :: "'a node set => 'a term list" + App :: "['a, ('a term)list] => 'a term" + Term_rec :: + "['a node set, ['a node set , 'a node set, 'b list]=>'b] => 'b" + term_rec :: "['a term, ['a ,'a term list, 'b list]=>'b] => 'b" +rules + Term_def "Term(A) == lfp(%Z. A <*> List(Z))" + (*faking a type definition for term...*) + Rep_Term "Rep_Term(n): Term(range(Leaf))" + Rep_Term_inverse "Abs_Term(Rep_Term(t)) = t" + Abs_Term_inverse "M: Term(range(Leaf)) ==> Rep_Term(Abs_Term(M)) = M" + (*defining abstraction/representation functions for term list...*) + Rep_TList_def "Rep_TList == Rep_map(Rep_Term)" + Abs_TList_def "Abs_TList == Abs_map(Abs_Term)" + (*defining the abstract constants*) + App_def "App(a,ts) == Abs_Term(Leaf(a) . Rep_TList(ts))" + (*list recursion*) + Term_rec_def + "Term_rec(M,d) == wfrec(trancl(pred_Sexp), M, \ +\ %M g. Split(M, %x y. d(x,y, Abs_map(g,y))))" + + term_rec_def + "term_rec(t,d) == \ +\ Term_rec(Rep_Term(t), %x y r. d(Inv(Leaf,x), Abs_TList(y), r))" +end diff -r 000000000000 -r 7949f97df77a ex/unsolved.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ex/unsolved.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,70 @@ +(* Title: HOL/ex/unsolved + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Problems that currently defeat the MESON procedure as well as best_tac +*) + +(*from Vladimir Lifschitz, What Is the Inverse Method?, JAR 5 (1989), 1--23*) +goal HOL.thy "? x x'. ! y. ? z z'. (~P(y,y) | P(x,x) | ~S(z,x)) & \ +\ (S(x,y) | ~S(y,z) | Q(z',z')) & \ +\ (Q(x',y) | ~Q(y,z') | S(x',x'))"; + + +writeln"Problem 47 Schubert's Steamroller"; +goal HOL.thy + "(! x. P1(x) --> P0(x)) & (? x.P1(x)) & \ +\ (! x. P2(x) --> P0(x)) & (? x.P2(x)) & \ +\ (! x. P3(x) --> P0(x)) & (? x.P3(x)) & \ +\ (! x. P4(x) --> P0(x)) & (? x.P4(x)) & \ +\ (! x. P5(x) --> P0(x)) & (? x.P5(x)) & \ +\ (! x. Q1(x) --> Q0(x)) & (? x.Q1(x)) & \ +\ (! x. P0(x) --> ((! y.Q0(y)-->R(x,y)) | \ +\ (! y.P0(y) & S(y,x) & \ +\ (? z.Q0(z)&R(y,z)) --> R(x,y)))) & \ +\ (! x y. P3(y) & (P5(x)|P4(x)) --> S(x,y)) & \ +\ (! x y. P3(x) & P2(y) --> S(x,y)) & \ +\ (! x y. P2(x) & P1(y) --> S(x,y)) & \ +\ (! x y. P1(x) & (P2(y)|Q1(y)) --> ~R(x,y)) & \ +\ (! x y. P3(x) & P4(y) --> R(x,y)) & \ +\ (! x y. P3(x) & P5(y) --> ~R(x,y)) & \ +\ (! x. (P4(x)|P5(x)) --> (? y.Q0(y) & R(x,y))) \ +\ --> (? x y. P0(x) & P0(y) & (? z. Q1(z) & R(y,z) & R(x,y)))"; + + +writeln"Problem 55"; + +(*Original, equational version by Len Schubert, via Pelletier *) +goal HOL.thy + "(? x. lives(x) & killed(x,agatha)) & \ +\ lives(agatha) & lives(butler) & lives(charles) & \ +\ (! x. lives(x) --> x=agatha | x=butler | x=charles) & \ +\ (! x y. killed(x,y) --> hates(x,y)) & \ +\ (! x y. killed(x,y) --> ~richer(x,y)) & \ +\ (! x. hates(agatha,x) --> ~hates(charles,x)) & \ +\ (! x. ~ x=butler --> hates(agatha,x)) & \ +\ (! x. ~richer(x,agatha) --> hates(butler,x)) & \ +\ (! x. hates(agatha,x) --> hates(butler,x)) & \ +\ (! x. ? y. ~hates(x,y)) & \ +\ ~ agatha=butler --> \ +\ killed(agatha,agatha)"; + +(** Charles Morgan's problems **) + +val axa = "! x y. T(i(x, i(y,x)))"; +val axb = "! x y z. T(i(i(x, i(y,z)), i(i(x,y), i(x,z))))"; +val axc = "! x y. T(i(i(n(x), n(y)), i(y,x)))"; +val axd = "! x y. T(i(x,y)) & T(x) --> T(y)"; + +fun axjoin ([], q) = q + | axjoin(p::ps, q) = "(" ^ p ^ ") --> (" ^ axjoin(ps,q) ^ ")"; + +goal HOL.thy (axjoin([axa,axb,axd], "! x. T(i(x,x))")); + +writeln"Problem 66"; +goal HOL.thy (axjoin([axa,axb,axc,axd], "! x. T(i(x, n(n(x))))")); + +writeln"Problem 67"; +goal HOL.thy (axjoin([axa,axb,axc,axd], "! x. T(i(n(n(x)), x))")); + diff -r 000000000000 -r 7949f97df77a fun.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/fun.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,192 @@ +(* Title: HOL/fun + ID: $Id$ + Author: Tobias Nipkow, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +Lemmas about functions. +*) + +goal Set.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); +val expand_fun_eq = result(); + +val prems = goal Set.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)); +val apply_inverse = result(); + + +(*** Range of a function ***) + +(*Frequently b does not have the syntactic form of f(x).*) +val [prem] = goalw Set.thy [range_def] "b=f(x) ==> b : range(f)"; +by (EVERY1 [rtac CollectI, rtac exI, rtac prem]); +val range_eqI = result(); + +val rangeI = refl RS range_eqI; + +val [major,minor] = goalw Set.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); +val rangeE = result(); + +(*** Image of a set under a function ***) + +val prems = goalw Set.thy [image_def] "[| b=f(x); x:A |] ==> b : f``A"; +by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1)); +val image_eqI = result(); + +val imageI = refl RS image_eqI; + +val major::prems = goalw Set.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)); +val imageE = result(); + +goal Set.thy "(f o g)``r = f``(g``r)"; +by (stac o_def 1); +by (rtac set_ext 1); +by (fast_tac (HOL_cs addIs [imageI] addSEs [imageE]) 1); +val image_compose = result(); + +goal Set.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); +val image_Un = result(); + +(*** inj(f): f is a one-to-one function ***) + +val prems = goalw Set.thy [inj_def] + "[| !! x y. f(x) = f(y) ==> x=y |] ==> inj(f)"; +by (fast_tac (HOL_cs addIs prems) 1); +val injI = result(); + +val [major] = goal Set.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); +val inj_inverseI = result(); + +val [major,minor] = goalw Set.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); +val injD = result(); + +(*Useful with the simplifier*) +val [major] = goal Set.thy "inj(f) ==> (f(x) = f(y)) = (x=y)"; +by (rtac iffI 1); +by (etac (major RS injD) 1); +by (etac arg_cong 1); +val inj_eq = result(); + +val [major] = goal Set.thy "inj(f) ==> (@x.f(x)=f(y)) = y"; +by (rtac (major RS injD) 1); +by (rtac selectI 1); +by (rtac refl 1); +val inj_select = result(); + +(*A one-to-one function has an inverse (given using select).*) +val [major] = goal Set.thy "inj(f) ==> Inv(f,f(x)) = x"; +by (EVERY1 [stac Inv_def, rtac (major RS inj_select)]); +val Inv_f_f = result(); + +(* Useful??? *) +val [oneone,minor] = goal Set.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); +val inj_transfer = result(); + + +(*** inj_onto(f,A): f is one-to-one over A ***) + +val prems = goalw Set.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); +val inj_ontoI = result(); + +val [major] = goal Set.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)); +val inj_onto_inverseI = result(); + +val major::prems = goalw Set.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)); +val inj_ontoD = result(); + +val major::prems = goal Set.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)); +val inj_onto_contraD = result(); + + +(*** Lemmas about inj ***) + +val prems = goal Set.thy + "[| inj(f); inj_onto(g,range(f)) |] ==> inj(g o f)"; +by (stac o_def 1); +by (cut_facts_tac prems 1); +by (fast_tac (HOL_cs addIs [injI,rangeI] + addEs [injD,inj_ontoD]) 1); +val comp_inj = result(); + +val [prem] = goal Set.thy "inj(f) ==> inj_onto(f,A)"; +by (fast_tac (HOL_cs addIs [prem RS injD, inj_ontoI]) 1); +val inj_imp = result(); + +val [prem] = goal Set.thy "y : range(f) ==> f(Inv(f,y)) = y"; +by (EVERY1 [stac Inv_def, rtac (prem RS rangeE), rtac selectI, etac sym]); +val f_Inv_f = result(); + +val prems = goal Set.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)); +val Inv_injective = result(); + +val prems = goal Set.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); +val inj_onto_Inv = result(); + + +(*** Set reasoning tools ***) + +val set_cs = HOL_cs + addSIs [ballI, subsetI, InterI, INT_I, INT1_I, CollectI, + ComplI, IntI, DiffI, UnCI, insertCI] + addIs [bexI, UnionI, UN_I, UN1_I, imageI, rangeI] + addSEs [bexE, 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 Set.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 set_ss = HOL_ss addsimps mem_simps; diff -r 000000000000 -r 7949f97df77a gfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/gfp.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,136 @@ +(* 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] "[| A <= f(A) |] ==> A <= gfp(f)"; +by (rtac (CollectI RS Union_upper) 1); +by (resolve_tac prems 1); +val gfp_upperbound = result(); + +val prems = goalw Gfp.thy [gfp_def] + "[| !!u. u <= f(u) ==> u<=A |] ==> gfp(f) <= A"; +by (REPEAT (ares_tac ([Union_least]@prems) 1)); +by (etac CollectD 1); +val gfp_least = result(); + +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]); +val gfp_lemma2 = result(); + +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]); +val gfp_lemma3 = result(); + +val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))"; +by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1)); +val gfp_Tarski = result(); + +(*** Coinduction rules for greatest fixed points ***) + +(*weak version*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) |] ==> a : gfp(f)"; +by (rtac (gfp_upperbound RS subsetD) 1); +by (REPEAT (ares_tac prems 1)); +val coinduct = result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(A) Un gfp(f); mono(f) |] ==> \ +\ A Un gfp(f) <= f(A Un gfp(f))"; +by (rtac subset_trans 1); +by (rtac (mono RS mono_Un) 2); +by (rtac (mono RS gfp_Tarski RS subst) 1); +by (rtac (prem RS Un_least) 1); +by (rtac Un_upper2 1); +val coinduct2_lemma = result(); + +(*strong version, thanks to Martin Coen*) +val prems = goal Gfp.thy + "[| a: A; A <= f(A) Un gfp(f); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct2_lemma RSN (2,coinduct)) 1); +by (REPEAT (resolve_tac (prems@[UnI1]) 1)); +val coinduct2 = result(); + +(*** Even Stronger version of coinduct [by Martin Coen] + - instead of the condition A <= f(A) + consider A <= (f(A) Un f(f(A)) ...) Un gfp(A) ***) + +val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un A Un B)"; +by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1)); +val coinduct3_mono_lemma= result(); + +val [prem,mono] = goal Gfp.thy + "[| A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> \ +\ lfp(%x.f(x) Un A Un gfp(f)) <= f(lfp(%x.f(x) Un A 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); +val coinduct3_lemma = result(); + +val prems = goal Gfp.thy + "[| a:A; A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> a : gfp(f)"; +by (rtac (coinduct3_lemma RSN (2,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)); +val coinduct3 = result(); + + +(** Definition forms of gfp_Tarski and coinduct, to control unfolding **) + +val [rew,mono] = goal Gfp.thy "[| h==gfp(f); mono(f) |] ==> h = f(h)"; +by (rewtac rew); +by (rtac (mono RS gfp_Tarski) 1); +val def_gfp_Tarski = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (prems @ [coinduct]) 1)); +val def_coinduct = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(A) Un h; mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct2]) 1)); +val def_coinduct2 = result(); + +val rew::prems = goal Gfp.thy + "[| h==gfp(f); a:A; A <= f(lfp(%x.f(x) Un A Un h)); mono(f) |] ==> a: h"; +by (rewtac rew); +by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1)); +val def_coinduct3 = result(); + +(*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; +val gfp_mono = result(); diff -r 000000000000 -r 7949f97df77a gfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/gfp.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,14 @@ +(* Title: HOL/gfp.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Greatest fixed points +*) + +Gfp = Lfp + +consts gfp :: "['a set=>'a set] => 'a set" +rules + (*greatest fixed point*) + gfp_def "gfp(f) == Union({u. u <= f(u)})" +end diff -r 000000000000 -r 7949f97df77a hol.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/hol.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,334 @@ +(* 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; + +signature HOL_LEMMAS = + sig + val allE: thm + val all_dupE: thm + val allI: thm + val arg_cong: thm + val fun_cong: thm + val box_equals: thm + val ccontr: thm + val classical: thm + val cong: thm + val conjunct1: thm + val conjunct2: thm + val conjE: thm + val conjI: thm + val contrapos: thm + val disjCI: thm + val disjE: thm + val disjI1: thm + val disjI2: thm + val eqTrueI: thm + val eqTrueE: thm + val ex1E: thm + val ex1I: thm + val exCI: thm + val exI: thm + val exE: thm + val excluded_middle: thm + val FalseE: thm + val False_neq_True: thm + val iffCE : thm + val iffD1: thm + val iffD2: thm + val iffE: thm + val iffI: thm + val impCE: thm + val impE: thm + val not_sym: thm + val notE: thm + val notI: thm + val notnotD : thm + val rev_mp: thm + val select_equality: thm + val spec: thm + val sstac: thm list -> int -> tactic + val ssubst: thm + val stac: thm -> int -> tactic + val strip_tac: int -> tactic + val swap: thm + val sym: thm + val trans: thm + val TrueI: thm + end; + +structure HOL_Lemmas : HOL_LEMMAS = + +struct + +(** Equality **) + +val sym = prove_goal 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*) +val ssubst = standard (sym RS subst); + +val trans = prove_goal 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 *) +val box_equals = prove_goal 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*) +val fun_cong = prove_goal 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*) +val arg_cong = prove_goal HOL.thy "x=y ==> f(x)=f(y)" + (fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]); + +val cong = prove_goal 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 **) + +val iffI = prove_goal HOL.thy + "[| P ==> Q; Q ==> P |] ==> P=Q" + (fn prems=> [ (REPEAT (ares_tac (prems@[impI, iff RS mp RS mp]) 1)) ]); + +val iffD2 = prove_goal HOL.thy "[| P=Q; Q |] ==> P" + (fn prems => + [rtac ssubst 1, resolve_tac prems 1, resolve_tac prems 1]); + +val iffD1 = sym RS iffD2; + +val iffE = prove_goal 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 **) + +val TrueI = refl RS (True_def RS iffD2); + +val eqTrueI = prove_goal HOL.thy "P ==> P=True" + (fn prems => [REPEAT(resolve_tac ([iffI,TrueI]@prems) 1)]); + +val eqTrueE = prove_goal HOL.thy "P=True ==> P" + (fn prems => [REPEAT(resolve_tac (prems@[TrueI,iffD2]) 1)]); + +(** Universal quantifier **) + +val allI = prove_goal HOL.thy "(!!x::'a. P(x)) ==> !x. P(x)" + (fn [asm] => [rtac (All_def RS ssubst) 1, rtac (asm RS (eqTrueI RS ext)) 1]); + +val spec = prove_goal HOL.thy "! x::'a.P(x) ==> P(x)" + (fn prems => + [ rtac eqTrueE 1, + resolve_tac (prems RL [All_def RS subst] RL [fun_cong]) 1 ]); + +val allE = prove_goal HOL.thy "[| !x.P(x); P(x) ==> R |] ==> R" + (fn major::prems=> + [ (REPEAT (resolve_tac (prems @ [major RS spec]) 1)) ]); + +val all_dupE = prove_goal 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! **) + +val FalseE = prove_goal HOL.thy "False ==> P" + (fn prems => [rtac spec 1, rtac (False_def RS subst) 1, resolve_tac prems 1]); + +val False_neq_True = prove_goal HOL.thy "False=True ==> P" + (fn [prem] => [rtac (prem RS eqTrueE RS FalseE) 1]); + + +(** Negation **) + +val notI = prove_goal HOL.thy "(P ==> False) ==> ~P" + (fn prems=> [rtac (not_def RS ssubst) 1, rtac impI 1, eresolve_tac prems 1]); + +val notE = prove_goal HOL.thy "[| ~P; P |] ==> R" + (fn prems => + [rtac (mp RS FalseE) 1, + resolve_tac prems 2, rtac (not_def RS subst) 1, + resolve_tac prems 1]); + +(** Implication **) + +val impE = prove_goal 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. *) +val rev_mp = prove_goal HOL.thy "[| P; P --> Q |] ==> Q" + (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); + +val contrapos = prove_goal 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 **) + +val exI = prove_goal HOL.thy "P(x) ==> ? x::'a.P(x)" + (fn prems => + [rtac (selectI RS (Ex_def RS ssubst)) 1, + resolve_tac prems 1]); + +val exE = prove_goal HOL.thy "[| ? x::'a.P(x); !!x. P(x) ==> Q |] ==> Q" + (fn prems => + [resolve_tac prems 1, res_inst_tac [("P","%C.C(P)")] subst 1, + rtac Ex_def 1, resolve_tac prems 1]); + + +(** Conjunction **) + +val conjI = prove_goal HOL.thy "[| P; Q |] ==> P&Q" + (fn prems => + [ (rtac (and_def RS ssubst) 1), + (REPEAT (resolve_tac (prems@[allI,impI]) 1 ORELSE etac (mp RS mp) 1)) ]); + +val conjunct1 = prove_goal HOL.thy "[| P & Q |] ==> P" + (fn prems => + [ (resolve_tac (prems RL [and_def RS subst] RL [spec] RL [mp]) 1), + (REPEAT(ares_tac [impI] 1)) ]); + +val conjunct2 = prove_goal HOL.thy "[| P & Q |] ==> Q" + (fn prems => + [ (resolve_tac (prems RL [and_def RS subst] RL [spec] RL [mp]) 1), + (REPEAT(ares_tac [impI] 1)) ]); + +val conjE = prove_goal 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 *) + +val disjI1 = prove_goal HOL.thy "P ==> P|Q" + (fn [prem] => + [rtac (or_def RS ssubst) 1, + REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]); + +val disjI2 = prove_goal HOL.thy "Q ==> P|Q" + (fn [prem] => + [rtac (or_def RS ssubst) 1, + REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]); + +val disjE = prove_goal HOL.thy "[| P | Q; P ==> R; Q ==> R |] ==> R" + (fn [a1,a2,a3] => + [rtac (mp RS mp) 1, rtac spec 1, rtac (or_def RS subst) 1, rtac a1 1, + rtac (a2 RS impI) 1, atac 1, rtac (a3 RS impI) 1, atac 1]); + +(** CCONTR -- classical logic **) + +val ccontr = prove_goal HOL.thy "(~P ==> False) ==> P" + (fn prems => + [rtac (True_or_False RS (disjE RS eqTrueE)) 1, atac 1, + rtac spec 1, rtac (False_def RS subst) 1, resolve_tac prems 1, + rtac ssubst 1, atac 1, rtac (not_def RS ssubst) 1, + REPEAT (ares_tac [impI] 1) ]); + +val classical = prove_goal HOL.thy "(~P ==> P) ==> P" + (fn prems => + [rtac ccontr 1, + REPEAT (ares_tac (prems@[notE]) 1)]); + + +(*Double negation law*) +val notnotD = prove_goal HOL.thy "~~P ==> P" + (fn [major]=> + [ (rtac classical 1), (eresolve_tac [major RS notE] 1) ]); + + +(** Unique existence **) + +val ex1I = prove_goal HOL.thy + "[| P(a); !!x. P(x) ==> x=a |] ==> ?! x. P(x)" + (fn prems => + [ (rtac (Ex1_def RS ssubst) 1), + (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); + +val ex1E = prove_goal HOL.thy + "[| ?! x.P(x); !!x. [| P(x); ! y. P(y) --> y=x |] ==> R |] ==> R" + (fn major::prems => + [ (resolve_tac ([major] RL [Ex1_def RS subst] RL [exE]) 1), + (REPEAT (etac conjE 1 ORELSE ares_tac prems 1)) ]); + + +(** Select: Hilbert's Epsilon-operator **) + +val select_equality = prove_goal HOL.thy + "[| P(a); !!x. P(x) ==> x=a |] ==> (@x.P(x)) = a" + (fn prems => [ resolve_tac prems 1, + rtac selectI 1, + resolve_tac prems 1 ]); + +(** Classical intro rules for disjunction and existential quantifiers *) + +val disjCI = prove_goal 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)) ]); + +val excluded_middle = prove_goal HOL.thy "~P | P" + (fn _ => [ (REPEAT (ares_tac [disjCI] 1)) ]); + +(*Classical implies (-->) elimination. *) +val impCE = prove_goal 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. *) +val iffCE = prove_goal 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))) ]); + +val exCI = prove_goal HOL.thy "(! x. ~P(x) ==> P(a)) ==> ? x.P(x)" + (fn prems=> + [ (rtac ccontr 1), + (REPEAT (ares_tac (prems@[exI,allI,notI,notE]) 1)) ]); + +(*Required by the "classical" module*) +val swap = prove_goal HOL.thy "~P ==> (~Q ==> P) ==> Q" + (fn major::prems=> + [ rtac ccontr 1, rtac (major RS notE) 1, REPEAT (ares_tac prems 1)]); + +(** 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); + +end; + +open HOL_Lemmas; + diff -r 000000000000 -r 7949f97df77a hol.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/hol.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,142 @@ +(* Title: HOL/hol.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge + +Higher-Order Logic +*) + +HOL = Pure + + +classes + term < logic + plus < term + minus < term + times < term + +default term + +types + bool 0 + +arities + fun :: (term, term) term + bool :: term + + +consts + + (* Constants *) + + Trueprop :: "bool => prop" ("(_)" [0] 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" + "@let" :: "[idt,'a,'b] => 'b" ("(let _ = (2_)/ in (2_))" 10) + + (* Alternative Quantifiers *) + + "*All" :: "[idts, bool] => bool" ("(3ALL _./ _)" [0,0] 10) + "*Ex" :: "[idts, bool] => bool" ("(3EX _./ _)" [0,0] 10) + "*Ex1" :: "[idts, bool] => bool" ("(3EX! _./ _)" [0,0] 10) + + (* Infixes *) + + o :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixr 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) + + (* Rewriting Gadget *) + + NORM :: "'a => 'a" + + +translations + "ALL xs. P" => "! xs. P" + "EX xs. P" => "? xs. P" + "EX! xs. P" => "?! xs. P" + + "let x = s in t" == "Let(s,%x.t)" + +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" + + (* Definitions *) + + True_def "True = ((%x.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 "not = (%P. P-->False)" + and_def "op & = (%P Q. !R. (P-->Q-->R) --> R)" + or_def "op | = (%P Q. !R. (P-->R) --> (Q-->R) --> R)" + Ex1_def "Ex1 = (%P. ? x. P(x) & (! y. P(y) --> y=x))" + Let_def "Let(s,f) = f(s)" + + (* Axioms *) + + iff "(P-->Q) --> (Q-->P) --> (P=Q)" + True_or_False "(P=True) | (P=False)" + + (* Misc Definitions *) + + Inv_def "Inv = (%(f::'a=>'b) y. @x. f(x)=y)" + o_def "op o = (%(f::'b=>'c) g (x::'a). f(g(x)))" + + if_def "if = (%P x y.@z::'a. (P=True --> z=x) & (P=False --> z=y))" + + (* Rewriting -- special constant to flag normalized terms *) + + NORM_def "NORM(x) = x" + +end + + +ML + +(** Choice between the HOL and Isabelle style of quantifiers **) + +val HOL_quantifiers = ref true; + +fun mk_alt_ast_tr' (name, alt_name) = + let + fun ast_tr' (*name*) args = + if ! HOL_quantifiers then raise Match + else Ast.mk_appl (Ast.Constant alt_name) args; + in + (name, ast_tr') + end; + + +val print_ast_translation = + map mk_alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")]; + diff -r 000000000000 -r 7949f97df77a lfp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/lfp.ML Thu Sep 16 12:21:07 1993 +0200 @@ -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); +val lfp_lowerbound = result(); + +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); +val lfp_greatest = result(); + +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]); +val lfp_lemma2 = result(); + +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]); +val lfp_lemma3 = result(); + +val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))"; +by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1)); +val lfp_Tarski = result(); + + +(*** 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]); +val induct = result(); + +(** 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); +val def_lfp_Tarski = result(); + +val rew::prems = goal Lfp.thy + "[| A == lfp(f); a:A; mono(f); \ +\ !!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))]); +val def_induct = result(); + +(*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; +val lfp_mono = result(); diff -r 000000000000 -r 7949f97df77a lfp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/lfp.thy Thu Sep 16 12:21:07 1993 +0200 @@ -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 = Sum + +consts lfp :: "['a set=>'a set] => 'a set" +rules + (*least fixed point*) + lfp_def "lfp(f) == Inter({u. f(u) <= u})" +end diff -r 000000000000 -r 7949f97df77a llist.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/llist.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,899 @@ +(* Title: HOL/llist + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For llist.thy. + +SHOULD LListD_Fun_CONS_I, etc., be equations (for rewriting) +*) + +open LList; + +(** Simplification **) + +val llist_simps = [case_Inl, case_Inr]; +val llist_ss = univ_ss addsimps llist_simps + setloop (split_tac [expand_split,expand_case]); + +(** the llist functional **) + +val LList_unfold = rewrite_rule [List_Fun_def] + (List_Fun_mono RS (LList_def RS def_gfp_Tarski)); + +(*This justifies using LList in other recursive type definitions*) +goalw LList.thy [LList_def] "!!A B. A<=B ==> LList(A) <= LList(B)"; +by (rtac gfp_mono 1); +by (etac List_Fun_mono2 1); +val LList_mono = result(); + +(*Elimination is case analysis, not induction.*) +val [major,prem1,prem2] = goalw LList.thy [NIL_def,CONS_def] + "[| L : LList(A); \ +\ L=NIL ==> P; \ +\ !!M N. [| M:A; N: LList(A); L=CONS(M,N) |] ==> P \ +\ |] ==> P"; +by (rtac (major RS (LList_unfold RS equalityD1 RS subsetD RS usumE)) 1); +by (etac uprodE 2); +by (rtac prem2 2); +by (rtac prem1 1); +by (REPEAT (ares_tac [refl] 1 + ORELSE eresolve_tac [singletonE,ssubst] 1)); +val LListE = result(); + + +(*** Type checking by co-induction, using List_Fun ***) + +val prems = goalw LList.thy [LList_def] + "[| M: X; X <= List_Fun(A,X) |] ==> M: LList(A)"; +by (REPEAT (resolve_tac (prems@[coinduct]) 1)); +val LList_coinduct = result(); + +(*stronger version*) +val prems = goalw LList.thy [LList_def] + "[| M : X; X <= List_Fun(A, X) Un LList(A) |] ==> M : LList(A)"; +by (REPEAT (resolve_tac (prems@[coinduct2,List_Fun_mono]) 1)); +val LList_coinduct2 = result(); + +(** Rules to prove the 2nd premise of LList_coinduct **) + +goalw LList.thy [List_Fun_def,NIL_def] "NIL: List_Fun(A,X)"; +by (resolve_tac [singletonI RS usum_In0I] 1); +val List_Fun_NIL_I = result(); + +goalw LList.thy [List_Fun_def,CONS_def] + "!!M N. [| M: A; N: X |] ==> CONS(M,N) : List_Fun(A,X)"; +by (REPEAT (ares_tac [uprodI RS usum_In1I] 1)); +val List_Fun_CONS_I = result(); + +(*** LList_corec satisfies the desired recurion equation ***) + +(*A continuity result?*) +goalw LList.thy [CONS_def] "CONS(M, UN x.f(x)) = (UN x. CONS(M, f(x)))"; +by(simp_tac (univ_ss addsimps [In1_UN1, Scons_UN1_y]) 1); +val CONS_UN1 = result(); + +goal Prod.thy "split(p, %x y.UN z.f(x,y,z)) = (UN z. split(p, %x y.f(x,y,z)))"; +by(simp_tac (pair_ss setloop (split_tac [expand_split])) 1); +val split_UN1 = result(); + +goal Sum.thy "case(s, f, %y. UN z.g(y,z)) = (UN z. case(s, f, %y. g(y,z)))"; +by(simp_tac (sum_ss setloop (split_tac [expand_case])) 1); +val case2_UN1 = result(); + +val prems = goalw LList.thy [CONS_def] + "[| M<=M'; N<=N' |] ==> CONS(M,N) <= CONS(M',N')"; +by (REPEAT (resolve_tac ([In1_mono,Scons_mono]@prems) 1)); +val CONS_mono = result(); + +val corec_fun_simps = [LList_corec_fun_def RS def_nat_rec_0, + LList_corec_fun_def RS def_nat_rec_Suc]; +val corec_fun_ss = llist_ss addsimps corec_fun_simps; + +(** The directions of the equality are proved separately **) + +goalw LList.thy [LList_corec_def] + "LList_corec(a,f) <= case(f(a), %u.NIL, \ +\ %v. split(v, %z w. CONS(z, LList_corec(w,f))))"; +by (rtac UN1_least 1); +by (nat_ind_tac "k" 1); +by(ALLGOALS(simp_tac corec_fun_ss)); +by (REPEAT (resolve_tac [allI, impI, subset_refl RS CONS_mono, UN1_upper] 1)); +val LList_corec_subset1 = result(); + +goalw LList.thy [LList_corec_def] + "case(f(a), %u.NIL, %v. split(v, %z w. CONS(z, LList_corec(w,f)))) <= \ +\ LList_corec(a,f)"; +by (simp_tac (corec_fun_ss addsimps [CONS_UN1]) 1); +by (safe_tac set_cs); +by (ALLGOALS (res_inst_tac [("x","Suc(?k)")] UN1_I THEN' + asm_simp_tac corec_fun_ss)); +val LList_corec_subset2 = result(); + +(*the recursion equation for LList_corec -- NOT SUITABLE FOR REWRITING!*) +goal LList.thy + "LList_corec(a,f) = case(f(a), %u. NIL, \ +\ %v. split(v, %z w. CONS(z, LList_corec(w,f))))"; +by (REPEAT (resolve_tac [equalityI, LList_corec_subset1, + LList_corec_subset2] 1)); +val LList_corec = result(); + +(*definitional version of same*) +val [rew] = goal LList.thy + "[| !!x. h(x) == LList_corec(x,f) |] ==> \ +\ h(a) = case(f(a), %u.NIL, %v. split(v, %z w. CONS(z, h(w))))"; +by (rewtac rew); +by (rtac LList_corec 1); +val def_LList_corec = result(); + +(*A typical use of co-induction to show membership in the gfp. + Bisimulation is range(%x. LList_corec(x,f)) *) +goal LList.thy "LList_corec(a,f) : LList({u.True})"; +by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] LList_coinduct 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac LList_corec 1); +by(simp_tac (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I, + CollectI, range_eqI]) 1); +(* 6.7 vs 3.4 !!! +by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I, + CollectI, rangeI]) 1); +*) +val LList_corec_type = result(); + +(*Lemma for the proof of llist_corec*) +goal LList.thy + "LList_corec(a, %z. case(f(z),Inl,%x. split(x,%v w. Inr()))) : \ +\ LList(range(Leaf))"; +by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] LList_coinduct 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac LList_corec 1); +(*nested "case"; requires an explicit split*) +by (res_inst_tac [("s", "f(xa)")] sumE 1); +by(asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_NIL_I])) 1); +by(asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_CONS_I, range_eqI]) + setloop (split_tac [expand_split])) 1); +(* FIXME: can the selection of the case split be automated? +by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_CONS_I, rangeI]) 1);*) +val LList_corec_type2 = result(); + +(**** LList equality as a gfp; the bisimulation principle ****) + +goalw LList.thy [LListD_Fun_def] "mono(LListD_Fun(r))"; +by (REPEAT (ares_tac [monoI, subset_refl, dsum_mono, dprod_mono] 1)); +val LListD_fun_mono = result(); + +val LListD_unfold = rewrite_rule [LListD_Fun_def] + (LListD_fun_mono RS (LListD_def RS def_gfp_Tarski)); + +goal LList.thy "!M N. : LListD(diag(A)) --> ntrunc(k,M) = ntrunc(k,N)"; +by (res_inst_tac [("n", "k")] less_induct 1); +by (safe_tac set_cs); +by (etac (LListD_unfold RS equalityD1 RS subsetD RS dsumE) 1); +by (safe_tac (set_cs addSEs [Pair_inject, dprodE, diagE])); +by (res_inst_tac [("n", "n")] natE 1); +by(asm_simp_tac (univ_ss addsimps [ntrunc_0]) 1); +by (res_inst_tac [("n", "xb")] natE 1); +by(asm_simp_tac (univ_ss addsimps [ntrunc_one_In1]) 1); +by(asm_simp_tac (univ_ss addsimps [ntrunc_In1, ntrunc_Scons]) 1); +val LListD_implies_ntrunc_equality = result(); + +goalw LList.thy [LList_def,List_Fun_def] "fst``LListD(diag(A)) <= LList(A)"; +by (rtac gfp_upperbound 1); +by (res_inst_tac [("P", "%x. fst``x <= ?B")] (LListD_unfold RS ssubst) 1); +by(simp_tac fst_image_ss 1); +val fst_image_LListD = result(); + +(*This inclusion justifies the use of coinduction to show M=N*) +goal LList.thy "LListD(diag(A)) <= diag(LList(A))"; +by (rtac subsetI 1); +by (res_inst_tac [("p","x")] PairE 1); +by (safe_tac HOL_cs); +by (res_inst_tac [("s","xa")] subst 1); +by (rtac (LListD_implies_ntrunc_equality RS spec RS spec RS mp RS + ntrunc_equality) 1); +by (assume_tac 1); +by (rtac diagI 1); +by (etac (fst_imageI RS (fst_image_LListD RS subsetD)) 1); +val LListD_subset_diag = result(); + +(*This converse inclusion helps to strengthen LList_equalityI*) +goalw LList.thy [LListD_def] "diag(LList(A)) <= LListD(diag(A))"; +by (rtac gfp_upperbound 1); +by (rtac subsetI 1); +by (etac diagE 1); +by (etac ssubst 1); +by (etac (LList_unfold RS equalityD1 RS subsetD RS usumE) 1); +by (rewtac LListD_Fun_def); +by (ALLGOALS (fast_tac (set_cs addIs [diagI,dsum_In0I,dsum_In1I,dprodI] + addSEs [uprodE]))); +val diag_subset_LListD = result(); + +goal LList.thy "LListD(diag(A)) = diag(LList(A))"; +by (REPEAT (resolve_tac [equalityI, LListD_subset_diag, + diag_subset_LListD] 1)); +val LListD_eq_diag = result(); + +(** To show two LLists are equal, exhibit a bisimulation! **) +(* Replace "A" by some particular set, like {x.True}??? *) +val prems = goal LList.thy + "[| : r; r <= LListD_Fun(diag(A), r) |] ==> M=N"; +by (rtac (rewrite_rule [LListD_def] + (LListD_subset_diag RS subsetD RS diagE)) 1); +by (REPEAT (resolve_tac (prems@[coinduct]) 1)); +by (safe_tac (set_cs addSEs [Pair_inject])); +val LList_equalityI = result(); + +(*Stronger notion of bisimulation -- also admits true equality*) +val prems = goal LList.thy + "[| : r; r <= LListD_Fun(diag(A), r) Un diag(LList(A)) |] ==> M=N"; +by (rtac (rewrite_rule [LListD_def] + (LListD_subset_diag RS subsetD RS diagE)) 1); +by (rtac coinduct2 1); +by (stac (rewrite_rule [LListD_def] LListD_eq_diag) 2); +by (REPEAT (resolve_tac (prems@[LListD_fun_mono]) 1)); +by (safe_tac (set_cs addSEs [Pair_inject])); +val LList_equalityI2 = result(); + +(** Rules to prove the 2nd premise of LList_equalityI **) + +goalw LList.thy [LListD_Fun_def,NIL_def] " : LListD_Fun(r,s)"; +by (rtac (singletonI RS diagI RS dsum_In0I) 1); +val LListD_Fun_NIL_I = result(); + +val prems = goalw LList.thy [LListD_Fun_def,CONS_def] + "[| x:A; :s |] ==> : LListD_Fun(diag(A),s)"; +by (rtac (dprodI RS dsum_In1I) 1); +by (REPEAT (resolve_tac (diagI::prems) 1)); +val LListD_Fun_CONS_I = result(); + + +(*** Finality of LList(A): Uniqueness of functions defined by corecursion ***) + +(*abstract proof using a bisimulation*) +val [prem1,prem2] = goal LList.thy + "[| !!x. h1(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h1(w)))); \ +\ !!x. h2(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h2(w)))) |] \ +\ ==> h1=h2"; +by (rtac ext 1); +(*next step avoids an unknown (and flexflex pair) in simplification*) +by (res_inst_tac [("A", "{u.True}"), + ("r", "range(%u. )")] LList_equalityI 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac prem1 1); +by (stac prem2 1); +by(simp_tac (llist_ss addsimps [LListD_Fun_NIL_I, range_eqI, + CollectI RS LListD_Fun_CONS_I]) 1); +(* 9.5 vs 9.2/4.1/4.3 +by (ASM_SIMP_TAC (llist_ss addsimps [LListD_Fun_NIL_I, rangeI, + CollectI RS LListD_Fun_CONS_I]) 1);*) +val LList_corec_unique = result(); + +val [prem] = goal LList.thy + "[| !!x. h(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h(w)))) |] \ +\ ==> h = (%x.LList_corec(x,f))"; +by (rtac (LList_corec RS (prem RS LList_corec_unique)) 1); +val equals_LList_corec = result(); + + +(** Obsolete LList_corec_unique proof: complete induction, not coinduction **) + +goalw LList.thy [CONS_def] "ntrunc(Suc(0), CONS(M,N)) = {}"; +by (rtac ntrunc_one_In1 1); +val ntrunc_one_CONS = result(); + +goalw LList.thy [CONS_def] + "ntrunc(Suc(Suc(k)), CONS(M,N)) = CONS (ntrunc(k,M), ntrunc(k,N))"; +by(simp_tac (HOL_ss addsimps [ntrunc_Scons,ntrunc_In1]) 1); +val ntrunc_CONS = result(); + +val [prem1,prem2] = goal LList.thy + "[| !!x. h1(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h1(w)))); \ +\ !!x. h2(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h2(w)))) |] \ +\ ==> h1=h2"; +by (rtac (ntrunc_equality RS ext) 1); +by (res_inst_tac [("x", "x")] spec 1); +by (res_inst_tac [("n", "k")] less_induct 1); +by (rtac allI 1); +by (stac prem1 1); +by (stac prem2 1); +by(simp_tac (sum_ss setloop (split_tac [expand_split,expand_case])) 1); +by (strip_tac 1); +by (res_inst_tac [("n", "n")] natE 1); +by (res_inst_tac [("n", "xc")] natE 2); +by(ALLGOALS(asm_simp_tac(nat_ss addsimps + [ntrunc_0,ntrunc_one_CONS,ntrunc_CONS]))); +val LList_corec_unique = result(); + + +(*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***) + +goal LList.thy "mono(CONS(M))"; +by (REPEAT (ares_tac [monoI, subset_refl, CONS_mono] 1)); +val Lconst_fun_mono = result(); + +(* Lconst(M) = CONS(M,Lconst(M)) *) +val Lconst = standard (Lconst_fun_mono RS (Lconst_def RS def_lfp_Tarski)); + +(*A typical use of co-induction to show membership in the gfp. + The containing set is simply the singleton {Lconst(M)}. *) +goal LList.thy "!!M A. M:A ==> Lconst(M): LList(A)"; +by (rtac (singletonI RS LList_coinduct) 1); +by (safe_tac set_cs); +by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1); +by (REPEAT (ares_tac [List_Fun_CONS_I, singletonI] 1)); +val Lconst_type = result(); + +goal LList.thy "Lconst(M) = LList_corec(M, %x.Inr())"; +by (rtac (equals_LList_corec RS fun_cong) 1); +by(simp_tac sum_ss 1); +by (rtac Lconst 1); +val Lconst_eq_LList_corec = result(); + +(*Thus we could have used gfp in the definition of Lconst*) +goal LList.thy "gfp(%N. CONS(M, N)) = LList_corec(M, %x.Inr())"; +by (rtac (equals_LList_corec RS fun_cong) 1); +by(simp_tac sum_ss 1); +by (rtac (Lconst_fun_mono RS gfp_Tarski) 1); +val gfp_Lconst_eq_LList_corec = result(); + + +(** Introduction rules for LList constructors **) + +(* c : {Numb(0)} <+> A <*> LList(A) ==> c : LList(A) *) +val LListI = LList_unfold RS equalityD2 RS subsetD; + +(*This justifies the type definition: LList(A) is nonempty.*) +goalw LList.thy [NIL_def] "NIL: LList(A)"; +by (rtac (singletonI RS usum_In0I RS LListI) 1); +val NIL_LListI = result(); + +val prems = goalw LList.thy [CONS_def] + "[| M: A; N: LList(A) |] ==> CONS(M,N) : LList(A)"; +by (rtac (uprodI RS usum_In1I RS LListI) 1); +by (REPEAT (resolve_tac prems 1)); +val CONS_LListI = result(); + +(*** Isomorphisms ***) + +goal LList.thy "inj(Rep_LList)"; +by (rtac inj_inverseI 1); +by (rtac Rep_LList_inverse 1); +val inj_Rep_LList = result(); + +goal LList.thy "inj_onto(Abs_LList,LList(range(Leaf)))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_LList_inverse 1); +val inj_onto_Abs_LList = result(); + +(** Distinctness of constructors **) + +goalw LList.thy [LNil_def,LCons_def] "~ LCons(x,xs) = LNil"; +by (rtac (CONS_not_NIL RS (inj_onto_Abs_LList RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, NIL_LListI, CONS_LListI, Rep_LList] 1)); +val LCons_not_LNil = result(); + +val LNil_not_LCons = standard (LCons_not_LNil RS not_sym); + +val LCons_neq_LNil = standard (LCons_not_LNil RS notE); +val LNil_neq_LCons = sym RS LCons_neq_LNil; + +(** llist constructors **) + +goalw LList.thy [LNil_def] + "Rep_LList(LNil) = NIL"; +by (rtac (NIL_LListI RS Abs_LList_inverse) 1); +val Rep_LList_LNil = result(); + +goalw LList.thy [LCons_def] + "Rep_LList(LCons(x,l)) = CONS(Leaf(x),Rep_LList(l))"; +by (REPEAT (resolve_tac [CONS_LListI RS Abs_LList_inverse, + rangeI, Rep_LList] 1)); +val Rep_LList_LCons = result(); + +(** Injectiveness of CONS and LCons **) + +goalw LList.thy [CONS_def] "(CONS(M,N)=CONS(M',N')) = (M=M' & N=N')"; +by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1); +val CONS_CONS_eq = result(); + +val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE); + + +(*For reasoning about abstract llist constructors*) +val LList_cs = set_cs addIs [Rep_LList, NIL_LListI, CONS_LListI] + addSEs [CONS_neq_NIL,NIL_neq_CONS,CONS_inject] + addSDs [inj_onto_Abs_LList RS inj_ontoD, + inj_Rep_LList RS injD, Leaf_inject]; + +goalw LList.thy [LCons_def] "(LCons(x,xs)=LCons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac LList_cs 1); +val LCons_LCons_eq = result(); +val LCons_inject = standard (LCons_LCons_eq RS iffD1 RS conjE); + +val [major] = goal LList.thy "CONS(M,N): LList(A) ==> M: A & N: LList(A)"; +by (rtac (major RS LListE) 1); +by (etac CONS_neq_NIL 1); +by (fast_tac LList_cs 1); +val CONS_D = result(); + + +(****** Reasoning about LList(A) ******) + +val List_case_simps = [List_case_NIL, List_case_CONS]; +val List_case_ss = llist_ss addsimps List_case_simps; + +(*A special case of list_equality for functions over lazy lists*) +val [MList,gMList,NILcase,CONScase] = goal LList.thy + "[| M: LList(A); g(NIL): LList(A); \ +\ f(NIL)=g(NIL); \ +\ !!x l. [| x:A; l: LList(A) |] ==> \ +\ : \ +\ LListD_Fun(diag(A), (%u.)``LList(A)) Un \ +\ diag(LList(A)) \ +\ |] ==> f(M) = g(M)"; +by (rtac LList_equalityI2 1); +br (MList RS imageI) 1; +by (rtac subsetI 1); +by (etac imageE 1); +by (etac ssubst 1); +by (etac LListE 1); +by (etac ssubst 1); +by (stac NILcase 1); +br (gMList RS diagI RS UnI2) 1; +by (etac ssubst 1); +by (REPEAT (ares_tac [CONScase] 1)); +val LList_fun_equalityI = result(); + + +(*** The functional "Lmap" ***) + +goal LList.thy "Lmap(f,NIL) = NIL"; +by (rtac (Lmap_def RS def_LList_corec RS trans) 1); +by(simp_tac List_case_ss 1); +val Lmap_NIL = result(); + +goal LList.thy "Lmap(f, CONS(M,N)) = CONS(f(M), Lmap(f,N))"; +by (rtac (Lmap_def RS def_LList_corec RS trans) 1); +by(simp_tac List_case_ss 1); +val Lmap_CONS = result(); + +(*Another type-checking proof by coinduction*) +val [major,minor] = goal LList.thy + "[| M: LList(A); !!x. x:A ==> f(x):B |] ==> Lmap(f,M): LList(B)"; +by (rtac (major RS imageI RS LList_coinduct) 1); +by (safe_tac set_cs); +by (etac LListE 1); +by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); +by (REPEAT (ares_tac [List_Fun_NIL_I, List_Fun_CONS_I, minor, imageI] 1)); +val Lmap_type = result(); + +(*This type checking rule synthesises a sufficiently large set for f*) +val [major] = goal LList.thy "M: LList(A) ==> Lmap(f,M): LList(f``A)"; +by (rtac (major RS Lmap_type) 1); +by (etac imageI 1); +val Lmap_type2 = result(); + +(** Two easy results about Lmap **) + +val [prem] = goal LList.thy + "M: LList(A) ==> Lmap(f o g, M) = Lmap(f, Lmap(g, M))"; +by (rtac (prem RS imageI RS LList_equalityI) 1); +by (stac o_def 1); +by (safe_tac set_cs); +by (etac LListE 1); +by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); +by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, + rangeI RS LListD_Fun_CONS_I] 1)); +val Lmap_compose = result(); + +val [prem] = goal LList.thy "M: LList(A) ==> Lmap(%x.x, M) = M"; +by (rtac (prem RS imageI RS LList_equalityI) 1); +by (safe_tac set_cs); +by (etac LListE 1); +by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); +by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, + rangeI RS LListD_Fun_CONS_I] 1)); +val Lmap_ident = result(); + + +(*** Lappend -- its two arguments cause some complications! ***) + +goalw LList.thy [Lappend_def] "Lappend(NIL,NIL) = NIL"; +by (rtac (LList_corec RS trans) 1); +(* takes 2.4(3.4 w NORM) vs 0.9 w/o NORM terms *) +by(simp_tac List_case_ss 1); +(*by (SIMP_TAC List_case_ss 1);*) +val Lappend_NIL_NIL = result(); + +goalw LList.thy [Lappend_def] + "Lappend(NIL,CONS(N,N')) = CONS(N, Lappend(NIL,N'))"; +by (rtac (LList_corec RS trans) 1); +(* takes 5(7 w NORM) vs 2.1 w/o NORM terms *) +by(simp_tac List_case_ss 1); +(*by (SIMP_TAC List_case_ss 1);*) +val Lappend_NIL_CONS = result(); + +goalw LList.thy [Lappend_def] + "Lappend(CONS(M,M'), N) = CONS(M, Lappend(M',N))"; +by (rtac (LList_corec RS trans) 1); +(* takes 4.9(6.7) vs 2.2 w/o NORM terms *) +by(simp_tac List_case_ss 1); +(*by (SIMP_TAC List_case_ss 1);*) +val Lappend_CONS = result(); + +val Lappend_ss = List_case_ss addsimps + [NIL_LListI, Lappend_NIL_NIL, Lappend_NIL_CONS, + Lappend_CONS, image_eqI, LListD_Fun_CONS_I]; + +goal LList.thy "!!M. M: LList(A) ==> Lappend(NIL,M) = M"; +by (etac LList_fun_equalityI 1); +by (ALLGOALS (asm_simp_tac Lappend_ss)); +val Lappend_NIL = result(); + +goal LList.thy "!!M. M: LList(A) ==> Lappend(M,NIL) = M"; +by (etac LList_fun_equalityI 1); +by (ALLGOALS (asm_simp_tac Lappend_ss)); +val Lappend_NIL2 = result(); + +(** Alternative type-checking proofs for Lappend **) + +(*weak co-induction: bisimulation and case analysis on both variables*) +goal LList.thy + "!!M N. [| M: LList(A); N: LList(A) |] ==> Lappend(M,N): LList(A)"; +by (res_inst_tac + [("X", "UN u:LList(A). UN v: LList(A). {Lappend(u,v)}")] LList_coinduct 1); +by (fast_tac set_cs 1); +by (safe_tac set_cs); +by (eres_inst_tac [("L", "u")] LListE 1); +by (eres_inst_tac [("L", "v")] LListE 1); +(* 7/12 vs 7.8/13.3/8.2/13.4 *) +by (ALLGOALS + (asm_simp_tac Lappend_ss THEN' + fast_tac (set_cs addSIs [NIL_LListI,List_Fun_NIL_I,List_Fun_CONS_I]) )); +(* +by (REPEAT + (ASM_SIMP_TAC Lappend_ss 1 THEN + fast_tac (set_cs addSIs [NIL_LListI,List_Fun_NIL_I,List_Fun_CONS_I])1)); +*) +val Lappend_type = result(); + +(*strong co-induction: bisimulation and case analysis on one variable*) +goal LList.thy + "!!M N. [| M: LList(A); N: LList(A) |] ==> Lappend(M,N): LList(A)"; +by (res_inst_tac [("X", "(%u.Lappend(u,N))``LList(A)")] LList_coinduct2 1); +fe imageI; +br subsetI 1; +be imageE 1; +by (eres_inst_tac [("L", "u")] LListE 1); +by (asm_simp_tac (Lappend_ss addsimps [Lappend_NIL]) 1); +by (asm_simp_tac Lappend_ss 1); +by (fast_tac (set_cs addSIs [List_Fun_CONS_I]) 1); +val Lappend_type = result(); + +(**** Lazy lists as the type 'a llist -- strongly typed versions of above ****) + +(** llist_case: case analysis for 'a llist **) + +val Rep_LList_simps = + [List_case_NIL, List_case_CONS, + Abs_LList_inverse, Rep_LList_inverse, NIL_LListI, CONS_LListI, + Rep_LList, rangeI, inj_Leaf, Inv_f_f]; +val Rep_LList_ss = llist_ss addsimps Rep_LList_simps; + +goalw LList.thy [llist_case_def,LNil_def] "llist_case(LNil, c, d) = c"; +by(simp_tac Rep_LList_ss 1); +val llist_case_LNil = result(); + +goalw LList.thy [llist_case_def,LCons_def] + "llist_case(LCons(M,N), c, d) = d(M,N)"; +by(simp_tac Rep_LList_ss 1); +val llist_case_LCons = result(); + +(*Elimination is case analysis, not induction.*) +val [prem1,prem2] = goalw LList.thy [NIL_def,CONS_def] + "[| l=LNil ==> P; !!x l'. l=LCons(x,l') ==> P \ +\ |] ==> P"; +by (rtac (Rep_LList RS LListE) 1); +by (rtac (inj_Rep_LList RS injD RS prem1) 1); +by (stac Rep_LList_LNil 1); +by (assume_tac 1); +by (etac rangeE 1); +by (rtac (inj_Rep_LList RS injD RS prem2) 1); +by(asm_simp_tac (HOL_ss addsimps [Rep_LList_LCons]) 1); +by (etac (Abs_LList_inverse RS ssubst) 1); +by (rtac refl 1); +val llistE = result(); + +(** llist_corec: corecursion for 'a llist **) + +goalw LList.thy [llist_corec_def,LNil_def,LCons_def] + "llist_corec(a,f) = case(f(a), %u. LNil, \ +\ %v. split(v, %z w. LCons(z, llist_corec(w,f))))"; +by (stac LList_corec 1); +by(res_inst_tac [("s","f(a)")] sumE 1); +by(asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1); +by(res_inst_tac [("p","y")] PairE 1); +by(asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1); +(*FIXME: correct case splits usd to be found automatically: +by(ASM_SIMP_TAC(llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);*) +val llist_corec = result(); + +(*definitional version of same*) +val [rew] = goal LList.thy + "[| !!x. h(x) == llist_corec(x,f) |] ==> \ +\ h(a) = case(f(a), %u.LNil, %v. split(v, %z w. LCons(z, h(w))))"; +by (rewtac rew); +by (rtac llist_corec 1); +val def_llist_corec = result(); + +(**** Proofs about type 'a llist functions ****) + +(*** Deriving llist_equalityI -- llist equality is a bisimulation ***) + +val prems = goalw LList.thy [LListD_Fun_def] + "r <= Sigma(LList(A), %x.LList(A)) ==> \ +\ LListD_Fun(diag(A),r) <= Sigma(LList(A), %x.LList(A))"; +by (stac LList_unfold 1); +by (cut_facts_tac prems 1); +by (fast_tac univ_cs 1); +val LListD_Fun_subset_Sigma_LList = result(); + +goal LList.thy + "prod_fun(Rep_LList,Rep_LList) `` r <= \ +\ Sigma(LList(range(Leaf)), %x.LList(range(Leaf)))"; +by (fast_tac (set_cs addSEs [prod_fun_imageE] addIs [SigmaI, Rep_LList]) 1); +val subset_Sigma_LList = result(); + +val [prem] = goal LList.thy + "r <= Sigma(LList(range(Leaf)), %x.LList(range(Leaf))) ==> \ +\ prod_fun(Rep_LList o Abs_LList, Rep_LList o Abs_LList) `` r <= r"; +by (safe_tac (set_cs addSEs [prod_fun_imageE])); +by (rtac (prem RS subsetD RS SigmaE2) 1); +by (assume_tac 1); +by(asm_simp_tac (HOL_ss addsimps [o_def,prod_fun,Abs_LList_inverse]) 1); +val prod_fun_lemma = result(); + +(** To show two llists are equal, exhibit a bisimulation! **) +val [prem1,prem2] = goalw LList.thy [llistD_Fun_def] + "[| : r; r <= llistD_Fun(r) |] ==> l1=l2"; +by (rtac (inj_Rep_LList RS injD) 1); +by (res_inst_tac [("r", "prod_fun(Rep_LList,Rep_LList)``r")] + LList_equalityI 1); +by (rtac (prem1 RS prod_fun_imageI) 1); +by (rtac (prem2 RS image_mono RS subset_trans) 1); +by (rtac (image_compose RS subst) 1); +by (rtac (prod_fun_compose RS subst) 1); +by (rtac (subset_Sigma_LList RS LListD_Fun_subset_Sigma_LList RS + prod_fun_lemma) 1); +val llist_equalityI = result(); + + +(*Stronger notion of bisimulation -- also admits true equality*) +val [prem1,prem2] = goalw LList.thy [llistD_Fun_def] + "[| : r; r <= llistD_Fun(r) Un range(%x.) |] ==> l1=l2"; +by (rtac (inj_Rep_LList RS injD) 1); +by (res_inst_tac [("r", "prod_fun(Rep_LList,Rep_LList)``r"), + ("A", "range(Leaf)")] + LList_equalityI2 1); +by (rtac (prem1 RS prod_fun_imageI) 1); +by (rtac (prem2 RS image_mono RS subset_trans) 1); +by (rtac (image_Un RS ssubst) 1); +by (rtac Un_least 1); +by (rtac (image_compose RS subst) 1); +by (rtac (prod_fun_compose RS subst) 1); +by (rtac (subset_Sigma_LList RS LListD_Fun_subset_Sigma_LList RS + prod_fun_lemma RS subset_trans) 1); +by (rtac Un_upper1 1); +by (fast_tac (set_cs addSEs [prod_fun_imageE, Pair_inject] + addIs [diagI,Rep_LList]) 1); +val llist_equalityI2 = result(); + +(** Rules to prove the 2nd premise of llist_equalityI **) +goalw LList.thy [llistD_Fun_def,LNil_def] " : llistD_Fun(r)"; +by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1); +val llistD_Fun_LNil_I = result(); + +val [prem] = goalw LList.thy [llistD_Fun_def,LCons_def] + ":r ==> : llistD_Fun(r)"; +by (rtac (rangeI RS LListD_Fun_CONS_I RS prod_fun_imageI) 1); +by (rtac (prem RS prod_fun_imageI) 1); +val llistD_Fun_LCons_I = result(); + + +(*A special case of list_equality for functions over lazy lists*) +val [prem1,prem2] = goal LList.thy + "[| f(LNil)=g(LNil); \ +\ !!x l. : \ +\ llistD_Fun(range(%u. )) Un range(%v. ) \ +\ |] ==> f(l) = g(l :: 'a llist) :: 'b llist"; +by (res_inst_tac [("r", "range(%u. )")] llist_equalityI2 1); +by (rtac rangeI 1); +by (rtac subsetI 1); +by (etac rangeE 1); +by (etac ssubst 1); +by (res_inst_tac [("l", "u")] llistE 1); +by (etac ssubst 1); +by (stac prem1 1); +by (fast_tac set_cs 1); +by (etac ssubst 1); +by (rtac prem2 1); +val llist_fun_equalityI = result(); + +(*simpset for llist bisimulations*) +val llistD_simps = [llist_case_LNil, llist_case_LCons, range_eqI, + llistD_Fun_LNil_I, llistD_Fun_LCons_I]; +val llistD_ss = llist_ss addsimps llistD_simps; + + +(*** The functional "lmap" ***) + +goal LList.thy "lmap(f,LNil) = LNil"; +by (rtac (lmap_def RS def_llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +val lmap_LNil = result(); + +goal LList.thy "lmap(f, LCons(M,N)) = LCons(f(M), lmap(f,N))"; +by (rtac (lmap_def RS def_llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +val lmap_LCons = result(); + + +(** Two easy results about lmap **) + +goal LList.thy "lmap(f o g, l) = lmap(f, lmap(g, l))"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons]))); +val lmap_compose = result(); + +goal LList.thy "lmap(%x.x, l) = l"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons]))); +val lmap_ident = result(); + + +(*** iterates -- llist_fun_equalityI cannot be used! ***) + +goal LList.thy "iterates(f,x) = LCons(x, iterates(f,f(x)))"; +by (rtac (iterates_def RS def_llist_corec RS trans) 1); +by(simp_tac sum_ss 1); +val iterates = result(); + +goal LList.thy "lmap(f, iterates(f,x)) = iterates(f,f(x))"; +by (res_inst_tac [("r", "range(%u.)")] + llist_equalityI 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (res_inst_tac [("x1", "f(u)")] (iterates RS ssubst) 1); +by (res_inst_tac [("x1", "u")] (iterates RS ssubst) 1); +by (simp_tac (llistD_ss addsimps [lmap_LCons]) 1); +val lmap_iterates = result(); + +goal LList.thy "iterates(f,x) = LCons(x, lmap(f, iterates(f,x)))"; +br (lmap_iterates RS ssubst) 1; +br iterates 1; +val iterates_lmap = result(); + +(*** A rather complex proof about iterates -- cf Andy Pitts ***) + +(** Two lemmas about natrec(n,x,%m.g), which is essentially (g^n)(x) **) + +goal LList.thy + "nat_rec(n, LCons(b, l), %m. lmap(f)) = \ +\ LCons(nat_rec(n, b, %m. f), nat_rec(n, l, %m. lmap(f)))"; +by (nat_ind_tac "n" 1); +by(ALLGOALS (asm_simp_tac (nat_ss addsimps [lmap_LCons]))); +val fun_power_lmap = result(); + +goal Nat.thy "nat_rec(n, g(x), %m. g) = nat_rec(Suc(n), x, %m. g)"; +by (nat_ind_tac "n" 1); +by(ALLGOALS (asm_simp_tac nat_ss)); +val fun_power_Suc = result(); + +val Pair_cong = read_instantiate_sg (sign_of Prod.thy) + [("f","Pair")] (standard(refl RS cong RS cong)); + +(*The bisimulation consists of {} + for all u and all n::nat.*) +val [prem] = goal LList.thy + "(!!x. h(x) = LCons(x, lmap(f,h(x)))) ==> h = iterates(f)"; +br ext 1; +by (res_inst_tac [("r", + "UN u. range(%n. )")] + llist_equalityI 1); +by (REPEAT (resolve_tac [UN1_I, range_eqI, Pair_cong, nat_rec_0 RS sym] 1)); +by (safe_tac set_cs); +by (stac iterates 1); +by (stac prem 1); +by (stac fun_power_lmap 1); +by (stac fun_power_lmap 1); +br llistD_Fun_LCons_I 1; +by (rtac (lmap_iterates RS subst) 1); +by (stac fun_power_Suc 1); +by (stac fun_power_Suc 1); +br UN1_I 1; +br rangeI 1; +val iterates_equality = result(); + + +(*** lappend -- its two arguments cause some complications! ***) + +goalw LList.thy [lappend_def] "lappend(LNil,LNil) = LNil"; +by (rtac (llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +val lappend_LNil_LNil = result(); + +goalw LList.thy [lappend_def] + "lappend(LNil,LCons(l,l')) = LCons(l, lappend(LNil,l'))"; +by (rtac (llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +(* 3.3(5.7) vs 1.3 !by (SIMP_TAC llistD_ss 1);*) +val lappend_LNil_LCons = result(); + +goalw LList.thy [lappend_def] + "lappend(LCons(l,l'), N) = LCons(l, lappend(l',N))"; +by (rtac (llist_corec RS trans) 1); +by(simp_tac llistD_ss 1); +(* 5(5.5) vs 1.3 !by (SIMP_TAC llistD_ss 1);*) +val lappend_LCons = result(); + +goal LList.thy "lappend(LNil,l) = l"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac (llistD_ss addsimps [lappend_LNil_LNil, + lappend_LNil_LCons]))); +val lappend_LNil = result(); + +goal LList.thy "lappend(l,LNil) = l"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (ALLGOALS (simp_tac(llistD_ss addsimps [lappend_LNil_LNil,lappend_LCons]))); +val lappend_LNil2 = result(); + +(*The infinite first argument blocks the second*) +goal LList.thy "lappend(iterates(f,x), N) = iterates(f,x)"; +by (res_inst_tac [("r", "range(%u.)")] + llist_equalityI 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (stac iterates 1); +by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1); +val lappend_iterates = result(); + +(** Two proofs that lmap distributes over lappend **) + +(*Long proof requiring case analysis on both both arguments*) +goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))"; +by (res_inst_tac + [("r", + "UN n. range(%l.)")] + llist_equalityI 1); +by (rtac UN1_I 1); +by (rtac rangeI 1); +by (safe_tac set_cs); +by (res_inst_tac [("l", "l")] llistE 1); +by (res_inst_tac [("l", "n")] llistE 1); +by (ALLGOALS (asm_simp_tac (llistD_ss addsimps + [lappend_LNil_LNil,lappend_LCons,lappend_LNil_LCons, + lmap_LNil,lmap_LCons]))); +by (REPEAT_SOME (ares_tac [llistD_Fun_LCons_I, UN1_I, rangeI])); +by (rtac range_eqI 1); +by (rtac (refl RS Pair_cong) 1); +by (stac lmap_LNil 1); +by (rtac refl 1); +val lmap_lappend_distrib = result(); + +(*Shorter proof of the theorem above using llist_equalityI2*) +goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))"; +by (res_inst_tac [("l","l")] llist_fun_equalityI 1); +by (simp_tac (llistD_ss addsimps [lappend_LNil, lmap_LNil])1); +by (simp_tac (llistD_ss addsimps [lappend_LCons, lmap_LCons]) 1); +val lmap_lappend_distrib = result(); + +(*Without llist_equalityI2, three case analyses might be needed*) +goal LList.thy "lappend(lappend(l1,l2) ,l3) = lappend(l1, lappend(l2,l3))"; +by (res_inst_tac [("l","l1")] llist_fun_equalityI 1); +by (simp_tac (llistD_ss addsimps [lappend_LNil])1); +by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1); +val lappend_assoc = result(); diff -r 000000000000 -r 7949f97df77a llist.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/llist.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,126 @@ +(* Title: HOL/llist.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Definition of type 'a llist by a greatest fixed point + +Shares NIL, CONS, List_Fun, List_case with list.thy + +Still needs filter and flatten functions -- hard because they need +bounds on the amount of lookahead required. + +Could try (but would it work for the gfp analogue of term?) + LListD_Fun_def "LListD_Fun(A) == (%Z.diag({Numb(0)}) <++> diag(A) <**> Z)" + +A nice but complex example would be [ML for the Working Programmer, page 176] + from(1) = enumerate (Lmap (Lmap(pack), makeqq(from(1),from(1)))) + +Previous definition of llistD_Fun was explicit: + llistD_Fun_def + "llistD_Fun(r) == \ +\ {} Un \ +\ (UN x. (%z.split(z, %l1 l2.))``r)" +*) + +LList = Gfp + List + +types llist 1 +arities llist :: (term)term +consts + LList :: "'a node set set => 'a node set set" + LListD_Fun :: + "[('a node set * 'a node set)set, ('a node set * 'a node set)set] => \ +\ ('a node set * 'a node set)set" + LListD :: + "('a node set * 'a node set)set => ('a node set * 'a node set)set" + llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set" + + Rep_LList :: "'a llist => 'a node set" + Abs_LList :: "'a node set => 'a llist" + LNil :: "'a llist" + LCons :: "['a, 'a llist] => 'a llist" + + llist_case :: "['a llist, 'b, ['a, 'a llist]=>'b] => 'b" + + LList_corec_fun :: "[nat, 'a=>unit+('b node set * 'a), 'a] => 'b node set" + LList_corec :: "['a, 'a => unit + ('b node set * 'a)] => 'b node set" + llist_corec :: "['a, 'a => unit + ('b * 'a)] => 'b llist" + + Lmap :: "('a node set => 'b node set) => ('a node set => 'b node set)" + lmap :: "('a=>'b) => ('a llist => 'b llist)" + + iterates :: "['a => 'a, 'a] => 'a llist" + + Lconst :: "'a node set => 'a node set" + Lappend :: "['a node set, 'a node set] => 'a node set" + lappend :: "['a llist, 'a llist] => 'a llist" + + +rules + LListD_Fun_def "LListD_Fun(r) == (%Z.diag({Numb(0)}) <++> r <**> Z)" + LList_def "LList(A) == gfp(List_Fun(A))" + LListD_def "LListD(r) == gfp(LListD_Fun(r))" + (*faking a type definition...*) + Rep_LList "Rep_LList(xs): LList(range(Leaf))" + Rep_LList_inverse "Abs_LList(Rep_LList(xs)) = xs" + Abs_LList_inverse "M: LList(range(Leaf)) ==> Rep_LList(Abs_LList(M)) = M" + (*defining the abstract constructors*) + LNil_def "LNil == Abs_LList(NIL)" + LCons_def "LCons(x,xs) == Abs_LList(CONS(Leaf(x), Rep_LList(xs)))" + + llist_case_def + "llist_case(l,c,d) == \ +\ List_case(Rep_LList(l), c, %x y. d(Inv(Leaf,x), Abs_LList(y)))" + + LList_corec_fun_def + "LList_corec_fun(k,f) == \ +\ nat_rec(k, %x. {}, \ +\ %j r x. case(f(x), %u.NIL, \ +\ %v. split(v, %z w. CONS(z, r(w)))))" + + LList_corec_def + "LList_corec(a,f) == UN k. LList_corec_fun(k,f,a)" + + llist_corec_def + "llist_corec(a,f) == \ +\ Abs_LList(LList_corec(a, %z.case(f(z), %x.Inl(x), \ +\ %y.split(y, %v w. Inr()))))" + + llistD_Fun_def + "llistD_Fun(r) == \ +\ prod_fun(Abs_LList,Abs_LList) `` \ +\ LListD_Fun(diag(range(Leaf)), \ +\ prod_fun(Rep_LList,Rep_LList) `` r)" + + Lconst_def "Lconst(M) == lfp(%N. CONS(M, N))" + + Lmap_def + "Lmap(f,M) == \ +\ LList_corec(M, %M. List_case(M, Inl(Unity), %x M'. Inr()))" + + lmap_def + "lmap(f,l) == \ +\ llist_corec(l, %l. llist_case(l, Inl(Unity), %y z. Inr()))" + + iterates_def "iterates(f,a) == llist_corec(a, %x. Inr())" + +(*Append generates its result by applying f, where + f() = Inl(Unity) + f() = Inr() + f() = Inr() +*) + + Lappend_def + "Lappend(M,N) == LList_corec(, %p. split(p, \ +\ %M N. List_case(M, List_case(N, Inl(Unity), \ +\ %N1 N2. Inr(>)), \ +\ %M1 M2. Inr(>))))" + + lappend_def + "lappend(l,n) == llist_corec(, %p. split(p, \ +\ %l n. llist_case(l, llist_case(n, Inl(Unity), \ +\ %n1 n2. Inr(>)), \ +\ %l1 l2. Inr(>))))" + + +end diff -r 000000000000 -r 7949f97df77a mono.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/mono.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,72 @@ +(* Title: HOL/mono + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +Monotonicity of various operations +*) + +val [prem] = goal Set.thy + "[| !!x. P(x) ==> Q(x) |] ==> Collect(P) <= Collect(Q)"; +by (fast_tac (set_cs addIs [prem]) 1); +val Collect_mono = result(); + +goal Set.thy "!!A B. A<=B ==> f``A <= f``B"; +by (fast_tac set_cs 1); +val image_mono = result(); + +goal Set.thy "!!A B. A<=B ==> Union(A) <= Union(B)"; +by (fast_tac set_cs 1); +val Union_mono = result(); + +goal Set.thy "!!A B. B<=A ==> Inter(A) <= Inter(B)"; +by (fast_tac set_cs 1); +val Inter_anti_mono = result(); + +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); +val UN_mono = result(); + +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); +val UN1_mono = result(); + +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); +val INT_anti_mono = result(); + +(*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); +val INT1_mono = result(); + +goal Set.thy "!!A B. [| A<=C; B<=D |] ==> A Un B <= C Un D"; +by (fast_tac set_cs 1); +val Un_mono = result(); + +goal Set.thy "!!A B. [| A<=C; B<=D |] ==> A Int B <= C Int D"; +by (fast_tac set_cs 1); +val Int_mono = result(); + +goal Set.thy "!!A::'a set. [| A<=C; D<=B |] ==> A-B <= C-D"; +by (fast_tac set_cs 1); +val Diff_mono = result(); + +goal Set.thy "!!A B. A<=B ==> Compl(B) <= Compl(A)"; +by (fast_tac set_cs 1); +val Compl_anti_mono = result(); + +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); +val Sigma_mono = result(); + diff -r 000000000000 -r 7949f97df77a nat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/nat.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,409 @@ +(* 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)); +val Nat_fun_mono = result(); + +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); +val Zero_RepI = result(); + +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); +val Suc_RepI = result(); + +(*** 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 (major RS (Nat_def RS def_induct)) 1); +by (rtac Nat_fun_mono 1); +by (fast_tac (set_cs addIs prems) 1); +val Nat_induct = result(); + +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)); +val nat_induct = result(); + +(*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)); +val diff_induct = result(); + +(*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); +val natE = result(); + +(*** 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); +val inj_Rep_Nat = result(); + +goal Nat.thy "inj_onto(Abs_Nat,Nat)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Nat_inverse 1); +val inj_onto_Abs_Nat = result(); + +(*** 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)); +val Suc_not_Zero = result(); + +val Zero_not_Suc = standard (Suc_not_Zero RS not_sym); + +val Suc_neq_Zero = standard (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); +val inj_Suc = result(); + +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]); +val Suc_Suc_eq = result(); + +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]))); +val n_not_Suc_n = result(); + +(*** nat_case -- the selection operator for nat ***) + +goalw Nat.thy [nat_case_def] "nat_case(0, a, f) = a"; +by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1); +val nat_case_0 = result(); + +goalw Nat.thy [nat_case_def] "nat_case(Suc(k), a, f) = f(k)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim Suc_inject, Suc_neq_Zero]) 1); +val nat_case_Suc = result(); + +(** Introduction rules for 'pred_nat' **) + +goalw Nat.thy [pred_nat_def] " : pred_nat"; +by (fast_tac set_cs 1); +val pred_natI = result(); + +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)); +val pred_natE = result(); + +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); +val wf_pred_nat = result(); + + +(*** nat_rec -- by wf recursion on pred_nat ***) + +val nat_rec_unfold = standard (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 (rtac nat_case_0 1); +val nat_rec_0 = result(); + +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 (rtac (nat_case_Suc RS trans) 1); +by(simp_tac (HOL_ss addsimps [pred_natI,cut_apply]) 1); +val nat_rec_Suc = result(); + +(*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); +val def_nat_rec_0 = result(); + +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); +val def_nat_rec_Suc = result(); + +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 i ~ m R *) +val less_anti_sym = standard (less_not_sym RS mp RS notE); + + +goalw Nat.thy [less_def] "~ n R *) +val less_anti_refl = standard (less_not_refl RS notE); + + +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); +val lessE = result(); + +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); +val not_less0 = result(); + +(* n<0 ==> R *) +val less_zeroE = standard (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); +val less_SucE = result(); + +goal Nat.thy "(m < Suc(n)) = (m < n | m = n)"; +by (fast_tac (HOL_cs addSIs [lessI] + addEs [less_trans, less_SucE]) 1); +val less_Suc_eq = result(); + + +(** 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); +val Suc_lessE = result(); + +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); +val Suc_mono = result(); + +goal Nat.thy "(Suc(m) < Suc(n)) = (m P"; +by (rtac (major RS Suc_lessD RS less_anti_refl) 1); +val not_Suc_n_less_n = result(); + +(*"Less than" is a linear ordering*) +goal Nat.thy "m P(m) |] ==> P(n) |] ==> P(n)"; +by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1); +by (eresolve_tac prems 1); +val less_induct = result(); + + +(*** Properties of <= ***) + +goalw Nat.thy [le_def] "0 <= n"; +by (rtac not_less0 1); +val le0 = result(); + +val nat_simps = [not_less0, less_not_refl, zero_less_Suc, lessI, + Suc_less_eq, less_Suc_eq, le0, + Suc_not_Zero, Zero_not_Suc, Suc_Suc_eq, + nat_case_0, nat_case_Suc, nat_rec_0, nat_rec_Suc]; + +val nat_ss = pair_ss addsimps nat_simps; + +val prems = goalw Nat.thy [le_def] "~(n m<=n::nat"; +by (resolve_tac prems 1); +val leI = result(); + +val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n n Suc(m) <= n"; +by(simp_tac (HOL_ss addsimps [less_Suc_eq]) 1); +by (fast_tac (HOL_cs addEs [less_anti_refl,less_anti_sym]) 1); +val lessD = result(); + +goalw Nat.thy [le_def] "!!m. m < n ==> m <= n::nat"; +by (fast_tac (HOL_cs addEs [less_anti_sym]) 1); +val less_imp_le = result(); + +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_anti_sym]) 1); +val le_imp_less_or_eq = result(); + +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_anti_sym]) 1); +by (flexflex_tac); +val less_or_eq_imp_le = result(); + +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)); +val le_eq_less_or_eq = result(); + +goal Nat.thy "n <= n::nat"; +by(simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1); +val le_refl = result(); + +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); +val le_less_trans = result(); + +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])]); +val le_trans = result(); + +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_anti_sym])]); +val le_anti_sym = result(); diff -r 000000000000 -r 7949f97df77a nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/nat.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,51 @@ +(* Title: HOL/nat + 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 + +types ind,nat 0 +arities ind,nat :: term + nat :: ord + +consts + Zero_Rep :: "ind" + Suc_Rep :: "ind => ind" + Nat :: "ind set" + Rep_Nat :: "nat => ind" + Abs_Nat :: "ind => nat" + Suc :: "nat => nat" + nat_case :: "[nat, 'a, nat=>'a] =>'a" + pred_nat :: "(nat*nat) set" + nat_rec :: "[nat, 'a, [nat, 'a]=>'a] => 'a" + "0" :: "nat" ("0") + +rules + (*the axiom of infinity in 2 parts*) + inj_Suc_Rep "inj(Suc_Rep)" + Suc_Rep_not_Zero_Rep "~(Suc_Rep(x) = Zero_Rep)" + Nat_def "Nat == lfp(%X. {Zero_Rep} Un (Suc_Rep``X))" + (*faking a type definition...*) + Rep_Nat "Rep_Nat(n): Nat" + Rep_Nat_inverse "Abs_Nat(Rep_Nat(n)) = n" + Abs_Nat_inverse "i: Nat ==> Rep_Nat(Abs_Nat(i)) = i" + (*defining the abstract constants*) + 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 == (%n a f. @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)); +val monoI = result(); + +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); +val monoD = result(); + diff -r 000000000000 -r 7949f97df77a ord.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ord.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,15 @@ +Ord = HOL + +classes + ord < term +consts + "<", "<=" :: "['a::ord, 'a] => bool" (infixl 50) + mono :: "['a::ord => 'b::ord] => bool" (*monotonicity*) + min,max :: "['a::ord,'a] => 'a" + +rules + +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 000000000000 -r 7949f97df77a prod.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/prod.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,204 @@ +(* 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]); +val ProdI = result(); + +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]); +val Pair_Rep_inject = result(); + +goal Prod.thy "inj_onto(Abs_Prod,Prod)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Prod_inverse 1); +val inj_onto_Abs_Prod = result(); + +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)); +val Pair_inject = result(); + +goal Prod.thy "( = ) = (a=a' & b=b')"; +by (fast_tac (set_cs addIs [Pair_inject]) 1); +val Pair_eq = result(); + +goalw Prod.thy [fst_def] "fst() = a"; +by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); +val fst_conv = result(); + +goalw Prod.thy [snd_def] "snd() = b"; +by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1); +val snd_conv = result(); + +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]); +val PairE_lemma = result(); + +val [prem] = goal Prod.thy "[| !!x y. p = ==> Q |] ==> Q"; +by (rtac (PairE_lemma RS exE) 1); +by (REPEAT (eresolve_tac [prem,exE] 1)); +val PairE = result(); + +goalw Prod.thy [split_def] "split(, c) = c(a,b)"; +by (sstac [fst_conv, snd_conv] 1); +by (rtac refl 1); +val split = result(); + +(*FIXME: split's congruence rule should only simplifies the pair*) +val pair_ss = set_ss addsimps [fst_conv, snd_conv, split]; + +goal Prod.thy "p = "; +by (res_inst_tac [("p","p")] PairE 1); +by(asm_simp_tac pair_ss 1); +val surjective_pairing = result(); + +goal Prod.thy "p = split(p, %x y.)"; +by (res_inst_tac [("p","p")] PairE 1); +by(asm_simp_tac pair_ss 1); +val surjective_pairing2 = result(); + +(** split used as a logical connective, with result type bool **) + +val prems = goal Prod.thy "c(a,b) ==> split(, c)"; +by (stac split 1); +by (resolve_tac prems 1); +val splitI = result(); + +val prems = goalw Prod.thy [split_def] + "[| split(p,c); !!x y. [| p = ; c(x,y) |] ==> Q |] ==> Q"; +by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1)); +val splitE = result(); + +goal Prod.thy "R(split(p,c)) = (! 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); +val expand_split = result(); + +(*** prod_fun -- action of the product functor upon functions ***) + +goalw Prod.thy [prod_fun_def] "prod_fun(f,g,) = "; +by (rtac split 1); +val prod_fun = result(); + +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 (pair_ss addsimps [prod_fun,o_def]) 1); +val prod_fun_compose = result(); + +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 (pair_ss addsimps [prod_fun]) 1); +val prod_fun_ident = result(); + +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); +val prod_fun_imageI = result(); + +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); +val prod_fun_imageE = result(); + +(*** Disjoint union of a family of sets - Sigma ***) + +val SigmaI = prove_goalw 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*) +val SigmaE = prove_goalw 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 **) +val SigmaD1 = prove_goal Prod.thy " : Sigma(A,B) ==> a : A" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +val SigmaD2 = prove_goal Prod.thy " : Sigma(A,B) ==> b : B(a)" + (fn [major]=> + [ (rtac (major RS SigmaE) 1), + (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]); + +val SigmaE2 = prove_goal 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); +val fst_imageI = result(); + +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); +val fst_imageE = result(); + +(*** 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); +val snd_imageI = result(); + +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); +val snd_imageE = result(); + +(** 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); +val unit_eq = result(); diff -r 000000000000 -r 7949f97df77a prod.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/prod.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,66 @@ +(* Title: HOL/prod + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +Ordered Pairs and the Cartesian product type +The unit type + +The type definition admits the following unused axiom: + Abs_Unit_inverse "f: Unit ==> Rep_Unit(Abs_Unit(f)) = f" +*) + +Prod = Set + +types + "*" 2 (infixr 20) + unit 0 +arities + "*" :: (term,term)term + unit :: term +consts + Pair_Rep :: "['a,'b] => ['a,'b] => bool" + Prod :: "('a => 'b => bool)set" + Rep_Prod :: "'a * 'b => ('a => 'b => bool)" + Abs_Prod :: "('a => 'b => bool) => 'a * 'b" + fst :: "'a * 'b => 'a" + snd :: "'a * 'b => 'b" + split :: "['a * 'b, ['a,'b]=>'c] => 'c" + prod_fun :: "['a=>'b, 'c=>'d, 'a*'c] => 'b*'d" + Pair :: "['a,'b] => 'a * 'b" + "@Tuple" :: "args => 'a*'b" ("(1<_>)") + Sigma :: "['a set, 'a => 'b set] => ('a*'b)set" + + Unit :: "bool set" + Rep_Unit :: "unit => bool" + Abs_Unit :: "bool => unit" + Unity :: "unit" ("<>") + +translations + + "" == ">" + "" == "Pair(x,y)" + "" => "x" + +rules + + Pair_Rep_def "Pair_Rep == (%a b. %x y. x=a & y=b)" + Prod_def "Prod == {f. ? a b. f = Pair_Rep(a,b)}" + (*faking a type definition...*) + Rep_Prod "Rep_Prod(p): Prod" + Rep_Prod_inverse "Abs_Prod(Rep_Prod(p)) = p" + Abs_Prod_inverse "f: Prod ==> Rep_Prod(Abs_Prod(f)) = f" + (*defining the abstract constants*) + 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(p,c) == c(fst(p),snd(p))" + prod_fun_def "prod_fun(f,g) == (%z.split(z, %x y.))" + Sigma_def "Sigma(A,B) == UN x:A. UN y:B(x). {}" + + Unit_def "Unit == {p. p=True}" + (*faking a type definition...*) + Rep_Unit "Rep_Unit(u): Unit" + Rep_Unit_inverse "Abs_Unit(Rep_Unit(u)) = u" + (*defining the abstract constants*) + Unity_def "Unity == Abs_Unit(True)" +end diff -r 000000000000 -r 7949f97df77a set.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/set.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,436 @@ +(* 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); +val CollectI = result(); + +val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)"; +by (resolve_tac (prems RL [mem_Collect_eq RS subst]) 1); +val CollectD = result(); + +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); +val set_ext = result(); + +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); +val Collect_cong = result(); + +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)); +val ballI = result(); + +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); +val bspec = result(); + +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)); +val ballE = result(); + +(*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)); +val bexI = result(); + +val bexCI = prove_goal 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)); +val bexE = result(); + +(*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)); +val ball_rew = result(); + +(** 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)); +val ball_cong = result(); + +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)); +val bex_cong = result(); + +(*** Subsets ***) + +val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B"; +by (REPEAT (ares_tac (prems @ [ballI]) 1)); +val subsetI = result(); + +(*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); +val subsetD = result(); + +(*The same, with reversed premises for use with etac -- cf rev_mp*) +val rev_subsetD = prove_goal 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)); +val subsetCE = result(); + +(*Takes assumptions A<=B; c:A and creates the assumption c:B *) +fun set_mp_tac i = etac subsetCE i THEN mp_tac i; + +val subset_refl = prove_goal 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)); +val subset_trans = result(); + + +(*** 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)); +val subset_antisym = result(); +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); +val equalityD1 = result(); + +val prems = goal Set.thy "A = B ==> B<=A::'a set"; +by (resolve_tac (prems RL [subst]) 1); +by (rtac subset_refl 1); +val equalityD2 = result(); + +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)); +val equalityE = result(); + +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)); +val equalityCE = result(); + +(*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)); +val setup_induction = result(); + + +(*** Set complement -- Compl ***) + +val prems = goalw Set.thy [Compl_def] + "[| c:A ==> False |] ==> c : Compl(A)"; +by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); +val ComplI = result(); + +(*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); +val ComplD = result(); + +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)); +val UnI1 = result(); + +val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B"; +by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1)); +val UnI2 = result(); + +(*Classical introduction rule: no commitment to A vs B*) +val UnCI = prove_goal 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)); +val UnE = result(); + + +(*** 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)); +val IntI = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A"; +by (rtac (major RS CollectD RS conjunct1) 1); +val IntD1 = result(); + +val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B"; +by (rtac (major RS CollectD RS conjunct2) 1); +val IntD2 = result(); + +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); +val IntE = result(); + + +(*** Set difference ***) + +val DiffI = prove_goalw Set.thy [set_diff_def] + "[| c : A; ~ c : B |] ==> c : A - B" + (fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)) ]); + +val DiffD1 = prove_goalw Set.thy [set_diff_def] + "c : A - B ==> c : A" + (fn [major]=> [ (rtac (major RS CollectD RS conjunct1) 1) ]); + +val DiffD2 = prove_goalw 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]); + +val DiffE = prove_goal 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)) ]); + +val Diff_iff = prove_goal Set.thy "(c : A-B) = (c:A & ~ c:B)" + (fn _ => [ (fast_tac (HOL_cs addSIs [DiffI] addSEs [DiffE]) 1) ]); + + +(*** The empty set -- {} ***) + +val emptyE = prove_goalw Set.thy [empty_def] "a:{} ==> P" + (fn [prem] => [rtac (prem RS CollectD RS FalseE) 1]); + +val empty_subsetI = prove_goal Set.thy "{} <= A" + (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]); + +val equals0I = prove_goal 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)) ]); + +val equals0D = prove_goal Set.thy "[| A={}; a:A |] ==> P" + (fn [major,minor]=> + [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]); + + +(*** Augmenting a set -- insert ***) + +val insertI1 = prove_goalw Set.thy [insert_def] "a : insert(a,B)" + (fn _ => [rtac (CollectI RS UnI1) 1, rtac refl 1]); + +val insertI2 = prove_goalw Set.thy [insert_def] "a : B ==> a : insert(b,B)" + (fn [prem]=> [ (rtac (prem RS UnI2) 1) ]); + +val insertE = prove_goalw 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)) ]); + +val insert_iff = prove_goal Set.thy "a : insert(b,A) = (a=b | a:A)" + (fn _ => [fast_tac (HOL_cs addIs [insertI1,insertI2] addSEs [insertE]) 1]); + +(*Classical introduction rule*) +val insertCI = prove_goal 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 ***) + +val singletonI = prove_goal Set.thy "a : {a}" + (fn _=> [ (rtac insertI1 1) ]); + +val singletonE = prove_goal 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); +val singletonD = result(); + +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); +val singleton_inject = result(); + +(*** 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)); +val UN_I = result(); + +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)); +val UN_E = result(); + +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)); +val UN_cong = result(); + + +(*** 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)); +val INT_I = result(); + +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); +val INT_D = result(); + +(*"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)); +val INT_E = result(); + +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)); +val INT_cong = result(); + + +(*** 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)); +val UN1_I = result(); + +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)); +val UN1_E = result(); + + +(*** 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)); +val INT1_I = result(); + +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); +val INT1_D = result(); + +(*** 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)); +val UnionI = result(); + +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)); +val UnionE = result(); + +(*** Inter ***) + +val prems = goalw Set.thy [Inter_def] + "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; +by (REPEAT (ares_tac ([INT_I] @ prems) 1)); +val InterI = result(); + +(*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); +val InterD = result(); + +(*"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)); +val InterE = result(); diff -r 000000000000 -r 7949f97df77a set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/set.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,111 @@ +(* Title: HOL/set.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 1993 University of Cambridge +*) + +Set = Ord + + +types + set 1 + +arities + set :: (term) term + set :: (term) ord + set :: (term) minus + + +consts + + (* Constants *) + + 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*) + range :: "('a => 'b) => 'b set" (*of function*) + Ball, Bex :: "['a set, 'a => bool] => bool" (*bounded quantifiers*) + inj, surj :: "('a => 'b) => bool" (*injective/surjective*) + inj_onto :: "['a => 'b, 'a set] => bool" + "``" :: "['a => 'b, 'a set] => ('b set)" (infixl 90) + ":" :: "['a, 'a set] => bool" (infixl 50) (*membership*) + + (* Finite Sets *) + + insert :: "['a, 'a set] => 'a set" + "{}" :: "'a set" ("{}") + "@Finset" :: "args => 'a set" ("{(_)}") + + + (** Binding Constants **) + + "@Coll" :: "[idt, bool] => 'a set" ("(1{_./ _})") (*collection*) + + (* 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. 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)" + + "{x, xs}" == "insert(x, {xs})" + "{x}" == "insert(x, {})" + + +rules + + (* Isomorphisms between Predicates and Sets *) + + mem_Collect_eq "(a : {x.P(x)}) = P(a)" + Collect_mem_eq "{x.x:A} = A" + + (* Definitions *) + + 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)" + 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 + +val print_ast_translation = + map HOL.mk_alt_ast_tr' [("@Ball", "*Ball"), ("@Bex", "*Bex")]; + diff -r 000000000000 -r 7949f97df77a sexp.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/sexp.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,166 @@ +(* Title: HOL/sexp + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +For sexp.thy. S-expressions. +*) + +open Sexp; + +(** the sexp functional **) + +goal Sexp.thy "mono(%Z. range(Leaf) Un range(Numb) Un Z<*>Z)"; +by (REPEAT (ares_tac [monoI, subset_refl, Un_mono, uprod_mono] 1)); +val Sexp_fun_mono = result(); + +val Sexp_unfold = Sexp_fun_mono RS (Sexp_def RS def_lfp_Tarski); + +(** Induction **) + +val major::prems = goal Sexp.thy + "[| ii: Sexp; !!a. P(Leaf(a)); !!k. P(Numb(k)); \ +\ !!i j. [| i: Sexp; j: Sexp; P(i); P(j) |] ==> P(i.j) \ +\ |] ==> P(ii)"; +by (rtac (major RS (Sexp_def RS def_induct)) 1); +by (rtac Sexp_fun_mono 1); +by (fast_tac (set_cs addIs prems addSEs [uprodE]) 1); +val Sexp_induct = result(); + +(** Sexp_case **) + +goalw Sexp.thy [Sexp_case_def] "Sexp_case(Leaf(a),c,d,e) = c(a)"; +by (fast_tac (HOL_cs addIs [select_equality] + addSDs [Leaf_inject] + addSEs [Leaf_neq_Scons, Leaf_neq_Numb]) 1); +val Sexp_case_Leaf = result(); + +goalw Sexp.thy [Sexp_case_def] "Sexp_case(Numb(k),c,d,e) = d(k)"; +by (fast_tac (HOL_cs addIs [select_equality] + addSDs [Numb_inject] + addSEs [Numb_neq_Scons, Numb_neq_Leaf]) 1); +val Sexp_case_Numb = result(); + +goalw Sexp.thy [Sexp_case_def] "Sexp_case(M.N, c, d, e) = e(M,N)"; +by (fast_tac (HOL_cs addIs [select_equality] + addSDs [Scons_inject] + addSEs [Scons_neq_Leaf, Scons_neq_Numb]) 1); +val Sexp_case_Scons = result(); + + +(** Introduction rules for Sexp constructors **) + +val SexpI = Sexp_unfold RS equalityD2 RS subsetD; + +goal Sexp.thy "Leaf(a): Sexp"; +by (fast_tac (set_cs addIs [SexpI]) 1); +val Sexp_LeafI = result(); + +goal Sexp.thy "Numb(a): Sexp"; +by (fast_tac (set_cs addIs [SexpI]) 1); +val Sexp_NumbI = result(); + +val prems = goal Sexp.thy + "[| M: Sexp; N: Sexp |] ==> M.N : Sexp"; +by (fast_tac (set_cs addIs ([uprodI,SexpI]@prems)) 1); +val Sexp_SconsI = result(); + +val [prem] = goalw Sexp.thy [In0_def] + "M: Sexp ==> In0(M) : Sexp"; +by (rtac (prem RS (Sexp_NumbI RS Sexp_SconsI)) 1); +val Sexp_In0I = result(); + +val [prem] = goalw Sexp.thy [In1_def] + "M: Sexp ==> In1(M) : Sexp"; +by (rtac (prem RS (Sexp_NumbI RS Sexp_SconsI)) 1); +val Sexp_In1I = result(); + +goal Sexp.thy "range(Leaf) <= Sexp"; +by (fast_tac (set_cs addIs [SexpI]) 1); +val range_Leaf_subset_Sexp = result(); + +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]))); +val Scons_D = result(); + +(** Introduction rules for 'pred_Sexp' **) + +val sexp_cs = set_cs addIs [SigmaI, uprodI, SexpI]; + +goalw Sexp.thy [pred_Sexp_def] "pred_Sexp <= Sigma(Sexp, %u.Sexp)"; +by (fast_tac sexp_cs 1); +val pred_Sexp_subset_Sigma = result(); + +(* : 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); +val pred_SexpI1 = result(); + +val prems = goalw Sexp.thy [pred_Sexp_def] + "[| M: Sexp; N: Sexp |] ==> : pred_Sexp"; +by (fast_tac (set_cs addIs prems) 1); +val pred_SexpI2 = result(); + +(*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_LeafI, Sexp_NumbI, Sexp_SconsI, + 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)); +val pred_SexpE = result(); + +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); +val wf_pred_Sexp = result(); + +(*** 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); +val Sexp_rec_Leaf = result(); + +goal Sexp.thy "Sexp_rec(Numb(k), c, d, h) = d(k)"; +by (stac Sexp_rec_unfold 1); +by (rtac Sexp_case_Numb 1); +val Sexp_rec_Numb = result(); + +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); +val Sexp_rec_Scons = result(); diff -r 000000000000 -r 7949f97df77a sexp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/sexp.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,36 @@ +(* 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 node set set" + + Sexp_case :: "['a node set, 'a=>'b, nat=>'b, \ +\ ['a node set,'a node set]=>'b] => 'b" + + Sexp_rec :: "['a node set, 'a=>'b, nat=>'b, \ +\ ['a node set,'a node set,'b,'b]=>'b] => 'b" + + pred_Sexp :: "('a node set * 'a node set)set" + +rules + Sexp_def "Sexp == lfp(%Z. range(Leaf) Un range(Numb) Un Z<*>Z)" + + Sexp_case_def + "Sexp_case(M,c,d,e) == @ 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(M, c, d, %N1 N2. e(N1, N2, g(N1), g(N2))))" +end + diff -r 000000000000 -r 7949f97df77a simpdata.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/simpdata.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,100 @@ +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 r = + case concl_of r of + Const("Trueprop",_) $ p => + (case p of + Const("op -->",_)$_$_ => atomize(r RS mp) + | Const("op &",_)$_$_ => atomize(r RS conjunct1) @ + atomize(r RS conjunct2) + | Const("All",_)$_ => atomize(r RS spec) + | Const("True",_) => [] + | Const("False",_) => [] + | _ => [r]) + | _ => [r]; + +fun mk_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; + +fun mk_rews thm = map mk_eq (atomize(gen_all thm)); + +val imp_cong_lemma = 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 imp_meta_cong = imp_cong_lemma RS eq_reflection; + +val o_apply = prove_goal HOL.thy "(f o g)(x) = f(g(x))" + (fn _ => [ (stac o_def 1), (rtac refl 1) ]); + +val simp_thms = map prover + [ "(x=x) = True", + "(~True) = False", "(~False) = True", "(~ ~ P) = P", + "(True=P) = P", "(P=True) = P", + "(True --> P) = P", "(False --> P) = True", + "(P --> True) = True", "(P --> P) = True", + "(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", + "(P|Q --> R) = ((P-->R)&(Q-->R))" ]; + +val meta_obj_reflection = prove_goal HOL.thy "x==y ==> x=y" + (fn [prem] => [rewtac prem, rtac refl 1]); + +in + + +val if_True = prove_goal HOL.thy "if(True,x,y) = x" + (fn _=>[stac if_def 1, fast_tac (HOL_cs addIs [select_equality]) 1]); + +val if_False = prove_goal HOL.thy "if(False,x,y) = y" + (fn _=>[stac if_def 1, 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_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]) RS eq_reflection; + + +val HOL_ss = empty_ss + setmksimps mk_rews + setsolver (fn prems => resolve_tac (TrueI::refl::prems)) + setsubgoaler asm_simp_tac + addsimps ([if_True, if_False, o_apply] @ simp_thms) + addcongs [imp_meta_cong]; + +fun split_tac splits = + mk_case_split_tac (meta_obj_reflection RS iffD2) (map mk_eq splits); + +end; diff -r 000000000000 -r 7949f97df77a subset.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/subset.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,142 @@ +(* 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 ***) + +val subset_insertI = prove_goal Set.thy "B <= insert(a,B)" + (fn _=> [ (rtac subsetI 1), (etac insertI2 1) ]); + +(*Useful for rewriting!*) +goalw Set.thy [subset_def,Ball_def] "insert(a,B)<=C = (a:C & B<=C)"; +(*by(SIMP_TAC (HOL_ss addsimps [insert_iff]) 1);*) +by(simp_tac (HOL_ss addsimps [insert_iff]) 1); +by(fast_tac HOL_cs 1); +val insert_subset_iff = result(); + +(*** 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)); +val Union_upper = result(); + +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)); +val Union_least = result(); + +(** 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)); +val UN_upper = result(); + +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)); +val UN_least = result(); + +goal Set.thy "B(a) <= (UN x. B(x))"; +by (REPEAT (ares_tac [UN1_I RS subsetI] 1)); +val UN1_upper = result(); + +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)); +val UN1_least = result(); + + +(*** 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)); +val Inter_lower = result(); + +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)); +val Inter_greatest = result(); + +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)); +val INT_lower = result(); + +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)); +val INT_greatest = result(); + +goal Set.thy "(INT x. B(x)) <= B(a)"; +br subsetI 1; +by (REPEAT (resolve_tac prems 1 ORELSE etac INT1_D 1)); +val INT1_lower = result(); + +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)); +val INT1_greatest = result(); + +(*** Finite Union -- the least upper bound of 2 sets ***) + +goal Set.thy "A <= A Un B"; +by (REPEAT (ares_tac [subsetI,UnI1] 1)); +val Un_upper1 = result(); + +goal Set.thy "B <= A Un B"; +by (REPEAT (ares_tac [subsetI,UnI2] 1)); +val Un_upper2 = result(); + +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)); +val Un_least = result(); + +(*** 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)); +val Int_lower1 = result(); + +goal Set.thy "A Int B <= B"; +by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)); +val Int_lower2 = result(); + +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)); +val Int_greatest = result(); + +(*** Set difference ***) + +val Diff_subset = prove_goal 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); +val mono_Un = result(); + +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); +val mono_Int = result(); diff -r 000000000000 -r 7949f97df77a sum.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/sum.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,130 @@ +(* Title: HOL/sum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1991 University of Cambridge + +For sum.ML. 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]); +val Inl_RepI = result(); + +goalw Sum.thy [Sum_def] "Inr_Rep(b) : Sum"; +by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]); +val Inr_RepI = result(); + +goal Sum.thy "inj_onto(Abs_Sum,Sum)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Sum_inverse 1); +val inj_onto_Abs_Sum = result(); + +(** 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]) ]); +val Inl_Rep_not_Inr_Rep = result(); + +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); +val Inl_not_Inr = result(); + +val Inl_neq_Inr = standard (Inl_not_Inr RS notE); +val Inr_neq_Inl = sym RS Inl_neq_Inr; + +(** 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); +val Inl_Rep_inject = result(); + +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); +val Inr_Rep_inject = result(); + +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); +val inj_Inl = result(); +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); +val inj_Inr = result(); +val Inr_inject = inj_Inr RS injD; + +goal Sum.thy "(Inl(x)=Inl(y)) = (x=y)"; +br iffI 1; +be (rewrite_rule [inj_def] Inl_inject) 1; +be ssubst 1; +br refl 1; +val Inl_inj = result(); + +goal Sum.thy "(Inr(x)=Inr(y)) = (x=y)"; +br iffI 1; +be (rewrite_rule [inj_def] Inr_inject) 1; +be ssubst 1; +br refl 1; +val Inr_inj = result(); + +(** case -- the selection operator for sums **) + +goalw Sum.thy [case_def] "case(Inl(x), f, g) = f(x)"; +by (fast_tac (set_cs addIs [select_equality] + addSEs [make_elim Inl_inject, Inl_neq_Inr]) 1); +val case_Inl = result(); + +goalw Sum.thy [case_def] "case(Inr(x), f, g) = g(x)"; +by (fast_tac (set_cs addIs [select_equality] + addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1); +val case_Inr = result(); + +(** 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)])); +val sumE = result(); + +goal Sum.thy "case(s, %x::'a. f(Inl(x)), %y::'b. f(Inr(y))) = f(s)"; +by (EVERY1 [res_inst_tac [("s","s")] sumE, + etac ssubst, rtac case_Inl, + etac ssubst, rtac case_Inr]); +val surjective_sum = result(); + +goal Sum.thy "R(case(s,f,g)) = \ +\ ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))"; +by (rtac sumE 1); +by (etac ssubst 1); +by (stac case_Inl 1); +by (fast_tac (set_cs addSEs [make_elim Inl_inject, Inl_neq_Inr]) 1); +by (etac ssubst 1); +by (stac case_Inr 1); +by (fast_tac (set_cs addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1); +val expand_case = result(); + +(*FIXME: case's congruence rule only should simplifies the first argument*) +val sum_ss = pair_ss addsimps [case_Inl, case_Inr]; diff -r 000000000000 -r 7949f97df77a sum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/sum.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,34 @@ +(* Title: HOL/sum + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1992 University of Cambridge + +The disjoint sum of two types +*) + +Sum = Prod + +types "+" 2 (infixl 10) +arities "+" :: (term,term)term +consts + Inl_Rep :: "['a,'a,'b,bool] => bool" + Inr_Rep :: "['b,'a,'b,bool] => bool" + Sum :: "(['a,'b,bool] => bool)set" + Rep_Sum :: "'a + 'b => (['a,'b,bool] => bool)" + Abs_Sum :: "(['a,'b,bool] => bool) => 'a+'b" + Inl :: "'a => 'a+'b" + Inr :: "'b => 'a+'b" + case :: "['a+'b, 'a=>'c,'b=>'c] =>'c" +rules + Inl_Rep_def "Inl_Rep == (%a. %x y p. x=a & p)" + Inr_Rep_def "Inr_Rep == (%b. %x y p. y=b & ~p)" + Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}" + (*faking a type definition...*) + Rep_Sum "Rep_Sum(s): Sum" + Rep_Sum_inverse "Abs_Sum(Rep_Sum(s)) = s" + Abs_Sum_inverse "f: Sum ==> Rep_Sum(Abs_Sum(f)) = f" + (*defining the abstract constants*) + Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))" + Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))" + case_def "case == (%p f g. @z. (!x. p=Inl(x) --> z=f(x))\ +\ & (!y. p=Inr(y) --> z=g(y)))" +end diff -r 000000000000 -r 7949f97df77a test.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/test.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,6 @@ +let + fun tac ss = resolve_tac(prems_of_ss ss) ORELSE' asm_simp_tac ss; + val ss = set_prove_tac(HOL_ss addsimps [Suc_lessD],tac) +in prove_goal Nat.thy "!!x. Suc(Suc(Suc(x))) x [asm_simp_tac ss 1]) +end; diff -r 000000000000 -r 7949f97df77a trancl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/trancl.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,240 @@ +(* 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)); +val transI = result(); + +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); +val transD = result(); + +(** Identity relation **) + +goalw Trancl.thy [id_def] " : id"; +by (rtac CollectI 1); +by (rtac exI 1); +by (rtac refl 1); +val idI = result(); + +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); +val idE = result(); + +(** Composition of two relations **) + +val prems = goalw Trancl.thy [comp_def] + "[| :s; :r |] ==> : r O s"; +by (fast_tac (set_cs addIs prems) 1); +val compI = result(); + +(*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)); +val compE = result(); + +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)); +val compEpair = result(); + +val comp_cs = set_cs addIs [compI,idI] + addSEs [compE,idE,Pair_inject]; + +val prems = goal Trancl.thy + "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +by (cut_facts_tac prems 1); +by (fast_tac comp_cs 1); +val comp_mono = result(); + +val prems = goal Trancl.thy + "[| s <= Sigma(A,%x.B); r <= Sigma(B,%x.C) |] ==> \ +\ (r O s) <= Sigma(A,%x.C)"; +by (cut_facts_tac prems 1); +by (fast_tac (comp_cs addIs [SigmaI] addSEs [SigmaE2]) 1); +val comp_subset_Sigma = result(); + + +(** 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)); +val rtrancl_fun_mono = result(); + +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); +val rtrancl_refl = result(); + +(*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); +val rtrancl_into_rtrancl = result(); + +(*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); +val r_into_rtrancl = result(); + +(*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)); +val rtrancl_mono = result(); + +(** standard induction rule **) + +val major::prems = goal Trancl.thy + "[| : r^*; \ +\ !!x. P(); \ +\ !!x y z.[| P(); : r^*; : r |] ==> P() |] \ +\ ==> P()"; +by (rtac (major RS (rtrancl_def RS def_induct)) 1); +by (rtac rtrancl_fun_mono 1); +by (fast_tac (comp_cs addIs prems) 1); +val rtrancl_full_induct = result(); + +(*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); +val rtrancl_induct = result(); + +(*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)); +val trans_rtrancl = result(); + +(*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)); +val rtranclE = result(); + + +(**** 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)); +val trancl_into_rtrancl = result(); + +(*r^+ contains r*) +val [prem] = goalw Trancl.thy [trancl_def] + "[| : r |] ==> : r^+"; +by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1)); +val r_into_trancl = result(); + +(*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)); +val rtrancl_into_trancl1 = result(); + +(*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)); +val rtrancl_into_trancl2 = result(); + +(*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); +val tranclE = result(); + +(*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)); +val trans_trancl = result(); + +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); +val trancl_into_trancl2 = result(); + + +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); +val trancl_subset_Sigma_lemma = result(); + +val prems = goalw Trancl.thy [trancl_def] + "r <= Sigma(A,%x.A) ==> trancl(r) <= Sigma(A,%x.A)"; +by (cut_facts_tac prems 1); +by (fast_tac (comp_cs addIs [SigmaI] + addSDs [trancl_subset_Sigma_lemma] + addSEs [SigmaE2]) 1); +val trancl_subset_Sigma = result(); + diff -r 000000000000 -r 7949f97df77a trancl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/trancl.thy Thu Sep 16 12:21:07 1993 +0200 @@ -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 + +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) +rules +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 000000000000 -r 7949f97df77a univ.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/univ.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,635 @@ +(* 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); +val Least_equality = result(); + +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); +val LeastI = result(); + +(*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); +val Least_le = result(); + +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); +val not_less_Least = result(); + + +(** apfst -- can be used in similar type definitions **) + +goalw Univ.thy [apfst_def] "apfst(f,) = "; +by (rtac split 1); +val apfst = result(); + +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); +val apfstE = result(); + +(** 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); +val Push_inject1 = result(); + +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); +val Push_inject2 = result(); + +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); +val Push_inject = result(); + +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); +val Push_neq_K0 = result(); + +(*** Isomorphisms ***) + +goal Univ.thy "inj(Rep_Node)"; +by (rtac inj_inverseI 1); (*cannot combine by RS: multiple unifiers*) +by (rtac Rep_Node_inverse 1); +val inj_Rep_Node = result(); + +goal Univ.thy "inj_onto(Abs_Node,Node)"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_Node_inverse 1); +val inj_onto_Abs_Node = result(); + +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); +val Node_K0_I = result(); + +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); +val Node_Push_I = result(); + + +(*** 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)); +val Scons_not_Atom = result(); +val Atom_not_Scons = standard (Scons_not_Atom RS not_sym); + +val Scons_neq_Atom = standard (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)); +val inj_Atom = result(); +val Atom_inject = inj_Atom RS injD; + +goalw Univ.thy [Leaf_def] "inj(Leaf)"; +by (stac o_def 1); +by (rtac injI 1); +by (etac (Atom_inject RS Inl_inject) 1); +val inj_Leaf = result(); + +val Leaf_inject = inj_Leaf RS injD; + +goalw Univ.thy [Numb_def] "inj(Numb)"; +by (stac o_def 1); +by (rtac injI 1); +by (etac (Atom_inject RS Inr_inject) 1); +val inj_Numb = result(); + +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])); +val Push_Node_inject = result(); + + +(** 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); +val Scons_inject_lemma1 = result(); + +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); +val Scons_inject_lemma2 = result(); + +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)); +val Scons_inject1 = result(); + +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)); +val Scons_inject2 = result(); + +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); +val Scons_inject = result(); + +(*rewrite rules*) +goal Univ.thy "(Atom(a)=Atom(b)) = (a=b)"; +by (fast_tac (HOL_cs addSEs [Atom_inject]) 1); +val Atom_Atom_eq = result(); + +goal Univ.thy "(M.N = M'.N') = (M=M' & N=N')"; +by (fast_tac (HOL_cs addSEs [Scons_inject]) 1); +val Scons_Scons_eq = result(); + +(*** Distinctness involving Leaf and Numb ***) + +(** Scons vs Leaf **) + +goalw Univ.thy [Leaf_def] "~ ((M.N) = Leaf(a))"; +by (stac o_def 1); +by (rtac Scons_not_Atom 1); +val Scons_not_Leaf = result(); +val Leaf_not_Scons = standard (Scons_not_Leaf RS not_sym); + +val Scons_neq_Leaf = standard (Scons_not_Leaf RS notE); +val Leaf_neq_Scons = sym RS Scons_neq_Leaf; + +(** Scons vs Numb **) + +goalw Univ.thy [Numb_def] "~ ((M.N) = Numb(k))"; +by (stac o_def 1); +by (rtac Scons_not_Atom 1); +val Scons_not_Numb = result(); +val Numb_not_Scons = standard (Scons_not_Numb RS not_sym); + +val Scons_neq_Numb = standard (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); +val Leaf_not_Numb = result(); +val Numb_not_Leaf = standard (Leaf_not_Numb RS not_sym); + +val Leaf_neq_Numb = standard (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); +val ndepth_K0 = result(); + +goal Univ.thy "k < Suc(LEAST x. f(x)=0) --> ~ nat_case(k, Suc(i), f) = 0"; +by (nat_ind_tac "k" 1); +by (ALLGOALS (simp_tac nat_ss)); +by (rtac impI 1); +by (etac not_less_Least 1); +val ndepth_Push_lemma = result(); + +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); +val ndepth_Push_Node = result(); + + +(*** 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])); +val ntrunc_0 = result(); + +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); +val ntrunc_Atom = result(); + +goalw Univ.thy [Leaf_def] "ntrunc(Suc(k), Leaf(a)) = Leaf(a)"; +by (stac o_def 1); +by (rtac ntrunc_Atom 1); +val ntrunc_Leaf = result(); + +goalw Univ.thy [Numb_def] "ntrunc(Suc(k), Numb(i)) = Numb(i)"; +by (stac o_def 1); +by (rtac ntrunc_Atom 1); +val ntrunc_Numb = result(); + +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)); +val ntrunc_Scons = result(); + +(** 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])); +val ntrunc_one_In0 = result(); + +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); +val ntrunc_In0 = result(); + +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])); +val ntrunc_one_In1 = result(); + +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); +val ntrunc_In1 = result(); + + +(*** 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)); +val uprodI = result(); + +(*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)); +val uprodE = result(); + +(*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)); +val uprodE2 = result(); + + +(*** Disjoint Sum ***) + +goalw Univ.thy [usum_def] "!!M. M:A ==> In0(M) : A<+>B"; +by (fast_tac set_cs 1); +val usum_In0I = result(); + +goalw Univ.thy [usum_def] "!!N. N:B ==> In1(N) : A<+>B"; +by (fast_tac set_cs 1); +val usum_In1I = result(); + +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)); +val usumE = result(); + + +(** 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); +val In0_not_In1 = result(); + +val In1_not_In0 = standard (In0_not_In1 RS not_sym); +val In0_neq_In1 = standard (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); +val In0_inject = result(); + +val [major] = goalw Univ.thy [In1_def] "In1(M) = In1(N) ==> M=N"; +by (rtac (major RS Scons_inject2) 1); +val In1_inject = result(); + + +(*** proving equality of sets and functions using ntrunc ***) + +goalw Univ.thy [ntrunc_def] "ntrunc(k,M) <= M"; +by (fast_tac set_cs 1); +val ntrunc_subsetI = result(); + +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); +val ntrunc_subsetD = result(); + +(*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); +val ntrunc_equality = result(); + +val [major] = goal Univ.thy + "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"; +by (rtac (ntrunc_equality RS ext) 1); +by (resolve_tac ([major RS fun_cong] RL [o_def RS subst]) 1); +val ntrunc_o_equality = result(); + +(*** Monotonicity ***) + +goalw Univ.thy [uprod_def] "!!A B. [| A<=A'; B<=B' |] ==> A<*>B <= A'<*>B'"; +by (fast_tac set_cs 1); +val uprod_mono = result(); + +goalw Univ.thy [usum_def] "!!A B. [| A<=A'; B<=B' |] ==> A<+>B <= A'<+>B'"; +by (fast_tac set_cs 1); +val usum_mono = result(); + +goalw Univ.thy [Scons_def] "!!M N. [| M<=M'; N<=N' |] ==> M.N <= M'.N'"; +by (fast_tac set_cs 1); +val Scons_mono = result(); + +goalw Univ.thy [In0_def] "!!M N. M<=N ==> In0(M) <= In0(N)"; +by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); +val In0_mono = result(); + +goalw Univ.thy [In1_def] "!!M N. M<=N ==> In1(M) <= In1(N)"; +by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); +val In1_mono = result(); + + +(*** Split and Case ***) + +goalw Univ.thy [Split_def] "Split(M.N, c) = c(M,N)"; +by (fast_tac (set_cs addIs [select_equality] addEs [Scons_inject]) 1); +val Split = result(); + +goalw Univ.thy [Case_def] "Case(In0(M), c, d) = c(M)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim In0_inject, In0_neq_In1]) 1); +val Case_In0 = result(); + +goalw Univ.thy [Case_def] "Case(In1(N), c, d) = d(N)"; +by (fast_tac (set_cs addIs [select_equality] + addEs [make_elim In1_inject, In1_neq_In0]) 1); +val Case_In1 = result(); + +(**** 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); +val ntrunc_UN1 = result(); + +goalw Univ.thy [Scons_def] "(UN x.f(x)) . M = (UN x. f(x) . M)"; +by (fast_tac (set_cs addIs [equalityI]) 1); +val Scons_UN1_x = result(); + +goalw Univ.thy [Scons_def] "M . (UN x.f(x)) = (UN x. M . f(x))"; +by (fast_tac (set_cs addIs [equalityI]) 1); +val Scons_UN1_y = result(); + +goalw Univ.thy [In0_def] "In0(UN x.f(x)) = (UN x. In0(f(x)))"; +br Scons_UN1_y 1; +val In0_UN1 = result(); + +goalw Univ.thy [In1_def] "In1(UN x.f(x)) = (UN x. In1(f(x)))"; +br Scons_UN1_y 1; +val In1_UN1 = result(); + + +(*** Equality : the diagonal relation ***) + +goalw Univ.thy [diag_def] "!!a A. a:A ==> : diag(A)"; +by (REPEAT (ares_tac [singletonI,UN_I] 1)); +val diagI = result(); + +(*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)); +val diagE = result(); + +(*** Equality for Cartesian Product ***) + +goal Univ.thy + "split(, %x x'. split(, %y y'. {})) = {}"; +by (simp_tac univ_ss 1); +val dprod_lemma = result(); + +goalw Univ.thy [dprod_def] + "!!r s. [| :r; :s |] ==> : r<**>s"; +by (REPEAT (ares_tac [UN_I] 1)); +by (rtac (singletonI RS (dprod_lemma RS equalityD2 RS subsetD)) 1); +val dprodI = result(); + +(*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 (eresolve_tac [asm_rl,singletonE,UN_E] 1)); +by (res_inst_tac [("p","u")] PairE 1); +by (res_inst_tac [("p","v")] PairE 1); +by (safe_tac HOL_cs); +by (REPEAT (ares_tac prems 1)); +by (safe_tac (set_cs addSDs [dprod_lemma RS equalityD1 RS subsetD])); +val dprodE = result(); + + +(*** Equality for Disjoint Sum ***) + +goalw Univ.thy [dsum_def] "!!r. :r ==> : r<++>s"; +by (fast_tac (set_cs addSIs [split RS equalityD2 RS subsetD]) 1); +val dsum_In0I = result(); + +goalw Univ.thy [dsum_def] "!!r. :s ==> : r<++>s"; +by (fast_tac (set_cs addSIs [split RS equalityD2 RS subsetD]) 1); +val dsum_In1I = result(); + +val major::prems = goalw Univ.thy [dsum_def] + "[| w : r<++>s; \ +\ !!x x'. [| : r; w = |] ==> P; \ +\ !!y y'. [| : s; w = |] ==> P \ +\ |] ==> P"; +by (rtac (major RS UnE) 1); +by (safe_tac set_cs); +by (res_inst_tac [("p","u")] PairE 1); +by (res_inst_tac [("p","v")] PairE 2); +by (safe_tac (set_cs addSEs prems + addSDs [split RS equalityD1 RS subsetD])); +val dsumE = result(); + + +(*** Monotonicity ***) + +goalw Univ.thy [dprod_def] "!!r s. [| r<=r'; s<=s' |] ==> r<**>s <= r'<**>s'"; +by (fast_tac set_cs 1); +val dprod_mono = result(); + +goalw Univ.thy [dsum_def] "!!r s. [| r<=r'; s<=s' |] ==> r<++>s <= r'<++>s'"; +by (fast_tac set_cs 1); +val dsum_mono = result(); + + +(*** Bounding theorems ***) + +goal Univ.thy "diag(A) <= Sigma(A,%x.A)"; +by (fast_tac (set_cs addIs [SigmaI] addSEs [diagE]) 1); +val diag_subset_Sigma = result(); + +val prems = goal Univ.thy + "[| r <= Sigma(A,%x.B); s <= Sigma(C,%x.D) |] ==> \ +\ (r<**>s) <= Sigma(A<*>C, %z. B<*>D)"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSIs [SigmaI,uprodI] + addSEs [dprodE,SigmaE2]) 1); +val dprod_subset_Sigma = result(); + +goal Univ.thy + "!!r s. [| r <= Sigma(A,B); s <= Sigma(C,D) |] ==> \ +\ (r<**>s) <= Sigma(A<*>C, %z. Split(z, %x y. B(x)<*>D(y)))"; +by (safe_tac (set_cs addSIs [SigmaI,uprodI] addSEs [dprodE])); +by (stac Split 3); +by (ALLGOALS (fast_tac (set_cs addSIs [uprodI] addSEs [SigmaE2]))); +val dprod_subset_Sigma2 = result(); + +goal Univ.thy + "!!r s. [| r <= Sigma(A,%x.B); s <= Sigma(C,%x.D) |] ==> \ +\ (r<++>s) <= Sigma(A<+>C, %z. B<+>D)"; +by (fast_tac (set_cs addSIs [SigmaI,usum_In0I,usum_In1I] + addSEs [dsumE,SigmaE2]) 1); +val dsum_subset_Sigma = result(); + + +(*** Domain ***) + +goal Univ.thy "fst `` diag(A) = A"; +by (fast_tac (set_cs addIs [equalityI, fst_imageI, diagI] + addSEs [fst_imageE, Pair_inject, diagE]) 1); +val fst_image_diag = result(); + +goal Univ.thy "fst `` (r<**>s) = (fst``r) <*> (fst``s)"; +by (fast_tac (set_cs addIs [equalityI, fst_imageI, uprodI, dprodI] + addSEs [fst_imageE, Pair_inject, uprodE, dprodE]) 1); +val fst_image_dprod = result(); + +goal Univ.thy "fst `` (r<++>s) = (fst``r) <+> (fst``s)"; +by (fast_tac (set_cs addIs [equalityI, fst_imageI, usum_In0I, usum_In1I, + dsum_In0I, dsum_In1I] + addSEs [fst_imageE, Pair_inject, usumE, dsumE]) 1); +val fst_image_dsum = result(); + +val fst_image_simps = [fst_image_diag, fst_image_dprod, fst_image_dsum]; +val fst_image_ss = univ_ss addsimps fst_image_simps; + +val univ_cs = + set_cs addSIs [SigmaI,uprodI,dprodI] + addIs [usum_In0I,usum_In1I,dsum_In0I,dsum_In1I] + addSEs [diagE,uprodE,dprodE,usumE,dsumE,SigmaE2,Pair_inject]; diff -r 000000000000 -r 7949f97df77a univ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/univ.thy Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,105 @@ +(* 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 + +types node 1 +arities node :: (term)term +consts + Least :: "(nat=>bool) => nat" (binder "LEAST " 10) + + apfst :: "['a=>'c, 'a*'b] => 'c*'b" + Push :: "[nat, nat=>nat] => (nat=>nat)" + + Node :: "((nat=>nat) * ('a+nat)) set" + Rep_Node :: "'a node => (nat=>nat) * ('a+nat)" + Abs_Node :: "(nat=>nat) * ('a+nat) => 'a node" + Push_Node :: "[nat, 'a node] => 'a node" + ndepth :: "'a node => nat" + + Atom :: "('a+nat) => 'a node set" + Leaf :: "'a => 'a node set" + Numb :: "nat => 'a node set" + "." :: "['a node set, 'a node set]=> 'a node set" (infixr 60) + In0,In1 :: "'a node set => 'a node set" + + ntrunc :: "[nat, 'a node set] => 'a node set" + + "<*>" :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 80) + "<+>" :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 70) + + Split :: "['a node set, ['a node set, 'a node set]=>'b] => 'b" + Case :: "['a node set, ['a node set]=>'b, ['a node set]=>'b] => 'b" + + diag :: "'a set => ('a * 'a)set" + "<**>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \ +\ => ('a node set * 'a node set)set" (infixr 80) + "<++>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \ +\ => ('a node set * 'a node set)set" (infixr 70) + +rules + + (*least number operator*) + Least_def "Least(P) == @k. P(k) & (ALL j. j ~P(j))" + + (** lists, trees will be sets of nodes **) + Node_def "Node == {p. EX f x k. p = & f(k)=0}" + (*faking the type definition 'a node == (nat=>nat) * ('a+nat) *) + Rep_Node "Rep_Node(n): Node" + Rep_Node_inverse "Abs_Node(Rep_Node(n)) = n" + Abs_Node_inverse "p: Node ==> Rep_Node(Abs_Node(p)) = p" + 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 p.split(p, %x y. ))" + Push_def "Push == (%b h n. nat_case(n,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(Rep_Node(n), %f x. LEAST k. f(k)=0)" + 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(M,c) == @u. ? x y. M = x.y & u = c(x,y)" + + Case_def "Case(M,c,d) == @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. UN v:s. \ +\ split(u, %x x'. split(v, %y y'. {}))" + + dsum_def "r<++>s == (UN u:r. split(u, %x x'. {})) Un \ +\ (UN v:s. split(v, %y y'. {}))" + +end diff -r 000000000000 -r 7949f97df77a wf.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/wf.ML Thu Sep 16 12:21:07 1993 +0200 @@ -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); +val wfI = result(); + +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); +val wf_induct = result(); + +(*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); +val wf_anti_sym = result(); + +val prems = goal WF.thy "[| wf(r); : r |] ==> P"; +by (rtac wf_anti_sym 1); +by (REPEAT (resolve_tac prems 1)); +val wf_anti_refl = result(); + +(*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); +val wf_trancl = result(); + + +(** 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); +val cut_cut_eq = result(); + +goalw WF.thy [cut_def] "!!x. :r ==> cut(f,r,a)(x) = f(x)"; +by(asm_simp_tac HOL_ss 1); +val cut_apply = result(); + + +(*** 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); +val is_recfun_undef = result(); + +(*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); +val is_recfun_equal_lemma = result(); +val is_recfun_equal = standard (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); +val is_recfun_cut = result(); + +(*** 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); +val is_the_recfun = result(); + +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); +val unfold_the_recfun = result(); + + +(*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)); +val the_recfun_equal = result(); + +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)); +val the_recfun_cut = result(); + +(*** 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); +val wftrec = result(); + +(*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)); +val wftrec2 = result(); + +(** 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); +val wfrec = result(); + +(*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)); +val def_wfrec = result(); diff -r 000000000000 -r 7949f97df77a wf.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/wf.thy Thu Sep 16 12:21:07 1993 +0200 @@ -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" + +rules + 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