--- a/src/HOLCF/cfun3.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,403 +0,0 @@
-(* Title: HOLCF/cfun3.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-*)
-
-open Cfun3;
-
-(* ------------------------------------------------------------------------ *)
-(* the contlub property for fapp its 'first' argument *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_fapp1 = prove_goal Cfun3.thy "contlub(fapp)"
-(fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (expand_fun_eq RS iffD2) 1),
- (strip_tac 1),
- (rtac (lub_cfun RS thelubI RS ssubst) 1),
- (atac 1),
- (rtac (Cfunapp2 RS ssubst) 1),
- (etac contX_lubcfun 1),
- (rtac (lub_fun RS thelubI RS ssubst) 1),
- (etac (monofun_fapp1 RS ch2ch_monofun) 1),
- (rtac refl 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* the contX property for fapp in its first argument *)
-(* ------------------------------------------------------------------------ *)
-
-val contX_fapp1 = prove_goal Cfun3.thy "contX(fapp)"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_fapp1 1),
- (rtac contlub_fapp1 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* contlub, contX properties of fapp in its first argument in mixfix _[_] *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_cfun_fun = prove_goal Cfun3.thy
-"is_chain(FY) ==>\
-\ lub(range(FY))[x] = lub(range(%i.FY(i)[x]))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac trans 1),
- (etac (contlub_fapp1 RS contlubE RS spec RS mp RS fun_cong) 1),
- (rtac (thelub_fun RS ssubst) 1),
- (etac (monofun_fapp1 RS ch2ch_monofun) 1),
- (rtac refl 1)
- ]);
-
-
-val contX_cfun_fun = prove_goal Cfun3.thy
-"is_chain(FY) ==>\
-\ range(%i.FY(i)[x]) <<| lub(range(FY))[x]"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac thelubE 1),
- (etac ch2ch_fappL 1),
- (etac (contlub_cfun_fun RS sym) 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* contlub, contX properties of fapp in both argument in mixfix _[_] *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_cfun = prove_goal Cfun3.thy
-"[|is_chain(FY);is_chain(TY)|] ==>\
-\ lub(range(FY))[lub(range(TY))] = lub(range(%i.FY(i)[TY(i)]))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contlub_CF2 1),
- (rtac contX_fapp1 1),
- (rtac allI 1),
- (rtac contX_fapp2 1),
- (atac 1),
- (atac 1)
- ]);
-
-val contX_cfun = prove_goal Cfun3.thy
-"[|is_chain(FY);is_chain(TY)|] ==>\
-\ range(%i.FY(i)[TY(i)]) <<| lub(range(FY))[lub(range(TY))]"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac thelubE 1),
- (rtac (monofun_fapp1 RS ch2ch_MF2LR) 1),
- (rtac allI 1),
- (rtac monofun_fapp2 1),
- (atac 1),
- (atac 1),
- (etac (contlub_cfun RS sym) 1),
- (atac 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* contX2contX lemma for fapp *)
-(* ------------------------------------------------------------------------ *)
-
-val contX2contX_fapp = prove_goal Cfun3.thy
- "[|contX(%x.ft(x));contX(%x.tt(x))|] ==> contX(%x.(ft(x))[tt(x)])"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contX2contX_app2 1),
- (rtac contX2contX_app2 1),
- (rtac contX_const 1),
- (rtac contX_fapp1 1),
- (atac 1),
- (rtac contX_fapp2 1),
- (atac 1)
- ]);
-
-
-
-(* ------------------------------------------------------------------------ *)
-(* contX2mono Lemma for %x. LAM y. c1(x,y) *)
-(* ------------------------------------------------------------------------ *)
-
-val contX2mono_LAM = prove_goal Cfun3.thy
- "[|!x.contX(c1(x)); !y.monofun(%x.c1(x,y))|] ==>\
-\ monofun(%x. LAM y. c1(x,y))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monofunI 1),
- (strip_tac 1),
- (rtac (less_cfun RS ssubst) 1),
- (rtac (less_fun RS ssubst) 1),
- (rtac allI 1),
- (rtac (beta_cfun RS ssubst) 1),
- (etac spec 1),
- (rtac (beta_cfun RS ssubst) 1),
- (etac spec 1),
- (etac ((hd (tl prems)) RS spec RS monofunE RS spec RS spec RS mp) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* contX2contX Lemma for %x. LAM y. c1(x,y) *)
-(* ------------------------------------------------------------------------ *)
-
-val contX2contX_LAM = prove_goal Cfun3.thy
- "[| !x.contX(c1(x)); !y.contX(%x.c1(x,y)) |] ==> contX(%x. LAM y. c1(x,y))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monocontlub2contX 1),
- (etac contX2mono_LAM 1),
- (rtac (contX2mono RS allI) 1),
- (etac spec 1),
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (thelub_cfun RS ssubst) 1),
- (rtac (contX2mono_LAM RS ch2ch_monofun) 1),
- (atac 1),
- (rtac (contX2mono RS allI) 1),
- (etac spec 1),
- (atac 1),
- (res_inst_tac [("f","fabs")] arg_cong 1),
- (rtac ext 1),
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (etac spec 1),
- (rtac (contX2contlub RS contlubE
- RS spec RS mp ) 1),
- (etac spec 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* elimination of quantifier in premisses of contX2contX_LAM yields good *)
-(* lemma for the contX tactic *)
-(* ------------------------------------------------------------------------ *)
-
-val contX2contX_LAM2 = (allI RSN (2,(allI RS contX2contX_LAM)));
-(*
- [| !!x. contX(?c1.0(x)); !!y. contX(%x. ?c1.0(x,y)) |] ==>
- contX(%x. LAM y. ?c1.0(x,y))
-*)
-
-(* ------------------------------------------------------------------------ *)
-(* contX2contX tactic *)
-(* ------------------------------------------------------------------------ *)
-
-val contX_lemmas = [contX_const, contX_id, contX_fapp2,
- contX2contX_fapp,contX2contX_LAM2];
-
-
-val contX_tac = (fn i => (resolve_tac contX_lemmas i));
-
-val contX_tacR = (fn i => (REPEAT (contX_tac i)));
-
-(* ------------------------------------------------------------------------ *)
-(* function application _[_] is strict in its first arguments *)
-(* ------------------------------------------------------------------------ *)
-
-val strict_fapp1 = prove_goal Cfun3.thy "UU[x] = UU"
- (fn prems =>
- [
- (rtac (inst_cfun_pcpo RS ssubst) 1),
- (rewrite_goals_tac [UU_cfun_def]),
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tac 1),
- (rtac refl 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* results about strictify *)
-(* ------------------------------------------------------------------------ *)
-
-val Istrictify1 = prove_goalw Cfun3.thy [Istrictify_def]
- "Istrictify(f)(UU)=UU"
- (fn prems =>
- [
- (rtac select_equality 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val Istrictify2 = prove_goalw Cfun3.thy [Istrictify_def]
- "~x=UU ==> Istrictify(f)(x)=f[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val monofun_Istrictify1 = prove_goal Cfun3.thy "monofun(Istrictify)"
- (fn prems =>
- [
- (rtac monofunI 1),
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (res_inst_tac [("Q","xa=UU")] (excluded_middle RS disjE) 1),
- (rtac (Istrictify2 RS ssubst) 1),
- (atac 1),
- (rtac (Istrictify2 RS ssubst) 1),
- (atac 1),
- (rtac monofun_cfun_fun 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac (Istrictify1 RS ssubst) 1),
- (rtac (Istrictify1 RS ssubst) 1),
- (rtac refl_less 1)
- ]);
-
-val monofun_Istrictify2 = prove_goal Cfun3.thy "monofun(Istrictify(f))"
- (fn prems =>
- [
- (rtac monofunI 1),
- (strip_tac 1),
- (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1),
- (rtac (Istrictify2 RS ssubst) 1),
- (etac notUU_I 1),
- (atac 1),
- (rtac (Istrictify2 RS ssubst) 1),
- (atac 1),
- (rtac monofun_cfun_arg 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac (Istrictify1 RS ssubst) 1),
- (rtac minimal 1)
- ]);
-
-
-val contlub_Istrictify1 = prove_goal Cfun3.thy "contlub(Istrictify)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (expand_fun_eq RS iffD2) 1),
- (strip_tac 1),
- (rtac (thelub_fun RS ssubst) 1),
- (etac (monofun_Istrictify1 RS ch2ch_monofun) 1),
- (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1),
- (rtac (Istrictify2 RS ssubst) 1),
- (atac 1),
- (rtac (Istrictify2 RS ext RS ssubst) 1),
- (atac 1),
- (rtac (thelub_cfun RS ssubst) 1),
- (atac 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_lubcfun 1),
- (atac 1),
- (rtac refl 1),
- (hyp_subst_tac 1),
- (rtac (Istrictify1 RS ssubst) 1),
- (rtac (Istrictify1 RS ext RS ssubst) 1),
- (rtac (chain_UU_I_inverse RS sym) 1),
- (rtac (refl RS allI) 1)
- ]);
-
-val contlub_Istrictify2 = prove_goal Cfun3.thy "contlub(Istrictify(f))"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (res_inst_tac [("Q","lub(range(Y))=UU")] classical2 1),
- (res_inst_tac [("t","lub(range(Y))")] subst 1),
- (rtac sym 1),
- (atac 1),
- (rtac (Istrictify1 RS ssubst) 1),
- (rtac sym 1),
- (rtac chain_UU_I_inverse 1),
- (strip_tac 1),
- (res_inst_tac [("t","Y(i)"),("s","UU")] subst 1),
- (rtac sym 1),
- (rtac (chain_UU_I RS spec) 1),
- (atac 1),
- (atac 1),
- (rtac Istrictify1 1),
- (rtac (Istrictify2 RS ssubst) 1),
- (atac 1),
- (res_inst_tac [("s","lub(range(%i. f[Y(i)]))")] trans 1),
- (rtac contlub_cfun_arg 1),
- (atac 1),
- (rtac lub_equal2 1),
- (rtac (chain_mono2 RS exE) 1),
- (atac 2),
- (rtac chain_UU_I_inverse2 1),
- (atac 1),
- (rtac exI 1),
- (strip_tac 1),
- (rtac (Istrictify2 RS sym) 1),
- (fast_tac HOL_cs 1),
- (rtac ch2ch_monofun 1),
- (rtac monofun_fapp2 1),
- (atac 1),
- (rtac ch2ch_monofun 1),
- (rtac monofun_Istrictify2 1),
- (atac 1)
- ]);
-
-
-val contX_Istrictify1 = (contlub_Istrictify1 RS
- (monofun_Istrictify1 RS monocontlub2contX));
-
-val contX_Istrictify2 = (contlub_Istrictify2 RS
- (monofun_Istrictify2 RS monocontlub2contX));
-
-
-val strictify1 = prove_goalw Cfun3.thy [strictify_def]
- "strictify[f][UU]=UU"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tac 1),
- (rtac contX_Istrictify2 1),
- (rtac contX2contX_CF1L 1),
- (rtac contX_Istrictify1 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Istrictify2 1),
- (rtac Istrictify1 1)
- ]);
-
-val strictify2 = prove_goalw Cfun3.thy [strictify_def]
- "~x=UU ==> strictify[f][x]=f[x]"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tac 1),
- (rtac contX_Istrictify2 1),
- (rtac contX2contX_CF1L 1),
- (rtac contX_Istrictify1 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Istrictify2 1),
- (rtac Istrictify2 1),
- (resolve_tac prems 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Instantiate the simplifier *)
-(* ------------------------------------------------------------------------ *)
-
-val Cfun_rews = [minimal,refl_less,beta_cfun,strict_fapp1,strictify1,
- strictify2];
-
-(* ------------------------------------------------------------------------ *)
-(* use contX_tac as autotac. *)
-(* ------------------------------------------------------------------------ *)
-
-val Cfun_ss = HOL_ss
- addsimps Cfun_rews
- setsolver
- (fn thms => (resolve_tac (TrueI::refl::thms)) ORELSE' atac ORELSE'
- (fn i => DEPTH_SOLVE_1 (contX_tac i))
- );
--- a/src/HOLCF/cfun3.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,31 +0,0 @@
-(* Title: HOLCF/cfun3.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class instance of -> for class pcpo
-
-*)
-
-Cfun3 = Cfun2 +
-
-arities "->" :: (pcpo,pcpo)pcpo (* Witness cfun2.ML *)
-
-consts
- Istrictify :: "('a->'b)=>'a=>'b"
- strictify :: "('a->'b)->'a->'b"
-
-rules
-
-inst_cfun_pcpo "UU::'a->'b = UU_cfun"
-
-Istrictify_def "Istrictify(f,x) == (@z.\
-\ ( x=UU --> z = UU)\
-\ & (~x=UU --> z = f[x]))"
-
-strictify_def "strictify == (LAM f x.Istrictify(f,x))"
-
-end
-
-
-
--- a/src/HOLCF/cinfix.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,73 +0,0 @@
-(* Title: HOLCF/cinfix.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-Some functions for the print and parse translation of continuous functions
-
-Suppose the user introduces the following notation for the continuous
-infixl . and the cont. infixr # with binding power 100
-
-consts
- "." :: "'a => 'b => ('a**'b)" ("_._" [100,101] 100)
- "cop ." :: "'a -> 'b -> ('a**'b)" ("spair")
-
- "#" :: "'a => 'b => ('a**'b)" ("_#_" [101,100] 100)
- "cop #" :: "'a -> 'b -> ('a**'b)" ("spair2")
-
-the following functions are needed to set up proper translations
-*)
-
-(* -----------------------------------------------------------------------
- a general purpose parse translation for continuous infix operators
- this functions must be used for every cont. infix
- ----------------------------------------------------------------------- *)
-
-fun mk_cinfixtr id =
- (fn ts =>
- let val Cfapp = Const("fapp",dummyT) in
- Cfapp $ (Cfapp$Const("cop "^id,dummyT)$(nth_elem (0,ts)))$
- (nth_elem (1,ts))
- end);
-
-
-
-(* -----------------------------------------------------------------------
- make a print translation for a cont. infix operator "cop ???"
- this is a print translation for fapp and is installed only once
- special translations for other mixfixes (e.g. If_then_else_fi) are also
- defined.
- ----------------------------------------------------------------------- *)
-
-fun fapptr' ts =
- case ts of
- [Const("fapp",T1)$Const(s,T2)$t1,t2] =>
- if ["c","o","p"," "] = take(4, explode s)
- then Const(implode(drop(4, explode s)),dummyT)$t1$t2
- else raise Match
- | [Const("fapp",dummyT)$
- (Const("fapp",T1)$Const("Icifte",T2)$t)$e1,e2]
- => Const("@cifte",dummyT)$t$e1$e2
- | _ => raise Match;
-
-
-(* -----------------------------------------------------------------------
-
-for the example above, the following must be setup in the ML section
-
-val parse_translation = [(".",mk_cinfixtr "."),
- ("#",mk_cinfixtr "#")];
-
-
-the print translation for fapp is setup only once in the system
-
-val print_translation = [("fapp",fapptr')];
-
- ----------------------------------------------------------------------- *)
-
-
-
-
-
-
--- a/src/HOLCF/cont.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,670 +0,0 @@
-(* Title: HOLCF/cont.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for cont.thy
-*)
-
-open Cont;
-
-(* ------------------------------------------------------------------------ *)
-(* access to definition *)
-(* ------------------------------------------------------------------------ *)
-
-val contlubI = prove_goalw Cont.thy [contlub]
- "! Y. is_chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))==>\
-\ contlub(f)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-val contlubE = prove_goalw Cont.thy [contlub]
- " contlub(f)==>\
-\ ! Y. is_chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-
-val contXI = prove_goalw Cont.thy [contX]
- "! Y. is_chain(Y) --> range(% i.f(Y(i))) <<| f(lub(range(Y))) ==> contX(f)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-val contXE = prove_goalw Cont.thy [contX]
- "contX(f) ==> ! Y. is_chain(Y) --> range(% i.f(Y(i))) <<| f(lub(range(Y)))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-
-val monofunI = prove_goalw Cont.thy [monofun]
- "! x y. x << y --> f(x) << f(y) ==> monofun(f)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-val monofunE = prove_goalw Cont.thy [monofun]
- "monofun(f) ==> ! x y. x << y --> f(x) << f(y)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the main purpose of cont.thy is to show: *)
-(* monofun(f) & contlub(f) <==> contX(f) *)
-(* ------------------------------------------------------------------------ *)
-
-(* ------------------------------------------------------------------------ *)
-(* monotone functions map chains to chains *)
-(* ------------------------------------------------------------------------ *)
-
-val ch2ch_monofun= prove_goal Cont.thy
- "[| monofun(f); is_chain(Y) |] ==> is_chain(%i. f(Y(i)))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_chainI 1),
- (rtac allI 1),
- (etac (monofunE RS spec RS spec RS mp) 1),
- (etac (is_chainE RS spec) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* monotone functions map upper bound to upper bounds *)
-(* ------------------------------------------------------------------------ *)
-
-val ub2ub_monofun = prove_goal Cont.thy
- "[| monofun(f); range(Y) <| u|] ==> range(%i.f(Y(i))) <| f(u)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (etac (monofunE RS spec RS spec RS mp) 1),
- (etac (ub_rangeE RS spec) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* left to right: monofun(f) & contlub(f) ==> contX(f) *)
-(* ------------------------------------------------------------------------ *)
-
-val monocontlub2contX = prove_goalw Cont.thy [contX]
- "[|monofun(f);contlub(f)|] ==> contX(f)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac thelubE 1),
- (etac ch2ch_monofun 1),
- (atac 1),
- (etac (contlubE RS spec RS mp RS sym) 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* first a lemma about binary chains *)
-(* ------------------------------------------------------------------------ *)
-
-val binchain_contX = prove_goal Cont.thy
-"[| contX(f); x << y |] ==> range(%i. f(if(i = 0,x,y))) <<| f(y)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac subst 1),
- (etac (contXE RS spec RS mp) 2),
- (etac bin_chain 2),
- (res_inst_tac [("y","y")] arg_cong 1),
- (etac (lub_bin_chain RS thelubI) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* right to left: contX(f) ==> monofun(f) & contlub(f) *)
-(* part1: contX(f) ==> monofun(f *)
-(* ------------------------------------------------------------------------ *)
-
-val contX2mono = prove_goalw Cont.thy [monofun]
- "contX(f) ==> monofun(f)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (res_inst_tac [("s","if(0 = 0,x,y)")] subst 1),
- (rtac (binchain_contX RS is_ub_lub) 2),
- (atac 2),
- (atac 2),
- (simp_tac nat_ss 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* right to left: contX(f) ==> monofun(f) & contlub(f) *)
-(* part2: contX(f) ==> contlub(f) *)
-(* ------------------------------------------------------------------------ *)
-
-val contX2contlub = prove_goalw Cont.thy [contlub]
- "contX(f) ==> contlub(f)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac (thelubI RS sym) 1),
- (etac (contXE RS spec RS mp) 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* The following results are about a curried function that is monotone *)
-(* in both arguments *)
-(* ------------------------------------------------------------------------ *)
-
-val ch2ch_MF2L = prove_goal Cont.thy
-"[|monofun(MF2::('a::po=>'b::po=>'c::po));\
-\ is_chain(F)|] ==> is_chain(%i. MF2(F(i),x))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac (ch2ch_monofun RS ch2ch_fun) 1),
- (atac 1)
- ]);
-
-
-val ch2ch_MF2R = prove_goal Cont.thy "[|monofun(MF2(f)::('b::po=>'c::po));\
-\ is_chain(Y)|] ==> is_chain(%i. MF2(f,Y(i)))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac ch2ch_monofun 1),
- (atac 1)
- ]);
-
-val ch2ch_MF2LR = prove_goal Cont.thy
-"[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\
-\ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\
-\ is_chain(F); is_chain(Y)|] ==> \
-\ is_chain(%i. MF2(F(i))(Y(i)))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_chainI 1),
- (strip_tac 1 ),
- (rtac trans_less 1),
- (etac (ch2ch_MF2L RS is_chainE RS spec) 1),
- (atac 1),
- ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)),
- (etac (is_chainE RS spec) 1)
- ]);
-
-val ch2ch_lubMF2R = prove_goal Cont.thy
-"[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\
-\ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\
-\ is_chain(F);is_chain(Y)|] ==> \
-\ is_chain(%j. lub(range(%i. MF2(F(j),Y(i)))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (lub_mono RS allI RS is_chainI) 1),
- ((rtac ch2ch_MF2R 1) THEN (etac spec 1)),
- (atac 1),
- ((rtac ch2ch_MF2R 1) THEN (etac spec 1)),
- (atac 1),
- (strip_tac 1),
- (rtac (is_chainE RS spec) 1),
- (etac ch2ch_MF2L 1),
- (atac 1)
- ]);
-
-
-val ch2ch_lubMF2L = prove_goal Cont.thy
-"[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\
-\ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\
-\ is_chain(F);is_chain(Y)|] ==> \
-\ is_chain(%i. lub(range(%j. MF2(F(j),Y(i)))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (lub_mono RS allI RS is_chainI) 1),
- (etac ch2ch_MF2L 1),
- (atac 1),
- (etac ch2ch_MF2L 1),
- (atac 1),
- (strip_tac 1),
- (rtac (is_chainE RS spec) 1),
- ((rtac ch2ch_MF2R 1) THEN (etac spec 1)),
- (atac 1)
- ]);
-
-
-val lub_MF2_mono = prove_goal Cont.thy
-"[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\
-\ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\
-\ is_chain(F)|] ==> \
-\ monofun(% x.lub(range(% j.MF2(F(j),x))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monofunI 1),
- (strip_tac 1),
- (rtac lub_mono 1),
- (etac ch2ch_MF2L 1),
- (atac 1),
- (etac ch2ch_MF2L 1),
- (atac 1),
- (strip_tac 1),
- ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)),
- (atac 1)
- ]);
-
-
-val ex_lubMF2 = prove_goal Cont.thy
-"[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\
-\ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\
-\ is_chain(F); is_chain(Y)|] ==> \
-\ lub(range(%j. lub(range(%i. MF2(F(j),Y(i)))))) =\
-\ lub(range(%i. lub(range(%j. MF2(F(j),Y(i))))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac antisym_less 1),
- (rtac is_lub_thelub 1),
- (etac ch2ch_lubMF2R 1),
- (atac 1),(atac 1),(atac 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (rtac lub_mono 1),
- ((rtac ch2ch_MF2R 1) THEN (etac spec 1)),
- (atac 1),
- (etac ch2ch_lubMF2L 1),
- (atac 1),(atac 1),(atac 1),
- (strip_tac 1),
- (rtac is_ub_thelub 1),
- (etac ch2ch_MF2L 1),(atac 1),
- (rtac is_lub_thelub 1),
- (etac ch2ch_lubMF2L 1),
- (atac 1),(atac 1),(atac 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (rtac lub_mono 1),
- (etac ch2ch_MF2L 1),(atac 1),
- (etac ch2ch_lubMF2R 1),
- (atac 1),(atac 1),(atac 1),
- (strip_tac 1),
- (rtac is_ub_thelub 1),
- ((rtac ch2ch_MF2R 1) THEN (etac spec 1)),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* The following results are about a curried function that is continuous *)
-(* in both arguments *)
-(* ------------------------------------------------------------------------ *)
-
-val diag_lubCF2_1 = prove_goal Cont.thy
-"[|contX(CF2);!f.contX(CF2(f));is_chain(FY);is_chain(TY)|] ==>\
-\ lub(range(%i. lub(range(%j. CF2(FY(j))(TY(i)))))) =\
-\ lub(range(%i. CF2(FY(i))(TY(i))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac antisym_less 1),
- (rtac is_lub_thelub 1),
- (rtac ch2ch_lubMF2L 1),
- (rtac contX2mono 1),
- (atac 1),
- (rtac allI 1),
- (rtac contX2mono 1),
- (etac spec 1),
- (atac 1),
- (atac 1),
- (rtac ub_rangeI 1),
- (strip_tac 1 ),
- (rtac is_lub_thelub 1),
- ((rtac ch2ch_MF2L 1) THEN (rtac contX2mono 1) THEN (atac 1)),
- (atac 1),
- (rtac ub_rangeI 1),
- (strip_tac 1 ),
- (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1),
- (rtac trans_less 1),
- (rtac is_ub_thelub 2),
- (rtac (chain_mono RS mp) 1),
- ((rtac ch2ch_MF2R 1) THEN (rtac contX2mono 1) THEN (etac spec 1)),
- (atac 1),
- (atac 1),
- ((rtac ch2ch_MF2LR 1) THEN (etac contX2mono 1)),
- (rtac allI 1),
- ((rtac contX2mono 1) THEN (etac spec 1)),
- (atac 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac is_ub_thelub 1),
- ((rtac ch2ch_MF2LR 1) THEN (etac contX2mono 1)),
- (rtac allI 1),
- ((rtac contX2mono 1) THEN (etac spec 1)),
- (atac 1),
- (atac 1),
- (rtac trans_less 1),
- (rtac is_ub_thelub 2),
- (res_inst_tac [("x1","ia")] (chain_mono RS mp) 1),
- ((rtac ch2ch_MF2L 1) THEN (etac contX2mono 1)),
- (atac 1),
- (atac 1),
- ((rtac ch2ch_MF2LR 1) THEN (etac contX2mono 1)),
- (rtac allI 1),
- ((rtac contX2mono 1) THEN (etac spec 1)),
- (atac 1),
- (atac 1),
- (rtac lub_mono 1),
- ((rtac ch2ch_MF2LR 1) THEN (etac contX2mono 1)),
- (rtac allI 1),
- ((rtac contX2mono 1) THEN (etac spec 1)),
- (atac 1),
- (atac 1),
- (rtac ch2ch_lubMF2L 1),
- (rtac contX2mono 1),
- (atac 1),
- (rtac allI 1),
- ((rtac contX2mono 1) THEN (etac spec 1)),
- (atac 1),
- (atac 1),
- (strip_tac 1 ),
- (rtac is_ub_thelub 1),
- ((rtac ch2ch_MF2L 1) THEN (etac contX2mono 1)),
- (atac 1)
- ]);
-
-
-val diag_lubCF2_2 = prove_goal Cont.thy
-"[|contX(CF2);!f.contX(CF2(f));is_chain(FY);is_chain(TY)|] ==>\
-\ lub(range(%j. lub(range(%i. CF2(FY(j))(TY(i)))))) =\
-\ lub(range(%i. CF2(FY(i))(TY(i))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac trans 1),
- (rtac ex_lubMF2 1),
- (rtac ((hd prems) RS contX2mono) 1),
- (rtac allI 1),
- (rtac (((hd (tl prems)) RS spec RS contX2mono)) 1),
- (atac 1),
- (atac 1),
- (rtac diag_lubCF2_1 1),
- (atac 1),
- (atac 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-val contlub_CF2 = prove_goal Cont.thy
-"[|contX(CF2);!f.contX(CF2(f));is_chain(FY);is_chain(TY)|] ==>\
-\ CF2(lub(range(FY)))(lub(range(TY))) = lub(range(%i.CF2(FY(i))(TY(i))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac ((hd prems) RS contX2contlub RS contlubE RS
- spec RS mp RS ssubst) 1),
- (atac 1),
- (rtac (thelub_fun RS ssubst) 1),
- (rtac ((hd prems) RS contX2mono RS ch2ch_monofun) 1),
- (atac 1),
- (rtac trans 1),
- (rtac (((hd (tl prems)) RS spec RS contX2contlub) RS
- contlubE RS spec RS mp RS ext RS arg_cong RS arg_cong) 1),
- (atac 1),
- (rtac diag_lubCF2_2 1),
- (atac 1),
- (atac 1),
- (atac 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* The following results are about application for functions in 'a=>'b *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_fun_fun = prove_goal Cont.thy
- "f1 << f2 ==> f1(x) << f2(x)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac (less_fun RS iffD1 RS spec) 1)
- ]);
-
-val monofun_fun_arg = prove_goal Cont.thy
- "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac (monofunE RS spec RS spec RS mp) 1),
- (atac 1)
- ]);
-
-val monofun_fun = prove_goal Cont.thy
-"[|monofun(f1); monofun(f2); f1 << f2; x1 << x2|] ==> f1(x1) << f2(x2)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac trans_less 1),
- (etac monofun_fun_arg 1),
- (atac 1),
- (etac monofun_fun_fun 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* The following results are about the propagation of monotonicity and *)
-(* continuity *)
-(* ------------------------------------------------------------------------ *)
-
-val mono2mono_MF1L = prove_goal Cont.thy
- "[|monofun(c1)|] ==> monofun(%x. c1(x,y))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monofunI 1),
- (strip_tac 1),
- (etac (monofun_fun_arg RS monofun_fun_fun) 1),
- (atac 1)
- ]);
-
-val contX2contX_CF1L = prove_goal Cont.thy
- "[|contX(c1)|] ==> contX(%x. c1(x,y))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monocontlub2contX 1),
- (etac (contX2mono RS mono2mono_MF1L) 1),
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac ((hd prems) RS contX2contlub RS
- contlubE RS spec RS mp RS ssubst) 1),
- (atac 1),
- (rtac (thelub_fun RS ssubst) 1),
- (rtac ch2ch_monofun 1),
- (etac contX2mono 1),
- (atac 1),
- (rtac refl 1)
- ]);
-
-(********* Note "(%x.%y.c1(x,y)) = c1" ***********)
-
-val mono2mono_MF1L_rev = prove_goal Cont.thy
- "!y.monofun(%x.c1(x,y)) ==> monofun(c1)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monofunI 1),
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (rtac ((hd prems) RS spec RS monofunE RS spec RS spec RS mp) 1),
- (atac 1)
- ]);
-
-val contX2contX_CF1L_rev = prove_goal Cont.thy
- "!y.contX(%x.c1(x,y)) ==> contX(c1)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monocontlub2contX 1),
- (rtac (contX2mono RS allI RS mono2mono_MF1L_rev ) 1),
- (etac spec 1),
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac ext 1),
- (rtac (thelub_fun RS ssubst) 1),
- (rtac (contX2mono RS allI RS mono2mono_MF1L_rev RS ch2ch_monofun) 1),
- (etac spec 1),
- (atac 1),
- (rtac
- ((hd prems) RS spec RS contX2contlub RS contlubE RS spec RS mp) 1),
- (atac 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* What D.A.Schmidt calls continuity of abstraction *)
-(* never used here *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_abstraction = prove_goal Cont.thy
-"[|is_chain(Y::nat=>'a);!y.contX(%x.(c::'a=>'b=>'c)(x,y))|] ==>\
-\ (%y.lub(range(%i.c(Y(i),y)))) = (lub(range(%i.%y.c(Y(i),y))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac trans 1),
- (rtac (contX2contlub RS contlubE RS spec RS mp) 2),
- (atac 3),
- (etac contX2contX_CF1L_rev 2),
- (rtac ext 1),
- (rtac (contX2contlub RS contlubE RS spec RS mp RS sym) 1),
- (etac spec 1),
- (atac 1)
- ]);
-
-
-val mono2mono_app = prove_goal Cont.thy
-"[|monofun(ft);!x.monofun(ft(x));monofun(tt)|] ==>\
-\ monofun(%x.(ft(x))(tt(x)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac monofunI 1),
- (strip_tac 1),
- (res_inst_tac [("f1.0","ft(x)"),("f2.0","ft(y)")] monofun_fun 1),
- (etac spec 1),
- (etac spec 1),
- (etac (monofunE RS spec RS spec RS mp) 1),
- (atac 1),
- (etac (monofunE RS spec RS spec RS mp) 1),
- (atac 1)
- ]);
-
-val contX2contlub_app = prove_goal Cont.thy
-"[|contX(ft);!x.contX(ft(x));contX(tt)|] ==>\
-\ contlub(%x.(ft(x))(tt(x)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contlubI 1),
- (strip_tac 1),
- (res_inst_tac [("f3","tt")] (contlubE RS spec RS mp RS ssubst) 1),
- (rtac contX2contlub 1),
- (resolve_tac prems 1),
- (atac 1),
- (rtac contlub_CF2 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (atac 1),
- (rtac (contX2mono RS ch2ch_monofun) 1),
- (resolve_tac prems 1),
- (atac 1)
- ]);
-
-
-val contX2contX_app = prove_goal Cont.thy
-"[|contX(ft);!x.contX(ft(x));contX(tt)|] ==>\
-\ contX(%x.(ft(x))(tt(x)))"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac mono2mono_app 1),
- (rtac contX2mono 1),
- (resolve_tac prems 1),
- (strip_tac 1),
- (rtac contX2mono 1),
- (cut_facts_tac prems 1),
- (etac spec 1),
- (rtac contX2mono 1),
- (resolve_tac prems 1),
- (rtac contX2contlub_app 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1)
- ]);
-
-
-val contX2contX_app2 = (allI RSN (2,contX2contX_app));
-(* [| contX(?ft); !!x. contX(?ft(x)); contX(?tt) |] ==> *)
-(* contX(%x. ?ft(x,?tt(x))) *)
-
-
-(* ------------------------------------------------------------------------ *)
-(* The identity function is continuous *)
-(* ------------------------------------------------------------------------ *)
-
-val contX_id = prove_goal Cont.thy "contX(% x.x)"
- (fn prems =>
- [
- (rtac contXI 1),
- (strip_tac 1),
- (etac thelubE 1),
- (rtac refl 1)
- ]);
-
-
-
-(* ------------------------------------------------------------------------ *)
-(* constant functions are continuous *)
-(* ------------------------------------------------------------------------ *)
-
-val contX_const = prove_goalw Cont.thy [contX] "contX(%x.c)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (rtac refl_less 1),
- (strip_tac 1),
- (dtac ub_rangeE 1),
- (etac spec 1)
- ]);
-
-
-val contX2contX_app3 = prove_goal Cont.thy
- "[|contX(f);contX(t) |] ==> contX(%x. f(t(x)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contX2contX_app2 1),
- (rtac contX_const 1),
- (atac 1),
- (atac 1)
- ]);
-
--- a/src/HOLCF/cont.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-(* Title: HOLCF/cont.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
- Results about continuity and monotonicity
-*)
-
-Cont = Fun3 +
-
-(*
-
- Now we change the default class! Form now on all untyped typevariables are
- of default class pcpo
-
-*)
-
-
-default pcpo
-
-consts
- monofun :: "('a::po => 'b::po) => bool" (* monotonicity *)
- contlub :: "('a => 'b) => bool" (* first cont. def *)
- contX :: "('a => 'b) => bool" (* secnd cont. def *)
-
-rules
-
-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))))"
-
-contX "contX(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) <==> contX(f) *)
-(* ------------------------------------------------------------------------ *)
-
-end
--- a/src/HOLCF/cprod1.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,117 +0,0 @@
-(* Title: HOLCF/cprod1.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for theory cprod1.thy
-*)
-
-open Cprod1;
-
-val less_cprod1b = prove_goalw Cprod1.thy [less_cprod_def]
- "less_cprod(p1,p2) = ( fst(p1) << fst(p2) & snd(p1) << snd(p2))"
- (fn prems =>
- [
- (rtac refl 1)
- ]);
-
-val less_cprod2a = prove_goalw Cprod1.thy [less_cprod_def]
- "less_cprod(<x,y>,<UU,UU>) ==> x = UU & y = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac conjE 1),
- (dtac (fst_conv RS subst) 1),
- (dtac (fst_conv RS subst) 1),
- (dtac (fst_conv RS subst) 1),
- (dtac (snd_conv RS subst) 1),
- (dtac (snd_conv RS subst) 1),
- (dtac (snd_conv RS subst) 1),
- (rtac conjI 1),
- (etac UU_I 1),
- (etac UU_I 1)
- ]);
-
-val less_cprod2b = prove_goal Cprod1.thy
- "less_cprod(p,<UU,UU>) ==> p=<UU,UU>"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p")] PairE 1),
- (hyp_subst_tac 1),
- (dtac less_cprod2a 1),
- (asm_simp_tac HOL_ss 1)
- ]);
-
-val less_cprod2c = prove_goalw Cprod1.thy [less_cprod_def]
- "less_cprod(<x1,y1>,<x2,y2>) ==> x1 << x2 & y1 << y2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac conjE 1),
- (dtac (fst_conv RS subst) 1),
- (dtac (fst_conv RS subst) 1),
- (dtac (fst_conv RS subst) 1),
- (dtac (snd_conv RS subst) 1),
- (dtac (snd_conv RS subst) 1),
- (dtac (snd_conv RS subst) 1),
- (rtac conjI 1),
- (atac 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* less_cprod is a partial order on 'a * 'b *)
-(* ------------------------------------------------------------------------ *)
-
-val refl_less_cprod = prove_goalw Cprod1.thy [less_cprod_def] "less_cprod(p,p)"
- (fn prems =>
- [
- (res_inst_tac [("p","p")] PairE 1),
- (hyp_subst_tac 1),
- (simp_tac pair_ss 1),
- (simp_tac Cfun_ss 1)
- ]);
-
-val antisym_less_cprod = prove_goal Cprod1.thy
- "[|less_cprod(p1,p2);less_cprod(p2,p1)|] ==> p1=p2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] PairE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] PairE 1),
- (hyp_subst_tac 1),
- (dtac less_cprod2c 1),
- (dtac less_cprod2c 1),
- (etac conjE 1),
- (etac conjE 1),
- (rtac (Pair_eq RS ssubst) 1),
- (fast_tac (HOL_cs addSIs [antisym_less]) 1)
- ]);
-
-
-val trans_less_cprod = prove_goal Cprod1.thy
- "[|less_cprod(p1,p2);less_cprod(p2,p3)|] ==> less_cprod(p1,p3)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] PairE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p3")] PairE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] PairE 1),
- (hyp_subst_tac 1),
- (dtac less_cprod2c 1),
- (dtac less_cprod2c 1),
- (rtac (less_cprod1b RS ssubst) 1),
- (simp_tac pair_ss 1),
- (etac conjE 1),
- (etac conjE 1),
- (rtac conjI 1),
- (etac trans_less 1),
- (atac 1),
- (etac trans_less 1),
- (atac 1)
- ]);
-
--- a/src/HOLCF/cprod1.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,23 +0,0 @@
-(* Title: HOLCF/cprod1.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-Partial ordering for cartesian product of HOL theory prod.thy
-
-*)
-
-Cprod1 = Cfun3 +
-
-
-consts
- less_cprod :: "[('a::pcpo * 'b::pcpo),('a * 'b)] => bool"
-
-rules
-
- less_cprod_def "less_cprod(p1,p2) == ( fst(p1) << fst(p2) &\
-\ snd(p1) << snd(p2))"
-
-end
-
--- a/src/HOLCF/cprod2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,181 +0,0 @@
-(* Title: HOLCF/cprod2.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for cprod2.thy
-*)
-
-open Cprod2;
-
-val less_cprod3a = prove_goal Cprod2.thy
- "p1=<UU,UU> ==> p1 << p2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (inst_cprod_po RS ssubst) 1),
- (rtac (less_cprod1b RS ssubst) 1),
- (hyp_subst_tac 1),
- (asm_simp_tac pair_ss 1),
- (rtac conjI 1),
- (rtac minimal 1),
- (rtac minimal 1)
- ]);
-
-val less_cprod3b = prove_goal Cprod2.thy
- "(p1 << p2) = (fst(p1)<<fst(p2) & snd(p1)<<snd(p2))"
- (fn prems =>
- [
- (rtac (inst_cprod_po RS ssubst) 1),
- (rtac less_cprod1b 1)
- ]);
-
-val less_cprod4a = prove_goal Cprod2.thy
- "<x1,x2> << <UU,UU> ==> x1=UU & x2=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac less_cprod2a 1),
- (etac (inst_cprod_po RS subst) 1)
- ]);
-
-val less_cprod4b = prove_goal Cprod2.thy
- "p << <UU,UU> ==> p = <UU,UU>"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac less_cprod2b 1),
- (etac (inst_cprod_po RS subst) 1)
- ]);
-
-val less_cprod4c = prove_goal Cprod2.thy
- " <xa,ya> << <x,y> ==> xa<<x & ya << y"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac less_cprod2c 1),
- (etac (inst_cprod_po RS subst) 1),
- (REPEAT (atac 1))
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* type cprod is pointed *)
-(* ------------------------------------------------------------------------ *)
-
-val minimal_cprod = prove_goal Cprod2.thy "<UU,UU><<p"
-(fn prems =>
- [
- (rtac less_cprod3a 1),
- (rtac refl 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Pair <_,_> is monotone in both arguments *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_pair1 = prove_goalw Cprod2.thy [monofun] "monofun(Pair)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (rtac (less_cprod3b RS iffD2) 1),
- (simp_tac pair_ss 1),
- (asm_simp_tac Cfun_ss 1)
- ]);
-
-val monofun_pair2 = prove_goalw Cprod2.thy [monofun] "monofun(Pair(x))"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac (less_cprod3b RS iffD2) 1),
- (simp_tac pair_ss 1),
- (asm_simp_tac Cfun_ss 1)
- ]);
-
-val monofun_pair = prove_goal Cprod2.thy
- "[|x1<<x2; y1<<y2|] ==> <x1,y1> << <x2,y2>"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac trans_less 1),
- (rtac (monofun_pair1 RS monofunE RS spec RS spec RS mp RS
- (less_fun RS iffD1 RS spec)) 1),
- (rtac (monofun_pair2 RS monofunE RS spec RS spec RS mp) 2),
- (atac 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* fst and snd are monotone *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_fst = prove_goalw Cprod2.thy [monofun] "monofun(fst)"
- (fn prems =>
- [
- (strip_tac 1),
- (res_inst_tac [("p","x")] PairE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","y")] PairE 1),
- (hyp_subst_tac 1),
- (asm_simp_tac pair_ss 1),
- (etac (less_cprod4c RS conjunct1) 1)
- ]);
-
-val monofun_snd = prove_goalw Cprod2.thy [monofun] "monofun(snd)"
- (fn prems =>
- [
- (strip_tac 1),
- (res_inst_tac [("p","x")] PairE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","y")] PairE 1),
- (hyp_subst_tac 1),
- (asm_simp_tac pair_ss 1),
- (etac (less_cprod4c RS conjunct2) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the type 'a * 'b is a cpo *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_cprod = prove_goal Cprod2.thy
-" is_chain(S) ==> range(S) <<| \
-\ < lub(range(%i.fst(S(i)))),lub(range(%i.snd(S(i))))> "
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (res_inst_tac [("t","S(i)")] (surjective_pairing RS ssubst) 1),
- (rtac monofun_pair 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_fst RS ch2ch_monofun) 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_snd RS ch2ch_monofun) 1),
- (strip_tac 1),
- (res_inst_tac [("t","u")] (surjective_pairing RS ssubst) 1),
- (rtac monofun_pair 1),
- (rtac is_lub_thelub 1),
- (etac (monofun_fst RS ch2ch_monofun) 1),
- (etac (monofun_fst RS ub2ub_monofun) 1),
- (rtac is_lub_thelub 1),
- (etac (monofun_snd RS ch2ch_monofun) 1),
- (etac (monofun_snd RS ub2ub_monofun) 1)
- ]);
-
-val thelub_cprod = (lub_cprod RS thelubI);
-(* "is_chain(?S1) ==> lub(range(?S1)) = *)
-(* <lub(range(%i. fst(?S1(i)))), lub(range(%i. snd(?S1(i))))>" *)
-
-
-val cpo_cprod = prove_goal Cprod2.thy
- "is_chain(S::nat=>'a*'b)==>? x.range(S)<<| x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac exI 1),
- (etac lub_cprod 1)
- ]);
-
--- a/src/HOLCF/cprod2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-(* Title: HOLCF/cprod2.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class Instance *::(pcpo,pcpo)po
-
-*)
-
-Cprod2 = Cprod1 +
-
-(* Witness for the above arity axiom is cprod1.ML *)
-
-arities "*" :: (pcpo,pcpo)po
-
-rules
-
-(* instance of << for type ['a * 'b] *)
-
-inst_cprod_po "(op <<)::['a * 'b,'a * 'b]=>bool = less_cprod"
-
-end
-
-
-
--- a/src/HOLCF/cprod3.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,315 +0,0 @@
-(* Title: HOLCF/cprod3.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for Cprod3.thy
-*)
-
-open Cprod3;
-
-(* ------------------------------------------------------------------------ *)
-(* continuity of <_,_> , fst, snd *)
-(* ------------------------------------------------------------------------ *)
-
-val Cprod3_lemma1 = prove_goal Cprod3.thy
-"is_chain(Y::(nat=>'a)) ==>\
-\ <lub(range(Y)),(x::'b)> =\
-\ <lub(range(%i. fst(<Y(i),x>))),lub(range(%i. snd(<Y(i),x>)))>"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("f1","Pair")] (arg_cong RS cong) 1),
- (rtac lub_equal 1),
- (atac 1),
- (rtac (monofun_fst RS ch2ch_monofun) 1),
- (rtac ch2ch_fun 1),
- (rtac (monofun_pair1 RS ch2ch_monofun) 1),
- (atac 1),
- (rtac allI 1),
- (simp_tac pair_ss 1),
- (rtac sym 1),
- (simp_tac pair_ss 1),
- (rtac (lub_const RS thelubI) 1)
- ]);
-
-val contlub_pair1 = prove_goal Cprod3.thy "contlub(Pair)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (expand_fun_eq RS iffD2) 1),
- (strip_tac 1),
- (rtac (lub_fun RS thelubI RS ssubst) 1),
- (etac (monofun_pair1 RS ch2ch_monofun) 1),
- (rtac trans 1),
- (rtac (thelub_cprod RS sym) 2),
- (rtac ch2ch_fun 2),
- (etac (monofun_pair1 RS ch2ch_monofun) 2),
- (etac Cprod3_lemma1 1)
- ]);
-
-val Cprod3_lemma2 = prove_goal Cprod3.thy
-"is_chain(Y::(nat=>'a)) ==>\
-\ <(x::'b),lub(range(Y))> =\
-\ <lub(range(%i. fst(<x,Y(i)>))),lub(range(%i. snd(<x,Y(i)>)))>"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("f1","Pair")] (arg_cong RS cong) 1),
- (rtac sym 1),
- (simp_tac pair_ss 1),
- (rtac (lub_const RS thelubI) 1),
- (rtac lub_equal 1),
- (atac 1),
- (rtac (monofun_snd RS ch2ch_monofun) 1),
- (rtac (monofun_pair2 RS ch2ch_monofun) 1),
- (atac 1),
- (rtac allI 1),
- (simp_tac pair_ss 1)
- ]);
-
-val contlub_pair2 = prove_goal Cprod3.thy "contlub(Pair(x))"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_cprod RS sym) 2),
- (etac (monofun_pair2 RS ch2ch_monofun) 2),
- (etac Cprod3_lemma2 1)
- ]);
-
-val contX_pair1 = prove_goal Cprod3.thy "contX(Pair)"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_pair1 1),
- (rtac contlub_pair1 1)
- ]);
-
-val contX_pair2 = prove_goal Cprod3.thy "contX(Pair(x))"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_pair2 1),
- (rtac contlub_pair2 1)
- ]);
-
-val contlub_fst = prove_goal Cprod3.thy "contlub(fst)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (lub_cprod RS thelubI RS ssubst) 1),
- (atac 1),
- (simp_tac pair_ss 1)
- ]);
-
-val contlub_snd = prove_goal Cprod3.thy "contlub(snd)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (lub_cprod RS thelubI RS ssubst) 1),
- (atac 1),
- (simp_tac pair_ss 1)
- ]);
-
-val contX_fst = prove_goal Cprod3.thy "contX(fst)"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_fst 1),
- (rtac contlub_fst 1)
- ]);
-
-val contX_snd = prove_goal Cprod3.thy "contX(snd)"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_snd 1),
- (rtac contlub_snd 1)
- ]);
-
-(*
- --------------------------------------------------------------------------
- more lemmas for Cprod3.thy
-
- --------------------------------------------------------------------------
-*)
-
-(* ------------------------------------------------------------------------ *)
-(* convert all lemmas to the continuous versions *)
-(* ------------------------------------------------------------------------ *)
-
-val beta_cfun_cprod = prove_goalw Cprod3.thy [cpair_def]
- "(LAM x y.<x,y>)[a][b] = <a,b>"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tac 1),
- (rtac contX_pair2 1),
- (rtac contX2contX_CF1L 1),
- (rtac contX_pair1 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_pair2 1),
- (rtac refl 1)
- ]);
-
-val inject_cpair = prove_goalw Cprod3.thy [cpair_def]
- " (a#b)=(aa#ba) ==> a=aa & b=ba"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (dtac (beta_cfun_cprod RS subst) 1),
- (dtac (beta_cfun_cprod RS subst) 1),
- (etac Pair_inject 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val inst_cprod_pcpo2 = prove_goalw Cprod3.thy [cpair_def] "UU = (UU#UU)"
- (fn prems =>
- [
- (rtac sym 1),
- (rtac trans 1),
- (rtac beta_cfun_cprod 1),
- (rtac sym 1),
- (rtac inst_cprod_pcpo 1)
- ]);
-
-val defined_cpair_rev = prove_goal Cprod3.thy
- "(a#b) = UU ==> a = UU & b = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (dtac (inst_cprod_pcpo2 RS subst) 1),
- (etac inject_cpair 1)
- ]);
-
-val Exh_Cprod2 = prove_goalw Cprod3.thy [cpair_def]
- "? a b. z=(a#b) "
- (fn prems =>
- [
- (rtac PairE 1),
- (rtac exI 1),
- (rtac exI 1),
- (etac (beta_cfun_cprod RS ssubst) 1)
- ]);
-
-val cprodE = prove_goalw Cprod3.thy [cpair_def]
-"[|!!x y. [|p=(x#y) |] ==> Q|] ==> Q"
- (fn prems =>
- [
- (rtac PairE 1),
- (resolve_tac prems 1),
- (etac (beta_cfun_cprod RS ssubst) 1)
- ]);
-
-val cfst2 = prove_goalw Cprod3.thy [cfst_def,cpair_def]
- "cfst[x#y]=x"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun_cprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_fst 1),
- (simp_tac pair_ss 1)
- ]);
-
-val csnd2 = prove_goalw Cprod3.thy [csnd_def,cpair_def]
- "csnd[x#y]=y"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun_cprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_snd 1),
- (simp_tac pair_ss 1)
- ]);
-
-val surjective_pairing_Cprod2 = prove_goalw Cprod3.thy
- [cfst_def,csnd_def,cpair_def] "(cfst[p] # csnd[p]) = p"
- (fn prems =>
- [
- (rtac (beta_cfun_cprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_snd 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_fst 1),
- (rtac (surjective_pairing RS sym) 1)
- ]);
-
-
-val less_cprod5b = prove_goalw Cprod3.thy [cfst_def,csnd_def,cpair_def]
- " (p1 << p2) = (cfst[p1]<<cfst[p2] & csnd[p1]<<csnd[p2])"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_snd 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_snd 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_fst 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_fst 1),
- (rtac less_cprod3b 1)
- ]);
-
-val less_cprod5c = prove_goalw Cprod3.thy [cfst_def,csnd_def,cpair_def]
- "xa#ya << x#y ==>xa<<x & ya << y"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac less_cprod4c 1),
- (dtac (beta_cfun_cprod RS subst) 1),
- (dtac (beta_cfun_cprod RS subst) 1),
- (atac 1)
- ]);
-
-
-val lub_cprod2 = prove_goalw Cprod3.thy [cfst_def,csnd_def,cpair_def]
-"[|is_chain(S)|] ==> range(S) <<| \
-\ (lub(range(%i.cfst[S(i)])) # lub(range(%i.csnd[S(i)])))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun_cprod RS ssubst) 1),
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (rtac contX_snd 1),
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (rtac contX_fst 1),
- (rtac lub_cprod 1),
- (atac 1)
- ]);
-
-val thelub_cprod2 = (lub_cprod2 RS thelubI);
-(* "is_chain(?S1) ==> lub(range(?S1)) = *)
-(* lub(range(%i. cfst[?S1(i)]))#lub(range(%i. csnd[?S1(i)]))" *)
-
-val csplit2 = prove_goalw Cprod3.thy [csplit_def]
- "csplit[f][x#y]=f[x][y]"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tacR 1),
- (simp_tac Cfun_ss 1),
- (simp_tac (Cfun_ss addsimps [cfst2,csnd2]) 1)
- ]);
-
-val csplit3 = prove_goalw Cprod3.thy [csplit_def]
- "csplit[cpair][z]=z"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tacR 1),
- (simp_tac (Cfun_ss addsimps [surjective_pairing_Cprod2]) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* install simplifier for Cprod *)
-(* ------------------------------------------------------------------------ *)
-
-val Cprod_rews = [cfst2,csnd2,csplit2];
-
-val Cprod_ss = Cfun_ss addsimps Cprod_rews;
--- a/src/HOLCF/cprod3.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,43 +0,0 @@
-(* Title: HOLCF/cprod3.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-Class instance of * for class pcpo
-
-*)
-
-Cprod3 = Cprod2 +
-
-arities "*" :: (pcpo,pcpo)pcpo (* Witness cprod2.ML *)
-
-consts
- "@cpair" :: "'a => 'b => ('a*'b)" ("_#_" [101,100] 100)
- "cop @cpair" :: "'a -> 'b -> ('a*'b)" ("cpair")
- (* continuous pairing *)
- cfst :: "('a*'b)->'a"
- csnd :: "('a*'b)->'b"
- csplit :: "('a->'b->'c)->('a*'b)->'c"
-
-rules
-
-inst_cprod_pcpo "UU::'a*'b = <UU,UU>"
-
-cpair_def "cpair == (LAM x y.<x,y>)"
-cfst_def "cfst == (LAM p.fst(p))"
-csnd_def "csnd == (LAM p.snd(p))"
-csplit_def "csplit == (LAM f p.f[cfst[p]][csnd[p]])"
-
-end
-
-ML
-
-(* ----------------------------------------------------------------------*)
-(* parse translations for the above mixfix *)
-(* ----------------------------------------------------------------------*)
-
-val parse_translation = [("@cpair",mk_cinfixtr "@cpair")];
-
-
-
--- a/src/HOLCF/dlist.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,563 +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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss addsimps
- (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1),
- (res_inst_tac [("Q","x=UU")] classical2 1),
- (asm_simp_tac HOLCF_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps [defined_spair]) 1)
- ]);
-
-val dlist_copy = temp :: dlist_copy;
-
-val dlist_rews = dlist_copy @ dlist_rews;
-
-(* ------------------------------------------------------------------------*)
-(* Exhaustion and elimination for dlists *)
-(* ------------------------------------------------------------------------*)
-
-val Exh_dlist = prove_goalw Dlist.thy [dcons_def,dnil_def]
- "l = UU | l = dnil | (? x xl. x~=UU &xl~=UU & l = dcons[x][xl])"
- (fn prems =>
- [
- (simp_tac HOLCF_ss 1),
- (rtac (dlist_rep_iso RS subst) 1),
- (res_inst_tac [("p","dlist_rep[l]")] ssumE 1),
- (rtac disjI1 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (rtac disjI2 1),
- (rtac disjI1 1),
- (res_inst_tac [("p","x")] oneE 1),
- (contr_tac 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_rews) 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val dlistE = prove_goal 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss addsimps dlist_rews) 1),
- (dtac sym 1),
- (asm_simp_tac HOLCF_ss 1),
- (simp_tac (HOLCF_ss addsimps (prems @ dlist_rews)) 1)
- ]);
-
-
-val dlist_constrdef = [
-prover "is_dnil[UU] ~= UU" "dnil~=UU",
-prover "is_dcons[UU] ~= UU" "[|x~=UU;xl~=UU|] ==> dcons[x][xl]~=UU"
- ];
-
-
-fun prover defs thm = prove_goalw Dlist.thy defs thm
- (fn prems =>
- [
- (simp_tac (HOLCF_ss 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][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 (HOLCF_ss addsimps dlist_rews) 1),
- (res_inst_tac [("Q","x=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (res_inst_tac [("Q","xl=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_rews) 1),
- (res_inst_tac [("Q","xl=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dlist_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss 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 *)
-(* ------------------------------------------------------------------------*)
-
-val dhd2 = prove_goal 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 (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1)
- ]);
-
-val dtl2 = prove_goal 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 (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 iterate_ss 1),
- (asm_simp_tac iterate_ss 1),
- (simp_tac (HOLCF_ss 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 iterate_ss 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 iterate_ss 1),
- (simp_tac (HOLCF_ss 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 iterate_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (res_inst_tac [("Q","xl=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (res_inst_tac [("n","n")] natE 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss 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 *)
-(* ------------------------------------------------------------------------*)
-
-val dlist_coind_lemma = prove_goalw 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 (HOLCF_ss 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 (HOLCF_ss addsimps dlist_rews) 1),
- (etac disjE 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (etac exE 1),
- (etac exE 1),
- (etac exE 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (REPEAT (etac conjE 1)),
- (rtac cfun_arg_cong 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val dlist_coind = prove_goal 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 *)
-(* ------------------------------------------------------------------------*)
-
-val dlist_finite_ind = prove_goal 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 (HOLCF_ss addsimps dlist_rews) 1),
- (resolve_tac prems 1),
- (rtac allI 1),
- (res_inst_tac [("l","l")] dlistE 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (resolve_tac prems 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (resolve_tac prems 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (res_inst_tac [("Q","dlist_take(n1)[xl]=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (atac 1),
- (atac 1),
- (etac spec 1)
- ]);
-
-val dlist_all_finite_lemma1 = prove_goal Dlist.thy
-"!l.dlist_take(n)[l]=UU |dlist_take(n)[l]=l"
- (fn prems =>
- [
- (nat_ind_tac "n" 1),
- (simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (rtac allI 1),
- (res_inst_tac [("l","l")] dlistE 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (eres_inst_tac [("x","xl")] allE 1),
- (etac disjE 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1)
- ]);
-
-val dlist_all_finite_lemma2 = prove_goal Dlist.thy "? n.dlist_take(n)[l]=l"
- (fn prems =>
- [
- (res_inst_tac [("Q","l=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss 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)
- ]);
-
-val dlist_all_finite= prove_goalw Dlist.thy [dlist_finite_def] "dlist_finite(l)"
- (fn prems =>
- [
- (rtac dlist_all_finite_lemma2 1)
- ]);
-
-val dlist_ind = prove_goal 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))
- ]);
-
-
-
-
--- a/src/HOLCF/dlist.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,111 +0,0 @@
-(* Title: HOLCF/dlist.thy
-
- Author: Franz Regensburger
- ID: $ $
- Copyright 1994 Technische Universitaet Muenchen
-
-Theory for lists
-*)
-
-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 domain equation *)
-(* X = one ++ 'a ** X *)
-(* ----------------------------------------------------------------------- *)
-(* 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 \
-\ (when[sinl][sinr oo (ssplit[LAM x y. x ## f[y]])])\
-\ oo dlist_rep)"
-dlist_reach "(fix[dlist_copy])[x]=x"
-
-(* ----------------------------------------------------------------------- *)
-(* 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.\
-\ when[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
-
-
-
-
--- a/src/HOLCF/dnat.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,532 +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 (HOLCF_ss 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 *)
-(* ------------------------------------------------------------------------*)
-
-val Exh_dnat = prove_goalw Dnat.thy [dsucc_def,dzero_def]
- "n = UU | n = dzero | (? x . x~=UU & n = dsucc[x])"
- (fn prems =>
- [
- (simp_tac HOLCF_ss 1),
- (rtac (dnat_rep_iso RS subst) 1),
- (res_inst_tac [("p","dnat_rep[n]")] ssumE 1),
- (rtac disjI1 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (rtac (disjI1 RS disjI2) 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (res_inst_tac [("p","x")] oneE 1),
- (contr_tac 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (rtac (disjI2 RS disjI2) 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val dnatE = prove_goal 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss addsimps dnat_rews) 1),
- (dtac sym 1),
- (asm_simp_tac HOLCF_ss 1),
- (simp_tac (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss addsimps dnat_rews) 1),
- (res_inst_tac [("Q","n=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss 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 iterate_ss 1),
- (asm_simp_tac iterate_ss 1),
- (simp_tac (HOLCF_ss 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 iterate_ss 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 iterate_ss 1),
- (simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (res_inst_tac [("n","n")] natE 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss 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 *)
-(* ------------------------------------------------------------------------*)
-
-val dnat_coind_lemma = prove_goalw 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 (HOLCF_ss 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 (HOLCF_ss addsimps dnat_take) 1),
- (etac disjE 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_take) 1),
- (etac exE 1),
- (etac exE 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_take) 1),
- (REPEAT (etac conjE 1)),
- (rtac cfun_arg_cong 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val dnat_coind = prove_goal 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
-val dnat_ind = prove_goal 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),
- (contX_tacR 1),
- (resolve_tac prems 1),
- (simp_tac HOLCF_ss 1),
- (resolve_tac prems 1),
- (strip_tac 1),
- (res_inst_tac [("n","xa")] dnatE 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_copy) 1),
- (resolve_tac prems 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_copy) 1),
- (resolve_tac prems 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_copy) 1),
- (res_inst_tac [("Q","x[xb]=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (resolve_tac prems 1),
- (eresolve_tac prems 1),
- (etac spec 1)
- ]);
-*)
-
-val dnat_finite_ind = prove_goal 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 (HOLCF_ss addsimps dnat_rews) 1),
- (resolve_tac prems 1),
- (rtac allI 1),
- (res_inst_tac [("n","s")] dnatE 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (resolve_tac prems 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (resolve_tac prems 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (res_inst_tac [("Q","dnat_take(n1)[x]=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (atac 1),
- (etac spec 1)
- ]);
-
-val dnat_all_finite_lemma1 = prove_goal Dnat.thy
-"!s.dnat_take(n)[s]=UU |dnat_take(n)[s]=s"
- (fn prems =>
- [
- (nat_ind_tac "n" 1),
- (simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (rtac allI 1),
- (res_inst_tac [("n","s")] dnatE 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (eres_inst_tac [("x","x")] allE 1),
- (etac disjE 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1)
- ]);
-
-val dnat_all_finite_lemma2 = prove_goal Dnat.thy "? n.dnat_take(n)[s]=s"
- (fn prems =>
- [
- (res_inst_tac [("Q","s=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss 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)
- ]);
-
-
-val dnat_ind = prove_goal 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))
- ]);
-
-
-val dnat_flat = prove_goalw 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 HOLCF_ss 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps dnat_dist_less) 1),
- (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1),
- (strip_tac 1),
- (subgoal_tac "s1<<xa" 1),
- (etac allE 1),
- (dtac mp 1),
- (atac 1),
- (etac disjE 1),
- (contr_tac 1),
- (asm_simp_tac HOLCF_ss 1),
- (resolve_tac dnat_invert 1),
- (REPEAT (atac 1))
- ]);
-
-
-
-
--- a/src/HOLCF/dnat.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,101 +0,0 @@
-(* Title: HOLCF/dnat.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Theory for the domain of natural numbers
-
-*)
-
-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 domain equation *)
-(* X = one ++ X *)
-(* ----------------------------------------------------------------------- *)
-(* 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 \
-\ (when[sinl][sinr oo f]) oo dnat_rep )"
-dnat_reach "(fix[dnat_copy])[x]=x"
-
-(* ----------------------------------------------------------------------- *)
-(* 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.when[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
-
-
-
-
--- a/src/HOLCF/dnat2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /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_prover Dnat2.thy iterator_def
- "iterator = (LAM n f x. dnat_when[x][LAM m.f[iterator[m][f][x]]][n])";
-
-(* ------------------------------------------------------------------------- *)
-(* recursive properties *)
-(* ------------------------------------------------------------------------- *)
-
-val iterator1 = prove_goal Dnat2.thy "iterator[UU][f][x] = UU"
- (fn prems =>
- [
- (rtac (iterator_def2 RS ssubst) 1),
- (simp_tac (HOLCF_ss addsimps dnat_when) 1)
- ]);
-
-val iterator2 = prove_goal Dnat2.thy "iterator[dzero][f][x] = x"
- (fn prems =>
- [
- (rtac (iterator_def2 RS ssubst) 1),
- (simp_tac (HOLCF_ss addsimps dnat_when) 1)
- ]);
-
-val iterator3 = prove_goal 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 (HOLCF_ss addsimps dnat_rews) 1),
- (rtac refl 1)
- ]);
-
-val dnat2_rews =
- [iterator1, iterator2, iterator3];
-
-
-
--- a/src/HOLCF/dnat2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,32 +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"
-
-
-rules
-
-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]]
-*)
-
--- a/src/HOLCF/ex/coind.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,37 +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"
-
-rules
-
-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]]]
-
-
-*)
-
-
--- a/src/HOLCF/ex/dagstuhl.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,149 +0,0 @@
-(*
- ID: $ $
-*)
-
-
-open Dagstuhl;
-
-val YS_def2 = fix_prover Dagstuhl.thy YS_def "YS = scons[y][YS]";
-val YYS_def2 = fix_prover Dagstuhl.thy YYS_def "YYS = scons[y][scons[y][YYS]]";
-
-
-val prems = goal Dagstuhl.thy "YYS << scons[y][YYS]";
-by (rtac (YYS_def RS ssubst) 1);
-by (rtac fix_ind 1);
-by (resolve_tac adm_thms 1);
-by (contX_tacR 1);
-by (rtac minimal 1);
-by (rtac (beta_cfun RS ssubst) 1);
-by (contX_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 (HOLCF_ss addsimps stream_rews) 1);
-by (res_inst_tac [("y1","y")] (YS_def2 RS ssubst) 1);
-by (res_inst_tac [("y1","y")] (YYS_def2 RS ssubst) 1);
-by (asm_simp_tac (HOLCF_ss 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 (rtac (YYS_def RS ssubst) 1);
-by (rtac fix_least 1);
-by (rtac (beta_cfun RS ssubst) 1);
-by (contX_tacR 1);
-by (simp_tac (HOLCF_ss addsimps [YS_def2 RS sym]) 1);
-val lemma6 = result();
-
-val prems = goal Dagstuhl.thy "YS << YYS";
-by (rtac (YS_def RS ssubst) 1);
-by (rtac fix_ind 1);
-by (resolve_tac adm_thms 1);
-by (contX_tacR 1);
-by (rtac minimal 1);
-by (rtac (beta_cfun RS ssubst) 1);
-by (contX_tacR 1);
-by (res_inst_tac [("y2","y10")] (lemma5 RS sym RS ssubst) 1);
-by (etac monofun_cfun_arg 1);
-val lemma7 = result();
-
-val wir_moel = lemma6 RS (lemma7 RS antisym_less);
-
-
-(* ------------------------------------------------------------------------ *)
-(* L"osung aus Dagstuhl (F.R.) *)
-(* Wie oben, jedoch ohne Konstantendefinition f"ur YS, YYS *)
-(* ------------------------------------------------------------------------ *)
-
-val prems = goal Stream2.thy
- "(fix[LAM x. scons[y][x]]) = scons[y][fix[LAM x. scons[y][x]]]";
-by (rtac (fix_eq RS ssubst) 1);
-back();
-back();
-by (rtac (beta_cfun RS ssubst) 1);
-by (contX_tacR 1);
-by (rtac refl 1);
-val lemma1 = result();
-
-val prems = goal Stream2.thy
- "(fix[ LAM z. scons[y][scons[y][z]]]) = \
-\ scons[y][scons[y][(fix[ LAM z. scons[y][scons[y][z]]])]]";
-by (rtac (fix_eq RS ssubst) 1);
-back();
-back();
-by (rtac (beta_cfun RS ssubst) 1);
-by (contX_tacR 1);
-by (rtac refl 1);
-val lemma2 = result();
-
-val prems = goal Stream2.thy
-"fix[LAM z. scons[y][scons[y][z]]] << \
-\ scons[y][fix[LAM z. scons[y][scons[y][z]]]]";
-by (rtac fix_ind 1);
-by (resolve_tac adm_thms 1);
-by (contX_tacR 1);
-by (rtac minimal 1);
-by (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1);
-by (rtac monofun_cfun_arg 1);
-by (rtac monofun_cfun_arg 1);
-by (atac 1);
-val lemma3 = result();
-
-val prems = goal Stream2.thy
-" scons[y][fix[LAM z. scons[y][scons[y][z]]]] <<\
-\ fix[LAM z. scons[y][scons[y][z]]]";
-by (rtac (lemma2 RS ssubst) 1);
-back();
-by (rtac monofun_cfun_arg 1);
-by (rtac lemma3 1);
-val lemma4 = result();
-
-val prems = goal Stream2.thy
-" scons[y][fix[LAM z. scons[y][scons[y][z]]]] =\
-\ fix[LAM z. scons[y][scons[y][z]]]";
-by (rtac antisym_less 1);
-by (rtac lemma4 1);
-by (rtac lemma3 1);
-val lemma5 = result();
-
-val prems = goal Stream2.thy
- "(fix[LAM x. scons[y][x]]) = (fix[ LAM z. scons[y][scons[y][z]]])";
-by (rtac stream_take_lemma 1);
-by (nat_ind_tac "n" 1);
-by (simp_tac (HOLCF_ss addsimps stream_rews) 1);
-by (rtac (lemma1 RS ssubst) 1);
-by (rtac (lemma2 RS ssubst) 1);
-by (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1);
-by (rtac (lemma5 RS sym RS subst) 1);
-by (rtac refl 1);
-val wir_moel = result();
-
-
-
--- a/src/HOLCF/ex/dagstuhl.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,17 +0,0 @@
-(*
- ID: $ $
-*)
-
-Dagstuhl = Stream2 +
-
-consts
- YS :: "'a stream"
- YYS :: "'a stream"
-
-rules
-
-YS_def "YS = fix[LAM x. scons[y][x]]"
-YYS_def "YYS = fix[LAM z. scons[y][scons[y][z]]]"
-
-end
-
--- a/src/HOLCF/ex/hoare.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,540 +0,0 @@
-(* Title: HOLCF/ex/hoare.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-*)
-
-open Hoare;
-
-(* --------- pure HOLCF logic, some little lemmas ------ *)
-
-val hoare_lemma2 = prove_goal HOLCF.thy "~b=TT ==> b=FF | b=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (Exh_tr RS disjE) 1),
- (fast_tac HOL_cs 1),
- (etac disjE 1),
- (contr_tac 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val hoare_lemma3 = prove_goal HOLCF.thy
-" (!k.b1[iterate(k,g,x)]=TT) | (? k.~ b1[iterate(k,g,x)]=TT)"
- (fn prems =>
- [
- (fast_tac HOL_cs 1)
- ]);
-
-val hoare_lemma4 = prove_goal HOLCF.thy
-"(? k.~ b1[iterate(k,g,x)]=TT) ==> \
-\ ? k.b1[iterate(k,g,x)]=FF | b1[iterate(k,g,x)]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exE 1),
- (rtac exI 1),
- (rtac hoare_lemma2 1),
- (atac 1)
- ]);
-
-val hoare_lemma5 = prove_goal HOLCF.thy
-"[|(? k.~ b1[iterate(k,g,x)]=TT);\
-\ k=theleast(%n.~ b1[iterate(n,g,x)]=TT)|] ==> \
-\ b1[iterate(k,g,x)]=FF | b1[iterate(k,g,x)]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (rtac hoare_lemma2 1),
- (etac exE 1),
- (etac theleast1 1)
- ]);
-
-val hoare_lemma6 = prove_goal HOLCF.thy "b=UU ==> ~b=TT"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (resolve_tac dist_eq_tr 1)
- ]);
-
-val hoare_lemma7 = prove_goal HOLCF.thy "b=FF ==> ~b=TT"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (resolve_tac dist_eq_tr 1)
- ]);
-
-val hoare_lemma8 = prove_goal HOLCF.thy
-"[|(? k.~ b1[iterate(k,g,x)]=TT);\
-\ k=theleast(%n.~ b1[iterate(n,g,x)]=TT)|] ==> \
-\ !m. m<k --> b1[iterate(m,g,x)]=TT"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (etac exE 1),
- (strip_tac 1),
- (res_inst_tac [("p","b1[iterate(m,g,x)]")] trE 1),
- (atac 2),
- (rtac (le_less_trans RS less_anti_refl) 1),
- (atac 2),
- (rtac theleast2 1),
- (etac hoare_lemma6 1),
- (rtac (le_less_trans RS less_anti_refl) 1),
- (atac 2),
- (rtac theleast2 1),
- (etac hoare_lemma7 1)
- ]);
-
-val hoare_lemma28 = prove_goal HOLCF.thy
-"b1[y::'a]=UU::tr ==> b1[UU] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac (flat_tr RS flat_codom RS disjE) 1),
- (atac 1),
- (etac spec 1)
- ]);
-
-
-(* ----- access to definitions ----- *)
-
-val p_def2 = prove_goalw Hoare.thy [p_def]
-"p = fix[LAM f x. If b1[x] then f[g[x]] else x fi]"
- (fn prems =>
- [
- (rtac refl 1)
- ]);
-
-val q_def2 = prove_goalw Hoare.thy [q_def]
-"q = fix[LAM f x. If b1[x] orelse b2[x] then \
-\ f[g[x]] else x fi]"
- (fn prems =>
- [
- (rtac refl 1)
- ]);
-
-
-val p_def3 = prove_goal Hoare.thy
-"p[x] = If b1[x] then p[g[x]] else x fi"
- (fn prems =>
- [
- (fix_tac3 p_def 1),
- (simp_tac HOLCF_ss 1)
- ]);
-
-val q_def3 = prove_goal Hoare.thy
-"q[x] = If b1[x] orelse b2[x] then q[g[x]] else x fi"
- (fn prems =>
- [
- (fix_tac3 q_def 1),
- (simp_tac HOLCF_ss 1)
- ]);
-
-(** --------- proves about iterations of p and q ---------- **)
-
-val hoare_lemma9 = prove_goal Hoare.thy
-"(! m. m<Suc(k) --> b1[iterate(m,g,x)]=TT) -->\
-\ p[iterate(k,g,x)]=p[x]"
- (fn prems =>
- [
- (nat_ind_tac "k" 1),
- (simp_tac iterate_ss 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (res_inst_tac [("s","p[iterate(k1,g,x)]")] trans 1),
- (rtac trans 1),
- (rtac (p_def3 RS sym) 2),
- (res_inst_tac [("s","TT"),("t","b1[iterate(k1,g,x)]")] ssubst 1),
- (rtac mp 1),
- (etac spec 1),
- (simp_tac nat_ss 1),
- (simp_tac HOLCF_ss 1),
- (etac mp 1),
- (strip_tac 1),
- (rtac mp 1),
- (etac spec 1),
- (etac less_trans 1),
- (simp_tac nat_ss 1)
- ]);
-
-val hoare_lemma24 = prove_goal Hoare.thy
-"(! m. m<Suc(k) --> b1[iterate(m,g,x)]=TT) --> \
-\ q[iterate(k,g,x)]=q[x]"
- (fn prems =>
- [
- (nat_ind_tac "k" 1),
- (simp_tac iterate_ss 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (res_inst_tac [("s","q[iterate(k1,g,x)]")] trans 1),
- (rtac trans 1),
- (rtac (q_def3 RS sym) 2),
- (res_inst_tac [("s","TT"),("t","b1[iterate(k1,g,x)]")] ssubst 1),
- (rtac mp 1),
- (etac spec 1),
- (simp_tac nat_ss 1),
- (simp_tac HOLCF_ss 1),
- (etac mp 1),
- (strip_tac 1),
- (rtac mp 1),
- (etac spec 1),
- (etac less_trans 1),
- (simp_tac nat_ss 1)
- ]);
-
-(* -------- results about p for case (? k.~ b1[iterate(k,g,x)]=TT) ------- *)
-
-
-val hoare_lemma10 = (hoare_lemma8 RS (hoare_lemma9 RS mp));
-(*
-[| ? k. ~ b1[iterate(k,g,?x1)] = TT;
- Suc(?k3) = theleast(%n. ~ b1[iterate(n,g,?x1)] = TT) |] ==>
-p[iterate(?k3,g,?x1)] = p[?x1]
-*)
-
-val hoare_lemma11 = prove_goal Hoare.thy
-"(? n.b1[iterate(n,g,x)]~=TT) ==>\
-\ k=theleast(%n.b1[iterate(n,g,x)]~=TT) & b1[iterate(k,g,x)]=FF \
-\ --> p[x] = iterate(k,g,x)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("n","k")] natE 1),
- (hyp_subst_tac 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (etac conjE 1),
- (rtac trans 1),
- (rtac p_def3 1),
- (asm_simp_tac HOLCF_ss 1),
- (eres_inst_tac [("s","0"),("t","theleast(%n. b1[iterate(n, g, x)] ~= TT)")]
- subst 1),
- (simp_tac iterate_ss 1),
- (hyp_subst_tac 1),
- (strip_tac 1),
- (etac conjE 1),
- (rtac trans 1),
- (etac (hoare_lemma10 RS sym) 1),
- (atac 1),
- (rtac trans 1),
- (rtac p_def3 1),
- (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 1),
- (rtac (hoare_lemma8 RS spec RS mp) 1),
- (atac 1),
- (atac 1),
- (simp_tac nat_ss 1),
- (simp_tac HOLCF_ss 1),
- (rtac trans 1),
- (rtac p_def3 1),
- (simp_tac (HOLCF_ss addsimps [iterate_Suc RS sym]) 1),
- (eres_inst_tac [("s","FF")] ssubst 1),
- (simp_tac HOLCF_ss 1)
- ]);
-
-val hoare_lemma12 = prove_goal Hoare.thy
-"(? n.~ b1[iterate(n,g,x)]=TT) ==>\
-\ k=theleast(%n.~ b1[iterate(n,g,x)]=TT) & b1[iterate(k,g,x)]=UU \
-\ --> p[x] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("n","k")] natE 1),
- (hyp_subst_tac 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (etac conjE 1),
- (rtac trans 1),
- (rtac p_def3 1),
- (asm_simp_tac HOLCF_ss 1),
- (hyp_subst_tac 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (etac conjE 1),
- (rtac trans 1),
- (rtac (hoare_lemma10 RS sym) 1),
- (atac 1),
- (atac 1),
- (rtac trans 1),
- (rtac p_def3 1),
- (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 1),
- (rtac (hoare_lemma8 RS spec RS mp) 1),
- (atac 1),
- (atac 1),
- (simp_tac nat_ss 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac trans 1),
- (rtac p_def3 1),
- (asm_simp_tac HOLCF_ss 1)
- ]);
-
-(* -------- results about p for case (! k. b1[iterate(k,g,x)]=TT) ------- *)
-
-val fernpass_lemma = prove_goal Hoare.thy
-"(! k. b1[iterate(k,g,x)]=TT) ==> !k.p[iterate(k,g,x)] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (p_def2 RS ssubst) 1),
- (rtac fix_ind 1),
- (rtac adm_all 1),
- (rtac allI 1),
- (rtac adm_eq 1),
- (contX_tacR 1),
- (rtac allI 1),
- (rtac (strict_fapp1 RS ssubst) 1),
- (rtac refl 1),
- (simp_tac iterate_ss 1),
- (rtac allI 1),
- (res_inst_tac [("s","TT"),("t","b1[iterate(k,g,x)]")] ssubst 1),
- (etac spec 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac (iterate_Suc RS subst) 1),
- (etac spec 1)
- ]);
-
-val hoare_lemma16 = prove_goal Hoare.thy
-"(! k. b1[iterate(k,g,x)]=TT) ==> p[x] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1),
- (etac (fernpass_lemma RS spec) 1)
- ]);
-
-(* -------- results about q for case (! k. b1[iterate(k,g,x)]=TT) ------- *)
-
-val hoare_lemma17 = prove_goal Hoare.thy
-"(! k. b1[iterate(k,g,x)]=TT) ==> !k.q[iterate(k,g,x)] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (q_def2 RS ssubst) 1),
- (rtac fix_ind 1),
- (rtac adm_all 1),
- (rtac allI 1),
- (rtac adm_eq 1),
- (contX_tacR 1),
- (rtac allI 1),
- (rtac (strict_fapp1 RS ssubst) 1),
- (rtac refl 1),
- (rtac allI 1),
- (simp_tac iterate_ss 1),
- (res_inst_tac [("s","TT"),("t","b1[iterate(k,g,x)]")] ssubst 1),
- (etac spec 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac (iterate_Suc RS subst) 1),
- (etac spec 1)
- ]);
-
-val hoare_lemma18 = prove_goal Hoare.thy
-"(! k. b1[iterate(k,g,x)]=TT) ==> q[x] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1),
- (etac (hoare_lemma17 RS spec) 1)
- ]);
-
-val hoare_lemma19 = prove_goal Hoare.thy
-"(!k. (b1::'a->tr)[iterate(k,g,x)]=TT) ==> b1[UU::'a] = UU | (!y.b1[y::'a]=TT)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (flat_tr RS flat_codom) 1),
- (res_inst_tac [("t","x1")] (iterate_0 RS subst) 1),
- (etac spec 1)
- ]);
-
-val hoare_lemma20 = prove_goal Hoare.thy
-"(! y. b1[y::'a]=TT) ==> !k.q[iterate(k,g,x::'a)] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (q_def2 RS ssubst) 1),
- (rtac fix_ind 1),
- (rtac adm_all 1),
- (rtac allI 1),
- (rtac adm_eq 1),
- (contX_tacR 1),
- (rtac allI 1),
- (rtac (strict_fapp1 RS ssubst) 1),
- (rtac refl 1),
- (rtac allI 1),
- (simp_tac iterate_ss 1),
- (res_inst_tac [("s","TT"),("t","b1[iterate(k,g,x::'a)]")] ssubst 1),
- (etac spec 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac (iterate_Suc RS subst) 1),
- (etac spec 1)
- ]);
-
-val hoare_lemma21 = prove_goal Hoare.thy
-"(! y. b1[y::'a]=TT) ==> q[x::'a] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1),
- (etac (hoare_lemma20 RS spec) 1)
- ]);
-
-val hoare_lemma22 = prove_goal Hoare.thy
-"b1[UU::'a]=UU ==> q[UU::'a] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (q_def3 RS ssubst) 1),
- (asm_simp_tac HOLCF_ss 1)
- ]);
-
-(* -------- results about q for case (? k.~ b1[iterate(k,g,x)]=TT) ------- *)
-
-val hoare_lemma25 = (hoare_lemma8 RS (hoare_lemma24 RS mp) );
-(*
-[| ? k. ~ ?b1.1[iterate(k,?g1,?x1)] = TT;
- Suc(?k3) = theleast(%n. ~ ?b1.1[iterate(n,?g1,?x1)] = TT) |] ==>
-q[iterate(?k3,?g1,?x1)] = q[?x1]
-*)
-
-val hoare_lemma26 = prove_goal Hoare.thy
-"(? n.~ b1[iterate(n,g,x)]=TT) ==>\
-\ k=theleast(%n.~ b1[iterate(n,g,x)]=TT) & b1[iterate(k,g,x)]=FF \
-\ --> q[x] = q[iterate(k,g,x)]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("n","k")] natE 1),
- (hyp_subst_tac 1),
- (strip_tac 1),
- (simp_tac iterate_ss 1),
- (hyp_subst_tac 1),
- (strip_tac 1),
- (etac conjE 1),
- (rtac trans 1),
- (rtac (hoare_lemma25 RS sym) 1),
- (atac 1),
- (atac 1),
- (rtac trans 1),
- (rtac q_def3 1),
- (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 1),
- (rtac (hoare_lemma8 RS spec RS mp) 1),
- (atac 1),
- (atac 1),
- (simp_tac nat_ss 1),
- (simp_tac (HOLCF_ss addsimps [iterate_Suc]) 1)
- ]);
-
-
-val hoare_lemma27 = prove_goal Hoare.thy
-"(? n.~ b1[iterate(n,g,x)]=TT) ==>\
-\ k=theleast(%n.~ b1[iterate(n,g,x)]=TT) & b1[iterate(k,g,x)]=UU \
-\ --> q[x] = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("n","k")] natE 1),
- (hyp_subst_tac 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (etac conjE 1),
- (rtac (q_def3 RS ssubst) 1),
- (asm_simp_tac HOLCF_ss 1),
- (hyp_subst_tac 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (etac conjE 1),
- (rtac trans 1),
- (rtac (hoare_lemma25 RS sym) 1),
- (atac 1),
- (atac 1),
- (rtac trans 1),
- (rtac q_def3 1),
- (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 1),
- (rtac (hoare_lemma8 RS spec RS mp) 1),
- (atac 1),
- (atac 1),
- (simp_tac nat_ss 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac trans 1),
- (rtac q_def3 1),
- (asm_simp_tac HOLCF_ss 1)
- ]);
-
-(* ------- (! k. b1[iterate(k,g,x)]=TT) ==> q o p = q ----- *)
-
-val hoare_lemma23 = prove_goal Hoare.thy
-"(! k. b1[iterate(k,g,x)]=TT) ==> q[p[x]] = q[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (hoare_lemma16 RS ssubst) 1),
- (atac 1),
- (rtac (hoare_lemma19 RS disjE) 1),
- (atac 1),
- (rtac (hoare_lemma18 RS ssubst) 1),
- (atac 1),
- (rtac (hoare_lemma22 RS ssubst) 1),
- (atac 1),
- (rtac refl 1),
- (rtac (hoare_lemma21 RS ssubst) 1),
- (atac 1),
- (rtac (hoare_lemma21 RS ssubst) 1),
- (atac 1),
- (rtac refl 1)
- ]);
-
-(* ------------ ? k. ~ b1[iterate(k,g,x)] = TT ==> q o p = q ----- *)
-
-val hoare_lemma29 = prove_goal Hoare.thy
-"? k. ~ b1[iterate(k,g,x)] = TT ==> q[p[x]] = q[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (hoare_lemma5 RS disjE) 1),
- (atac 1),
- (rtac refl 1),
- (rtac (hoare_lemma11 RS mp RS ssubst) 1),
- (atac 1),
- (rtac conjI 1),
- (rtac refl 1),
- (atac 1),
- (rtac (hoare_lemma26 RS mp RS subst) 1),
- (atac 1),
- (rtac conjI 1),
- (rtac refl 1),
- (atac 1),
- (rtac refl 1),
- (rtac (hoare_lemma12 RS mp RS ssubst) 1),
- (atac 1),
- (rtac conjI 1),
- (rtac refl 1),
- (atac 1),
- (rtac (hoare_lemma22 RS ssubst) 1),
- (rtac (hoare_lemma28 RS ssubst) 1),
- (atac 1),
- (rtac refl 1),
- (rtac sym 1),
- (rtac (hoare_lemma27 RS mp RS ssubst) 1),
- (atac 1),
- (rtac conjI 1),
- (rtac refl 1),
- (atac 1),
- (rtac refl 1)
- ]);
-
-(* ------ the main prove q o p = q ------ *)
-
-val hoare_main = prove_goal Hoare.thy "q oo p = q"
- (fn prems =>
- [
- (rtac ext_cfun 1),
- (rtac (cfcomp2 RS ssubst) 1),
- (rtac (hoare_lemma3 RS disjE) 1),
- (etac hoare_lemma23 1),
- (etac hoare_lemma29 1)
- ]);
-
-
--- a/src/HOLCF/ex/hoare.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,43 +0,0 @@
-(* Title: HOLCF/ex/hoare.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Theory for an example by C.A.R. Hoare
-
-p x = if b1(x)
- then p(g(x))
- else x fi
-
-q x = if b1(x) orelse b2(x)
- then q (g(x))
- else x fi
-
-Prove: for all b1 b2 g .
- q o p = q
-
-In order to get a nice notation we fix the functions b1,b2 and g in the
-signature of this example
-
-*)
-
-Hoare = Tr2 +
-
-consts
- b1:: "'a -> tr"
- b2:: "'a -> tr"
- g:: "'a -> 'a"
- p :: "'a -> 'a"
- q :: "'a -> 'a"
-
-rules
-
- p_def "p == fix[LAM f. LAM x.\
-\ If b1[x] then f[g[x]] else x fi]"
-
- q_def "q == fix[LAM f. LAM x.\
-\ If b1[x] orelse b2[x] then f[g[x]] else x fi]"
-
-
-end
-
--- a/src/HOLCF/ex/loop.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,282 +0,0 @@
-(* Title: HOLCF/ex/loop.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for theory loop.thy
-*)
-
-open Loop;
-
-(* --------------------------------------------------------------------------- *)
-(* access to definitions *)
-(* --------------------------------------------------------------------------- *)
-
-val step_def2 = prove_goalw Loop.thy [step_def]
-"step[b][g][x] = If b[x] then g[x] else x fi"
- (fn prems =>
- [
- (simp_tac Cfun_ss 1)
- ]);
-
-val while_def2 = prove_goalw Loop.thy [while_def]
-"while[b][g] = fix[LAM f x. If b[x] then f[g[x]] else x fi]"
- (fn prems =>
- [
- (simp_tac Cfun_ss 1)
- ]);
-
-
-(* ------------------------------------------------------------------------- *)
-(* rekursive properties of while *)
-(* ------------------------------------------------------------------------- *)
-
-val while_unfold = prove_goal Loop.thy
-"while[b][g][x] = If b[x] then while[b][g][g[x]] else x fi"
- (fn prems =>
- [
- (fix_tac5 while_def2 1),
- (simp_tac Cfun_ss 1)
- ]);
-
-val while_unfold2 = prove_goal Loop.thy
- "!x.while[b][g][x] = while[b][g][iterate(k,step[b][g],x)]"
- (fn prems =>
- [
- (nat_ind_tac "k" 1),
- (simp_tac (HOLCF_ss addsimps [iterate_0,iterate_Suc]) 1),
- (rtac allI 1),
- (rtac trans 1),
- (rtac (while_unfold RS ssubst) 1),
- (rtac refl 2),
- (rtac (iterate_Suc2 RS ssubst) 1),
- (rtac trans 1),
- (etac spec 2),
- (rtac (step_def2 RS ssubst) 1),
- (res_inst_tac [("p","b[x]")] trE 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac (while_unfold RS ssubst) 1),
- (res_inst_tac [("s","UU"),("t","b[UU]")] ssubst 1),
- (etac (flat_tr RS flat_codom RS disjE) 1),
- (atac 1),
- (etac spec 1),
- (simp_tac HOLCF_ss 1),
- (asm_simp_tac HOLCF_ss 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac (while_unfold RS ssubst) 1),
- (asm_simp_tac HOLCF_ss 1)
- ]);
-
-val while_unfold3 = prove_goal Loop.thy
- "while[b][g][x] = while[b][g][step[b][g][x]]"
- (fn prems =>
- [
- (res_inst_tac [("s","while[b][g][iterate(Suc(0),step[b][g],x)]")] trans 1),
- (rtac (while_unfold2 RS spec) 1),
- (simp_tac iterate_ss 1)
- ]);
-
-
-(* --------------------------------------------------------------------------- *)
-(* properties of while and iterations *)
-(* --------------------------------------------------------------------------- *)
-
-val loop_lemma1 = prove_goal Loop.thy
-"[|? y.b[y]=FF; iterate(k,step[b][g],x)=UU|]==>iterate(Suc(k),step[b][g],x)=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (simp_tac iterate_ss 1),
- (rtac trans 1),
- (rtac step_def2 1),
- (asm_simp_tac HOLCF_ss 1),
- (etac exE 1),
- (etac (flat_tr RS flat_codom RS disjE) 1),
- (asm_simp_tac HOLCF_ss 1),
- (asm_simp_tac HOLCF_ss 1)
- ]);
-
-val loop_lemma2 = prove_goal Loop.thy
-"[|? y.b[y]=FF;~iterate(Suc(k),step[b][g],x)=UU |]==>\
-\~iterate(k,step[b][g],x)=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contrapos 1),
- (etac loop_lemma1 2),
- (atac 1),
- (atac 1)
- ]);
-
-val loop_lemma3 = prove_goal Loop.thy
-"[|!x. INV(x) & b[x]=TT & ~g[x]=UU --> INV(g[x]);\
-\? y.b[y]=FF; INV(x)|] ==>\
-\~iterate(k,step[b][g],x)=UU --> INV(iterate(k,step[b][g],x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (nat_ind_tac "k" 1),
- (asm_simp_tac iterate_ss 1),
- (strip_tac 1),
- (simp_tac (iterate_ss addsimps [step_def2]) 1),
- (res_inst_tac [("p","b[iterate(k1, step[b][g], x)]")] trE 1),
- (etac notE 1),
- (asm_simp_tac (HOLCF_ss addsimps [step_def2,iterate_Suc] ) 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac mp 1),
- (etac spec 1),
- (asm_simp_tac (HOLCF_ss addsimps [loop_lemma2] ) 1),
- (res_inst_tac [("s","iterate(Suc(k1), step[b][g], x)"),
- ("t","g[iterate(k1, step[b][g], x)]")] ssubst 1),
- (atac 2),
- (asm_simp_tac (HOLCF_ss addsimps [iterate_Suc,step_def2] ) 1),
- (asm_simp_tac (HOLCF_ss addsimps [loop_lemma2] ) 1)
- ]);
-
-
-val loop_lemma4 = prove_goal Loop.thy
-"!x. b[iterate(k,step[b][g],x)]=FF --> while[b][g][x]=iterate(k,step[b][g],x)"
- (fn prems =>
- [
- (nat_ind_tac "k" 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (rtac (while_unfold RS ssubst) 1),
- (asm_simp_tac HOLCF_ss 1),
- (rtac allI 1),
- (rtac (iterate_Suc2 RS ssubst) 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac while_unfold3 1),
- (asm_simp_tac HOLCF_ss 1)
- ]);
-
-val loop_lemma5 = prove_goal Loop.thy
-"!k. ~b[iterate(k,step[b][g],x)]=FF ==>\
-\ !m. while[b][g][iterate(m,step[b][g],x)]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (while_def2 RS ssubst) 1),
- (rtac fix_ind 1),
- (rtac (allI RS adm_all) 1),
- (rtac adm_eq 1),
- (contX_tacR 1),
- (simp_tac HOLCF_ss 1),
- (rtac allI 1),
- (simp_tac HOLCF_ss 1),
- (res_inst_tac [("p","b[iterate(m, step[b][g],x)]")] trE 1),
- (asm_simp_tac HOLCF_ss 1),
- (asm_simp_tac HOLCF_ss 1),
- (res_inst_tac [("s","xa[iterate(Suc(m), step[b][g], x)]")] trans 1),
- (etac spec 2),
- (rtac cfun_arg_cong 1),
- (rtac trans 1),
- (rtac (iterate_Suc RS sym) 2),
- (asm_simp_tac (HOLCF_ss addsimps [step_def2]) 1),
- (dtac spec 1),
- (contr_tac 1)
- ]);
-
-
-val loop_lemma6 = prove_goal Loop.thy
-"!k. ~b[iterate(k,step[b][g],x)]=FF ==> while[b][g][x]=UU"
- (fn prems =>
- [
- (res_inst_tac [("t","x")] (iterate_0 RS subst) 1),
- (rtac (loop_lemma5 RS spec) 1),
- (resolve_tac prems 1)
- ]);
-
-val loop_lemma7 = prove_goal Loop.thy
-"~while[b][g][x]=UU ==> ? k. b[iterate(k,step[b][g],x)]=FF"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac swap 1),
- (rtac loop_lemma6 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val loop_lemma8 = prove_goal Loop.thy
-"~while[b][g][x]=UU ==> ? y. b[y]=FF"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (dtac loop_lemma7 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-(* --------------------------------------------------------------------------- *)
-(* an invariant rule for loops *)
-(* --------------------------------------------------------------------------- *)
-
-val loop_inv2 = prove_goal Loop.thy
-"[| (!y. INV(y) & b[y]=TT & ~g[y]=UU --> INV(g[y]));\
-\ (!y. INV(y) & b[y]=FF --> Q(y));\
-\ INV(x); ~while[b][g][x]=UU |] ==> Q(while[b][g][x])"
- (fn prems =>
- [
- (res_inst_tac [("P","%k.b[iterate(k,step[b][g],x)]=FF")] exE 1),
- (rtac loop_lemma7 1),
- (resolve_tac prems 1),
- (rtac ((loop_lemma4 RS spec RS mp) RS ssubst) 1),
- (atac 1),
- (rtac (nth_elem (1,prems) RS spec RS mp) 1),
- (rtac conjI 1),
- (atac 2),
- (rtac (loop_lemma3 RS mp) 1),
- (resolve_tac prems 1),
- (rtac loop_lemma8 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (rtac classical3 1),
- (resolve_tac prems 1),
- (etac box_equals 1),
- (rtac (loop_lemma4 RS spec RS mp RS sym) 1),
- (atac 1),
- (rtac refl 1)
- ]);
-
-val loop_inv3 = prove_goal Loop.thy
-"[| !!y.[| INV(y); b[y]=TT; ~g[y]=UU|] ==> INV(g[y]);\
-\ !!y.[| INV(y); b[y]=FF|]==> Q(y);\
-\ INV(x); ~while[b][g][x]=UU |] ==> Q(while[b][g][x])"
- (fn prems =>
- [
- (rtac loop_inv2 1),
- (rtac allI 1),
- (rtac impI 1),
- (resolve_tac prems 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1),
- (rtac allI 1),
- (rtac impI 1),
- (resolve_tac prems 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1)
- ]);
-
-val loop_inv = prove_goal Loop.thy
-"[| P(x);\
-\ !!y.P(y) ==> INV(y);\
-\ !!y.[| INV(y); b[y]=TT; ~g[y]=UU|] ==> INV(g[y]);\
-\ !!y.[| INV(y); b[y]=FF|]==> Q(y);\
-\ ~while[b][g][x]=UU |] ==> Q(while[b][g][x])"
- (fn prems =>
- [
- (rtac loop_inv3 1),
- (eresolve_tac prems 1),
- (atac 1),
- (atac 1),
- (resolve_tac prems 1),
- (atac 1),
- (atac 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1)
- ]);
--- a/src/HOLCF/ex/loop.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-(* Title: HOLCF/ex/loop.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Theory for a loop primitive like while
-*)
-
-Loop = Tr2 +
-
-consts
-
- step :: "('a -> tr)->('a -> 'a)->'a->'a"
- while :: "('a -> tr)->('a -> 'a)->'a->'a"
-
-rules
-
- step_def "step == (LAM b g x. If b[x] then g[x] else x fi)"
- while_def "while == (LAM b g. fix[LAM f x.\
-\ If b[x] then f[g[x]] else x fi])"
-
-
-end
-
--- a/src/HOLCF/fix.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1153 +0,0 @@
-(* Title: HOLCF/fix.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for fix.thy
-*)
-
-open Fix;
-
-(* ------------------------------------------------------------------------ *)
-(* derive inductive properties of iterate from primitive recursion *)
-(* ------------------------------------------------------------------------ *)
-
-val iterate_0 = prove_goal Fix.thy "iterate(0,F,x) = x"
- (fn prems =>
- [
- (resolve_tac (nat_recs iterate_def) 1)
- ]);
-
-val iterate_Suc = prove_goal Fix.thy "iterate(Suc(n),F,x) = F[iterate(n,F,x)]"
- (fn prems =>
- [
- (resolve_tac (nat_recs iterate_def) 1)
- ]);
-
-val iterate_ss = Cfun_ss addsimps [iterate_0,iterate_Suc];
-
-val iterate_Suc2 = prove_goal Fix.thy "iterate(Suc(n),F,x) = iterate(n,F,F[x])"
- (fn prems =>
- [
- (nat_ind_tac "n" 1),
- (simp_tac iterate_ss 1),
- (asm_simp_tac iterate_ss 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the sequence of function itertaions is a chain *)
-(* This property is essential since monotonicity of iterate makes no sense *)
-(* ------------------------------------------------------------------------ *)
-
-val is_chain_iterate2 = prove_goalw Fix.thy [is_chain]
- " x << F[x] ==> is_chain(%i.iterate(i,F,x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (simp_tac iterate_ss 1),
- (nat_ind_tac "i" 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac iterate_ss 1),
- (etac monofun_cfun_arg 1)
- ]);
-
-
-val is_chain_iterate = prove_goal Fix.thy
- "is_chain(%i.iterate(i,F,UU))"
- (fn prems =>
- [
- (rtac is_chain_iterate2 1),
- (rtac minimal 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Kleene's fixed point theorems for continuous functions in pointed *)
-(* omega cpo's *)
-(* ------------------------------------------------------------------------ *)
-
-
-val Ifix_eq = prove_goalw Fix.thy [Ifix_def] "Ifix(F)=F[Ifix(F)]"
- (fn prems =>
- [
- (rtac (contlub_cfun_arg RS ssubst) 1),
- (rtac is_chain_iterate 1),
- (rtac antisym_less 1),
- (rtac lub_mono 1),
- (rtac is_chain_iterate 1),
- (rtac ch2ch_fappR 1),
- (rtac is_chain_iterate 1),
- (rtac allI 1),
- (rtac (iterate_Suc RS subst) 1),
- (rtac (is_chain_iterate RS is_chainE RS spec) 1),
- (rtac is_lub_thelub 1),
- (rtac ch2ch_fappR 1),
- (rtac is_chain_iterate 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (rtac (iterate_Suc RS subst) 1),
- (rtac is_ub_thelub 1),
- (rtac is_chain_iterate 1)
- ]);
-
-
-val Ifix_least = prove_goalw Fix.thy [Ifix_def] "F[x]=x ==> Ifix(F) << x"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lub_thelub 1),
- (rtac is_chain_iterate 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (nat_ind_tac "i" 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac iterate_ss 1),
- (res_inst_tac [("t","x")] subst 1),
- (atac 1),
- (etac monofun_cfun_arg 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* monotonicity and continuity of iterate *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_iterate = prove_goalw Fix.thy [monofun] "monofun(iterate(i))"
- (fn prems =>
- [
- (strip_tac 1),
- (nat_ind_tac "i" 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac iterate_ss 1),
- (rtac (less_fun RS iffD2) 1),
- (rtac allI 1),
- (rtac monofun_cfun 1),
- (atac 1),
- (rtac (less_fun RS iffD1 RS spec) 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the following lemma uses contlub_cfun which itself is based on a *)
-(* diagonalisation lemma for continuous functions with two arguments. *)
-(* In this special case it is the application function fapp *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_iterate = prove_goalw Fix.thy [contlub] "contlub(iterate(i))"
- (fn prems =>
- [
- (strip_tac 1),
- (nat_ind_tac "i" 1),
- (asm_simp_tac iterate_ss 1),
- (rtac (lub_const RS thelubI RS sym) 1),
- (asm_simp_tac iterate_ss 1),
- (rtac ext 1),
- (rtac (thelub_fun RS ssubst) 1),
- (rtac is_chainI 1),
- (rtac allI 1),
- (rtac (less_fun RS iffD2) 1),
- (rtac allI 1),
- (rtac (is_chainE RS spec) 1),
- (rtac (monofun_fapp1 RS ch2ch_MF2LR) 1),
- (rtac allI 1),
- (rtac monofun_fapp2 1),
- (atac 1),
- (rtac ch2ch_fun 1),
- (rtac (monofun_iterate RS ch2ch_monofun) 1),
- (atac 1),
- (rtac (thelub_fun RS ssubst) 1),
- (rtac (monofun_iterate RS ch2ch_monofun) 1),
- (atac 1),
- (rtac contlub_cfun 1),
- (atac 1),
- (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1)
- ]);
-
-
-val contX_iterate = prove_goal Fix.thy "contX(iterate(i))"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_iterate 1),
- (rtac contlub_iterate 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* a lemma about continuity of iterate in its third argument *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_iterate2 = prove_goal Fix.thy "monofun(iterate(n,F))"
- (fn prems =>
- [
- (rtac monofunI 1),
- (strip_tac 1),
- (nat_ind_tac "n" 1),
- (asm_simp_tac iterate_ss 1),
- (asm_simp_tac iterate_ss 1),
- (etac monofun_cfun_arg 1)
- ]);
-
-val contlub_iterate2 = prove_goal Fix.thy "contlub(iterate(n,F))"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (nat_ind_tac "n" 1),
- (simp_tac iterate_ss 1),
- (simp_tac iterate_ss 1),
- (res_inst_tac [("t","iterate(n1, F, lub(range(%u. Y(u))))"),
- ("s","lub(range(%i. iterate(n1, F, Y(i))))")] ssubst 1),
- (atac 1),
- (rtac contlub_cfun_arg 1),
- (etac (monofun_iterate2 RS ch2ch_monofun) 1)
- ]);
-
-val contX_iterate2 = prove_goal Fix.thy "contX(iterate(n,F))"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_iterate2 1),
- (rtac contlub_iterate2 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* monotonicity and continuity of Ifix *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_Ifix = prove_goalw Fix.thy [monofun,Ifix_def] "monofun(Ifix)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac lub_mono 1),
- (rtac is_chain_iterate 1),
- (rtac is_chain_iterate 1),
- (rtac allI 1),
- (rtac (less_fun RS iffD1 RS spec) 1),
- (etac (monofun_iterate RS monofunE RS spec RS spec RS mp) 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* since iterate is not monotone in its first argument, special lemmas must *)
-(* be derived for lubs in this argument *)
-(* ------------------------------------------------------------------------ *)
-
-val is_chain_iterate_lub = prove_goal Fix.thy
-"is_chain(Y) ==> is_chain(%i. lub(range(%ia. iterate(ia,Y(i),UU))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_chainI 1),
- (strip_tac 1),
- (rtac lub_mono 1),
- (rtac is_chain_iterate 1),
- (rtac is_chain_iterate 1),
- (strip_tac 1),
- (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun RS is_chainE
- RS spec) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* this exchange lemma is analog to the one for monotone functions *)
-(* observe that monotonicity is not really needed. The propagation of *)
-(* chains is the essential argument which is usually derived from monot. *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_Ifix_lemma1 = prove_goal Fix.thy
-"is_chain(Y) ==> iterate(n,lub(range(Y)),y) = lub(range(%i. iterate(n,Y(i),y)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (thelub_fun RS subst) 1),
- (rtac (monofun_iterate RS ch2ch_monofun) 1),
- (atac 1),
- (rtac fun_cong 1),
- (rtac (contlub_iterate RS contlubE RS spec RS mp RS ssubst) 1),
- (atac 1),
- (rtac refl 1)
- ]);
-
-
-val ex_lub_iterate = prove_goal Fix.thy "is_chain(Y) ==>\
-\ lub(range(%i. lub(range(%ia. iterate(i,Y(ia),UU))))) =\
-\ lub(range(%i. lub(range(%ia. iterate(ia,Y(i),UU)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac antisym_less 1),
- (rtac is_lub_thelub 1),
- (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1),
- (atac 1),
- (rtac is_chain_iterate 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (rtac lub_mono 1),
- (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1),
- (etac is_chain_iterate_lub 1),
- (strip_tac 1),
- (rtac is_ub_thelub 1),
- (rtac is_chain_iterate 1),
- (rtac is_lub_thelub 1),
- (etac is_chain_iterate_lub 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (rtac lub_mono 1),
- (rtac is_chain_iterate 1),
- (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1),
- (atac 1),
- (rtac is_chain_iterate 1),
- (strip_tac 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1)
- ]);
-
-
-val contlub_Ifix = prove_goalw Fix.thy [contlub,Ifix_def] "contlub(Ifix)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac (contlub_Ifix_lemma1 RS ext RS ssubst) 1),
- (atac 1),
- (etac ex_lub_iterate 1)
- ]);
-
-
-val contX_Ifix = prove_goal Fix.thy "contX(Ifix)"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Ifix 1),
- (rtac contlub_Ifix 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* propagate properties of Ifix to its continuous counterpart *)
-(* ------------------------------------------------------------------------ *)
-
-val fix_eq = prove_goalw Fix.thy [fix_def] "fix[F]=F[fix[F]]"
- (fn prems =>
- [
- (asm_simp_tac (Cfun_ss addsimps [contX_Ifix]) 1),
- (rtac Ifix_eq 1)
- ]);
-
-val fix_least = prove_goalw Fix.thy [fix_def] "F[x]=x ==> fix[F] << x"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac (Cfun_ss addsimps [contX_Ifix]) 1),
- (etac Ifix_least 1)
- ]);
-
-
-val fix_eq2 = prove_goal Fix.thy "f == fix[F] ==> f = F[f]"
- (fn prems =>
- [
- (rewrite_goals_tac prems),
- (rtac fix_eq 1)
- ]);
-
-val fix_eq3 = prove_goal Fix.thy "f == fix[F] ==> f[x] = F[f][x]"
- (fn prems =>
- [
- (rtac trans 1),
- (rtac ((hd prems) RS fix_eq2 RS cfun_fun_cong) 1),
- (rtac refl 1)
- ]);
-
-fun fix_tac3 thm i = ((rtac trans i) THEN (rtac (thm RS fix_eq3) i));
-
-val fix_eq4 = prove_goal Fix.thy "f = fix[F] ==> f = F[f]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (rtac fix_eq 1)
- ]);
-
-val fix_eq5 = prove_goal Fix.thy "f = fix[F] ==> f[x] = F[f][x]"
- (fn prems =>
- [
- (rtac trans 1),
- (rtac ((hd prems) RS fix_eq4 RS cfun_fun_cong) 1),
- (rtac refl 1)
- ]);
-
-fun fix_tac5 thm i = ((rtac trans i) THEN (rtac (thm RS fix_eq5) i));
-
-fun fix_prover thy fixdef thm = prove_goal thy thm
- (fn prems =>
- [
- (rtac trans 1),
- (rtac (fixdef RS fix_eq4) 1),
- (rtac trans 1),
- (rtac beta_cfun 1),
- (contX_tacR 1),
- (rtac refl 1)
- ]);
-
-(* ------------------------------------------------------------------------
-
-given the definition
-
-smap_def
- "smap = fix[LAM h f s. stream_when[LAM x l.scons[f[x]][h[f][l]]][s]]"
-
-use fix_prover for
-
-val smap_def2 = fix_prover Stream2.thy smap_def
- "smap = (LAM f s. stream_when[LAM x l.scons[f[x]][smap[f][l]]][s])";
-
- ------------------------------------------------------------------------ *)
-
-(* ------------------------------------------------------------------------ *)
-(* better access to definitions *)
-(* ------------------------------------------------------------------------ *)
-
-
-val Ifix_def2 = prove_goal Fix.thy "Ifix=(%x. lub(range(%i. iterate(i,x,UU))))"
- (fn prems =>
- [
- (rtac ext 1),
- (rewrite_goals_tac [Ifix_def]),
- (rtac refl 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* direct connection between fix and iteration without Ifix *)
-(* ------------------------------------------------------------------------ *)
-
-val fix_def2 = prove_goalw Fix.thy [fix_def]
- "fix[F] = lub(range(%i. iterate(i,F,UU)))"
- (fn prems =>
- [
- (fold_goals_tac [Ifix_def]),
- (asm_simp_tac (Cfun_ss addsimps [contX_Ifix]) 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Lemmas about admissibility and fixed point induction *)
-(* ------------------------------------------------------------------------ *)
-
-(* ------------------------------------------------------------------------ *)
-(* access to definitions *)
-(* ------------------------------------------------------------------------ *)
-
-val adm_def2 = prove_goalw Fix.thy [adm_def]
- "adm(P) = (!Y. is_chain(Y) --> (!i.P(Y(i))) --> P(lub(range(Y))))"
- (fn prems =>
- [
- (rtac refl 1)
- ]);
-
-val admw_def2 = prove_goalw Fix.thy [admw_def]
- "admw(P) = (!F.((!n.P(iterate(n,F,UU)))-->\
-\ P(lub(range(%i.iterate(i,F,UU))))))"
- (fn prems =>
- [
- (rtac refl 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* an admissible formula is also weak admissible *)
-(* ------------------------------------------------------------------------ *)
-
-val adm_impl_admw = prove_goalw Fix.thy [admw_def] "adm(P)==>admw(P)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1),
- (atac 1),
- (rtac is_chain_iterate 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* fixed point induction *)
-(* ------------------------------------------------------------------------ *)
-
-val fix_ind = prove_goal Fix.thy
-"[| adm(P);P(UU);!!x. P(x) ==> P(F[x])|] ==> P(fix[F])"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (fix_def2 RS ssubst) 1),
- (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1),
- (atac 1),
- (rtac is_chain_iterate 1),
- (rtac allI 1),
- (nat_ind_tac "i" 1),
- (rtac (iterate_0 RS ssubst) 1),
- (atac 1),
- (rtac (iterate_Suc RS ssubst) 1),
- (resolve_tac prems 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* computational induction for weak admissible formulae *)
-(* ------------------------------------------------------------------------ *)
-
-val wfix_ind = prove_goal Fix.thy
-"[| admw(P); !n. P(iterate(n,F,UU))|] ==> P(fix[F])"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (fix_def2 RS ssubst) 1),
- (rtac (admw_def2 RS iffD1 RS spec RS mp) 1),
- (atac 1),
- (rtac allI 1),
- (etac spec 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* for chain-finite (easy) types every formula is admissible *)
-(* ------------------------------------------------------------------------ *)
-
-val adm_max_in_chain = prove_goalw Fix.thy [adm_def]
-"!Y. is_chain(Y::nat=>'a) --> (? n.max_in_chain(n,Y)) ==> adm(P::'a=>bool)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac exE 1),
- (rtac mp 1),
- (etac spec 1),
- (atac 1),
- (rtac (lub_finch1 RS thelubI RS ssubst) 1),
- (atac 1),
- (atac 1),
- (etac spec 1)
- ]);
-
-
-val adm_chain_finite = prove_goalw Fix.thy [chain_finite_def]
- "chain_finite(x::'a) ==> adm(P::'a=>bool)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac adm_max_in_chain 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* flat types are chain_finite *)
-(* ------------------------------------------------------------------------ *)
-
-val flat_imp_chain_finite = prove_goalw Fix.thy [flat_def,chain_finite_def]
- "flat(x::'a)==>chain_finite(x::'a)"
- (fn prems =>
- [
- (rewrite_goals_tac [max_in_chain_def]),
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (res_inst_tac [("Q","!i.Y(i)=UU")] classical2 1),
- (res_inst_tac [("x","0")] exI 1),
- (strip_tac 1),
- (rtac trans 1),
- (etac spec 1),
- (rtac sym 1),
- (etac spec 1),
- (rtac (chain_mono2 RS exE) 1),
- (fast_tac HOL_cs 1),
- (atac 1),
- (res_inst_tac [("x","Suc(x)")] exI 1),
- (strip_tac 1),
- (rtac disjE 1),
- (atac 3),
- (rtac mp 1),
- (dtac spec 1),
- (etac spec 1),
- (etac (le_imp_less_or_eq RS disjE) 1),
- (etac (chain_mono RS mp) 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac refl_less 1),
- (res_inst_tac [("P","Y(Suc(x)) = UU")] notE 1),
- (atac 2),
- (rtac mp 1),
- (etac spec 1),
- (asm_simp_tac nat_ss 1)
- ]);
-
-
-val adm_flat = flat_imp_chain_finite RS adm_chain_finite;
-(* flat(?x::?'a) ==> adm(?P::?'a => bool) *)
-
-val flat_void = prove_goalw Fix.thy [flat_def] "flat(UU::void)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac disjI1 1),
- (rtac unique_void2 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* continuous isomorphisms are strict *)
-(* a prove for embedding projection pairs is similar *)
-(* ------------------------------------------------------------------------ *)
-
-val iso_strict = prove_goal Fix.thy
-"!!f g.[|!y.f[g[y]]=(y::'b) ; !x.g[f[x]]=(x::'a) |] \
-\ ==> f[UU]=UU & g[UU]=UU"
- (fn prems =>
- [
- (rtac conjI 1),
- (rtac UU_I 1),
- (res_inst_tac [("s","f[g[UU::'b]]"),("t","UU::'b")] subst 1),
- (etac spec 1),
- (rtac (minimal RS monofun_cfun_arg) 1),
- (rtac UU_I 1),
- (res_inst_tac [("s","g[f[UU::'a]]"),("t","UU::'a")] subst 1),
- (etac spec 1),
- (rtac (minimal RS monofun_cfun_arg) 1)
- ]);
-
-
-val isorep_defined = prove_goal Fix.thy
- "[|!x.rep[abs[x]]=x;!y.abs[rep[y]]=y;z~=UU|] ==> rep[z]~=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac swap 1),
- (dtac notnotD 1),
- (dres_inst_tac [("f","abs")] cfun_arg_cong 1),
- (etac box_equals 1),
- (fast_tac HOL_cs 1),
- (etac (iso_strict RS conjunct1) 1),
- (atac 1)
- ]);
-
-val isoabs_defined = prove_goal Fix.thy
- "[|!x.rep[abs[x]]=x;!y.abs[rep[y]]=y;z~=UU|] ==> abs[z]~=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac swap 1),
- (dtac notnotD 1),
- (dres_inst_tac [("f","rep")] cfun_arg_cong 1),
- (etac box_equals 1),
- (fast_tac HOL_cs 1),
- (etac (iso_strict RS conjunct2) 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* propagation of flatness and chainfiniteness by continuous isomorphisms *)
-(* ------------------------------------------------------------------------ *)
-
-val chfin2chfin = prove_goalw Fix.thy [chain_finite_def]
-"!!f g.[|chain_finite(x::'a); !y.f[g[y]]=(y::'b) ; !x.g[f[x]]=(x::'a) |] \
-\ ==> chain_finite(y::'b)"
- (fn prems =>
- [
- (rewrite_goals_tac [max_in_chain_def]),
- (strip_tac 1),
- (rtac exE 1),
- (res_inst_tac [("P","is_chain(%i.g[Y(i)])")] mp 1),
- (etac spec 1),
- (etac ch2ch_fappR 1),
- (rtac exI 1),
- (strip_tac 1),
- (res_inst_tac [("s","f[g[Y(x)]]"),("t","Y(x)")] subst 1),
- (etac spec 1),
- (res_inst_tac [("s","f[g[Y(j)]]"),("t","Y(j)")] subst 1),
- (etac spec 1),
- (rtac cfun_arg_cong 1),
- (rtac mp 1),
- (etac spec 1),
- (atac 1)
- ]);
-
-val flat2flat = prove_goalw Fix.thy [flat_def]
-"!!f g.[|flat(x::'a); !y.f[g[y]]=(y::'b) ; !x.g[f[x]]=(x::'a) |] \
-\ ==> flat(y::'b)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac disjE 1),
- (res_inst_tac [("P","g[x]<<g[y]")] mp 1),
- (etac monofun_cfun_arg 2),
- (dtac spec 1),
- (etac spec 1),
- (rtac disjI1 1),
- (rtac trans 1),
- (res_inst_tac [("s","f[g[x]]"),("t","x")] subst 1),
- (etac spec 1),
- (etac cfun_arg_cong 1),
- (rtac (iso_strict RS conjunct1) 1),
- (atac 1),
- (atac 1),
- (rtac disjI2 1),
- (res_inst_tac [("s","f[g[x]]"),("t","x")] subst 1),
- (etac spec 1),
- (res_inst_tac [("s","f[g[y]]"),("t","y")] subst 1),
- (etac spec 1),
- (etac cfun_arg_cong 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* admissibility of special formulae and propagation *)
-(* ------------------------------------------------------------------------ *)
-
-val adm_less = prove_goalw Fix.thy [adm_def]
- "[|contX(u);contX(v)|]==> adm(%x.u(x)<<v(x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (etac (contX2contlub RS contlubE RS spec RS mp RS ssubst) 1),
- (atac 1),
- (etac (contX2contlub RS contlubE RS spec RS mp RS ssubst) 1),
- (atac 1),
- (rtac lub_mono 1),
- (cut_facts_tac prems 1),
- (etac (contX2mono RS ch2ch_monofun) 1),
- (atac 1),
- (cut_facts_tac prems 1),
- (etac (contX2mono RS ch2ch_monofun) 1),
- (atac 1),
- (atac 1)
- ]);
-
-val adm_conj = prove_goal Fix.thy
- "[| adm(P); adm(Q) |] ==> adm(%x.P(x)&Q(x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (adm_def2 RS iffD2) 1),
- (strip_tac 1),
- (rtac conjI 1),
- (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1),
- (atac 1),
- (atac 1),
- (fast_tac HOL_cs 1),
- (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1),
- (atac 1),
- (atac 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val adm_cong = prove_goal Fix.thy
- "(!x. P(x) = Q(x)) ==> adm(P)=adm(Q)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("s","P"),("t","Q")] subst 1),
- (rtac refl 2),
- (rtac ext 1),
- (etac spec 1)
- ]);
-
-val adm_not_free = prove_goalw Fix.thy [adm_def] "adm(%x.t)"
- (fn prems =>
- [
- (fast_tac HOL_cs 1)
- ]);
-
-val adm_not_less = prove_goalw Fix.thy [adm_def]
- "contX(t) ==> adm(%x.~ t(x) << u)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac contrapos 1),
- (etac spec 1),
- (rtac trans_less 1),
- (atac 2),
- (etac (contX2mono RS monofun_fun_arg) 1),
- (rtac is_ub_thelub 1),
- (atac 1)
- ]);
-
-val adm_all = prove_goal Fix.thy
- " !y.adm(P(y)) ==> adm(%x.!y.P(y,x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (adm_def2 RS iffD2) 1),
- (strip_tac 1),
- (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1),
- (etac spec 1),
- (atac 1),
- (rtac allI 1),
- (dtac spec 1),
- (etac spec 1)
- ]);
-
-val adm_subst = prove_goal Fix.thy
- "[|contX(t); adm(P)|] ==> adm(%x.P(t(x)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (adm_def2 RS iffD2) 1),
- (strip_tac 1),
- (rtac (contX2contlub RS contlubE RS spec RS mp RS ssubst) 1),
- (atac 1),
- (atac 1),
- (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1),
- (atac 1),
- (rtac (contX2mono RS ch2ch_monofun) 1),
- (atac 1),
- (atac 1),
- (atac 1)
- ]);
-
-val adm_UU_not_less = prove_goal Fix.thy "adm(%x.~ UU << t(x))"
- (fn prems =>
- [
- (res_inst_tac [("P2","%x.False")] (adm_cong RS iffD1) 1),
- (asm_simp_tac Cfun_ss 1),
- (rtac adm_not_free 1)
- ]);
-
-val adm_not_UU = prove_goalw Fix.thy [adm_def]
- "contX(t)==> adm(%x.~ t(x) = UU)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac contrapos 1),
- (etac spec 1),
- (rtac (chain_UU_I RS spec) 1),
- (rtac (contX2mono RS ch2ch_monofun) 1),
- (atac 1),
- (atac 1),
- (rtac (contX2contlub RS contlubE RS spec RS mp RS subst) 1),
- (atac 1),
- (atac 1),
- (atac 1)
- ]);
-
-val adm_eq = prove_goal Fix.thy
- "[|contX(u);contX(v)|]==> adm(%x.u(x)= v(x))"
- (fn prems =>
- [
- (rtac (adm_cong RS iffD1) 1),
- (rtac allI 1),
- (rtac iffI 1),
- (rtac antisym_less 1),
- (rtac antisym_less_inverse 3),
- (atac 3),
- (etac conjunct1 1),
- (etac conjunct2 1),
- (rtac adm_conj 1),
- (rtac adm_less 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (rtac adm_less 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* admissibility for disjunction is hard to prove. It takes 10 Lemmas *)
-(* ------------------------------------------------------------------------ *)
-
-val adm_disj_lemma1 = prove_goal Pcpo.thy
-"[| is_chain(Y); !n.P(Y(n))|Q(Y(n))|]\
-\ ==> (? i.!j. i<j --> Q(Y(j))) | (!i.? j.i<j & P(Y(j)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val adm_disj_lemma2 = prove_goal Fix.thy
-"[| adm(Q); ? X.is_chain(X) & (!n.Q(X(n))) &\
-\ lub(range(Y))=lub(range(X))|] ==> Q(lub(range(Y)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exE 1),
- (etac conjE 1),
- (etac conjE 1),
- (res_inst_tac [("s","lub(range(X))"),("t","lub(range(Y))")] ssubst 1),
- (atac 1),
- (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1),
- (atac 1),
- (atac 1),
- (atac 1)
- ]);
-
-val adm_disj_lemma3 = prove_goal Fix.thy
-"[| is_chain(Y); ! j. i < j --> Q(Y(j)) |] ==>\
-\ is_chain(%m. if(m < Suc(i),Y(Suc(i)),Y(m)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_chainI 1),
- (rtac allI 1),
- (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1),
- (res_inst_tac [("s","False"),("t","ia < Suc(i)")] ssubst 1),
- (rtac iffI 1),
- (etac FalseE 2),
- (rtac notE 1),
- (rtac (not_less_eq RS iffD2) 1),
- (atac 1),
- (atac 1),
- (res_inst_tac [("s","False"),("t","Suc(ia) < Suc(i)")] ssubst 1),
- (asm_simp_tac nat_ss 1),
- (rtac iffI 1),
- (etac FalseE 2),
- (rtac notE 1),
- (etac less_not_sym 1),
- (atac 1),
- (asm_simp_tac Cfun_ss 1),
- (etac (is_chainE RS spec) 1),
- (hyp_subst_tac 1),
- (asm_simp_tac nat_ss 1),
- (rtac refl_less 1),
- (asm_simp_tac nat_ss 1),
- (rtac refl_less 1)
- ]);
-
-val adm_disj_lemma4 = prove_goal Fix.thy
-"[| ! j. i < j --> Q(Y(j)) |] ==>\
-\ ! n. Q(if(n < Suc(i),Y(Suc(i)),Y(n)))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac allI 1),
- (res_inst_tac [("m","n"),("n","Suc(i)")] nat_less_cases 1),
- (res_inst_tac[("s","Y(Suc(i))"),("t","if(n<Suc(i),Y(Suc(i)),Y(n))")]
- ssubst 1),
- (asm_simp_tac nat_ss 1),
- (etac allE 1),
- (rtac mp 1),
- (atac 1),
- (asm_simp_tac nat_ss 1),
- (res_inst_tac[("s","Y(n)"),("t","if(n<Suc(i),Y(Suc(i)),Y(n))")]
- ssubst 1),
- (asm_simp_tac nat_ss 1),
- (hyp_subst_tac 1),
- (dtac spec 1),
- (rtac mp 1),
- (atac 1),
- (asm_simp_tac nat_ss 1),
- (res_inst_tac [("s","Y(n)"),("t","if(n < Suc(i),Y(Suc(i)),Y(n))")]
- ssubst 1),
- (res_inst_tac [("s","False"),("t","n < Suc(i)")] ssubst 1),
- (rtac iffI 1),
- (etac FalseE 2),
- (rtac notE 1),
- (etac less_not_sym 1),
- (atac 1),
- (asm_simp_tac nat_ss 1),
- (dtac spec 1),
- (rtac mp 1),
- (atac 1),
- (etac Suc_lessD 1)
- ]);
-
-val adm_disj_lemma5 = prove_goal Fix.thy
-"[| is_chain(Y::nat=>'a); ! j. i < j --> Q(Y(j)) |] ==>\
-\ lub(range(Y)) = lub(range(%m. if(m < Suc(i),Y(Suc(i)),Y(m))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac lub_equal2 1),
- (atac 2),
- (rtac adm_disj_lemma3 2),
- (atac 2),
- (atac 2),
- (res_inst_tac [("x","i")] exI 1),
- (strip_tac 1),
- (res_inst_tac [("s","False"),("t","ia < Suc(i)")] ssubst 1),
- (rtac iffI 1),
- (etac FalseE 2),
- (rtac notE 1),
- (rtac (not_less_eq RS iffD2) 1),
- (atac 1),
- (atac 1),
- (rtac (if_False RS ssubst) 1),
- (rtac refl 1)
- ]);
-
-val adm_disj_lemma6 = prove_goal Fix.thy
-"[| is_chain(Y::nat=>'a); ? i. ! j. i < j --> Q(Y(j)) |] ==>\
-\ ? X. is_chain(X) & (! n. Q(X(n))) & lub(range(Y)) = lub(range(X))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exE 1),
- (res_inst_tac [("x","%m.if(m< Suc(i),Y(Suc(i)),Y(m))")] exI 1),
- (rtac conjI 1),
- (rtac adm_disj_lemma3 1),
- (atac 1),
- (atac 1),
- (rtac conjI 1),
- (rtac adm_disj_lemma4 1),
- (atac 1),
- (rtac adm_disj_lemma5 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-val adm_disj_lemma7 = prove_goal Fix.thy
-"[| is_chain(Y::nat=>'a); ! i. ? j. i < j & P(Y(j)) |] ==>\
-\ is_chain(%m. Y(theleast(%j. m<j & P(Y(j)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_chainI 1),
- (rtac allI 1),
- (rtac chain_mono3 1),
- (atac 1),
- (rtac theleast2 1),
- (rtac conjI 1),
- (rtac Suc_lessD 1),
- (etac allE 1),
- (etac exE 1),
- (rtac (theleast1 RS conjunct1) 1),
- (atac 1),
- (etac allE 1),
- (etac exE 1),
- (rtac (theleast1 RS conjunct2) 1),
- (atac 1)
- ]);
-
-val adm_disj_lemma8 = prove_goal Fix.thy
-"[| ! i. ? j. i < j & P(Y(j)) |] ==> ! m. P(Y(theleast(%j. m<j & P(Y(j)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (etac allE 1),
- (etac exE 1),
- (etac (theleast1 RS conjunct2) 1)
- ]);
-
-val adm_disj_lemma9 = prove_goal Fix.thy
-"[| is_chain(Y::nat=>'a); ! i. ? j. i < j & P(Y(j)) |] ==>\
-\ lub(range(Y)) = lub(range(%m. Y(theleast(%j. m<j & P(Y(j))))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac antisym_less 1),
- (rtac lub_mono 1),
- (atac 1),
- (rtac adm_disj_lemma7 1),
- (atac 1),
- (atac 1),
- (strip_tac 1),
- (rtac (chain_mono RS mp) 1),
- (atac 1),
- (etac allE 1),
- (etac exE 1),
- (rtac (theleast1 RS conjunct1) 1),
- (atac 1),
- (rtac lub_mono3 1),
- (rtac adm_disj_lemma7 1),
- (atac 1),
- (atac 1),
- (atac 1),
- (strip_tac 1),
- (rtac exI 1),
- (rtac (chain_mono RS mp) 1),
- (atac 1),
- (rtac lessI 1)
- ]);
-
-val adm_disj_lemma10 = prove_goal Fix.thy
-"[| is_chain(Y::nat=>'a); ! i. ? j. i < j & P(Y(j)) |] ==>\
-\ ? X. is_chain(X) & (! n. P(X(n))) & lub(range(Y)) = lub(range(X))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("x","%m. Y(theleast(%j. m<j & P(Y(j))))")] exI 1),
- (rtac conjI 1),
- (rtac adm_disj_lemma7 1),
- (atac 1),
- (atac 1),
- (rtac conjI 1),
- (rtac adm_disj_lemma8 1),
- (atac 1),
- (rtac adm_disj_lemma9 1),
- (atac 1),
- (atac 1)
- ]);
-
-val adm_disj = prove_goal Fix.thy
- "[| adm(P); adm(Q) |] ==> adm(%x.P(x)|Q(x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (adm_def2 RS iffD2) 1),
- (strip_tac 1),
- (rtac (adm_disj_lemma1 RS disjE) 1),
- (atac 1),
- (atac 1),
- (rtac disjI2 1),
- (rtac adm_disj_lemma2 1),
- (atac 1),
- (rtac adm_disj_lemma6 1),
- (atac 1),
- (atac 1),
- (rtac disjI1 1),
- (rtac adm_disj_lemma2 1),
- (atac 1),
- (rtac adm_disj_lemma10 1),
- (atac 1),
- (atac 1)
- ]);
-
-val adm_impl = prove_goal Fix.thy
- "[| adm(%x.~P(x)); adm(Q) |] ==> adm(%x.P(x)-->Q(x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("P2","%x.~P(x)|Q(x)")] (adm_cong RS iffD1) 1),
- (fast_tac HOL_cs 1),
- (rtac adm_disj 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-val adm_all2 = (allI RS adm_all);
-
-val adm_thms = [adm_impl,adm_disj,adm_eq,adm_not_UU,adm_UU_not_less,
- adm_all2,adm_not_less,adm_not_free,adm_conj,adm_less
- ];
-
-(* ------------------------------------------------------------------------- *)
-(* a result about functions with flat codomain *)
-(* ------------------------------------------------------------------------- *)
-
-val flat_codom = prove_goalw Fix.thy [flat_def]
-"[|flat(y::'b);f[x::'a]=(c::'b)|] ==> f[UU::'a]=UU::'b | (!z.f[z::'a]=c)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("Q","f[x::'a]=UU::'b")] classical2 1),
- (rtac disjI1 1),
- (rtac UU_I 1),
- (res_inst_tac [("s","f[x]"),("t","UU::'b")] subst 1),
- (atac 1),
- (rtac (minimal RS monofun_cfun_arg) 1),
- (res_inst_tac [("Q","f[UU::'a]=UU::'b")] classical2 1),
- (etac disjI1 1),
- (rtac disjI2 1),
- (rtac allI 1),
- (res_inst_tac [("s","f[x]"),("t","c")] subst 1),
- (atac 1),
- (res_inst_tac [("a","f[UU::'a]")] (refl RS box_equals) 1),
- (etac allE 1),(etac allE 1),
- (dtac mp 1),
- (res_inst_tac [("fo5","f")] (minimal RS monofun_cfun_arg) 1),
- (etac disjE 1),
- (contr_tac 1),
- (atac 1),
- (etac allE 1),
- (etac allE 1),
- (dtac mp 1),
- (res_inst_tac [("fo5","f")] (minimal RS monofun_cfun_arg) 1),
- (etac disjE 1),
- (contr_tac 1),
- (atac 1)
- ]);
--- a/src/HOLCF/fix.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,42 +0,0 @@
-(* Title: HOLCF/fix.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-definitions for fixed point operator and admissibility
-
-*)
-
-Fix = Cfun3 +
-
-consts
-
-iterate :: "nat=>('a->'a)=>'a=>'a"
-Ifix :: "('a->'a)=>'a"
-fix :: "('a->'a)->'a"
-adm :: "('a=>bool)=>bool"
-admw :: "('a=>bool)=>bool"
-chain_finite :: "'a=>bool"
-flat :: "'a=>bool"
-
-rules
-
-iterate_def "iterate(n,F,c) == nat_rec(n,c,%n x.F[x])"
-Ifix_def "Ifix(F) == lub(range(%i.iterate(i,F,UU)))"
-fix_def "fix == (LAM f. Ifix(f))"
-
-adm_def "adm(P) == !Y. is_chain(Y) --> \
-\ (!i.P(Y(i))) --> P(lub(range(Y)))"
-
-admw_def "admw(P)== (!F.((!n.P(iterate(n,F,UU)))-->\
-\ P(lub(range(%i.iterate(i,F,UU))))))"
-
-chain_finite_def "chain_finite(x::'a)==\
-\ !Y. is_chain(Y::nat=>'a) --> (? n.max_in_chain(n,Y))"
-
-flat_def "flat(x::'a) ==\
-\ ! x y. x::'a << y --> (x = UU) | (x=y)"
-
-end
-
--- a/src/HOLCF/fun1.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,49 +0,0 @@
-(* Title: HOLCF/fun1.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for fun1.thy
-*)
-
-open Fun1;
-
-(* ------------------------------------------------------------------------ *)
-(* less_fun is a partial order on 'a => 'b *)
-(* ------------------------------------------------------------------------ *)
-
-val refl_less_fun = prove_goalw Fun1.thy [less_fun_def] "less_fun(f,f)"
-(fn prems =>
- [
- (fast_tac (HOL_cs addSIs [refl_less]) 1)
- ]);
-
-val antisym_less_fun = prove_goalw Fun1.thy [less_fun_def]
- "[|less_fun(f1,f2); less_fun(f2,f1)|] ==> f1 = f2"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (expand_fun_eq RS ssubst) 1),
- (fast_tac (HOL_cs addSIs [antisym_less]) 1)
- ]);
-
-val trans_less_fun = prove_goalw Fun1.thy [less_fun_def]
- "[|less_fun(f1,f2); less_fun(f2,f3)|] ==> less_fun(f1,f3)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac trans_less 1),
- (etac allE 1),
- (atac 1),
- ((etac allE 1) THEN (atac 1))
- ]);
-
-(*
- --------------------------------------------------------------------------
- Since less_fun :: "['a::term=>'b::po,'a::term=>'b::po] => bool" the
- lemmas refl_less_fun, antisym_less_fun, trans_less_fun justify
- the class arity fun::(term,po)po !!
- --------------------------------------------------------------------------
-*)
-
--- a/src/HOLCF/fun1.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,30 +0,0 @@
-(* Title: HOLCF/fun1.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-Definition of the partial ordering for the type of all functions => (fun)
-
-REMARK: The ordering on 'a => 'b is only defined if 'b is in class po !!
-
-*)
-
-Fun1 = Pcpo +
-
-(* default class is still term *)
-
-consts
- less_fun :: "['a=>'b::po,'a=>'b] => bool"
-
-rules
- (* definition of the ordering less_fun *)
- (* in fun1.ML it is proved that less_fun is a po *)
-
- less_fun_def "less_fun(f1,f2) == ! x. f1(x) << f2(x)"
-
-end
-
-
-
-
--- a/src/HOLCF/fun2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,106 +0,0 @@
-(* Title: HOLCF/fun2.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for fun2.thy
-*)
-
-open Fun2;
-
-(* ------------------------------------------------------------------------ *)
-(* Type 'a::term => 'b::pcpo is pointed *)
-(* ------------------------------------------------------------------------ *)
-
-val minimal_fun = prove_goalw Fun2.thy [UU_fun_def] "UU_fun << f"
-(fn prems =>
- [
- (rtac (inst_fun_po RS ssubst) 1),
- (rewrite_goals_tac [less_fun_def]),
- (fast_tac (HOL_cs addSIs [minimal]) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* make the symbol << accessible for type fun *)
-(* ------------------------------------------------------------------------ *)
-
-val less_fun = prove_goal Fun2.thy "(f1 << f2) = (! x. f1(x) << f2(x))"
-(fn prems =>
- [
- (rtac (inst_fun_po RS ssubst) 1),
- (fold_goals_tac [less_fun_def]),
- (rtac refl 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* chains of functions yield chains in the po range *)
-(* ------------------------------------------------------------------------ *)
-
-val ch2ch_fun = prove_goal Fun2.thy
- "is_chain(S::nat=>('a::term => 'b::po)) ==> is_chain(% i.S(i)(x))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rewrite_goals_tac [is_chain]),
- (rtac allI 1),
- (rtac spec 1),
- (rtac (less_fun RS subst) 1),
- (etac allE 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* upper bounds of function chains yield upper bound in the po range *)
-(* ------------------------------------------------------------------------ *)
-
-val ub2ub_fun = prove_goal Fun2.thy
- " range(S::nat=>('a::term => 'b::po)) <| u ==> range(%i. S(i,x)) <| u(x)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (rtac allE 1),
- (rtac (less_fun RS subst) 1),
- (etac (ub_rangeE RS spec) 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Type 'a::term => 'b::pcpo is chain complete *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_fun = prove_goal Fun2.thy
- "is_chain(S::nat=>('a::term => 'b::pcpo)) ==> \
-\ range(S) <<| (% x.lub(range(% i.S(i)(x))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (rtac (less_fun RS ssubst) 1),
- (rtac allI 1),
- (rtac is_ub_thelub 1),
- (etac ch2ch_fun 1),
- (strip_tac 1),
- (rtac (less_fun RS ssubst) 1),
- (rtac allI 1),
- (rtac is_lub_thelub 1),
- (etac ch2ch_fun 1),
- (etac ub2ub_fun 1)
- ]);
-
-val thelub_fun = (lub_fun RS thelubI);
-(* is_chain(?S1) ==> lub(range(?S1)) = (%x. lub(range(%i. ?S1(i,x)))) *)
-
-val cpo_fun = prove_goal Fun2.thy
- "is_chain(S::nat=>('a::term => 'b::pcpo)) ==> ? x. range(S) <<| x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac exI 1),
- (etac lub_fun 1)
- ]);
-
--- a/src/HOLCF/fun2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,40 +0,0 @@
-(* Title: HOLCF/fun2.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class Instance =>::(term,po)po
-Definiton of least element
-*)
-
-Fun2 = Fun1 +
-
-(* default class is still term !*)
-
-(* Witness for the above arity axiom is fun1.ML *)
-
-arities fun :: (term,po)po
-
-consts
- UU_fun :: "'a::term => 'b::pcpo"
-
-rules
-
-(* instance of << for type ['a::term => 'b::po] *)
-
-inst_fun_po "(op <<)::['a=>'b::po,'a=>'b::po ]=>bool = less_fun"
-
-(* definitions *)
-(* The least element in type 'a::term => 'b::pcpo *)
-
-UU_fun_def "UU_fun == (% x.UU)"
-
-end
-
-
-
-
-
-
-
-
--- a/src/HOLCF/fun3.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,7 +0,0 @@
-(* Title: HOLCF/fun3.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-*)
-
-open Fun3;
--- a/src/HOLCF/fun3.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,23 +0,0 @@
-(* Title: HOLCF/fun3.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class instance of => (fun) for class pcpo
-
-*)
-
-Fun3 = Fun2 +
-
-(* default class is still term *)
-
-arities fun :: (term,pcpo)pcpo (* Witness fun2.ML *)
-
-rules
-
-inst_fun_pcpo "UU::'a=>'b::pcpo = UU_fun"
-
-end
-
-
-
--- a/src/HOLCF/holcf.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,20 +0,0 @@
-(* Title: HOLCF/HOLCF.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-*)
-
-open HOLCF;
-
-val HOLCF_ss = ccc1_ss
- addsimps one_when
- addsimps dist_less_one
- addsimps dist_eq_one
- addsimps dist_less_tr
- addsimps dist_eq_tr
- addsimps tr_when
- addsimps andalso_thms
- addsimps orelse_thms
- addsimps ifte_thms;
-
-
--- a/src/HOLCF/holcf.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(* Title: HOLCF/HOLCF.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-Top theory for HOLCF system
-
-*)
-
-HOLCF = Tr2
-
-
--- a/src/HOLCF/holcfb.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,152 +0,0 @@
-(* Title: HOLCF/holcfb.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for Holcfb.thy
-*)
-
-open Holcfb;
-
-(* ------------------------------------------------------------------------ *)
-(* <::nat=>nat=>bool is well-founded *)
-(* ------------------------------------------------------------------------ *)
-
-val well_founded_nat = prove_goal Nat.thy
- "!P. P(x::nat) --> (? y. P(y) & (! x. P(x) --> y <= x))"
- (fn prems =>
- [
- (res_inst_tac [("n","x")] less_induct 1),
- (strip_tac 1),
- (res_inst_tac [("Q","? k.k<n & P(k)")] (excluded_middle RS disjE) 1),
- (etac exE 2),
- (etac conjE 2),
- (rtac mp 2),
- (etac allE 2),
- (etac impE 2),
- (atac 2),
- (etac spec 2),
- (atac 2),
- (res_inst_tac [("x","n")] exI 1),
- (rtac conjI 1),
- (atac 1),
- (strip_tac 1),
- (rtac leI 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Lemmas for selecting the least element in a nonempty set *)
-(* ------------------------------------------------------------------------ *)
-
-val theleast = prove_goalw Holcfb.thy [theleast_def]
-"P(x) ==> P(theleast(P)) & (!x.P(x)--> theleast(P) <= x)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (well_founded_nat RS spec RS mp RS exE) 1),
- (atac 1),
- (rtac selectI 1),
- (atac 1)
- ]);
-
-val theleast1= theleast RS conjunct1;
-(* ?P1(?x1) ==> ?P1(theleast(?P1)) *)
-
-val theleast2 = prove_goal Holcfb.thy "P(x) ==> theleast(P) <= x"
- (fn prems =>
- [
- (rtac (theleast RS conjunct2 RS spec RS mp) 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Some lemmas in HOL.thy *)
-(* ------------------------------------------------------------------------ *)
-
-
-val de_morgan1 = prove_goal HOL.thy "(~a & ~b)=(~(a | b))"
-(fn prems =>
- [
- (rtac iffI 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val de_morgan2 = prove_goal HOL.thy "(~a | ~b)=(~(a & b))"
-(fn prems =>
- [
- (rtac iffI 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val notall2ex = prove_goal HOL.thy "(~ (!x.P(x))) = (? x.~P(x))"
-(fn prems =>
- [
- (rtac iffI 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val notex2all = prove_goal HOL.thy "(~ (? x.P(x))) = (!x.~P(x))"
-(fn prems =>
- [
- (rtac iffI 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val selectI2 = prove_goal HOL.thy "(? x. P(x)) ==> P(@ x.P(x))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exE 1),
- (etac selectI 1)
- ]);
-
-val selectE = prove_goal HOL.thy "P(@ x.P(x)) ==> (? x. P(x))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exI 1)
- ]);
-
-val select_eq_Ex = prove_goal HOL.thy "(P(@ x.P(x))) = (? x. P(x))"
-(fn prems =>
- [
- (rtac (iff RS mp RS mp) 1),
- (strip_tac 1),
- (etac selectE 1),
- (strip_tac 1),
- (etac selectI2 1)
- ]);
-
-
-val notnotI = prove_goal HOL.thy "P ==> ~~P"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val classical2 = prove_goal HOL.thy "[|Q ==> R; ~Q ==> R|]==>R"
- (fn prems =>
- [
- (rtac disjE 1),
- (rtac excluded_middle 1),
- (eresolve_tac prems 1),
- (eresolve_tac prems 1)
- ]);
-
-
-
-val classical3 = (notE RS notI);
-(* [| ?P ==> ~ ?P1; ?P ==> ?P1 |] ==> ~ ?P *)
-
--- a/src/HOLCF/holcfb.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-(* Title: HOLCF/holcfb.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Basic definitions for the embedding of LCF into HOL.
-
-*)
-
-Holcfb = Nat +
-
-consts
-
-theleast :: "(nat=>bool)=>nat"
-
-rules
-
-theleast_def "theleast(P) == (@z.(P(z) & (!n.P(n)-->z<=n)))"
-
-end
-
-
-
-
-
--- a/src/HOLCF/lift1.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,188 +0,0 @@
-(* Title: HOLCF/lift1.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-*)
-
-open Lift1;
-
-val Exh_Lift = prove_goalw Lift1.thy [UU_lift_def,Iup_def ]
- "z = UU_lift | (? x. z = Iup(x))"
- (fn prems =>
- [
- (rtac (Rep_Lift_inverse RS subst) 1),
- (res_inst_tac [("s","Rep_Lift(z)")] sumE 1),
- (rtac disjI1 1),
- (res_inst_tac [("f","Abs_Lift")] arg_cong 1),
- (rtac (unique_void2 RS subst) 1),
- (atac 1),
- (rtac disjI2 1),
- (rtac exI 1),
- (res_inst_tac [("f","Abs_Lift")] arg_cong 1),
- (atac 1)
- ]);
-
-val inj_Abs_Lift = prove_goal Lift1.thy "inj(Abs_Lift)"
- (fn prems =>
- [
- (rtac inj_inverseI 1),
- (rtac Abs_Lift_inverse 1)
- ]);
-
-val inj_Rep_Lift = prove_goal Lift1.thy "inj(Rep_Lift)"
- (fn prems =>
- [
- (rtac inj_inverseI 1),
- (rtac Rep_Lift_inverse 1)
- ]);
-
-val inject_Iup = prove_goalw Lift1.thy [Iup_def] "Iup(x)=Iup(y) ==> x=y"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (inj_Inr RS injD) 1),
- (rtac (inj_Abs_Lift RS injD) 1),
- (atac 1)
- ]);
-
-val defined_Iup=prove_goalw Lift1.thy [Iup_def,UU_lift_def] "~ Iup(x)=UU_lift"
- (fn prems =>
- [
- (rtac notI 1),
- (rtac notE 1),
- (rtac Inl_not_Inr 1),
- (rtac sym 1),
- (etac (inj_Abs_Lift RS injD) 1)
- ]);
-
-
-val liftE = prove_goal Lift1.thy
- "[| p=UU_lift ==> Q; !!x. p=Iup(x)==>Q|] ==>Q"
- (fn prems =>
- [
- (rtac (Exh_Lift RS disjE) 1),
- (eresolve_tac prems 1),
- (etac exE 1),
- (eresolve_tac prems 1)
- ]);
-
-val Ilift1 = prove_goalw Lift1.thy [Ilift_def,UU_lift_def]
- "Ilift(f)(UU_lift)=UU"
- (fn prems =>
- [
- (rtac (Abs_Lift_inverse RS ssubst) 1),
- (rtac (sum_case_Inl RS ssubst) 1),
- (rtac refl 1)
- ]);
-
-val Ilift2 = prove_goalw Lift1.thy [Ilift_def,Iup_def]
- "Ilift(f)(Iup(x))=f[x]"
- (fn prems =>
- [
- (rtac (Abs_Lift_inverse RS ssubst) 1),
- (rtac (sum_case_Inr RS ssubst) 1),
- (rtac refl 1)
- ]);
-
-val Lift_ss = Cfun_ss addsimps [Ilift1,Ilift2];
-
-val less_lift1a = prove_goalw Lift1.thy [less_lift_def,UU_lift_def]
- "less_lift(UU_lift)(z)"
- (fn prems =>
- [
- (rtac (Abs_Lift_inverse RS ssubst) 1),
- (rtac (sum_case_Inl RS ssubst) 1),
- (rtac TrueI 1)
- ]);
-
-val less_lift1b = prove_goalw Lift1.thy [Iup_def,less_lift_def,UU_lift_def]
- "~less_lift(Iup(x),UU_lift)"
- (fn prems =>
- [
- (rtac notI 1),
- (rtac iffD1 1),
- (atac 2),
- (rtac (Abs_Lift_inverse RS ssubst) 1),
- (rtac (Abs_Lift_inverse RS ssubst) 1),
- (rtac (sum_case_Inr RS ssubst) 1),
- (rtac (sum_case_Inl RS ssubst) 1),
- (rtac refl 1)
- ]);
-
-val less_lift1c = prove_goalw Lift1.thy [Iup_def,less_lift_def,UU_lift_def]
- "less_lift(Iup(x),Iup(y))=(x<<y)"
- (fn prems =>
- [
- (rtac (Abs_Lift_inverse RS ssubst) 1),
- (rtac (Abs_Lift_inverse RS ssubst) 1),
- (rtac (sum_case_Inr RS ssubst) 1),
- (rtac (sum_case_Inr RS ssubst) 1),
- (rtac refl 1)
- ]);
-
-
-val refl_less_lift = prove_goal Lift1.thy "less_lift(p,p)"
- (fn prems =>
- [
- (res_inst_tac [("p","p")] liftE 1),
- (hyp_subst_tac 1),
- (rtac less_lift1a 1),
- (hyp_subst_tac 1),
- (rtac (less_lift1c RS iffD2) 1),
- (rtac refl_less 1)
- ]);
-
-val antisym_less_lift = prove_goal Lift1.thy
- "[|less_lift(p1,p2);less_lift(p2,p1)|] ==> p1=p2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] liftE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] liftE 1),
- (hyp_subst_tac 1),
- (rtac refl 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("P","less_lift(Iup(x),UU_lift)")] notE 1),
- (rtac less_lift1b 1),
- (atac 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] liftE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("P","less_lift(Iup(x),UU_lift)")] notE 1),
- (rtac less_lift1b 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac arg_cong 1),
- (rtac antisym_less 1),
- (etac (less_lift1c RS iffD1) 1),
- (etac (less_lift1c RS iffD1) 1)
- ]);
-
-val trans_less_lift = prove_goal Lift1.thy
- "[|less_lift(p1,p2);less_lift(p2,p3)|] ==> less_lift(p1,p3)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] liftE 1),
- (hyp_subst_tac 1),
- (rtac less_lift1a 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] liftE 1),
- (hyp_subst_tac 1),
- (rtac notE 1),
- (rtac less_lift1b 1),
- (atac 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p3")] liftE 1),
- (hyp_subst_tac 1),
- (rtac notE 1),
- (rtac less_lift1b 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac (less_lift1c RS iffD2) 1),
- (rtac trans_less 1),
- (etac (less_lift1c RS iffD1) 1),
- (etac (less_lift1c RS iffD1) 1)
- ]);
-
--- a/src/HOLCF/lift1.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,55 +0,0 @@
-(* Title: HOLCF/lift1.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-Lifting
-
-*)
-
-Lift1 = Cfun3 +
-
-(* new type for lifting *)
-
-types "u" 1
-
-arities "u" :: (pcpo)term
-
-consts
-
- Rep_Lift :: "('a)u => (void + 'a)"
- Abs_Lift :: "(void + 'a) => ('a)u"
-
- Iup :: "'a => ('a)u"
- UU_lift :: "('a)u"
- Ilift :: "('a->'b)=>('a)u => 'b"
- less_lift :: "('a)u => ('a)u => bool"
-
-rules
-
- (*faking a type definition... *)
- (* ('a)u is isomorphic to void + 'a *)
-
- Rep_Lift_inverse "Abs_Lift(Rep_Lift(p)) = p"
- Abs_Lift_inverse "Rep_Lift(Abs_Lift(p)) = p"
-
- (*defining the abstract constants*)
-
- UU_lift_def "UU_lift == Abs_Lift(Inl(UU))"
- Iup_def "Iup(x) == Abs_Lift(Inr(x))"
-
- Ilift_def "Ilift(f)(x)==\
-\ sum_case (Rep_Lift(x)) (%y.UU) (%z.f[z])"
-
- less_lift_def "less_lift(x1)(x2) == \
-\ (sum_case (Rep_Lift(x1))\
-\ (% y1.True)\
-\ (% y2.sum_case (Rep_Lift(x2))\
-\ (% z1.False)\
-\ (% z2.y2<<z2)))"
-
-end
-
-
-
--- a/src/HOLCF/lift2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,182 +0,0 @@
-(* Title: HOLCF/lift2.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for lift2.thy
-*)
-
-open Lift2;
-
-(* -------------------------------------------------------------------------*)
-(* type ('a)u is pointed *)
-(* ------------------------------------------------------------------------ *)
-
-val minimal_lift = prove_goal Lift2.thy "UU_lift << z"
- (fn prems =>
- [
- (rtac (inst_lift_po RS ssubst) 1),
- (rtac less_lift1a 1)
- ]);
-
-(* -------------------------------------------------------------------------*)
-(* access to less_lift in class po *)
-(* ------------------------------------------------------------------------ *)
-
-val less_lift2b = prove_goal Lift2.thy "~ Iup(x) << UU_lift"
- (fn prems =>
- [
- (rtac (inst_lift_po RS ssubst) 1),
- (rtac less_lift1b 1)
- ]);
-
-val less_lift2c = prove_goal Lift2.thy "(Iup(x)<<Iup(y)) = (x<<y)"
- (fn prems =>
- [
- (rtac (inst_lift_po RS ssubst) 1),
- (rtac less_lift1c 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Iup and Ilift are monotone *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_Iup = prove_goalw Lift2.thy [monofun] "monofun(Iup)"
- (fn prems =>
- [
- (strip_tac 1),
- (etac (less_lift2c RS iffD2) 1)
- ]);
-
-val monofun_Ilift1 = prove_goalw Lift2.thy [monofun] "monofun(Ilift)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (res_inst_tac [("p","xa")] liftE 1),
- (asm_simp_tac Lift_ss 1),
- (asm_simp_tac Lift_ss 1),
- (etac monofun_cfun_fun 1)
- ]);
-
-val monofun_Ilift2 = prove_goalw Lift2.thy [monofun] "monofun(Ilift(f))"
- (fn prems =>
- [
- (strip_tac 1),
- (res_inst_tac [("p","x")] liftE 1),
- (asm_simp_tac Lift_ss 1),
- (asm_simp_tac Lift_ss 1),
- (res_inst_tac [("p","y")] liftE 1),
- (hyp_subst_tac 1),
- (hyp_subst_tac 1),
- (rtac notE 1),
- (rtac less_lift2b 1),
- (atac 1),
- (asm_simp_tac Lift_ss 1),
- (rtac monofun_cfun_arg 1),
- (hyp_subst_tac 1),
- (hyp_subst_tac 1),
- (etac (less_lift2c RS iffD1) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Some kind of surjectivity lemma *)
-(* ------------------------------------------------------------------------ *)
-
-
-val lift_lemma1 = prove_goal Lift2.thy "z=Iup(x) ==> Iup(Ilift(LAM x.x)(z)) = z"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac Lift_ss 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* ('a)u is a cpo *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_lift1a = prove_goal Lift2.thy
-"[|is_chain(Y);? i x.Y(i)=Iup(x)|] ==>\
-\ range(Y) <<| Iup(lub(range(%i.(Ilift (LAM x.x) (Y(i))))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (res_inst_tac [("p","Y(i)")] liftE 1),
- (res_inst_tac [("s","UU_lift"),("t","Y(i)")] subst 1),
- (etac sym 1),
- (rtac minimal_lift 1),
- (res_inst_tac [("t","Y(i)")] (lift_lemma1 RS subst) 1),
- (atac 1),
- (rtac (less_lift2c RS iffD2) 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_Ilift2 RS ch2ch_monofun) 1),
- (strip_tac 1),
- (res_inst_tac [("p","u")] liftE 1),
- (etac exE 1),
- (etac exE 1),
- (res_inst_tac [("P","Y(i)<<UU_lift")] notE 1),
- (res_inst_tac [("s","Iup(x)"),("t","Y(i)")] ssubst 1),
- (atac 1),
- (rtac less_lift2b 1),
- (hyp_subst_tac 1),
- (etac (ub_rangeE RS spec) 1),
- (res_inst_tac [("t","u")] (lift_lemma1 RS subst) 1),
- (atac 1),
- (rtac (less_lift2c RS iffD2) 1),
- (rtac is_lub_thelub 1),
- (etac (monofun_Ilift2 RS ch2ch_monofun) 1),
- (etac (monofun_Ilift2 RS ub2ub_monofun) 1)
- ]);
-
-val lub_lift1b = prove_goal Lift2.thy
-"[|is_chain(Y);!i x.~Y(i)=Iup(x)|] ==>\
-\ range(Y) <<| UU_lift"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (res_inst_tac [("p","Y(i)")] liftE 1),
- (res_inst_tac [("s","UU_lift"),("t","Y(i)")] ssubst 1),
- (atac 1),
- (rtac refl_less 1),
- (rtac notE 1),
- (dtac spec 1),
- (dtac spec 1),
- (atac 1),
- (atac 1),
- (strip_tac 1),
- (rtac minimal_lift 1)
- ]);
-
-val thelub_lift1a = lub_lift1a RS thelubI;
-(* [| is_chain(?Y1); ? i x. ?Y1(i) = Iup(x) |] ==> *)
-(* lub(range(?Y1)) = Iup(lub(range(%i. Ilift(LAM x. x,?Y1(i))))) *)
-
-val thelub_lift1b = lub_lift1b RS thelubI;
-(* [| is_chain(?Y1); ! i x. ~ ?Y1(i) = Iup(x) |] ==> *)
-(* lub(range(?Y1)) = UU_lift *)
-
-
-val cpo_lift = prove_goal Lift2.thy
- "is_chain(Y::nat=>('a)u) ==> ? x.range(Y) <<|x"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac disjE 1),
- (rtac exI 2),
- (etac lub_lift1a 2),
- (atac 2),
- (rtac exI 2),
- (etac lub_lift1b 2),
- (atac 2),
- (fast_tac HOL_cs 1)
- ]);
-
--- a/src/HOLCF/lift2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-(* Title: HOLCF/lift2.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class Instance u::(pcpo)po
-
-*)
-
-Lift2 = Lift1 +
-
-(* Witness for the above arity axiom is lift1.ML *)
-
-arities "u" :: (pcpo)po
-
-rules
-
-(* instance of << for type ('a)u *)
-
-inst_lift_po "(op <<)::[('a)u,('a)u]=>bool = less_lift"
-
-end
-
-
-
--- a/src/HOLCF/lift3.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,349 +0,0 @@
-(* Title: HOLCF/lift3.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for lift3.thy
-*)
-
-open Lift3;
-
-(* -------------------------------------------------------------------------*)
-(* some lemmas restated for class pcpo *)
-(* ------------------------------------------------------------------------ *)
-
-val less_lift3b = prove_goal Lift3.thy "~ Iup(x) << UU"
- (fn prems =>
- [
- (rtac (inst_lift_pcpo RS ssubst) 1),
- (rtac less_lift2b 1)
- ]);
-
-val defined_Iup2 = prove_goal Lift3.thy "~ Iup(x) = UU"
- (fn prems =>
- [
- (rtac (inst_lift_pcpo RS ssubst) 1),
- (rtac defined_Iup 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* continuity for Iup *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_Iup = prove_goal Lift3.thy "contlub(Iup)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_lift1a RS sym) 2),
- (fast_tac HOL_cs 3),
- (etac (monofun_Iup RS ch2ch_monofun) 2),
- (res_inst_tac [("f","Iup")] arg_cong 1),
- (rtac lub_equal 1),
- (atac 1),
- (rtac (monofun_Ilift2 RS ch2ch_monofun) 1),
- (etac (monofun_Iup RS ch2ch_monofun) 1),
- (asm_simp_tac Lift_ss 1)
- ]);
-
-val contX_Iup = prove_goal Lift3.thy "contX(Iup)"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Iup 1),
- (rtac contlub_Iup 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* continuity for Ilift *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_Ilift1 = prove_goal Lift3.thy "contlub(Ilift)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_fun RS sym) 2),
- (etac (monofun_Ilift1 RS ch2ch_monofun) 2),
- (rtac ext 1),
- (res_inst_tac [("p","x")] liftE 1),
- (asm_simp_tac Lift_ss 1),
- (rtac (lub_const RS thelubI RS sym) 1),
- (asm_simp_tac Lift_ss 1),
- (etac contlub_cfun_fun 1)
- ]);
-
-
-val contlub_Ilift2 = prove_goal Lift3.thy "contlub(Ilift(f))"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac disjE 1),
- (rtac (thelub_lift1a RS ssubst) 2),
- (atac 2),
- (atac 2),
- (asm_simp_tac Lift_ss 2),
- (rtac (thelub_lift1b RS ssubst) 3),
- (atac 3),
- (atac 3),
- (fast_tac HOL_cs 1),
- (asm_simp_tac Lift_ss 2),
- (rtac (chain_UU_I_inverse RS sym) 2),
- (rtac allI 2),
- (res_inst_tac [("p","Y(i)")] liftE 2),
- (asm_simp_tac Lift_ss 2),
- (rtac notE 2),
- (dtac spec 2),
- (etac spec 2),
- (atac 2),
- (rtac (contlub_cfun_arg RS ssubst) 1),
- (etac (monofun_Ilift2 RS ch2ch_monofun) 1),
- (rtac lub_equal2 1),
- (rtac (monofun_fapp2 RS ch2ch_monofun) 2),
- (etac (monofun_Ilift2 RS ch2ch_monofun) 2),
- (etac (monofun_Ilift2 RS ch2ch_monofun) 2),
- (rtac (chain_mono2 RS exE) 1),
- (atac 2),
- (etac exE 1),
- (etac exE 1),
- (rtac exI 1),
- (res_inst_tac [("s","Iup(x)"),("t","Y(i)")] ssubst 1),
- (atac 1),
- (rtac defined_Iup2 1),
- (rtac exI 1),
- (strip_tac 1),
- (res_inst_tac [("p","Y(i)")] liftE 1),
- (asm_simp_tac Lift_ss 2),
- (res_inst_tac [("P","Y(i) = UU")] notE 1),
- (fast_tac HOL_cs 1),
- (rtac (inst_lift_pcpo RS ssubst) 1),
- (atac 1)
- ]);
-
-val contX_Ilift1 = prove_goal Lift3.thy "contX(Ilift)"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Ilift1 1),
- (rtac contlub_Ilift1 1)
- ]);
-
-val contX_Ilift2 = prove_goal Lift3.thy "contX(Ilift(f))"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Ilift2 1),
- (rtac contlub_Ilift2 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* continuous versions of lemmas for ('a)u *)
-(* ------------------------------------------------------------------------ *)
-
-val Exh_Lift1 = prove_goalw Lift3.thy [up_def] "z = UU | (? x. z = up[x])"
- (fn prems =>
- [
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1),
- (rtac (inst_lift_pcpo RS ssubst) 1),
- (rtac Exh_Lift 1)
- ]);
-
-val inject_up = prove_goalw Lift3.thy [up_def] "up[x]=up[y] ==> x=y"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac inject_Iup 1),
- (etac box_equals 1),
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1),
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1)
- ]);
-
-val defined_up = prove_goalw Lift3.thy [up_def] "~ up[x]=UU"
- (fn prems =>
- [
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1),
- (rtac defined_Iup2 1)
- ]);
-
-val liftE1 = prove_goalw Lift3.thy [up_def]
- "[| p=UU ==> Q; !!x. p=up[x]==>Q|] ==>Q"
- (fn prems =>
- [
- (rtac liftE 1),
- (resolve_tac prems 1),
- (etac (inst_lift_pcpo RS ssubst) 1),
- (resolve_tac (tl prems) 1),
- (asm_simp_tac (Lift_ss addsimps [contX_Iup]) 1)
- ]);
-
-
-val lift1 = prove_goalw Lift3.thy [up_def,lift_def] "lift[f][UU]=UU"
- (fn prems =>
- [
- (rtac (inst_lift_pcpo RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1,
- contX_Ilift2,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1,
- contX_Ilift2,contX2contX_CF1L]) 1)),
- (simp_tac (Lift_ss addsimps [contX_Iup,contX_Ilift1,contX_Ilift2]) 1)
- ]);
-
-val lift2 = prove_goalw Lift3.thy [up_def,lift_def] "lift[f][up[x]]=f[x]"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Iup 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1,
- contX_Ilift2,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Ilift2 1),
- (simp_tac (Lift_ss addsimps [contX_Iup,contX_Ilift1,contX_Ilift2]) 1)
- ]);
-
-val less_lift4b = prove_goalw Lift3.thy [up_def,lift_def] "~ up[x] << UU"
- (fn prems =>
- [
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1),
- (rtac less_lift3b 1)
- ]);
-
-val less_lift4c = prove_goalw Lift3.thy [up_def,lift_def]
- "(up[x]<<up[y]) = (x<<y)"
- (fn prems =>
- [
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1),
- (rtac less_lift2c 1)
- ]);
-
-val thelub_lift2a = prove_goalw Lift3.thy [up_def,lift_def]
-"[| is_chain(Y); ? i x. Y(i) = up[x] |] ==>\
-\ lub(range(Y)) = up[lub(range(%i. lift[LAM x. x][Y(i)]))]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1,
- contX_Ilift2,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1,
- contX_Ilift2,contX2contX_CF1L]) 1)),
-
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1,
- contX_Ilift2,contX2contX_CF1L]) 1)),
- (rtac thelub_lift1a 1),
- (atac 1),
- (etac exE 1),
- (etac exE 1),
- (rtac exI 1),
- (rtac exI 1),
- (etac box_equals 1),
- (rtac refl 1),
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1)
- ]);
-
-
-
-val thelub_lift2b = prove_goalw Lift3.thy [up_def,lift_def]
-"[| is_chain(Y); ! i x. ~ Y(i) = up[x] |] ==> lub(range(Y)) = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (inst_lift_pcpo RS ssubst) 1),
- (rtac thelub_lift1b 1),
- (atac 1),
- (strip_tac 1),
- (dtac spec 1),
- (dtac spec 1),
- (rtac swap 1),
- (atac 1),
- (dtac notnotD 1),
- (etac box_equals 1),
- (rtac refl 1),
- (simp_tac (Lift_ss addsimps [contX_Iup]) 1)
- ]);
-
-
-val lift_lemma2 = prove_goal Lift3.thy " (? x.z = up[x]) = (~z=UU)"
- (fn prems =>
- [
- (rtac iffI 1),
- (etac exE 1),
- (hyp_subst_tac 1),
- (rtac defined_up 1),
- (res_inst_tac [("p","z")] liftE1 1),
- (etac notE 1),
- (atac 1),
- (etac exI 1)
- ]);
-
-
-val thelub_lift2a_rev = prove_goal Lift3.thy
-"[| is_chain(Y); lub(range(Y)) = up[x] |] ==> ? i x. Y(i) = up[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac exE 1),
- (rtac chain_UU_I_inverse2 1),
- (rtac (lift_lemma2 RS iffD1) 1),
- (etac exI 1),
- (rtac exI 1),
- (rtac (lift_lemma2 RS iffD2) 1),
- (atac 1)
- ]);
-
-val thelub_lift2b_rev = prove_goal Lift3.thy
-"[| is_chain(Y); lub(range(Y)) = UU |] ==> ! i x. ~ Y(i) = up[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac allI 1),
- (rtac (notex2all RS iffD1) 1),
- (rtac contrapos 1),
- (etac (lift_lemma2 RS iffD1) 2),
- (rtac notnotI 1),
- (rtac (chain_UU_I RS spec) 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-val thelub_lift3 = prove_goal Lift3.thy
-"is_chain(Y) ==> lub(range(Y)) = UU |\
-\ lub(range(Y)) = up[lub(range(%i. lift[LAM x. x][Y(i)]))]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac disjE 1),
- (rtac disjI1 2),
- (rtac thelub_lift2b 2),
- (atac 2),
- (atac 2),
- (rtac disjI2 2),
- (rtac thelub_lift2a 2),
- (atac 2),
- (atac 2),
- (fast_tac HOL_cs 1)
- ]);
-
-val lift3 = prove_goal Lift3.thy "lift[up][x]=x"
- (fn prems =>
- [
- (res_inst_tac [("p","x")] liftE1 1),
- (asm_simp_tac (Cfun_ss addsimps [lift1,lift2]) 1),
- (asm_simp_tac (Cfun_ss addsimps [lift1,lift2]) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* install simplifier for ('a)u *)
-(* ------------------------------------------------------------------------ *)
-
-val lift_rews = [lift1,lift2,defined_up];
--- a/src/HOLCF/lift3.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,29 +0,0 @@
-(* Title: HOLCF/lift3.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-
-Class instance of ('a)u for class pcpo
-
-*)
-
-Lift3 = Lift2 +
-
-arities "u" :: (pcpo)pcpo (* Witness lift2.ML *)
-
-consts
- up :: "'a -> ('a)u"
- lift :: "('a->'c)-> ('a)u -> 'c"
-
-rules
-
-inst_lift_pcpo "UU::('a)u = UU_lift"
-
-up_def "up == (LAM x.Iup(x))"
-lift_def "lift == (LAM f p.Ilift(f)(p))"
-
-end
-
-
-
--- a/src/HOLCF/one.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,96 +0,0 @@
-(* Title: HOLCF/one.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for one.thy
-*)
-
-open One;
-
-(* ------------------------------------------------------------------------ *)
-(* Exhaustion and Elimination for type one *)
-(* ------------------------------------------------------------------------ *)
-
-val Exh_one = prove_goalw One.thy [one_def] "z=UU | z = one"
- (fn prems =>
- [
- (res_inst_tac [("p","rep_one[z]")] liftE1 1),
- (rtac disjI1 1),
- (rtac ((abs_one_iso RS allI) RS ((rep_one_iso RS allI) RS iso_strict )
- RS conjunct2 RS subst) 1),
- (rtac (abs_one_iso RS subst) 1),
- (etac cfun_arg_cong 1),
- (rtac disjI2 1),
- (rtac (abs_one_iso RS subst) 1),
- (rtac cfun_arg_cong 1),
- (rtac (unique_void2 RS subst) 1),
- (atac 1)
- ]);
-
-val oneE = prove_goal One.thy
- "[| p=UU ==> Q; p = one ==>Q|] ==>Q"
- (fn prems =>
- [
- (rtac (Exh_one RS disjE) 1),
- (eresolve_tac prems 1),
- (eresolve_tac prems 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* distinctness for type one : stored in a list *)
-(* ------------------------------------------------------------------------ *)
-
-val dist_less_one = [
-prove_goalw One.thy [one_def] "~one << UU"
- (fn prems =>
- [
- (rtac classical3 1),
- (rtac less_lift4b 1),
- (rtac (rep_one_iso RS subst) 1),
- (rtac (rep_one_iso RS subst) 1),
- (rtac monofun_cfun_arg 1),
- (etac ((abs_one_iso RS allI) RS ((rep_one_iso RS allI) RS iso_strict )
- RS conjunct2 RS ssubst) 1)
- ])
-];
-
-val dist_eq_one = [prove_goal One.thy "~one=UU"
- (fn prems =>
- [
- (rtac not_less2not_eq 1),
- (resolve_tac dist_less_one 1)
- ])];
-
-val dist_eq_one = dist_eq_one @ (map (fn thm => (thm RS not_sym)) dist_eq_one);
-
-(* ------------------------------------------------------------------------ *)
-(* one is flat *)
-(* ------------------------------------------------------------------------ *)
-
-val prems = goalw One.thy [flat_def] "flat(one)";
-by (rtac allI 1);
-by (rtac allI 1);
-by (res_inst_tac [("p","x")] oneE 1);
-by (asm_simp_tac ccc1_ss 1);
-by (res_inst_tac [("p","y")] oneE 1);
-by (asm_simp_tac (ccc1_ss addsimps dist_less_one) 1);
-by (asm_simp_tac ccc1_ss 1);
-val flat_one = result();
-
-
-(* ------------------------------------------------------------------------ *)
-(* properties of one_when *)
-(* here I tried a generic prove procedure *)
-(* ------------------------------------------------------------------------ *)
-
-fun prover s = prove_goalw One.thy [one_when_def,one_def] s
- (fn prems =>
- [
- (simp_tac (ccc1_ss addsimps [(rep_one_iso ),
- (abs_one_iso RS allI) RS ((rep_one_iso RS allI)
- RS iso_strict) RS conjunct1] )1)
- ]);
-
-val one_when = map prover ["one_when[x][UU] = UU","one_when[x][one] = x"];
-
--- a/src/HOLCF/one.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,53 +0,0 @@
-(* Title: HOLCF/one.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Introduve atomic type one = (void)u
-
-This is the first type that is introduced using a domain isomorphism.
-The type axiom
-
- arities one :: pcpo
-
-and the continuity of the Isomorphisms are taken for granted. Since the
-type is not recursive, it could be easily introduced in a purely conservative
-style as it was used for the types sprod, ssum, lift. The definition of the
-ordering is canonical using abstraction and representation, but this would take
-again a lot of internal constants. It can easily be seen that the axioms below
-are consistent.
-
-The partial ordering on type one is implicitly defined via the
-isomorphism axioms and the continuity of abs_one and rep_one.
-
-We could also introduce the function less_one with definition
-
-less_one(x,y) = rep_one[x] << rep_one[y]
-
-
-*)
-
-One = ccc1+
-
-types one 0
-arities one :: pcpo
-
-consts
- abs_one :: "(void)u -> one"
- rep_one :: "one -> (void)u"
- one :: "one"
- one_when :: "'c -> one -> 'c"
-
-rules
- abs_one_iso "abs_one[rep_one[u]] = u"
- rep_one_iso "rep_one[abs_one[x]] = x"
-
- one_def "one == abs_one[up[UU]]"
- one_when_def "one_when == (LAM c u.lift[LAM x.c][rep_one[u]])"
-
-end
-
-
-
-
-
--- a/src/HOLCF/pcpo.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,272 +0,0 @@
-(* Title: HOLCF/pcpo.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for pcpo.thy
-*)
-
-open Pcpo;
-
-(* ------------------------------------------------------------------------ *)
-(* in pcpo's everthing equal to THE lub has lub properties for every chain *)
-(* ------------------------------------------------------------------------ *)
-
-val thelubE = prove_goal Pcpo.thy
- "[| is_chain(S);lub(range(S)) = l::'a::pcpo|] ==> range(S) <<| l "
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (rtac lubI 1),
- (etac cpo 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Properties of the lub *)
-(* ------------------------------------------------------------------------ *)
-
-
-val is_ub_thelub = (cpo RS lubI RS is_ub_lub);
-(* is_chain(?S1) ==> ?S1(?x) << lub(range(?S1)) *)
-
-val is_lub_thelub = (cpo RS lubI RS is_lub_lub);
-(* [| is_chain(?S5); range(?S5) <| ?x1 |] ==> lub(range(?S5)) << ?x1 *)
-
-
-(* ------------------------------------------------------------------------ *)
-(* the << relation between two chains is preserved by their lubs *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_mono = prove_goal Pcpo.thy
- "[|is_chain(C1::(nat=>'a::pcpo));is_chain(C2); ! k. C1(k) << C2(k)|]\
-\ ==> lub(range(C1)) << lub(range(C2))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac is_lub_thelub 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (rtac trans_less 1),
- (etac spec 1),
- (etac is_ub_thelub 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the = relation between two chains is preserved by their lubs *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_equal = prove_goal Pcpo.thy
-"[| is_chain(C1::(nat=>'a::pcpo));is_chain(C2);!k.C1(k)=C2(k)|]\
-\ ==> lub(range(C1))=lub(range(C2))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac antisym_less 1),
- (rtac lub_mono 1),
- (atac 1),
- (atac 1),
- (strip_tac 1),
- (rtac (antisym_less_inverse RS conjunct1) 1),
- (etac spec 1),
- (rtac lub_mono 1),
- (atac 1),
- (atac 1),
- (strip_tac 1),
- (rtac (antisym_less_inverse RS conjunct2) 1),
- (etac spec 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* more results about mono and = of lubs of chains *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_mono2 = prove_goal Pcpo.thy
-"[|? j.!i. j<i --> X(i::nat)=Y(i);is_chain(X::nat=>'a::pcpo);is_chain(Y)|]\
-\ ==> lub(range(X))<<lub(range(Y))"
- (fn prems =>
- [
- (rtac exE 1),
- (resolve_tac prems 1),
- (rtac is_lub_thelub 1),
- (resolve_tac prems 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (res_inst_tac [("Q","x<i")] classical2 1),
- (res_inst_tac [("s","Y(i)"),("t","X(i)")] subst 1),
- (rtac sym 1),
- (fast_tac HOL_cs 1),
- (rtac is_ub_thelub 1),
- (resolve_tac prems 1),
- (res_inst_tac [("y","X(Suc(x))")] trans_less 1),
- (rtac (chain_mono RS mp) 1),
- (resolve_tac prems 1),
- (rtac (not_less_eq RS subst) 1),
- (atac 1),
- (res_inst_tac [("s","Y(Suc(x))"),("t","X(Suc(x))")] subst 1),
- (rtac sym 1),
- (asm_simp_tac nat_ss 1),
- (rtac is_ub_thelub 1),
- (resolve_tac prems 1)
- ]);
-
-val lub_equal2 = prove_goal Pcpo.thy
-"[|? j.!i. j<i --> X(i)=Y(i);is_chain(X::nat=>'a::pcpo);is_chain(Y)|]\
-\ ==> lub(range(X))=lub(range(Y))"
- (fn prems =>
- [
- (rtac antisym_less 1),
- (rtac lub_mono2 1),
- (REPEAT (resolve_tac prems 1)),
- (cut_facts_tac prems 1),
- (rtac lub_mono2 1),
- (safe_tac HOL_cs),
- (step_tac HOL_cs 1),
- (safe_tac HOL_cs),
- (rtac sym 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val lub_mono3 = prove_goal Pcpo.thy "[|is_chain(Y::nat=>'a::pcpo);is_chain(X);\
-\! i. ? j. Y(i)<< X(j)|]==> lub(range(Y))<<lub(range(X))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lub_thelub 1),
- (atac 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (etac allE 1),
- (etac exE 1),
- (rtac trans_less 1),
- (rtac is_ub_thelub 2),
- (atac 2),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* usefull lemmas about UU *)
-(* ------------------------------------------------------------------------ *)
-
-val eq_UU_iff = prove_goal Pcpo.thy "(x=UU)=(x<<UU)"
- (fn prems =>
- [
- (rtac iffI 1),
- (hyp_subst_tac 1),
- (rtac refl_less 1),
- (rtac antisym_less 1),
- (atac 1),
- (rtac minimal 1)
- ]);
-
-val UU_I = prove_goal Pcpo.thy "x << UU ==> x = UU"
- (fn prems =>
- [
- (rtac (eq_UU_iff RS ssubst) 1),
- (resolve_tac prems 1)
- ]);
-
-val not_less2not_eq = prove_goal Pcpo.thy "~x<<y ==> ~x=y"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac classical3 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac refl_less 1)
- ]);
-
-
-val chain_UU_I = prove_goal Pcpo.thy
- "[|is_chain(Y);lub(range(Y))=UU|] ==> ! i.Y(i)=UU"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac allI 1),
- (rtac antisym_less 1),
- (rtac minimal 2),
- (res_inst_tac [("t","UU")] subst 1),
- (atac 1),
- (etac is_ub_thelub 1)
- ]);
-
-
-val chain_UU_I_inverse = prove_goal Pcpo.thy
- "!i.Y(i::nat)=UU ==> lub(range(Y::(nat=>'a::pcpo)))=UU"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac lub_chain_maxelem 1),
- (rtac is_chainI 1),
- (rtac allI 1),
- (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1),
- (rtac sym 1),
- (etac spec 1),
- (rtac minimal 1),
- (rtac exI 1),
- (etac spec 1),
- (rtac allI 1),
- (rtac (antisym_less_inverse RS conjunct1) 1),
- (etac spec 1)
- ]);
-
-val chain_UU_I_inverse2 = prove_goal Pcpo.thy
- "~lub(range(Y::(nat=>'a::pcpo)))=UU ==> ? i.~ Y(i)=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (notall2ex RS iffD1) 1),
- (rtac swap 1),
- (rtac chain_UU_I_inverse 2),
- (etac notnotD 2),
- (atac 1)
- ]);
-
-
-val notUU_I = prove_goal Pcpo.thy "[| x<<y; ~x=UU |] ==> ~y=UU"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac contrapos 1),
- (rtac UU_I 1),
- (hyp_subst_tac 1),
- (atac 1)
- ]);
-
-
-val chain_mono2 = prove_goal Pcpo.thy
-"[|? j.~Y(j)=UU;is_chain(Y::nat=>'a::pcpo)|]\
-\ ==> ? j.!i.j<i-->~Y(i)=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (safe_tac HOL_cs),
- (step_tac HOL_cs 1),
- (strip_tac 1),
- (rtac notUU_I 1),
- (atac 2),
- (etac (chain_mono RS mp) 1),
- (atac 1)
- ]);
-
-
-
-
-(* ------------------------------------------------------------------------ *)
-(* uniqueness in void *)
-(* ------------------------------------------------------------------------ *)
-
-val unique_void2 = prove_goal Pcpo.thy "x::void=UU"
- (fn prems =>
- [
- (rtac (inst_void_pcpo RS ssubst) 1),
- (rtac (Rep_Void_inverse RS subst) 1),
- (rtac (Rep_Void_inverse RS subst) 1),
- (rtac arg_cong 1),
- (rtac box_equals 1),
- (rtac refl 1),
- (rtac (unique_void RS sym) 1),
- (rtac (unique_void RS sym) 1)
- ]);
-
-
--- a/src/HOLCF/pcpo.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,39 +0,0 @@
-(* Title: HOLCF/pcpo.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Definition of class pcpo (pointed complete partial order)
-
-The prototype theory for this class is porder.thy
-
-*)
-
-Pcpo = Porder +
-
-(* Introduction of new class. The witness is type void. *)
-
-classes pcpo < po
-
-(* default class is still term *)
-(* void is the prototype in pcpo *)
-
-arities void :: pcpo
-
-consts
- UU :: "'a::pcpo" (* UU is the least element *)
-rules
-
-(* class axioms: justification is theory Porder *)
-
-minimal "UU << x" (* witness is minimal_void *)
-
-cpo "is_chain(S) ==> ? x. range(S) <<| x::('a::pcpo)"
- (* witness is cpo_void *)
- (* needs explicit type *)
-
-(* instance of UU for the prototype void *)
-
-inst_void_pcpo "UU::void = UU_void"
-
-end
--- a/src/HOLCF/porder.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,427 +0,0 @@
-(* Title: HOLCF/porder.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for theory porder.thy
-*)
-
-open Porder0;
-open Porder;
-
-val box_less = prove_goal Porder.thy
-"[| a << b; c << a; b << d|] ==> c << d"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac trans_less 1),
- (etac trans_less 1),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* lubs are unique *)
-(* ------------------------------------------------------------------------ *)
-
-val unique_lub = prove_goalw Porder.thy [is_lub, is_ub]
- "[| S <<| x ; S <<| y |] ==> x=y"
-( fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac conjE 1),
- (etac conjE 1),
- (rtac antisym_less 1),
- (rtac mp 1),((etac allE 1) THEN (atac 1) THEN (atac 1)),
- (rtac mp 1),((etac allE 1) THEN (atac 1) THEN (atac 1))
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* chains are monotone functions *)
-(* ------------------------------------------------------------------------ *)
-
-val chain_mono = prove_goalw Porder.thy [is_chain]
- " is_chain(F) ==> x<y --> F(x)<<F(y)"
-( fn prems =>
- [
- (cut_facts_tac prems 1),
- (nat_ind_tac "y" 1),
- (rtac impI 1),
- (etac less_zeroE 1),
- (rtac (less_Suc_eq RS ssubst) 1),
- (strip_tac 1),
- (etac disjE 1),
- (rtac trans_less 1),
- (etac allE 2),
- (atac 2),
- (fast_tac HOL_cs 1),
- (hyp_subst_tac 1),
- (etac allE 1),
- (atac 1)
- ]);
-
-val chain_mono3 = prove_goal Porder.thy
- "[| is_chain(F); x <= y |] ==> F(x) << F(y)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (le_imp_less_or_eq RS disjE) 1),
- (atac 1),
- (etac (chain_mono RS mp) 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac refl_less 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Lemma for reasoning by cases on the natural numbers *)
-(* ------------------------------------------------------------------------ *)
-
-val nat_less_cases = prove_goal Porder.thy
- "[| m::nat < n ==> P(n,m); m=n ==> P(n,m);n < m ==> P(n,m)|]==>P(n,m)"
-( fn prems =>
- [
- (res_inst_tac [("m1","n"),("n1","m")] (less_linear RS disjE) 1),
- (etac disjE 2),
- (etac (hd (tl (tl prems))) 1),
- (etac (sym RS hd (tl prems)) 1),
- (etac (hd prems) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* The range of a chain is a totaly ordered << *)
-(* ------------------------------------------------------------------------ *)
-
-val chain_is_tord = prove_goalw Porder.thy [is_tord]
- "is_chain(F) ==> is_tord(range(F))"
-( fn prems =>
- [
- (cut_facts_tac prems 1),
- (rewrite_goals_tac [range_def]),
- (rtac allI 1 ),
- (rtac allI 1 ),
- (rtac (mem_Collect_eq RS ssubst) 1),
- (rtac (mem_Collect_eq RS ssubst) 1),
- (strip_tac 1),
- (etac conjE 1),
- (etac exE 1),
- (etac exE 1),
- (hyp_subst_tac 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("m","xa"),("n","xb")] (nat_less_cases) 1),
- (rtac disjI1 1),
- (rtac (chain_mono RS mp) 1),
- (atac 1),
- (atac 1),
- (rtac disjI1 1),
- (hyp_subst_tac 1),
- (rtac refl_less 1),
- (rtac disjI2 1),
- (rtac (chain_mono RS mp) 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* technical lemmas about lub and is_lub, use above results about @ *)
-(* ------------------------------------------------------------------------ *)
-
-val lubI = prove_goal Porder.thy "(? x. M <<| x) ==> M <<| lub(M)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (lub RS ssubst) 1),
- (etac selectI2 1)
- ]);
-
-val lubE = prove_goal Porder.thy " M <<| lub(M) ==> ? x. M <<| x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exI 1)
- ]);
-
-val lub_eq = prove_goal Porder.thy
- "(? x. M <<| x) = M <<| lub(M)"
-(fn prems =>
- [
- (rtac (lub RS ssubst) 1),
- (rtac (select_eq_Ex RS subst) 1),
- (rtac refl 1)
- ]);
-
-
-val thelubI = prove_goal Porder.thy " M <<| l ==> lub(M) = l"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac unique_lub 1),
- (rtac (lub RS ssubst) 1),
- (etac selectI 1),
- (atac 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* access to some definition as inference rule *)
-(* ------------------------------------------------------------------------ *)
-
-val is_lubE = prove_goalw Porder.thy [is_lub]
- "S <<| x ==> S <| x & (! u. S <| u --> x << u)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-val is_lubI = prove_goalw Porder.thy [is_lub]
- "S <| x & (! u. S <| u --> x << u) ==> S <<| x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)
- ]);
-
-val is_chainE = prove_goalw Porder.thy [is_chain]
- "is_chain(F) ==> ! i. F(i) << F(Suc(i))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)]);
-
-val is_chainI = prove_goalw Porder.thy [is_chain]
- "! i. F(i) << F(Suc(i)) ==> is_chain(F) "
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (atac 1)]);
-
-(* ------------------------------------------------------------------------ *)
-(* technical lemmas about (least) upper bounds of chains *)
-(* ------------------------------------------------------------------------ *)
-
-val ub_rangeE = prove_goalw Porder.thy [is_ub]
- "range(S) <| x ==> ! i. S(i) << x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (rtac mp 1),
- (etac spec 1),
- (rtac rangeI 1)
- ]);
-
-val ub_rangeI = prove_goalw Porder.thy [is_ub]
- "! i. S(i) << x ==> range(S) <| x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (etac rangeE 1),
- (hyp_subst_tac 1),
- (etac spec 1)
- ]);
-
-val is_ub_lub = (is_lubE RS conjunct1 RS ub_rangeE RS spec);
-(* range(?S1) <<| ?x1 ==> ?S1(?x) << ?x1 *)
-
-val is_lub_lub = (is_lubE RS conjunct2 RS spec RS mp);
-(* [| ?S3 <<| ?x3; ?S3 <| ?x1 |] ==> ?x3 << ?x1 *)
-
-(* ------------------------------------------------------------------------ *)
-(* Prototype lemmas for class pcpo *)
-(* ------------------------------------------------------------------------ *)
-
-(* ------------------------------------------------------------------------ *)
-(* a technical argument about << on void *)
-(* ------------------------------------------------------------------------ *)
-
-val less_void = prove_goal Porder.thy "(u1::void << u2) = (u1 = u2)"
-(fn prems =>
- [
- (rtac (inst_void_po RS ssubst) 1),
- (rewrite_goals_tac [less_void_def]),
- (rtac iffI 1),
- (rtac injD 1),
- (atac 2),
- (rtac inj_inverseI 1),
- (rtac Rep_Void_inverse 1),
- (etac arg_cong 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* void is pointed. The least element is UU_void *)
-(* ------------------------------------------------------------------------ *)
-
-val minimal_void = prove_goal Porder.thy "UU_void << x"
-(fn prems =>
- [
- (rtac (inst_void_po RS ssubst) 1),
- (rewrite_goals_tac [less_void_def]),
- (simp_tac (HOL_ss addsimps [unique_void]) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* UU_void is the trivial lub of all chains in void *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_void = prove_goalw Porder.thy [is_lub] "M <<| UU_void"
-(fn prems =>
- [
- (rtac conjI 1),
- (rewrite_goals_tac [is_ub]),
- (strip_tac 1),
- (rtac (inst_void_po RS ssubst) 1),
- (rewrite_goals_tac [less_void_def]),
- (simp_tac (HOL_ss addsimps [unique_void]) 1),
- (strip_tac 1),
- (rtac minimal_void 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* lub(?M) = UU_void *)
-(* ------------------------------------------------------------------------ *)
-
-val thelub_void = (lub_void RS thelubI);
-
-(* ------------------------------------------------------------------------ *)
-(* void is a cpo wrt. countable chains *)
-(* ------------------------------------------------------------------------ *)
-
-val cpo_void = prove_goal Porder.thy
- "is_chain(S::nat=>void) ==> ? x. range(S) <<| x "
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("x","UU_void")] exI 1),
- (rtac lub_void 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* end of prototype lemmas for class pcpo *)
-(* ------------------------------------------------------------------------ *)
-
-
-(* ------------------------------------------------------------------------ *)
-(* the reverse law of anti--symmetrie of << *)
-(* ------------------------------------------------------------------------ *)
-
-val antisym_less_inverse = prove_goal Porder.thy "x=y ==> x << y & y << x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac conjI 1),
- ((rtac subst 1) THEN (rtac refl_less 2) THEN (atac 1)),
- ((rtac subst 1) THEN (rtac refl_less 2) THEN (etac sym 1))
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* results about finite chains *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_finch1 = prove_goalw Porder.thy [max_in_chain_def]
- "[| is_chain(C) ; max_in_chain(i,C)|] ==> range(C) <<| C(i)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (res_inst_tac [("m","i")] nat_less_cases 1),
- (rtac (antisym_less_inverse RS conjunct2) 1),
- (etac (disjI1 RS less_or_eq_imp_le RS rev_mp) 1),
- (etac spec 1),
- (rtac (antisym_less_inverse RS conjunct2) 1),
- (etac (disjI2 RS less_or_eq_imp_le RS rev_mp) 1),
- (etac spec 1),
- (etac (chain_mono RS mp) 1),
- (atac 1),
- (strip_tac 1),
- (etac (ub_rangeE RS spec) 1)
- ]);
-
-val lub_finch2 = prove_goalw Porder.thy [finite_chain_def]
- "finite_chain(C) ==> range(C) <<| C(@ i. max_in_chain(i,C))"
- (fn prems=>
- [
- (cut_facts_tac prems 1),
- (rtac lub_finch1 1),
- (etac conjunct1 1),
- (rtac selectI2 1),
- (etac conjunct2 1)
- ]);
-
-
-val bin_chain = prove_goal Porder.thy "x<<y ==> is_chain(%i. if(i=0,x,y))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_chainI 1),
- (rtac allI 1),
- (nat_ind_tac "i" 1),
- (asm_simp_tac nat_ss 1),
- (asm_simp_tac nat_ss 1),
- (rtac refl_less 1)
- ]);
-
-val bin_chainmax = prove_goalw Porder.thy [max_in_chain_def,le_def]
- "x<<y ==> max_in_chain(Suc(0),%i. if(i=0,x,y))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac allI 1),
- (nat_ind_tac "j" 1),
- (asm_simp_tac nat_ss 1),
- (asm_simp_tac nat_ss 1)
- ]);
-
-val lub_bin_chain = prove_goal Porder.thy
- "x << y ==> range(%i. if(i = 0,x,y)) <<| y"
-(fn prems=>
- [ (cut_facts_tac prems 1),
- (res_inst_tac [("s","if(Suc(0) = 0,x,y)")] subst 1),
- (rtac lub_finch1 2),
- (etac bin_chain 2),
- (etac bin_chainmax 2),
- (simp_tac nat_ss 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the maximal element in a chain is its lub *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_chain_maxelem = prove_goal Porder.thy
-"[|is_chain(Y);? i.Y(i)=c;!i.Y(i)<<c|] ==> lub(range(Y)) = c"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac thelubI 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (etac ub_rangeI 1),
- (strip_tac 1),
- (res_inst_tac [("P","%i.Y(i)=c")] exE 1),
- (atac 1),
- (hyp_subst_tac 1),
- (etac (ub_rangeE RS spec) 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the lub of a constant chain is the constant *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_const = prove_goal Porder.thy "range(%x.c) <<| c"
- (fn prems =>
- [
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (strip_tac 1),
- (rtac refl_less 1),
- (strip_tac 1),
- (etac (ub_rangeE RS spec) 1)
- ]);
-
-
-
--- a/src/HOLCF/porder.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,42 +0,0 @@
-(* Title: HOLCF/porder.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Conservative extension of theory Porder0 by constant definitions
-
-*)
-
-Porder = Porder0 +
-
-consts
- "<|" :: "['a set,'a::po] => bool" (infixl 55)
- "<<|" :: "['a set,'a::po] => bool" (infixl 55)
- lub :: "'a set => 'a::po"
- is_tord :: "'a::po set => bool"
- is_chain :: "(nat=>'a::po) => bool"
- max_in_chain :: "[nat,nat=>'a::po]=>bool"
- finite_chain :: "(nat=>'a::po)=>bool"
-
-rules
-
-(* class definitions *)
-
-is_ub "S <| x == ! y.y:S --> y<<x"
-is_lub "S <<| x == S <| x & (! u. S <| u --> x << u)"
-
-lub "lub(S) = (@x. S <<| x)"
-
-(* Arbitrary chains are total orders *)
-is_tord "is_tord(S) == ! x y. x:S & y:S --> (x<<y | y<<x)"
-
-(* Here we use countable chains and I prefer to code them as functions! *)
-is_chain "is_chain(F) == (! i.F(i) << F(Suc(i)))"
-
-(* finite chains, needed for monotony of continouous functions *)
-
-max_in_chain_def "max_in_chain(i,C) == ! j. i <= j --> C(i) = C(j)"
-
-finite_chain_def "finite_chain(C) == is_chain(C) & (? i. max_in_chain(i,C))"
-
-end
--- a/src/HOLCF/porder0.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,42 +0,0 @@
-(* Title: HOLCF/porder0.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Definition of class porder (partial order)
-
-The prototype theory for this class is void.thy
-
-*)
-
-Porder0 = Void +
-
-(* Introduction of new class. The witness is type void. *)
-
-classes po < term
-
-(* default type is still term ! *)
-(* void is the prototype in po *)
-
-arities void :: po
-
-consts "<<" :: "['a,'a::po] => bool" (infixl 55)
-
-rules
-
-(* class axioms: justification is theory Void *)
-
-refl_less "x << x"
- (* witness refl_less_void *)
-
-antisym_less "[|x<<y ; y<<x |] ==> x = y"
- (* witness antisym_less_void *)
-
-trans_less "[|x<<y ; y<<z |] ==> x<<z"
- (* witness trans_less_void *)
-
-(* instance of << for the prototype void *)
-
-inst_void_po "(op <<)::[void,void]=>bool = less_void"
-
-end
--- a/src/HOLCF/sprod0.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,389 +0,0 @@
-(* Title: HOLCF/sprod0.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for theory sprod0.thy
-*)
-
-open Sprod0;
-
-(* ------------------------------------------------------------------------ *)
-(* A non-emptyness result for Sprod *)
-(* ------------------------------------------------------------------------ *)
-
-val SprodI = prove_goalw Sprod0.thy [Sprod_def]
- "Spair_Rep(a,b):Sprod"
-(fn prems =>
- [
- (EVERY1 [rtac CollectI, rtac exI,rtac exI, rtac refl])
- ]);
-
-
-val inj_onto_Abs_Sprod = prove_goal Sprod0.thy
- "inj_onto(Abs_Sprod,Sprod)"
-(fn prems =>
- [
- (rtac inj_onto_inverseI 1),
- (etac Abs_Sprod_inverse 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Strictness and definedness of Spair_Rep *)
-(* ------------------------------------------------------------------------ *)
-
-
-val strict_Spair_Rep = prove_goalw Sprod0.thy [Spair_Rep_def]
- "(a=UU | b=UU) ==> (Spair_Rep(a,b) = Spair_Rep(UU,UU))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac ext 1),
- (rtac ext 1),
- (rtac iffI 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val defined_Spair_Rep_rev = prove_goalw Sprod0.thy [Spair_Rep_def]
- "(Spair_Rep(a,b) = Spair_Rep(UU,UU)) ==> (a=UU | b=UU)"
- (fn prems =>
- [
- (res_inst_tac [("Q","a=UU|b=UU")] classical2 1),
- (atac 1),
- (rtac disjI1 1),
- (rtac ((hd prems) RS fun_cong RS fun_cong RS iffD2 RS mp RS
- conjunct1 RS sym) 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* injectivity of Spair_Rep and Ispair *)
-(* ------------------------------------------------------------------------ *)
-
-val inject_Spair_Rep = prove_goalw Sprod0.thy [Spair_Rep_def]
-"[|~aa=UU ; ~ba=UU ; Spair_Rep(a,b)=Spair_Rep(aa,ba) |] ==> a=aa & b=ba"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong
- RS iffD1 RS mp) 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val inject_Ispair = prove_goalw Sprod0.thy [Ispair_def]
- "[|~aa=UU ; ~ba=UU ; Ispair(a,b)=Ispair(aa,ba) |] ==> a=aa & b=ba"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac inject_Spair_Rep 1),
- (atac 1),
- (etac (inj_onto_Abs_Sprod RS inj_ontoD) 1),
- (rtac SprodI 1),
- (rtac SprodI 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* strictness and definedness of Ispair *)
-(* ------------------------------------------------------------------------ *)
-
-val strict_Ispair = prove_goalw Sprod0.thy [Ispair_def]
- "(a=UU | b=UU) ==> Ispair(a,b)=Ispair(UU,UU)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac (strict_Spair_Rep RS arg_cong) 1)
- ]);
-
-val strict_Ispair1 = prove_goalw Sprod0.thy [Ispair_def]
- "Ispair(UU,b) = Ispair(UU,UU)"
-(fn prems =>
- [
- (rtac (strict_Spair_Rep RS arg_cong) 1),
- (rtac disjI1 1),
- (rtac refl 1)
- ]);
-
-val strict_Ispair2 = prove_goalw Sprod0.thy [Ispair_def]
- "Ispair(a,UU) = Ispair(UU,UU)"
-(fn prems =>
- [
- (rtac (strict_Spair_Rep RS arg_cong) 1),
- (rtac disjI2 1),
- (rtac refl 1)
- ]);
-
-val strict_Ispair_rev = prove_goal Sprod0.thy
- "~Ispair(x,y)=Ispair(UU,UU) ==> ~x=UU & ~y=UU"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (de_morgan1 RS ssubst) 1),
- (etac contrapos 1),
- (etac strict_Ispair 1)
- ]);
-
-val defined_Ispair_rev = prove_goalw Sprod0.thy [Ispair_def]
- "Ispair(a,b) = Ispair(UU,UU) ==> (a = UU | b = UU)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac defined_Spair_Rep_rev 1),
- (rtac (inj_onto_Abs_Sprod RS inj_ontoD) 1),
- (atac 1),
- (rtac SprodI 1),
- (rtac SprodI 1)
- ]);
-
-val defined_Ispair = prove_goal Sprod0.thy
-"[|~a=UU; ~b=UU|] ==> ~(Ispair(a,b) = Ispair(UU,UU))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contrapos 1),
- (etac defined_Ispair_rev 2),
- (rtac (de_morgan1 RS iffD1) 1),
- (etac conjI 1),
- (atac 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Exhaustion of the strict product ** *)
-(* ------------------------------------------------------------------------ *)
-
-val Exh_Sprod = prove_goalw Sprod0.thy [Ispair_def]
- "z=Ispair(UU,UU) | (? a b. z=Ispair(a,b) & ~a=UU & ~b=UU)"
-(fn prems =>
- [
- (rtac (rewrite_rule [Sprod_def] Rep_Sprod RS CollectE) 1),
- (etac exE 1),
- (etac exE 1),
- (rtac (excluded_middle RS disjE) 1),
- (rtac disjI2 1),
- (rtac exI 1),
- (rtac exI 1),
- (rtac conjI 1),
- (rtac (Rep_Sprod_inverse RS sym RS trans) 1),
- (etac arg_cong 1),
- (rtac (de_morgan1 RS ssubst) 1),
- (atac 1),
- (rtac disjI1 1),
- (rtac (Rep_Sprod_inverse RS sym RS trans) 1),
- (res_inst_tac [("f","Abs_Sprod")] arg_cong 1),
- (etac trans 1),
- (etac strict_Spair_Rep 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* general elimination rule for strict product *)
-(* ------------------------------------------------------------------------ *)
-
-val IsprodE = prove_goal Sprod0.thy
-"[|p=Ispair(UU,UU) ==> Q ;!!x y. [|p=Ispair(x,y); ~x=UU ; ~y=UU|] ==> Q|] ==> Q"
-(fn prems =>
- [
- (rtac (Exh_Sprod RS disjE) 1),
- (etac (hd prems) 1),
- (etac exE 1),
- (etac exE 1),
- (etac conjE 1),
- (etac conjE 1),
- (etac (hd (tl prems)) 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* some results about the selectors Isfst, Issnd *)
-(* ------------------------------------------------------------------------ *)
-
-val strict_Isfst = prove_goalw Sprod0.thy [Isfst_def]
- "p=Ispair(UU,UU)==>Isfst(p)=UU"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (rtac conjI 1),
- (fast_tac HOL_cs 1),
- (strip_tac 1),
- (res_inst_tac [("P","Ispair(UU,UU) = Ispair(a,b)")] notE 1),
- (rtac not_sym 1),
- (rtac defined_Ispair 1),
- (REPEAT (fast_tac HOL_cs 1))
- ]);
-
-
-val strict_Isfst1 = prove_goal Sprod0.thy
- "Isfst(Ispair(UU,y)) = UU"
-(fn prems =>
- [
- (rtac (strict_Ispair1 RS ssubst) 1),
- (rtac strict_Isfst 1),
- (rtac refl 1)
- ]);
-
-val strict_Isfst2 = prove_goal Sprod0.thy
- "Isfst(Ispair(x,UU)) = UU"
-(fn prems =>
- [
- (rtac (strict_Ispair2 RS ssubst) 1),
- (rtac strict_Isfst 1),
- (rtac refl 1)
- ]);
-
-
-val strict_Issnd = prove_goalw Sprod0.thy [Issnd_def]
- "p=Ispair(UU,UU)==>Issnd(p)=UU"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (rtac conjI 1),
- (fast_tac HOL_cs 1),
- (strip_tac 1),
- (res_inst_tac [("P","Ispair(UU,UU) = Ispair(a,b)")] notE 1),
- (rtac not_sym 1),
- (rtac defined_Ispair 1),
- (REPEAT (fast_tac HOL_cs 1))
- ]);
-
-val strict_Issnd1 = prove_goal Sprod0.thy
- "Issnd(Ispair(UU,y)) = UU"
-(fn prems =>
- [
- (rtac (strict_Ispair1 RS ssubst) 1),
- (rtac strict_Issnd 1),
- (rtac refl 1)
- ]);
-
-val strict_Issnd2 = prove_goal Sprod0.thy
- "Issnd(Ispair(x,UU)) = UU"
-(fn prems =>
- [
- (rtac (strict_Ispair2 RS ssubst) 1),
- (rtac strict_Issnd 1),
- (rtac refl 1)
- ]);
-
-val Isfst = prove_goalw Sprod0.thy [Isfst_def]
- "[|~x=UU ;~y=UU |] ==> Isfst(Ispair(x,y)) = x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (rtac conjI 1),
- (strip_tac 1),
- (res_inst_tac [("P","Ispair(x,y) = Ispair(UU,UU)")] notE 1),
- (etac defined_Ispair 1),
- (atac 1),
- (atac 1),
- (strip_tac 1),
- (rtac (inject_Ispair RS conjunct1) 1),
- (fast_tac HOL_cs 3),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val Issnd = prove_goalw Sprod0.thy [Issnd_def]
- "[|~x=UU ;~y=UU |] ==> Issnd(Ispair(x,y)) = y"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (rtac conjI 1),
- (strip_tac 1),
- (res_inst_tac [("P","Ispair(x,y) = Ispair(UU,UU)")] notE 1),
- (etac defined_Ispair 1),
- (atac 1),
- (atac 1),
- (strip_tac 1),
- (rtac (inject_Ispair RS conjunct2) 1),
- (fast_tac HOL_cs 3),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val Isfst2 = prove_goal Sprod0.thy "~y=UU ==>Isfst(Ispair(x,y))=x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1),
- (etac Isfst 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac strict_Isfst1 1)
- ]);
-
-val Issnd2 = prove_goal Sprod0.thy "~x=UU ==>Issnd(Ispair(x,y))=y"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("Q","y=UU")] (excluded_middle RS disjE) 1),
- (etac Issnd 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac strict_Issnd2 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* instantiate the simplifier *)
-(* ------------------------------------------------------------------------ *)
-
-val Sprod_ss =
- HOL_ss
- addsimps [strict_Isfst1,strict_Isfst2,strict_Issnd1,strict_Issnd2,
- Isfst2,Issnd2];
-
-
-val defined_IsfstIssnd = prove_goal Sprod0.thy
- "~p=Ispair(UU,UU) ==> ~Isfst(p)=UU & ~Issnd(p)=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p")] IsprodE 1),
- (contr_tac 1),
- (hyp_subst_tac 1),
- (rtac conjI 1),
- (asm_simp_tac Sprod_ss 1),
- (asm_simp_tac Sprod_ss 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Surjective pairing: equivalent to Exh_Sprod *)
-(* ------------------------------------------------------------------------ *)
-
-val surjective_pairing_Sprod = prove_goal Sprod0.thy
- "z = Ispair(Isfst(z))(Issnd(z))"
-(fn prems =>
- [
- (res_inst_tac [("z1","z")] (Exh_Sprod RS disjE) 1),
- (asm_simp_tac Sprod_ss 1),
- (etac exE 1),
- (etac exE 1),
- (asm_simp_tac Sprod_ss 1)
- ]);
-
-
-
-
-
-
-
-
-
-
-
-
-
--- a/src/HOLCF/sprod0.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,53 +0,0 @@
-(* Title: HOLCF/sprod0.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Strict product
-*)
-
-Sprod0 = Cfun3 +
-
-(* new type for strict product *)
-
-types "**" 2 (infixr 20)
-
-arities "**" :: (pcpo,pcpo)term
-
-consts
- Sprod :: "('a => 'b => bool)set"
- Spair_Rep :: "['a,'b] => ['a,'b] => bool"
- Rep_Sprod :: "('a ** 'b) => ('a => 'b => bool)"
- Abs_Sprod :: "('a => 'b => bool) => ('a ** 'b)"
- Ispair :: "['a,'b] => ('a ** 'b)"
- Isfst :: "('a ** 'b) => 'a"
- Issnd :: "('a ** 'b) => 'b"
-
-rules
-
- Spair_Rep_def "Spair_Rep == (%a b. %x y.\
-\ (~a=UU & ~b=UU --> x=a & y=b ))"
-
- Sprod_def "Sprod == {f. ? a b. f = Spair_Rep(a,b)}"
-
- (*faking a type definition... *)
- (* "**" is isomorphic to Sprod *)
-
- Rep_Sprod "Rep_Sprod(p):Sprod"
- Rep_Sprod_inverse "Abs_Sprod(Rep_Sprod(p)) = p"
- Abs_Sprod_inverse "f:Sprod ==> Rep_Sprod(Abs_Sprod(f)) = f"
-
- (*defining the abstract constants*)
-
- Ispair_def "Ispair(a,b) == Abs_Sprod(Spair_Rep(a,b))"
-
- Isfst_def "Isfst(p) == @z.\
-\ (p=Ispair(UU,UU) --> z=UU)\
-\ &(! a b. ~a=UU & ~b=UU & p=Ispair(a,b) --> z=a)"
-
- Issnd_def "Issnd(p) == @z.\
-\ (p=Ispair(UU,UU) --> z=UU)\
-\ &(! a b. ~a=UU & ~b=UU & p=Ispair(a,b) --> z=b)"
-
-end
-
--- a/src/HOLCF/sprod1.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,204 +0,0 @@
-(* Title: HOLCF/sprod1.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for theory sprod1.thy
-*)
-
-open Sprod1;
-
-(* ------------------------------------------------------------------------ *)
-(* reduction properties for less_sprod *)
-(* ------------------------------------------------------------------------ *)
-
-
-val less_sprod1a = prove_goalw Sprod1.thy [less_sprod_def]
- "p1=Ispair(UU,UU) ==> less_sprod(p1,p2)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac eqTrueE 1),
- (rtac select_equality 1),
- (rtac conjI 1),
- (fast_tac HOL_cs 1),
- (strip_tac 1),
- (contr_tac 1),
- (dtac conjunct1 1),
- (etac rev_mp 1),
- (atac 1)
- ]);
-
-val less_sprod1b = prove_goalw Sprod1.thy [less_sprod_def]
- "~p1=Ispair(UU,UU) ==> \
-\ less_sprod(p1,p2) = ( Isfst(p1) << Isfst(p2) & Issnd(p1) << Issnd(p2))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (rtac conjI 1),
- (strip_tac 1),
- (contr_tac 1),
- (fast_tac HOL_cs 1),
- (dtac conjunct2 1),
- (etac rev_mp 1),
- (atac 1)
- ]);
-
-val less_sprod2a = prove_goal Sprod1.thy
- "less_sprod(Ispair(x,y),Ispair(UU,UU)) ==> x = UU | y = UU"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (excluded_middle RS disjE) 1),
- (atac 2),
- (rtac disjI1 1),
- (rtac antisym_less 1),
- (rtac minimal 2),
- (res_inst_tac [("s","Isfst(Ispair(x,y))"),("t","x")] subst 1),
- (rtac Isfst 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1),
- (res_inst_tac [("s","Isfst(Ispair(UU,UU))"),("t","UU")] subst 1),
- (simp_tac Sprod_ss 1),
- (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct1) 1),
- (REPEAT (fast_tac HOL_cs 1))
- ]);
-
-val less_sprod2b = prove_goal Sprod1.thy
- "less_sprod(p,Ispair(UU,UU)) ==> p = Ispair(UU,UU)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p")] IsprodE 1),
- (atac 1),
- (hyp_subst_tac 1),
- (rtac strict_Ispair 1),
- (etac less_sprod2a 1)
- ]);
-
-val less_sprod2c = prove_goal Sprod1.thy
- "[|less_sprod(Ispair(xa,ya),Ispair(x,y));\
-\~ xa = UU ; ~ ya = UU;~ x = UU ; ~ y = UU |] ==> xa << x & ya << y"
-(fn prems =>
- [
- (rtac conjI 1),
- (res_inst_tac [("s","Isfst(Ispair(xa,ya))"),("t","xa")] subst 1),
- (simp_tac (Sprod_ss addsimps prems)1),
- (res_inst_tac [("s","Isfst(Ispair(x,y))"),("t","x")] subst 1),
- (simp_tac (Sprod_ss addsimps prems)1),
- (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct1) 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (simp_tac (Sprod_ss addsimps prems)1),
- (res_inst_tac [("s","Issnd(Ispair(xa,ya))"),("t","ya")] subst 1),
- (simp_tac (Sprod_ss addsimps prems)1),
- (res_inst_tac [("s","Issnd(Ispair(x,y))"),("t","y")] subst 1),
- (simp_tac (Sprod_ss addsimps prems)1),
- (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct2) 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (simp_tac (Sprod_ss addsimps prems)1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* less_sprod is a partial order on Sprod *)
-(* ------------------------------------------------------------------------ *)
-
-val refl_less_sprod = prove_goal Sprod1.thy "less_sprod(p,p)"
-(fn prems =>
- [
- (res_inst_tac [("p","p")] IsprodE 1),
- (etac less_sprod1a 1),
- (hyp_subst_tac 1),
- (rtac (less_sprod1b RS ssubst) 1),
- (rtac defined_Ispair 1),
- (REPEAT (fast_tac (HOL_cs addIs [refl_less]) 1))
- ]);
-
-
-val antisym_less_sprod = prove_goal Sprod1.thy
- "[|less_sprod(p1,p2);less_sprod(p2,p1)|] ==> p1=p2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] IsprodE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] IsprodE 1),
- (hyp_subst_tac 1),
- (rtac refl 1),
- (hyp_subst_tac 1),
- (rtac (strict_Ispair RS sym) 1),
- (etac less_sprod2a 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] IsprodE 1),
- (hyp_subst_tac 1),
- (rtac (strict_Ispair) 1),
- (etac less_sprod2a 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("x1","x"),("y1","xa"),("x","y"),("y","ya")] (arg_cong RS cong) 1),
- (rtac antisym_less 1),
- (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct1]) 1),
- (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct1]) 1),
- (rtac antisym_less 1),
- (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct2]) 1),
- (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct2]) 1)
- ]);
-
-val trans_less_sprod = prove_goal Sprod1.thy
- "[|less_sprod(p1,p2);less_sprod(p2,p3)|] ==> less_sprod(p1,p3)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] IsprodE 1),
- (etac less_sprod1a 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p3")] IsprodE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("s","p2"),("t","Ispair(UU,UU)")] subst 1),
- (etac less_sprod2b 1),
- (atac 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("Q","p2=Ispair(UU,UU)")]
- (excluded_middle RS disjE) 1),
- (rtac (defined_Ispair RS less_sprod1b RS ssubst) 1),
- (atac 1),
- (atac 1),
- (rtac conjI 1),
- (res_inst_tac [("y","Isfst(p2)")] trans_less 1),
- (rtac conjunct1 1),
- (rtac (less_sprod1b RS subst) 1),
- (rtac defined_Ispair 1),
- (atac 1),
- (atac 1),
- (atac 1),
- (rtac conjunct1 1),
- (rtac (less_sprod1b RS subst) 1),
- (atac 1),
- (atac 1),
- (res_inst_tac [("y","Issnd(p2)")] trans_less 1),
- (rtac conjunct2 1),
- (rtac (less_sprod1b RS subst) 1),
- (rtac defined_Ispair 1),
- (atac 1),
- (atac 1),
- (atac 1),
- (rtac conjunct2 1),
- (rtac (less_sprod1b RS subst) 1),
- (atac 1),
- (atac 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("s","Ispair(UU,UU)"),("t","Ispair(x,y)")] subst 1),
- (etac (less_sprod2b RS sym) 1),
- (atac 1)
- ]);
-
-
-
-
-
-
-
-
-
-
--- a/src/HOLCF/sprod1.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,22 +0,0 @@
-(* Title: HOLCF/sprod1.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Partial ordering for the strict product
-*)
-
-Sprod1 = Sprod0 +
-
-consts
- less_sprod :: "[('a ** 'b),('a ** 'b)] => bool"
-
-rules
-
- less_sprod_def "less_sprod(p1,p2) == @z.\
-\ ( p1=Ispair(UU,UU) --> z = True)\
-\ &(~p1=Ispair(UU,UU) --> z = ( Isfst(p1) << Isfst(p2) &\
-\ Issnd(p1) << Issnd(p2)))"
-
-end
-
--- a/src/HOLCF/sprod2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,274 +0,0 @@
-(* Title: HOLCF/sprod2.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for sprod2.thy
-*)
-
-
-open Sprod2;
-
-(* ------------------------------------------------------------------------ *)
-(* access to less_sprod in class po *)
-(* ------------------------------------------------------------------------ *)
-
-val less_sprod3a = prove_goal Sprod2.thy
- "p1=Ispair(UU,UU) ==> p1 << p2"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (inst_sprod_po RS ssubst) 1),
- (etac less_sprod1a 1)
- ]);
-
-
-val less_sprod3b = prove_goal Sprod2.thy
- "~p1=Ispair(UU,UU) ==>\
-\ (p1<<p2) = (Isfst(p1)<<Isfst(p2) & Issnd(p1)<<Issnd(p2))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (inst_sprod_po RS ssubst) 1),
- (etac less_sprod1b 1)
- ]);
-
-val less_sprod4b = prove_goal Sprod2.thy
- "p << Ispair(UU,UU) ==> p = Ispair(UU,UU)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac less_sprod2b 1),
- (etac (inst_sprod_po RS subst) 1)
- ]);
-
-val less_sprod4a = (less_sprod4b RS defined_Ispair_rev);
-(* Ispair(?a,?b) << Ispair(UU,UU) ==> ?a = UU | ?b = UU *)
-
-val less_sprod4c = prove_goal Sprod2.thy
- "[|Ispair(xa,ya)<<Ispair(x,y);~xa=UU;~ya=UU;~x=UU;~y=UU|] ==>\
-\ xa<<x & ya << y"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac less_sprod2c 1),
- (etac (inst_sprod_po RS subst) 1),
- (REPEAT (atac 1))
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* type sprod is pointed *)
-(* ------------------------------------------------------------------------ *)
-
-val minimal_sprod = prove_goal Sprod2.thy "Ispair(UU,UU)<<p"
-(fn prems =>
- [
- (rtac less_sprod3a 1),
- (rtac refl 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Ispair is monotone in both arguments *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_Ispair1 = prove_goalw Sprod2.thy [monofun] "monofun(Ispair)"
-(fn prems =>
- [
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (res_inst_tac [("Q",
- " Ispair(y,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1),
- (res_inst_tac [("Q",
- " Ispair(x,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1),
- (rtac (less_sprod3b RS iffD2) 1),
- (atac 1),
- (rtac conjI 1),
- (rtac (Isfst RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (rtac (Isfst RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (atac 1),
- (rtac (Issnd RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (rtac (Issnd RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (rtac refl_less 1),
- (etac less_sprod3a 1),
- (res_inst_tac [("Q",
- " Ispair(x,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1),
- (etac less_sprod3a 2),
- (res_inst_tac [("P","Ispair(y,xa) = Ispair(UU,UU)")] notE 1),
- (atac 2),
- (rtac defined_Ispair 1),
- (etac notUU_I 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1)
- ]);
-
-
-val monofun_Ispair2 = prove_goalw Sprod2.thy [monofun] "monofun(Ispair(x))"
-(fn prems =>
- [
- (strip_tac 1),
- (res_inst_tac [("Q",
- " Ispair(x,y) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1),
- (res_inst_tac [("Q",
- " Ispair(x,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1),
- (rtac (less_sprod3b RS iffD2) 1),
- (atac 1),
- (rtac conjI 1),
- (rtac (Isfst RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (rtac (Isfst RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (rtac refl_less 1),
- (rtac (Issnd RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (rtac (Issnd RS ssubst) 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac (strict_Ispair_rev RS conjunct2) 1),
- (atac 1),
- (etac less_sprod3a 1),
- (res_inst_tac [("Q",
- " Ispair(x,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1),
- (etac less_sprod3a 2),
- (res_inst_tac [("P","Ispair(x,y) = Ispair(UU,UU)")] notE 1),
- (atac 2),
- (rtac defined_Ispair 1),
- (etac (strict_Ispair_rev RS conjunct1) 1),
- (etac notUU_I 1),
- (etac (strict_Ispair_rev RS conjunct2) 1)
- ]);
-
-val monofun_Ispair = prove_goal Sprod2.thy
- "[|x1<<x2; y1<<y2|] ==> Ispair(x1,y1)<<Ispair(x2,y2)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac trans_less 1),
- (rtac (monofun_Ispair1 RS monofunE RS spec RS spec RS mp RS
- (less_fun RS iffD1 RS spec)) 1),
- (rtac (monofun_Ispair2 RS monofunE RS spec RS spec RS mp) 2),
- (atac 1),
- (atac 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Isfst and Issnd are monotone *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_Isfst = prove_goalw Sprod2.thy [monofun] "monofun(Isfst)"
-(fn prems =>
- [
- (strip_tac 1),
- (res_inst_tac [("p","x")] IsprodE 1),
- (hyp_subst_tac 1),
- (rtac trans_less 1),
- (rtac minimal 2),
- (rtac (strict_Isfst1 RS ssubst) 1),
- (rtac refl_less 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","y")] IsprodE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("t","Isfst(Ispair(xa,ya))")] subst 1),
- (rtac refl_less 2),
- (etac (less_sprod4b RS sym RS arg_cong) 1),
- (hyp_subst_tac 1),
- (rtac (Isfst RS ssubst) 1),
- (atac 1),
- (atac 1),
- (rtac (Isfst RS ssubst) 1),
- (atac 1),
- (atac 1),
- (etac (less_sprod4c RS conjunct1) 1),
- (REPEAT (atac 1))
- ]);
-
-val monofun_Issnd = prove_goalw Sprod2.thy [monofun] "monofun(Issnd)"
-(fn prems =>
- [
- (strip_tac 1),
- (res_inst_tac [("p","x")] IsprodE 1),
- (hyp_subst_tac 1),
- (rtac trans_less 1),
- (rtac minimal 2),
- (rtac (strict_Issnd1 RS ssubst) 1),
- (rtac refl_less 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","y")] IsprodE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("t","Issnd(Ispair(xa,ya))")] subst 1),
- (rtac refl_less 2),
- (etac (less_sprod4b RS sym RS arg_cong) 1),
- (hyp_subst_tac 1),
- (rtac (Issnd RS ssubst) 1),
- (atac 1),
- (atac 1),
- (rtac (Issnd RS ssubst) 1),
- (atac 1),
- (atac 1),
- (etac (less_sprod4c RS conjunct2) 1),
- (REPEAT (atac 1))
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* the type 'a ** 'b is a cpo *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_sprod = prove_goal Sprod2.thy
-"[|is_chain(S)|] ==> range(S) <<| \
-\ Ispair(lub(range(%i.Isfst(S(i)))),lub(range(%i.Issnd(S(i)))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (res_inst_tac [("t","S(i)")] (surjective_pairing_Sprod RS ssubst) 1),
- (rtac monofun_Ispair 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_Isfst RS ch2ch_monofun) 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_Issnd RS ch2ch_monofun) 1),
- (strip_tac 1),
- (res_inst_tac [("t","u")] (surjective_pairing_Sprod RS ssubst) 1),
- (rtac monofun_Ispair 1),
- (rtac is_lub_thelub 1),
- (etac (monofun_Isfst RS ch2ch_monofun) 1),
- (etac (monofun_Isfst RS ub2ub_monofun) 1),
- (rtac is_lub_thelub 1),
- (etac (monofun_Issnd RS ch2ch_monofun) 1),
- (etac (monofun_Issnd RS ub2ub_monofun) 1)
- ]);
-
-val thelub_sprod = (lub_sprod RS thelubI);
-(* is_chain(?S1) ==> lub(range(?S1)) = *)
-(* Ispair(lub(range(%i. Isfst(?S1(i)))),lub(range(%i. Issnd(?S1(i))))) *)
-
-val cpo_sprod = prove_goal Sprod2.thy
- "is_chain(S::nat=>'a**'b)==>? x.range(S)<<| x"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac exI 1),
- (etac lub_sprod 1)
- ]);
-
-
-
-
-
-
-
-
--- a/src/HOLCF/sprod2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-(* Title: HOLCF/sprod2.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class Instance **::(pcpo,pcpo)po
-*)
-
-Sprod2 = Sprod1 +
-
-arities "**" :: (pcpo,pcpo)po
-
-(* Witness for the above arity axiom is sprod1.ML *)
-
-rules
-
-(* instance of << for type ['a ** 'b] *)
-
-inst_sprod_po "(op <<)::['a ** 'b,'a ** 'b]=>bool = less_sprod"
-
-end
-
-
-
--- a/src/HOLCF/sprod3.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,695 +0,0 @@
-(* Title: HOLCF/sprod3.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for Sprod3.thy
-*)
-
-open Sprod3;
-
-(* ------------------------------------------------------------------------ *)
-(* continuity of Ispair, Isfst, Issnd *)
-(* ------------------------------------------------------------------------ *)
-
-val sprod3_lemma1 = prove_goal Sprod3.thy
-"[| is_chain(Y); x~= UU; lub(range(Y))~= UU |] ==>\
-\ Ispair(lub(range(Y)),x) =\
-\ Ispair(lub(range(%i. Isfst(Ispair(Y(i),x)))),\
-\ lub(range(%i. Issnd(Ispair(Y(i),x)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("f1","Ispair")] (arg_cong RS cong) 1),
- (rtac lub_equal 1),
- (atac 1),
- (rtac (monofun_Isfst RS ch2ch_monofun) 1),
- (rtac ch2ch_fun 1),
- (rtac (monofun_Ispair1 RS ch2ch_monofun) 1),
- (atac 1),
- (rtac allI 1),
- (asm_simp_tac Sprod_ss 1),
- (rtac sym 1),
- (rtac lub_chain_maxelem 1),
- (rtac (monofun_Issnd RS ch2ch_monofun) 1),
- (rtac ch2ch_fun 1),
- (rtac (monofun_Ispair1 RS ch2ch_monofun) 1),
- (atac 1),
- (res_inst_tac [("P","%j.~Y(j)=UU")] exE 1),
- (rtac (notall2ex RS iffD1) 1),
- (res_inst_tac [("Q","lub(range(Y)) = UU")] contrapos 1),
- (atac 1),
- (rtac chain_UU_I_inverse 1),
- (atac 1),
- (rtac exI 1),
- (etac Issnd2 1),
- (rtac allI 1),
- (res_inst_tac [("Q","Y(i)=UU")] (excluded_middle RS disjE) 1),
- (asm_simp_tac Sprod_ss 1),
- (rtac refl_less 1),
- (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1),
- (etac sym 1),
- (asm_simp_tac Sprod_ss 1),
- (rtac minimal 1)
- ]);
-
-
-val sprod3_lemma2 = prove_goal Sprod3.thy
-"[| is_chain(Y); ~ x = UU; lub(range(Y)) = UU |] ==>\
-\ Ispair(lub(range(Y)),x) =\
-\ Ispair(lub(range(%i. Isfst(Ispair(Y(i),x)))),\
-\ lub(range(%i. Issnd(Ispair(Y(i),x)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1),
- (atac 1),
- (rtac trans 1),
- (rtac strict_Ispair1 1),
- (rtac (strict_Ispair RS sym) 1),
- (rtac disjI1 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (asm_simp_tac Sprod_ss 1),
- (etac (chain_UU_I RS spec) 1),
- (atac 1)
- ]);
-
-
-val sprod3_lemma3 = prove_goal Sprod3.thy
-"[| is_chain(Y); x = UU |] ==>\
-\ Ispair(lub(range(Y)),x) =\
-\ Ispair(lub(range(%i. Isfst(Ispair(Y(i),x)))),\
-\ lub(range(%i. Issnd(Ispair(Y(i),x)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("s","UU"),("t","x")] ssubst 1),
- (atac 1),
- (rtac trans 1),
- (rtac strict_Ispair2 1),
- (rtac (strict_Ispair RS sym) 1),
- (rtac disjI1 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (simp_tac Sprod_ss 1)
- ]);
-
-
-val contlub_Ispair1 = prove_goal Sprod3.thy "contlub(Ispair)"
-(fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (expand_fun_eq RS iffD2) 1),
- (strip_tac 1),
- (rtac (lub_fun RS thelubI RS ssubst) 1),
- (etac (monofun_Ispair1 RS ch2ch_monofun) 1),
- (rtac trans 1),
- (rtac (thelub_sprod RS sym) 2),
- (rtac ch2ch_fun 2),
- (etac (monofun_Ispair1 RS ch2ch_monofun) 2),
- (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1),
- (res_inst_tac
- [("Q","lub(range(Y))=UU")] (excluded_middle RS disjE) 1),
- (etac sprod3_lemma1 1),
- (atac 1),
- (atac 1),
- (etac sprod3_lemma2 1),
- (atac 1),
- (atac 1),
- (etac sprod3_lemma3 1),
- (atac 1)
- ]);
-
-val sprod3_lemma4 = prove_goal Sprod3.thy
-"[| is_chain(Y); ~ x = UU; ~ lub(range(Y)) = UU |] ==>\
-\ Ispair(x,lub(range(Y))) =\
-\ Ispair(lub(range(%i. Isfst(Ispair(x,Y(i))))),\
-\ lub(range(%i. Issnd(Ispair(x,Y(i))))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("f1","Ispair")] (arg_cong RS cong) 1),
- (rtac sym 1),
- (rtac lub_chain_maxelem 1),
- (rtac (monofun_Isfst RS ch2ch_monofun) 1),
- (rtac (monofun_Ispair2 RS ch2ch_monofun) 1),
- (atac 1),
- (res_inst_tac [("P","%j.~Y(j)=UU")] exE 1),
- (rtac (notall2ex RS iffD1) 1),
- (res_inst_tac [("Q","lub(range(Y)) = UU")] contrapos 1),
- (atac 1),
- (rtac chain_UU_I_inverse 1),
- (atac 1),
- (rtac exI 1),
- (etac Isfst2 1),
- (rtac allI 1),
- (res_inst_tac [("Q","Y(i)=UU")] (excluded_middle RS disjE) 1),
- (asm_simp_tac Sprod_ss 1),
- (rtac refl_less 1),
- (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1),
- (etac sym 1),
- (asm_simp_tac Sprod_ss 1),
- (rtac minimal 1),
- (rtac lub_equal 1),
- (atac 1),
- (rtac (monofun_Issnd RS ch2ch_monofun) 1),
- (rtac (monofun_Ispair2 RS ch2ch_monofun) 1),
- (atac 1),
- (rtac allI 1),
- (asm_simp_tac Sprod_ss 1)
- ]);
-
-val sprod3_lemma5 = prove_goal Sprod3.thy
-"[| is_chain(Y); ~ x = UU; lub(range(Y)) = UU |] ==>\
-\ Ispair(x,lub(range(Y))) =\
-\ Ispair(lub(range(%i. Isfst(Ispair(x,Y(i))))),\
-\ lub(range(%i. Issnd(Ispair(x,Y(i))))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1),
- (atac 1),
- (rtac trans 1),
- (rtac strict_Ispair2 1),
- (rtac (strict_Ispair RS sym) 1),
- (rtac disjI2 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (asm_simp_tac Sprod_ss 1),
- (etac (chain_UU_I RS spec) 1),
- (atac 1)
- ]);
-
-val sprod3_lemma6 = prove_goal Sprod3.thy
-"[| is_chain(Y); x = UU |] ==>\
-\ Ispair(x,lub(range(Y))) =\
-\ Ispair(lub(range(%i. Isfst(Ispair(x,Y(i))))),\
-\ lub(range(%i. Issnd(Ispair(x,Y(i))))))"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("s","UU"),("t","x")] ssubst 1),
- (atac 1),
- (rtac trans 1),
- (rtac strict_Ispair1 1),
- (rtac (strict_Ispair RS sym) 1),
- (rtac disjI1 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (simp_tac Sprod_ss 1)
- ]);
-
-val contlub_Ispair2 = prove_goal Sprod3.thy "contlub(Ispair(x))"
-(fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_sprod RS sym) 2),
- (etac (monofun_Ispair2 RS ch2ch_monofun) 2),
- (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1),
- (res_inst_tac [("Q","lub(range(Y))=UU")]
- (excluded_middle RS disjE) 1),
- (etac sprod3_lemma4 1),
- (atac 1),
- (atac 1),
- (etac sprod3_lemma5 1),
- (atac 1),
- (atac 1),
- (etac sprod3_lemma6 1),
- (atac 1)
- ]);
-
-
-val contX_Ispair1 = prove_goal Sprod3.thy "contX(Ispair)"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Ispair1 1),
- (rtac contlub_Ispair1 1)
- ]);
-
-
-val contX_Ispair2 = prove_goal Sprod3.thy "contX(Ispair(x))"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Ispair2 1),
- (rtac contlub_Ispair2 1)
- ]);
-
-val contlub_Isfst = prove_goal Sprod3.thy "contlub(Isfst)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (lub_sprod RS thelubI RS ssubst) 1),
- (atac 1),
- (res_inst_tac [("Q","lub(range(%i. Issnd(Y(i))))=UU")]
- (excluded_middle RS disjE) 1),
- (asm_simp_tac Sprod_ss 1),
- (res_inst_tac [("s","UU"),("t","lub(range(%i. Issnd(Y(i))))")]
- ssubst 1),
- (atac 1),
- (rtac trans 1),
- (asm_simp_tac Sprod_ss 1),
- (rtac sym 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (rtac strict_Isfst 1),
- (rtac swap 1),
- (etac (defined_IsfstIssnd RS conjunct2) 2),
- (rtac notnotI 1),
- (rtac (chain_UU_I RS spec) 1),
- (rtac (monofun_Issnd RS ch2ch_monofun) 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-val contlub_Issnd = prove_goal Sprod3.thy "contlub(Issnd)"
-(fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac (lub_sprod RS thelubI RS ssubst) 1),
- (atac 1),
- (res_inst_tac [("Q","lub(range(%i. Isfst(Y(i))))=UU")]
- (excluded_middle RS disjE) 1),
- (asm_simp_tac Sprod_ss 1),
- (res_inst_tac [("s","UU"),("t","lub(range(%i. Isfst(Y(i))))")]
- ssubst 1),
- (atac 1),
- (asm_simp_tac Sprod_ss 1),
- (rtac sym 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (rtac strict_Issnd 1),
- (rtac swap 1),
- (etac (defined_IsfstIssnd RS conjunct1) 2),
- (rtac notnotI 1),
- (rtac (chain_UU_I RS spec) 1),
- (rtac (monofun_Isfst RS ch2ch_monofun) 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-val contX_Isfst = prove_goal Sprod3.thy "contX(Isfst)"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Isfst 1),
- (rtac contlub_Isfst 1)
- ]);
-
-val contX_Issnd = prove_goal Sprod3.thy "contX(Issnd)"
-(fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Issnd 1),
- (rtac contlub_Issnd 1)
- ]);
-
-(*
- --------------------------------------------------------------------------
- more lemmas for Sprod3.thy
-
- --------------------------------------------------------------------------
-*)
-
-val spair_eq = prove_goal Sprod3.thy "[|x1=x2;y1=y2|] ==> x1##y1 = x2##y2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (fast_tac HOL_cs 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* convert all lemmas to the continuous versions *)
-(* ------------------------------------------------------------------------ *)
-
-val beta_cfun_sprod = prove_goalw Sprod3.thy [spair_def]
- "(LAM x y.Ispair(x,y))[a][b] = Ispair(a,b)"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tac 1),
- (rtac contX_Ispair2 1),
- (rtac contX2contX_CF1L 1),
- (rtac contX_Ispair1 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Ispair2 1),
- (rtac refl 1)
- ]);
-
-val inject_spair = prove_goalw Sprod3.thy [spair_def]
- "[|~aa=UU ; ~ba=UU ; (a##b)=(aa##ba) |] ==> a=aa & b=ba"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac inject_Ispair 1),
- (atac 1),
- (etac box_equals 1),
- (rtac beta_cfun_sprod 1),
- (rtac beta_cfun_sprod 1)
- ]);
-
-val inst_sprod_pcpo2 = prove_goalw Sprod3.thy [spair_def] "UU = (UU##UU)"
- (fn prems =>
- [
- (rtac sym 1),
- (rtac trans 1),
- (rtac beta_cfun_sprod 1),
- (rtac sym 1),
- (rtac inst_sprod_pcpo 1)
- ]);
-
-val strict_spair = prove_goalw Sprod3.thy [spair_def]
- "(a=UU | b=UU) ==> (a##b)=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac trans 1),
- (rtac beta_cfun_sprod 1),
- (rtac trans 1),
- (rtac (inst_sprod_pcpo RS sym) 2),
- (etac strict_Ispair 1)
- ]);
-
-val strict_spair1 = prove_goalw Sprod3.thy [spair_def] "(UU##b) = UU"
- (fn prems =>
- [
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac trans 1),
- (rtac (inst_sprod_pcpo RS sym) 2),
- (rtac strict_Ispair1 1)
- ]);
-
-val strict_spair2 = prove_goalw Sprod3.thy [spair_def] "(a##UU) = UU"
- (fn prems =>
- [
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac trans 1),
- (rtac (inst_sprod_pcpo RS sym) 2),
- (rtac strict_Ispair2 1)
- ]);
-
-val strict_spair_rev = prove_goalw Sprod3.thy [spair_def]
- "~(x##y)=UU ==> ~x=UU & ~y=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac strict_Ispair_rev 1),
- (rtac (beta_cfun_sprod RS subst) 1),
- (rtac (inst_sprod_pcpo RS subst) 1),
- (atac 1)
- ]);
-
-val defined_spair_rev = prove_goalw Sprod3.thy [spair_def]
- "(a##b) = UU ==> (a = UU | b = UU)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac defined_Ispair_rev 1),
- (rtac (beta_cfun_sprod RS subst) 1),
- (rtac (inst_sprod_pcpo RS subst) 1),
- (atac 1)
- ]);
-
-val defined_spair = prove_goalw Sprod3.thy [spair_def]
- "[|~a=UU; ~b=UU|] ==> ~(a##b) = UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (inst_sprod_pcpo RS ssubst) 1),
- (etac defined_Ispair 1),
- (atac 1)
- ]);
-
-val Exh_Sprod2 = prove_goalw Sprod3.thy [spair_def]
- "z=UU | (? a b. z=(a##b) & ~a=UU & ~b=UU)"
- (fn prems =>
- [
- (rtac (Exh_Sprod RS disjE) 1),
- (rtac disjI1 1),
- (rtac (inst_sprod_pcpo RS ssubst) 1),
- (atac 1),
- (rtac disjI2 1),
- (etac exE 1),
- (etac exE 1),
- (rtac exI 1),
- (rtac exI 1),
- (rtac conjI 1),
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val sprodE = prove_goalw Sprod3.thy [spair_def]
-"[|p=UU ==> Q;!!x y. [|p=(x##y);~x=UU ; ~y=UU|] ==> Q|] ==> Q"
-(fn prems =>
- [
- (rtac IsprodE 1),
- (resolve_tac prems 1),
- (rtac (inst_sprod_pcpo RS ssubst) 1),
- (atac 1),
- (resolve_tac prems 1),
- (atac 2),
- (atac 2),
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (atac 1)
- ]);
-
-
-val strict_sfst = prove_goalw Sprod3.thy [sfst_def]
- "p=UU==>sfst[p]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac strict_Isfst 1),
- (rtac (inst_sprod_pcpo RS subst) 1),
- (atac 1)
- ]);
-
-val strict_sfst1 = prove_goalw Sprod3.thy [sfst_def,spair_def]
- "sfst[UU##y] = UU"
- (fn prems =>
- [
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac strict_Isfst1 1)
- ]);
-
-val strict_sfst2 = prove_goalw Sprod3.thy [sfst_def,spair_def]
- "sfst[x##UU] = UU"
- (fn prems =>
- [
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac strict_Isfst2 1)
- ]);
-
-val strict_ssnd = prove_goalw Sprod3.thy [ssnd_def]
- "p=UU==>ssnd[p]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac strict_Issnd 1),
- (rtac (inst_sprod_pcpo RS subst) 1),
- (atac 1)
- ]);
-
-val strict_ssnd1 = prove_goalw Sprod3.thy [ssnd_def,spair_def]
- "ssnd[UU##y] = UU"
- (fn prems =>
- [
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac strict_Issnd1 1)
- ]);
-
-val strict_ssnd2 = prove_goalw Sprod3.thy [ssnd_def,spair_def]
- "ssnd[x##UU] = UU"
- (fn prems =>
- [
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac strict_Issnd2 1)
- ]);
-
-val sfst2 = prove_goalw Sprod3.thy [sfst_def,spair_def]
- "~y=UU ==>sfst[x##y]=x"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (etac Isfst2 1)
- ]);
-
-val ssnd2 = prove_goalw Sprod3.thy [ssnd_def,spair_def]
- "~x=UU ==>ssnd[x##y]=y"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (etac Issnd2 1)
- ]);
-
-
-val defined_sfstssnd = prove_goalw Sprod3.thy [sfst_def,ssnd_def,spair_def]
- "~p=UU ==> ~sfst[p]=UU & ~ssnd[p]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac defined_IsfstIssnd 1),
- (rtac (inst_sprod_pcpo RS subst) 1),
- (atac 1)
- ]);
-
-
-val surjective_pairing_Sprod2 = prove_goalw Sprod3.thy
- [sfst_def,ssnd_def,spair_def] "(sfst[p] ## ssnd[p]) = p"
- (fn prems =>
- [
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac (surjective_pairing_Sprod RS sym) 1)
- ]);
-
-
-val less_sprod5b = prove_goalw Sprod3.thy [sfst_def,ssnd_def,spair_def]
- "~p1=UU ==> (p1<<p2) = (sfst[p1]<<sfst[p2] & ssnd[p1]<<ssnd[p2])"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac less_sprod3b 1),
- (rtac (inst_sprod_pcpo RS subst) 1),
- (atac 1)
- ]);
-
-
-val less_sprod5c = prove_goalw Sprod3.thy [sfst_def,ssnd_def,spair_def]
- "[|xa##ya<<x##y;~xa=UU;~ya=UU;~x=UU;~y=UU|] ==>xa<<x & ya << y"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac less_sprod4c 1),
- (REPEAT (atac 2)),
- (rtac (beta_cfun_sprod RS subst) 1),
- (rtac (beta_cfun_sprod RS subst) 1),
- (atac 1)
- ]);
-
-val lub_sprod2 = prove_goalw Sprod3.thy [sfst_def,ssnd_def,spair_def]
-"[|is_chain(S)|] ==> range(S) <<| \
-\ (lub(range(%i.sfst[S(i)])) ## lub(range(%i.ssnd[S(i)])))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun_sprod RS ssubst) 1),
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (rtac contX_Issnd 1),
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (rtac contX_Isfst 1),
- (rtac lub_sprod 1),
- (resolve_tac prems 1)
- ]);
-
-
-val thelub_sprod2 = (lub_sprod2 RS thelubI);
-(* is_chain(?S1) ==> lub(range(?S1)) = *)
-(* (lub(range(%i. sfst[?S1(i)]))## lub(range(%i. ssnd[?S1(i)]))) *)
-
-
-
-val ssplit1 = prove_goalw Sprod3.thy [ssplit_def]
- "ssplit[f][UU]=UU"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tacR 1),
- (rtac (strictify1 RS ssubst) 1),
- (rtac refl 1)
- ]);
-
-val ssplit2 = prove_goalw Sprod3.thy [ssplit_def]
- "[|~x=UU;~y=UU|] ==> ssplit[f][x##y]=f[x][y]"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tacR 1),
- (rtac (strictify2 RS ssubst) 1),
- (rtac defined_spair 1),
- (resolve_tac prems 1),
- (resolve_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tacR 1),
- (rtac (sfst2 RS ssubst) 1),
- (resolve_tac prems 1),
- (rtac (ssnd2 RS ssubst) 1),
- (resolve_tac prems 1),
- (rtac refl 1)
- ]);
-
-
-val ssplit3 = prove_goalw Sprod3.thy [ssplit_def]
- "ssplit[spair][z]=z"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tacR 1),
- (res_inst_tac [("Q","z=UU")] classical2 1),
- (hyp_subst_tac 1),
- (rtac strictify1 1),
- (rtac trans 1),
- (rtac strictify2 1),
- (atac 1),
- (rtac (beta_cfun RS ssubst) 1),
- (contX_tacR 1),
- (rtac surjective_pairing_Sprod2 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* install simplifier for Sprod *)
-(* ------------------------------------------------------------------------ *)
-
-val Sprod_rews = [strict_spair1,strict_spair2,strict_sfst1,strict_sfst2,
- strict_ssnd1,strict_ssnd2,sfst2,ssnd2,
- ssplit1,ssplit2];
-
-val Sprod_ss = Cfun_ss addsimps Sprod_rews;
-
-
--- a/src/HOLCF/sprod3.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,39 +0,0 @@
-(* Title: HOLCF/sprod3.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class instance of ** for class pcpo
-*)
-
-Sprod3 = Sprod2 +
-
-arities "**" :: (pcpo,pcpo)pcpo (* Witness sprod2.ML *)
-
-consts
- "@spair" :: "'a => 'b => ('a**'b)" ("_##_" [101,100] 100)
- "cop @spair" :: "'a -> 'b -> ('a**'b)" ("spair")
- (* continuous strict pairing *)
- sfst :: "('a**'b)->'a"
- ssnd :: "('a**'b)->'b"
- ssplit :: "('a->'b->'c)->('a**'b)->'c"
-
-rules
-
-inst_sprod_pcpo "UU::'a**'b = Ispair(UU,UU)"
-spair_def "spair == (LAM x y.Ispair(x,y))"
-sfst_def "sfst == (LAM p.Isfst(p))"
-ssnd_def "ssnd == (LAM p.Issnd(p))"
-ssplit_def "ssplit == (LAM f. strictify[LAM p.f[sfst[p]][ssnd[p]]])"
-
-end
-
-ML
-
-(* ----------------------------------------------------------------------*)
-(* parse translations for the above mixfix *)
-(* ----------------------------------------------------------------------*)
-
-val parse_translation = [("@spair",mk_cinfixtr "@spair")];
-
-
--- a/src/HOLCF/ssum0.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,394 +0,0 @@
-(* Title: HOLCF/ssum0.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for theory ssum0.thy
-*)
-
-open Ssum0;
-
-(* ------------------------------------------------------------------------ *)
-(* A non-emptyness result for Sssum *)
-(* ------------------------------------------------------------------------ *)
-
-val SsumIl = prove_goalw Ssum0.thy [Ssum_def] "Sinl_Rep(a):Ssum"
- (fn prems =>
- [
- (rtac CollectI 1),
- (rtac disjI1 1),
- (rtac exI 1),
- (rtac refl 1)
- ]);
-
-val SsumIr = prove_goalw Ssum0.thy [Ssum_def] "Sinr_Rep(a):Ssum"
- (fn prems =>
- [
- (rtac CollectI 1),
- (rtac disjI2 1),
- (rtac exI 1),
- (rtac refl 1)
- ]);
-
-val inj_onto_Abs_Ssum = prove_goal Ssum0.thy "inj_onto(Abs_Ssum,Ssum)"
-(fn prems =>
- [
- (rtac inj_onto_inverseI 1),
- (etac Abs_Ssum_inverse 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Strictness of Sinr_Rep, Sinl_Rep and Isinl, Isinr *)
-(* ------------------------------------------------------------------------ *)
-
-val strict_SinlSinr_Rep = prove_goalw Ssum0.thy [Sinr_Rep_def,Sinl_Rep_def]
- "Sinl_Rep(UU) = Sinr_Rep(UU)"
- (fn prems =>
- [
- (rtac ext 1),
- (rtac ext 1),
- (rtac ext 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val strict_IsinlIsinr = prove_goalw Ssum0.thy [Isinl_def,Isinr_def]
- "Isinl(UU) = Isinr(UU)"
- (fn prems =>
- [
- (rtac (strict_SinlSinr_Rep RS arg_cong) 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* distinctness of Sinl_Rep, Sinr_Rep and Isinl, Isinr *)
-(* ------------------------------------------------------------------------ *)
-
-val noteq_SinlSinr_Rep = prove_goalw Ssum0.thy [Sinl_Rep_def,Sinr_Rep_def]
- "(Sinl_Rep(a) = Sinr_Rep(b)) ==> a=UU & b=UU"
- (fn prems =>
- [
- (rtac conjI 1),
- (res_inst_tac [("Q","a=UU")] classical2 1),
- (atac 1),
- (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD2
- RS mp RS conjunct1 RS sym) 1),
- (fast_tac HOL_cs 1),
- (atac 1),
- (res_inst_tac [("Q","b=UU")] classical2 1),
- (atac 1),
- (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD1
- RS mp RS conjunct1 RS sym) 1),
- (fast_tac HOL_cs 1),
- (atac 1)
- ]);
-
-
-val noteq_IsinlIsinr = prove_goalw Ssum0.thy [Isinl_def,Isinr_def]
- "Isinl(a)=Isinr(b) ==> a=UU & b=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac noteq_SinlSinr_Rep 1),
- (etac (inj_onto_Abs_Ssum RS inj_ontoD) 1),
- (rtac SsumIl 1),
- (rtac SsumIr 1)
- ]);
-
-
-
-(* ------------------------------------------------------------------------ *)
-(* injectivity of Sinl_Rep, Sinr_Rep and Isinl, Isinr *)
-(* ------------------------------------------------------------------------ *)
-
-val inject_Sinl_Rep1 = prove_goalw Ssum0.thy [Sinl_Rep_def]
- "(Sinl_Rep(a) = Sinl_Rep(UU)) ==> a=UU"
- (fn prems =>
- [
- (res_inst_tac [("Q","a=UU")] classical2 1),
- (atac 1),
- (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong
- RS iffD2 RS mp RS conjunct1 RS sym) 1),
- (fast_tac HOL_cs 1),
- (atac 1)
- ]);
-
-val inject_Sinr_Rep1 = prove_goalw Ssum0.thy [Sinr_Rep_def]
- "(Sinr_Rep(b) = Sinr_Rep(UU)) ==> b=UU"
- (fn prems =>
- [
- (res_inst_tac [("Q","b=UU")] classical2 1),
- (atac 1),
- (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong
- RS iffD2 RS mp RS conjunct1 RS sym) 1),
- (fast_tac HOL_cs 1),
- (atac 1)
- ]);
-
-val inject_Sinl_Rep2 = prove_goalw Ssum0.thy [Sinl_Rep_def]
-"[|~a1=UU ; ~a2=UU ; Sinl_Rep(a1)=Sinl_Rep(a2) |] ==> a1=a2"
- (fn prems =>
- [
- (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong RS fun_cong
- RS iffD1 RS mp RS conjunct1) 1),
- (fast_tac HOL_cs 1),
- (resolve_tac prems 1)
- ]);
-
-val inject_Sinr_Rep2 = prove_goalw Ssum0.thy [Sinr_Rep_def]
-"[|~b1=UU ; ~b2=UU ; Sinr_Rep(b1)=Sinr_Rep(b2) |] ==> b1=b2"
- (fn prems =>
- [
- (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong RS fun_cong
- RS iffD1 RS mp RS conjunct1) 1),
- (fast_tac HOL_cs 1),
- (resolve_tac prems 1)
- ]);
-
-val inject_Sinl_Rep = prove_goal Ssum0.thy
- "Sinl_Rep(a1)=Sinl_Rep(a2) ==> a1=a2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("Q","a1=UU")] classical2 1),
- (hyp_subst_tac 1),
- (rtac (inject_Sinl_Rep1 RS sym) 1),
- (etac sym 1),
- (res_inst_tac [("Q","a2=UU")] classical2 1),
- (hyp_subst_tac 1),
- (etac inject_Sinl_Rep1 1),
- (etac inject_Sinl_Rep2 1),
- (atac 1),
- (atac 1)
- ]);
-
-val inject_Sinr_Rep = prove_goal Ssum0.thy
- "Sinr_Rep(b1)=Sinr_Rep(b2) ==> b1=b2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("Q","b1=UU")] classical2 1),
- (hyp_subst_tac 1),
- (rtac (inject_Sinr_Rep1 RS sym) 1),
- (etac sym 1),
- (res_inst_tac [("Q","b2=UU")] classical2 1),
- (hyp_subst_tac 1),
- (etac inject_Sinr_Rep1 1),
- (etac inject_Sinr_Rep2 1),
- (atac 1),
- (atac 1)
- ]);
-
-val inject_Isinl = prove_goalw Ssum0.thy [Isinl_def]
-"Isinl(a1)=Isinl(a2)==> a1=a2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac inject_Sinl_Rep 1),
- (etac (inj_onto_Abs_Ssum RS inj_ontoD) 1),
- (rtac SsumIl 1),
- (rtac SsumIl 1)
- ]);
-
-val inject_Isinr = prove_goalw Ssum0.thy [Isinr_def]
-"Isinr(b1)=Isinr(b2) ==> b1=b2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac inject_Sinr_Rep 1),
- (etac (inj_onto_Abs_Ssum RS inj_ontoD) 1),
- (rtac SsumIr 1),
- (rtac SsumIr 1)
- ]);
-
-val inject_Isinl_rev = prove_goal Ssum0.thy
-"~a1=a2 ==> ~Isinl(a1) = Isinl(a2)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contrapos 1),
- (etac inject_Isinl 2),
- (atac 1)
- ]);
-
-val inject_Isinr_rev = prove_goal Ssum0.thy
-"~b1=b2 ==> ~Isinr(b1) = Isinr(b2)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac contrapos 1),
- (etac inject_Isinr 2),
- (atac 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* Exhaustion of the strict sum ++ *)
-(* choice of the bottom representation is arbitrary *)
-(* ------------------------------------------------------------------------ *)
-
-val Exh_Ssum = prove_goalw Ssum0.thy [Isinl_def,Isinr_def]
- "z=Isinl(UU) | (? a. z=Isinl(a) & ~a=UU) | (? b. z=Isinr(b) & ~b=UU)"
- (fn prems =>
- [
- (rtac (rewrite_rule [Ssum_def] Rep_Ssum RS CollectE) 1),
- (etac disjE 1),
- (etac exE 1),
- (res_inst_tac [("Q","z= Abs_Ssum(Sinl_Rep(UU))")] classical2 1),
- (etac disjI1 1),
- (rtac disjI2 1),
- (rtac disjI1 1),
- (rtac exI 1),
- (rtac conjI 1),
- (rtac (Rep_Ssum_inverse RS sym RS trans) 1),
- (etac arg_cong 1),
- (res_inst_tac [("Q","Sinl_Rep(a)=Sinl_Rep(UU)")] contrapos 1),
- (etac arg_cong 2),
- (etac contrapos 1),
- (rtac (Rep_Ssum_inverse RS sym RS trans) 1),
- (rtac trans 1),
- (etac arg_cong 1),
- (etac arg_cong 1),
- (etac exE 1),
- (res_inst_tac [("Q","z= Abs_Ssum(Sinl_Rep(UU))")] classical2 1),
- (etac disjI1 1),
- (rtac disjI2 1),
- (rtac disjI2 1),
- (rtac exI 1),
- (rtac conjI 1),
- (rtac (Rep_Ssum_inverse RS sym RS trans) 1),
- (etac arg_cong 1),
- (res_inst_tac [("Q","Sinr_Rep(b)=Sinl_Rep(UU)")] contrapos 1),
- (hyp_subst_tac 2),
- (rtac (strict_SinlSinr_Rep RS sym) 2),
- (etac contrapos 1),
- (rtac (Rep_Ssum_inverse RS sym RS trans) 1),
- (rtac trans 1),
- (etac arg_cong 1),
- (etac arg_cong 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* elimination rules for the strict sum ++ *)
-(* ------------------------------------------------------------------------ *)
-
-val IssumE = prove_goal Ssum0.thy
- "[|p=Isinl(UU) ==> Q ;\
-\ !!x.[|p=Isinl(x); ~x=UU |] ==> Q;\
-\ !!y.[|p=Isinr(y); ~y=UU |] ==> Q|] ==> Q"
- (fn prems =>
- [
- (rtac (Exh_Ssum RS disjE) 1),
- (etac disjE 2),
- (eresolve_tac prems 1),
- (etac exE 1),
- (etac conjE 1),
- (eresolve_tac prems 1),
- (atac 1),
- (etac exE 1),
- (etac conjE 1),
- (eresolve_tac prems 1),
- (atac 1)
- ]);
-
-val IssumE2 = prove_goal Ssum0.thy
-"[| !!x. [| p = Isinl(x) |] ==> Q; !!y. [| p = Isinr(y) |] ==> Q |] ==>Q"
- (fn prems =>
- [
- (rtac IssumE 1),
- (eresolve_tac prems 1),
- (eresolve_tac prems 1),
- (eresolve_tac prems 1)
- ]);
-
-
-
-
-(* ------------------------------------------------------------------------ *)
-(* rewrites for Iwhen *)
-(* ------------------------------------------------------------------------ *)
-
-val Iwhen1 = prove_goalw Ssum0.thy [Iwhen_def]
- "Iwhen(f)(g)(Isinl(UU)) = UU"
- (fn prems =>
- [
- (rtac select_equality 1),
- (rtac conjI 1),
- (fast_tac HOL_cs 1),
- (rtac conjI 1),
- (strip_tac 1),
- (res_inst_tac [("P","a=UU")] notE 1),
- (fast_tac HOL_cs 1),
- (rtac inject_Isinl 1),
- (rtac sym 1),
- (fast_tac HOL_cs 1),
- (strip_tac 1),
- (res_inst_tac [("P","b=UU")] notE 1),
- (fast_tac HOL_cs 1),
- (rtac inject_Isinr 1),
- (rtac sym 1),
- (rtac (strict_IsinlIsinr RS subst) 1),
- (fast_tac HOL_cs 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val Iwhen2 = prove_goalw Ssum0.thy [Iwhen_def]
- "~x=UU ==> Iwhen(f)(g)(Isinl(x)) = f[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (fast_tac HOL_cs 2),
- (rtac conjI 1),
- (strip_tac 1),
- (res_inst_tac [("P","x=UU")] notE 1),
- (atac 1),
- (rtac inject_Isinl 1),
- (atac 1),
- (rtac conjI 1),
- (strip_tac 1),
- (rtac cfun_arg_cong 1),
- (rtac inject_Isinl 1),
- (fast_tac HOL_cs 1),
- (strip_tac 1),
- (res_inst_tac [("P","Isinl(x) = Isinr(b)")] notE 1),
- (fast_tac HOL_cs 2),
- (rtac contrapos 1),
- (etac noteq_IsinlIsinr 2),
- (fast_tac HOL_cs 1)
- ]);
-
-val Iwhen3 = prove_goalw Ssum0.thy [Iwhen_def]
- "~y=UU ==> Iwhen(f)(g)(Isinr(y)) = g[y]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (fast_tac HOL_cs 2),
- (rtac conjI 1),
- (strip_tac 1),
- (res_inst_tac [("P","y=UU")] notE 1),
- (atac 1),
- (rtac inject_Isinr 1),
- (rtac (strict_IsinlIsinr RS subst) 1),
- (atac 1),
- (rtac conjI 1),
- (strip_tac 1),
- (res_inst_tac [("P","Isinr(y) = Isinl(a)")] notE 1),
- (fast_tac HOL_cs 2),
- (rtac contrapos 1),
- (etac (sym RS noteq_IsinlIsinr) 2),
- (fast_tac HOL_cs 1),
- (strip_tac 1),
- (rtac cfun_arg_cong 1),
- (rtac inject_Isinr 1),
- (fast_tac HOL_cs 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* instantiate the simplifier *)
-(* ------------------------------------------------------------------------ *)
-
-val Ssum_ss = Cfun_ss addsimps
- [(strict_IsinlIsinr RS sym),Iwhen1,Iwhen2,Iwhen3];
-
-
--- a/src/HOLCF/ssum0.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,54 +0,0 @@
-(* Title: HOLCF/ssum0.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Strict sum
-*)
-
-Ssum0 = Cfun3 +
-
-(* new type for strict sum *)
-
-types "++" 2 (infixr 10)
-
-arities "++" :: (pcpo,pcpo)term
-
-consts
- Ssum :: "(['a,'b,bool]=>bool)set"
- Sinl_Rep :: "['a,'a,'b,bool]=>bool"
- Sinr_Rep :: "['b,'a,'b,bool]=>bool"
- Rep_Ssum :: "('a ++ 'b) => (['a,'b,bool]=>bool)"
- Abs_Ssum :: "(['a,'b,bool]=>bool) => ('a ++ 'b)"
- Isinl :: "'a => ('a ++ 'b)"
- Isinr :: "'b => ('a ++ 'b)"
- Iwhen :: "('a->'c)=>('b->'c)=>('a ++ 'b)=> 'c"
-
-rules
-
- Sinl_Rep_def "Sinl_Rep == (%a.%x y p.\
-\ (~a=UU --> x=a & p))"
-
- Sinr_Rep_def "Sinr_Rep == (%b.%x y p.\
-\ (~b=UU --> y=b & ~p))"
-
- Ssum_def "Ssum =={f.(? a.f=Sinl_Rep(a))|(? b.f=Sinr_Rep(b))}"
-
- (*faking a type definition... *)
- (* "++" is isomorphic to Ssum *)
-
- Rep_Ssum "Rep_Ssum(p):Ssum"
- Rep_Ssum_inverse "Abs_Ssum(Rep_Ssum(p)) = p"
- Abs_Ssum_inverse "f:Ssum ==> Rep_Ssum(Abs_Ssum(f)) = f"
-
- (*defining the abstract constants*)
- Isinl_def "Isinl(a) == Abs_Ssum(Sinl_Rep(a))"
- Isinr_def "Isinr(b) == Abs_Ssum(Sinr_Rep(b))"
-
- Iwhen_def "Iwhen(f)(g)(s) == @z.\
-\ (s=Isinl(UU) --> z=UU)\
-\ &(!a. ~a=UU & s=Isinl(a) --> z=f[a])\
-\ &(!b. ~b=UU & s=Isinr(b) --> z=g[b])"
-
-end
-
--- a/src/HOLCF/ssum1.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,353 +0,0 @@
-(* Title: HOLCF/ssum1.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for theory ssum1.thy
-*)
-
-open Ssum1;
-
-local
-
-fun eq_left s1 s2 =
- (
- (res_inst_tac [("s",s1),("t",s2)] (inject_Isinl RS subst) 1)
- THEN (rtac trans 1)
- THEN (atac 2)
- THEN (etac sym 1));
-
-fun eq_right s1 s2 =
- (
- (res_inst_tac [("s",s1),("t",s2)] (inject_Isinr RS subst) 1)
- THEN (rtac trans 1)
- THEN (atac 2)
- THEN (etac sym 1));
-
-fun UU_left s1 =
- (
- (res_inst_tac [("t",s1)](noteq_IsinlIsinr RS conjunct1 RS ssubst)1)
- THEN (rtac trans 1)
- THEN (atac 2)
- THEN (etac sym 1));
-
-fun UU_right s1 =
- (
- (res_inst_tac [("t",s1)](noteq_IsinlIsinr RS conjunct2 RS ssubst)1)
- THEN (rtac trans 1)
- THEN (atac 2)
- THEN (etac sym 1))
-
-in
-
-val less_ssum1a = prove_goalw Ssum1.thy [less_ssum_def]
-"[|s1=Isinl(x); s2=Isinl(y)|] ==> less_ssum(s1,s2) = (x << y)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (dtac conjunct1 2),
- (dtac spec 2),
- (dtac spec 2),
- (etac mp 2),
- (fast_tac HOL_cs 2),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (eq_left "x" "u"),
- (eq_left "y" "xa"),
- (rtac refl 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_left "x"),
- (UU_right "v"),
- (simp_tac Cfun_ss 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (eq_left "x" "u"),
- (UU_left "y"),
- (rtac iffI 1),
- (etac UU_I 1),
- (res_inst_tac [("s","x"),("t","UU")] subst 1),
- (atac 1),
- (rtac refl_less 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_left "x"),
- (UU_right "v"),
- (simp_tac Cfun_ss 1)
- ]);
-
-
-val less_ssum1b = prove_goalw Ssum1.thy [less_ssum_def]
-"[|s1=Isinr(x); s2=Isinr(y)|] ==> less_ssum(s1,s2) = (x << y)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (dtac conjunct2 2),
- (dtac conjunct1 2),
- (dtac spec 2),
- (dtac spec 2),
- (etac mp 2),
- (fast_tac HOL_cs 2),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_right "x"),
- (UU_left "u"),
- (simp_tac Cfun_ss 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (eq_right "x" "v"),
- (eq_right "y" "ya"),
- (rtac refl 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_right "x"),
- (UU_left "u"),
- (simp_tac Cfun_ss 1),
- (strip_tac 1),
- (etac conjE 1),
- (eq_right "x" "v"),
- (UU_right "y"),
- (rtac iffI 1),
- (etac UU_I 1),
- (res_inst_tac [("s","UU"),("t","x")] subst 1),
- (etac sym 1),
- (rtac refl_less 1)
- ]);
-
-
-val less_ssum1c = prove_goalw Ssum1.thy [less_ssum_def]
-"[|s1=Isinl(x); s2=Isinr(y)|] ==> less_ssum(s1,s2) = (x = UU)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (eq_left "x" "u"),
- (UU_left "xa"),
- (rtac iffI 1),
- (res_inst_tac [("s","x"),("t","UU")] subst 1),
- (atac 1),
- (rtac refl_less 1),
- (etac UU_I 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_left "x"),
- (UU_right "v"),
- (simp_tac Cfun_ss 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (eq_left "x" "u"),
- (rtac refl 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_left "x"),
- (UU_right "v"),
- (simp_tac Cfun_ss 1),
- (dtac conjunct2 1),
- (dtac conjunct2 1),
- (dtac conjunct1 1),
- (dtac spec 1),
- (dtac spec 1),
- (etac mp 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val less_ssum1d = prove_goalw Ssum1.thy [less_ssum_def]
-"[|s1=Isinr(x); s2=Isinl(y)|] ==> less_ssum(s1,s2) = (x = UU)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac select_equality 1),
- (dtac conjunct2 2),
- (dtac conjunct2 2),
- (dtac conjunct2 2),
- (dtac spec 2),
- (dtac spec 2),
- (etac mp 2),
- (fast_tac HOL_cs 2),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_right "x"),
- (UU_left "u"),
- (simp_tac Cfun_ss 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_right "ya"),
- (eq_right "x" "v"),
- (rtac iffI 1),
- (etac UU_I 2),
- (res_inst_tac [("s","UU"),("t","x")] subst 1),
- (etac sym 1),
- (rtac refl_less 1),
- (rtac conjI 1),
- (strip_tac 1),
- (etac conjE 1),
- (UU_right "x"),
- (UU_left "u"),
- (simp_tac HOL_ss 1),
- (strip_tac 1),
- (etac conjE 1),
- (eq_right "x" "v"),
- (rtac refl 1)
- ])
-end;
-
-
-(* ------------------------------------------------------------------------ *)
-(* optimize lemmas about less_ssum *)
-(* ------------------------------------------------------------------------ *)
-
-val less_ssum2a = prove_goal Ssum1.thy
- "less_ssum(Isinl(x),Isinl(y)) = (x << y)"
- (fn prems =>
- [
- (rtac less_ssum1a 1),
- (rtac refl 1),
- (rtac refl 1)
- ]);
-
-val less_ssum2b = prove_goal Ssum1.thy
- "less_ssum(Isinr(x),Isinr(y)) = (x << y)"
- (fn prems =>
- [
- (rtac less_ssum1b 1),
- (rtac refl 1),
- (rtac refl 1)
- ]);
-
-val less_ssum2c = prove_goal Ssum1.thy
- "less_ssum(Isinl(x),Isinr(y)) = (x = UU)"
- (fn prems =>
- [
- (rtac less_ssum1c 1),
- (rtac refl 1),
- (rtac refl 1)
- ]);
-
-val less_ssum2d = prove_goal Ssum1.thy
- "less_ssum(Isinr(x),Isinl(y)) = (x = UU)"
- (fn prems =>
- [
- (rtac less_ssum1d 1),
- (rtac refl 1),
- (rtac refl 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* less_ssum is a partial order on ++ *)
-(* ------------------------------------------------------------------------ *)
-
-val refl_less_ssum = prove_goal Ssum1.thy "less_ssum(p,p)"
- (fn prems =>
- [
- (res_inst_tac [("p","p")] IssumE2 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum2a RS iffD2) 1),
- (rtac refl_less 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum2b RS iffD2) 1),
- (rtac refl_less 1)
- ]);
-
-val antisym_less_ssum = prove_goal Ssum1.thy
- "[|less_ssum(p1,p2);less_ssum(p2,p1)|] ==> p1=p2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] IssumE2 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] IssumE2 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("f","Isinl")] arg_cong 1),
- (rtac antisym_less 1),
- (etac (less_ssum2a RS iffD1) 1),
- (etac (less_ssum2a RS iffD1) 1),
- (hyp_subst_tac 1),
- (etac (less_ssum2d RS iffD1 RS ssubst) 1),
- (etac (less_ssum2c RS iffD1 RS ssubst) 1),
- (rtac strict_IsinlIsinr 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p2")] IssumE2 1),
- (hyp_subst_tac 1),
- (etac (less_ssum2c RS iffD1 RS ssubst) 1),
- (etac (less_ssum2d RS iffD1 RS ssubst) 1),
- (rtac (strict_IsinlIsinr RS sym) 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("f","Isinr")] arg_cong 1),
- (rtac antisym_less 1),
- (etac (less_ssum2b RS iffD1) 1),
- (etac (less_ssum2b RS iffD1) 1)
- ]);
-
-val trans_less_ssum = prove_goal Ssum1.thy
- "[|less_ssum(p1,p2);less_ssum(p2,p3)|] ==> less_ssum(p1,p3)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","p1")] IssumE2 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p3")] IssumE2 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum2a RS iffD2) 1),
- (res_inst_tac [("p","p2")] IssumE2 1),
- (hyp_subst_tac 1),
- (rtac trans_less 1),
- (etac (less_ssum2a RS iffD1) 1),
- (etac (less_ssum2a RS iffD1) 1),
- (hyp_subst_tac 1),
- (etac (less_ssum2c RS iffD1 RS ssubst) 1),
- (rtac minimal 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum2c RS iffD2) 1),
- (res_inst_tac [("p","p2")] IssumE2 1),
- (hyp_subst_tac 1),
- (rtac UU_I 1),
- (rtac trans_less 1),
- (etac (less_ssum2a RS iffD1) 1),
- (rtac (antisym_less_inverse RS conjunct1) 1),
- (etac (less_ssum2c RS iffD1) 1),
- (hyp_subst_tac 1),
- (etac (less_ssum2c RS iffD1) 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","p3")] IssumE2 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum2d RS iffD2) 1),
- (res_inst_tac [("p","p2")] IssumE2 1),
- (hyp_subst_tac 1),
- (etac (less_ssum2d RS iffD1) 1),
- (hyp_subst_tac 1),
- (rtac UU_I 1),
- (rtac trans_less 1),
- (etac (less_ssum2b RS iffD1) 1),
- (rtac (antisym_less_inverse RS conjunct1) 1),
- (etac (less_ssum2d RS iffD1) 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum2b RS iffD2) 1),
- (res_inst_tac [("p","p2")] IssumE2 1),
- (hyp_subst_tac 1),
- (etac (less_ssum2d RS iffD1 RS ssubst) 1),
- (rtac minimal 1),
- (hyp_subst_tac 1),
- (rtac trans_less 1),
- (etac (less_ssum2b RS iffD1) 1),
- (etac (less_ssum2b RS iffD1) 1)
- ]);
-
-
-
--- a/src/HOLCF/ssum1.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-(* Title: HOLCF/ssum1.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Partial ordering for the strict sum ++
-*)
-
-Ssum1 = Ssum0 +
-
-consts
-
- less_ssum :: "[('a ++ 'b),('a ++ 'b)] => bool"
-
-rules
-
- less_ssum_def "less_ssum(s1,s2) == (@z.\
-\ (! u x.s1=Isinl(u) & s2=Isinl(x) --> z = (u << x))\
-\ &(! v y.s1=Isinr(v) & s2=Isinr(y) --> z = (v << y))\
-\ &(! u y.s1=Isinl(u) & s2=Isinr(y) --> z = (u = UU))\
-\ &(! v x.s1=Isinr(v) & s2=Isinl(x) --> z = (v = UU)))"
-
-end
-
-
--- a/src/HOLCF/ssum2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,419 +0,0 @@
-(* Title: HOLCF/ssum2.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for ssum2.thy
-*)
-
-open Ssum2;
-
-(* ------------------------------------------------------------------------ *)
-(* access to less_ssum in class po *)
-(* ------------------------------------------------------------------------ *)
-
-val less_ssum3a = prove_goal Ssum2.thy
- "(Isinl(x) << Isinl(y)) = (x << y)"
- (fn prems =>
- [
- (rtac (inst_ssum_po RS ssubst) 1),
- (rtac less_ssum2a 1)
- ]);
-
-val less_ssum3b = prove_goal Ssum2.thy
- "(Isinr(x) << Isinr(y)) = (x << y)"
- (fn prems =>
- [
- (rtac (inst_ssum_po RS ssubst) 1),
- (rtac less_ssum2b 1)
- ]);
-
-val less_ssum3c = prove_goal Ssum2.thy
- "(Isinl(x) << Isinr(y)) = (x = UU)"
- (fn prems =>
- [
- (rtac (inst_ssum_po RS ssubst) 1),
- (rtac less_ssum2c 1)
- ]);
-
-val less_ssum3d = prove_goal Ssum2.thy
- "(Isinr(x) << Isinl(y)) = (x = UU)"
- (fn prems =>
- [
- (rtac (inst_ssum_po RS ssubst) 1),
- (rtac less_ssum2d 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* type ssum ++ is pointed *)
-(* ------------------------------------------------------------------------ *)
-
-val minimal_ssum = prove_goal Ssum2.thy "Isinl(UU) << s"
- (fn prems =>
- [
- (res_inst_tac [("p","s")] IssumE2 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum3a RS iffD2) 1),
- (rtac minimal 1),
- (hyp_subst_tac 1),
- (rtac (strict_IsinlIsinr RS ssubst) 1),
- (rtac (less_ssum3b RS iffD2) 1),
- (rtac minimal 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Isinl, Isinr are monotone *)
-(* ------------------------------------------------------------------------ *)
-
-val monofun_Isinl = prove_goalw Ssum2.thy [monofun] "monofun(Isinl)"
- (fn prems =>
- [
- (strip_tac 1),
- (etac (less_ssum3a RS iffD2) 1)
- ]);
-
-val monofun_Isinr = prove_goalw Ssum2.thy [monofun] "monofun(Isinr)"
- (fn prems =>
- [
- (strip_tac 1),
- (etac (less_ssum3b RS iffD2) 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* Iwhen is monotone in all arguments *)
-(* ------------------------------------------------------------------------ *)
-
-
-val monofun_Iwhen1 = prove_goalw Ssum2.thy [monofun] "monofun(Iwhen)"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (res_inst_tac [("p","xb")] IssumE 1),
- (hyp_subst_tac 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1),
- (etac monofun_cfun_fun 1),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-val monofun_Iwhen2 = prove_goalw Ssum2.thy [monofun] "monofun(Iwhen(f))"
- (fn prems =>
- [
- (strip_tac 1),
- (rtac (less_fun RS iffD2) 1),
- (strip_tac 1),
- (res_inst_tac [("p","xa")] IssumE 1),
- (hyp_subst_tac 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1),
- (etac monofun_cfun_fun 1)
- ]);
-
-val monofun_Iwhen3 = prove_goalw Ssum2.thy [monofun] "monofun(Iwhen(f)(g))"
- (fn prems =>
- [
- (strip_tac 1),
- (res_inst_tac [("p","x")] IssumE 1),
- (hyp_subst_tac 1),
- (asm_simp_tac Ssum_ss 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","y")] IssumE 1),
- (hyp_subst_tac 1),
- (asm_simp_tac Ssum_ss 1),
- (res_inst_tac [("P","xa=UU")] notE 1),
- (atac 1),
- (rtac UU_I 1),
- (rtac (less_ssum3a RS iffD1) 1),
- (atac 1),
- (hyp_subst_tac 1),
- (asm_simp_tac Ssum_ss 1),
- (rtac monofun_cfun_arg 1),
- (etac (less_ssum3a RS iffD1) 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("s","UU"),("t","xa")] subst 1),
- (etac (less_ssum3c RS iffD1 RS sym) 1),
- (asm_simp_tac Ssum_ss 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("p","y")] IssumE 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("s","UU"),("t","ya")] subst 1),
- (etac (less_ssum3d RS iffD1 RS sym) 1),
- (asm_simp_tac Ssum_ss 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("s","UU"),("t","ya")] subst 1),
- (etac (less_ssum3d RS iffD1 RS sym) 1),
- (asm_simp_tac Ssum_ss 1),
- (hyp_subst_tac 1),
- (asm_simp_tac Ssum_ss 1),
- (rtac monofun_cfun_arg 1),
- (etac (less_ssum3b RS iffD1) 1)
- ]);
-
-
-
-
-(* ------------------------------------------------------------------------ *)
-(* some kind of exhaustion rules for chains in 'a ++ 'b *)
-(* ------------------------------------------------------------------------ *)
-
-
-val ssum_lemma1 = prove_goal Ssum2.thy
-"[|~(!i.? x.Y(i::nat)=Isinl(x))|] ==> (? i.! x.~Y(i)=Isinl(x))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val ssum_lemma2 = prove_goal Ssum2.thy
-"[|(? i.!x.~(Y::nat => 'a++'b)(i::nat)=Isinl(x::'a))|] ==>\
-\ (? i y. (Y::nat => 'a++'b)(i::nat)=Isinr(y::'b) & ~y=UU)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exE 1),
- (res_inst_tac [("p","Y(i)")] IssumE 1),
- (dtac spec 1),
- (contr_tac 1),
- (dtac spec 1),
- (contr_tac 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
-val ssum_lemma3 = prove_goal Ssum2.thy
-"[|is_chain(Y);(? i x. Y(i)=Isinr(x) & ~x=UU)|] ==> (!i.? y.Y(i)=Isinr(y))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac exE 1),
- (etac exE 1),
- (rtac allI 1),
- (res_inst_tac [("p","Y(ia)")] IssumE 1),
- (rtac exI 1),
- (rtac trans 1),
- (rtac strict_IsinlIsinr 2),
- (atac 1),
- (etac exI 2),
- (etac conjE 1),
- (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1),
- (hyp_subst_tac 2),
- (etac exI 2),
- (res_inst_tac [("P","x=UU")] notE 1),
- (atac 1),
- (rtac (less_ssum3d RS iffD1) 1),
- (res_inst_tac [("s","Y(i)"),("t","Isinr(x)")] subst 1),
- (atac 1),
- (res_inst_tac [("s","Y(ia)"),("t","Isinl(xa)")] subst 1),
- (atac 1),
- (etac (chain_mono RS mp) 1),
- (atac 1),
- (res_inst_tac [("P","xa=UU")] notE 1),
- (atac 1),
- (rtac (less_ssum3c RS iffD1) 1),
- (res_inst_tac [("s","Y(i)"),("t","Isinr(x)")] subst 1),
- (atac 1),
- (res_inst_tac [("s","Y(ia)"),("t","Isinl(xa)")] subst 1),
- (atac 1),
- (etac (chain_mono RS mp) 1),
- (atac 1)
- ]);
-
-val ssum_lemma4 = prove_goal Ssum2.thy
-"is_chain(Y) ==> (!i.? x.Y(i)=Isinl(x))|(!i.? y.Y(i)=Isinr(y))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac classical2 1),
- (etac disjI1 1),
- (rtac disjI2 1),
- (etac ssum_lemma3 1),
- (rtac ssum_lemma2 1),
- (etac ssum_lemma1 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* restricted surjectivity of Isinl *)
-(* ------------------------------------------------------------------------ *)
-
-val ssum_lemma5 = prove_goal Ssum2.thy
-"z=Isinl(x)==> Isinl((Iwhen (LAM x.x) (LAM y.UU))(z)) = z"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("Q","x=UU")] classical2 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* restricted surjectivity of Isinr *)
-(* ------------------------------------------------------------------------ *)
-
-val ssum_lemma6 = prove_goal Ssum2.thy
-"z=Isinr(x)==> Isinr((Iwhen (LAM y.UU) (LAM x.x))(z)) = z"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (hyp_subst_tac 1),
- (res_inst_tac [("Q","x=UU")] classical2 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* technical lemmas *)
-(* ------------------------------------------------------------------------ *)
-
-val ssum_lemma7 = prove_goal Ssum2.thy
-"[|Isinl(x) << z; ~x=UU|] ==> ? y.z=Isinl(y) & ~y=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","z")] IssumE 1),
- (hyp_subst_tac 1),
- (etac notE 1),
- (rtac antisym_less 1),
- (etac (less_ssum3a RS iffD1) 1),
- (rtac minimal 1),
- (fast_tac HOL_cs 1),
- (hyp_subst_tac 1),
- (rtac notE 1),
- (etac (less_ssum3c RS iffD1) 2),
- (atac 1)
- ]);
-
-val ssum_lemma8 = prove_goal Ssum2.thy
-"[|Isinr(x) << z; ~x=UU|] ==> ? y.z=Isinr(y) & ~y=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (res_inst_tac [("p","z")] IssumE 1),
- (hyp_subst_tac 1),
- (etac notE 1),
- (etac (less_ssum3d RS iffD1) 1),
- (hyp_subst_tac 1),
- (rtac notE 1),
- (etac (less_ssum3d RS iffD1) 2),
- (atac 1),
- (fast_tac HOL_cs 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* the type 'a ++ 'b is a cpo in three steps *)
-(* ------------------------------------------------------------------------ *)
-
-val lub_ssum1a = prove_goal Ssum2.thy
-"[|is_chain(Y);(!i.? x.Y(i)=Isinl(x))|] ==>\
-\ range(Y) <<|\
-\ Isinl(lub(range(%i.(Iwhen (LAM x.x) (LAM y.UU))(Y(i)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (etac allE 1),
- (etac exE 1),
- (res_inst_tac [("t","Y(i)")] (ssum_lemma5 RS subst) 1),
- (atac 1),
- (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (strip_tac 1),
- (res_inst_tac [("p","u")] IssumE2 1),
- (res_inst_tac [("t","u")] (ssum_lemma5 RS subst) 1),
- (atac 1),
- (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1),
- (rtac is_lub_thelub 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (etac (monofun_Iwhen3 RS ub2ub_monofun) 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum3c RS iffD2) 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (res_inst_tac [("p","Y(i)")] IssumE 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 2),
- (etac notE 1),
- (rtac (less_ssum3c RS iffD1) 1),
- (res_inst_tac [("t","Isinl(x)")] subst 1),
- (atac 1),
- (etac (ub_rangeE RS spec) 1)
- ]);
-
-
-val lub_ssum1b = prove_goal Ssum2.thy
-"[|is_chain(Y);(!i.? x.Y(i)=Isinr(x))|] ==>\
-\ range(Y) <<|\
-\ Isinr(lub(range(%i.(Iwhen (LAM y.UU) (LAM x.x))(Y(i)))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac is_lubI 1),
- (rtac conjI 1),
- (rtac ub_rangeI 1),
- (rtac allI 1),
- (etac allE 1),
- (etac exE 1),
- (res_inst_tac [("t","Y(i)")] (ssum_lemma6 RS subst) 1),
- (atac 1),
- (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1),
- (rtac is_ub_thelub 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (strip_tac 1),
- (res_inst_tac [("p","u")] IssumE2 1),
- (hyp_subst_tac 1),
- (rtac (less_ssum3d RS iffD2) 1),
- (rtac chain_UU_I_inverse 1),
- (rtac allI 1),
- (res_inst_tac [("p","Y(i)")] IssumE 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1),
- (etac notE 1),
- (rtac (less_ssum3d RS iffD1) 1),
- (res_inst_tac [("t","Isinr(y)")] subst 1),
- (atac 1),
- (etac (ub_rangeE RS spec) 1),
- (res_inst_tac [("t","u")] (ssum_lemma6 RS subst) 1),
- (atac 1),
- (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1),
- (rtac is_lub_thelub 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (etac (monofun_Iwhen3 RS ub2ub_monofun) 1)
- ]);
-
-
-val thelub_ssum1a = lub_ssum1a RS thelubI;
-(* [| is_chain(?Y1); ! i. ? x. ?Y1(i) = Isinl(x) |] ==> *)
-(* lub(range(?Y1)) = Isinl(lub(range(%i. Iwhen(LAM x. x,LAM y. UU,?Y1(i)))))*)
-
-val thelub_ssum1b = lub_ssum1b RS thelubI;
-(* [| is_chain(?Y1); ! i. ? x. ?Y1(i) = Isinr(x) |] ==> *)
-(* lub(range(?Y1)) = Isinr(lub(range(%i. Iwhen(LAM y. UU,LAM x. x,?Y1(i)))))*)
-
-val cpo_ssum = prove_goal Ssum2.thy
- "is_chain(Y::nat=>'a ++'b) ==> ? x.range(Y) <<|x"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (ssum_lemma4 RS disjE) 1),
- (atac 1),
- (rtac exI 1),
- (etac lub_ssum1a 1),
- (atac 1),
- (rtac exI 1),
- (etac lub_ssum1b 1),
- (atac 1)
- ]);
--- a/src/HOLCF/ssum2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,23 +0,0 @@
-(* Title: HOLCF/ssum2.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class Instance ++::(pcpo,pcpo)po
-*)
-
-Ssum2 = Ssum1 +
-
-arities "++" :: (pcpo,pcpo)po
-(* Witness for the above arity axiom is ssum1.ML *)
-
-rules
-
-(* instance of << for type ['a ++ 'b] *)
-
-inst_ssum_po "(op <<)::['a ++ 'b,'a ++ 'b]=>bool = less_ssum"
-
-end
-
-
-
--- a/src/HOLCF/ssum3.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,728 +0,0 @@
-(* Title: HOLCF/ssum3.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for ssum3.thy
-*)
-
-open Ssum3;
-
-(* ------------------------------------------------------------------------ *)
-(* continuity for Isinl and Isinr *)
-(* ------------------------------------------------------------------------ *)
-
-
-val contlub_Isinl = prove_goal Ssum3.thy "contlub(Isinl)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_ssum1a RS sym) 2),
- (rtac allI 3),
- (rtac exI 3),
- (rtac refl 3),
- (etac (monofun_Isinl RS ch2ch_monofun) 2),
- (res_inst_tac [("Q","lub(range(Y))=UU")] classical2 1),
- (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1),
- (atac 1),
- (res_inst_tac [("f","Isinl")] arg_cong 1),
- (rtac (chain_UU_I_inverse RS sym) 1),
- (rtac allI 1),
- (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1),
- (etac (chain_UU_I RS spec ) 1),
- (atac 1),
- (rtac Iwhen1 1),
- (res_inst_tac [("f","Isinl")] arg_cong 1),
- (rtac lub_equal 1),
- (atac 1),
- (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (etac (monofun_Isinl RS ch2ch_monofun) 1),
- (rtac allI 1),
- (res_inst_tac [("Q","Y(k)=UU")] classical2 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-val contlub_Isinr = prove_goal Ssum3.thy "contlub(Isinr)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_ssum1b RS sym) 2),
- (rtac allI 3),
- (rtac exI 3),
- (rtac refl 3),
- (etac (monofun_Isinr RS ch2ch_monofun) 2),
- (res_inst_tac [("Q","lub(range(Y))=UU")] classical2 1),
- (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1),
- (atac 1),
- ((rtac arg_cong 1) THEN (rtac (chain_UU_I_inverse RS sym) 1)),
- (rtac allI 1),
- (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1),
- (etac (chain_UU_I RS spec ) 1),
- (atac 1),
- (rtac (strict_IsinlIsinr RS subst) 1),
- (rtac Iwhen1 1),
- ((rtac arg_cong 1) THEN (rtac lub_equal 1)),
- (atac 1),
- (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (etac (monofun_Isinr RS ch2ch_monofun) 1),
- (rtac allI 1),
- (res_inst_tac [("Q","Y(k)=UU")] classical2 1),
- (asm_simp_tac Ssum_ss 1),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-val contX_Isinl = prove_goal Ssum3.thy "contX(Isinl)"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Isinl 1),
- (rtac contlub_Isinl 1)
- ]);
-
-val contX_Isinr = prove_goal Ssum3.thy "contX(Isinr)"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Isinr 1),
- (rtac contlub_Isinr 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* continuity for Iwhen in the firts two arguments *)
-(* ------------------------------------------------------------------------ *)
-
-val contlub_Iwhen1 = prove_goal Ssum3.thy "contlub(Iwhen)"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_fun RS sym) 2),
- (etac (monofun_Iwhen1 RS ch2ch_monofun) 2),
- (rtac (expand_fun_eq RS iffD2) 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_fun RS sym) 2),
- (rtac ch2ch_fun 2),
- (etac (monofun_Iwhen1 RS ch2ch_monofun) 2),
- (rtac (expand_fun_eq RS iffD2) 1),
- (strip_tac 1),
- (res_inst_tac [("p","xa")] IssumE 1),
- (asm_simp_tac Ssum_ss 1),
- (rtac (lub_const RS thelubI RS sym) 1),
- (asm_simp_tac Ssum_ss 1),
- (etac contlub_cfun_fun 1),
- (asm_simp_tac Ssum_ss 1),
- (rtac (lub_const RS thelubI RS sym) 1)
- ]);
-
-val contlub_Iwhen2 = prove_goal Ssum3.thy "contlub(Iwhen(f))"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (rtac trans 1),
- (rtac (thelub_fun RS sym) 2),
- (etac (monofun_Iwhen2 RS ch2ch_monofun) 2),
- (rtac (expand_fun_eq RS iffD2) 1),
- (strip_tac 1),
- (res_inst_tac [("p","x")] IssumE 1),
- (asm_simp_tac Ssum_ss 1),
- (rtac (lub_const RS thelubI RS sym) 1),
- (asm_simp_tac Ssum_ss 1),
- (rtac (lub_const RS thelubI RS sym) 1),
- (asm_simp_tac Ssum_ss 1),
- (etac contlub_cfun_fun 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* continuity for Iwhen in its third argument *)
-(* ------------------------------------------------------------------------ *)
-
-(* ------------------------------------------------------------------------ *)
-(* first 5 ugly lemmas *)
-(* ------------------------------------------------------------------------ *)
-
-val ssum_lemma9 = prove_goal Ssum3.thy
-"[| is_chain(Y); lub(range(Y)) = Isinl(x)|] ==> !i.? x.Y(i)=Isinl(x)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (res_inst_tac [("p","Y(i)")] IssumE 1),
- (etac exI 1),
- (etac exI 1),
- (res_inst_tac [("P","y=UU")] notE 1),
- (atac 1),
- (rtac (less_ssum3d RS iffD1) 1),
- (etac subst 1),
- (etac subst 1),
- (etac is_ub_thelub 1)
- ]);
-
-
-val ssum_lemma10 = prove_goal Ssum3.thy
-"[| is_chain(Y); lub(range(Y)) = Isinr(x)|] ==> !i.? x.Y(i)=Isinr(x)"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (strip_tac 1),
- (res_inst_tac [("p","Y(i)")] IssumE 1),
- (rtac exI 1),
- (etac trans 1),
- (rtac strict_IsinlIsinr 1),
- (etac exI 2),
- (res_inst_tac [("P","xa=UU")] notE 1),
- (atac 1),
- (rtac (less_ssum3c RS iffD1) 1),
- (etac subst 1),
- (etac subst 1),
- (etac is_ub_thelub 1)
- ]);
-
-val ssum_lemma11 = prove_goal Ssum3.thy
-"[| is_chain(Y); lub(range(Y)) = Isinl(UU) |] ==>\
-\ Iwhen(f,g,lub(range(Y))) = lub(range(%i. Iwhen(f,g,Y(i))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac Ssum_ss 1),
- (rtac (chain_UU_I_inverse RS sym) 1),
- (rtac allI 1),
- (res_inst_tac [("s","Isinl(UU)"),("t","Y(i)")] subst 1),
- (rtac (inst_ssum_pcpo RS subst) 1),
- (rtac (chain_UU_I RS spec RS sym) 1),
- (atac 1),
- (etac (inst_ssum_pcpo RS ssubst) 1),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-val ssum_lemma12 = prove_goal Ssum3.thy
-"[| is_chain(Y); lub(range(Y)) = Isinl(x); ~ x = UU |] ==>\
-\ Iwhen(f,g,lub(range(Y))) = lub(range(%i. Iwhen(f,g,Y(i))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac Ssum_ss 1),
- (res_inst_tac [("t","x")] subst 1),
- (rtac inject_Isinl 1),
- (rtac trans 1),
- (atac 2),
- (rtac (thelub_ssum1a RS sym) 1),
- (atac 1),
- (etac ssum_lemma9 1),
- (atac 1),
- (rtac trans 1),
- (rtac contlub_cfun_arg 1),
- (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (atac 1),
- (rtac lub_equal2 1),
- (rtac (chain_mono2 RS exE) 1),
- (atac 2),
- (rtac chain_UU_I_inverse2 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (etac swap 1),
- (rtac inject_Isinl 1),
- (rtac trans 1),
- (etac sym 1),
- (etac notnotD 1),
- (rtac exI 1),
- (strip_tac 1),
- (rtac (ssum_lemma9 RS spec RS exE) 1),
- (atac 1),
- (atac 1),
- (res_inst_tac [("t","Y(i)")] ssubst 1),
- (atac 1),
- (rtac trans 1),
- (rtac cfun_arg_cong 1),
- (rtac Iwhen2 1),
- (res_inst_tac [("P","Y(i)=UU")] swap 1),
- (fast_tac HOL_cs 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (res_inst_tac [("t","Y(i)")] ssubst 1),
- (atac 1),
- (fast_tac HOL_cs 1),
- (rtac (Iwhen2 RS ssubst) 1),
- (res_inst_tac [("P","Y(i)=UU")] swap 1),
- (fast_tac HOL_cs 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (res_inst_tac [("t","Y(i)")] ssubst 1),
- (atac 1),
- (fast_tac HOL_cs 1),
- (simp_tac Cfun_ss 1),
- (rtac (monofun_fapp2 RS ch2ch_monofun) 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1)
- ]);
-
-
-val ssum_lemma13 = prove_goal Ssum3.thy
-"[| is_chain(Y); lub(range(Y)) = Isinr(x); ~ x = UU |] ==>\
-\ Iwhen(f,g,lub(range(Y))) = lub(range(%i. Iwhen(f,g,Y(i))))"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac Ssum_ss 1),
- (res_inst_tac [("t","x")] subst 1),
- (rtac inject_Isinr 1),
- (rtac trans 1),
- (atac 2),
- (rtac (thelub_ssum1b RS sym) 1),
- (atac 1),
- (etac ssum_lemma10 1),
- (atac 1),
- (rtac trans 1),
- (rtac contlub_cfun_arg 1),
- (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (atac 1),
- (rtac lub_equal2 1),
- (rtac (chain_mono2 RS exE) 1),
- (atac 2),
- (rtac chain_UU_I_inverse2 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (etac swap 1),
- (rtac inject_Isinr 1),
- (rtac trans 1),
- (etac sym 1),
- (rtac (strict_IsinlIsinr RS subst) 1),
- (etac notnotD 1),
- (rtac exI 1),
- (strip_tac 1),
- (rtac (ssum_lemma10 RS spec RS exE) 1),
- (atac 1),
- (atac 1),
- (res_inst_tac [("t","Y(i)")] ssubst 1),
- (atac 1),
- (rtac trans 1),
- (rtac cfun_arg_cong 1),
- (rtac Iwhen3 1),
- (res_inst_tac [("P","Y(i)=UU")] swap 1),
- (fast_tac HOL_cs 1),
- (dtac notnotD 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (rtac (strict_IsinlIsinr RS ssubst) 1),
- (res_inst_tac [("t","Y(i)")] ssubst 1),
- (atac 1),
- (fast_tac HOL_cs 1),
- (rtac (Iwhen3 RS ssubst) 1),
- (res_inst_tac [("P","Y(i)=UU")] swap 1),
- (fast_tac HOL_cs 1),
- (dtac notnotD 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (rtac (strict_IsinlIsinr RS ssubst) 1),
- (res_inst_tac [("t","Y(i)")] ssubst 1),
- (atac 1),
- (fast_tac HOL_cs 1),
- (simp_tac Cfun_ss 1),
- (rtac (monofun_fapp2 RS ch2ch_monofun) 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1),
- (etac (monofun_Iwhen3 RS ch2ch_monofun) 1)
- ]);
-
-
-val contlub_Iwhen3 = prove_goal Ssum3.thy "contlub(Iwhen(f)(g))"
- (fn prems =>
- [
- (rtac contlubI 1),
- (strip_tac 1),
- (res_inst_tac [("p","lub(range(Y))")] IssumE 1),
- (etac ssum_lemma11 1),
- (atac 1),
- (etac ssum_lemma12 1),
- (atac 1),
- (atac 1),
- (etac ssum_lemma13 1),
- (atac 1),
- (atac 1)
- ]);
-
-val contX_Iwhen1 = prove_goal Ssum3.thy "contX(Iwhen)"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Iwhen1 1),
- (rtac contlub_Iwhen1 1)
- ]);
-
-val contX_Iwhen2 = prove_goal Ssum3.thy "contX(Iwhen(f))"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Iwhen2 1),
- (rtac contlub_Iwhen2 1)
- ]);
-
-val contX_Iwhen3 = prove_goal Ssum3.thy "contX(Iwhen(f)(g))"
- (fn prems =>
- [
- (rtac monocontlub2contX 1),
- (rtac monofun_Iwhen3 1),
- (rtac contlub_Iwhen3 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* continuous versions of lemmas for 'a ++ 'b *)
-(* ------------------------------------------------------------------------ *)
-
-val strict_sinl = prove_goalw Ssum3.thy [sinl_def] "sinl[UU]=UU"
- (fn prems =>
- [
- (simp_tac (Ssum_ss addsimps [contX_Isinl]) 1),
- (rtac (inst_ssum_pcpo RS sym) 1)
- ]);
-
-val strict_sinr = prove_goalw Ssum3.thy [sinr_def] "sinr[UU]=UU"
- (fn prems =>
- [
- (simp_tac (Ssum_ss addsimps [contX_Isinr]) 1),
- (rtac (inst_ssum_pcpo RS sym) 1)
- ]);
-
-val noteq_sinlsinr = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "sinl[a]=sinr[b] ==> a=UU & b=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac noteq_IsinlIsinr 1),
- (etac box_equals 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1)
- ]);
-
-val inject_sinl = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "sinl[a1]=sinl[a2]==> a1=a2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac inject_Isinl 1),
- (etac box_equals 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1)
- ]);
-
-val inject_sinr = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "sinr[a1]=sinr[a2]==> a1=a2"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac inject_Isinr 1),
- (etac box_equals 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1)
- ]);
-
-
-val defined_sinl = prove_goal Ssum3.thy
- "~x=UU ==> ~sinl[x]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac swap 1),
- (rtac inject_sinl 1),
- (rtac (strict_sinl RS ssubst) 1),
- (etac notnotD 1)
- ]);
-
-val defined_sinr = prove_goal Ssum3.thy
- "~x=UU ==> ~sinr[x]=UU"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (etac swap 1),
- (rtac inject_sinr 1),
- (rtac (strict_sinr RS ssubst) 1),
- (etac notnotD 1)
- ]);
-
-val Exh_Ssum1 = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "z=UU | (? a. z=sinl[a] & ~a=UU) | (? b. z=sinr[b] & ~b=UU)"
- (fn prems =>
- [
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (rtac Exh_Ssum 1)
- ]);
-
-
-val ssumE = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "[|p=UU ==> Q ;\
-\ !!x.[|p=sinl[x]; ~x=UU |] ==> Q;\
-\ !!y.[|p=sinr[y]; ~y=UU |] ==> Q|] ==> Q"
- (fn prems =>
- [
- (rtac IssumE 1),
- (resolve_tac prems 1),
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (atac 1),
- (resolve_tac prems 1),
- (atac 2),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1),
- (resolve_tac prems 1),
- (atac 2),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1)
- ]);
-
-
-val ssumE2 = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "[|!!x.[|p=sinl[x]|] ==> Q;\
-\ !!y.[|p=sinr[y]|] ==> Q|] ==> Q"
- (fn prems =>
- [
- (rtac IssumE2 1),
- (resolve_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isinl 1),
- (atac 1),
- (resolve_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (rtac contX_Isinr 1),
- (atac 1)
- ]);
-
-val when1 = prove_goalw Ssum3.thy [when_def,sinl_def,sinr_def]
- "when[f][g][UU] = UU"
- (fn prems =>
- [
- (rtac (inst_ssum_pcpo RS ssubst) 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX2contX_CF1L]) 1)),
- (simp_tac Ssum_ss 1)
- ]);
-
-val when2 = prove_goalw Ssum3.thy [when_def,sinl_def,sinr_def]
- "~x=UU==>when[f][g][sinl[x]] = f[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-
-
-val when3 = prove_goalw Ssum3.thy [when_def,sinl_def,sinr_def]
- "~x=UU==>when[f][g][sinr[x]] = g[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (asm_simp_tac Ssum_ss 1)
- ]);
-
-
-val less_ssum4a = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "(sinl[x] << sinl[y]) = (x << y)"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac less_ssum3a 1)
- ]);
-
-val less_ssum4b = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "(sinr[x] << sinr[y]) = (x << y)"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac less_ssum3b 1)
- ]);
-
-val less_ssum4c = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "(sinl[x] << sinr[y]) = (x = UU)"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac less_ssum3c 1)
- ]);
-
-val less_ssum4d = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "(sinr[x] << sinl[y]) = (x = UU)"
- (fn prems =>
- [
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac less_ssum3d 1)
- ]);
-
-val ssum_chainE = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "is_chain(Y) ==> (!i.? x.Y(i)=sinl[x])|(!i.? y.Y(i)=sinr[y])"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1),
- (etac ssum_lemma4 1)
- ]);
-
-
-val thelub_ssum2a = prove_goalw Ssum3.thy [sinl_def,sinr_def,when_def]
-"[| is_chain(Y); !i.? x. Y(i) = sinl[x] |] ==>\
-\ lub(range(Y)) = sinl[lub(range(%i. when[LAM x. x][LAM y. UU][Y(i)]))]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac thelub_ssum1a 1),
- (atac 1),
- (rtac allI 1),
- (etac allE 1),
- (etac exE 1),
- (rtac exI 1),
- (etac box_equals 1),
- (rtac refl 1),
- (asm_simp_tac (Ssum_ss addsimps [contX_Isinl]) 1)
- ]);
-
-val thelub_ssum2b = prove_goalw Ssum3.thy [sinl_def,sinr_def,when_def]
-"[| is_chain(Y); !i.? x. Y(i) = sinr[x] |] ==>\
-\ lub(range(Y)) = sinr[lub(range(%i. when[LAM y. UU][LAM x. x][Y(i)]))]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac (beta_cfun RS ext RS ssubst) 1),
- (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)),
- (rtac thelub_ssum1b 1),
- (atac 1),
- (rtac allI 1),
- (etac allE 1),
- (etac exE 1),
- (rtac exI 1),
- (etac box_equals 1),
- (rtac refl 1),
- (asm_simp_tac (Ssum_ss addsimps
- [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3]) 1)
- ]);
-
-val thelub_ssum2a_rev = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "[| is_chain(Y); lub(range(Y)) = sinl[x]|] ==> !i.? x.Y(i)=sinl[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac (Ssum_ss addsimps
- [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3]) 1),
- (etac ssum_lemma9 1),
- (asm_simp_tac (Ssum_ss addsimps
- [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3]) 1)
- ]);
-
-val thelub_ssum2b_rev = prove_goalw Ssum3.thy [sinl_def,sinr_def]
- "[| is_chain(Y); lub(range(Y)) = sinr[x]|] ==> !i.? x.Y(i)=sinr[x]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (asm_simp_tac (Ssum_ss addsimps
- [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3]) 1),
- (etac ssum_lemma10 1),
- (asm_simp_tac (Ssum_ss addsimps
- [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2,
- contX_Iwhen3]) 1)
- ]);
-
-val thelub_ssum3 = prove_goal Ssum3.thy
-"is_chain(Y) ==>\
-\ lub(range(Y)) = sinl[lub(range(%i. when[LAM x. x][LAM y. UU][Y(i)]))]\
-\ | lub(range(Y)) = sinr[lub(range(%i. when[LAM y. UU][LAM x. x][Y(i)]))]"
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (ssum_chainE RS disjE) 1),
- (atac 1),
- (rtac disjI1 1),
- (etac thelub_ssum2a 1),
- (atac 1),
- (rtac disjI2 1),
- (etac thelub_ssum2b 1),
- (atac 1)
- ]);
-
-
-val when4 = prove_goal Ssum3.thy
- "when[sinl][sinr][z]=z"
- (fn prems =>
- [
- (res_inst_tac [("p","z")] ssumE 1),
- (asm_simp_tac (Cfun_ss addsimps [when1,when2,when3]) 1),
- (asm_simp_tac (Cfun_ss addsimps [when1,when2,when3]) 1),
- (asm_simp_tac (Cfun_ss addsimps [when1,when2,when3]) 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* install simplifier for Ssum *)
-(* ------------------------------------------------------------------------ *)
-
-val Ssum_rews = [strict_sinl,strict_sinr,when1,when2,when3];
-val Ssum_ss = Cfun_ss addsimps Ssum_rews;
--- a/src/HOLCF/ssum3.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,29 +0,0 @@
-(* Title: HOLCF/ssum3.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Class instance of ++ for class pcpo
-*)
-
-Ssum3 = Ssum2 +
-
-arities "++" :: (pcpo,pcpo)pcpo (* Witness ssum2.ML *)
-
-consts
- sinl :: "'a -> ('a++'b)"
- sinr :: "'b -> ('a++'b)"
- when :: "('a->'c)->('b->'c)->('a ++ 'b)-> 'c"
-
-rules
-
-inst_ssum_pcpo "UU::'a++'b = Isinl(UU)"
-
-sinl_def "sinl == (LAM x.Isinl(x))"
-sinr_def "sinr == (LAM x.Isinr(x))"
-when_def "when == (LAM f g s.Iwhen(f)(g)(s))"
-
-end
-
-
-
--- a/src/HOLCF/stream.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,791 +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 (HOLCF_ss 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 *)
-(* ------------------------------------------------------------------------*)
-
-val Exh_stream = prove_goalw Stream.thy [scons_def]
- "s = UU | (? x xs. x~=UU & s = scons[x][xs])"
- (fn prems =>
- [
- (simp_tac HOLCF_ss 1),
- (rtac (stream_rep_iso RS subst) 1),
- (res_inst_tac [("p","stream_rep[s]")] sprodE 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (asm_simp_tac HOLCF_ss 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 HOLCF_ss 1)
- ]);
-
-val streamE = prove_goal 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss 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 (HOLCF_ss addsimps stream_rews) 1),
- (dtac sym 1),
- (asm_simp_tac HOLCF_ss 1),
- (simp_tac (HOLCF_ss addsimps (prems @ stream_rews)) 1)
- ]);
-
-val stream_constrdef = [
- prover "is_scons[UU] ~= UU" "x~=UU ==> scons[x][xs]~=UU"
- ];
-
-fun prover defs thm = prove_goalw Stream.thy defs thm
- (fn prems =>
- [
- (simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps stream_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps stream_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps stream_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps stream_when) 1),
- (asm_simp_tac (HOLCF_ss 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 (HOLCF_ss 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 iterate_ss 1),
- (asm_simp_tac iterate_ss 1),
- (simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]),
-prove_goalw Stream.thy [stream_take_def] "stream_take(0)[xs]=UU"
- (fn prems =>
- [
- (asm_simp_tac iterate_ss 1)
- ])];
-
-fun prover thm = prove_goalw Stream.thy [stream_take_def] thm
- (fn prems =>
- [
- (cut_facts_tac prems 1),
- (simp_tac iterate_ss 1),
- (asm_simp_tac (HOLCF_ss 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 *)
-(* ------------------------------------------------------------------------*)
-
-val stream_copy2 = prove_goal 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 (HOLCF_ss addsimps stream_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]);
-
-val shd2 = prove_goal Stream.thy "shd[scons[x][xs]]=x"
- (fn prems =>
- [
- (res_inst_tac [("Q","x=UU")] classical2 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]);
-
-val stream_take2 = prove_goal 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 (HOLCF_ss addsimps stream_rews) 1),
- (asm_simp_tac (HOLCF_ss 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";
-
-
-(* ------------------------------------------------------------------------*)
-(* Co -induction for streams *)
-(* ------------------------------------------------------------------------*)
-
-val stream_coind_lemma = prove_goalw 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 (HOLCF_ss 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 (HOLCF_ss addsimps stream_rews) 1),
- (etac exE 1),
- (etac exE 1),
- (etac exE 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (REPEAT (etac conjE 1)),
- (rtac cfun_arg_cong 1),
- (fast_tac HOL_cs 1)
- ]);
-
-val stream_coind = prove_goal 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 *)
-(* ------------------------------------------------------------------------*)
-
-val stream_finite_ind = prove_goal 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 (HOLCF_ss addsimps stream_rews) 1),
- (resolve_tac prems 1),
- (rtac allI 1),
- (res_inst_tac [("s","s")] streamE 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (resolve_tac prems 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (resolve_tac prems 1),
- (atac 1),
- (etac spec 1)
- ]);
-
-val stream_finite_ind2 = prove_goalw 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)
- ]);
-
-val stream_finite_ind3 = prove_goal 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))
- ]);
-
-val stream_ind = prove_goal 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),
- (contX_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 *)
-(* ------------------------------------------------------------------------*)
-
-val surjectiv_scons = prove_goal Stream.thy "scons[shd[s]][stl[s]]=s"
- (fn prems =>
- [
- (res_inst_tac [("s","s")] streamE 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]);
-
-
-val stream_coind_lemma2 = prove_goalw 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 (HOLCF_ss addsimps stream_rews addsimps [surjectiv_scons]) 1),
- (eres_inst_tac [("s","shd[s1]"),("t","shd[s2]")] subst 1),
- (simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps stream_rews addsimps [surjectiv_scons]) 1),
- (res_inst_tac [("s","shd[s1]"),("t","shd[s2]")] ssubst 1),
- (etac sym 1),
- (simp_tac (HOLCF_ss addsimps stream_rews addsimps [surjectiv_scons]) 1)
- ]);
-
-
-(* ------------------------------------------------------------------------*)
-(* theorems about finite and infinite streams *)
-(* ------------------------------------------------------------------------*)
-
-(* ----------------------------------------------------------------------- *)
-(* 2 lemmas about stream_finite *)
-(* ----------------------------------------------------------------------- *)
-
-val stream_finite_UU = prove_goalw Stream.thy [stream_finite_def]
- "stream_finite(UU)"
- (fn prems =>
- [
- (rtac exI 1),
- (simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]);
-
-val inf_stream_not_UU = prove_goal 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 *)
-(* ----------------------------------------------------------------------- *)
-
-val stream_shd_lemma1 = prove_goal Stream.thy "shd[s]=UU --> s=UU"
- (fn prems =>
- [
- (res_inst_tac [("s","s")] streamE 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (hyp_subst_tac 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]);
-
-
-(* ----------------------------------------------------------------------- *)
-(* lemmas about stream_take *)
-(* ----------------------------------------------------------------------- *)
-
-val stream_take_lemma1 = prove_goal 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 (HOLCF_ss addsimps stream_rews) 1),
- (strip_tac 1),
- (rtac ((hd stream_inject) RS conjunct2) 1),
- (atac 1),
- (atac 1),
- (atac 1)
- ]);
-
-
-val stream_take_lemma2 = prove_goal Stream.thy
- "! s2. stream_take(n)[s2] = s2 --> stream_take(Suc(n))[s2]=s2"
- (fn prems =>
- [
- (nat_ind_tac "n" 1),
- (simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (strip_tac 1 ),
- (hyp_subst_tac 1),
- (simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (rtac allI 1),
- (res_inst_tac [("s","s2")] streamE 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (asm_simp_tac (HOLCF_ss 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)
- ]);
-
-val stream_take_lemma3 = prove_goal 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 (HOLCF_ss 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)
- ]);
-
-val stream_take_lemma4 = prove_goal 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 (HOLCF_ss addsimps stream_rews) 1),
- (simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]);
-
-(* ---- *)
-
-val stream_take_lemma5 = prove_goal Stream.thy
-"!s. stream_take(n)[s]=s --> iterate(n,stl,s)=UU"
- (fn prems =>
- [
- (nat_ind_tac "n" 1),
- (simp_tac iterate_ss 1),
- (simp_tac (HOLCF_ss 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 (HOLCF_ss addsimps stream_rews) 1),
- (rtac (iterate_Suc2 RS ssubst) 1),
- (asm_simp_tac (HOLCF_ss 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)
- ]);
-
-val stream_take_lemma6 = prove_goal Stream.thy
-"!s.iterate(n,stl,s)=UU --> stream_take(n)[s]=s"
- (fn prems =>
- [
- (nat_ind_tac "n" 1),
- (simp_tac iterate_ss 1),
- (strip_tac 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (rtac allI 1),
- (res_inst_tac [("s","s")] streamE 1),
- (hyp_subst_tac 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1),
- (hyp_subst_tac 1),
- (rtac (iterate_Suc2 RS ssubst) 1),
- (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1)
- ]);
-
-val stream_take_lemma7 = prove_goal 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)
- ]);
-
-
-(* ----------------------------------------------------------------------- *)
-(* lemmas stream_finite *)
-(* ----------------------------------------------------------------------- *)
-
-val stream_finite_lemma1 = prove_goalw 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)
- ]);
-
-val stream_finite_lemma2 = prove_goalw 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)
- ]);
-
-val stream_finite_lemma3 = prove_goal 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)
- ]);
-
-
-val stream_finite_lemma5 = prove_goalw 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)
- ]);
-
-val stream_finite_lemma6 = prove_goal Stream.thy
- "!s1 s2. s1 << s2 --> stream_take(n)[s2] = s2 --> stream_finite(s1)"
- (fn prems =>
- [
- (nat_ind_tac "n" 1),
- (simp_tac (HOLCF_ss 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(HOLCF_ss addsimps (stream_rews @ [stream_finite_UU])) 1),
- (hyp_subst_tac 1),
- (simp_tac (HOLCF_ss 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)
- ]);
-
-val stream_finite_lemma7 = prove_goal 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)
- ]);
-
-val stream_finite_lemma8 = prove_goalw Stream.thy [stream_finite_def]
-"stream_finite(s) = (? n. iterate(n,stl,s)=UU)"
- (fn prems =>
- [
- (simp_tac (HOL_ss addsimps [stream_take_lemma7]) 1)
- ]);
-
-
-(* ----------------------------------------------------------------------- *)
-(* admissibility of ~stream_finite *)
-(* ----------------------------------------------------------------------- *)
-
-val adm_not_stream_finite = prove_goalw 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 *)
-(* ----------------------------------------------------------------------- *)
-
-
-val adm_not_stream_finite2=prove_goal 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 contX_iterate2 1),
- (rtac allI 1),
- (rtac (stream_finite_lemma8 RS ssubst) 1),
- (fast_tac HOL_cs 1)
- ]);
-
-
--- a/src/HOLCF/stream.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,102 +0,0 @@
-(* Title: HOLCF/stream.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Theory for streams without defined empty stream
-*)
-
-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 domain equation *)
-(* X = 'a ** (X)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"
-
-(* ----------------------------------------------------------------------- *)
-(* 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
-
-
-
-
--- a/src/HOLCF/stream2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /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_prover Stream2.thy smap_def
- "smap = (LAM f s. stream_when[LAM x l.scons[f[x]][smap[f][l]]][s])";
-
-
-(* ------------------------------------------------------------------------- *)
-(* recursive properties *)
-(* ------------------------------------------------------------------------- *)
-
-
-val smap1 = prove_goal Stream2.thy "smap[f][UU] = UU"
- (fn prems =>
- [
- (rtac (smap_def2 RS ssubst) 1),
- (simp_tac (HOLCF_ss addsimps stream_when) 1)
- ]);
-
-val smap2 = prove_goal 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 (HOLCF_ss addsimps stream_rews) 1),
- (rtac refl 1)
- ]);
-
-
-val stream2_rews = [smap1, smap2];
--- a/src/HOLCF/stream2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /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"
-
-rules
-
-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]]
-
-*)
-
--- a/src/HOLCF/test Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-Test examples ran successfully
--- a/src/HOLCF/tr1.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,164 +0,0 @@
-(* Title: HOLCF/tr1.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for tr1.thy
-*)
-
-open Tr1;
-
-(* -------------------------------------------------------------------------- *)
-(* distinctness for type tr *)
-(* -------------------------------------------------------------------------- *)
-
-val dist_less_tr = [
-prove_goalw Tr1.thy [TT_def] "~TT << UU"
- (fn prems =>
- [
- (rtac classical3 1),
- (rtac defined_sinl 1),
- (rtac not_less2not_eq 1),
- (resolve_tac dist_less_one 1),
- (rtac (rep_tr_iso RS subst) 1),
- (rtac (rep_tr_iso RS subst) 1),
- (rtac cfun_arg_cong 1),
- (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) RS conjunct2 RS ssubst) 1),
- (etac (eq_UU_iff RS ssubst) 1)
- ]),
-prove_goalw Tr1.thy [FF_def] "~FF << UU"
- (fn prems =>
- [
- (rtac classical3 1),
- (rtac defined_sinr 1),
- (rtac not_less2not_eq 1),
- (resolve_tac dist_less_one 1),
- (rtac (rep_tr_iso RS subst) 1),
- (rtac (rep_tr_iso RS subst) 1),
- (rtac cfun_arg_cong 1),
- (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) RS conjunct2 RS ssubst) 1),
- (etac (eq_UU_iff RS ssubst) 1)
- ]),
-prove_goalw Tr1.thy [FF_def,TT_def] "~TT << FF"
- (fn prems =>
- [
- (rtac classical3 1),
- (rtac (less_ssum4c RS iffD1) 2),
- (rtac not_less2not_eq 1),
- (resolve_tac dist_less_one 1),
- (rtac (rep_tr_iso RS subst) 1),
- (rtac (rep_tr_iso RS subst) 1),
- (etac monofun_cfun_arg 1)
- ]),
-prove_goalw Tr1.thy [FF_def,TT_def] "~FF << TT"
- (fn prems =>
- [
- (rtac classical3 1),
- (rtac (less_ssum4d RS iffD1) 2),
- (rtac not_less2not_eq 1),
- (resolve_tac dist_less_one 1),
- (rtac (rep_tr_iso RS subst) 1),
- (rtac (rep_tr_iso RS subst) 1),
- (etac monofun_cfun_arg 1)
- ])
-];
-
-fun prover s = prove_goal Tr1.thy s
- (fn prems =>
- [
- (rtac not_less2not_eq 1),
- (resolve_tac dist_less_tr 1)
- ]);
-
-val dist_eq_tr = map prover ["~TT=UU","~FF=UU","~TT=FF"];
-val dist_eq_tr = dist_eq_tr @ (map (fn thm => (thm RS not_sym)) dist_eq_tr);
-
-(* ------------------------------------------------------------------------ *)
-(* Exhaustion and elimination for type tr *)
-(* ------------------------------------------------------------------------ *)
-
-val Exh_tr = prove_goalw Tr1.thy [FF_def,TT_def] "t=UU | t = TT | t = FF"
- (fn prems =>
- [
- (res_inst_tac [("p","rep_tr[t]")] ssumE 1),
- (rtac disjI1 1),
- (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict )
- RS conjunct2 RS subst) 1),
- (rtac (abs_tr_iso RS subst) 1),
- (etac cfun_arg_cong 1),
- (rtac disjI2 1),
- (rtac disjI1 1),
- (rtac (abs_tr_iso RS subst) 1),
- (rtac cfun_arg_cong 1),
- (etac trans 1),
- (rtac cfun_arg_cong 1),
- (rtac (Exh_one RS disjE) 1),
- (contr_tac 1),
- (atac 1),
- (rtac disjI2 1),
- (rtac disjI2 1),
- (rtac (abs_tr_iso RS subst) 1),
- (rtac cfun_arg_cong 1),
- (etac trans 1),
- (rtac cfun_arg_cong 1),
- (rtac (Exh_one RS disjE) 1),
- (contr_tac 1),
- (atac 1)
- ]);
-
-
-val trE = prove_goal Tr1.thy
- "[| p=UU ==> Q; p = TT ==>Q; p = FF ==>Q|] ==>Q"
- (fn prems =>
- [
- (rtac (Exh_tr RS disjE) 1),
- (eresolve_tac prems 1),
- (etac disjE 1),
- (eresolve_tac prems 1),
- (eresolve_tac prems 1)
- ]);
-
-
-(* ------------------------------------------------------------------------ *)
-(* type tr is flat *)
-(* ------------------------------------------------------------------------ *)
-
-val prems = goalw Tr1.thy [flat_def] "flat(TT)";
-by (rtac allI 1);
-by (rtac allI 1);
-by (res_inst_tac [("p","x")] trE 1);
-by (asm_simp_tac ccc1_ss 1);
-by (res_inst_tac [("p","y")] trE 1);
-by (asm_simp_tac (ccc1_ss addsimps dist_less_tr) 1);
-by (asm_simp_tac (ccc1_ss addsimps dist_less_tr) 1);
-by (asm_simp_tac (ccc1_ss addsimps dist_less_tr) 1);
-by (res_inst_tac [("p","y")] trE 1);
-by (asm_simp_tac (ccc1_ss addsimps dist_less_tr) 1);
-by (asm_simp_tac (ccc1_ss addsimps dist_less_tr) 1);
-by (asm_simp_tac (ccc1_ss addsimps dist_less_tr) 1);
-val flat_tr = result();
-
-
-(* ------------------------------------------------------------------------ *)
-(* properties of tr_when *)
-(* ------------------------------------------------------------------------ *)
-
-fun prover s = prove_goalw Tr1.thy [tr_when_def,TT_def,FF_def] s
- (fn prems =>
- [
- (simp_tac Cfun_ss 1),
- (simp_tac (Ssum_ss addsimps [(rep_tr_iso ),
- (abs_tr_iso RS allI) RS ((rep_tr_iso RS allI)
- RS iso_strict) RS conjunct1]@dist_eq_one)1)
- ]);
-
-val tr_when = map prover [
- "tr_when[x][y][UU] = UU",
- "tr_when[x][y][TT] = x",
- "tr_when[x][y][FF] = y"
- ];
-
-
-
-
-
--- a/src/HOLCF/tr1.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,57 +0,0 @@
-(* Title: HOLCF/tr1.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Introduve the domain of truth values tr = {UU,TT,FF}
-
-This type is introduced using a domain isomorphism.
-
-The type axiom
-
- arities tr :: pcpo
-
-and the continuity of the Isomorphisms are taken for granted. Since the
-type is not recursive, it could be easily introduced in a purely conservative
-style as it was used for the types sprod, ssum, lift. The definition of the
-ordering is canonical using abstraction and representation, but this would take
-again a lot of internal constants. It can be easily seen that the axioms below
-are consistent.
-
-Partial Ordering is implicit in isomorphims abs_tr,rep_tr and their continuity
-
-*)
-
-Tr1 = One +
-
-types tr 0
-arities tr :: pcpo
-
-consts
- abs_tr :: "one ++ one -> tr"
- rep_tr :: "tr -> one ++ one"
- TT :: "tr"
- FF :: "tr"
- tr_when :: "'c -> 'c -> tr -> 'c"
-
-rules
-
- abs_tr_iso "abs_tr[rep_tr[u]] = u"
- rep_tr_iso "rep_tr[abs_tr[x]] = x"
-
- TT_def "TT == abs_tr[sinl[one]]"
- FF_def "FF == abs_tr[sinr[one]]"
-
- tr_when_def "tr_when == (LAM e1 e2 t.when[LAM x.e1][LAM y.e2][rep_tr[t]])"
-
-end
-
-
-
-
-
-
-
-
-
-
--- a/src/HOLCF/tr2.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,82 +0,0 @@
-(* Title: HOLCF/tr2.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for tr2.thy
-*)
-
-open Tr2;
-
-(* ------------------------------------------------------------------------ *)
-(* lemmas about andalso *)
-(* ------------------------------------------------------------------------ *)
-
-fun prover s = prove_goalw Tr2.thy [andalso_def] s
- (fn prems =>
- [
- (simp_tac (ccc1_ss addsimps tr_when) 1)
- ]);
-
-val andalso_thms = map prover [
- "TT andalso y = y",
- "FF andalso y = FF",
- "UU andalso y = UU"
- ];
-
-val andalso_thms = andalso_thms @
- [prove_goalw Tr2.thy [andalso_def] "x andalso TT = x"
- (fn prems =>
- [
- (res_inst_tac [("p","x")] trE 1),
- (asm_simp_tac (ccc1_ss addsimps tr_when) 1),
- (asm_simp_tac (ccc1_ss addsimps tr_when) 1),
- (asm_simp_tac (ccc1_ss addsimps tr_when) 1)
- ])];
-
-(* ------------------------------------------------------------------------ *)
-(* lemmas about orelse *)
-(* ------------------------------------------------------------------------ *)
-
-fun prover s = prove_goalw Tr2.thy [orelse_def] s
- (fn prems =>
- [
- (simp_tac (ccc1_ss addsimps tr_when) 1)
- ]);
-
-val orelse_thms = map prover [
- "TT orelse y = TT",
- "FF orelse y = y",
- "UU orelse y = UU"
- ];
-
-val orelse_thms = orelse_thms @
- [prove_goalw Tr2.thy [orelse_def] "x orelse FF = x"
- (fn prems =>
- [
- (res_inst_tac [("p","x")] trE 1),
- (asm_simp_tac (ccc1_ss addsimps tr_when) 1),
- (asm_simp_tac (ccc1_ss addsimps tr_when) 1),
- (asm_simp_tac (ccc1_ss addsimps tr_when) 1)
- ])];
-
-
-(* ------------------------------------------------------------------------ *)
-(* lemmas about If_then_else_fi *)
-(* ------------------------------------------------------------------------ *)
-
-fun prover s = prove_goalw Tr2.thy [ifte_def] s
- (fn prems =>
- [
- (simp_tac (ccc1_ss addsimps tr_when) 1)
- ]);
-
-val ifte_thms = map prover [
- "If UU then e1 else e2 fi = UU",
- "If FF then e1 else e2 fi = e2",
- "If TT then e1 else e2 fi = e1"];
-
-
-
-
-
--- a/src/HOLCF/tr2.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,56 +0,0 @@
-(* Title: HOLCF/tr2.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Introduce infix if_then_else_fi and boolean connectives andalso, orelse
-*)
-
-Tr2 = Tr1 +
-
-consts
- "@cifte" :: "tr=>'c=>'c=>'c"
- ("(3If _/ (then _/ else _) fi)" [60,60,60] 60)
- "Icifte" :: "tr->'c->'c->'c"
-
- "@andalso" :: "tr => tr => tr" ("_ andalso _" [61,60] 60)
- "cop @andalso" :: "tr -> tr -> tr" ("trand")
- "@orelse" :: "tr => tr => tr" ("_ orelse _" [61,60] 60)
- "cop @orelse" :: "tr -> tr -> tr" ("tror")
-
-
-rules
-
- ifte_def "Icifte == (LAM t e1 e2.tr_when[e1][e2][t])"
- andalso_def "trand == (LAM t1 t2.tr_when[t2][FF][t1])"
- orelse_def "tror == (LAM t1 t2.tr_when[TT][t2][t1])"
-
-
-end
-
-ML
-
-(* ----------------------------------------------------------------------*)
-(* translations for the above mixfixes *)
-(* ----------------------------------------------------------------------*)
-
-fun ciftetr ts =
- let val Cfapp = Const("fapp",dummyT) in
- Cfapp $
- (Cfapp $
- (Cfapp$Const("Icifte",dummyT)$(nth_elem (0,ts)))$
- (nth_elem (1,ts)))$
- (nth_elem (2,ts))
- end;
-
-
-val parse_translation = [("@andalso",mk_cinfixtr "@andalso"),
- ("@orelse",mk_cinfixtr "@orelse"),
- ("@cifte",ciftetr)];
-
-val print_translation = [];
-
-
-
-
-
--- a/src/HOLCF/void.ML Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,69 +0,0 @@
-(* Title: HOLCF/void.ML
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Lemmas for void.thy.
-
-These lemmas are prototype lemmas for class porder
-see class theory porder.thy
-*)
-
-open Void;
-
-(* ------------------------------------------------------------------------ *)
-(* A non-emptyness result for Void *)
-(* ------------------------------------------------------------------------ *)
-
-val VoidI = prove_goalw Void.thy [UU_void_Rep_def,Void_def]
- " UU_void_Rep : Void"
-(fn prems =>
- [
- (rtac (mem_Collect_eq RS ssubst) 1),
- (rtac refl 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* less_void is a partial ordering on void *)
-(* ------------------------------------------------------------------------ *)
-
-val refl_less_void = prove_goalw Void.thy [ less_void_def ] "less_void(x,x)"
-(fn prems =>
- [
- (fast_tac HOL_cs 1)
- ]);
-
-val antisym_less_void = prove_goalw Void.thy [ less_void_def ]
- "[|less_void(x,y); less_void(y,x)|] ==> x = y"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (rtac (Rep_Void_inverse RS subst) 1),
- (etac subst 1),
- (rtac (Rep_Void_inverse RS sym) 1)
- ]);
-
-val trans_less_void = prove_goalw Void.thy [ less_void_def ]
- "[|less_void(x,y); less_void(y,z)|] ==> less_void(x,z)"
-(fn prems =>
- [
- (cut_facts_tac prems 1),
- (fast_tac HOL_cs 1)
- ]);
-
-(* ------------------------------------------------------------------------ *)
-(* a technical lemma about void: *)
-(* every element in void is represented by UU_void_Rep *)
-(* ------------------------------------------------------------------------ *)
-
-val unique_void = prove_goal Void.thy "Rep_Void(x) = UU_void_Rep"
-(fn prems =>
- [
- (rtac (mem_Collect_eq RS subst) 1),
- (fold_goals_tac [Void_def]),
- (rtac Rep_Void 1)
- ]);
-
-
-
-
--- a/src/HOLCF/void.thy Sat Apr 05 17:03:38 2003 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,48 +0,0 @@
-(* Title: HOLCF/void.thy
- ID: $Id$
- Author: Franz Regensburger
- Copyright 1993 Technische Universitaet Muenchen
-
-Definition of type void with partial order. Void is the prototype for
-all types in class 'po'
-
-Type void is defined as a set Void over type bool.
-*)
-
-Void = Holcfb +
-
-types void 0
-
-arities void :: term
-
-consts
- Void :: "bool set"
- UU_void_Rep :: "bool"
- Rep_Void :: "void => bool"
- Abs_Void :: "bool => void"
- UU_void :: "void"
- less_void :: "[void,void] => bool"
-
-rules
-
- (* The unique element in Void is False:bool *)
-
- UU_void_Rep_def "UU_void_Rep == False"
- Void_def "Void == {x. x = UU_void_Rep}"
-
- (*faking a type definition... *)
- (* void is isomorphic to Void *)
-
- Rep_Void "Rep_Void(x):Void"
- Rep_Void_inverse "Abs_Void(Rep_Void(x)) = x"
- Abs_Void_inverse "y:Void ==> Rep_Void(Abs_Void(y)) = y"
-
- (*defining the abstract constants*)
-
- UU_void_def "UU_void == Abs_Void(UU_void_Rep)"
- less_void_def "less_void(x,y) == (Rep_Void(x) = Rep_Void(y))"
-end
-
-
-
-