# HG changeset patch # User regensbu # Date 812996724 -3600 # Node ID ea0668a1c0ba751e239cb8474f65baf56677ffe1 # Parent 6960ec882bca8146fca07b26dde551b66ca9105b added 8bit pragmas added directory ax_ops for sections axioms and ops added directory domain for sections domain and generated this is the type definition package of David Oheimb diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Cfun1.thy --- a/src/HOLCF/Cfun1.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Cfun1.thy Fri Oct 06 17:25:24 1995 +0100 @@ -47,4 +47,11 @@ (*defining the abstract constants*) less_cfun_def "less_cfun fo1 fo2 == ( fapp fo1 << fapp fo2 )" +(* start 8bit 1 *) +(* end 8bit 1 *) + + end + +(* start 8bit 2 *) +(* end 8bit 2 *) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Cont.thy --- a/src/HOLCF/Cont.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Cont.thy Fri Oct 06 17:25:24 1995 +0100 @@ -27,15 +27,15 @@ monofun "monofun(f) == ! x y. x << y --> f(x) << f(y)" -contlub "contlub(f) == ! Y. is_chain(Y) --> \ -\ f(lub(range(Y))) = lub(range(% i.f(Y(i))))" +contlub "contlub(f) == ! Y. is_chain(Y) --> + f(lub(range(Y))) = lub(range(% i.f(Y(i))))" -cont "cont(f) == ! Y. is_chain(Y) --> \ -\ range(% i.f(Y(i))) <<| f(lub(range(Y)))" +cont "cont(f) == ! Y. is_chain(Y) --> + range(% i.f(Y(i))) <<| f(lub(range(Y)))" (* ------------------------------------------------------------------------ *) (* the main purpose of cont.thy is to show: *) -(* monofun(f) & contlub(f) <==> cont(f) *) +(* monofun(f) & contlub(f) <==> cont(f) *) (* ------------------------------------------------------------------------ *) end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Cprod3.ML --- a/src/HOLCF/Cprod3.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Cprod3.ML Fri Oct 06 17:25:24 1995 +0100 @@ -314,3 +314,4 @@ Addsimps [cfst2,csnd2,csplit2]; +val Cprod_rews = [cfst2,csnd2,csplit2]; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Cprod3.thy --- a/src/HOLCF/Cprod3.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Cprod3.thy Fri Oct 06 17:25:24 1995 +0100 @@ -36,6 +36,53 @@ csnd_def "csnd == (LAM p.snd(p))" csplit_def "csplit == (LAM f p.f`(cfst`p)`(csnd`p))" + + +(* introduce syntax for + + Let = e1; z = E2 in E3 + + and + + ¤.e +*) + +types + Cletbinds Cletbind + +consts + CLet :: "'a -> ('a -> 'b) -> 'b" + +syntax + (* syntax for Let *) + + "_Cbind" :: "[pttrn, 'a] => Cletbind" ("(2_ =/ _)" 10) + "" :: "Cletbind => Cletbinds" ("_") + "_Cbinds" :: "[Cletbind, Cletbinds] => Cletbinds" ("_;/ _") + "_CLet" :: "[Cletbinds, 'a] => 'a" ("(Let (_)/ in (_))" 10) + +translations + (* translation for Let *) + "_CLet (_Cbinds b bs) e" == "_CLet b (_CLet bs e)" + "Let x = a in e" == "CLet`a`(LAM x.e)" + +defs + (* Misc Definitions *) + CLet_def "CLet == LAM s. LAM f.f`s" + + +syntax + (* syntax for LAM .E *) + "@Cpttrn" :: "[pttrn,pttrns] => pttrn" ("<_,/_>") + +translations + (* translations for LAM .E *) + "LAM .b" == "csplit`(LAM x.LAM .b)" + "LAM .b" == "csplit`(LAM x.LAM y.b)" + (* reverse translation <= does not work yet !! *) + +(* start 8bit 1 *) +(* end 8bit 1 *) end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Dlist.ML --- a/src/HOLCF/Dlist.ML Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,564 +0,0 @@ -(* Title: HOLCF/dlist.ML - Author: Franz Regensburger - ID: $ $ - Copyright 1994 Technische Universitaet Muenchen - -Lemmas for dlist.thy -*) - -open Dlist; - -(* ------------------------------------------------------------------------*) -(* The isomorphisms dlist_rep_iso and dlist_abs_iso are strict *) -(* ------------------------------------------------------------------------*) - -val dlist_iso_strict= dlist_rep_iso RS (dlist_abs_iso RS - (allI RSN (2,allI RS iso_strict))); - -val dlist_rews = [dlist_iso_strict RS conjunct1, - dlist_iso_strict RS conjunct2]; - -(* ------------------------------------------------------------------------*) -(* Properties of dlist_copy *) -(* ------------------------------------------------------------------------*) - -val temp = prove_goalw Dlist.thy [dlist_copy_def] "dlist_copy`f`UU=UU" - (fn prems => - [ - (asm_simp_tac (!simpset addsimps - (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) - ]); - -val dlist_copy = [temp]; - - -val temp = prove_goalw Dlist.thy [dlist_copy_def,dnil_def] - "dlist_copy`f`dnil=dnil" - (fn prems => - [ - (asm_simp_tac (!simpset addsimps - (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) - ]); - -val dlist_copy = temp :: dlist_copy; - - -val temp = prove_goalw Dlist.thy [dlist_copy_def,dcons_def] - "xl~=UU ==> dlist_copy`f`(dcons`x`xl)= dcons`x`(f`xl)" - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1), - (res_inst_tac [("Q","x=UU")] classical2 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps [defined_spair]) 1) - ]); - -val dlist_copy = temp :: dlist_copy; - -val dlist_rews = dlist_copy @ dlist_rews; - -(* ------------------------------------------------------------------------*) -(* Exhaustion and elimination for dlists *) -(* ------------------------------------------------------------------------*) - -qed_goalw "Exh_dlist" Dlist.thy [dcons_def,dnil_def] - "l = UU | l = dnil | (? x xl. x~=UU &xl~=UU & l = dcons`x`xl)" - (fn prems => - [ - (Simp_tac 1), - (rtac (dlist_rep_iso RS subst) 1), - (res_inst_tac [("p","dlist_rep`l")] ssumE 1), - (rtac disjI1 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (rtac disjI2 1), - (rtac disjI1 1), - (res_inst_tac [("p","x")] oneE 1), - (contr_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (rtac disjI2 1), - (rtac disjI2 1), - (res_inst_tac [("p","y")] sprodE 1), - (contr_tac 1), - (rtac exI 1), - (rtac exI 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (fast_tac HOL_cs 1) - ]); - - -qed_goal "dlistE" Dlist.thy -"[| l=UU ==> Q; l=dnil ==> Q;!!x xl.[|l=dcons`x`xl;x~=UU;xl~=UU|]==>Q|]==>Q" - (fn prems => - [ - (rtac (Exh_dlist RS disjE) 1), - (eresolve_tac prems 1), - (etac disjE 1), - (eresolve_tac prems 1), - (etac exE 1), - (etac exE 1), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); - -(* ------------------------------------------------------------------------*) -(* Properties of dlist_when *) -(* ------------------------------------------------------------------------*) - -val temp = prove_goalw Dlist.thy [dlist_when_def] "dlist_when`f1`f2`UU=UU" - (fn prems => - [ - (asm_simp_tac (!simpset addsimps [dlist_iso_strict]) 1) - ]); - -val dlist_when = [temp]; - -val temp = prove_goalw Dlist.thy [dlist_when_def,dnil_def] - "dlist_when`f1`f2`dnil= f1" - (fn prems => - [ - (asm_simp_tac (!simpset addsimps [dlist_abs_iso]) 1) - ]); - -val dlist_when = temp::dlist_when; - -val temp = prove_goalw Dlist.thy [dlist_when_def,dcons_def] - "[|x~=UU;xl~=UU|] ==> dlist_when`f1`f2`(dcons`x`xl)= f2`x`xl" - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps [dlist_abs_iso,defined_spair]) 1) - ]); - -val dlist_when = temp::dlist_when; - -val dlist_rews = dlist_when @ dlist_rews; - -(* ------------------------------------------------------------------------*) -(* Rewrites for discriminators and selectors *) -(* ------------------------------------------------------------------------*) - -fun prover defs thm = prove_goalw Dlist.thy defs thm - (fn prems => - [ - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_discsel = [ - prover [is_dnil_def] "is_dnil`UU=UU", - prover [is_dcons_def] "is_dcons`UU=UU", - prover [dhd_def] "dhd`UU=UU", - prover [dtl_def] "dtl`UU=UU" - ]; - -fun prover defs thm = prove_goalw Dlist.thy defs thm - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_discsel = [ -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "is_dnil`dnil=TT", -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> is_dnil`(dcons`x`xl)=FF", -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "is_dcons`dnil=FF", -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> is_dcons`(dcons`x`xl)=TT", -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "dhd`dnil=UU", -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> dhd`(dcons`x`xl)=x", -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "dtl`dnil=UU", -prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> dtl`(dcons`x`xl)=xl"] @ dlist_discsel; - -val dlist_rews = dlist_discsel @ dlist_rews; - -(* ------------------------------------------------------------------------*) -(* Definedness and strictness *) -(* ------------------------------------------------------------------------*) - -fun prover contr thm = prove_goal Dlist.thy thm - (fn prems => - [ - (res_inst_tac [("P1",contr)] classical3 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (dtac sym 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps (prems @ dlist_rews)) 1) - ]); - - -val dlist_constrdef = [ -prover "is_dnil`(UU::'a dlist) ~= UU" "dnil~=(UU::'a dlist)", -prover "is_dcons`(UU::'a dlist) ~= UU" - "[|x~=UU;xl~=UU|]==>dcons`(x::'a)`xl ~= UU" - ]; - - -fun prover defs thm = prove_goalw Dlist.thy defs thm - (fn prems => - [ - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_constrdef = [ - prover [dcons_def] "dcons`UU`xl=UU", - prover [dcons_def] "dcons`x`UU=UU" - ] @ dlist_constrdef; - -val dlist_rews = dlist_constrdef @ dlist_rews; - - -(* ------------------------------------------------------------------------*) -(* Distinctness wrt. << and = *) -(* ------------------------------------------------------------------------*) - -val temp = prove_goal Dlist.thy "~dnil << dcons`(x::'a)`xl" - (fn prems => - [ - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dnil")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","(x::'a)=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","(xl ::'a dlist)=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_dist_less = [temp]; - -val temp = prove_goal Dlist.thy "[|x~=UU;xl~=UU|]==>~ dcons`x`xl << dnil" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dcons")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_dist_less = temp::dlist_dist_less; - -val temp = prove_goal Dlist.thy "dnil ~= dcons`x`xl" - (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("P1","TT = FF")] classical3 1), - (resolve_tac dist_eq_tr 1), - (dres_inst_tac [("f","is_dnil")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_dist_eq = [temp,temp RS not_sym]; - -val dlist_rews = dlist_dist_less @ dlist_dist_eq @ dlist_rews; - -(* ------------------------------------------------------------------------*) -(* Invertibility *) -(* ------------------------------------------------------------------------*) - -val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ -\ dcons`x1`x2 << dcons`y1`y2 |] ==> x1<< y1 & x2 << y2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.x)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.l)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1) - ]); - -val dlist_invert =[temp]; - -(* ------------------------------------------------------------------------*) -(* Injectivity *) -(* ------------------------------------------------------------------------*) - -val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ -\ dcons`x1`x2 = dcons`y1`y2 |] ==> x1= y1 & x2 = y2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("f","dlist_when`UU`(LAM x l.x)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (dres_inst_tac [("f","dlist_when`UU`(LAM x l.l)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1) - ]); - -val dlist_inject = [temp]; - - -(* ------------------------------------------------------------------------*) -(* definedness for discriminators and selectors *) -(* ------------------------------------------------------------------------*) - -fun prover thm = prove_goal Dlist.thy thm - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac dlistE 1), - (contr_tac 1), - (REPEAT (asm_simp_tac (!simpset addsimps dlist_discsel) 1)) - ]); - -val dlist_discsel_def = - [ - prover "l~=UU ==> is_dnil`l~=UU", - prover "l~=UU ==> is_dcons`l~=UU" - ]; - -val dlist_rews = dlist_discsel_def @ dlist_rews; - -(* ------------------------------------------------------------------------*) -(* enhance the simplifier *) -(* ------------------------------------------------------------------------*) - -qed_goal "dhd2" Dlist.thy "xl~=UU ==> dhd`(dcons`x`xl)=x" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -qed_goal "dtl2" Dlist.thy "x~=UU ==> dtl`(dcons`x`xl)=xl" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_rews = dhd2 :: dtl2 :: dlist_rews; - -(* ------------------------------------------------------------------------*) -(* Properties dlist_take *) -(* ------------------------------------------------------------------------*) - -val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take n`UU=UU" - (fn prems => - [ - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_take = [temp]; - -val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take 0`xs=UU" - (fn prems => - [ - (Asm_simp_tac 1) - ]); - -val dlist_take = temp::dlist_take; - -val temp = prove_goalw Dlist.thy [dlist_take_def] - "dlist_take (Suc n)`dnil=dnil" - (fn prems => - [ - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_take = temp::dlist_take; - -val temp = prove_goalw Dlist.thy [dlist_take_def] - "dlist_take (Suc n)`(dcons`x`xl)= dcons`x`(dlist_take n`xl)" - (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -val dlist_take = temp::dlist_take; - -val dlist_rews = dlist_take @ dlist_rews; - -(* ------------------------------------------------------------------------*) -(* take lemma for dlists *) -(* ------------------------------------------------------------------------*) - -fun prover reach defs thm = prove_goalw Dlist.thy defs thm - (fn prems => - [ - (res_inst_tac [("t","l1")] (reach RS subst) 1), - (res_inst_tac [("t","l2")] (reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac lub_equal 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); - -val dlist_take_lemma = prover dlist_reach [dlist_take_def] - "(!!n.dlist_take n`l1 = dlist_take n`l2) ==> l1=l2"; - - -(* ------------------------------------------------------------------------*) -(* Co -induction for dlists *) -(* ------------------------------------------------------------------------*) - -qed_goalw "dlist_coind_lemma" Dlist.thy [dlist_bisim_def] -"dlist_bisim R ==> ! p q. R p q --> dlist_take n`p = dlist_take n`q" - (fn prems => - [ - (cut_facts_tac prems 1), - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (strip_tac 1), - ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), - (atac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (etac exE 1), - (etac exE 1), - (etac exE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (REPEAT (etac conjE 1)), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); - -qed_goal "dlist_coind" Dlist.thy "[|dlist_bisim R ; R p q |] ==> p = q" - (fn prems => - [ - (rtac dlist_take_lemma 1), - (rtac (dlist_coind_lemma RS spec RS spec RS mp) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); - -(* ------------------------------------------------------------------------*) -(* structural induction *) -(* ------------------------------------------------------------------------*) - -qed_goal "dlist_finite_ind" Dlist.thy -"[|P(UU);P(dnil);\ -\ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons`x`l1)\ -\ |] ==> !l.P(dlist_take n`l)" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (rtac allI 1), - (res_inst_tac [("l","l")] dlistE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","dlist_take n1`xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (atac 1), - (atac 1), - (etac spec 1) - ]); - -qed_goal "dlist_all_finite_lemma1" Dlist.thy -"!l.dlist_take n`l=UU |dlist_take n`l=l" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (rtac allI 1), - (res_inst_tac [("l","l")] dlistE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (eres_inst_tac [("x","xl")] allE 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); - -qed_goal "dlist_all_finite_lemma2" Dlist.thy "? n.dlist_take n`l=l" - (fn prems => - [ - (res_inst_tac [("Q","l=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (subgoal_tac "(!n.dlist_take n`l=UU) |(? n.dlist_take n`l = l)" 1), - (etac disjE 1), - (eres_inst_tac [("P","l=UU")] notE 1), - (rtac dlist_take_lemma 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (atac 1), - (subgoal_tac "!n.!l.dlist_take n`l=UU |dlist_take n`l=l" 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (rtac dlist_all_finite_lemma1 1) - ]); - -qed_goalw "dlist_all_finite" Dlist.thy [dlist_finite_def] "dlist_finite(l)" - (fn prems => - [ - (rtac dlist_all_finite_lemma2 1) - ]); - -qed_goal "dlist_ind" Dlist.thy -"[|P(UU);P(dnil);\ -\ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons`x`l1)|] ==> P(l)" - (fn prems => - [ - (rtac (dlist_all_finite_lemma2 RS exE) 1), - (etac subst 1), - (rtac (dlist_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); - - - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Dlist.thy --- a/src/HOLCF/Dlist.thy Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,126 +0,0 @@ -(* Title: HOLCF/dlist.thy - - Author: Franz Regensburger - ID: $ $ - Copyright 1994 Technische Universitaet Muenchen - -Theory for finite lists 'a dlist = one ++ ('a ** 'a dlist) - -The type is axiomatized as the least solution of the domain equation above. -The functor term that specifies the domain equation is: - - FT = <++,K_{one},<**,K_{'a},I>> - -For details see chapter 5 of: - -[Franz Regensburger] HOLCF: Eine konservative Erweiterung von HOL um LCF, - Dissertation, Technische Universit"at M"unchen, 1994 - - -*) - -Dlist = Stream2 + - -types dlist 1 - -(* ----------------------------------------------------------------------- *) -(* arity axiom is validated by semantic reasoning *) -(* partial ordering is implicit in the isomorphism axioms and their cont. *) - -arities dlist::(pcpo)pcpo - -consts - -(* ----------------------------------------------------------------------- *) -(* essential constants *) - -dlist_rep :: "('a dlist) -> (one ++ 'a ** 'a dlist)" -dlist_abs :: "(one ++ 'a ** 'a dlist) -> ('a dlist)" - -(* ----------------------------------------------------------------------- *) -(* abstract constants and auxiliary constants *) - -dlist_copy :: "('a dlist -> 'a dlist) ->'a dlist -> 'a dlist" - -dnil :: "'a dlist" -dcons :: "'a -> 'a dlist -> 'a dlist" -dlist_when :: " 'b -> ('a -> 'a dlist -> 'b) -> 'a dlist -> 'b" -is_dnil :: "'a dlist -> tr" -is_dcons :: "'a dlist -> tr" -dhd :: "'a dlist -> 'a" -dtl :: "'a dlist -> 'a dlist" -dlist_take :: "nat => 'a dlist -> 'a dlist" -dlist_finite :: "'a dlist => bool" -dlist_bisim :: "('a dlist => 'a dlist => bool) => bool" - -rules - -(* ----------------------------------------------------------------------- *) -(* axiomatization of recursive type 'a dlist *) -(* ----------------------------------------------------------------------- *) -(* ('a dlist,dlist_abs) is the initial F-algebra where *) -(* F is the locally continuous functor determined by functor term FT. *) -(* domain equation: 'a dlist = one ++ ('a ** 'a dlist) *) -(* functor term: FT = <++,K_{one},<**,K_{'a},I>> *) -(* ----------------------------------------------------------------------- *) -(* dlist_abs is an isomorphism with inverse dlist_rep *) -(* identity is the least endomorphism on 'a dlist *) - -dlist_abs_iso "dlist_rep`(dlist_abs`x) = x" -dlist_rep_iso "dlist_abs`(dlist_rep`x) = x" -dlist_copy_def "dlist_copy == (LAM f. dlist_abs oo \ -\ (sswhen`sinl`(sinr oo (ssplit`(LAM x y. (|x,f`y|) ))))\ -\ oo dlist_rep)" -dlist_reach "(fix`dlist_copy)`x=x" - - -defs -(* ----------------------------------------------------------------------- *) -(* properties of additional constants *) -(* ----------------------------------------------------------------------- *) -(* constructors *) - -dnil_def "dnil == dlist_abs`(sinl`one)" -dcons_def "dcons == (LAM x l. dlist_abs`(sinr`(|x,l|) ))" - -(* ----------------------------------------------------------------------- *) -(* discriminator functional *) - -dlist_when_def -"dlist_when == (LAM f1 f2 l.\ -\ sswhen`(LAM x.f1) `(ssplit`(LAM x l.f2`x`l)) `(dlist_rep`l))" - -(* ----------------------------------------------------------------------- *) -(* discriminators and selectors *) - -is_dnil_def "is_dnil == dlist_when`TT`(LAM x l.FF)" -is_dcons_def "is_dcons == dlist_when`FF`(LAM x l.TT)" -dhd_def "dhd == dlist_when`UU`(LAM x l.x)" -dtl_def "dtl == dlist_when`UU`(LAM x l.l)" - -(* ----------------------------------------------------------------------- *) -(* the taker for dlists *) - -dlist_take_def "dlist_take == (%n.iterate n dlist_copy UU)" - -(* ----------------------------------------------------------------------- *) - -dlist_finite_def "dlist_finite == (%s.? n.dlist_take n`s=s)" - -(* ----------------------------------------------------------------------- *) -(* definition of bisimulation is determined by domain equation *) -(* simplification and rewriting for abstract constants yields def below *) - -dlist_bisim_def "dlist_bisim == - ( %R.!l1 l2. - R l1 l2 --> - ((l1=UU & l2=UU) | - (l1=dnil & l2=dnil) | - (? x l11 l21. x~=UU & l11~=UU & l21~=UU & - l1=dcons`x`l11 & l2 = dcons`x`l21 & R l11 l21)))" - -end - - - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Dnat.ML --- a/src/HOLCF/Dnat.ML Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,534 +0,0 @@ -(* Title: HOLCF/dnat.ML - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Lemmas for dnat.thy -*) - -open Dnat; - -(* ------------------------------------------------------------------------*) -(* The isomorphisms dnat_rep_iso and dnat_abs_iso are strict *) -(* ------------------------------------------------------------------------*) - -val dnat_iso_strict = dnat_rep_iso RS (dnat_abs_iso RS - (allI RSN (2,allI RS iso_strict))); - -val dnat_rews = [dnat_iso_strict RS conjunct1, - dnat_iso_strict RS conjunct2]; - -(* ------------------------------------------------------------------------*) -(* Properties of dnat_copy *) -(* ------------------------------------------------------------------------*) - -fun prover defs thm = prove_goalw Dnat.thy defs thm - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) - ]); - -val dnat_copy = - [ - prover [dnat_copy_def] "dnat_copy`f`UU=UU", - prover [dnat_copy_def,dzero_def] "dnat_copy`f`dzero= dzero", - prover [dnat_copy_def,dsucc_def] - "n~=UU ==> dnat_copy`f`(dsucc`n) = dsucc`(f`n)" - ]; - -val dnat_rews = dnat_copy @ dnat_rews; - -(* ------------------------------------------------------------------------*) -(* Exhaustion and elimination for dnat *) -(* ------------------------------------------------------------------------*) - -qed_goalw "Exh_dnat" Dnat.thy [dsucc_def,dzero_def] - "n = UU | n = dzero | (? x . x~=UU & n = dsucc`x)" - (fn prems => - [ - (Simp_tac 1), - (rtac (dnat_rep_iso RS subst) 1), - (res_inst_tac [("p","dnat_rep`n")] ssumE 1), - (rtac disjI1 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (rtac (disjI1 RS disjI2) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("p","x")] oneE 1), - (contr_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (rtac (disjI2 RS disjI2) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (fast_tac HOL_cs 1) - ]); - -qed_goal "dnatE" Dnat.thy - "[| n=UU ==> Q; n=dzero ==> Q; !!x.[|n=dsucc`x;x~=UU|]==>Q|]==>Q" - (fn prems => - [ - (rtac (Exh_dnat RS disjE) 1), - (eresolve_tac prems 1), - (etac disjE 1), - (eresolve_tac prems 1), - (REPEAT (etac exE 1)), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); - -(* ------------------------------------------------------------------------*) -(* Properties of dnat_when *) -(* ------------------------------------------------------------------------*) - -fun prover defs thm = prove_goalw Dnat.thy defs thm - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) - ]); - - -val dnat_when = [ - prover [dnat_when_def] "dnat_when`c`f`UU=UU", - prover [dnat_when_def,dzero_def] "dnat_when`c`f`dzero=c", - prover [dnat_when_def,dsucc_def] - "n~=UU ==> dnat_when`c`f`(dsucc`n)=f`n" - ]; - -val dnat_rews = dnat_when @ dnat_rews; - -(* ------------------------------------------------------------------------*) -(* Rewrites for discriminators and selectors *) -(* ------------------------------------------------------------------------*) - -fun prover defs thm = prove_goalw Dnat.thy defs thm - (fn prems => - [ - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_discsel = [ - prover [is_dzero_def] "is_dzero`UU=UU", - prover [is_dsucc_def] "is_dsucc`UU=UU", - prover [dpred_def] "dpred`UU=UU" - ]; - - -fun prover defs thm = prove_goalw Dnat.thy defs thm - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_discsel = [ - prover [is_dzero_def] "is_dzero`dzero=TT", - prover [is_dzero_def] "n~=UU ==>is_dzero`(dsucc`n)=FF", - prover [is_dsucc_def] "is_dsucc`dzero=FF", - prover [is_dsucc_def] "n~=UU ==> is_dsucc`(dsucc`n)=TT", - prover [dpred_def] "dpred`dzero=UU", - prover [dpred_def] "n~=UU ==> dpred`(dsucc`n)=n" - ] @ dnat_discsel; - -val dnat_rews = dnat_discsel @ dnat_rews; - -(* ------------------------------------------------------------------------*) -(* Definedness and strictness *) -(* ------------------------------------------------------------------------*) - -fun prover contr thm = prove_goal Dnat.thy thm - (fn prems => - [ - (res_inst_tac [("P1",contr)] classical3 1), - (simp_tac (!simpset addsimps dnat_rews) 1), - (dtac sym 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps (prems @ dnat_rews)) 1) - ]); - -val dnat_constrdef = [ - prover "is_dzero`UU ~= UU" "dzero~=UU", - prover "is_dsucc`UU ~= UU" "n~=UU ==> dsucc`n~=UU" - ]; - - -fun prover defs thm = prove_goalw Dnat.thy defs thm - (fn prems => - [ - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_constrdef = [ - prover [dsucc_def] "dsucc`UU=UU" - ] @ dnat_constrdef; - -val dnat_rews = dnat_constrdef @ dnat_rews; - - -(* ------------------------------------------------------------------------*) -(* Distinctness wrt. << and = *) -(* ------------------------------------------------------------------------*) - -val temp = prove_goal Dnat.thy "~dzero << dsucc`n" - (fn prems => - [ - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dzero")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("Q","n=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_dist_less = [temp]; - -val temp = prove_goal Dnat.thy "n~=UU ==> ~dsucc`n << dzero" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dsucc")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_dist_less = temp::dnat_dist_less; - -val temp = prove_goal Dnat.thy "dzero ~= dsucc`n" - (fn prems => - [ - (res_inst_tac [("Q","n=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("P1","TT = FF")] classical3 1), - (resolve_tac dist_eq_tr 1), - (dres_inst_tac [("f","is_dzero")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_dist_eq = [temp, temp RS not_sym]; - -val dnat_rews = dnat_dist_less @ dnat_dist_eq @ dnat_rews; - -(* ------------------------------------------------------------------------*) -(* Invertibility *) -(* ------------------------------------------------------------------------*) - -val dnat_invert = - [ -prove_goal Dnat.thy -"[|x1~=UU; y1~=UU; dsucc`x1 << dsucc`y1 |] ==> x1<< y1" - (fn prems => - [ - (cut_facts_tac prems 1), - (dres_inst_tac [("fo5","dnat_when`c`(LAM x.x)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]) - ]; - -(* ------------------------------------------------------------------------*) -(* Injectivity *) -(* ------------------------------------------------------------------------*) - -val dnat_inject = - [ -prove_goal Dnat.thy -"[|x1~=UU; y1~=UU; dsucc`x1 = dsucc`y1 |] ==> x1= y1" - (fn prems => - [ - (cut_facts_tac prems 1), - (dres_inst_tac [("f","dnat_when`c`(LAM x.x)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]) - ]; - -(* ------------------------------------------------------------------------*) -(* definedness for discriminators and selectors *) -(* ------------------------------------------------------------------------*) - - -fun prover thm = prove_goal Dnat.thy thm - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac dnatE 1), - (contr_tac 1), - (REPEAT (asm_simp_tac (!simpset addsimps dnat_rews) 1)) - ]); - -val dnat_discsel_def = - [ - prover "n~=UU ==> is_dzero`n ~= UU", - prover "n~=UU ==> is_dsucc`n ~= UU" - ]; - -val dnat_rews = dnat_discsel_def @ dnat_rews; - - -(* ------------------------------------------------------------------------*) -(* Properties dnat_take *) -(* ------------------------------------------------------------------------*) -val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take n`UU = UU" - (fn prems => - [ - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_take = [temp]; - -val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take 0`xs = UU" - (fn prems => - [ - (Asm_simp_tac 1) - ]); - -val dnat_take = temp::dnat_take; - -val temp = prove_goalw Dnat.thy [dnat_take_def] - "dnat_take (Suc n)`dzero=dzero" - (fn prems => - [ - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_take = temp::dnat_take; - -val temp = prove_goalw Dnat.thy [dnat_take_def] - "dnat_take (Suc n)`(dsucc`xs)=dsucc`(dnat_take n ` xs)" - (fn prems => - [ - (res_inst_tac [("Q","xs=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -val dnat_take = temp::dnat_take; - -val dnat_rews = dnat_take @ dnat_rews; - - -(* ------------------------------------------------------------------------*) -(* take lemma for dnats *) -(* ------------------------------------------------------------------------*) - -fun prover reach defs thm = prove_goalw Dnat.thy defs thm - (fn prems => - [ - (res_inst_tac [("t","s1")] (reach RS subst) 1), - (res_inst_tac [("t","s2")] (reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac lub_equal 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); - -val dnat_take_lemma = prover dnat_reach [dnat_take_def] - "(!!n.dnat_take n`s1 = dnat_take n`s2) ==> s1=s2"; - - -(* ------------------------------------------------------------------------*) -(* Co -induction for dnats *) -(* ------------------------------------------------------------------------*) - -qed_goalw "dnat_coind_lemma" Dnat.thy [dnat_bisim_def] -"dnat_bisim R ==> ! p q. R p q --> dnat_take n`p = dnat_take n`q" - (fn prems => - [ - (cut_facts_tac prems 1), - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dnat_take) 1), - (strip_tac 1), - ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), - (atac 1), - (asm_simp_tac (!simpset addsimps dnat_take) 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dnat_take) 1), - (etac exE 1), - (etac exE 1), - (asm_simp_tac (!simpset addsimps dnat_take) 1), - (REPEAT (etac conjE 1)), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); - -qed_goal "dnat_coind" Dnat.thy "[|dnat_bisim R;R p q|] ==> p = q" - (fn prems => - [ - (rtac dnat_take_lemma 1), - (rtac (dnat_coind_lemma RS spec RS spec RS mp) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); - - -(* ------------------------------------------------------------------------*) -(* structural induction for admissible predicates *) -(* ------------------------------------------------------------------------*) - -(* not needed any longer -qed_goal "dnat_ind" Dnat.thy -"[| adm(P);\ -\ P(UU);\ -\ P(dzero);\ -\ !! s1.[|s1~=UU ; P(s1)|] ==> P(dsucc`s1)|] ==> P(s)" - (fn prems => - [ - (rtac (dnat_reach RS subst) 1), - (res_inst_tac [("x","s")] spec 1), - (rtac fix_ind 1), - (rtac adm_all2 1), - (rtac adm_subst 1), - (cont_tacR 1), - (resolve_tac prems 1), - (Simp_tac 1), - (resolve_tac prems 1), - (strip_tac 1), - (res_inst_tac [("n","xa")] dnatE 1), - (asm_simp_tac (!simpset addsimps dnat_copy) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_copy) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_copy) 1), - (res_inst_tac [("Q","x`xb=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (eresolve_tac prems 1), - (etac spec 1) - ]); -*) - -qed_goal "dnat_finite_ind" Dnat.thy -"[|P(UU);P(dzero);\ -\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ -\ |] ==> !s.P(dnat_take n`s)" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (rtac allI 1), - (res_inst_tac [("n","s")] dnatE 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("Q","dnat_take n1`x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (atac 1), - (etac spec 1) - ]); - -qed_goal "dnat_all_finite_lemma1" Dnat.thy -"!s.dnat_take n`s=UU |dnat_take n`s=s" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dnat_rews) 1), - (rtac allI 1), - (res_inst_tac [("n","s")] dnatE 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (eres_inst_tac [("x","x")] allE 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); - -qed_goal "dnat_all_finite_lemma2" Dnat.thy "? n.dnat_take n`s=s" - (fn prems => - [ - (res_inst_tac [("Q","s=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (subgoal_tac "(!n.dnat_take(n)`s=UU) |(? n.dnat_take(n)`s=s)" 1), - (etac disjE 1), - (eres_inst_tac [("P","s=UU")] notE 1), - (rtac dnat_take_lemma 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (atac 1), - (subgoal_tac "!n.!s.dnat_take(n)`s=UU |dnat_take(n)`s=s" 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (rtac dnat_all_finite_lemma1 1) - ]); - - -qed_goal "dnat_ind" Dnat.thy -"[|P(UU);P(dzero);\ -\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ -\ |] ==> P(s)" - (fn prems => - [ - (rtac (dnat_all_finite_lemma2 RS exE) 1), - (etac subst 1), - (rtac (dnat_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); - - -qed_goalw "dnat_flat" Dnat.thy [flat_def] "flat(dzero)" - (fn prems => - [ - (rtac allI 1), - (res_inst_tac [("s","x")] dnat_ind 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (res_inst_tac [("n","y")] dnatE 1), - (fast_tac (HOL_cs addSIs [UU_I]) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_dist_less) 1), - (rtac allI 1), - (res_inst_tac [("n","y")] dnatE 1), - (fast_tac (HOL_cs addSIs [UU_I]) 1), - (asm_simp_tac (!simpset addsimps dnat_dist_less) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (strip_tac 1), - (subgoal_tac "s1< - -For details see chapter 5 of: - -[Franz Regensburger] HOLCF: Eine konservative Erweiterung von HOL um LCF, - Dissertation, Technische Universit"at M"unchen, 1994 - -*) - -Dnat = HOLCF + - -types dnat 0 - -(* ----------------------------------------------------------------------- *) -(* arrity axiom is valuated by semantical reasoning *) - -arities dnat::pcpo - -consts - -(* ----------------------------------------------------------------------- *) -(* essential constants *) - -dnat_rep :: " dnat -> (one ++ dnat)" -dnat_abs :: "(one ++ dnat) -> dnat" - -(* ----------------------------------------------------------------------- *) -(* abstract constants and auxiliary constants *) - -dnat_copy :: "(dnat -> dnat) -> dnat -> dnat" - -dzero :: "dnat" -dsucc :: "dnat -> dnat" -dnat_when :: "'b -> (dnat -> 'b) -> dnat -> 'b" -is_dzero :: "dnat -> tr" -is_dsucc :: "dnat -> tr" -dpred :: "dnat -> dnat" -dnat_take :: "nat => dnat -> dnat" -dnat_bisim :: "(dnat => dnat => bool) => bool" - -rules - -(* ----------------------------------------------------------------------- *) -(* axiomatization of recursive type dnat *) -(* ----------------------------------------------------------------------- *) -(* (dnat,dnat_abs) is the initial F-algebra where *) -(* F is the locally continuous functor determined by functor term FT. *) -(* domain equation: dnat = one ++ dnat *) -(* functor term: FT = <++,K_{one},I> *) -(* ----------------------------------------------------------------------- *) -(* dnat_abs is an isomorphism with inverse dnat_rep *) -(* identity is the least endomorphism on dnat *) - -dnat_abs_iso "dnat_rep`(dnat_abs`x) = x" -dnat_rep_iso "dnat_abs`(dnat_rep`x) = x" -dnat_copy_def "dnat_copy == (LAM f. dnat_abs oo - (sswhen`sinl`(sinr oo f)) oo dnat_rep )" -dnat_reach "(fix`dnat_copy)`x=x" - - -defs -(* ----------------------------------------------------------------------- *) -(* properties of additional constants *) -(* ----------------------------------------------------------------------- *) -(* constructors *) - -dzero_def "dzero == dnat_abs`(sinl`one)" -dsucc_def "dsucc == (LAM n. dnat_abs`(sinr`n))" - -(* ----------------------------------------------------------------------- *) -(* discriminator functional *) - -dnat_when_def "dnat_when == (LAM f1 f2 n.sswhen`(LAM x.f1)`f2`(dnat_rep`n))" - - -(* ----------------------------------------------------------------------- *) -(* discriminators and selectors *) - -is_dzero_def "is_dzero == dnat_when`TT`(LAM x.FF)" -is_dsucc_def "is_dsucc == dnat_when`FF`(LAM x.TT)" -dpred_def "dpred == dnat_when`UU`(LAM x.x)" - - -(* ----------------------------------------------------------------------- *) -(* the taker for dnats *) - -dnat_take_def "dnat_take == (%n.iterate n dnat_copy UU)" - -(* ----------------------------------------------------------------------- *) -(* definition of bisimulation is determined by domain equation *) -(* simplification and rewriting for abstract constants yields def below *) - -dnat_bisim_def "dnat_bisim == -(%R.!s1 s2. - R s1 s2 --> - ((s1=UU & s2=UU) |(s1=dzero & s2=dzero) | - (? s11 s21. s11~=UU & s21~=UU & s1=dsucc`s11 & - s2 = dsucc`s21 & R s11 s21)))" - -end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Dnat2.ML --- a/src/HOLCF/Dnat2.ML Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,52 +0,0 @@ -(* Title: HOLCF/dnat2.ML - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Lemmas for theory Dnat2.thy -*) - -open Dnat2; - - -(* ------------------------------------------------------------------------- *) -(* expand fixed point properties *) -(* ------------------------------------------------------------------------- *) - -val iterator_def2 = fix_prover2 Dnat2.thy iterator_def - "iterator = (LAM n f x. dnat_when`x`(LAM m.f`(iterator`m`f`x)) `n)"; - -(* ------------------------------------------------------------------------- *) -(* recursive properties *) -(* ------------------------------------------------------------------------- *) - -qed_goal "iterator1" Dnat2.thy "iterator`UU`f`x = UU" - (fn prems => - [ - (rtac (iterator_def2 RS ssubst) 1), - (simp_tac (!simpset addsimps dnat_when) 1) - ]); - -qed_goal "iterator2" Dnat2.thy "iterator`dzero`f`x = x" - (fn prems => - [ - (rtac (iterator_def2 RS ssubst) 1), - (simp_tac (!simpset addsimps dnat_when) 1) - ]); - -qed_goal "iterator3" Dnat2.thy -"n~=UU ==> iterator`(dsucc`n)`f`x = f`(iterator`n`f`x)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (rtac (iterator_def2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (rtac refl 1) - ]); - -val dnat2_rews = - [iterator1, iterator2, iterator3]; - - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Dnat2.thy --- a/src/HOLCF/Dnat2.thy Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,29 +0,0 @@ -(* Title: HOLCF/dnat2.thy - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Additional constants for dnat - -*) - -Dnat2 = Dnat + - -consts - -iterator :: "dnat -> ('a -> 'a) -> 'a -> 'a" - - -defs - -iterator_def "iterator == fix`(LAM h n f x. - dnat_when `x `(LAM m.f`(h`m`f`x)) `n)" -end - -(* - - iterator`UU`f`x = UU - iterator`dzero`f`x = x - n~=UU --> iterator`(dsucc`n)`f`x = f`(iterator`n`f`x) -*) - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Fix.ML --- a/src/HOLCF/Fix.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Fix.ML Fri Oct 06 17:25:24 1995 +0100 @@ -341,6 +341,19 @@ ]); +qed_goal "fix_eqI" Fix.thy +"[| F`x = x; !z. F`z = z --> x << z |] ==> x = fix`F" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac antisym_less 1), + (etac allE 1), + (etac mp 1), + (rtac (fix_eq RS sym) 1), + (etac fix_least 1) + ]); + + qed_goal "fix_eq2" Fix.thy "f == fix`F ==> f = F`f" (fn prems => [ diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/HOLCF.ML --- a/src/HOLCF/HOLCF.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/HOLCF.ML Fri Oct 06 17:25:24 1995 +0100 @@ -8,3 +8,5 @@ Addsimps (one_when @ dist_less_one @ dist_eq_one @ dist_less_tr @ dist_eq_tr @ tr_when @ andalso_thms @ orelse_thms @ neg_thms @ ifte_thms); + +val HOLCF_ss = !simpset; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/HOLCF.thy --- a/src/HOLCF/HOLCF.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/HOLCF.thy Fri Oct 06 17:25:24 1995 +0100 @@ -8,6 +8,5 @@ *) -HOLCF = Tr2 +HOLCF = Tr2 - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Holcfb.thy --- a/src/HOLCF/Holcfb.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Holcfb.thy Fri Oct 06 17:25:24 1995 +0100 @@ -10,16 +10,12 @@ Holcfb = Nat + consts - -theleast :: "(nat=>bool)=>nat" - + theleast :: "(nat=>bool)=>nat" defs theleast_def "theleast P == (@z.(P z & (!n.P n --> z<=n)))" -end - +(* start 8bit 1 *) +(* end 8bit 1 *) - - - +end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Lift3.ML --- a/src/HOLCF/Lift3.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Lift3.ML Fri Oct 06 17:25:24 1995 +0100 @@ -347,3 +347,4 @@ (* ------------------------------------------------------------------------ *) val lift_rews = [lift1,lift2,defined_up]; + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Lift3.thy --- a/src/HOLCF/Lift3.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Lift3.thy Fri Oct 06 17:25:24 1995 +0100 @@ -24,6 +24,12 @@ up_def "up == (LAM x.Iup(x))" lift_def "lift == (LAM f p.Ilift(f)(p))" +translations +"case l of up`x => t1" == "lift`(LAM x.t1)`l" + +(* start 8bit 1 *) +(* end 8bit 1 *) + end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Makefile --- a/src/HOLCF/Makefile Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Makefile Fri Oct 06 17:25:24 1995 +0100 @@ -23,8 +23,7 @@ Sprod0.thy Sprod1.thy Sprod2.thy Sprod3.thy \ Ssum0.thy Ssum1.thy Ssum2.thy Ssum3.thy \ Lift1.thy Lift2.thy Lift3.thy Fix.thy ccc1.thy One.thy \ - Tr1.thy Tr2.thy HOLCF.thy Dnat.thy Dnat2.thy \ - Stream.thy Stream2.thy Dlist.thy + Tr1.thy Tr2.thy HOLCF.thy FILES = ROOT.ML Porder0.thy $(THYS) $(THYS:.thy=.ML) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/One.thy --- a/src/HOLCF/One.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/One.thy Fri Oct 06 17:25:24 1995 +0100 @@ -35,6 +35,10 @@ defs one_def "one == abs_one`(up`UU)" one_when_def "one_when == (LAM c u.lift`(LAM x.c)`(rep_one`u))" + +translations + "case l of one => t1" == "one_when`t1`l" + end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Pcpo.thy --- a/src/HOLCF/Pcpo.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Pcpo.thy Fri Oct 06 17:25:24 1995 +0100 @@ -5,11 +5,14 @@ consts UU :: "'a::pcpo" + rules -minimal "UU << x" -cpo "is_chain(S) ==> ? x. range(S) <<| (x::'a::pcpo)" + minimal "UU << x" + cpo "is_chain(S) ==> ? x. range(S) <<| (x::'a::pcpo)" inst_void_pcpo "(UU::void) = UU_void" +(* start 8bit 1 *) +(* end 8bit 1 *) end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Porder.thy --- a/src/HOLCF/Porder.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Porder.thy Fri Oct 06 17:25:24 1995 +0100 @@ -42,6 +42,9 @@ lub "lub(S) = (@x. S <<| x)" +(* start 8bit 1 *) +(* end 8bit 1 *) + end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Porder0.thy --- a/src/HOLCF/Porder0.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Porder0.thy Fri Oct 06 17:25:24 1995 +0100 @@ -39,4 +39,12 @@ inst_void_po "((op <<)::[void,void]=>bool) = less_void" +(* start 8bit 1 *) +(* end 8bit 1 *) + end + + + + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/README --- a/src/HOLCF/README Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/README Fri Oct 06 17:25:24 1995 +0100 @@ -2,10 +2,10 @@ ========================================================== Author: Franz Regensburger -Copyright 1993,1994 Technische Universitaet Muenchen +Copyright 1995 Technische Universitaet Muenchen -Version: 1.5 -Date: 14.10.94 +Version: 2.0 +Date: 16.08.95 A detailed description of the entire development can be found in @@ -19,3 +19,6 @@ 28.06.95 The old uncurried version of HOLCF is no longer supported in the distribution. +18.08.95 added sections axioms, ops, domain, genertated + and 8bit support + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ROOT.ML --- a/src/HOLCF/ROOT.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/ROOT.ML Fri Oct 06 17:25:24 1995 +0100 @@ -7,16 +7,37 @@ Should be executed in subdirectory HOLCF. *) -val banner = "Higher-order Logic of Computable Functions; curried version"; +val banner = "HOLCF with sections axioms,ops,domain,generated"; +init_thy_reader(); + +(* start 8bit 1 *) +(* end 8bit 1 *) + writeln banner; print_depth 1; -init_thy_reader(); +use_thy "HOLCF"; + +(* install sections: axioms, ops *) + +use "ax_ops/holcflogic.ML"; +use "ax_ops/thy_axioms.ML"; +use "ax_ops/thy_ops.ML"; +use "ax_ops/thy_syntax.ML"; + + +(* install sections: domain, generated *) -use_thy "Fix"; -use_thy "Dlist"; +use "domain/library"; +use "domain/syntax"; +use "domain/axioms"; +use "domain/theorems"; +use "domain/extender"; +use "domain/interface"; -use "../Pure/install_pp.ML"; -print_depth 8; +init_thy_reader(); +init_pps (); + +print_depth 100; val HOLCF_build_completed = (); (*indicate successful build*) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Sprod0.ML --- a/src/HOLCF/Sprod0.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Sprod0.ML Fri Oct 06 17:25:24 1995 +0100 @@ -343,7 +343,6 @@ Addsimps [strict_Isfst1,strict_Isfst2,strict_Issnd1,strict_Issnd2, Isfst2,Issnd2]; - qed_goal "defined_IsfstIssnd" Sprod0.thy "p~=Ispair UU UU ==> Isfst p ~= UU & Issnd p ~= UU" (fn prems => diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Sprod0.thy --- a/src/HOLCF/Sprod0.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Sprod0.thy Fri Oct 06 17:25:24 1995 +0100 @@ -50,5 +50,8 @@ (p=Ispair UU UU --> z=UU) &(! a b. ~a=UU & ~b=UU & p=Ispair a b --> z=b)" +(* start 8bit 1 *) +(* end 8bit 1 *) + end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Sprod3.ML --- a/src/HOLCF/Sprod3.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Sprod3.ML Fri Oct 06 17:25:24 1995 +0100 @@ -675,8 +675,10 @@ (* install simplifier for Sprod *) (* ------------------------------------------------------------------------ *) +val Sprod_rews = [strict_spair1,strict_spair2,strict_sfst1,strict_sfst2, + strict_ssnd1,strict_ssnd2,sfst2,ssnd2,defined_spair, + ssplit1,ssplit2]; + Addsimps [strict_spair1,strict_spair2,strict_sfst1,strict_sfst2, strict_ssnd1,strict_ssnd2,sfst2,ssnd2, ssplit1,ssplit2]; - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Sprod3.thy --- a/src/HOLCF/Sprod3.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Sprod3.thy Fri Oct 06 17:25:24 1995 +0100 @@ -33,6 +33,9 @@ ssnd_def "ssnd == (LAM p.Issnd p)" ssplit_def "ssplit == (LAM f. strictify`(LAM p.f`(sfst`p)`(ssnd`p)))" +(* start 8bit 1 *) +(* end 8bit 1 *) + end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Ssum0.ML --- a/src/HOLCF/Ssum0.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Ssum0.ML Fri Oct 06 17:25:24 1995 +0100 @@ -389,5 +389,3 @@ (* ------------------------------------------------------------------------ *) Addsimps [(strict_IsinlIsinr RS sym),Iwhen1,Iwhen2,Iwhen3]; - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Ssum0.thy --- a/src/HOLCF/Ssum0.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Ssum0.thy Fri Oct 06 17:25:24 1995 +0100 @@ -51,5 +51,7 @@ &(!a. a~=UU & s=Isinl(a) --> z=f`a) &(!b. b~=UU & s=Isinr(b) --> z=g`b)" +(* start 8bit 1 *) +(* end 8bit 1 *) end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Ssum3.ML --- a/src/HOLCF/Ssum3.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Ssum3.ML Fri Oct 06 17:25:24 1995 +0100 @@ -724,4 +724,8 @@ (* install simplifier for Ssum *) (* ------------------------------------------------------------------------ *) -Addsimps [strict_sinl,strict_sinr,sswhen1,sswhen2,sswhen3]; +val Ssum_rews = [strict_sinl,strict_sinr,defined_sinl,defined_sinr, + sswhen1,sswhen2,sswhen3]; + +Addsimps [strict_sinl,strict_sinr,defined_sinl,defined_sinr, + sswhen1,sswhen2,sswhen3]; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Ssum3.thy --- a/src/HOLCF/Ssum3.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Ssum3.thy Fri Oct 06 17:25:24 1995 +0100 @@ -26,4 +26,10 @@ sinr_def "sinr == (LAM x.Isinr(x))" sswhen_def "sswhen == (LAM f g s.Iwhen(f)(g)(s))" +translations +"case s of sinl`x => t1 | sinr`y => t2" == "sswhen`(LAM x.t1)`(LAM y.t2)`s" + +(* start 8bit 1 *) +(* end 8bit 1 *) + end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Stream.ML --- a/src/HOLCF/Stream.ML Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,840 +0,0 @@ -(* Title: HOLCF/stream.ML - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Lemmas for stream.thy -*) - -open Stream; - -(* ------------------------------------------------------------------------*) -(* The isomorphisms stream_rep_iso and stream_abs_iso are strict *) -(* ------------------------------------------------------------------------*) - -val stream_iso_strict= stream_rep_iso RS (stream_abs_iso RS - (allI RSN (2,allI RS iso_strict))); - -val stream_rews = [stream_iso_strict RS conjunct1, - stream_iso_strict RS conjunct2]; - -(* ------------------------------------------------------------------------*) -(* Properties of stream_copy *) -(* ------------------------------------------------------------------------*) - -fun prover defs thm = prove_goalw Stream.thy defs thm - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) - ]); - -val stream_copy = - [ - prover [stream_copy_def] "stream_copy`f`UU=UU", - prover [stream_copy_def,scons_def] - "x~=UU ==> stream_copy`f`(scons`x`xs)= scons`x`(f`xs)" - ]; - -val stream_rews = stream_copy @ stream_rews; - -(* ------------------------------------------------------------------------*) -(* Exhaustion and elimination for streams *) -(* ------------------------------------------------------------------------*) - -qed_goalw "Exh_stream" Stream.thy [scons_def] - "s = UU | (? x xs. x~=UU & s = scons`x`xs)" - (fn prems => - [ - (Simp_tac 1), - (rtac (stream_rep_iso RS subst) 1), - (res_inst_tac [("p","stream_rep`s")] sprodE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (Asm_simp_tac 1), - (res_inst_tac [("p","y")] liftE1 1), - (contr_tac 1), - (rtac disjI2 1), - (rtac exI 1), - (rtac exI 1), - (etac conjI 1), - (Asm_simp_tac 1) - ]); - -qed_goal "streamE" Stream.thy - "[| s=UU ==> Q; !!x xs.[|s=scons`x`xs;x~=UU|]==>Q|]==>Q" - (fn prems => - [ - (rtac (Exh_stream RS disjE) 1), - (eresolve_tac prems 1), - (etac exE 1), - (etac exE 1), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); - -(* ------------------------------------------------------------------------*) -(* Properties of stream_when *) -(* ------------------------------------------------------------------------*) - -fun prover defs thm = prove_goalw Stream.thy defs thm - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) - ]); - - -val stream_when = [ - prover [stream_when_def] "stream_when`f`UU=UU", - prover [stream_when_def,scons_def] - "x~=UU ==> stream_when`f`(scons`x`xs)= f`x`xs" - ]; - -val stream_rews = stream_when @ stream_rews; - -(* ------------------------------------------------------------------------*) -(* Rewrites for discriminators and selectors *) -(* ------------------------------------------------------------------------*) - -fun prover defs thm = prove_goalw Stream.thy defs thm - (fn prems => - [ - (simp_tac (!simpset addsimps stream_rews) 1) - ]); - -val stream_discsel = [ - prover [is_scons_def] "is_scons`UU=UU", - prover [shd_def] "shd`UU=UU", - prover [stl_def] "stl`UU=UU" - ]; - -fun prover defs thm = prove_goalw Stream.thy defs thm - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - -val stream_discsel = [ -prover [is_scons_def,shd_def,stl_def] "x~=UU ==> is_scons`(scons`x`xs)=TT", -prover [is_scons_def,shd_def,stl_def] "x~=UU ==> shd`(scons`x`xs)=x", -prover [is_scons_def,shd_def,stl_def] "x~=UU ==> stl`(scons`x`xs)=xs" - ] @ stream_discsel; - -val stream_rews = stream_discsel @ stream_rews; - -(* ------------------------------------------------------------------------*) -(* Definedness and strictness *) -(* ------------------------------------------------------------------------*) - -fun prover contr thm = prove_goal Stream.thy thm - (fn prems => - [ - (res_inst_tac [("P1",contr)] classical3 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (dtac sym 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps (prems @ stream_rews)) 1) - ]); - -val stream_constrdef = [ - prover "is_scons`(UU::'a stream)~=UU" "x~=UU ==> scons`(x::'a)`xs~=UU" - ]; - -fun prover defs thm = prove_goalw Stream.thy defs thm - (fn prems => - [ - (simp_tac (!simpset addsimps stream_rews) 1) - ]); - -val stream_constrdef = [ - prover [scons_def] "scons`UU`xs=UU" - ] @ stream_constrdef; - -val stream_rews = stream_constrdef @ stream_rews; - - -(* ------------------------------------------------------------------------*) -(* Distinctness wrt. << and = *) -(* ------------------------------------------------------------------------*) - - -(* ------------------------------------------------------------------------*) -(* Invertibility *) -(* ------------------------------------------------------------------------*) - -val stream_invert = - [ -prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ -\ scons`x1`x2 << scons`y1`y2|] ==> x1<< y1 & x2 << y2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("fo5","stream_when`(LAM x l.x)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (dres_inst_tac [("fo5","stream_when`(LAM x l.l)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1) - ]) - ]; - -(* ------------------------------------------------------------------------*) -(* Injectivity *) -(* ------------------------------------------------------------------------*) - -val stream_inject = - [ -prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ -\ scons`x1`x2 = scons`y1`y2 |] ==> x1= y1 & x2 = y2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("f","stream_when`(LAM x l.x)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (dres_inst_tac [("f","stream_when`(LAM x l.l)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1) - ]) - ]; - -(* ------------------------------------------------------------------------*) -(* definedness for discriminators and selectors *) -(* ------------------------------------------------------------------------*) - -fun prover thm = prove_goal Stream.thy thm - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac streamE 1), - (contr_tac 1), - (REPEAT (asm_simp_tac (!simpset addsimps stream_discsel) 1)) - ]); - -val stream_discsel_def = - [ - prover "s~=UU ==> is_scons`s ~= UU", - prover "s~=UU ==> shd`s ~=UU" - ]; - -val stream_rews = stream_discsel_def @ stream_rews; - - -(* ------------------------------------------------------------------------*) -(* Properties stream_take *) -(* ------------------------------------------------------------------------*) - -val stream_take = - [ -prove_goalw Stream.thy [stream_take_def] "stream_take n`UU = UU" - (fn prems => - [ - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1) - ]), -prove_goalw Stream.thy [stream_take_def] "stream_take 0`xs=UU" - (fn prems => - [ - (Asm_simp_tac 1) - ])]; - -fun prover thm = prove_goalw Stream.thy [stream_take_def] thm - (fn prems => - [ - (cut_facts_tac prems 1), - (Simp_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - -val stream_take = [ -prover - "x~=UU ==> stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" - ] @ stream_take; - -val stream_rews = stream_take @ stream_rews; - -(* ------------------------------------------------------------------------*) -(* enhance the simplifier *) -(* ------------------------------------------------------------------------*) - -qed_goal "stream_copy2" Stream.thy - "stream_copy`f`(scons`x`xs) = scons`x`(f`xs)" - (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - -qed_goal "shd2" Stream.thy "shd`(scons`x`xs) = x" - (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - -qed_goal "stream_take2" Stream.thy - "stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" - (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - -val stream_rews = [stream_iso_strict RS conjunct1, - stream_iso_strict RS conjunct2, - hd stream_copy, stream_copy2] - @ stream_when - @ [hd stream_discsel,shd2] @ (tl (tl stream_discsel)) - @ stream_constrdef - @ stream_discsel_def - @ [ stream_take2] @ (tl stream_take); - - -(* ------------------------------------------------------------------------*) -(* take lemma for streams *) -(* ------------------------------------------------------------------------*) - -fun prover reach defs thm = prove_goalw Stream.thy defs thm - (fn prems => - [ - (res_inst_tac [("t","s1")] (reach RS subst) 1), - (res_inst_tac [("t","s2")] (reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac lub_equal 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); - -val stream_take_lemma = prover stream_reach [stream_take_def] - "(!!n.stream_take n`s1 = stream_take n`s2) ==> s1=s2"; - - -qed_goal "stream_reach2" Stream.thy "lub(range(%i.stream_take i`s))=s" - (fn prems => - [ - (res_inst_tac [("t","s")] (stream_reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rewrite_goals_tac [stream_take_def]), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac refl 1) - ]); - -(* ------------------------------------------------------------------------*) -(* Co -induction for streams *) -(* ------------------------------------------------------------------------*) - -qed_goalw "stream_coind_lemma" Stream.thy [stream_bisim_def] -"stream_bisim R ==> ! p q. R p q --> stream_take n`p = stream_take n`q" - (fn prems => - [ - (cut_facts_tac prems 1), - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1), - ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), - (atac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (etac exE 1), - (etac exE 1), - (etac exE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (REPEAT (etac conjE 1)), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); - -qed_goal "stream_coind" Stream.thy "[|stream_bisim R ;R p q|] ==> p = q" - (fn prems => - [ - (rtac stream_take_lemma 1), - (rtac (stream_coind_lemma RS spec RS spec RS mp) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); - -(* ------------------------------------------------------------------------*) -(* structural induction for admissible predicates *) -(* ------------------------------------------------------------------------*) - -qed_goal "stream_finite_ind" Stream.thy -"[|P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ -\ |] ==> !s.P(stream_take n`s)" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (resolve_tac prems 1), - (rtac allI 1), - (res_inst_tac [("s","s")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (resolve_tac prems 1), - (atac 1), - (etac spec 1) - ]); - -qed_goalw "stream_finite_ind2" Stream.thy [stream_finite_def] -"(!!n.P(stream_take n`s)) ==> stream_finite(s) -->P(s)" - (fn prems => - [ - (strip_tac 1), - (etac exE 1), - (etac subst 1), - (resolve_tac prems 1) - ]); - -qed_goal "stream_finite_ind3" Stream.thy -"[|P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ -\ |] ==> stream_finite(s) --> P(s)" - (fn prems => - [ - (rtac stream_finite_ind2 1), - (rtac (stream_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); - -(* prove induction using definition of admissibility - stream_reach rsp. stream_reach2 - and finite induction stream_finite_ind *) - -qed_goal "stream_ind" Stream.thy -"[|adm(P);\ -\ P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ -\ |] ==> P(s)" - (fn prems => - [ - (rtac (stream_reach2 RS subst) 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (resolve_tac prems 1), - (SELECT_GOAL (rewrite_goals_tac [stream_take_def]) 1), - (rtac ch2ch_fappL 1), - (rtac is_chain_iterate 1), - (rtac allI 1), - (rtac (stream_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); - -(* prove induction with usual LCF-Method using fixed point induction *) -qed_goal "stream_ind" Stream.thy -"[|adm(P);\ -\ P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ -\ |] ==> P(s)" - (fn prems => - [ - (rtac (stream_reach RS subst) 1), - (res_inst_tac [("x","s")] spec 1), - (rtac wfix_ind 1), - (rtac adm_impl_admw 1), - (REPEAT (resolve_tac adm_thms 1)), - (rtac adm_subst 1), - (cont_tacR 1), - (resolve_tac prems 1), - (rtac allI 1), - (rtac (rewrite_rule [stream_take_def] stream_finite_ind) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); - - -(* ------------------------------------------------------------------------*) -(* simplify use of Co-induction *) -(* ------------------------------------------------------------------------*) - -qed_goal "surjectiv_scons" Stream.thy "scons`(shd`s)`(stl`s)=s" - (fn prems => - [ - (res_inst_tac [("s","s")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - - -qed_goalw "stream_coind_lemma2" Stream.thy [stream_bisim_def] -"!s1 s2. R s1 s2 --> shd`s1 = shd`s2 & R (stl`s1) (stl`s2) ==> stream_bisim R" - (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (etac allE 1), - (etac allE 1), - (dtac mp 1), - (atac 1), - (etac conjE 1), - (res_inst_tac [("Q","s1 = UU & s2 = UU")] classical2 1), - (rtac disjI1 1), - (fast_tac HOL_cs 1), - (rtac disjI2 1), - (rtac disjE 1), - (etac (de_morgan2 RS ssubst) 1), - (res_inst_tac [("x","shd`s1")] exI 1), - (res_inst_tac [("x","stl`s1")] exI 1), - (res_inst_tac [("x","stl`s2")] exI 1), - (rtac conjI 1), - (eresolve_tac stream_discsel_def 1), - (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), - (eres_inst_tac [("s","shd`s1"),("t","shd`s2")] subst 1), - (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), - (res_inst_tac [("x","shd`s2")] exI 1), - (res_inst_tac [("x","stl`s1")] exI 1), - (res_inst_tac [("x","stl`s2")] exI 1), - (rtac conjI 1), - (eresolve_tac stream_discsel_def 1), - (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), - (res_inst_tac [("s","shd`s1"),("t","shd`s2")] ssubst 1), - (etac sym 1), - (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1) - ]); - - -(* ------------------------------------------------------------------------*) -(* theorems about finite and infinite streams *) -(* ------------------------------------------------------------------------*) - -(* ----------------------------------------------------------------------- *) -(* 2 lemmas about stream_finite *) -(* ----------------------------------------------------------------------- *) - -qed_goalw "stream_finite_UU" Stream.thy [stream_finite_def] - "stream_finite(UU)" - (fn prems => - [ - (rtac exI 1), - (simp_tac (!simpset addsimps stream_rews) 1) - ]); - -qed_goal "inf_stream_not_UU" Stream.thy "~stream_finite(s) ==> s ~= UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac swap 1), - (dtac notnotD 1), - (hyp_subst_tac 1), - (rtac stream_finite_UU 1) - ]); - -(* ----------------------------------------------------------------------- *) -(* a lemma about shd *) -(* ----------------------------------------------------------------------- *) - -qed_goal "stream_shd_lemma1" Stream.thy "shd`s=UU --> s=UU" - (fn prems => - [ - (res_inst_tac [("s","s")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (hyp_subst_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - - -(* ----------------------------------------------------------------------- *) -(* lemmas about stream_take *) -(* ----------------------------------------------------------------------- *) - -qed_goal "stream_take_lemma1" Stream.thy - "!x xs.x~=UU --> \ -\ stream_take (Suc n)`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" - (fn prems => - [ - (rtac allI 1), - (rtac allI 1), - (rtac impI 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1), - (rtac ((hd stream_inject) RS conjunct2) 1), - (atac 1), - (atac 1), - (atac 1) - ]); - - -qed_goal "stream_take_lemma2" Stream.thy - "! s2. stream_take n`s2 = s2 --> stream_take (Suc n)`s2=s2" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (hyp_subst_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (rtac allI 1), - (res_inst_tac [("s","s2")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (subgoal_tac "stream_take n1`xs = xs" 1), - (rtac ((hd stream_inject) RS conjunct2) 2), - (atac 4), - (atac 2), - (atac 2), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); - -qed_goal "stream_take_lemma3" Stream.thy - "!x xs.x~=UU --> \ -\ stream_take n`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" - (fn prems => - [ - (nat_ind_tac "n" 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (res_inst_tac [("P","scons`x`xs=UU")] notE 1), - (eresolve_tac stream_constrdef 1), - (etac sym 1), - (strip_tac 1 ), - (rtac (stream_take_lemma2 RS spec RS mp) 1), - (res_inst_tac [("x1.1","x")] ((hd stream_inject) RS conjunct2) 1), - (atac 1), - (atac 1), - (etac (stream_take2 RS subst) 1) - ]); - -qed_goal "stream_take_lemma4" Stream.thy - "!x xs.\ -\stream_take n`xs=xs --> stream_take (Suc n)`(scons`x`xs) = scons`x`xs" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (simp_tac (!simpset addsimps stream_rews) 1) - ]); - -(* ---- *) - -qed_goal "stream_take_lemma5" Stream.thy -"!s. stream_take n`s=s --> iterate n stl s=UU" - (fn prems => - [ - (nat_ind_tac "n" 1), - (Simp_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1), - (res_inst_tac [("s","s")] streamE 1), - (hyp_subst_tac 1), - (rtac (iterate_Suc2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac (iterate_Suc2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (etac allE 1), - (etac mp 1), - (hyp_subst_tac 1), - (etac (stream_take_lemma1 RS spec RS spec RS mp RS mp) 1), - (atac 1) - ]); - -qed_goal "stream_take_lemma6" Stream.thy -"!s.iterate n stl s =UU --> stream_take n`s=s" - (fn prems => - [ - (nat_ind_tac "n" 1), - (Simp_tac 1), - (strip_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac allI 1), - (res_inst_tac [("s","s")] streamE 1), - (hyp_subst_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (hyp_subst_tac 1), - (rtac (iterate_Suc2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); - -qed_goal "stream_take_lemma7" Stream.thy -"(iterate n stl s=UU) = (stream_take n`s=s)" - (fn prems => - [ - (rtac iffI 1), - (etac (stream_take_lemma6 RS spec RS mp) 1), - (etac (stream_take_lemma5 RS spec RS mp) 1) - ]); - - -qed_goal "stream_take_lemma8" Stream.thy -"[|adm(P); !n. ? m. n < m & P (stream_take m`s)|] ==> P(s)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (stream_reach2 RS subst) 1), - (rtac adm_disj_lemma11 1), - (atac 1), - (atac 2), - (rewrite_goals_tac [stream_take_def]), - (rtac ch2ch_fappL 1), - (rtac is_chain_iterate 1) - ]); - -(* ----------------------------------------------------------------------- *) -(* lemmas stream_finite *) -(* ----------------------------------------------------------------------- *) - -qed_goalw "stream_finite_lemma1" Stream.thy [stream_finite_def] - "stream_finite(xs) ==> stream_finite(scons`x`xs)" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (rtac exI 1), - (etac (stream_take_lemma4 RS spec RS spec RS mp) 1) - ]); - -qed_goalw "stream_finite_lemma2" Stream.thy [stream_finite_def] - "[|x~=UU; stream_finite(scons`x`xs)|] ==> stream_finite(xs)" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (rtac exI 1), - (etac (stream_take_lemma3 RS spec RS spec RS mp RS mp) 1), - (atac 1) - ]); - -qed_goal "stream_finite_lemma3" Stream.thy - "x~=UU ==> stream_finite(scons`x`xs) = stream_finite(xs)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac iffI 1), - (etac stream_finite_lemma2 1), - (atac 1), - (etac stream_finite_lemma1 1) - ]); - - -qed_goalw "stream_finite_lemma5" Stream.thy [stream_finite_def] - "(!n. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1))\ -\=(s1 << s2 --> stream_finite(s2) --> stream_finite(s1))" - (fn prems => - [ - (rtac iffI 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); - -qed_goal "stream_finite_lemma6" Stream.thy - "!s1 s2. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1)" - (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (hyp_subst_tac 1), - (dtac UU_I 1), - (hyp_subst_tac 1), - (rtac stream_finite_UU 1), - (rtac allI 1), - (rtac allI 1), - (res_inst_tac [("s","s1")] streamE 1), - (hyp_subst_tac 1), - (strip_tac 1 ), - (rtac stream_finite_UU 1), - (hyp_subst_tac 1), - (res_inst_tac [("s","s2")] streamE 1), - (hyp_subst_tac 1), - (strip_tac 1 ), - (dtac UU_I 1), - (asm_simp_tac(!simpset addsimps (stream_rews @ [stream_finite_UU])) 1), - (hyp_subst_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (rtac stream_finite_lemma1 1), - (subgoal_tac "xs << xsa" 1), - (subgoal_tac "stream_take n1`xsa = xsa" 1), - (fast_tac HOL_cs 1), - (res_inst_tac [("x1.1","xa"),("y1.1","xa")] - ((hd stream_inject) RS conjunct2) 1), - (atac 1), - (atac 1), - (atac 1), - (res_inst_tac [("x1.1","x"),("y1.1","xa")] - ((hd stream_invert) RS conjunct2) 1), - (atac 1), - (atac 1), - (atac 1) - ]); - -qed_goal "stream_finite_lemma7" Stream.thy -"s1 << s2 --> stream_finite(s2) --> stream_finite(s1)" - (fn prems => - [ - (rtac (stream_finite_lemma5 RS iffD1) 1), - (rtac allI 1), - (rtac (stream_finite_lemma6 RS spec RS spec) 1) - ]); - -qed_goalw "stream_finite_lemma8" Stream.thy [stream_finite_def] -"stream_finite(s) = (? n. iterate n stl s = UU)" - (fn prems => - [ - (simp_tac (!simpset addsimps [stream_take_lemma7]) 1) - ]); - - -(* ----------------------------------------------------------------------- *) -(* admissibility of ~stream_finite *) -(* ----------------------------------------------------------------------- *) - -qed_goalw "adm_not_stream_finite" Stream.thy [adm_def] - "adm(%s. ~ stream_finite(s))" - (fn prems => - [ - (strip_tac 1 ), - (res_inst_tac [("P1","!i. ~ stream_finite(Y(i))")] classical3 1), - (atac 2), - (subgoal_tac "!i.stream_finite(Y(i))" 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (rtac (stream_finite_lemma7 RS mp RS mp) 1), - (etac is_ub_thelub 1), - (atac 1) - ]); - -(* ----------------------------------------------------------------------- *) -(* alternative prove for admissibility of ~stream_finite *) -(* show that stream_finite(s) = (? n. iterate n stl s = UU) *) -(* and prove adm. of ~(? n. iterate n stl s = UU) *) -(* proof uses theorems stream_take_lemma5-7; stream_finite_lemma8 *) -(* ----------------------------------------------------------------------- *) - - -qed_goal "adm_not_stream_finite" Stream.thy "adm(%s. ~ stream_finite(s))" - (fn prems => - [ - (subgoal_tac "(!s.(~stream_finite(s))=(!n.iterate n stl s ~=UU))" 1), - (etac (adm_cong RS iffD2)1), - (REPEAT(resolve_tac adm_thms 1)), - (rtac cont_iterate2 1), - (rtac allI 1), - (rtac (stream_finite_lemma8 RS ssubst) 1), - (fast_tac HOL_cs 1) - ]); - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Stream.thy --- a/src/HOLCF/Stream.thy Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,115 +0,0 @@ -(* Title: HOLCF/stream.thy - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Theory for streams without defined empty stream - 'a stream = 'a ** ('a stream)u - -The type is axiomatized as the least solution of the domain equation above. -The functor term that specifies the domain equation is: - - FT = <**,K_{'a},U> - -For details see chapter 5 of: - -[Franz Regensburger] HOLCF: Eine konservative Erweiterung von HOL um LCF, - Dissertation, Technische Universit"at M"unchen, 1994 -*) - -Stream = Dnat2 + - -types stream 1 - -(* ----------------------------------------------------------------------- *) -(* arity axiom is validated by semantic reasoning *) -(* partial ordering is implicit in the isomorphism axioms and their cont. *) - -arities stream::(pcpo)pcpo - -consts - -(* ----------------------------------------------------------------------- *) -(* essential constants *) - -stream_rep :: "('a stream) -> ('a ** ('a stream)u)" -stream_abs :: "('a ** ('a stream)u) -> ('a stream)" - -(* ----------------------------------------------------------------------- *) -(* abstract constants and auxiliary constants *) - -stream_copy :: "('a stream -> 'a stream) ->'a stream -> 'a stream" - -scons :: "'a -> 'a stream -> 'a stream" -stream_when :: "('a -> 'a stream -> 'b) -> 'a stream -> 'b" -is_scons :: "'a stream -> tr" -shd :: "'a stream -> 'a" -stl :: "'a stream -> 'a stream" -stream_take :: "nat => 'a stream -> 'a stream" -stream_finite :: "'a stream => bool" -stream_bisim :: "('a stream => 'a stream => bool) => bool" - -rules - -(* ----------------------------------------------------------------------- *) -(* axiomatization of recursive type 'a stream *) -(* ----------------------------------------------------------------------- *) -(* ('a stream,stream_abs) is the initial F-algebra where *) -(* F is the locally continuous functor determined by functor term FT. *) -(* domain equation: 'a stream = 'a ** ('a stream)u *) -(* functor term: FT = <**,K_{'a},U> *) -(* ----------------------------------------------------------------------- *) -(* stream_abs is an isomorphism with inverse stream_rep *) -(* identity is the least endomorphism on 'a stream *) - -stream_abs_iso "stream_rep`(stream_abs`x) = x" -stream_rep_iso "stream_abs`(stream_rep`x) = x" -stream_copy_def "stream_copy == (LAM f. stream_abs oo - (ssplit`(LAM x y. (|x , (lift`(up oo f))`y|) )) oo stream_rep)" -stream_reach "(fix`stream_copy)`x = x" - -defs -(* ----------------------------------------------------------------------- *) -(* properties of additional constants *) -(* ----------------------------------------------------------------------- *) -(* constructors *) - -scons_def "scons == (LAM x l. stream_abs`(| x, up`l |))" - -(* ----------------------------------------------------------------------- *) -(* discriminator functional *) - -stream_when_def -"stream_when == (LAM f l.ssplit `(LAM x l.f`x`(lift`ID`l)) `(stream_rep`l))" - -(* ----------------------------------------------------------------------- *) -(* discriminators and selectors *) - -is_scons_def "is_scons == stream_when`(LAM x l.TT)" -shd_def "shd == stream_when`(LAM x l.x)" -stl_def "stl == stream_when`(LAM x l.l)" - -(* ----------------------------------------------------------------------- *) -(* the taker for streams *) - -stream_take_def "stream_take == (%n.iterate n stream_copy UU)" - -(* ----------------------------------------------------------------------- *) - -stream_finite_def "stream_finite == (%s.? n.stream_take n `s=s)" - -(* ----------------------------------------------------------------------- *) -(* definition of bisimulation is determined by domain equation *) -(* simplification and rewriting for abstract constants yields def below *) - -stream_bisim_def "stream_bisim == -(%R.!s1 s2. - R s1 s2 --> - ((s1=UU & s2=UU) | - (? x s11 s21. x~=UU & s1=scons`x`s11 & s2 = scons`x`s21 & R s11 s21)))" - -end - - - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Stream2.ML --- a/src/HOLCF/Stream2.ML Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,43 +0,0 @@ -(* Title: HOLCF/stream2.ML - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Lemmas for theory Stream2.thy -*) - -open Stream2; - -(* ------------------------------------------------------------------------- *) -(* expand fixed point properties *) -(* ------------------------------------------------------------------------- *) - -val smap_def2 = fix_prover2 Stream2.thy smap_def - "smap = (LAM f s. stream_when`(LAM x l.scons`(f`x) `(smap`f`l)) `s)"; - - -(* ------------------------------------------------------------------------- *) -(* recursive properties *) -(* ------------------------------------------------------------------------- *) - - -qed_goal "smap1" Stream2.thy "smap`f`UU = UU" - (fn prems => - [ - (rtac (smap_def2 RS ssubst) 1), - (simp_tac (!simpset addsimps stream_when) 1) - ]); - -qed_goal "smap2" Stream2.thy - "x~=UU ==> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (rtac (smap_def2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac refl 1) - ]); - - -val stream2_rews = [smap1, smap2]; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Stream2.thy --- a/src/HOLCF/Stream2.thy Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,29 +0,0 @@ -(* Title: HOLCF/stream2.thy - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Additional constants for stream -*) - -Stream2 = Stream + - -consts - -smap :: "('a -> 'b) -> 'a stream -> 'b stream" - -defs - -smap_def - "smap == fix`(LAM h f s. stream_when`(LAM x l.scons `(f`x) `(h`f`l)) `s)" - - -end - - -(* - smap`f`UU = UU - x~=UU --> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs) - -*) - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/Tr1.thy --- a/src/HOLCF/Tr1.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/Tr1.thy Fri Oct 06 17:25:24 1995 +0100 @@ -41,4 +41,9 @@ tr_when_def "tr_when == (LAM e1 e2 t. sswhen`(LAM x.e1)`(LAM y.e2)`(rep_tr`t))" + +(* start 8bit 1 *) +(* end 8bit 1 *) + + end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ax_ops/holcflogic.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ax_ops/holcflogic.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,64 @@ +(* + ID: $Id$ + Author: Tobias Mayr + +Additional term and type constructors, extension of Pure/term.ML, logic.ML +and HOL/hologic.ML + +TODO: + +*) + +signature HOLCFLOGIC = +sig + val True : term; + val False : term; + val Imp : term; + val And : term; + val Not : term; + val mkNot : term -> term; (* negates, no Trueprop *) + val mkNotEqUU : term -> term; (* Trueprop(x ~= UU) *) + val mkNotEqUU_in : term -> term -> term; (* "v~=UU ==> t" *) + val ==> : typ * typ -> typ; (* Infix operation typ constructor *) + val mkOpApp : term -> term -> term; (* Ops application (f ` x) *) + val mkCPair : term -> term -> term; (* cpair constructor *) +end; + +structure HOLCFlogic : HOLCFLOGIC = +struct +open Logic +open HOLogic + +val True = Const("True",boolT); +val False = Const("False",boolT); +val Imp = Const("op -->",boolT --> boolT --> boolT); +val And = Const("op &",boolT --> boolT --> boolT); +val Not = Const("not",boolT --> boolT); + +fun mkNot A = Not $ A; (* negates, no Trueprop *) + +(* Trueprop(x ~= UU) *) +fun mkNotEqUU v = mk_Trueprop(mkNot(mk_eq(v,Const("UU",fastype_of v)))); + +(* mkNotEqUU_in v t = "v~=UU ==> t" *) +fun mkNotEqUU_in vterm term = + mk_implies(mkNotEqUU vterm,term) + + +infixr 6 ==>; (* the analogon to --> for operations *) +fun a ==> b = Type("->",[a,b]); + +(* Ops application (f ` x) *) +fun mkOpApp (f as Const(_,ft as Type("->",[xt,rt]))) x = + Const("fapp",ft --> xt --> rt) $ f $ x + | mkOpApp f x = (print(f);error("Internal error: mkOpApp: wrong args")); + +(* cpair constructor *) +fun mkCPair x y = let val tx = fastype_of x + val ty = fastype_of y + in + Const("fapp",(ty==>Type("*",[tx,ty]))-->ty-->Type("*",[tx,ty])) $ + (mkOpApp (Const("cpair",tx ==> ty ==> Type("*",[tx,ty]))) x) $ y + end; + +end; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ax_ops/install.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ax_ops/install.tex Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,25 @@ + +Um diese Erweiterung zu installieren, gen\"ugt es nat\"urlich nicht, die +Sourcefiles nach {\tt isabelle/src/HOLCF} zu kopieren, +es mu\ss\ au\ss erdem ROOT.ML +folgenderma\ss en abge\"andert werden:\\ +Die Zeile +\begin{verbatim} +init_thy_reader(); +\end{verbatim} +wird ersetzt durch die Zeilen +\begin{verbatim} +use "holcflogic.ML"; +use "thy_axioms.ML"; +use "thy_ops.ML"; +use "thy_syntax.ML"; +\end{verbatim} +abschliessend wird die {\tt HOLCF}--Database neu erzeugt:\\ +{\tt make -f Makefile}\\ +(Vorraussetzung ist nat\"urlich, da\ss\ die {\tt ISABELLE...}--Environment +Variablen korrekt, wie im Makefile beschrieben, gesetzt sind.) + +Die Installation ist damit abgeschlossen. + + + \ No newline at end of file diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ax_ops/thy_axioms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ax_ops/thy_axioms.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,216 @@ +(* + ID: $Id$ + Author: Tobias Mayr + +Additional theory file section for HOLCF: axioms +There's an elaborate but german description of this program +and a short english description of the new sections, +write to mayrt@informatik.tu-muenchen.de. + +TODO: + +*) + +(*** new section of HOLCF : axioms + since rules are already internally called axioms, + the new section is internally called ext_axioms res. eaxms *) + +signature THY_AXIOMS = +sig + (* theory extenders : *) + val add_ext_axioms : (string * string) list -> (string * string) list + -> (string * string) list -> theory -> theory; + val add_ext_axioms_i : (string * (typ option)) list -> + (string * (typ option)) list -> + (string * term) list -> theory -> theory; + val axioms_keywords : string list; + val axioms_sections : (string * (ThyParse.token list -> + (string * string) * ThyParse.token list)) list; +end; + +structure ThyAxioms : THY_AXIOMS = +struct + +open HOLCFlogic; + +(** library ******************************************************) + +fun apsnd_of_three f = fn (a,b,c) => (a,f b,c); + +fun is_elem e list = exists (fn y => e=y) list + +fun without l1 l2 = (* l1 without the elements of l2 *) + filter (fn x => (not (is_elem x l2))) l1; + +fun conc [e:string] = e + | conc (head::tail) = head^", "^(conc tail) + | conc [] = ""; + +fun appear varlist = (* all (x,_) for which (x,_) is in varlist *) + filter (fn x => (exists (fn y => (fst x)=(fst y)) varlist)) + + +(* all non unique elements of a list *) +fun doubles (hd::tl) = if (is_elem hd tl) + then (hd::(doubles tl)) + else (doubles tl) + | doubles _ = []; + + +(* The main functions are the section parser ext_axiom_decls and the + theory extender add_ext_axioms. *) + +(** theory extender : add_ext_axioms *) + +(* forms a function from constrained varnames to their constraints + these constraints are then used local to each axiom, as an argument + of read_def_cterm. Called by add_ext_axioms. *) +fun get_constraints_of_str sign ((vname,vtyp)::tail) = + if (vtyp <> "") + then ((fn (x,_)=> if x=vname + then Some (#T (rep_ctyp (read_ctyp sign vtyp))) + else raise Match) + orelf (get_constraints_of_str sign tail)) + else (get_constraints_of_str sign tail) + | get_constraints_of_str sign [] = K None; + +(* does the same job for allready parsed optional constraints. + Called by add_ext_axioms_i. *) +fun get_constraints_of_typ sign ((vname,vtyp)::tail) = + if (is_some vtyp) + then ((fn (x,_)=> if x=vname + then vtyp + else raise Match) + orelf (get_constraints_of_typ sign tail)) + else (get_constraints_of_typ sign tail) + | get_constraints_of_typ sign [] = K None; + + +(* applies mkNotEqUU_in on the axiom and every Free that appears in the list + and in the axiom. Called by check_and_extend. *) +fun add_prem axiom [] = axiom + | add_prem axiom (vname::tl) = + let val vterm = find_first (fn x => ((#1 o dest_Free) x = vname)) + (term_frees axiom) + in + add_prem + (if (is_some vterm) + then mkNotEqUU_in (the vterm) axiom + else axiom) + tl + end + +(* checks for uniqueness and completeness of var/defvar declarations, + and enriches the axiom term with premises. Called by add_ext_axioms(_i).*) +fun check_and_extend sign defvarl varl axiom = + let + val names_of_frees = map (fst o dest_Free) + (term_frees axiom); + val all_decl_varnames = (defvarl @ varl); + val undeclared = without names_of_frees all_decl_varnames; + val doubles = doubles all_decl_varnames + in + if (doubles <> []) + then + (error("Multiple declarations of one identifier in section axioms :\n" + ^(conc doubles))) + else (); + if (undeclared <> []) + then + (error("Undeclared identifiers in section axioms : \n" + ^(conc undeclared))) + else (); + add_prem axiom (rev defvarl) + end; + +(* the next five only differ from the original add_axioms' subfunctions + in the constraints argument for read_def_cterm *) +local + fun err_in_axm name = + error ("The error(s) above occurred in axiom " ^ quote name); + + fun no_vars tm = + if null (term_vars tm) andalso null (term_tvars tm) then tm + else error "Illegal schematic variable(s) in term"; + + fun read_ext_cterm sign constraints = + #1 o read_def_cterm (sign, constraints, K None) [] true; + + (* only for add_ext_axioms (working on strings) *) + fun read_ext_axm sg constraints (name,str) = + (name, no_vars (term_of (read_ext_cterm sg constraints (str, propT)))) + handle ERROR => err_in_axm name; + + (* only for add_ext_axioms_i (working on terms) *) + fun read_ext_axm_terms sg constraints (name,term) = + (name, no_vars (#2(Sign.infer_types sg constraints (K None) [] true + ([term], propT)))) + handle ERROR => err_in_axm name; + +in + +(******* THE THEORY EXTENDERS THEMSELVES *****) + fun add_ext_axioms varlist defvarlist axioms theory = + let val {sign, ...} = rep_theory theory; + val constraints = get_constraints_of_str sign (defvarlist@varlist) + in + add_axioms_i (map (apsnd + (check_and_extend sign (map fst defvarlist) (map fst varlist)) o + (read_ext_axm sign constraints)) axioms) theory + end + + fun add_ext_axioms_i varlist defvarlist axiom_terms theory = + let val {sign, ...} = rep_theory theory; + val constraints = get_constraints_of_typ sign (defvarlist@varlist) + in + add_axioms_i (map (apsnd (check_and_extend sign + (map fst defvarlist) (map fst varlist)) o + (read_ext_axm_terms sign constraints)) axiom_terms) theory + end +end; + + +(******** SECTION PARSER : ext_axiom_decls **********) +local + open ThyParse + + (* as in the pure section 'rules' : + making the "val thmname = get_axiom thy thmname" list *) + val mk_list_of_pairs = mk_big_list o map (mk_pair o apfst quote); + fun mk_val ax = "val " ^ ax ^ " = get_axiom thy " ^ quote ax ^ ";"; + val mk_vals = cat_lines o map mk_val; + + (* making the call for the theory extender *) + fun mk_eaxms_decls ((vars,defvars),axms) = + ( "|> ThyAxioms.add_ext_axioms \n " ^ + (mk_list_of_pairs vars) ^ "\n " ^ + (mk_list_of_pairs defvars) ^ "\n " ^ + (mk_list_of_pairs axms), + mk_vals (map fst axms)); + + (* parsing the concrete syntax *) + + val axiom_decls = (repeat1 (ident -- !! string)); + + val varlist = "vars" $$-- + repeat1 (ident -- optional ("::" $$-- string) "\"\""); + + val defvarlist = "defvars" $$-- + repeat1 (ident -- optional ("::" $$-- string) "\"\""); + +in + + val ext_axiom_decls = (optional varlist []) -- (optional defvarlist []) + -- ("in" $$-- axiom_decls) >> mk_eaxms_decls; +end; (* local *) + + +(**** new keywords and sections ************************************) + +val axioms_keywords = ["vars", "defvars","in"]; + (* "::" is already a pure keyword *) + +val axioms_sections = [("axioms" , ext_axiom_decls)] + +end; (* the structure *) + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ax_ops/thy_ops.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ax_ops/thy_ops.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,439 @@ +(* TTITLEF/thy_ops.ML + ID: $Id$ + Author: Tobias Mayr + +Additional theory file section for HOLCF: ops +There's an elaborate but german description of this program, +write to mayrt@informatik.tu-muenchen.de. +For a short english description of the new sections +write to regensbu@informatik.tu-muenchen.de. + +TODO: vielleicht AST-Darstellung mit "op name" statt _I_... + +*) + +signature THY_OPS = +sig + (* continuous mixfixes (extension of datatype PrivateSyntax.Mixfix.mixfix) *) + datatype cmixfix = + Mixfix of PrivateSyntax.Mixfix.mixfix | + CInfixl of int | + CInfixr of int | + CMixfix of string * int list *int; + + exception CINFIX of cmixfix; + val cmixfix_to_mixfix : cmixfix -> PrivateSyntax.Mixfix.mixfix; + + (* theory extenders : *) + val add_ops : {curried: bool, total: bool, strict: bool} -> + (string * string * cmixfix) list -> theory -> theory; + val add_ops_i : {curried: bool, total: bool, strict: bool} -> + (string * typ * cmixfix) list -> theory -> theory; + val ops_keywords : string list; + val ops_sections : (string * (ThyParse.token list -> + (string * string) * ThyParse.token list)) list; + val opt_cmixfix: ThyParse.token list -> (string * ThyParse.token list); + val const_name : string -> cmixfix -> string; +end; + +structure ThyOps : THY_OPS = +struct + +open HOLCFlogic; + +(** library ******************************************************) + +(* abbreviations *) +val internal = fst; (* cinfix-ops will have diffrent internal/external names *) +val external = snd; + +fun apsnd_of_three f = fn (a,b,c) => (a,f b,c); + + +(******** ops ********************) + +(* the extended copy of mixfix *) +datatype cmixfix = + Mixfix of PrivateSyntax.Mixfix.mixfix | + CInfixl of int | + CInfixr of int | + CMixfix of string * int list *int; + +exception CINFIX of cmixfix; + +fun cmixfix_to_mixfix (Mixfix x) = x + | cmixfix_to_mixfix x = raise CINFIX x; + + +(** theory extender : add_ops *) + +(* generating the declarations of the new constants. ************* + cinfix names x are internally non infix (renamed by mk_internal_name) + and externally non continous infix function names (changed by op_to_fun). + Thus the cinfix declaration is splitted in an 'oldstyle' decl, + which is NoSyn (non infix) and is added by add_consts_i, + and an syn(tactic) decl, which is an infix function (not operation) + added by add_syntax_i, so that it can appear in input strings, but + not in terms. + The interface between internal and external names is realized by + transrules A x B <=> _x ' A ' B (generated by xrules_of) + The boolean argument 'curried' distinguishes between curried and + tupeled syntax of operation application *) + +local + fun strip ("'" :: c :: cs) = c :: strip cs + | strip ["'"] = [] + | strip (c :: cs) = c :: strip cs + | strip [] = []; + + val strip_esc = implode o strip o explode; + + fun infix_name c = "op " ^ strip_esc c; +in + val mk_internal_name = infix_name; +(* +(* changing e.g. 'ab' to '_I_97_98'. + Called by oldstyle, xrules_of, strictness_axms and totality_axms. *) + fun mk_internal_name name = + let fun alphanum (s::ss) = "_"^(string_of_int (ord s))^(alphanum ss) + | alphanum [] = ""; + in + "_I"^(alphanum o explode) name + end; +*) + (* extension of Pure/Syntax/mixfix.ML: SynExt.const_name *) + fun const_name c (CInfixl _) = mk_internal_name c + | const_name c (CInfixr _) = mk_internal_name c + | const_name c (CMixfix _) = c + | const_name c (Mixfix x) = Syntax.const_name c x; +end; + +(* Changing a->b->c res. a*b->c to a=>b=>c. Called by syn_decls. *) +(*####*) +fun op_to_fun true sign (Type("->" ,[larg,t]))= + Type("fun",[larg,op_to_fun true sign t]) + | op_to_fun false sign (Type("->",[args,res])) = let + fun otf (Type("*",[larg,rargs])) = Type("fun",[larg,otf rargs]) + | otf t = Type("fun",[t,res]); + in otf args end + | op_to_fun _ sign t = t(*error("Wrong type for cinfix/cmixfix : "^ + (Sign.string_of_typ sign t))*); +(*####*) + +(* oldstyle is called by add_ext_axioms(_i) *) + (* the first part is just copying the homomorphic part of the structures *) +fun oldstyle ((name,typ,Mixfix(x))::tl) = + (name,typ,x)::(oldstyle tl) + | oldstyle ((name,typ,CInfixl(i))::tl) = + (mk_internal_name name,typ,PrivateSyntax.Mixfix.NoSyn):: + (oldstyle tl) + | oldstyle ((name,typ,CInfixr(i))::tl) = + (mk_internal_name name,typ,PrivateSyntax.Mixfix.NoSyn):: + (oldstyle tl) + | oldstyle ((name,typ,CMixfix(x))::tl) = + (name,typ,PrivateSyntax.Mixfix.NoSyn):: + (oldstyle tl) + | oldstyle [] = []; + +(* generating the external purely syntactical infix functions. + Called by add_ext_axioms(_i) *) +fun syn_decls curried sign ((name,typ,CInfixl(i))::tl) = + (name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Infixl(i)):: + (syn_decls curried sign tl) + | syn_decls curried sign ((name,typ,CInfixr(i))::tl) = + (name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Infixr(i)):: + (syn_decls curried sign tl) + | syn_decls curried sign ((name,typ,CMixfix(x))::tl) = +(*#### + ("@"^name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Mixfix(x)):: +####**) + (name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Mixfix(x)):: + + (syn_decls curried sign tl) + | syn_decls curried sign (_::tl) = syn_decls curried sign tl + | syn_decls _ _ [] = []; + +(* generating the translation rules. Called by add_ext_axioms(_i) *) +local open PrivateSyntax.Ast in +fun xrules_of true ((name,typ,CInfixl(i))::tail) = + ((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <-> + (mk_appl (Constant "@fapp") [(mk_appl (Constant "@fapp") [ + Constant (mk_internal_name name),Variable "A"]),Variable "B"])) + ::xrules_of true tail + | xrules_of true ((name,typ,CInfixr(i))::tail) = + ((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <-> + (mk_appl (Constant "@fapp") [(mk_appl (Constant "@fapp") [ + Constant (mk_internal_name name),Variable "A"]),Variable "B"])) + ::xrules_of true tail +(*####*) + | xrules_of true ((name,typ,CMixfix(_))::tail) = let + fun argnames n (Type("->" ,[_,t]))= chr n :: argnames (n+1) t + | argnames _ _ = []; + val names = argnames (ord"A") typ; + in if names = [] then [] else [mk_appl (Constant name) (map Variable names)<-> + foldl (fn (t,arg) => (mk_appl (Constant "@fapp") [t,Variable arg])) + (Constant name,names)] end + @xrules_of true tail +(*####*) + | xrules_of false ((name,typ,CInfixl(i))::tail) = + ((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <-> + (mk_appl (Constant "@fapp") [ Constant(mk_internal_name name), + (mk_appl (Constant "@ctuple") [Variable "A",Variable "B"])])) + ::xrules_of false tail + | xrules_of false ((name,typ,CInfixr(i))::tail) = + ((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <-> + (mk_appl (Constant "@fapp") [ Constant(mk_internal_name name), + (mk_appl (Constant "@ctuple") [Variable "A",Variable "B"])])) + ::xrules_of false tail +(*####*) + | xrules_of false ((name,typ,CMixfix(_))::tail) = let + fun foldr' f l = + let fun itr [] = raise LIST "foldr'" + | itr [a] = a + | itr (a::l) = f(a, itr l) + in itr l end; + fun argnames n (Type("*" ,[_,t]))= chr n :: argnames (n+1) t + | argnames n _ = [chr n]; + val vars = map Variable (case typ of (Type("->" ,[t,_])) =>argnames (ord"A") t + | _ => []); + in if vars = [] then [] else [mk_appl (Constant name) vars <-> + (mk_appl (Constant "@fapp") [Constant name, case vars of [v] => v + | args => mk_appl (Constant "@ctuple") [hd args,foldr' (fn (t,arg) => + mk_appl (Constant "_args") [t,arg]) (tl args)]])] + end + @xrules_of false tail +(*####*) + | xrules_of c ((name,typ,Mixfix(_))::tail) = xrules_of c tail + | xrules_of _ [] = []; +end; +(**** producing the new axioms ****************) + +datatype arguments = Curried_args of ((typ*typ) list) | + Tupeled_args of (typ list); + +fun num_of_args (Curried_args l) = length l + | num_of_args (Tupeled_args l) = length l; + +fun types_of (Curried_args l) = map fst l + | types_of (Tupeled_args l) = l; + +fun mk_mkNotEqUU_vars (typ::tl) cnt = mkNotEqUU (Free("x"^(string_of_int cnt),typ)):: + (mk_mkNotEqUU_vars tl (cnt+1)) + | mk_mkNotEqUU_vars [] _ = []; + +local + (* T1*...*Tn goes to [T1,...,Tn] *) + fun args_of_tupel (Type("*",[left,right])) = left::(args_of_tupel right) + | args_of_tupel T = [T]; + + (* A1->...->An->R goes to [(A1,B1),...,(An,Bn)] where Bi=Ai->...->An->R + Bi is the Type of the function that is applied to an Ai type argument *) + fun args_of_curried (typ as (Type("->",[S,T]))) = + (S,typ) :: args_of_curried T + | args_of_curried _ = []; +in + fun args_of_op true typ = Curried_args(rev(args_of_curried typ)) + | args_of_op false (typ as (Type("->",[S,T]))) = + Tupeled_args(args_of_tupel S) + | args_of_op false _ = Tupeled_args([]); +end; + +(* generates for the type t the type of the fapp constant + that will be applied to t *) +fun mk_fapp_typ (typ as Type("->",argl)) = Type("fun",[typ,Type("fun",argl)]) + | mk_fapp_typ t = (print t; + error("Internal error:mk_fapp_typ: wrong argument\n")); + +fun mk_arg_tupel_UU uu_pos [typ] n = + if n<>uu_pos then Free("x"^(string_of_int n),typ) + else Const("UU",typ) + | mk_arg_tupel_UU uu_pos (typ::tail) n = + mkCPair + (if n<>uu_pos then Free("x"^(string_of_int n),typ) + else Const("UU",typ)) + (mk_arg_tupel_UU uu_pos tail (n+1)) + | mk_arg_tupel_UU _ [] _ = error("Internal error:mk_arg_tupel: empty list"); + +fun mk_app_UU cnt uu_pos fname (Curried_args((typ,ftyp)::tl)) = + Const("fapp",mk_fapp_typ ftyp) $ + (mk_app_UU (cnt-1) uu_pos fname (Curried_args tl))$ + (if cnt = uu_pos then Const("UU",typ) + else Free("x"^(string_of_int cnt),typ)) + | mk_app_UU _ _ (name,typ) (Curried_args []) = Const(name,typ) + | mk_app_UU cnt uu_pos (name,typ) (Tupeled_args []) = Const(name,typ) + | mk_app_UU cnt uu_pos (name,typ) (Tupeled_args list) = + Const("fapp",mk_fapp_typ typ) $ Const(name,typ) $ + mk_arg_tupel_UU uu_pos list 0; + +fun mk_app cnt fname args = mk_app_UU cnt (~1) fname args; + +(* producing the strictness axioms *) +local + fun s_axm_of curried name typ args num cnt = + if cnt = num then + error("Internal error: s_axm_of: arg is no operation "^(external name)) + else + let val app = mk_app_UU (num-1) cnt (internal name,typ) args + val equation = HOLogic.mk_eq(app,Const("UU",fastype_of app)) + in + if cnt = num-1 then equation + else And $ equation $ + s_axm_of curried name typ args num (cnt+1) + end; +in + fun strictness_axms curried ((rawname,typ,cmixfix)::tail) = + let val name = case cmixfix of + (CInfixl _) => (mk_internal_name rawname,rawname) + | (CInfixr _) => (mk_internal_name rawname,rawname) + | _ => (rawname,rawname) + val args = args_of_op curried typ; + val num = num_of_args args; + in + ((external name)^"_strict", + if num <> 0 + then HOLogic.mk_Trueprop(s_axm_of curried name typ args num 0) + else HOLogic.mk_Trueprop(True)) :: strictness_axms curried tail + end + | strictness_axms _ [] = []; +end; (*local*) + +(* producing the totality axioms *) + +fun totality_axms curried ((rawname,typ,cmixfix)::tail) = + let val name = case cmixfix of + (CInfixl _) => (mk_internal_name rawname,rawname) + | (CInfixr _) => (mk_internal_name rawname,rawname) + | _ => (rawname,rawname) + val args = args_of_op curried typ; + val prems = mk_mkNotEqUU_vars (if curried then rev (types_of args) + else (types_of args)) 0; + val term = mk_app (num_of_args args - 1) (internal name,typ) args; + in + ((external name)^"_total", + if num_of_args args <> 0 + then Logic.list_implies (prems,mkNotEqUU term) + else HOLogic.mk_Trueprop(True)) :: totality_axms curried tail + end + | totality_axms _ [] = []; + + + +(* the theory extenders ****************************) + +fun add_ops {curried,strict,total} raw_decls thy = + let val {sign,...} = rep_theory thy; + val decls = map (apsnd_of_three (typ_of o read_ctyp sign)) raw_decls; + val oldstyledecls = oldstyle decls; + val syndecls = syn_decls curried sign decls; + val xrules = xrules_of curried decls; + val s_axms = (if strict then strictness_axms curried decls else []); + val t_axms = (if total then totality_axms curried decls else []); + in + add_trrules_i xrules (add_axioms_i (s_axms @ t_axms) + (add_syntax_i syndecls (add_consts_i oldstyledecls thy))) + end; + +fun add_ops_i {curried,strict,total} decls thy = + let val {sign,...} = rep_theory thy; + val oldstyledecls = oldstyle decls; + val syndecls = syn_decls curried sign decls; + val xrules = xrules_of curried decls; + val s_axms = (if strict then strictness_axms curried decls else []); + val t_axms = (if total then totality_axms curried decls else []); + in + add_trrules_i xrules (add_axioms_i (s_axms @ t_axms) + (add_syntax_i syndecls (add_consts_i oldstyledecls thy))) + end; + + +(* parser: ops_decls ********************************) + +local open ThyParse +in +(* the following is an adapted version of const_decls from thy_parse.ML *) + +val names1 = list1 name; + +val split_decls = flat o map (fn (xs, y) => map (rpair y) xs); + +fun mk_triple2 (x, (y, z)) = mk_triple (x, y, z); + +fun mk_strict_vals [] = "" + | mk_strict_vals [name] = + "get_axiom thy \""^name^"_strict\"\n" + | mk_strict_vals (name::tail) = + "get_axiom thy \""^name^"_strict\",\n"^ + mk_strict_vals tail; + +fun mk_total_vals [] = "" + | mk_total_vals [name] = + "get_axiom thy \""^name^"_total\"\n" + | mk_total_vals (name::tail) = + "get_axiom thy \""^name^"_total\",\n"^ + mk_total_vals tail; + +fun mk_ops_decls (((curried,strict),total),list) = + (* call for the theory extender *) + ("|> ThyOps.add_ops \n"^ + "{ curried = "^curried^" , strict = "^strict^ + " , total = "^total^" } \n"^ + (mk_big_list o map mk_triple2) list^";\n"^ + "val strict_axms = []; val total_axms = [];\nval thy = thy\n", + (* additional declarations *) + (if strict="true" then "val strict_axms = strict_axms @ [\n"^ + mk_strict_vals (map (strip_quotes o fst) list)^ + "];\n" + else "")^ + (if total="true" then "val total_axms = total_axms @ [\n"^ + mk_total_vals (map (strip_quotes o fst) list)^ + "];\n" + else "")); + +(* mixfix annotations *) + +fun cat_parens pre1 pre2 s = cat pre1 (parens (cat pre2 s)); + +val infxl = "infixl" $$-- !! nat >> cat_parens "ThyOps.Mixfix" "Infixl"; +val infxr = "infixr" $$-- !! nat >> cat_parens "ThyOps.Mixfix" "Infixr"; + +val cinfxl = "cinfixl" $$-- !! nat >> cat "ThyOps.CInfixl"; +val cinfxr = "cinfixr" $$-- !! nat >> cat "ThyOps.CInfixr"; + +val opt_pris = optional ("[" $$-- !! (list nat --$$ "]")) [] >> mk_list; + +val cmixfx = "cmixfix" $$-- string -- !! (opt_pris -- optional nat "max_pri") + >> (cat "ThyOps.CMixfix" o mk_triple2); + +val bindr = "binder" $$-- + !! (string -- ( ("[" $$-- nat --$$ "]") -- nat + || nat >> (fn n => (n,n)) + ) ) + >> (cat_parens "ThyOps.Mixfix" "Binder" o mk_triple2); + +val mixfx = string -- !! (opt_pris -- optional nat "max_pri") + >> (cat_parens "ThyOps.Mixfix" "Mixfix" o mk_triple2); + +fun opt_syn fx = optional ("(" $$-- fx --$$ ")") "ThyOps.Mixfix NoSyn"; + +val opt_cmixfix = opt_syn (mixfx || infxl || infxr || bindr || + cinfxl || cinfxr || cmixfx); + +fun ops_decls toks= + (optional ($$ "curried" >> K "true") "false" -- + optional ($$ "strict" >> K "true") "false" -- + optional ($$ "total" >> K "true") "false" -- + (repeat1 (names1 --$$ "::" -- !! (string -- opt_cmixfix)) + >> split_decls) + >> mk_ops_decls) toks + +end; + +(*** new keywords and sections: ******************************************) + +val ops_keywords = ["curried","strict","total","cinfixl","cinfixr","cmixfix"]; + (* "::" is already a pure keyword *) + +val ops_sections = [("ops" , ops_decls) ]; + +end; (* the structure ThyOps *) + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ax_ops/thy_syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ax_ops/thy_syntax.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,36 @@ +(* Title: HOLCF/thy_syntax.ML + ID: $Id$ + Author: Tobias Mayr + +Installation of the additional theory file sections for HOLCF: axioms , ops +There's an elaborate but german description of this extension +and a short english description of the new sections, +write to mayrt@informatik.tu-muenchen.de. + +TODO: + +*) + +(* use "holcflogics.ML"; + use "thy_axioms.ML"; + use "thy_ops.ML"; should already have been done in ROOT.ML *) + +structure ThySynData : THY_SYN_DATA = +struct + +open HOLCFlogic; + +val user_keywords = (*####*)filter_out (fn s => s mem (ThyAxioms.axioms_keywords@ + ThyOps.ops_keywords)) (*####*)ThySynData.user_keywords @ + ThyAxioms.axioms_keywords @ + ThyOps.ops_keywords; + +val user_sections = (*####*)filter_out (fn (s,_) => s mem (map fst ( + ThyAxioms.axioms_sections@ ThyOps.ops_sections))) (*####*) + ThySynData.user_sections @ + ThyAxioms.axioms_sections @ + ThyOps.ops_sections; +end; + +structure ThySyn = ThySynFun(ThySynData); +init_thy_reader (); diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/domain/axioms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/domain/axioms.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,161 @@ +(* axioms.ML + ID: $Id$ + Author : David von Oheimb + Created: 31-May-95 + Updated: 12-Jun-95 axioms for discriminators, selectors and induction + Updated: 19-Jun-95 axiom for bisimulation + Updated: 28-Jul-95 gen_by-section + Updated: 29-Aug-95 simultaneous domain equations + Copyright 1995 TU Muenchen +*) + + +structure Domain_Axioms = struct + +local + +open Domain_Library; +infixr 0 ===>;infixr 0 ==>;infix 0 == ; +infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<; +infix 9 ` ; infix 9 `% ; infix 9 `%%; infixr 9 oo; + +fun infer_types thy' = map (inferT_axm (sign_of thy')); + +fun calc_axioms comp_dname (eqs : eq list) n (((dname,_),cons) : eq)= +let + +(* ----- axioms and definitions concerning the isomorphism ------------------------ *) + + val dc_abs = %%(dname^"_abs"); + val dc_rep = %%(dname^"_rep"); + val x_name'= "x"; + val x_name = idx_name eqs x_name' (n+1); + + val ax_abs_iso = (dname^"_abs_iso",mk_trp(dc_rep`(dc_abs`%x_name') === %x_name')); + val ax_rep_iso = (dname^"_rep_iso",mk_trp(dc_abs`(dc_rep`%x_name') === %x_name')); + + val ax_when_def = (dname^"_when_def",%%(dname^"_when") == + foldr (uncurry /\ ) (when_funs cons, /\x_name'((when_body cons (fn (x,y) => + Bound(1+length cons+x-y)))`(dc_rep`Bound 0)))); + + val ax_copy_def = let + fun simp_oo (Const ("fapp", _) $ (Const ("fapp", _) $ + Const ("cfcomp", _) $ fc) $ Const ("ID", _)) = fc + | simp_oo t = t; + fun simp_app (Const ("fapp", _) $ Const ("ID", _) $ t) = t + | simp_app t = t; + fun mk_arg m n arg = (if is_lazy arg + then fn t => %%"lift"`(simp_oo (%%"up" oo t)) else Id) + (if_rec arg (cproj (Bound (2*min[n,m])) eqs) (%%"ID")); + fun mk_prod (t1,t2) = %%"ssplit"`(/\ "x" (/\ "y" (%%"spair"` + simp_app(t1`Bound 1)`simp_app(t2`Bound 0)))); + fun one_con (_,args) = if args = [] then %%"ID" else + foldr' mk_prod (mapn (mk_arg (length args-1)) 1 args); + fun mk_sum (t1,t2) = %%"sswhen"`(simp_oo (%%"sinl" oo t1)) + `(simp_oo (%%"sinr" oo t2)); + in (dname^"_copy_def", %%(dname^"_copy") == /\"f" + (dc_abs oo foldr' mk_sum (map one_con cons) oo dc_rep)) end; + +(* ----- definitions concerning the constructors, discriminators and selectors ---- *) + + val axs_con_def = let + fun idxs z x arg = (if is_lazy arg then fn x => %%"up"`x else Id)(Bound(z-x)); + fun prms [] = %%"one" + | prms vs = foldr' (fn(x,t)=> %%"spair"`x`t) (mapn (idxs (length vs)) 1 vs); + val injs = bin_branchr (fn l=> l@["l"]) (fn l=> l@["r"]); + fun cdef ((con,args),injs) = (extern_name con ^"_def",%%con == + foldr /\# (args,dc_abs` + (foldr (fn (i,t) => %%("sin"^i)`t ) (injs, prms args)))); + in map cdef (cons~~(mapn (fn n => K(injs [] cons n)) 0 cons)) end; + + val axs_dis_def = let + fun ddef (con,_) = (dis_name con ^"_def",%%(dis_name con) == + mk_cfapp(%%(dname^"_when"),map + (fn (con',args) => (foldr /\# + (args,if con'=con then %%"TT" else %%"FF"))) cons)) + in map ddef cons end; + + val axs_sel_def = let + fun sdef con n arg = (sel_of arg^"_def",%%(sel_of arg) == + mk_cfapp(%%(dname^"_when"),map + (fn (con',args) => if con'<>con then %%"UU" else + foldr /\# (args,Bound (length args - n))) cons)); + in flat(map (fn (con,args) => mapn (sdef con) 1 args) cons) end; + + +(* ----- axiom and definitions concerning induction ------------------------------- *) + + fun cproj' T = cproj T eqs n; + val ax_reach = (dname^"_reach" , mk_trp(cproj'(%%"fix"`%%(comp_dname^"_copy")) + `%x_name === %x_name)); + val ax_take_def = (dname^"_take_def",%%(dname^"_take") == mk_lam("n", + cproj'(%%"iterate" $ Bound 0 $ %%(comp_dname^"_copy") $ %%"UU"))); + val ax_finite_def = (dname^"_finite_def",%%(dname^"_finite") == mk_lam(x_name, + mk_ex("n",(%%(dname^"_take") $ Bound 0)`Bound 1 === Bound 1))); + +in [ax_abs_iso, ax_rep_iso, ax_when_def, ax_copy_def] @ + axs_con_def @ axs_dis_def @ axs_sel_def @ + [ax_reach, ax_take_def, ax_finite_def] end; + + +in (* local *) + +fun add_axioms (comp_dname, eqs : eq list) thy' =let + val dnames = map (fst o fst) eqs; + val x_name = idx_name dnames "x"; + fun copy_app dname = %%(dname^"_copy")`Bound 0; + val ax_copy_def = (comp_dname^"_copy_def" , %%(comp_dname^"_copy") == + /\"f"(foldr' cpair (map copy_app dnames))); + val ax_bisim_def = (comp_dname^"_bisim_def",%%(comp_dname^"_bisim") == mk_lam("R", + let + fun one_con (con,args) = let + val nonrec_args = filter_out is_rec args; + val rec_args = filter is_rec args; + val nonrecs_cnt = length nonrec_args; + val recs_cnt = length rec_args; + val allargs = nonrec_args @ rec_args + @ map (upd_vname (fn s=> s^"'")) rec_args; + val allvns = map vname allargs; + fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg; + val vns1 = map (vname_arg "" ) args; + val vns2 = map (vname_arg "'") args; + val allargs_cnt = nonrecs_cnt + 2*recs_cnt; + val rec_idxs = (recs_cnt-1) downto 0; + val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg) + (allargs~~((allargs_cnt-1) downto 0))); + fun rel_app i ra = proj (Bound(allargs_cnt+2)) dnames (rec_of ra) $ + Bound (2*recs_cnt-i) $ Bound (recs_cnt-i); + val capps = foldr mk_conj (mapn rel_app 1 rec_args, + mk_conj(Bound(allargs_cnt+1)===mk_cfapp(%%con,map (bound_arg allvns) vns1), + Bound(allargs_cnt+0)===mk_cfapp(%%con,map (bound_arg allvns) vns2))); + in foldr mk_ex (allvns, foldr mk_conj + (map (defined o Bound) nonlazy_idxs,capps)) end; + fun one_comp n (_,cons) = mk_all(x_name (n+1), mk_all(x_name (n+1)^"'", mk_imp( + proj (Bound 2) dnames n $ Bound 1 $ Bound 0, + foldr' mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU)::map one_con cons)))); + in foldr' mk_conj (mapn one_comp 0 eqs)end )); + val thy_axs = flat (mapn (calc_axioms comp_dname eqs) 0 eqs) @ + (if length eqs>1 then [ax_copy_def] else []) @ [ax_bisim_def]; +in thy' |> add_axioms_i (infer_types thy' thy_axs) end; + + +fun add_gen_by ((tname,finite),(typs,cnstrs)) thy' = let + fun pred_name typ ="P"^(if typs=[typ] then "" else string_of_int(1+find(typ,typs))); + fun lift_adm t = lift (fn typ => %%"adm" $ %(pred_name typ)) + (if finite then [] else typs,t); + fun lift_pred_UU t = lift (fn typ => %(pred_name typ) $ UU) (typs,t); + fun one_cnstr (cnstr,vns,(args,res)) = let + val rec_args = filter (fn (_,typ) => typ mem typs)(vns~~args); + val app = mk_cfapp(%%cnstr,map (bound_arg vns) vns); + in foldr mk_All (vns, + lift (fn (vn,typ) => %(pred_name typ) $ bound_arg vns vn) + (rec_args,defined app ==> %(pred_name res)$app)) end; + fun one_conc typ = let val pn = pred_name typ in + %pn $ %("x"^implode(tl(explode pn))) end; + val concl = mk_trp(foldr' mk_conj (map one_conc typs)); + val induct =(tname^"_induct",lift_adm(lift_pred_UU( + foldr (op ===>) (map one_cnstr cnstrs,concl)))); +in thy' |> add_axioms_i (infer_types thy' [induct]) end; + +end; (* local *) +end; (* struct *) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/domain/extender.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/domain/extender.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,123 @@ +(* extender.ML + ID: $Id$ + Author : David von Oheimb + Created: 17-May-95 + Updated: 31-May-95 extracted syntax.ML, theorems.ML + Updated: 07-Jul-95 streamlined format of cons list + Updated: 21-Jul-95 gen_by-section + Updated: 28-Aug-95 simultaneous domain equations + Copyright 1995 TU Muenchen +*) + + +structure Extender = +struct + +local + +open Domain_Library; + +(* ----- general testing and preprocessing of constructor list -------------------- *) + + fun check_domain(eqs':((string * typ list) * + (string * ThyOps.cmixfix * (bool*string*typ) list) list) list) = let + val dtnvs = map fst eqs'; + val cons' = flat (map snd eqs'); + val test_dupl_typs = (case duplicates (map fst dtnvs) of + [] => false | dups => error ("Duplicate types: " ^ commas_quote dups)); + val test_dupl_cons = (case duplicates (map first cons') of + [] => false | dups => error ("Duplicate constructors: " ^ commas_quote dups)); + val test_dupl_sels = (case duplicates(map second (flat(map third cons'))) of + [] => false | dups => error ("Duplicate selectors: "^commas_quote dups)); + val test_dupl_tvars = let fun vname (TFree(v,_)) = v + | vname _ = Imposs "extender:vname"; + in exists (fn tvars => case duplicates (map vname tvars) of + [] => false | dups => error ("Duplicate type arguments: " ^commas_quote dups)) + (map snd dtnvs) end; + (*test for free type variables and invalid use of recursive type*) + val analyse_types = forall (fn ((_,typevars),cons') => + forall (fn con' => let + val types = map third (third con'); + fun analyse(t as TFree(v,_)) = t mem typevars orelse + error ("Free type variable " ^ v ^ " on rhs.") + | analyse(Type(s,typl)) = (case assoc (dtnvs,s) of None => analyses typl + | Some tvs => tvs = typl orelse + error ("Recursion of type " ^ s ^ " with different arguments")) + | analyse(TVar _) = Imposs "extender:analyse" + and analyses ts = forall analyse ts; + in analyses types end) cons' + ) eqs'; + in true end; (* let *) + + fun check_gen_by thy' (typs': string list,cnstrss': string list list) = let + val test_dupl_typs = (case duplicates typs' of [] => false + | dups => error ("Duplicate types: " ^ commas_quote dups)); + val test_dupl_cnstrs = map (fn cnstrs' => (case duplicates cnstrs' of [] => false + | dups => error ("Duplicate constructors: " ^ commas_quote dups))) cnstrss'; + val tsig = #tsig(Sign.rep_sg(sign_of thy')); + val tycons = map fst (#tycons(Type.rep_tsig tsig)); + val test_types = forall(fn t=>t mem tycons orelse error("Unknown type: "^t))typs'; + val cnstrss = let + fun type_of c = case (Sign.const_type(sign_of thy') c) of Some t => t + | None => error ("Unknown constructor: "^c); + fun args_result_type (t as (Type(tn,[arg,rest]))) = + if tn = "->" orelse tn = "=>" + then let val (ts,r) = args_result_type rest in (arg::ts,r) end + else ([],t) + | args_result_type t = ([],t); + in map (map (fn cn => let val (args,res) = args_result_type (type_of cn) in + (cn,mk_var_names args,(args,res)) end)) cnstrss' + : (string * (* operator name of constr *) + string list * (* argument name list *) + (typ list * (* argument types *) + typ)) (* result type *) + list list end; + fun test_equal_type tn (cn,_,(_,rt)) = fst (type_name_vars rt) = tn orelse + error("Inappropriate result type for constructor "^cn); + val typs = map (fn (tn, cnstrs) => + (map (test_equal_type tn) cnstrs; snd(third(hd(cnstrs))))) + (typs'~~cnstrss); + val test_typs = map (fn (typ,cnstrs) => + if not (Type.typ_instance(tsig,typ,TVar(("'a",0),["pcpo"]))) + then error("Not a pcpo type: "^fst(type_name_vars typ)) + else map (fn (cn,_,(_,rt)) => rt=typ + orelse error("Non-identical result types for constructors "^ + first(hd cnstrs)^" and "^ cn )) cnstrs) (typs~~cnstrss); + val proper_args = let + fun occurs tn (Type(tn',ts)) = (tn'=tn) orelse exists (occurs tn) ts + | occurs _ _ = false; + fun proper_arg cn atyp = forall (fn typ => let + val tn = fst(type_name_vars typ) + in atyp=typ orelse not (occurs tn atyp) orelse + error("Illegal use of type "^ tn ^ + " as argument of constructor " ^cn)end )typs; + fun proper_curry (cn,_,(args,_)) = forall (proper_arg cn) args; + in map (map proper_curry) cnstrss end; + in (typs, flat cnstrss) end; + +(* ----- calls for building new thy and thms -------------------------------------- *) + +in + + fun add_domain (comp_dname,eqs') thy'' = let + val ok_dummy = check_domain eqs'; + val thy' = thy'' |> Domain_Syntax.add_syntax (comp_dname,eqs'); + val dts = map (Type o fst) eqs'; + fun cons cons' = (map (fn (con,syn,args) => + (ThyOps.const_name con syn, + map (fn ((lazy,sel,tp),vn) => ((lazy, + find (tp,dts) handle LIST "find" => ~1), + sel,vn)) + (args~~(mk_var_names(map third args))) + )) cons') : cons list; + val eqs = map (fn (dtnvs,cons') => (dtnvs,cons cons')) eqs' : eq list; + val thy = thy' |> Domain_Axioms.add_axioms (comp_dname,eqs); + in (thy,eqs) end; + + fun add_gen_by ((tname,finite),(typs',cnstrss')) thy' = let + val (typs,cnstrs) = check_gen_by thy' (typs',cnstrss'); + in + Domain_Axioms.add_gen_by ((tname,finite),(typs,cnstrs)) thy' end; + +end (* local *) +end (* struct *) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/domain/interface.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/domain/interface.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,165 @@ +(* interface.ML + ID: $Id$ + Author: David von Oheimb + Created: 17-May-95 + Updated: 24-May-95 + Updated: 03-Jun-95 incremental change of ThySyn + Updated: 11-Jul-95 use of ThyOps for cinfixes + Updated: 21-Jul-95 gen_by-section + Updated: 29-Aug-95 simultaneous domain equations + Updated: 25-Aug-95 better syntax for simultaneous domain equations + Copyright 1995 TU Muenchen +*) + +local + +structure ThySynData: THY_SYN_DATA = (* overwrites old version of ThySynData!!!!!!! *) +struct + +local + open ThyParse; + open Domain_Library; + +(* ----- generation of bindings for axioms and theorems in trailer of .thy.ML ----- *) + + fun gt_ax name = "get_axiom thy "^quote name; + fun gen_val dname name = "val "^name^" = " ^gt_ax (dname^"_"^name)^";"; + fun gen_vall name l = "val "^name^" = " ^mk_list l^";"; + val rews1 = "iso_rews @ when_rews @\n\ + \con_rews @ sel_rews @ dis_rews @ dists_eq @ dists_le @\n\ + \copy_rews"; + + fun gen_domain eqname num ((dname,_), cons') = let + val axioms1 = ["abs_iso", "rep_iso", "when_def"]; + (* con_defs , sel_defs, dis_defs *) + val axioms2 = ["copy_def"]; + val theorems = + "iso_rews, exhaust, cases, when_rews, con_rews, sel_rews, dis_rews,\n \ + \dists_eq, dists_le, inverts, injects, copy_rews"; + in + "structure "^dname^" = struct\n"^ + cat_lines(map (gen_val dname) axioms1)^"\n"^ + gen_vall"con_defs"(map(fn (con,_,_) => gt_ax(strip_esc con^"_def")) cons')^"\n"^ + gen_vall"sel_defs"(flat(map (fn (_,_,args) => map (fn (_,sel,_) => + gt_ax(sel^"_def")) args) cons'))^"\n"^ + gen_vall"dis_defs"(map(fn (con,_,_) => gt_ax(dis_name_ con^"_def")) + cons')^"\n"^ + cat_lines(map (gen_val dname) axioms2)^"\n"^ + "val ("^ theorems ^") =\n\ + \Domain_Theorems.theorems thy "^eqname^";\n" ^ + (if num > 1 then "val rews = " ^rews1 ^";\n" else "") + end; + + fun mk_domain (eqs'') = let + val dtnvs = map (type_name_vars o fst) eqs''; + val dnames = map fst dtnvs; + val num = length dnames; + val comp_dname = implode (separate "_" dnames); + val conss' = map (fn (dname,cons'') => + let + fun sel n m = upd_second (fn None => dname^"_sel_"^(string_of_int n)^ + "_"^(string_of_int m) + | Some s => s) + fun fill_sels n con = upd_third (mapn (sel n) 1) con; + in mapn fill_sels 1 cons'' end) (dnames~~(map snd eqs'')); + val eqs' = dtnvs~~conss'; + +(* ----- generation of argument string for calling add_domain --------------------- *) + + fun mk_tnv (n,v) = mk_pair(quote n,mk_list(map mk_typ v)) + and mk_typ (TFree(name,sort)) = "TFree"^mk_pair(quote name,makestring sort) + | mk_typ (Type (name,args)) = "Type" ^mk_tnv(name,args) + | mk_typ _ = Imposs "interface:mk_typ"; + fun mk_conslist cons' = mk_list (map + (fn (c,syn,ts)=>mk_triple(quote c,syn,mk_list + (map (fn (b,s ,tp) =>mk_triple(makestring(b:bool),quote s, + mk_typ tp)) ts))) cons'); + in + ("val (thy, "^comp_dname^"_equations) = thy |> Extender.add_domain \n" + ^ mk_pair(quote comp_dname, + mk_list(map (fn (t,cs)=> mk_pair (mk_tnv t,mk_conslist cs)) eqs')) + ^ ";\nval thy = thy", + let + fun plural s = if num > 1 then s^"s" else "["^s^"]"; + val comp_axioms = [(* copy, *) "take_def", "finite_def", "reach" + (*, "bisim_def" *)]; + val comp_theorems = "take_rews, " ^ implode (separate ", " (map plural + ["take_lemma","finite"]))^", finite_ind, ind, coind"; + fun eqname n = "(hd(funpow "^string_of_int n^" tl "^comp_dname^"_equations), " + ^comp_dname^"_equations)"; + fun collect sep name = if num = 1 then name else + implode (separate sep (map (fn s => s^"."^name) dnames)); + in + implode (separate "end; (* struct *)\n\n" + (mapn (fn n => gen_domain (eqname n) num) 0 eqs'))^(if num > 1 then + "end; (* struct *)\n\n\ + \structure "^comp_dname^" = struct\n" else "") ^ + (if num > 1 then gen_val comp_dname "copy_def" ^"\n" else "") ^ + implode ((map (fn s => gen_vall (plural s) + (map (fn dn => gt_ax(dn ^ "_" ^ s)) dnames) ^"\n") comp_axioms)) ^ + gen_val comp_dname "bisim_def" ^"\n\ + \val ("^ comp_theorems ^") =\n\ + \Domain_Theorems.comp_theorems thy \ + \(" ^ quote comp_dname ^ ","^ comp_dname ^"_equations,\n\ + \ ["^collect "," "cases" ^"],\n\ + \ "^ collect "@" "con_rews " ^",\n\ + \ "^ collect "@" "copy_rews" ^");\n\ + \val rews = "^(if num>1 then collect" @ " "rews"else rews1)^ " @ take_rews;\n\ + \end; (* struct *)" + end + ) end; + + fun mk_gen_by (finite,eqs) = let + val typs = map fst eqs; + val cnstrss = map snd eqs; + val tname = implode (separate "_" typs) in + ("|> Extender.add_gen_by " + ^ mk_pair(mk_pair(quote tname,makestring (finite:bool)), + mk_pair(mk_list(map quote typs), + mk_list (map (fn cns => mk_list(map quote cns)) cnstrss))), + "val "^tname^"_induct = " ^gt_ax (tname^"_induct")^";") end; + +(* ----- parser for domain declaration equation ----------------------------------- *) + +(** + val sort = name >> (fn s => [strip_quotes s]) + || "{" $$-- !! (list (name >> strip_quotes) --$$ "}"); + val tvar = (type_var -- (optional ("::" $$-- !! sort) ["pcpo"])) >> TFree +**) + val tvar = type_var >> (fn tv => TFree(strip_quotes tv,["pcpo"])); + + val type_args = "(" $$-- !! (list1 tvar --$$ ")") + || tvar >> (fn x => [x]) + || empty >> K []; + val con_typ = type_args -- ident >> (fn (x,y) => Type(y,x)); + val typ = con_typ + || tvar; + val domain_arg = "(" $$-- (optional ($$ "lazy" >> K true) false) + -- (optional ((ident >> Some) --$$ "::") None) + -- typ --$$ ")" >> (fn ((lazy,sel),tp) => (lazy,sel,tp)) + || ident >> (fn x => (false,None,Type(x,[]))) + || tvar >> (fn x => (false,None,x)); + val domain_cons = (name >> strip_quotes) -- !! (repeat domain_arg) + -- ThyOps.opt_cmixfix + >> (fn ((con,args),syn) => (con,syn,args)); +in + val domain_decl = (enum1 "," (con_typ --$$ "=" -- !! + (enum1 "|" domain_cons))) >> mk_domain; + val gen_by_decl = (optional ($$ "finite" >> K true) false) -- + (enum1 "," (ident --$$ "by" -- !! + (enum1 "|" (name >> strip_quotes)))) >> mk_gen_by; +end; + +val user_keywords = "lazy"::"by"::"finite":: + (**)filter_out (fn s=>s="lazy" orelse s="by" orelse s="finite")(**) + ThySynData.user_keywords; +val user_sections = ("domain", domain_decl)::("generated", gen_by_decl):: + (**)filter_out (fn (s,_)=>s="domain" orelse s="generated")(**) + ThySynData.user_sections; +end; + +in + +structure ThySyn = ThySynFun(ThySynData); (* overwrites old version of ThySyn!!!!!! *) + +end; (* local *) \ No newline at end of file diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/domain/library.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/domain/library.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,202 @@ +(* library.ML + ID: $Id$ + Author: David von Oheimb + Created: 18-Jul-95 extracted from syntax.ML, axioms.ML, extender.ML, interface.ML + Updated: 30-Aug-95 + Copyright 1995 TU Muenchen +*) + +(* ----- general support ---------------------------------------------------------- *) + +fun Id x = x; + +fun mapn f n [] = [] +| mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs; + +fun foldr'' f (l,f2) = + let fun itr [] = raise LIST "foldr''" + | itr [a] = f2 a + | itr (a::l) = f(a, itr l) + in itr l end; +fun foldr' f l = foldr'' f (l,Id); +fun map_cumulr f start xs = foldr (fn (x,(ys,res)) => case f(x,res) of (y,res2) => + (y::ys,res2)) (xs,([],start)); + + +fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x; +fun upd_first f (x,y,z) = (f x, y, z); +fun upd_second f (x,y,z) = ( x, f y, z); +fun upd_third f (x,y,z) = ( x, y, f z); + +(* fn : ('a -> 'a) -> ('a -> 'a) -> 'a -> 'b list -> int -> 'a *) +fun bin_branchr f1 f2 y is j = let +fun bb y 1 _ = y +| bb y _ 0 = f1 y +| bb y i j = if i=2 then (f2 y) else bb (f2 y) (i-1) (j-1) +in bb y (length is) j end; + +fun atomize thm = case concl_of thm of + _ $ (Const("op &",_) $ _ $ _) => atomize (thm RS conjunct1) @ + atomize (thm RS conjunct2) + | _ $ (Const("All" ,_) $ Abs (s,_,_)) => atomize (thm RS + (read_instantiate [("x","?"^s)] spec)) + | _ => [thm]; + +(* ----- specific support for domain ---------------------------------------------- *) + +structure Domain_Library = struct + +exception Impossible of string; +fun Imposs msg = raise Impossible ("Domain:"^msg); + +(* ----- name handling ----- *) + +val strip_esc = let + fun strip ("'" :: c :: cs) = c :: strip cs + | strip ["'"] = [] + | strip (c :: cs) = c :: strip cs + | strip [] = []; +in implode o strip o explode end; + +fun extern_name con = case explode con of + ("o"::"p"::" "::rest) => implode rest + | _ => con; +fun dis_name con = "is_"^ (extern_name con); +fun dis_name_ con = "is_"^ (strip_esc con); + +(*make distinct names out of the type list, + forbidding "o", "x..","f..","P.." as names *) +(*a number string is added if necessary *) +fun mk_var_names types : string list = let + fun typid (Type (id,_) ) = hd (explode id) + | typid (TFree (id,_) ) = hd (tl (explode id)) + | typid (TVar ((id,_),_)) = hd (tl (explode id)); + fun nonreserved id = let val cs = explode id in + if not(hd cs mem ["x","f","P"]) then id + else implode(chr(1+ord (hd cs))::tl cs) end; + fun index_vnames(vn::vns,tab) = + (case assoc(tab,vn) of + None => if vn mem vns + then (vn^"1") :: index_vnames(vns,(vn,2)::tab) + else vn :: index_vnames(vns, tab) + | Some(i) => (vn^(string_of_int i)) :: index_vnames(vns,(vn,i+1)::tab)) + | index_vnames([],tab) = []; +in index_vnames(map (nonreserved o typid) types,[("o",1)]) end; + +fun type_name_vars (Type(name,typevars)) = (name,typevars) +| type_name_vars _ = Imposs "library:type_name_vars"; + +(* ----- support for type and mixfix expressions ----- *) + +fun mk_tvar s = TFree("'"^s,["pcpo"]); +fun mk_typ t (S,T) = Type(t,[S,T]); +infixr 5 -->; +infixr 6 ~>; val op ~> = mk_typ "->"; +val NoSyn' = ThyOps.Mixfix NoSyn; + +(* ----- constructor list handling ----- *) + +type cons = (string * (* operator name of constr *) + ((bool*int)* (* (lazy,recursive element or ~1) *) + string* (* selector name *) + string) (* argument name *) + list); (* argument list *) +type eq = (string * (* name of abstracted type *) + typ list) * (* arguments of abstracted type *) + cons list; (* represented type, as a constructor list *) + +val rec_of = snd o first; +val is_lazy = fst o first; +val sel_of = second; +val vname = third; +val upd_vname = upd_third; +fun is_rec arg = rec_of arg >=0; +fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); +fun nonlazy args = map vname (filter_out is_lazy args); +fun is_one_con_one_arg p cons = length cons = 1 andalso let val args = snd(hd cons) in + length args = 1 andalso p (hd args) end; + +(* ----- support for term expressions ----- *) + +fun % s = Free(s,dummyT); +fun %# arg = %(vname arg); +fun %% s = Const(s,dummyT); + +local open HOLogic in +val mk_trp = mk_Trueprop; +fun mk_conj (S,T) = conj $ S $ T; +fun mk_disj (S,T) = disj $ S $ T; +fun mk_imp (S,T) = imp $ S $ T; +fun mk_lam (x,T) = Abs(x,dummyT,T); +fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P); +local + fun tf (Type (s,args)) = foldl (op $) (%s,map tf args) + | tf (TFree(s,_ )) = %s + | tf _ = Imposs "mk_constrainall"; +in +fun mk_constrain (typ,T) = %%"_constrain" $ T $ tf typ; +fun mk_constrainall (x,typ,P) = %%"All" $ (%%"_constrainAbs"$Abs(x,dummyT,P)$tf typ); +end; + +fun mk_ex (x,P) = mk_exists (x,dummyT,P); +fun mk_not P = Const("not" ,boolT --> boolT) $ P; +end; + +fun mk_All (x,P) = %%"all" $ mk_lam(x,P); (* meta universal quantification *) + +infixr 0 ===>;fun S ===> T = Const("==>", dummyT) $ S $ T; +infixr 0 ==>;fun S ==> T = mk_trp S ===> mk_trp T; +infix 0 ==; fun S == T = Const("==", dummyT) $ S $ T; +infix 1 ===; fun S === T = Const("op =", dummyT) $ S $ T; +infix 1 ~=; fun S ~= T = mk_not (S === T); +infix 1 <<; fun S << T = Const("op <<", dummyT) $ S $ T; +infix 1 ~<<; fun S ~<< T = mk_not (S << T); + +infix 9 ` ; fun f` x = %%"fapp" $ f $ x; +infix 9 `% ; fun f`% s = f` % s; +infix 9 `%%; fun f`%%s = f` %%s; +fun mk_cfapp (F,As) = foldl (op `) (F,As); +fun con_app2 con f args = mk_cfapp(%%con,map f args); +fun con_app con = con_app2 con %#; +fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y; +fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) Id (%# arg); +val cproj = bin_branchr (fn S => %%"cfst"`S) (fn S => %%"csnd"`S); +val proj = bin_branchr (fn S => %%"fst" $S) (fn S => %%"snd" $S); +fun lift tfn = foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); + +fun /\ v T = %%"fabs" $ mk_lam(v,T); +fun /\# (arg,T) = /\ (vname arg) T; +infixr 9 oo; fun S oo T = %%"cfcomp"`S`T; +val UU = %%"UU"; +fun strict f = f`UU === UU; +fun defined t = t ~= UU; +fun cpair (S,T) = %%"cpair"`S`T; +fun lift_defined f = lift (fn x => defined (f x)); +fun bound_arg vns v = Bound(length vns-find(v,vns)-1); + +fun cont_eta_contract (Const("fabs",TT) $ Abs(a,T,body)) = + (case cont_eta_contract body of + body' as (Const("fapp",Ta) $ f $ Bound 0) => + if not (0 mem loose_bnos f) then incr_boundvars ~1 f + else Const("fabs",TT) $ Abs(a,T,body') + | body' => Const("fabs",TT) $ Abs(a,T,body')) +| cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t +| cont_eta_contract t = t; + +fun idx_name dnames s n = s ^ (if length dnames = 1 then "" else string_of_int n); +fun when_funs cons = if length cons = 1 then ["f"] + else mapn (fn n => K("f"^(string_of_int n))) 1 cons; +fun when_body cons funarg = let + fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) + | one_fun n (_,args) = let + val l2 = length args; + fun idxs m arg = (if is_lazy arg then fn x=> %%"lift"`%%"ID"`x + else Id) (Bound(l2-m)); + in cont_eta_contract (foldr'' + (fn (a,t) => %%"ssplit"`(/\# (a,t))) + (args, + fn a => /\# (a,(mk_cfapp(funarg (l2,n),mapn idxs 1 args)))) + ) end; +in foldr' (fn (x,y)=> %%"sswhen"`x`y) (mapn one_fun 1 cons) end; + +end; (* struct *) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/domain/syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/domain/syntax.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,128 @@ +(* syntax.ML + ID: $Id$ + Author: David von Oheimb + Created: 31-May-95 + Updated: 16-Aug-95 case translation + Updated: 28-Aug-95 corrections for case translation, simultaneous domain equations + Copyright 1995 TU Muenchen +*) + + +structure Domain_Syntax = struct + +local + +open Domain_Library; +infixr 5 -->; infixr 6 ~>; +fun calc_syntax dtypeprod ((dname,typevars), + (cons':(string*ThyOps.cmixfix*(bool*string*typ) list) list))= +let +(* ----- constants concerning the isomorphism ------------------------------------- *) + +local + fun opt_lazy (lazy,_,t) = if lazy then Type("u",[t]) else t + fun prod (_,_,args) = if args = [] then Type("one",[]) + else foldr' (mk_typ "**") (map opt_lazy args); + +in + val dtype = Type(dname,typevars); + val dtype2 = foldr' (mk_typ "++") (map prod cons'); + val const_rep = (dname^"_rep" , dtype ~> dtype2, NoSyn'); + val const_abs = (dname^"_abs" , dtype2 ~> dtype , NoSyn'); + val const_copy = (dname^"_copy", dtypeprod ~> dtype ~> dtype , NoSyn'); +end; + +(* ----- constants concerning the constructors, discriminators and selectors ------ *) + +fun is_infix (ThyOps.CInfixl _ ) = true +| is_infix (ThyOps.CInfixr _ ) = true +| is_infix (ThyOps.Mixfix(Infixl _)) = true +| is_infix (ThyOps.Mixfix(Infixr _)) = true +| is_infix _ = false; + +local + val escape = let + fun esc (c :: cs) = if c mem ["'","_","(",")","/"] then "'" :: c :: esc cs + else c :: esc cs + | esc [] = [] + in implode o esc o explode end; + fun freetvar s = if (mk_tvar s) mem typevars then freetvar ("t"^s) else mk_tvar s; + fun when_type (_ ,_,args) = foldr (op ~>) (map third args,freetvar "t"); + fun con (name,s,args) = (name,foldr (op ~>) (map third args,dtype),s); + fun dis (con ,s,_ ) = (dis_name_ con, dtype~>Type("tr",[]), + ThyOps.Mixfix(Mixfix("is'_"^ + (if is_infix s then Id else escape)con,[],max_pri))); + fun sel (_ ,_,args) = map (fn(_,sel,typ)=>(sel,dtype ~> typ,NoSyn'))args; +in + val const_when = (dname^"_when", foldr (op ~>) ((map when_type cons'), + dtype ~> freetvar "t"), NoSyn'); + val consts_con = map con cons'; + val consts_dis = map dis cons'; + val consts_sel = flat(map sel cons'); +end; + +(* ----- constants concerning induction ------------------------------------------- *) + + val const_take = (dname^"_take" ,Type("nat",[]) --> dtype ~> dtype ,NoSyn'); + val const_finite = (dname^"_finite",dtype-->HOLogic.boolT ,NoSyn'); + +(* ----- case translation --------------------------------------------------------- *) + +local open Syntax in + val case_trans = let + fun c_ast con syn = Constant (ThyOps.const_name con syn); + fun expvar n = Variable ("e"^(string_of_int n)); + fun argvar n m _ = Variable ("a"^(string_of_int n)^"_"^(string_of_int m)); + fun app s (l,r) = mk_appl (Constant s) [l,r]; + fun case1 n (con,syn,args) = mk_appl (Constant "@case1") + [if is_infix syn + then mk_appl (c_ast con syn) (mapn (argvar n) 1 args) + else foldl (app "@fapp") (c_ast con syn, (mapn (argvar n) 1 args)), + expvar n]; + fun arg1 n (con,_,args) = if args = [] then expvar n + else mk_appl (Constant "LAM ") + [foldr' (app "_idts") (mapn (argvar n) 1 args) , expvar n]; + in mk_appl (Constant "@case") [Variable "x", foldr' + (fn (c,cs) => mk_appl (Constant "@case2") [c,cs]) + (mapn case1 1 cons')] <-> + mk_appl (Constant "@fapp") [foldl + (fn (w,a ) => mk_appl (Constant "@fapp" ) [w,a ]) + (Constant (dname^"_when"),mapn arg1 1 cons'), + Variable "x"] + end; +end; + +in ([const_rep, const_abs, const_when, const_copy] @ + consts_con @ consts_dis @ consts_sel @ + [const_take, const_finite], + [case_trans]) +end; (* let *) + +(* ----- putting all the syntax stuff together ------------------------------------ *) + +in (* local *) + +fun add_syntax (comp_dname,eqs': ((string * typ list) * + (string * ThyOps.cmixfix * (bool*string*typ) list) list) list) thy'' = +let + fun thy_type (dname,typevars) = (dname, length typevars, NoSyn); + fun thy_arity (dname,typevars) = (dname, map (K ["pcpo"]) typevars, ["pcpo"]); + (** (fn TFree(_,sort) => sort | _ => Imposs "syntax:thy_arities")**) + val thy_types = map (thy_type o fst) eqs'; + val thy_arities = map (thy_arity o fst) eqs'; + val dtypes = map (Type o fst) eqs'; + val funprod = foldr' (mk_typ "*") (map (fn tp => tp ~> tp )dtypes); + val relprod = foldr' (mk_typ "*") (map (fn tp => tp--> tp --> HOLogic.boolT)dtypes); + val const_copy = (comp_dname^"_copy" ,funprod ~> funprod , NoSyn'); + val const_bisim = (comp_dname^"_bisim" ,relprod --> HOLogic.boolT, NoSyn'); + val ctt = map (calc_syntax funprod) eqs'; + val add_cur_ops_i = ThyOps.add_ops_i {curried=true, strict=false, total=false}; +in thy'' |> add_types thy_types + |> add_arities thy_arities + |> add_cur_ops_i (flat(map fst ctt)) + |> add_cur_ops_i ((if length eqs'>1 then [const_copy] else[])@[const_bisim]) + |> add_trrules_i (flat(map snd ctt)) +end; (* let *) + +end; (* local *) +end; (* struct *) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/domain/theorems.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/domain/theorems.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,596 @@ +(* theorems.ML + ID: $Id$ + Author : David von Oheimb + Created: 06-Jun-95 + Updated: 08-Jun-95 first proof from cterms + Updated: 26-Jun-95 proofs for exhaustion thms + Updated: 27-Jun-95 proofs for discriminators, constructors and selectors + Updated: 06-Jul-95 proofs for distinctness, invertibility and injectivity + Updated: 17-Jul-95 proofs for induction rules + Updated: 19-Jul-95 proof for co-induction rule + Updated: 28-Aug-95 definedness theorems for selectors (completion) + Updated: 05-Sep-95 simultaneous domain equations (main part) + Updated: 11-Sep-95 simultaneous domain equations (coding finished) + Updated: 13-Sep-95 simultaneous domain equations (debugging) + Copyright 1995 TU Muenchen +*) + + +structure Domain_Theorems = struct + +local + +open Domain_Library; +infixr 0 ===>;infixr 0 ==>;infix 0 == ; +infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<; +infix 9 ` ; infix 9 `% ; infix 9 `%%; infixr 9 oo; + +(* ----- general proof facilities ------------------------------------------------- *) + +fun inferT sg pre_tm = #2(Sign.infer_types sg (K None)(K None)[]true([pre_tm],propT)); + +(* +infix 0 y; +val b=0; +fun _ y t = by t; +fun g defs t = let val sg = sign_of thy; + val ct = Thm.cterm_of sg (inferT sg t); + in goalw_cterm defs ct end; +*) + +fun pg'' thy defs t = let val sg = sign_of thy; + val ct = Thm.cterm_of sg (inferT sg t); + in prove_goalw_cterm defs ct end; +fun pg' thy defs t tacsf=pg'' thy defs t (fn [] => tacsf + | prems=> (cut_facts_tac prems 1)::tacsf); + +fun REPEAT_DETERM_UNTIL p tac = +let fun drep st = if p st then Sequence.single st + else (case Sequence.pull(tapply(tac,st)) of + None => Sequence.null + | Some(st',_) => drep st') +in Tactic drep end; +val UNTIL_SOLVED = REPEAT_DETERM_UNTIL (has_fewer_prems 1); + +local val trueI2 = prove_goal HOL.thy "f~=x ==> True" (fn prems => [rtac TrueI 1]) in +val kill_neq_tac = dtac trueI2 end; +fun case_UU_tac rews i v = res_inst_tac [("Q",v^"=UU")] classical2 i THEN + asm_simp_tac (HOLCF_ss addsimps rews) i; + +val chain_tac = REPEAT_DETERM o resolve_tac + [is_chain_iterate, ch2ch_fappR, ch2ch_fappL]; + +(* ----- general proofs ----------------------------------------------------------- *) + +val swap3 = prove_goal HOL.thy "[| Q ==> P; ~P |] ==> ~Q" (fn prems => [ + cut_facts_tac prems 1, + etac swap 1, + dtac notnotD 1, + etac (hd prems) 1]); + +val dist_eqI = prove_goal Porder0.thy "~ x << y ==> x ~= y" (fn prems => [ + cut_facts_tac prems 1, + etac swap 1, + dtac notnotD 1, + asm_simp_tac HOLCF_ss 1]); +val cfst_strict = prove_goal Cprod3.thy "cfst`UU = UU" (fn _ => [ + (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]); +val csnd_strict = prove_goal Cprod3.thy "csnd`UU = UU" (fn _ => [ + (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]); + +in + + +fun theorems thy (((dname,_),cons) : eq, eqs :eq list) = +let + +val dummy = writeln ("Proving isomorphism properties of domain "^dname^"..."); +val pg = pg' thy; + +(* ----- getting the axioms and definitions --------------------------------------- *) + +local val ga = get_axiom thy in +val ax_abs_iso = ga (dname^"_abs_iso" ); +val ax_rep_iso = ga (dname^"_rep_iso" ); +val ax_when_def = ga (dname^"_when_def" ); +val axs_con_def = map (fn (con,_) => ga (extern_name con ^"_def")) cons; +val axs_dis_def = map (fn (con,_) => ga ( dis_name con ^"_def")) cons; +val axs_sel_def = flat(map (fn (_,args) => + map (fn arg => ga (sel_of arg ^"_def")) args) cons); +val ax_copy_def = ga (dname^"_copy_def" ); +end; (* local *) + +(* ----- theorems concerning the isomorphism -------------------------------------- *) + +val dc_abs = %%(dname^"_abs"); +val dc_rep = %%(dname^"_rep"); +val dc_copy = %%(dname^"_copy"); +val x_name = "x"; + +val (rep_strict, abs_strict) = let + val r = ax_rep_iso RS (ax_abs_iso RS (allI RSN(2,allI RS iso_strict))) + in (r RS conjunct1, r RS conjunct2) end; +val abs_defin' = pg [] ((dc_abs`%x_name === UU) ==> (%x_name === UU)) [ + res_inst_tac [("t",x_name)] (ax_abs_iso RS subst) 1, + etac ssubst 1, + rtac rep_strict 1]; +val rep_defin' = pg [] ((dc_rep`%x_name === UU) ==> (%x_name === UU)) [ + res_inst_tac [("t",x_name)] (ax_rep_iso RS subst) 1, + etac ssubst 1, + rtac abs_strict 1]; +val iso_rews = [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict]; + +local +val iso_swap = pg [] (dc_rep`%"x" === %"y" ==> %"x" === dc_abs`%"y") [ + dres_inst_tac [("f",dname^"_abs")] cfun_arg_cong 1, + etac (ax_rep_iso RS subst) 1]; +fun exh foldr1 cn quant foldr2 var = let + fun one_con (con,args) = let val vns = map vname args in + foldr quant (vns, foldr2 ((%x_name === con_app2 con (var vns) vns):: + map (defined o (var vns)) (nonlazy args))) end + in foldr1 ((cn(%x_name===UU))::map one_con cons) end; +in +val cases = let + fun common_tac thm = rtac thm 1 THEN contr_tac 1; + fun unit_tac true = common_tac liftE1 + | unit_tac _ = all_tac; + fun prod_tac [] = common_tac oneE + | prod_tac [arg] = unit_tac (is_lazy arg) + | prod_tac (arg::args) = + common_tac sprodE THEN + kill_neq_tac 1 THEN + unit_tac (is_lazy arg) THEN + prod_tac args; + fun sum_one_tac p = SELECT_GOAL(EVERY[ + rtac p 1, + rewrite_goals_tac axs_con_def, + dtac iso_swap 1, + simp_tac HOLCF_ss 1, + UNTIL_SOLVED(fast_tac HOL_cs 1)]) 1; + fun sum_tac [(_,args)] [p] = + prod_tac args THEN sum_one_tac p + | sum_tac ((_,args)::cons') (p::prems) = DETERM( + common_tac ssumE THEN + kill_neq_tac 1 THEN kill_neq_tac 2 THEN + prod_tac args THEN sum_one_tac p) THEN + sum_tac cons' prems + | sum_tac _ _ = Imposs "theorems:sum_tac"; + in pg'' thy [] (exh (fn l => foldr (op ===>) (l,mk_trp(%"P"))) + (fn T => T ==> %"P") mk_All + (fn l => foldr (op ===>) (map mk_trp l,mk_trp(%"P"))) + bound_arg) + (fn prems => [ + cut_facts_tac [excluded_middle] 1, + etac disjE 1, + rtac (hd prems) 2, + etac rep_defin' 2, + if is_one_con_one_arg (not o is_lazy) cons + then rtac (hd (tl prems)) 1 THEN atac 2 THEN + rewrite_goals_tac axs_con_def THEN + simp_tac (HOLCF_ss addsimps [ax_rep_iso]) 1 + else sum_tac cons (tl prems)])end; +val exhaust = pg [] (mk_trp(exh (foldr' mk_disj) Id mk_ex (foldr' mk_conj) (K %))) [ + rtac cases 1, + UNTIL_SOLVED(fast_tac HOL_cs 1)]; +end; + +local +val when_app = foldl (op `) (%%(dname^"_when"), map % (when_funs cons)); +val when_appl = pg [ax_when_def] (mk_trp(when_app`%x_name===when_body cons + (fn (_,n) => %(nth_elem(n-1,when_funs cons)))`(dc_rep`%x_name))) [ + simp_tac HOLCF_ss 1]; +in +val when_strict = pg [] ((if is_one_con_one_arg (K true) cons + then fn t => mk_trp(strict(%"f")) ===> t else Id)(mk_trp(strict when_app))) [ + simp_tac(HOLCF_ss addsimps [when_appl,rep_strict]) 1]; +val when_apps = let fun one_when n (con,args) = pg axs_con_def + (lift_defined % (nonlazy args, mk_trp(when_app`(con_app con args) === + mk_cfapp(%(nth_elem(n,when_funs cons)),map %# args))))[ + asm_simp_tac (HOLCF_ss addsimps [when_appl,ax_abs_iso]) 1]; + in mapn one_when 0 cons end; +end; +val when_rews = when_strict::when_apps; + +(* ----- theorems concerning the constructors, discriminators and selectors ------- *) + +val dis_stricts = map (fn (con,_) => pg axs_dis_def (mk_trp( + (if is_one_con_one_arg (K true) cons then mk_not else Id) + (strict(%%(dis_name con))))) [ + simp_tac (HOLCF_ss addsimps (if is_one_con_one_arg (K true) cons + then [ax_when_def] else when_rews)) 1]) cons; +val dis_apps = let fun one_dis c (con,args)= pg (axs_dis_def) + (lift_defined % (nonlazy args, (*(if is_one_con_one_arg is_lazy cons + then curry (lift_defined %#) args else Id) +#################*) + (mk_trp((%%(dis_name c))`(con_app con args) === + %%(if con=c then "TT" else "FF"))))) [ + asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; + in flat(map (fn (c,_) => map (one_dis c) cons) cons) end; +val dis_defins = map (fn (con,args) => pg [] (defined(%x_name)==> + defined(%%(dis_name con)`%x_name)) [ + rtac cases 1, + contr_tac 1, + UNTIL_SOLVED (CHANGED(asm_simp_tac + (HOLCF_ss addsimps dis_apps) 1))]) cons; +val dis_rews = dis_stricts @ dis_defins @ dis_apps; + +val con_stricts = flat(map (fn (con,args) => map (fn vn => + pg (axs_con_def) + (mk_trp(con_app2 con (fn arg => if vname arg = vn + then UU else %# arg) args === UU))[ + asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1] + ) (nonlazy args)) cons); +val con_defins = map (fn (con,args) => pg [] + (lift_defined % (nonlazy args, + mk_trp(defined(con_app con args)))) ([ + rtac swap3 1] @ (if is_one_con_one_arg (K true) cons + then [ + if is_lazy (hd args) then rtac defined_up 2 + else atac 2, + rtac abs_defin' 1, + asm_full_simp_tac (HOLCF_ss addsimps axs_con_def) 1] + else [ + eres_inst_tac [("f",dis_name con)] cfun_arg_cong 1, + asm_simp_tac (HOLCF_ss addsimps dis_rews) 1])))cons; +val con_rews = con_stricts @ con_defins; + +val sel_stricts = let fun one_sel sel = pg axs_sel_def (mk_trp(strict(%%sel))) [ + simp_tac (HOLCF_ss addsimps when_rews) 1]; +in flat(map (fn (_,args) => map (fn arg => one_sel (sel_of arg)) args) cons) end; +val sel_apps = let fun one_sel c n sel = map (fn (con,args) => + let val nlas = nonlazy args; + val vns = map vname args; + in pg axs_sel_def (lift_defined % + (filter (fn v => con=c andalso (v<>nth_elem(n,vns))) nlas, + mk_trp((%%sel)`(con_app con args) === (if con=c then %(nth_elem(n,vns)) else UU)))) + ( (if con=c then [] + else map(case_UU_tac(when_rews@con_stricts)1) nlas) + @(if con=c andalso ((nth_elem(n,vns)) mem nlas) + then[case_UU_tac (when_rews @ con_stricts) 1 + (nth_elem(n,vns))] else []) + @ [asm_simp_tac(HOLCF_ss addsimps when_rews)1])end) cons; +in flat(map (fn (c,args) => + flat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end; +val sel_defins = if length cons = 1 then map (fn arg => pg [] (defined(%x_name) ==> + defined(%%(sel_of arg)`%x_name)) [ + rtac cases 1, + contr_tac 1, + UNTIL_SOLVED (CHANGED(asm_simp_tac + (HOLCF_ss addsimps sel_apps) 1))]) + (filter_out is_lazy (snd(hd cons))) else []; +val sel_rews = sel_stricts @ sel_defins @ sel_apps; + +val distincts_le = let + fun dist (con1, args1) (con2, args2) = pg [] + (lift_defined % ((nonlazy args1), + (mk_trp (con_app con1 args1 ~<< con_app con2 args2))))([ + rtac swap3 1, + eres_inst_tac [("fo5",dis_name con1)] monofun_cfun_arg 1] + @ map (case_UU_tac (con_stricts @ dis_rews) 1) (nonlazy args2) + @[asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]); + fun distinct (con1,args1) (con2,args2) = + let val arg1 = (con1, args1); + val arg2 = (con2, (map (fn (arg,vn) => upd_vname (K vn) arg) + (args2~~variantlist(map vname args2,map vname args1)))); + in [dist arg1 arg2, dist arg2 arg1] end; + fun distincts [] = [] + | distincts (c::cs) = (map (distinct c) cs) :: distincts cs; +in distincts cons end; +val dists_le = flat (flat distincts_le); +val dists_eq = let + fun distinct (_,args1) ((_,args2),leqs) = let + val (le1,le2) = (hd leqs, hd(tl leqs)); + val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) in + if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else + if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else + [eq1, eq2] end; + fun distincts [] = [] + | distincts ((c,leqs)::cs) = flat(map (distinct c) ((map fst cs)~~leqs)) @ + distincts cs; + in distincts (cons~~distincts_le) end; + +local + fun pgterm rel con args = let + fun append s = upd_vname(fn v => v^s); + val (largs,rargs) = (args, map (append "'") args); + in pg [] (mk_trp (rel(con_app con largs,con_app con rargs)) ===> + lift_defined % ((nonlazy largs),lift_defined % ((nonlazy rargs), + mk_trp (foldr' mk_conj + (map rel (map %# largs ~~ map %# rargs)))))) end; + val cons' = filter (fn (_,args) => args<>[]) cons; +in +val inverts = map (fn (con,args) => + pgterm (op <<) con args (flat(map (fn arg => [ + TRY(rtac conjI 1), + dres_inst_tac [("fo5",sel_of arg)] monofun_cfun_arg 1, + asm_full_simp_tac (HOLCF_ss addsimps sel_apps) 1] + ) args))) cons'; +val injects = map (fn ((con,args),inv_thm) => + pgterm (op ===) con args [ + etac (antisym_less_inverse RS conjE) 1, + dtac inv_thm 1, REPEAT(atac 1), + dtac inv_thm 1, REPEAT(atac 1), + TRY(safe_tac HOL_cs), + REPEAT(rtac antisym_less 1 ORELSE atac 1)] ) + (cons'~~inverts); +end; + +(* ----- theorems concerning one induction step ----------------------------------- *) + +val copy_strict = pg [ax_copy_def] ((if is_one_con_one_arg (K true) cons then fn t => + mk_trp(strict(cproj (%"f") eqs (rec_of (hd(snd(hd cons)))))) ===> t + else Id) (mk_trp(strict(dc_copy`%"f")))) [ + asm_simp_tac(HOLCF_ss addsimps [abs_strict,rep_strict, + cfst_strict,csnd_strict]) 1]; +val copy_apps = map (fn (con,args) => pg (ax_copy_def::axs_con_def) + (lift_defined %# (filter is_nonlazy_rec args, + mk_trp(dc_copy`%"f"`(con_app con args) === + (con_app2 con (app_rec_arg (cproj (%"f") eqs)) args)))) + (map (case_UU_tac [ax_abs_iso] 1 o vname) + (filter(fn a=>not(is_rec a orelse is_lazy a))args)@ + [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1]) + )cons; +val copy_stricts = map(fn(con,args)=>pg[](mk_trp(dc_copy`UU`(con_app con args) ===UU)) + (let val rews = cfst_strict::csnd_strict::copy_strict::copy_apps@con_rews + in map (case_UU_tac rews 1) (nonlazy args) @ [ + asm_simp_tac (HOLCF_ss addsimps rews) 1] end)) + (filter (fn (_,args)=>exists is_nonlazy_rec args) cons); +val copy_rews = copy_strict::copy_apps @ copy_stricts; + +in (iso_rews, exhaust, cases, when_rews, + con_rews, sel_rews, dis_rews, dists_eq, dists_le, inverts, injects, + copy_rews) +end; (* let *) + + +fun comp_theorems thy (comp_dname, eqs: eq list, casess, con_rews, copy_rews) = +let + +val dummy = writeln ("Proving induction properties of domain "^comp_dname^"..."); +val pg = pg' thy; + +val dnames = map (fst o fst) eqs; +val conss = map snd eqs; + +(* ----- getting the composite axiom and definitions ------------------------------ *) + +local val ga = get_axiom thy in +val axs_reach = map (fn dn => ga (dn ^ "_reach" )) dnames; +val axs_take_def = map (fn dn => ga (dn ^ "_take_def")) dnames; +val axs_finite_def = map (fn dn => ga (dn ^"_finite_def")) dnames; +val ax_copy2_def = ga (comp_dname^ "_copy_def"); +val ax_bisim_def = ga (comp_dname^"_bisim_def"); +end; (* local *) + +(* ----- theorems concerning finiteness and induction ----------------------------- *) + +fun dc_take dn = %%(dn^"_take"); +val x_name = idx_name dnames "x"; +val P_name = idx_name dnames "P"; + +local + val iterate_ss = simpset_of "Fix"; + val iterate_Cprod_strict_ss = iterate_ss addsimps [cfst_strict, csnd_strict]; + val iterate_Cprod_ss = iterate_ss addsimps [cfst2,csnd2,csplit2]; + val copy_con_rews = copy_rews @ con_rews; + val copy_take_defs = (if length dnames=1 then [] else [ax_copy2_def]) @axs_take_def; + val take_stricts = pg copy_take_defs (mk_trp(foldr' mk_conj (map (fn ((dn,args),_)=> + (dc_take dn $ %"n")`UU === mk_constrain(Type(dn,args),UU)) eqs)))([ + nat_ind_tac "n" 1, + simp_tac iterate_ss 1, + simp_tac iterate_Cprod_strict_ss 1, + asm_simp_tac iterate_Cprod_ss 1, + TRY(safe_tac HOL_cs)] @ + map(K(asm_simp_tac (HOL_ss addsimps copy_rews)1))dnames); + val take_stricts' = rewrite_rule copy_take_defs take_stricts; + val take_0s = mapn (fn n => fn dn => pg axs_take_def(mk_trp((dc_take dn $ %%"0") + `%x_name n === UU))[ + simp_tac iterate_Cprod_strict_ss 1]) 1 dnames; + val take_apps = pg copy_take_defs (mk_trp(foldr' mk_conj + (flat(map (fn ((dn,_),cons) => map (fn (con,args) => foldr mk_all + (map vname args,(dc_take dn $ (%%"Suc" $ %"n"))`(con_app con args) === + con_app2 con (app_rec_arg (fn n=>dc_take (nth_elem(n,dnames))$ %"n")) + args)) cons) eqs)))) ([ + nat_ind_tac "n" 1, + simp_tac iterate_Cprod_strict_ss 1, + simp_tac (HOLCF_ss addsimps copy_con_rews) 1, + TRY(safe_tac HOL_cs)] @ + (flat(map (fn ((dn,_),cons) => map (fn (con,args) => EVERY ( + asm_full_simp_tac iterate_Cprod_ss 1:: + map (case_UU_tac (take_stricts'::copy_con_rews) 1) + (nonlazy args) @[ + asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1]) + ) cons) eqs))); +in +val take_rews = atomize take_stricts @ take_0s @ atomize take_apps; +end; (* local *) + +val take_lemmas = mapn (fn n => fn(dn,ax_reach) => pg'' thy axs_take_def (mk_All("n", + mk_trp(dc_take dn $ Bound 0 `%(x_name n) === + dc_take dn $ Bound 0 `%(x_name n^"'"))) + ===> mk_trp(%(x_name n) === %(x_name n^"'"))) (fn prems => [ + res_inst_tac[("t",x_name n )](ax_reach RS subst) 1, + res_inst_tac[("t",x_name n^"'")](ax_reach RS subst) 1, + rtac (fix_def2 RS ssubst) 1, + REPEAT(CHANGED(rtac (contlub_cfun_arg RS ssubst) 1 + THEN chain_tac 1)), + rtac (contlub_cfun_fun RS ssubst) 1, + rtac (contlub_cfun_fun RS ssubst) 2, + rtac lub_equal 3, + chain_tac 1, + rtac allI 1, + resolve_tac prems 1])) 1 (dnames~~axs_reach); + +local + fun one_con p (con,args) = foldr mk_All (map vname args, + lift_defined (bound_arg (map vname args)) (nonlazy args, + lift (fn arg => %(P_name (1+rec_of arg)) $ bound_arg args arg) + (filter is_rec args,mk_trp(%p $ con_app2 con (bound_arg args) args)))); + fun one_eq ((p,cons),concl) = (mk_trp(%p $ UU) ===> + foldr (op ===>) (map (one_con p) cons,concl)); + fun ind_term concf = foldr one_eq (mapn (fn n => fn x => (P_name n, x)) 1 conss, + mk_trp(foldr' mk_conj (mapn (fn n => concf (P_name n,x_name n)) 1 dnames))); + val take_ss = HOL_ss addsimps take_rews; + fun ind_tacs tacsf thms1 thms2 prems = TRY(safe_tac HOL_cs):: + flat (mapn (fn n => fn (thm1,thm2) => + tacsf (n,prems) (thm1,thm2) @ + flat (map (fn cons => + (resolve_tac prems 1 :: + flat (map (fn (_,args) => + resolve_tac prems 1:: + map (K(atac 1)) (nonlazy args) @ + map (K(atac 1)) (filter is_rec args)) + cons))) + conss)) + 0 (thms1~~thms2)); + local + fun all_rec_to ns lazy_rec (n,cons) = forall (exists (fn arg => + is_rec arg andalso not(rec_of arg mem ns) andalso + ((rec_of arg = n andalso not(lazy_rec orelse is_lazy arg)) orelse + rec_of arg <> n andalso all_rec_to (rec_of arg::ns) + (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss)))) + ) o snd) cons; + fun warn (n,cons) = if all_rec_to [] false (n,cons) then (writeln + ("WARNING: domain "^nth_elem(n,dnames)^" is empty!"); true) + else false; + fun lazy_rec_to ns lazy_rec (n,cons) = exists (exists (fn arg => + is_rec arg andalso not(rec_of arg mem ns) andalso + ((rec_of arg = n andalso (lazy_rec orelse is_lazy arg)) orelse + rec_of arg <> n andalso lazy_rec_to (rec_of arg::ns) + (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss)))) + ) o snd) cons; + in val is_emptys = map warn (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs); + val is_finite = forall (not o lazy_rec_to [] false) + (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs) + end; +in +val finite_ind = pg'' thy [] (ind_term (fn (P,x) => fn dn => + mk_all(x,%P $ (dc_take dn $ %"n" `Bound 0)))) (fn prems=> [ + nat_ind_tac "n" 1, + simp_tac (take_ss addsimps prems) 1, + TRY(safe_tac HOL_cs)] + @ flat(mapn (fn n => fn (cons,cases) => [ + res_inst_tac [("x",x_name n)] cases 1, + asm_simp_tac (take_ss addsimps prems) 1] + @ flat(map (fn (con,args) => + asm_simp_tac take_ss 1 :: + map (fn arg => + case_UU_tac (prems@con_rews) 1 ( + nth_elem(rec_of arg,dnames)^"_take n1`"^vname arg)) + (filter is_nonlazy_rec args) @ [ + resolve_tac prems 1] @ + map (K (atac 1)) (nonlazy args) @ + map (K (etac spec 1)) (filter is_rec args)) + cons)) + 1 (conss~~casess))); + +val (finites,ind) = if is_finite then +let + fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %"x" === %"x"); + val finite_lemmas1a = map (fn dn => pg [] (mk_trp(defined (%"x")) ===> + mk_trp(mk_disj(mk_all("n",dc_take dn $ Bound 0 ` %"x" === UU), + take_enough dn)) ===> mk_trp(take_enough dn)) [ + etac disjE 1, + etac notE 1, + resolve_tac take_lemmas 1, + asm_simp_tac take_ss 1, + atac 1]) dnames; + val finite_lemma1b = pg [] (mk_trp (mk_all("n",foldr' mk_conj (mapn + (fn n => fn ((dn,args),_) => mk_constrainall(x_name n,Type(dn,args), + mk_disj(dc_take dn $ Bound 1 ` Bound 0 === UU, + dc_take dn $ Bound 1 ` Bound 0 === Bound 0))) 1 eqs)))) ([ + rtac allI 1, + nat_ind_tac "n" 1, + simp_tac take_ss 1, + TRY(safe_tac(empty_cs addSEs[conjE] addSIs[conjI]))] @ + flat(mapn (fn n => fn (cons,cases) => [ + simp_tac take_ss 1, + rtac allI 1, + res_inst_tac [("x",x_name n)] cases 1, + asm_simp_tac take_ss 1] @ + flat(map (fn (con,args) => + asm_simp_tac take_ss 1 :: + flat(map (fn arg => [ + eres_inst_tac [("x",vname arg)] all_dupE 1, + etac disjE 1, + asm_simp_tac (HOL_ss addsimps con_rews) 1, + asm_simp_tac take_ss 1]) + (filter is_nonlazy_rec args))) + cons)) + 1 (conss~~casess))) handle ERROR => raise ERROR; + val all_finite=map (fn(dn,l1b)=>pg axs_finite_def (mk_trp(%%(dn^"_finite") $ %"x"))[ + case_UU_tac take_rews 1 "x", + eresolve_tac finite_lemmas1a 1, + step_tac HOL_cs 1, + step_tac HOL_cs 1, + cut_facts_tac [l1b] 1, + fast_tac HOL_cs 1]) (dnames~~atomize finite_lemma1b); +in +(all_finite, + pg'' thy [] (ind_term (fn (P,x) => fn dn => %P $ %x)) + (ind_tacs (fn _ => fn (all_fin,finite_ind) => [ + rtac (rewrite_rule axs_finite_def all_fin RS exE) 1, + etac subst 1, + rtac finite_ind 1]) all_finite (atomize finite_ind)) +) end (* let *) else +(mapn (fn n => fn dn => read_instantiate_sg (sign_of thy) + [("P",dn^"_finite "^x_name n)] excluded_middle) 1 dnames, + pg'' thy [] (foldr (op ===>) (mapn (fn n =>K(mk_trp(%%"adm" $ %(P_name n))))1 + dnames,ind_term (fn(P,x)=>fn dn=> %P $ %x))) + (ind_tacs (fn (n,prems) => fn (ax_reach,finite_ind) =>[ + rtac (ax_reach RS subst) 1, + res_inst_tac [("x",x_name n)] spec 1, + rtac wfix_ind 1, + rtac adm_impl_admw 1, + resolve_tac adm_thms 1, + rtac adm_subst 1, + cont_tacR 1, + resolve_tac prems 1, + strip_tac 1, + rtac(rewrite_rule axs_take_def finite_ind) 1]) + axs_reach (atomize finite_ind)) +) +end; (* local *) + +local + val xs = mapn (fn n => K (x_name n)) 1 dnames; + fun bnd_arg n i = Bound(2*(length dnames - n)-i-1); + val take_ss = HOL_ss addsimps take_rews; + val sproj = bin_branchr (fn s => "fst("^s^")") (fn s => "snd("^s^")"); + val coind_lemma = pg [ax_bisim_def] (mk_trp(mk_imp(%%(comp_dname^"_bisim") $ %"R", + foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs, + foldr mk_imp (mapn (fn n => K(proj (%"R") dnames n $ + bnd_arg n 0 $ bnd_arg n 1)) 0 dnames, + foldr' mk_conj (mapn (fn n => fn dn => + (dc_take dn $ %"n" `bnd_arg n 0 === + (dc_take dn $ %"n" `bnd_arg n 1))) 0 dnames)))))) ([ + rtac impI 1, + nat_ind_tac "n" 1, + simp_tac take_ss 1, + safe_tac HOL_cs] @ + flat(mapn (fn n => fn x => [ + etac allE 1, etac allE 1, + eres_inst_tac [("P1",sproj "R" dnames n^ + " "^x^" "^x^"'")](mp RS disjE) 1, + TRY(safe_tac HOL_cs), + REPEAT(CHANGED(asm_simp_tac take_ss 1))]) + 0 xs)); +in +val coind = pg [] (mk_trp(%%(comp_dname^"_bisim") $ %"R") ===> + foldr (op ===>) (mapn (fn n => fn x => + mk_trp(proj (%"R") dnames n $ %x $ %(x^"'"))) 0 xs, + mk_trp(foldr' mk_conj (map (fn x => %x === %(x^"'")) xs)))) ([ + TRY(safe_tac HOL_cs)] @ + flat(map (fn take_lemma => [ + rtac take_lemma 1, + cut_facts_tac [coind_lemma] 1, + fast_tac HOL_cs 1]) + take_lemmas)); +end; (* local *) + + +in (take_rews, take_lemmas, finites, finite_ind, ind, coind) + +end; (* let *) +end; (* local *) +end; (* struct *) diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Coind.ML --- a/src/HOLCF/ex/Coind.ML Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,138 +0,0 @@ -(* Title: HOLCF/coind.ML - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen -*) - -open Coind; - -(* ------------------------------------------------------------------------- *) -(* expand fixed point properties *) -(* ------------------------------------------------------------------------- *) - - -val nats_def2 = fix_prover2 Coind.thy nats_def - "nats = scons`dzero`(smap`dsucc`nats)"; - -val from_def2 = fix_prover2 Coind.thy from_def - "from = (LAM n.scons`n`(from`(dsucc`n)))"; - - - -(* ------------------------------------------------------------------------- *) -(* recursive properties *) -(* ------------------------------------------------------------------------- *) - - -val from = prove_goal Coind.thy "from`n = scons`n`(from`(dsucc`n))" - (fn prems => - [ - (rtac trans 1), - (rtac (from_def2 RS ssubst) 1), - (Simp_tac 1), - (rtac refl 1) - ]); - - -val from1 = prove_goal Coind.thy "from`UU = UU" - (fn prems => - [ - (rtac trans 1), - (rtac (from RS ssubst) 1), - (resolve_tac stream_constrdef 1), - (rtac refl 1) - ]); - -val coind_rews = - [iterator1, iterator2, iterator3, smap1, smap2,from1]; - - -(* ------------------------------------------------------------------------- *) -(* the example *) -(* prove: nats = from`dzero *) -(* ------------------------------------------------------------------------- *) - - -val coind_lemma1 = prove_goal Coind.thy "iterator`n`(smap`dsucc)`nats =\ -\ scons`n`(iterator`(dsucc`n)`(smap`dsucc)`nats)" - (fn prems => - [ - (res_inst_tac [("s","n")] dnat_ind 1), - (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), - (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), - (rtac trans 1), - (rtac nats_def2 1), - (simp_tac (!simpset addsimps (coind_rews @ dnat_rews)) 1), - (rtac trans 1), - (etac iterator3 1), - (rtac trans 1), - (Asm_simp_tac 1), - (rtac trans 1), - (etac smap2 1), - (rtac cfun_arg_cong 1), - (asm_simp_tac (!simpset addsimps ([iterator3 RS sym] @ dnat_rews)) 1) - ]); - - -val nats_eq_from = prove_goal Coind.thy "nats = from`dzero" - (fn prems => - [ - (res_inst_tac [("R", -"% p q.? n. p = iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), - (res_inst_tac [("x","dzero")] exI 2), - (asm_simp_tac (!simpset addsimps coind_rews) 2), - (rewrite_goals_tac [stream_bisim_def]), - (strip_tac 1), - (etac exE 1), - (res_inst_tac [("Q","n=UU")] classical2 1), - (rtac disjI1 1), - (asm_simp_tac (!simpset addsimps coind_rews) 1), - (rtac disjI2 1), - (etac conjE 1), - (hyp_subst_tac 1), - (res_inst_tac [("x","n")] exI 1), - (res_inst_tac [("x","iterator`(dsucc`n)`(smap`dsucc)`nats")] exI 1), - (res_inst_tac [("x","from`(dsucc`n)")] exI 1), - (etac conjI 1), - (rtac conjI 1), - (rtac coind_lemma1 1), - (rtac conjI 1), - (rtac from 1), - (res_inst_tac [("x","dsucc`n")] exI 1), - (fast_tac HOL_cs 1) - ]); - -(* another proof using stream_coind_lemma2 *) - -val nats_eq_from = prove_goal Coind.thy "nats = from`dzero" - (fn prems => - [ - (res_inst_tac [("R","% p q.? n. p = \ -\ iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), - (rtac stream_coind_lemma2 1), - (strip_tac 1), - (etac exE 1), - (res_inst_tac [("Q","n=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps coind_rews) 1), - (res_inst_tac [("x","UU::dnat")] exI 1), - (simp_tac (!simpset addsimps coind_rews addsimps stream_rews) 1), - (etac conjE 1), - (hyp_subst_tac 1), - (rtac conjI 1), - (rtac (coind_lemma1 RS ssubst) 1), - (rtac (from RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (res_inst_tac [("x","dsucc`n")] exI 1), - (rtac conjI 1), - (rtac trans 1), - (rtac (coind_lemma1 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac refl 1), - (rtac trans 1), - (rtac (from RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac refl 1), - (res_inst_tac [("x","dzero")] exI 1), - (asm_simp_tac (!simpset addsimps coind_rews) 1) - ]); - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Coind.thy --- a/src/HOLCF/ex/Coind.thy Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,33 +0,0 @@ -(* Title: HOLCF/coind.thy - ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen - -Example for co-induction on streams -*) - -Coind = Stream2 + - - -consts - - nats :: "dnat stream" - from :: "dnat -> dnat stream" - -defs - nats_def "nats == fix`(LAM h.scons`dzero`(smap`dsucc`h))" - - from_def "from == fix`(LAM h n.scons`n`(h`(dsucc`n)))" - -end - -(* - smap`f`UU = UU - x~=UU --> smap`f`(scons`x`xs) = scons`(f`x)`(smap`f`xs) - - nats = scons`dzero`(smap`dsucc`nats) - - from`n = scons`n`(from`(dsucc`n)) -*) - - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Dagstuhl.ML --- a/src/HOLCF/ex/Dagstuhl.ML Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,75 +0,0 @@ -(* $Id$ *) - -open Dagstuhl; - -val YS_def2 = fix_prover2 Dagstuhl.thy YS_def "YS = scons`y`YS"; -val YYS_def2 = fix_prover2 Dagstuhl.thy YYS_def "YYS = scons`y`(scons`y`YYS)"; - - -val prems = goal Dagstuhl.thy "YYS << scons`y`YYS"; -by (rewrite_goals_tac [YYS_def]); -by (rtac fix_ind 1); -by (resolve_tac adm_thms 1); -by (cont_tacR 1); -by (rtac minimal 1); -by (rtac (beta_cfun RS ssubst) 1); -by (cont_tacR 1); -by (rtac monofun_cfun_arg 1); -by (rtac monofun_cfun_arg 1); -by (atac 1); -val lemma3 = result(); - -val prems = goal Dagstuhl.thy "scons`y`YYS << YYS"; -by (rtac (YYS_def2 RS ssubst) 1); -back(); -by (rtac monofun_cfun_arg 1); -by (rtac lemma3 1); -val lemma4=result(); - -(* val lemma5 = lemma3 RS (lemma4 RS antisym_less) *) - -val prems = goal Dagstuhl.thy "scons`y`YYS = YYS"; -by (rtac antisym_less 1); -by (rtac lemma4 1); -by (rtac lemma3 1); -val lemma5=result(); - -val prems = goal Dagstuhl.thy "YS = YYS"; -by (rtac stream_take_lemma 1); -by (nat_ind_tac "n" 1); -by (simp_tac (!simpset addsimps stream_rews) 1); -by (rtac (YS_def2 RS ssubst) 1); -by (rtac (YYS_def2 RS ssubst) 1); -by (asm_simp_tac (!simpset addsimps stream_rews) 1); -by (rtac (lemma5 RS sym RS subst) 1); -by (rtac refl 1); -val wir_moel=result(); - -(* ------------------------------------------------------------------------ *) -(* Zweite L"osung: Bernhard M"oller *) -(* statt Beweis von wir_moel "uber take_lemma beidseitige Inclusion *) -(* verwendet lemma5 *) -(* ------------------------------------------------------------------------ *) - -val prems = goal Dagstuhl.thy "YYS << YS"; -by (rewrite_goals_tac [YYS_def]); -by (rtac fix_least 1); -by (rtac (beta_cfun RS ssubst) 1); -by (cont_tacR 1); -by (simp_tac (!simpset addsimps [YS_def2 RS sym]) 1); -val lemma6=result(); - -val prems = goal Dagstuhl.thy "YS << YYS"; -by (rewrite_goals_tac [YS_def]); -by (rtac fix_ind 1); -by (resolve_tac adm_thms 1); -by (cont_tacR 1); -by (rtac minimal 1); -by (rtac (beta_cfun RS ssubst) 1); -by (cont_tacR 1); -by (rtac (lemma5 RS sym RS ssubst) 1); -by (etac monofun_cfun_arg 1); -val lemma7 = result(); - -val wir_moel = lemma6 RS (lemma7 RS antisym_less); - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Dagstuhl.thy --- a/src/HOLCF/ex/Dagstuhl.thy Fri Oct 06 16:17:08 1995 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,17 +0,0 @@ -(* $Id$ *) - - -Dagstuhl = Stream2 + - -consts - y :: "'a" - YS :: "'a stream" - YYS :: "'a stream" - -defs - -YS_def "YS == fix`(LAM x. scons`y`x)" -YYS_def "YYS == fix`(LAM z. scons`y`(scons`y`z))" - -end - diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Fix2.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ex/Fix2.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,28 @@ +(* Title: HOLCF/ex/Fix2.ML + ID: $Id$ + Author: Franz Regensburger + Copyright 1995 Technische Universitaet Muenchen +*) + +open Fix2; + +val lemma1 = prove_goal Fix2.thy "fix = gix" + (fn prems => + [ + (rtac ext_cfun 1), + (rtac antisym_less 1), + (rtac fix_least 1), + (rtac gix1_def 1), + (rtac gix2_def 1), + (rtac (fix_eq RS sym) 1) + ]); + + +val lemma2 = prove_goal Fix2.thy "gix`F=lub(range(%i. iterate i F UU))" + (fn prems => + [ + (rtac (lemma1 RS subst) 1), + (rtac fix_def2 1) + ]); + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Fix2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/ex/Fix2.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,38 @@ +(* Title: HOLCF/ex/Fix2.thy + ID: $Id$ + Author: Franz Regensburger + Copyright 1995 Technische Universitaet Muenchen + + Show that fix is the unique least fixed-point operator. + From axioms gix1_def,gix2_def it follows that fix = gix + +*) + +Fix2 = Fix + + +consts + + gix :: "('a->'a)->'a" + +rules + +gix1_def "F`(gix`F) = gix`F" +gix2_def "F`y=y ==> gix`F << y" + +end + + + + + + + + + + + + + + + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Hoare.ML --- a/src/HOLCF/ex/Hoare.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/ex/Hoare.ML Fri Oct 06 17:25:24 1995 +0100 @@ -137,8 +137,6 @@ (** --------- proves about iterations of p and q ---------- **) -val HOLCF_ss = simpset_of "HOLCF"; - val hoare_lemma9 = prove_goal Hoare.thy "(! m. m< Suc k --> b1`(iterate m g x)=TT) -->\ \ p`(iterate k g x)=p`x" diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Loop.ML --- a/src/HOLCF/ex/Loop.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/ex/Loop.ML Fri Oct 06 17:25:24 1995 +0100 @@ -1,4 +1,4 @@ -(* Title: HOLCF/ex/loop.ML +(* Title: HOLCF/ex/Loop.ML ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen @@ -39,8 +39,6 @@ (Simp_tac 1) ]); -val HOLCF_ss = simpset_of "HOLCF"; - val while_unfold2 = prove_goal Loop.thy "!x.while`b`g`x = while`b`g`(iterate k (step`b`g) x)" (fn prems => diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/Loop.thy --- a/src/HOLCF/ex/Loop.thy Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/ex/Loop.thy Fri Oct 06 17:25:24 1995 +0100 @@ -1,4 +1,4 @@ -(* Title: HOLCF/ex/loop.thy +(* Title: HOLCF/ex/Loop.thy ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/ex/ROOT.ML --- a/src/HOLCF/ex/ROOT.ML Fri Oct 06 16:17:08 1995 +0100 +++ b/src/HOLCF/ex/ROOT.ML Fri Oct 06 17:25:24 1995 +0100 @@ -10,10 +10,9 @@ writeln"Root file for HOLCF examples"; proof_timing := true; -time_use_thy "ex/Coind"; time_use_thy "ex/Hoare"; time_use_thy "ex/Loop"; -time_use_thy "ex/Dagstuhl"; +time_use_thy "ex/Fix2"; time_use "ex/loeckx.ML"; maketest "END: Root file for HOLCF examples"; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Coind.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Coind.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,138 @@ +(* Title: HOLCF/Coind.ML + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen +*) + +open Coind; + +(* ------------------------------------------------------------------------- *) +(* expand fixed point properties *) +(* ------------------------------------------------------------------------- *) + + +val nats_def2 = fix_prover2 Coind.thy nats_def + "nats = scons`dzero`(smap`dsucc`nats)"; + +val from_def2 = fix_prover2 Coind.thy from_def + "from = (LAM n.scons`n`(from`(dsucc`n)))"; + + + +(* ------------------------------------------------------------------------- *) +(* recursive properties *) +(* ------------------------------------------------------------------------- *) + + +val from = prove_goal Coind.thy "from`n = scons`n`(from`(dsucc`n))" + (fn prems => + [ + (rtac trans 1), + (rtac (from_def2 RS ssubst) 1), + (Simp_tac 1), + (rtac refl 1) + ]); + + +val from1 = prove_goal Coind.thy "from`UU = UU" + (fn prems => + [ + (rtac trans 1), + (rtac (from RS ssubst) 1), + (resolve_tac stream_constrdef 1), + (rtac refl 1) + ]); + +val coind_rews = + [iterator1, iterator2, iterator3, smap1, smap2,from1]; + + +(* ------------------------------------------------------------------------- *) +(* the example *) +(* prove: nats = from`dzero *) +(* ------------------------------------------------------------------------- *) + + +val coind_lemma1 = prove_goal Coind.thy "iterator`n`(smap`dsucc)`nats =\ +\ scons`n`(iterator`(dsucc`n)`(smap`dsucc)`nats)" + (fn prems => + [ + (res_inst_tac [("s","n")] dnat_ind 1), + (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), + (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), + (rtac trans 1), + (rtac nats_def2 1), + (simp_tac (!simpset addsimps (coind_rews @ dnat_rews)) 1), + (rtac trans 1), + (etac iterator3 1), + (rtac trans 1), + (Asm_simp_tac 1), + (rtac trans 1), + (etac smap2 1), + (rtac cfun_arg_cong 1), + (asm_simp_tac (!simpset addsimps ([iterator3 RS sym] @ dnat_rews)) 1) + ]); + + +val nats_eq_from = prove_goal Coind.thy "nats = from`dzero" + (fn prems => + [ + (res_inst_tac [("R", +"% p q.? n. p = iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), + (res_inst_tac [("x","dzero")] exI 2), + (asm_simp_tac (!simpset addsimps coind_rews) 2), + (rewrite_goals_tac [stream_bisim_def]), + (strip_tac 1), + (etac exE 1), + (res_inst_tac [("Q","n=UU")] classical2 1), + (rtac disjI1 1), + (asm_simp_tac (!simpset addsimps coind_rews) 1), + (rtac disjI2 1), + (etac conjE 1), + (hyp_subst_tac 1), + (res_inst_tac [("x","n")] exI 1), + (res_inst_tac [("x","iterator`(dsucc`n)`(smap`dsucc)`nats")] exI 1), + (res_inst_tac [("x","from`(dsucc`n)")] exI 1), + (etac conjI 1), + (rtac conjI 1), + (rtac coind_lemma1 1), + (rtac conjI 1), + (rtac from 1), + (res_inst_tac [("x","dsucc`n")] exI 1), + (fast_tac HOL_cs 1) + ]); + +(* another proof using stream_coind_lemma2 *) + +val nats_eq_from = prove_goal Coind.thy "nats = from`dzero" + (fn prems => + [ + (res_inst_tac [("R","% p q.? n. p = \ +\ iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), + (rtac stream_coind_lemma2 1), + (strip_tac 1), + (etac exE 1), + (res_inst_tac [("Q","n=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps coind_rews) 1), + (res_inst_tac [("x","UU::dnat")] exI 1), + (simp_tac (!simpset addsimps coind_rews addsimps stream_rews) 1), + (etac conjE 1), + (hyp_subst_tac 1), + (rtac conjI 1), + (rtac (coind_lemma1 RS ssubst) 1), + (rtac (from RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (res_inst_tac [("x","dsucc`n")] exI 1), + (rtac conjI 1), + (rtac trans 1), + (rtac (coind_lemma1 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac refl 1), + (rtac trans 1), + (rtac (from RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac refl 1), + (res_inst_tac [("x","dzero")] exI 1), + (asm_simp_tac (!simpset addsimps coind_rews) 1) + ]); + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Coind.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Coind.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,33 @@ +(* Title: HOLCF/Coind.thy + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Example for co-induction on streams +*) + +Coind = Stream2 + + + +consts + + nats :: "dnat stream" + from :: "dnat -> dnat stream" + +defs + nats_def "nats == fix`(LAM h.scons`dzero`(smap`dsucc`h))" + + from_def "from == fix`(LAM h n.scons`n`(h`(dsucc`n)))" + +end + +(* + smap`f`UU = UU + x~=UU --> smap`f`(scons`x`xs) = scons`(f`x)`(smap`f`xs) + + nats = scons`dzero`(smap`dsucc`nats) + + from`n = scons`n`(from`(dsucc`n)) +*) + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Dagstuhl.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Dagstuhl.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,75 @@ +(* $Id$ *) + +open Dagstuhl; + +val YS_def2 = fix_prover2 Dagstuhl.thy YS_def "YS = scons`y`YS"; +val YYS_def2 = fix_prover2 Dagstuhl.thy YYS_def "YYS = scons`y`(scons`y`YYS)"; + + +val prems = goal Dagstuhl.thy "YYS << scons`y`YYS"; +by (rewrite_goals_tac [YYS_def]); +by (rtac fix_ind 1); +by (resolve_tac adm_thms 1); +by (cont_tacR 1); +by (rtac minimal 1); +by (rtac (beta_cfun RS ssubst) 1); +by (cont_tacR 1); +by (rtac monofun_cfun_arg 1); +by (rtac monofun_cfun_arg 1); +by (atac 1); +val lemma3 = result(); + +val prems = goal Dagstuhl.thy "scons`y`YYS << YYS"; +by (rtac (YYS_def2 RS ssubst) 1); +back(); +by (rtac monofun_cfun_arg 1); +by (rtac lemma3 1); +val lemma4=result(); + +(* val lemma5 = lemma3 RS (lemma4 RS antisym_less) *) + +val prems = goal Dagstuhl.thy "scons`y`YYS = YYS"; +by (rtac antisym_less 1); +by (rtac lemma4 1); +by (rtac lemma3 1); +val lemma5=result(); + +val prems = goal Dagstuhl.thy "YS = YYS"; +by (rtac stream_take_lemma 1); +by (nat_ind_tac "n" 1); +by (simp_tac (!simpset addsimps stream_rews) 1); +by (rtac (YS_def2 RS ssubst) 1); +by (rtac (YYS_def2 RS ssubst) 1); +by (asm_simp_tac (!simpset addsimps stream_rews) 1); +by (rtac (lemma5 RS sym RS subst) 1); +by (rtac refl 1); +val wir_moel=result(); + +(* ------------------------------------------------------------------------ *) +(* Zweite L"osung: Bernhard M"oller *) +(* statt Beweis von wir_moel "uber take_lemma beidseitige Inclusion *) +(* verwendet lemma5 *) +(* ------------------------------------------------------------------------ *) + +val prems = goal Dagstuhl.thy "YYS << YS"; +by (rewrite_goals_tac [YYS_def]); +by (rtac fix_least 1); +by (rtac (beta_cfun RS ssubst) 1); +by (cont_tacR 1); +by (simp_tac (!simpset addsimps [YS_def2 RS sym]) 1); +val lemma6=result(); + +val prems = goal Dagstuhl.thy "YS << YYS"; +by (rewrite_goals_tac [YS_def]); +by (rtac fix_ind 1); +by (resolve_tac adm_thms 1); +by (cont_tacR 1); +by (rtac minimal 1); +by (rtac (beta_cfun RS ssubst) 1); +by (cont_tacR 1); +by (rtac (lemma5 RS sym RS ssubst) 1); +by (etac monofun_cfun_arg 1); +val lemma7 = result(); + +val wir_moel = lemma6 RS (lemma7 RS antisym_less); + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Dagstuhl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Dagstuhl.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,17 @@ +(* $Id$ *) + + +Dagstuhl = Stream2 + + +consts + y :: "'a" + YS :: "'a stream" + YYS :: "'a stream" + +defs + +YS_def "YS == fix`(LAM x. scons`y`x)" +YYS_def "YYS == fix`(LAM z. scons`y`(scons`y`z))" + +end + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Dlist.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Dlist.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,564 @@ +(* Title: HOLCF/Dlist.ML + Author: Franz Regensburger + ID: $ $ + Copyright 1994 Technische Universitaet Muenchen + +Lemmas for dlist.thy +*) + +open Dlist; + +(* ------------------------------------------------------------------------*) +(* The isomorphisms dlist_rep_iso and dlist_abs_iso are strict *) +(* ------------------------------------------------------------------------*) + +val dlist_iso_strict= dlist_rep_iso RS (dlist_abs_iso RS + (allI RSN (2,allI RS iso_strict))); + +val dlist_rews = [dlist_iso_strict RS conjunct1, + dlist_iso_strict RS conjunct2]; + +(* ------------------------------------------------------------------------*) +(* Properties of dlist_copy *) +(* ------------------------------------------------------------------------*) + +val temp = prove_goalw Dlist.thy [dlist_copy_def] "dlist_copy`f`UU=UU" + (fn prems => + [ + (asm_simp_tac (!simpset addsimps + (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) + ]); + +val dlist_copy = [temp]; + + +val temp = prove_goalw Dlist.thy [dlist_copy_def,dnil_def] + "dlist_copy`f`dnil=dnil" + (fn prems => + [ + (asm_simp_tac (!simpset addsimps + (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) + ]); + +val dlist_copy = temp :: dlist_copy; + + +val temp = prove_goalw Dlist.thy [dlist_copy_def,dcons_def] + "xl~=UU ==> dlist_copy`f`(dcons`x`xl)= dcons`x`(f`xl)" + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1), + (res_inst_tac [("Q","x=UU")] classical2 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps [defined_spair]) 1) + ]); + +val dlist_copy = temp :: dlist_copy; + +val dlist_rews = dlist_copy @ dlist_rews; + +(* ------------------------------------------------------------------------*) +(* Exhaustion and elimination for dlists *) +(* ------------------------------------------------------------------------*) + +qed_goalw "Exh_dlist" Dlist.thy [dcons_def,dnil_def] + "l = UU | l = dnil | (? x xl. x~=UU &xl~=UU & l = dcons`x`xl)" + (fn prems => + [ + (Simp_tac 1), + (rtac (dlist_rep_iso RS subst) 1), + (res_inst_tac [("p","dlist_rep`l")] ssumE 1), + (rtac disjI1 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (rtac disjI2 1), + (rtac disjI1 1), + (res_inst_tac [("p","x")] oneE 1), + (contr_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (rtac disjI2 1), + (rtac disjI2 1), + (res_inst_tac [("p","y")] sprodE 1), + (contr_tac 1), + (rtac exI 1), + (rtac exI 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (fast_tac HOL_cs 1) + ]); + + +qed_goal "dlistE" Dlist.thy +"[| l=UU ==> Q; l=dnil ==> Q;!!x xl.[|l=dcons`x`xl;x~=UU;xl~=UU|]==>Q|]==>Q" + (fn prems => + [ + (rtac (Exh_dlist RS disjE) 1), + (eresolve_tac prems 1), + (etac disjE 1), + (eresolve_tac prems 1), + (etac exE 1), + (etac exE 1), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); + +(* ------------------------------------------------------------------------*) +(* Properties of dlist_when *) +(* ------------------------------------------------------------------------*) + +val temp = prove_goalw Dlist.thy [dlist_when_def] "dlist_when`f1`f2`UU=UU" + (fn prems => + [ + (asm_simp_tac (!simpset addsimps [dlist_iso_strict]) 1) + ]); + +val dlist_when = [temp]; + +val temp = prove_goalw Dlist.thy [dlist_when_def,dnil_def] + "dlist_when`f1`f2`dnil= f1" + (fn prems => + [ + (asm_simp_tac (!simpset addsimps [dlist_abs_iso]) 1) + ]); + +val dlist_when = temp::dlist_when; + +val temp = prove_goalw Dlist.thy [dlist_when_def,dcons_def] + "[|x~=UU;xl~=UU|] ==> dlist_when`f1`f2`(dcons`x`xl)= f2`x`xl" + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps [dlist_abs_iso,defined_spair]) 1) + ]); + +val dlist_when = temp::dlist_when; + +val dlist_rews = dlist_when @ dlist_rews; + +(* ------------------------------------------------------------------------*) +(* Rewrites for discriminators and selectors *) +(* ------------------------------------------------------------------------*) + +fun prover defs thm = prove_goalw Dlist.thy defs thm + (fn prems => + [ + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_discsel = [ + prover [is_dnil_def] "is_dnil`UU=UU", + prover [is_dcons_def] "is_dcons`UU=UU", + prover [dhd_def] "dhd`UU=UU", + prover [dtl_def] "dtl`UU=UU" + ]; + +fun prover defs thm = prove_goalw Dlist.thy defs thm + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_discsel = [ +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "is_dnil`dnil=TT", +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "[|x~=UU;xl~=UU|] ==> is_dnil`(dcons`x`xl)=FF", +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "is_dcons`dnil=FF", +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "[|x~=UU;xl~=UU|] ==> is_dcons`(dcons`x`xl)=TT", +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "dhd`dnil=UU", +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "[|x~=UU;xl~=UU|] ==> dhd`(dcons`x`xl)=x", +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "dtl`dnil=UU", +prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] + "[|x~=UU;xl~=UU|] ==> dtl`(dcons`x`xl)=xl"] @ dlist_discsel; + +val dlist_rews = dlist_discsel @ dlist_rews; + +(* ------------------------------------------------------------------------*) +(* Definedness and strictness *) +(* ------------------------------------------------------------------------*) + +fun prover contr thm = prove_goal Dlist.thy thm + (fn prems => + [ + (res_inst_tac [("P1",contr)] classical3 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (dtac sym 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps (prems @ dlist_rews)) 1) + ]); + + +val dlist_constrdef = [ +prover "is_dnil`(UU::'a dlist) ~= UU" "dnil~=(UU::'a dlist)", +prover "is_dcons`(UU::'a dlist) ~= UU" + "[|x~=UU;xl~=UU|]==>dcons`(x::'a)`xl ~= UU" + ]; + + +fun prover defs thm = prove_goalw Dlist.thy defs thm + (fn prems => + [ + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_constrdef = [ + prover [dcons_def] "dcons`UU`xl=UU", + prover [dcons_def] "dcons`x`UU=UU" + ] @ dlist_constrdef; + +val dlist_rews = dlist_constrdef @ dlist_rews; + + +(* ------------------------------------------------------------------------*) +(* Distinctness wrt. << and = *) +(* ------------------------------------------------------------------------*) + +val temp = prove_goal Dlist.thy "~dnil << dcons`(x::'a)`xl" + (fn prems => + [ + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dnil")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","(x::'a)=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","(xl ::'a dlist)=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_dist_less = [temp]; + +val temp = prove_goal Dlist.thy "[|x~=UU;xl~=UU|]==>~ dcons`x`xl << dnil" + (fn prems => + [ + (cut_facts_tac prems 1), + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dcons")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_dist_less = temp::dlist_dist_less; + +val temp = prove_goal Dlist.thy "dnil ~= dcons`x`xl" + (fn prems => + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("P1","TT = FF")] classical3 1), + (resolve_tac dist_eq_tr 1), + (dres_inst_tac [("f","is_dnil")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_dist_eq = [temp,temp RS not_sym]; + +val dlist_rews = dlist_dist_less @ dlist_dist_eq @ dlist_rews; + +(* ------------------------------------------------------------------------*) +(* Invertibility *) +(* ------------------------------------------------------------------------*) + +val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ +\ dcons`x1`x2 << dcons`y1`y2 |] ==> x1<< y1 & x2 << y2" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.x)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.l)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1) + ]); + +val dlist_invert =[temp]; + +(* ------------------------------------------------------------------------*) +(* Injectivity *) +(* ------------------------------------------------------------------------*) + +val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ +\ dcons`x1`x2 = dcons`y1`y2 |] ==> x1= y1 & x2 = y2" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("f","dlist_when`UU`(LAM x l.x)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (dres_inst_tac [("f","dlist_when`UU`(LAM x l.l)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1) + ]); + +val dlist_inject = [temp]; + + +(* ------------------------------------------------------------------------*) +(* definedness for discriminators and selectors *) +(* ------------------------------------------------------------------------*) + +fun prover thm = prove_goal Dlist.thy thm + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac dlistE 1), + (contr_tac 1), + (REPEAT (asm_simp_tac (!simpset addsimps dlist_discsel) 1)) + ]); + +val dlist_discsel_def = + [ + prover "l~=UU ==> is_dnil`l~=UU", + prover "l~=UU ==> is_dcons`l~=UU" + ]; + +val dlist_rews = dlist_discsel_def @ dlist_rews; + +(* ------------------------------------------------------------------------*) +(* enhance the simplifier *) +(* ------------------------------------------------------------------------*) + +qed_goal "dhd2" Dlist.thy "xl~=UU ==> dhd`(dcons`x`xl)=x" + (fn prems => + [ + (cut_facts_tac prems 1), + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +qed_goal "dtl2" Dlist.thy "x~=UU ==> dtl`(dcons`x`xl)=xl" + (fn prems => + [ + (cut_facts_tac prems 1), + (res_inst_tac [("Q","xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_rews = dhd2 :: dtl2 :: dlist_rews; + +(* ------------------------------------------------------------------------*) +(* Properties dlist_take *) +(* ------------------------------------------------------------------------*) + +val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take n`UU=UU" + (fn prems => + [ + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_take = [temp]; + +val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take 0`xs=UU" + (fn prems => + [ + (Asm_simp_tac 1) + ]); + +val dlist_take = temp::dlist_take; + +val temp = prove_goalw Dlist.thy [dlist_take_def] + "dlist_take (Suc n)`dnil=dnil" + (fn prems => + [ + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_take = temp::dlist_take; + +val temp = prove_goalw Dlist.thy [dlist_take_def] + "dlist_take (Suc n)`(dcons`x`xl)= dcons`x`(dlist_take n`xl)" + (fn prems => + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +val dlist_take = temp::dlist_take; + +val dlist_rews = dlist_take @ dlist_rews; + +(* ------------------------------------------------------------------------*) +(* take lemma for dlists *) +(* ------------------------------------------------------------------------*) + +fun prover reach defs thm = prove_goalw Dlist.thy defs thm + (fn prems => + [ + (res_inst_tac [("t","l1")] (reach RS subst) 1), + (res_inst_tac [("t","l2")] (reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac lub_equal 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac allI 1), + (resolve_tac prems 1) + ]); + +val dlist_take_lemma = prover dlist_reach [dlist_take_def] + "(!!n.dlist_take n`l1 = dlist_take n`l2) ==> l1=l2"; + + +(* ------------------------------------------------------------------------*) +(* Co -induction for dlists *) +(* ------------------------------------------------------------------------*) + +qed_goalw "dlist_coind_lemma" Dlist.thy [dlist_bisim_def] +"dlist_bisim R ==> ! p q. R p q --> dlist_take n`p = dlist_take n`q" + (fn prems => + [ + (cut_facts_tac prems 1), + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (strip_tac 1), + ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), + (atac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (etac exE 1), + (etac exE 1), + (etac exE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (REPEAT (etac conjE 1)), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); + +qed_goal "dlist_coind" Dlist.thy "[|dlist_bisim R ; R p q |] ==> p = q" + (fn prems => + [ + (rtac dlist_take_lemma 1), + (rtac (dlist_coind_lemma RS spec RS spec RS mp) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); + +(* ------------------------------------------------------------------------*) +(* structural induction *) +(* ------------------------------------------------------------------------*) + +qed_goal "dlist_finite_ind" Dlist.thy +"[|P(UU);P(dnil);\ +\ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons`x`l1)\ +\ |] ==> !l.P(dlist_take n`l)" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (rtac allI 1), + (res_inst_tac [("l","l")] dlistE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","dlist_take n1`xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (atac 1), + (atac 1), + (etac spec 1) + ]); + +qed_goal "dlist_all_finite_lemma1" Dlist.thy +"!l.dlist_take n`l=UU |dlist_take n`l=l" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (rtac allI 1), + (res_inst_tac [("l","l")] dlistE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (eres_inst_tac [("x","xl")] allE 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); + +qed_goal "dlist_all_finite_lemma2" Dlist.thy "? n.dlist_take n`l=l" + (fn prems => + [ + (res_inst_tac [("Q","l=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (subgoal_tac "(!n.dlist_take n`l=UU) |(? n.dlist_take n`l = l)" 1), + (etac disjE 1), + (eres_inst_tac [("P","l=UU")] notE 1), + (rtac dlist_take_lemma 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (atac 1), + (subgoal_tac "!n.!l.dlist_take n`l=UU |dlist_take n`l=l" 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (rtac dlist_all_finite_lemma1 1) + ]); + +qed_goalw "dlist_all_finite" Dlist.thy [dlist_finite_def] "dlist_finite(l)" + (fn prems => + [ + (rtac dlist_all_finite_lemma2 1) + ]); + +qed_goal "dlist_ind" Dlist.thy +"[|P(UU);P(dnil);\ +\ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons`x`l1)|] ==> P(l)" + (fn prems => + [ + (rtac (dlist_all_finite_lemma2 RS exE) 1), + (etac subst 1), + (rtac (dlist_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); + + + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Dlist.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Dlist.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,126 @@ +(* Title: HOLCF/Dlist.thy + + Author: Franz Regensburger + ID: $ $ + Copyright 1994 Technische Universitaet Muenchen + +Theory for finite lists 'a dlist = one ++ ('a ** 'a dlist) + +The type is axiomatized as the least solution of the domain equation above. +The functor term that specifies the domain equation is: + + FT = <++,K_{one},<**,K_{'a},I>> + +For details see chapter 5 of: + +[Franz Regensburger] HOLCF: Eine konservative Erweiterung von HOL um LCF, + Dissertation, Technische Universit"at M"unchen, 1994 + + +*) + +Dlist = Stream2 + + +types dlist 1 + +(* ----------------------------------------------------------------------- *) +(* arity axiom is validated by semantic reasoning *) +(* partial ordering is implicit in the isomorphism axioms and their cont. *) + +arities dlist::(pcpo)pcpo + +consts + +(* ----------------------------------------------------------------------- *) +(* essential constants *) + +dlist_rep :: "('a dlist) -> (one ++ 'a ** 'a dlist)" +dlist_abs :: "(one ++ 'a ** 'a dlist) -> ('a dlist)" + +(* ----------------------------------------------------------------------- *) +(* abstract constants and auxiliary constants *) + +dlist_copy :: "('a dlist -> 'a dlist) ->'a dlist -> 'a dlist" + +dnil :: "'a dlist" +dcons :: "'a -> 'a dlist -> 'a dlist" +dlist_when :: " 'b -> ('a -> 'a dlist -> 'b) -> 'a dlist -> 'b" +is_dnil :: "'a dlist -> tr" +is_dcons :: "'a dlist -> tr" +dhd :: "'a dlist -> 'a" +dtl :: "'a dlist -> 'a dlist" +dlist_take :: "nat => 'a dlist -> 'a dlist" +dlist_finite :: "'a dlist => bool" +dlist_bisim :: "('a dlist => 'a dlist => bool) => bool" + +rules + +(* ----------------------------------------------------------------------- *) +(* axiomatization of recursive type 'a dlist *) +(* ----------------------------------------------------------------------- *) +(* ('a dlist,dlist_abs) is the initial F-algebra where *) +(* F is the locally continuous functor determined by functor term FT. *) +(* domain equation: 'a dlist = one ++ ('a ** 'a dlist) *) +(* functor term: FT = <++,K_{one},<**,K_{'a},I>> *) +(* ----------------------------------------------------------------------- *) +(* dlist_abs is an isomorphism with inverse dlist_rep *) +(* identity is the least endomorphism on 'a dlist *) + +dlist_abs_iso "dlist_rep`(dlist_abs`x) = x" +dlist_rep_iso "dlist_abs`(dlist_rep`x) = x" +dlist_copy_def "dlist_copy == (LAM f. dlist_abs oo \ +\ (sswhen`sinl`(sinr oo (ssplit`(LAM x y. (|x,f`y|) ))))\ +\ oo dlist_rep)" +dlist_reach "(fix`dlist_copy)`x=x" + + +defs +(* ----------------------------------------------------------------------- *) +(* properties of additional constants *) +(* ----------------------------------------------------------------------- *) +(* constructors *) + +dnil_def "dnil == dlist_abs`(sinl`one)" +dcons_def "dcons == (LAM x l. dlist_abs`(sinr`(|x,l|) ))" + +(* ----------------------------------------------------------------------- *) +(* discriminator functional *) + +dlist_when_def +"dlist_when == (LAM f1 f2 l.\ +\ sswhen`(LAM x.f1) `(ssplit`(LAM x l.f2`x`l)) `(dlist_rep`l))" + +(* ----------------------------------------------------------------------- *) +(* discriminators and selectors *) + +is_dnil_def "is_dnil == dlist_when`TT`(LAM x l.FF)" +is_dcons_def "is_dcons == dlist_when`FF`(LAM x l.TT)" +dhd_def "dhd == dlist_when`UU`(LAM x l.x)" +dtl_def "dtl == dlist_when`UU`(LAM x l.l)" + +(* ----------------------------------------------------------------------- *) +(* the taker for dlists *) + +dlist_take_def "dlist_take == (%n.iterate n dlist_copy UU)" + +(* ----------------------------------------------------------------------- *) + +dlist_finite_def "dlist_finite == (%s.? n.dlist_take n`s=s)" + +(* ----------------------------------------------------------------------- *) +(* definition of bisimulation is determined by domain equation *) +(* simplification and rewriting for abstract constants yields def below *) + +dlist_bisim_def "dlist_bisim == + ( %R.!l1 l2. + R l1 l2 --> + ((l1=UU & l2=UU) | + (l1=dnil & l2=dnil) | + (? x l11 l21. x~=UU & l11~=UU & l21~=UU & + l1=dcons`x`l11 & l2 = dcons`x`l21 & R l11 l21)))" + +end + + + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Dnat.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Dnat.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,534 @@ +(* Title: HOLCF/Dnat.ML + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Lemmas for dnat.thy +*) + +open Dnat; + +(* ------------------------------------------------------------------------*) +(* The isomorphisms dnat_rep_iso and dnat_abs_iso are strict *) +(* ------------------------------------------------------------------------*) + +val dnat_iso_strict = dnat_rep_iso RS (dnat_abs_iso RS + (allI RSN (2,allI RS iso_strict))); + +val dnat_rews = [dnat_iso_strict RS conjunct1, + dnat_iso_strict RS conjunct2]; + +(* ------------------------------------------------------------------------*) +(* Properties of dnat_copy *) +(* ------------------------------------------------------------------------*) + +fun prover defs thm = prove_goalw Dnat.thy defs thm + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) + ]); + +val dnat_copy = + [ + prover [dnat_copy_def] "dnat_copy`f`UU=UU", + prover [dnat_copy_def,dzero_def] "dnat_copy`f`dzero= dzero", + prover [dnat_copy_def,dsucc_def] + "n~=UU ==> dnat_copy`f`(dsucc`n) = dsucc`(f`n)" + ]; + +val dnat_rews = dnat_copy @ dnat_rews; + +(* ------------------------------------------------------------------------*) +(* Exhaustion and elimination for dnat *) +(* ------------------------------------------------------------------------*) + +qed_goalw "Exh_dnat" Dnat.thy [dsucc_def,dzero_def] + "n = UU | n = dzero | (? x . x~=UU & n = dsucc`x)" + (fn prems => + [ + (Simp_tac 1), + (rtac (dnat_rep_iso RS subst) 1), + (res_inst_tac [("p","dnat_rep`n")] ssumE 1), + (rtac disjI1 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (rtac (disjI1 RS disjI2) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("p","x")] oneE 1), + (contr_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (rtac (disjI2 RS disjI2) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (fast_tac HOL_cs 1) + ]); + +qed_goal "dnatE" Dnat.thy + "[| n=UU ==> Q; n=dzero ==> Q; !!x.[|n=dsucc`x;x~=UU|]==>Q|]==>Q" + (fn prems => + [ + (rtac (Exh_dnat RS disjE) 1), + (eresolve_tac prems 1), + (etac disjE 1), + (eresolve_tac prems 1), + (REPEAT (etac exE 1)), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); + +(* ------------------------------------------------------------------------*) +(* Properties of dnat_when *) +(* ------------------------------------------------------------------------*) + +fun prover defs thm = prove_goalw Dnat.thy defs thm + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) + ]); + + +val dnat_when = [ + prover [dnat_when_def] "dnat_when`c`f`UU=UU", + prover [dnat_when_def,dzero_def] "dnat_when`c`f`dzero=c", + prover [dnat_when_def,dsucc_def] + "n~=UU ==> dnat_when`c`f`(dsucc`n)=f`n" + ]; + +val dnat_rews = dnat_when @ dnat_rews; + +(* ------------------------------------------------------------------------*) +(* Rewrites for discriminators and selectors *) +(* ------------------------------------------------------------------------*) + +fun prover defs thm = prove_goalw Dnat.thy defs thm + (fn prems => + [ + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_discsel = [ + prover [is_dzero_def] "is_dzero`UU=UU", + prover [is_dsucc_def] "is_dsucc`UU=UU", + prover [dpred_def] "dpred`UU=UU" + ]; + + +fun prover defs thm = prove_goalw Dnat.thy defs thm + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_discsel = [ + prover [is_dzero_def] "is_dzero`dzero=TT", + prover [is_dzero_def] "n~=UU ==>is_dzero`(dsucc`n)=FF", + prover [is_dsucc_def] "is_dsucc`dzero=FF", + prover [is_dsucc_def] "n~=UU ==> is_dsucc`(dsucc`n)=TT", + prover [dpred_def] "dpred`dzero=UU", + prover [dpred_def] "n~=UU ==> dpred`(dsucc`n)=n" + ] @ dnat_discsel; + +val dnat_rews = dnat_discsel @ dnat_rews; + +(* ------------------------------------------------------------------------*) +(* Definedness and strictness *) +(* ------------------------------------------------------------------------*) + +fun prover contr thm = prove_goal Dnat.thy thm + (fn prems => + [ + (res_inst_tac [("P1",contr)] classical3 1), + (simp_tac (!simpset addsimps dnat_rews) 1), + (dtac sym 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps (prems @ dnat_rews)) 1) + ]); + +val dnat_constrdef = [ + prover "is_dzero`UU ~= UU" "dzero~=UU", + prover "is_dsucc`UU ~= UU" "n~=UU ==> dsucc`n~=UU" + ]; + + +fun prover defs thm = prove_goalw Dnat.thy defs thm + (fn prems => + [ + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_constrdef = [ + prover [dsucc_def] "dsucc`UU=UU" + ] @ dnat_constrdef; + +val dnat_rews = dnat_constrdef @ dnat_rews; + + +(* ------------------------------------------------------------------------*) +(* Distinctness wrt. << and = *) +(* ------------------------------------------------------------------------*) + +val temp = prove_goal Dnat.thy "~dzero << dsucc`n" + (fn prems => + [ + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dzero")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("Q","n=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_dist_less = [temp]; + +val temp = prove_goal Dnat.thy "n~=UU ==> ~dsucc`n << dzero" + (fn prems => + [ + (cut_facts_tac prems 1), + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dsucc")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_dist_less = temp::dnat_dist_less; + +val temp = prove_goal Dnat.thy "dzero ~= dsucc`n" + (fn prems => + [ + (res_inst_tac [("Q","n=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("P1","TT = FF")] classical3 1), + (resolve_tac dist_eq_tr 1), + (dres_inst_tac [("f","is_dzero")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_dist_eq = [temp, temp RS not_sym]; + +val dnat_rews = dnat_dist_less @ dnat_dist_eq @ dnat_rews; + +(* ------------------------------------------------------------------------*) +(* Invertibility *) +(* ------------------------------------------------------------------------*) + +val dnat_invert = + [ +prove_goal Dnat.thy +"[|x1~=UU; y1~=UU; dsucc`x1 << dsucc`y1 |] ==> x1<< y1" + (fn prems => + [ + (cut_facts_tac prems 1), + (dres_inst_tac [("fo5","dnat_when`c`(LAM x.x)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]) + ]; + +(* ------------------------------------------------------------------------*) +(* Injectivity *) +(* ------------------------------------------------------------------------*) + +val dnat_inject = + [ +prove_goal Dnat.thy +"[|x1~=UU; y1~=UU; dsucc`x1 = dsucc`y1 |] ==> x1= y1" + (fn prems => + [ + (cut_facts_tac prems 1), + (dres_inst_tac [("f","dnat_when`c`(LAM x.x)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]) + ]; + +(* ------------------------------------------------------------------------*) +(* definedness for discriminators and selectors *) +(* ------------------------------------------------------------------------*) + + +fun prover thm = prove_goal Dnat.thy thm + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac dnatE 1), + (contr_tac 1), + (REPEAT (asm_simp_tac (!simpset addsimps dnat_rews) 1)) + ]); + +val dnat_discsel_def = + [ + prover "n~=UU ==> is_dzero`n ~= UU", + prover "n~=UU ==> is_dsucc`n ~= UU" + ]; + +val dnat_rews = dnat_discsel_def @ dnat_rews; + + +(* ------------------------------------------------------------------------*) +(* Properties dnat_take *) +(* ------------------------------------------------------------------------*) +val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take n`UU = UU" + (fn prems => + [ + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_take = [temp]; + +val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take 0`xs = UU" + (fn prems => + [ + (Asm_simp_tac 1) + ]); + +val dnat_take = temp::dnat_take; + +val temp = prove_goalw Dnat.thy [dnat_take_def] + "dnat_take (Suc n)`dzero=dzero" + (fn prems => + [ + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_take = temp::dnat_take; + +val temp = prove_goalw Dnat.thy [dnat_take_def] + "dnat_take (Suc n)`(dsucc`xs)=dsucc`(dnat_take n ` xs)" + (fn prems => + [ + (res_inst_tac [("Q","xs=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +val dnat_take = temp::dnat_take; + +val dnat_rews = dnat_take @ dnat_rews; + + +(* ------------------------------------------------------------------------*) +(* take lemma for dnats *) +(* ------------------------------------------------------------------------*) + +fun prover reach defs thm = prove_goalw Dnat.thy defs thm + (fn prems => + [ + (res_inst_tac [("t","s1")] (reach RS subst) 1), + (res_inst_tac [("t","s2")] (reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac lub_equal 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac allI 1), + (resolve_tac prems 1) + ]); + +val dnat_take_lemma = prover dnat_reach [dnat_take_def] + "(!!n.dnat_take n`s1 = dnat_take n`s2) ==> s1=s2"; + + +(* ------------------------------------------------------------------------*) +(* Co -induction for dnats *) +(* ------------------------------------------------------------------------*) + +qed_goalw "dnat_coind_lemma" Dnat.thy [dnat_bisim_def] +"dnat_bisim R ==> ! p q. R p q --> dnat_take n`p = dnat_take n`q" + (fn prems => + [ + (cut_facts_tac prems 1), + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dnat_take) 1), + (strip_tac 1), + ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), + (atac 1), + (asm_simp_tac (!simpset addsimps dnat_take) 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dnat_take) 1), + (etac exE 1), + (etac exE 1), + (asm_simp_tac (!simpset addsimps dnat_take) 1), + (REPEAT (etac conjE 1)), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); + +qed_goal "dnat_coind" Dnat.thy "[|dnat_bisim R;R p q|] ==> p = q" + (fn prems => + [ + (rtac dnat_take_lemma 1), + (rtac (dnat_coind_lemma RS spec RS spec RS mp) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); + + +(* ------------------------------------------------------------------------*) +(* structural induction for admissible predicates *) +(* ------------------------------------------------------------------------*) + +(* not needed any longer +qed_goal "dnat_ind" Dnat.thy +"[| adm(P);\ +\ P(UU);\ +\ P(dzero);\ +\ !! s1.[|s1~=UU ; P(s1)|] ==> P(dsucc`s1)|] ==> P(s)" + (fn prems => + [ + (rtac (dnat_reach RS subst) 1), + (res_inst_tac [("x","s")] spec 1), + (rtac fix_ind 1), + (rtac adm_all2 1), + (rtac adm_subst 1), + (cont_tacR 1), + (resolve_tac prems 1), + (Simp_tac 1), + (resolve_tac prems 1), + (strip_tac 1), + (res_inst_tac [("n","xa")] dnatE 1), + (asm_simp_tac (!simpset addsimps dnat_copy) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_copy) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_copy) 1), + (res_inst_tac [("Q","x`xb=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (eresolve_tac prems 1), + (etac spec 1) + ]); +*) + +qed_goal "dnat_finite_ind" Dnat.thy +"[|P(UU);P(dzero);\ +\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ +\ |] ==> !s.P(dnat_take n`s)" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (rtac allI 1), + (res_inst_tac [("n","s")] dnatE 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("Q","dnat_take n1`x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (atac 1), + (etac spec 1) + ]); + +qed_goal "dnat_all_finite_lemma1" Dnat.thy +"!s.dnat_take n`s=UU |dnat_take n`s=s" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dnat_rews) 1), + (rtac allI 1), + (res_inst_tac [("n","s")] dnatE 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (eres_inst_tac [("x","x")] allE 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); + +qed_goal "dnat_all_finite_lemma2" Dnat.thy "? n.dnat_take n`s=s" + (fn prems => + [ + (res_inst_tac [("Q","s=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (subgoal_tac "(!n.dnat_take(n)`s=UU) |(? n.dnat_take(n)`s=s)" 1), + (etac disjE 1), + (eres_inst_tac [("P","s=UU")] notE 1), + (rtac dnat_take_lemma 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (atac 1), + (subgoal_tac "!n.!s.dnat_take(n)`s=UU |dnat_take(n)`s=s" 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (rtac dnat_all_finite_lemma1 1) + ]); + + +qed_goal "dnat_ind" Dnat.thy +"[|P(UU);P(dzero);\ +\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ +\ |] ==> P(s)" + (fn prems => + [ + (rtac (dnat_all_finite_lemma2 RS exE) 1), + (etac subst 1), + (rtac (dnat_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); + + +qed_goalw "dnat_flat" Dnat.thy [flat_def] "flat(dzero)" + (fn prems => + [ + (rtac allI 1), + (res_inst_tac [("s","x")] dnat_ind 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (res_inst_tac [("n","y")] dnatE 1), + (fast_tac (HOL_cs addSIs [UU_I]) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_dist_less) 1), + (rtac allI 1), + (res_inst_tac [("n","y")] dnatE 1), + (fast_tac (HOL_cs addSIs [UU_I]) 1), + (asm_simp_tac (!simpset addsimps dnat_dist_less) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (strip_tac 1), + (subgoal_tac "s1< + +For details see chapter 5 of: + +[Franz Regensburger] HOLCF: Eine konservative Erweiterung von HOL um LCF, + Dissertation, Technische Universit"at M"unchen, 1994 + +*) + +Dnat = HOLCF + + +types dnat 0 + +(* ----------------------------------------------------------------------- *) +(* arrity axiom is valuated by semantical reasoning *) + +arities dnat::pcpo + +consts + +(* ----------------------------------------------------------------------- *) +(* essential constants *) + +dnat_rep :: " dnat -> (one ++ dnat)" +dnat_abs :: "(one ++ dnat) -> dnat" + +(* ----------------------------------------------------------------------- *) +(* abstract constants and auxiliary constants *) + +dnat_copy :: "(dnat -> dnat) -> dnat -> dnat" + +dzero :: "dnat" +dsucc :: "dnat -> dnat" +dnat_when :: "'b -> (dnat -> 'b) -> dnat -> 'b" +is_dzero :: "dnat -> tr" +is_dsucc :: "dnat -> tr" +dpred :: "dnat -> dnat" +dnat_take :: "nat => dnat -> dnat" +dnat_bisim :: "(dnat => dnat => bool) => bool" + +rules + +(* ----------------------------------------------------------------------- *) +(* axiomatization of recursive type dnat *) +(* ----------------------------------------------------------------------- *) +(* (dnat,dnat_abs) is the initial F-algebra where *) +(* F is the locally continuous functor determined by functor term FT. *) +(* domain equation: dnat = one ++ dnat *) +(* functor term: FT = <++,K_{one},I> *) +(* ----------------------------------------------------------------------- *) +(* dnat_abs is an isomorphism with inverse dnat_rep *) +(* identity is the least endomorphism on dnat *) + +dnat_abs_iso "dnat_rep`(dnat_abs`x) = x" +dnat_rep_iso "dnat_abs`(dnat_rep`x) = x" +dnat_copy_def "dnat_copy == (LAM f. dnat_abs oo + (sswhen`sinl`(sinr oo f)) oo dnat_rep )" +dnat_reach "(fix`dnat_copy)`x=x" + + +defs +(* ----------------------------------------------------------------------- *) +(* properties of additional constants *) +(* ----------------------------------------------------------------------- *) +(* constructors *) + +dzero_def "dzero == dnat_abs`(sinl`one)" +dsucc_def "dsucc == (LAM n. dnat_abs`(sinr`n))" + +(* ----------------------------------------------------------------------- *) +(* discriminator functional *) + +dnat_when_def "dnat_when == (LAM f1 f2 n.sswhen`(LAM x.f1)`f2`(dnat_rep`n))" + + +(* ----------------------------------------------------------------------- *) +(* discriminators and selectors *) + +is_dzero_def "is_dzero == dnat_when`TT`(LAM x.FF)" +is_dsucc_def "is_dsucc == dnat_when`FF`(LAM x.TT)" +dpred_def "dpred == dnat_when`UU`(LAM x.x)" + + +(* ----------------------------------------------------------------------- *) +(* the taker for dnats *) + +dnat_take_def "dnat_take == (%n.iterate n dnat_copy UU)" + +(* ----------------------------------------------------------------------- *) +(* definition of bisimulation is determined by domain equation *) +(* simplification and rewriting for abstract constants yields def below *) + +dnat_bisim_def "dnat_bisim == +(%R.!s1 s2. + R s1 s2 --> + ((s1=UU & s2=UU) |(s1=dzero & s2=dzero) | + (? s11 s21. s11~=UU & s21~=UU & s1=dsucc`s11 & + s2 = dsucc`s21 & R s11 s21)))" + +end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Dnat2.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Dnat2.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,52 @@ +(* Title: HOLCF/Dnat2.ML + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Lemmas for theory Dnat2.thy +*) + +open Dnat2; + + +(* ------------------------------------------------------------------------- *) +(* expand fixed point properties *) +(* ------------------------------------------------------------------------- *) + +val iterator_def2 = fix_prover2 Dnat2.thy iterator_def + "iterator = (LAM n f x. dnat_when`x`(LAM m.f`(iterator`m`f`x)) `n)"; + +(* ------------------------------------------------------------------------- *) +(* recursive properties *) +(* ------------------------------------------------------------------------- *) + +qed_goal "iterator1" Dnat2.thy "iterator`UU`f`x = UU" + (fn prems => + [ + (rtac (iterator_def2 RS ssubst) 1), + (simp_tac (!simpset addsimps dnat_when) 1) + ]); + +qed_goal "iterator2" Dnat2.thy "iterator`dzero`f`x = x" + (fn prems => + [ + (rtac (iterator_def2 RS ssubst) 1), + (simp_tac (!simpset addsimps dnat_when) 1) + ]); + +qed_goal "iterator3" Dnat2.thy +"n~=UU ==> iterator`(dsucc`n)`f`x = f`(iterator`n`f`x)" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac trans 1), + (rtac (iterator_def2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (rtac refl 1) + ]); + +val dnat2_rews = + [iterator1, iterator2, iterator3]; + + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Dnat2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Dnat2.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,29 @@ +(* Title: HOLCF/Dnat2.thy + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Additional constants for dnat + +*) + +Dnat2 = Dnat + + +consts + +iterator :: "dnat -> ('a -> 'a) -> 'a -> 'a" + + +defs + +iterator_def "iterator == fix`(LAM h n f x. + dnat_when `x `(LAM m.f`(h`m`f`x)) `n)" +end + +(* + + iterator`UU`f`x = UU + iterator`dzero`f`x = x + n~=UU --> iterator`(dsucc`n)`f`x = f`(iterator`n`f`x) +*) + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Focus_ex.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Focus_ex.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,151 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1995 Technische Universitaet Muenchen + +*) + +open Focus_ex; + +(* first some logical trading *) + +val prems = goal Focus_ex.thy +"is_g(g) = \ +\ (? f. is_f(f) & (!x.(? z. = f` & \ +\ (! w y. = f` --> z << w))))"; +by (simp_tac (!simpset addsimps [is_g,is_net_g]) 1); +by (fast_tac HOL_cs 1); +val lemma1 = result(); + +val prems = goal Focus_ex.thy +"(? f. is_f(f) & (!x. (? z. = f` & \ +\ (! w y. = f` --> z << w)))) \ +\ = \ +\ (? f. is_f(f) & (!x. ? z.\ +\ g`x = cfst`(f`) & \ +\ z = csnd`(f`) & \ +\ (! w y. = f` --> z << w)))"; +by (rtac iffI 1); +by (etac exE 1); +by (res_inst_tac [("x","f")] exI 1); +by (REPEAT (etac conjE 1)); +by (etac conjI 1); +by (strip_tac 1); +by (etac allE 1); +by (etac exE 1); +by (res_inst_tac [("x","z")] exI 1); +by (REPEAT (etac conjE 1)); +by (rtac conjI 1); +by (rtac conjI 2); +by (atac 3); +by (dtac sym 1); +by (Asm_simp_tac 1); +by (dtac sym 1); +by (Asm_simp_tac 1); +by (etac exE 1); +by (res_inst_tac [("x","f")] exI 1); +by (REPEAT (etac conjE 1)); +by (etac conjI 1); +by (strip_tac 1); +by (etac allE 1); +by (etac exE 1); +by (res_inst_tac [("x","z")] exI 1); +by (REPEAT (etac conjE 1)); +by (rtac conjI 1); +by (atac 2); +by (rtac trans 1); +by (rtac (surjective_pairing_Cprod2) 2); +by (etac subst 1); +by (etac subst 1); +by (rtac refl 1); +val lemma2 = result(); + +(* direction def_g(g) --> is_g(g) *) + +val prems = goal Focus_ex.thy "def_g(g) --> is_g(g)"; +by (simp_tac (!simpset addsimps [def_g,lemma1, lemma2]) 1); +by (rtac impI 1); +by (etac exE 1); +by (res_inst_tac [("x","f")] exI 1); +by (REPEAT (etac conjE 1)); +by (etac conjI 1); +by (strip_tac 1); +by (res_inst_tac [("x","fix`(LAM k.csnd`(f`))")] exI 1); +by (rtac conjI 1); +by (Asm_simp_tac 1); +by (rtac conjI 1); +by (rtac trans 1); +by (rtac fix_eq 1); +by (Simp_tac 1); +by (strip_tac 1); +by (rtac fix_least 1); +by (dtac sym 1); +back(); +by (Asm_simp_tac 1); +val lemma3 = result(); + +(* direction is_g(g) --> def_g(g) *) +val prems = goal Focus_ex.thy "is_g(g) --> def_g(g)"; +by (simp_tac (!simpset addsimps [lemma1,lemma2,def_g]) 1); +by (rtac impI 1); +by (etac exE 1); +by (res_inst_tac [("x","f")] exI 1); +by (REPEAT (etac conjE 1)); +by (etac conjI 1); +by (rtac ext_cfun 1); +by (etac allE 1); +by (etac exE 1); +by (REPEAT (etac conjE 1)); +by (subgoal_tac "fix`(LAM k. csnd`(f`)) = z" 1); +by (Asm_simp_tac 1); +by (subgoal_tac "! w y. f` = --> z << w" 1); +by (rtac sym 1); +by (rtac fix_eqI 1); +by (Asm_simp_tac 1); +by (etac sym 1); +by (rtac allI 1); +by (Simp_tac 1); +by (strip_tac 1); +by (subgoal_tac "f` = ),za>" 1); +by (fast_tac HOL_cs 1); +by (rtac trans 1); +by (rtac (surjective_pairing_Cprod2 RS sym) 1); +by (etac cfun_arg_cong 1); +by (strip_tac 1); +by (REPEAT (etac allE 1)); +by (etac mp 1); +by (etac sym 1); +val lemma4 = result(); + +(* now we assemble the result *) + +val prems = goal Focus_ex.thy "def_g = is_g"; +by (rtac ext 1); +by (rtac iffI 1); +by (etac (lemma3 RS mp) 1); +by (etac (lemma4 RS mp) 1); +val loopback_eq = result(); + +val prems = goal Focus_ex.thy +"(? f.\ +\ is_f(f::'b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream))\ +\ -->\ +\ (? g. def_g(g::'b stream -> 'c stream ))"; +by (simp_tac (!simpset addsimps [def_g]) 1); +by (strip_tac 1); +by (etac exE 1); +by (rtac exI 1); +by (rtac exI 1); +by (etac conjI 1); +by (rtac refl 1); +val L2 = result(); + +val prems = goal Focus_ex.thy +"(? f.\ +\ is_f(f::'b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream))\ +\ -->\ +\ (? g. is_g(g::'b stream -> 'c stream ))"; +by (rtac (loopback_eq RS subst) 1); +by (rtac L2 1); +val conservative_loopback = result(); + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Focus_ex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Focus_ex.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,143 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1995 Technische Universitaet Muenchen + +*) + +(* Specification of the following loop back device + + + g + -------------------- + | ------- | + x | | | | y + ------|---->| |------| -----> + | z | f | z | + | -->| |--- | + | | | | | | + | | ------- | | + | | | | + | <-------------- | + | | + -------------------- + + +First step: Notation in Agent Network Description Language (ANDL) +----------------------------------------------------------------- + +agent f + input channel i1:'b i2: ('b,'c) tc + output channel o1:'c o2: ('b,'c) tc +is + Rf(i1,i2,o1,o2) (left open in the example) +end f + +agent g + input channel x:'b + output channel y:'c +is network + = f` +end network +end g + + +Remark: the type of the feedback depends at most on the types of the input and + output of g. (No type miracles inside g) + +Second step: Translation of ANDL specification to HOLCF Specification +--------------------------------------------------------------------- + +Specification of agent f ist translated to predicate is_f + +is_f :: ('b stream * ('b,'c) tc stream -> + 'c stream * ('b,'c) tc stream) => bool + +is_f f = ! i1 i2 o1 o2. + f` = --> Rf(i1,i2,o1,o2) + +Specification of agent g is translated to predicate is_g which uses +predicate is_net_g + +is_net_g :: ('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => + 'b stream => 'c stream => bool + +is_net_g f x y = + ? z. = f` & + ! oy hz. = f` --> z << hz + + +is_g :: ('b stream -> 'c stream) => bool + +is_g g = ? f. is_f f & (! x y. g`x = y --> is_net_g f x y + +Third step: (show conservativity) +----------- + +Suppose we have a model for the theory TH1 which contains the axiom + + ? f. is_f f + +In this case there is also a model for the theory TH2 that enriches TH1 by +axiom + + ? g. is_g g + +The result is proved by showing that there is a definitional extension +that extends TH1 by a definition of g. + + +We define: + +def_g g = + (? f. is_f f & + g = (LAM x. cfst`(f`))>)) ) + +Now we prove: + + (?f. is_f f ) --> (? g. is_g g) + +using the theorems + +loopback_eq) def_g = is_g (real work) + +L1) (? f. is_f f ) --> (? g. def_g g) (trivial) + +*) + +Focus_ex = Stream + + +types tc 2 + +arities tc:: (pcpo,pcpo)pcpo + +consts + +is_f :: + "('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => bool" +is_net_g :: "('b stream *('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => + 'b stream => 'c stream => bool" +is_g :: "('b stream -> 'c stream) => bool" +def_g :: "('b stream -> 'c stream) => bool" +Rf :: +"('b stream * ('b,'c) tc stream * 'c stream * ('b,'c) tc stream) => bool" + +defs + +is_f "is_f f == (! i1 i2 o1 o2. + f` = --> Rf(i1,i2,o1,o2))" + +is_net_g "is_net_g f x y == (? z. + = f` & + (! oy hz. = f` --> z << hz))" + +is_g "is_g g == (? f. + is_f f & + (!x y. g`x = y --> is_net_g f x y))" + + +def_g "def_g g == (? f. + is_f f & + g = (LAM x. cfst`(f`))>)))" + +end diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/README Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,22 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1995 Technische Universitaet Muenchen + +*) + +The files contained in this directory are examples for the +explicit construction of domains. The technique used is described +in the thesis + + HOLCF: Eine konservative Erweiterung von HOL um LCF + +The thesis is available via the web using URL + + http://www4.informatik.tu-muenchen.de/~regensbu/papers.html + + +The same construction is automatically performed if you use the +type definition package of David Oheimb. See subdirectory HOLCF/domains +for more details. + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/ROOT.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,22 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1995 Technische Universitaet Muenchen + +*) + +HOLCF_build_completed; (*Cause examples to fail if HOLCF did*) + +writeln"Root file for HOLCF examples: explicit domain axiomatisation"; +proof_timing := true; +time_use_thy "explicit_domains/Dnat"; +time_use_thy "explicit_domains/Dnat2"; +time_use_thy "explicit_domains/Stream"; +time_use_thy "explicit_domains/Stream2"; +time_use_thy "explicit_domains/Dlist"; + +time_use_thy "explicit_domains/Coind"; +time_use_thy "explicit_domains/Dagstuhl"; +time_use_thy "explicit_domains/Focus_ex"; + +maketest "END: Root file for HOLCF examples: explicit domain axiomatization"; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Stream.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Stream.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,840 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Lemmas for stream.thy +*) + +open Stream; + +(* ------------------------------------------------------------------------*) +(* The isomorphisms stream_rep_iso and stream_abs_iso are strict *) +(* ------------------------------------------------------------------------*) + +val stream_iso_strict= stream_rep_iso RS (stream_abs_iso RS + (allI RSN (2,allI RS iso_strict))); + +val stream_rews = [stream_iso_strict RS conjunct1, + stream_iso_strict RS conjunct2]; + +(* ------------------------------------------------------------------------*) +(* Properties of stream_copy *) +(* ------------------------------------------------------------------------*) + +fun prover defs thm = prove_goalw Stream.thy defs thm + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) + ]); + +val stream_copy = + [ + prover [stream_copy_def] "stream_copy`f`UU=UU", + prover [stream_copy_def,scons_def] + "x~=UU ==> stream_copy`f`(scons`x`xs)= scons`x`(f`xs)" + ]; + +val stream_rews = stream_copy @ stream_rews; + +(* ------------------------------------------------------------------------*) +(* Exhaustion and elimination for streams *) +(* ------------------------------------------------------------------------*) + +qed_goalw "Exh_stream" Stream.thy [scons_def] + "s = UU | (? x xs. x~=UU & s = scons`x`xs)" + (fn prems => + [ + (Simp_tac 1), + (rtac (stream_rep_iso RS subst) 1), + (res_inst_tac [("p","stream_rep`s")] sprodE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (Asm_simp_tac 1), + (res_inst_tac [("p","y")] liftE1 1), + (contr_tac 1), + (rtac disjI2 1), + (rtac exI 1), + (rtac exI 1), + (etac conjI 1), + (Asm_simp_tac 1) + ]); + +qed_goal "streamE" Stream.thy + "[| s=UU ==> Q; !!x xs.[|s=scons`x`xs;x~=UU|]==>Q|]==>Q" + (fn prems => + [ + (rtac (Exh_stream RS disjE) 1), + (eresolve_tac prems 1), + (etac exE 1), + (etac exE 1), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); + +(* ------------------------------------------------------------------------*) +(* Properties of stream_when *) +(* ------------------------------------------------------------------------*) + +fun prover defs thm = prove_goalw Stream.thy defs thm + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) + ]); + + +val stream_when = [ + prover [stream_when_def] "stream_when`f`UU=UU", + prover [stream_when_def,scons_def] + "x~=UU ==> stream_when`f`(scons`x`xs)= f`x`xs" + ]; + +val stream_rews = stream_when @ stream_rews; + +(* ------------------------------------------------------------------------*) +(* Rewrites for discriminators and selectors *) +(* ------------------------------------------------------------------------*) + +fun prover defs thm = prove_goalw Stream.thy defs thm + (fn prems => + [ + (simp_tac (!simpset addsimps stream_rews) 1) + ]); + +val stream_discsel = [ + prover [is_scons_def] "is_scons`UU=UU", + prover [shd_def] "shd`UU=UU", + prover [stl_def] "stl`UU=UU" + ]; + +fun prover defs thm = prove_goalw Stream.thy defs thm + (fn prems => + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + +val stream_discsel = [ +prover [is_scons_def,shd_def,stl_def] "x~=UU ==> is_scons`(scons`x`xs)=TT", +prover [is_scons_def,shd_def,stl_def] "x~=UU ==> shd`(scons`x`xs)=x", +prover [is_scons_def,shd_def,stl_def] "x~=UU ==> stl`(scons`x`xs)=xs" + ] @ stream_discsel; + +val stream_rews = stream_discsel @ stream_rews; + +(* ------------------------------------------------------------------------*) +(* Definedness and strictness *) +(* ------------------------------------------------------------------------*) + +fun prover contr thm = prove_goal Stream.thy thm + (fn prems => + [ + (res_inst_tac [("P1",contr)] classical3 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (dtac sym 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps (prems @ stream_rews)) 1) + ]); + +val stream_constrdef = [ + prover "is_scons`(UU::'a stream)~=UU" "x~=UU ==> scons`(x::'a)`xs~=UU" + ]; + +fun prover defs thm = prove_goalw Stream.thy defs thm + (fn prems => + [ + (simp_tac (!simpset addsimps stream_rews) 1) + ]); + +val stream_constrdef = [ + prover [scons_def] "scons`UU`xs=UU" + ] @ stream_constrdef; + +val stream_rews = stream_constrdef @ stream_rews; + + +(* ------------------------------------------------------------------------*) +(* Distinctness wrt. << and = *) +(* ------------------------------------------------------------------------*) + + +(* ------------------------------------------------------------------------*) +(* Invertibility *) +(* ------------------------------------------------------------------------*) + +val stream_invert = + [ +prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ +\ scons`x1`x2 << scons`y1`y2|] ==> x1<< y1 & x2 << y2" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("fo5","stream_when`(LAM x l.x)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (dres_inst_tac [("fo5","stream_when`(LAM x l.l)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1) + ]) + ]; + +(* ------------------------------------------------------------------------*) +(* Injectivity *) +(* ------------------------------------------------------------------------*) + +val stream_inject = + [ +prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ +\ scons`x1`x2 = scons`y1`y2 |] ==> x1= y1 & x2 = y2" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("f","stream_when`(LAM x l.x)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (dres_inst_tac [("f","stream_when`(LAM x l.l)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1) + ]) + ]; + +(* ------------------------------------------------------------------------*) +(* definedness for discriminators and selectors *) +(* ------------------------------------------------------------------------*) + +fun prover thm = prove_goal Stream.thy thm + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac streamE 1), + (contr_tac 1), + (REPEAT (asm_simp_tac (!simpset addsimps stream_discsel) 1)) + ]); + +val stream_discsel_def = + [ + prover "s~=UU ==> is_scons`s ~= UU", + prover "s~=UU ==> shd`s ~=UU" + ]; + +val stream_rews = stream_discsel_def @ stream_rews; + + +(* ------------------------------------------------------------------------*) +(* Properties stream_take *) +(* ------------------------------------------------------------------------*) + +val stream_take = + [ +prove_goalw Stream.thy [stream_take_def] "stream_take n`UU = UU" + (fn prems => + [ + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1) + ]), +prove_goalw Stream.thy [stream_take_def] "stream_take 0`xs=UU" + (fn prems => + [ + (Asm_simp_tac 1) + ])]; + +fun prover thm = prove_goalw Stream.thy [stream_take_def] thm + (fn prems => + [ + (cut_facts_tac prems 1), + (Simp_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + +val stream_take = [ +prover + "x~=UU ==> stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" + ] @ stream_take; + +val stream_rews = stream_take @ stream_rews; + +(* ------------------------------------------------------------------------*) +(* enhance the simplifier *) +(* ------------------------------------------------------------------------*) + +qed_goal "stream_copy2" Stream.thy + "stream_copy`f`(scons`x`xs) = scons`x`(f`xs)" + (fn prems => + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + +qed_goal "shd2" Stream.thy "shd`(scons`x`xs) = x" + (fn prems => + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + +qed_goal "stream_take2" Stream.thy + "stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" + (fn prems => + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + +val stream_rews = [stream_iso_strict RS conjunct1, + stream_iso_strict RS conjunct2, + hd stream_copy, stream_copy2] + @ stream_when + @ [hd stream_discsel,shd2] @ (tl (tl stream_discsel)) + @ stream_constrdef + @ stream_discsel_def + @ [ stream_take2] @ (tl stream_take); + + +(* ------------------------------------------------------------------------*) +(* take lemma for streams *) +(* ------------------------------------------------------------------------*) + +fun prover reach defs thm = prove_goalw Stream.thy defs thm + (fn prems => + [ + (res_inst_tac [("t","s1")] (reach RS subst) 1), + (res_inst_tac [("t","s2")] (reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac lub_equal 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac allI 1), + (resolve_tac prems 1) + ]); + +val stream_take_lemma = prover stream_reach [stream_take_def] + "(!!n.stream_take n`s1 = stream_take n`s2) ==> s1=s2"; + + +qed_goal "stream_reach2" Stream.thy "lub(range(%i.stream_take i`s))=s" + (fn prems => + [ + (res_inst_tac [("t","s")] (stream_reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rewrite_goals_tac [stream_take_def]), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac refl 1) + ]); + +(* ------------------------------------------------------------------------*) +(* Co -induction for streams *) +(* ------------------------------------------------------------------------*) + +qed_goalw "stream_coind_lemma" Stream.thy [stream_bisim_def] +"stream_bisim R ==> ! p q. R p q --> stream_take n`p = stream_take n`q" + (fn prems => + [ + (cut_facts_tac prems 1), + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1), + ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), + (atac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (etac exE 1), + (etac exE 1), + (etac exE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (REPEAT (etac conjE 1)), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); + +qed_goal "stream_coind" Stream.thy "[|stream_bisim R ;R p q|] ==> p = q" + (fn prems => + [ + (rtac stream_take_lemma 1), + (rtac (stream_coind_lemma RS spec RS spec RS mp) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); + +(* ------------------------------------------------------------------------*) +(* structural induction for admissible predicates *) +(* ------------------------------------------------------------------------*) + +qed_goal "stream_finite_ind" Stream.thy +"[|P(UU);\ +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ +\ |] ==> !s.P(stream_take n`s)" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (resolve_tac prems 1), + (rtac allI 1), + (res_inst_tac [("s","s")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (resolve_tac prems 1), + (atac 1), + (etac spec 1) + ]); + +qed_goalw "stream_finite_ind2" Stream.thy [stream_finite_def] +"(!!n.P(stream_take n`s)) ==> stream_finite(s) -->P(s)" + (fn prems => + [ + (strip_tac 1), + (etac exE 1), + (etac subst 1), + (resolve_tac prems 1) + ]); + +qed_goal "stream_finite_ind3" Stream.thy +"[|P(UU);\ +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ +\ |] ==> stream_finite(s) --> P(s)" + (fn prems => + [ + (rtac stream_finite_ind2 1), + (rtac (stream_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); + +(* prove induction using definition of admissibility + stream_reach rsp. stream_reach2 + and finite induction stream_finite_ind *) + +qed_goal "stream_ind" Stream.thy +"[|adm(P);\ +\ P(UU);\ +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ +\ |] ==> P(s)" + (fn prems => + [ + (rtac (stream_reach2 RS subst) 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (resolve_tac prems 1), + (SELECT_GOAL (rewrite_goals_tac [stream_take_def]) 1), + (rtac ch2ch_fappL 1), + (rtac is_chain_iterate 1), + (rtac allI 1), + (rtac (stream_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); + +(* prove induction with usual LCF-Method using fixed point induction *) +qed_goal "stream_ind" Stream.thy +"[|adm(P);\ +\ P(UU);\ +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ +\ |] ==> P(s)" + (fn prems => + [ + (rtac (stream_reach RS subst) 1), + (res_inst_tac [("x","s")] spec 1), + (rtac wfix_ind 1), + (rtac adm_impl_admw 1), + (REPEAT (resolve_tac adm_thms 1)), + (rtac adm_subst 1), + (cont_tacR 1), + (resolve_tac prems 1), + (rtac allI 1), + (rtac (rewrite_rule [stream_take_def] stream_finite_ind) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); + + +(* ------------------------------------------------------------------------*) +(* simplify use of Co-induction *) +(* ------------------------------------------------------------------------*) + +qed_goal "surjectiv_scons" Stream.thy "scons`(shd`s)`(stl`s)=s" + (fn prems => + [ + (res_inst_tac [("s","s")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + + +qed_goalw "stream_coind_lemma2" Stream.thy [stream_bisim_def] +"!s1 s2. R s1 s2 --> shd`s1 = shd`s2 & R (stl`s1) (stl`s2) ==> stream_bisim R" + (fn prems => + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (etac allE 1), + (etac allE 1), + (dtac mp 1), + (atac 1), + (etac conjE 1), + (res_inst_tac [("Q","s1 = UU & s2 = UU")] classical2 1), + (rtac disjI1 1), + (fast_tac HOL_cs 1), + (rtac disjI2 1), + (rtac disjE 1), + (etac (de_morgan2 RS ssubst) 1), + (res_inst_tac [("x","shd`s1")] exI 1), + (res_inst_tac [("x","stl`s1")] exI 1), + (res_inst_tac [("x","stl`s2")] exI 1), + (rtac conjI 1), + (eresolve_tac stream_discsel_def 1), + (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), + (eres_inst_tac [("s","shd`s1"),("t","shd`s2")] subst 1), + (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), + (res_inst_tac [("x","shd`s2")] exI 1), + (res_inst_tac [("x","stl`s1")] exI 1), + (res_inst_tac [("x","stl`s2")] exI 1), + (rtac conjI 1), + (eresolve_tac stream_discsel_def 1), + (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), + (res_inst_tac [("s","shd`s1"),("t","shd`s2")] ssubst 1), + (etac sym 1), + (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1) + ]); + + +(* ------------------------------------------------------------------------*) +(* theorems about finite and infinite streams *) +(* ------------------------------------------------------------------------*) + +(* ----------------------------------------------------------------------- *) +(* 2 lemmas about stream_finite *) +(* ----------------------------------------------------------------------- *) + +qed_goalw "stream_finite_UU" Stream.thy [stream_finite_def] + "stream_finite(UU)" + (fn prems => + [ + (rtac exI 1), + (simp_tac (!simpset addsimps stream_rews) 1) + ]); + +qed_goal "inf_stream_not_UU" Stream.thy "~stream_finite(s) ==> s ~= UU" + (fn prems => + [ + (cut_facts_tac prems 1), + (etac swap 1), + (dtac notnotD 1), + (hyp_subst_tac 1), + (rtac stream_finite_UU 1) + ]); + +(* ----------------------------------------------------------------------- *) +(* a lemma about shd *) +(* ----------------------------------------------------------------------- *) + +qed_goal "stream_shd_lemma1" Stream.thy "shd`s=UU --> s=UU" + (fn prems => + [ + (res_inst_tac [("s","s")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (hyp_subst_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + + +(* ----------------------------------------------------------------------- *) +(* lemmas about stream_take *) +(* ----------------------------------------------------------------------- *) + +qed_goal "stream_take_lemma1" Stream.thy + "!x xs.x~=UU --> \ +\ stream_take (Suc n)`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" + (fn prems => + [ + (rtac allI 1), + (rtac allI 1), + (rtac impI 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1), + (rtac ((hd stream_inject) RS conjunct2) 1), + (atac 1), + (atac 1), + (atac 1) + ]); + + +qed_goal "stream_take_lemma2" Stream.thy + "! s2. stream_take n`s2 = s2 --> stream_take (Suc n)`s2=s2" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (hyp_subst_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (rtac allI 1), + (res_inst_tac [("s","s2")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (subgoal_tac "stream_take n1`xs = xs" 1), + (rtac ((hd stream_inject) RS conjunct2) 2), + (atac 4), + (atac 2), + (atac 2), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); + +qed_goal "stream_take_lemma3" Stream.thy + "!x xs.x~=UU --> \ +\ stream_take n`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" + (fn prems => + [ + (nat_ind_tac "n" 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (res_inst_tac [("P","scons`x`xs=UU")] notE 1), + (eresolve_tac stream_constrdef 1), + (etac sym 1), + (strip_tac 1 ), + (rtac (stream_take_lemma2 RS spec RS mp) 1), + (res_inst_tac [("x1.1","x")] ((hd stream_inject) RS conjunct2) 1), + (atac 1), + (atac 1), + (etac (stream_take2 RS subst) 1) + ]); + +qed_goal "stream_take_lemma4" Stream.thy + "!x xs.\ +\stream_take n`xs=xs --> stream_take (Suc n)`(scons`x`xs) = scons`x`xs" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (simp_tac (!simpset addsimps stream_rews) 1) + ]); + +(* ---- *) + +qed_goal "stream_take_lemma5" Stream.thy +"!s. stream_take n`s=s --> iterate n stl s=UU" + (fn prems => + [ + (nat_ind_tac "n" 1), + (Simp_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1), + (res_inst_tac [("s","s")] streamE 1), + (hyp_subst_tac 1), + (rtac (iterate_Suc2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac (iterate_Suc2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (etac allE 1), + (etac mp 1), + (hyp_subst_tac 1), + (etac (stream_take_lemma1 RS spec RS spec RS mp RS mp) 1), + (atac 1) + ]); + +qed_goal "stream_take_lemma6" Stream.thy +"!s.iterate n stl s =UU --> stream_take n`s=s" + (fn prems => + [ + (nat_ind_tac "n" 1), + (Simp_tac 1), + (strip_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac allI 1), + (res_inst_tac [("s","s")] streamE 1), + (hyp_subst_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (hyp_subst_tac 1), + (rtac (iterate_Suc2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); + +qed_goal "stream_take_lemma7" Stream.thy +"(iterate n stl s=UU) = (stream_take n`s=s)" + (fn prems => + [ + (rtac iffI 1), + (etac (stream_take_lemma6 RS spec RS mp) 1), + (etac (stream_take_lemma5 RS spec RS mp) 1) + ]); + + +qed_goal "stream_take_lemma8" Stream.thy +"[|adm(P); !n. ? m. n < m & P (stream_take m`s)|] ==> P(s)" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac (stream_reach2 RS subst) 1), + (rtac adm_disj_lemma11 1), + (atac 1), + (atac 2), + (rewrite_goals_tac [stream_take_def]), + (rtac ch2ch_fappL 1), + (rtac is_chain_iterate 1) + ]); + +(* ----------------------------------------------------------------------- *) +(* lemmas stream_finite *) +(* ----------------------------------------------------------------------- *) + +qed_goalw "stream_finite_lemma1" Stream.thy [stream_finite_def] + "stream_finite(xs) ==> stream_finite(scons`x`xs)" + (fn prems => + [ + (cut_facts_tac prems 1), + (etac exE 1), + (rtac exI 1), + (etac (stream_take_lemma4 RS spec RS spec RS mp) 1) + ]); + +qed_goalw "stream_finite_lemma2" Stream.thy [stream_finite_def] + "[|x~=UU; stream_finite(scons`x`xs)|] ==> stream_finite(xs)" + (fn prems => + [ + (cut_facts_tac prems 1), + (etac exE 1), + (rtac exI 1), + (etac (stream_take_lemma3 RS spec RS spec RS mp RS mp) 1), + (atac 1) + ]); + +qed_goal "stream_finite_lemma3" Stream.thy + "x~=UU ==> stream_finite(scons`x`xs) = stream_finite(xs)" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac iffI 1), + (etac stream_finite_lemma2 1), + (atac 1), + (etac stream_finite_lemma1 1) + ]); + + +qed_goalw "stream_finite_lemma5" Stream.thy [stream_finite_def] + "(!n. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1))\ +\=(s1 << s2 --> stream_finite(s2) --> stream_finite(s1))" + (fn prems => + [ + (rtac iffI 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); + +qed_goal "stream_finite_lemma6" Stream.thy + "!s1 s2. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1)" + (fn prems => + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (hyp_subst_tac 1), + (dtac UU_I 1), + (hyp_subst_tac 1), + (rtac stream_finite_UU 1), + (rtac allI 1), + (rtac allI 1), + (res_inst_tac [("s","s1")] streamE 1), + (hyp_subst_tac 1), + (strip_tac 1 ), + (rtac stream_finite_UU 1), + (hyp_subst_tac 1), + (res_inst_tac [("s","s2")] streamE 1), + (hyp_subst_tac 1), + (strip_tac 1 ), + (dtac UU_I 1), + (asm_simp_tac(!simpset addsimps (stream_rews @ [stream_finite_UU])) 1), + (hyp_subst_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (rtac stream_finite_lemma1 1), + (subgoal_tac "xs << xsa" 1), + (subgoal_tac "stream_take n1`xsa = xsa" 1), + (fast_tac HOL_cs 1), + (res_inst_tac [("x1.1","xa"),("y1.1","xa")] + ((hd stream_inject) RS conjunct2) 1), + (atac 1), + (atac 1), + (atac 1), + (res_inst_tac [("x1.1","x"),("y1.1","xa")] + ((hd stream_invert) RS conjunct2) 1), + (atac 1), + (atac 1), + (atac 1) + ]); + +qed_goal "stream_finite_lemma7" Stream.thy +"s1 << s2 --> stream_finite(s2) --> stream_finite(s1)" + (fn prems => + [ + (rtac (stream_finite_lemma5 RS iffD1) 1), + (rtac allI 1), + (rtac (stream_finite_lemma6 RS spec RS spec) 1) + ]); + +qed_goalw "stream_finite_lemma8" Stream.thy [stream_finite_def] +"stream_finite(s) = (? n. iterate n stl s = UU)" + (fn prems => + [ + (simp_tac (!simpset addsimps [stream_take_lemma7]) 1) + ]); + + +(* ----------------------------------------------------------------------- *) +(* admissibility of ~stream_finite *) +(* ----------------------------------------------------------------------- *) + +qed_goalw "adm_not_stream_finite" Stream.thy [adm_def] + "adm(%s. ~ stream_finite(s))" + (fn prems => + [ + (strip_tac 1 ), + (res_inst_tac [("P1","!i. ~ stream_finite(Y(i))")] classical3 1), + (atac 2), + (subgoal_tac "!i.stream_finite(Y(i))" 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (rtac (stream_finite_lemma7 RS mp RS mp) 1), + (etac is_ub_thelub 1), + (atac 1) + ]); + +(* ----------------------------------------------------------------------- *) +(* alternative prove for admissibility of ~stream_finite *) +(* show that stream_finite(s) = (? n. iterate n stl s = UU) *) +(* and prove adm. of ~(? n. iterate n stl s = UU) *) +(* proof uses theorems stream_take_lemma5-7; stream_finite_lemma8 *) +(* ----------------------------------------------------------------------- *) + + +qed_goal "adm_not_stream_finite" Stream.thy "adm(%s. ~ stream_finite(s))" + (fn prems => + [ + (subgoal_tac "(!s.(~stream_finite(s))=(!n.iterate n stl s ~=UU))" 1), + (etac (adm_cong RS iffD2)1), + (REPEAT(resolve_tac adm_thms 1)), + (rtac cont_iterate2 1), + (rtac allI 1), + (rtac (stream_finite_lemma8 RS ssubst) 1), + (fast_tac HOL_cs 1) + ]); + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Stream.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Stream.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,115 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Theory for streams without defined empty stream + 'a stream = 'a ** ('a stream)u + +The type is axiomatized as the least solution of the domain equation above. +The functor term that specifies the domain equation is: + + FT = <**,K_{'a},U> + +For details see chapter 5 of: + +[Franz Regensburger] HOLCF: Eine konservative Erweiterung von HOL um LCF, + Dissertation, Technische Universit"at M"unchen, 1994 +*) + +Stream = Dnat2 + + +types stream 1 + +(* ----------------------------------------------------------------------- *) +(* arity axiom is validated by semantic reasoning *) +(* partial ordering is implicit in the isomorphism axioms and their cont. *) + +arities stream::(pcpo)pcpo + +consts + +(* ----------------------------------------------------------------------- *) +(* essential constants *) + +stream_rep :: "('a stream) -> ('a ** ('a stream)u)" +stream_abs :: "('a ** ('a stream)u) -> ('a stream)" + +(* ----------------------------------------------------------------------- *) +(* abstract constants and auxiliary constants *) + +stream_copy :: "('a stream -> 'a stream) ->'a stream -> 'a stream" + +scons :: "'a -> 'a stream -> 'a stream" +stream_when :: "('a -> 'a stream -> 'b) -> 'a stream -> 'b" +is_scons :: "'a stream -> tr" +shd :: "'a stream -> 'a" +stl :: "'a stream -> 'a stream" +stream_take :: "nat => 'a stream -> 'a stream" +stream_finite :: "'a stream => bool" +stream_bisim :: "('a stream => 'a stream => bool) => bool" + +rules + +(* ----------------------------------------------------------------------- *) +(* axiomatization of recursive type 'a stream *) +(* ----------------------------------------------------------------------- *) +(* ('a stream,stream_abs) is the initial F-algebra where *) +(* F is the locally continuous functor determined by functor term FT. *) +(* domain equation: 'a stream = 'a ** ('a stream)u *) +(* functor term: FT = <**,K_{'a},U> *) +(* ----------------------------------------------------------------------- *) +(* stream_abs is an isomorphism with inverse stream_rep *) +(* identity is the least endomorphism on 'a stream *) + +stream_abs_iso "stream_rep`(stream_abs`x) = x" +stream_rep_iso "stream_abs`(stream_rep`x) = x" +stream_copy_def "stream_copy == (LAM f. stream_abs oo + (ssplit`(LAM x y. (|x , (lift`(up oo f))`y|) )) oo stream_rep)" +stream_reach "(fix`stream_copy)`x = x" + +defs +(* ----------------------------------------------------------------------- *) +(* properties of additional constants *) +(* ----------------------------------------------------------------------- *) +(* constructors *) + +scons_def "scons == (LAM x l. stream_abs`(| x, up`l |))" + +(* ----------------------------------------------------------------------- *) +(* discriminator functional *) + +stream_when_def +"stream_when == (LAM f l.ssplit `(LAM x l.f`x`(lift`ID`l)) `(stream_rep`l))" + +(* ----------------------------------------------------------------------- *) +(* discriminators and selectors *) + +is_scons_def "is_scons == stream_when`(LAM x l.TT)" +shd_def "shd == stream_when`(LAM x l.x)" +stl_def "stl == stream_when`(LAM x l.l)" + +(* ----------------------------------------------------------------------- *) +(* the taker for streams *) + +stream_take_def "stream_take == (%n.iterate n stream_copy UU)" + +(* ----------------------------------------------------------------------- *) + +stream_finite_def "stream_finite == (%s.? n.stream_take n `s=s)" + +(* ----------------------------------------------------------------------- *) +(* definition of bisimulation is determined by domain equation *) +(* simplification and rewriting for abstract constants yields def below *) + +stream_bisim_def "stream_bisim == +(%R.!s1 s2. + R s1 s2 --> + ((s1=UU & s2=UU) | + (? x s11 s21. x~=UU & s1=scons`x`s11 & s2 = scons`x`s21 & R s11 s21)))" + +end + + + + diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Stream2.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Stream2.ML Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,43 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Lemmas for theory Stream2.thy +*) + +open Stream2; + +(* ------------------------------------------------------------------------- *) +(* expand fixed point properties *) +(* ------------------------------------------------------------------------- *) + +val smap_def2 = fix_prover2 Stream2.thy smap_def + "smap = (LAM f s. stream_when`(LAM x l.scons`(f`x) `(smap`f`l)) `s)"; + + +(* ------------------------------------------------------------------------- *) +(* recursive properties *) +(* ------------------------------------------------------------------------- *) + + +qed_goal "smap1" Stream2.thy "smap`f`UU = UU" + (fn prems => + [ + (rtac (smap_def2 RS ssubst) 1), + (simp_tac (!simpset addsimps stream_when) 1) + ]); + +qed_goal "smap2" Stream2.thy + "x~=UU ==> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs)" + (fn prems => + [ + (cut_facts_tac prems 1), + (rtac trans 1), + (rtac (smap_def2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac refl 1) + ]); + + +val stream2_rews = [smap1, smap2]; diff -r 6960ec882bca -r ea0668a1c0ba src/HOLCF/explicit_domains/Stream2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/explicit_domains/Stream2.thy Fri Oct 06 17:25:24 1995 +0100 @@ -0,0 +1,29 @@ +(* + ID: $Id$ + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen + +Additional constants for stream +*) + +Stream2 = Stream + + +consts + +smap :: "('a -> 'b) -> 'a stream -> 'b stream" + +defs + +smap_def + "smap == fix`(LAM h f s. stream_when`(LAM x l.scons `(f`x) `(h`f`l)) `s)" + + +end + + +(* + smap`f`UU = UU + x~=UU --> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs) + +*) +