src/HOLCF/domain/theorems.ML
changeset 1461 6bcb44e4d6e5
parent 1274 ea0668a1c0ba
child 1512 ce37c64244c0
equal deleted inserted replaced
1460:5a6f2aabd538 1461:6bcb44e4d6e5
    32 (*
    32 (*
    33 infix 0 y;
    33 infix 0 y;
    34 val b=0;
    34 val b=0;
    35 fun _ y t = by t;
    35 fun _ y t = by t;
    36 fun  g  defs t = let val sg = sign_of thy;
    36 fun  g  defs t = let val sg = sign_of thy;
    37 		     val ct = Thm.cterm_of sg (inferT sg t);
    37                      val ct = Thm.cterm_of sg (inferT sg t);
    38 		 in goalw_cterm defs ct end;
    38                  in goalw_cterm defs ct end;
    39 *)
    39 *)
    40 
    40 
    41 fun pg'' thy defs t = let val sg = sign_of thy;
    41 fun pg'' thy defs t = let val sg = sign_of thy;
    42 		          val ct = Thm.cterm_of sg (inferT sg t);
    42                           val ct = Thm.cterm_of sg (inferT sg t);
    43 		      in prove_goalw_cterm defs ct end;
    43                       in prove_goalw_cterm defs ct end;
    44 fun pg'  thy defs t tacsf=pg'' thy defs t (fn []   => tacsf 
    44 fun pg'  thy defs t tacsf=pg'' thy defs t (fn []   => tacsf 
    45 					    | prems=> (cut_facts_tac prems 1)::tacsf);
    45                                             | prems=> (cut_facts_tac prems 1)::tacsf);
    46 
    46 
    47 fun REPEAT_DETERM_UNTIL p tac = 
    47 fun REPEAT_DETERM_UNTIL p tac = 
    48 let fun drep st = if p st then Sequence.single st
    48 let fun drep st = if p st then Sequence.single st
    49 			  else (case Sequence.pull(tapply(tac,st)) of
    49                           else (case Sequence.pull(tapply(tac,st)) of
    50 		                  None        => Sequence.null
    50                                   None        => Sequence.null
    51 				| Some(st',_) => drep st')
    51                                 | Some(st',_) => drep st')
    52 in Tactic drep end;
    52 in Tactic drep end;
    53 val UNTIL_SOLVED = REPEAT_DETERM_UNTIL (has_fewer_prems 1);
    53 val UNTIL_SOLVED = REPEAT_DETERM_UNTIL (has_fewer_prems 1);
    54 
    54 
    55 local val trueI2 = prove_goal HOL.thy "f~=x ==> True" (fn prems => [rtac TrueI 1]) in
    55 local val trueI2 = prove_goal HOL.thy "f~=x ==> True" (fn prems => [rtac TrueI 1]) in
    56 val kill_neq_tac = dtac trueI2 end;
    56 val kill_neq_tac = dtac trueI2 end;
    57 fun case_UU_tac rews i v =	res_inst_tac [("Q",v^"=UU")] classical2 i THEN
    57 fun case_UU_tac rews i v =      res_inst_tac [("Q",v^"=UU")] classical2 i THEN
    58 				asm_simp_tac (HOLCF_ss addsimps rews) i;
    58                                 asm_simp_tac (HOLCF_ss addsimps rews) i;
    59 
    59 
    60 val chain_tac = REPEAT_DETERM o resolve_tac 
    60 val chain_tac = REPEAT_DETERM o resolve_tac 
    61 		[is_chain_iterate, ch2ch_fappR, ch2ch_fappL];
    61                 [is_chain_iterate, ch2ch_fappR, ch2ch_fappL];
    62 
    62 
    63 (* ----- general proofs ----------------------------------------------------------- *)
    63 (* ----- general proofs ----------------------------------------------------------- *)
    64 
    64 
    65 val swap3 = prove_goal HOL.thy "[| Q ==> P; ~P |] ==> ~Q" (fn prems => [
    65 val swap3 = prove_goal HOL.thy "[| Q ==> P; ~P |] ==> ~Q" (fn prems => [
    66                                 cut_facts_tac prems 1,
    66                                 cut_facts_tac prems 1,
    67                                 etac swap 1,
    67                                 etac swap 1,
    68                                 dtac notnotD 1,
    68                                 dtac notnotD 1,
    69 				etac (hd prems) 1]);
    69                                 etac (hd prems) 1]);
    70 
    70 
    71 val dist_eqI = prove_goal Porder0.thy "~ x << y ==> x ~= y" (fn prems => [
    71 val dist_eqI = prove_goal Porder0.thy "~ x << y ==> x ~= y" (fn prems => [
    72 				cut_facts_tac prems 1,
    72                                 cut_facts_tac prems 1,
    73 				etac swap 1,
    73                                 etac swap 1,
    74 				dtac notnotD 1,
    74                                 dtac notnotD 1,
    75 				asm_simp_tac HOLCF_ss 1]);
    75                                 asm_simp_tac HOLCF_ss 1]);
    76 val cfst_strict  = prove_goal Cprod3.thy "cfst`UU = UU" (fn _ => [
    76 val cfst_strict  = prove_goal Cprod3.thy "cfst`UU = UU" (fn _ => [
    77 				(simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]);
    77                                 (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]);
    78 val csnd_strict  = prove_goal Cprod3.thy "csnd`UU = UU" (fn _ => [
    78 val csnd_strict  = prove_goal Cprod3.thy "csnd`UU = UU" (fn _ => [
    79 			(simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]);
    79                         (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]);
    80 
    80 
    81 in
    81 in
    82 
    82 
    83 
    83 
    84 fun theorems thy (((dname,_),cons) : eq, eqs :eq list) =
    84 fun theorems thy (((dname,_),cons) : eq, eqs :eq list) =
    94 val ax_rep_iso    = ga (dname^"_rep_iso"   );
    94 val ax_rep_iso    = ga (dname^"_rep_iso"   );
    95 val ax_when_def   = ga (dname^"_when_def"  );
    95 val ax_when_def   = ga (dname^"_when_def"  );
    96 val axs_con_def   = map (fn (con,_) => ga (extern_name con ^"_def")) cons;
    96 val axs_con_def   = map (fn (con,_) => ga (extern_name con ^"_def")) cons;
    97 val axs_dis_def   = map (fn (con,_) => ga (   dis_name con ^"_def")) cons;
    97 val axs_dis_def   = map (fn (con,_) => ga (   dis_name con ^"_def")) cons;
    98 val axs_sel_def   = flat(map (fn (_,args) => 
    98 val axs_sel_def   = flat(map (fn (_,args) => 
    99 		    map (fn     arg => ga (sel_of arg      ^"_def")) args) cons);
    99                     map (fn     arg => ga (sel_of arg      ^"_def")) args) cons);
   100 val ax_copy_def   = ga (dname^"_copy_def"  );
   100 val ax_copy_def   = ga (dname^"_copy_def"  );
   101 end; (* local *)
   101 end; (* local *)
   102 
   102 
   103 (* ----- theorems concerning the isomorphism -------------------------------------- *)
   103 (* ----- theorems concerning the isomorphism -------------------------------------- *)
   104 
   104 
   106 val dc_rep  = %%(dname^"_rep");
   106 val dc_rep  = %%(dname^"_rep");
   107 val dc_copy = %%(dname^"_copy");
   107 val dc_copy = %%(dname^"_copy");
   108 val x_name = "x";
   108 val x_name = "x";
   109 
   109 
   110 val (rep_strict, abs_strict) = let 
   110 val (rep_strict, abs_strict) = let 
   111 	       val r = ax_rep_iso RS (ax_abs_iso RS (allI  RSN(2,allI RS iso_strict)))
   111                val r = ax_rep_iso RS (ax_abs_iso RS (allI  RSN(2,allI RS iso_strict)))
   112 	       in (r RS conjunct1, r RS conjunct2) end;
   112                in (r RS conjunct1, r RS conjunct2) end;
   113 val abs_defin' = pg [] ((dc_abs`%x_name === UU) ==> (%x_name === UU)) [
   113 val abs_defin' = pg [] ((dc_abs`%x_name === UU) ==> (%x_name === UU)) [
   114 				res_inst_tac [("t",x_name)] (ax_abs_iso RS subst) 1,
   114                                 res_inst_tac [("t",x_name)] (ax_abs_iso RS subst) 1,
   115 				etac ssubst 1,
   115                                 etac ssubst 1,
   116 				rtac rep_strict 1];
   116                                 rtac rep_strict 1];
   117 val rep_defin' = pg [] ((dc_rep`%x_name === UU) ==> (%x_name === UU)) [
   117 val rep_defin' = pg [] ((dc_rep`%x_name === UU) ==> (%x_name === UU)) [
   118 				res_inst_tac [("t",x_name)] (ax_rep_iso RS subst) 1,
   118                                 res_inst_tac [("t",x_name)] (ax_rep_iso RS subst) 1,
   119 				etac ssubst 1,
   119                                 etac ssubst 1,
   120 				rtac abs_strict 1];
   120                                 rtac abs_strict 1];
   121 val iso_rews = [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
   121 val iso_rews = [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
   122 
   122 
   123 local 
   123 local 
   124 val iso_swap = pg [] (dc_rep`%"x" === %"y" ==> %"x" === dc_abs`%"y") [
   124 val iso_swap = pg [] (dc_rep`%"x" === %"y" ==> %"x" === dc_abs`%"y") [
   125 				dres_inst_tac [("f",dname^"_abs")] cfun_arg_cong 1,
   125                                 dres_inst_tac [("f",dname^"_abs")] cfun_arg_cong 1,
   126 				etac (ax_rep_iso RS subst) 1];
   126                                 etac (ax_rep_iso RS subst) 1];
   127 fun exh foldr1 cn quant foldr2 var = let
   127 fun exh foldr1 cn quant foldr2 var = let
   128   fun one_con (con,args) = let val vns = map vname args in
   128   fun one_con (con,args) = let val vns = map vname args in
   129     foldr quant (vns, foldr2 ((%x_name === con_app2 con (var vns) vns)::
   129     foldr quant (vns, foldr2 ((%x_name === con_app2 con (var vns) vns)::
   130 			      map (defined o (var vns)) (nonlazy args))) end
   130                               map (defined o (var vns)) (nonlazy args))) end
   131   in foldr1 ((cn(%x_name===UU))::map one_con cons) end;
   131   in foldr1 ((cn(%x_name===UU))::map one_con cons) end;
   132 in
   132 in
   133 val cases = let 
   133 val cases = let 
   134 	    fun common_tac thm = rtac thm 1 THEN contr_tac 1;
   134             fun common_tac thm = rtac thm 1 THEN contr_tac 1;
   135 	    fun unit_tac true = common_tac liftE1
   135             fun unit_tac true = common_tac liftE1
   136 	    |   unit_tac _    = all_tac;
   136             |   unit_tac _    = all_tac;
   137 	    fun prod_tac []          = common_tac oneE
   137             fun prod_tac []          = common_tac oneE
   138 	    |   prod_tac [arg]       = unit_tac (is_lazy arg)
   138             |   prod_tac [arg]       = unit_tac (is_lazy arg)
   139 	    |   prod_tac (arg::args) = 
   139             |   prod_tac (arg::args) = 
   140 				common_tac sprodE THEN
   140                                 common_tac sprodE THEN
   141 				kill_neq_tac 1 THEN
   141                                 kill_neq_tac 1 THEN
   142 				unit_tac (is_lazy arg) THEN
   142                                 unit_tac (is_lazy arg) THEN
   143 				prod_tac args;
   143                                 prod_tac args;
   144 	    fun sum_one_tac p = SELECT_GOAL(EVERY[
   144             fun sum_one_tac p = SELECT_GOAL(EVERY[
   145 				rtac p 1,
   145                                 rtac p 1,
   146 				rewrite_goals_tac axs_con_def,
   146                                 rewrite_goals_tac axs_con_def,
   147 				dtac iso_swap 1,
   147                                 dtac iso_swap 1,
   148 				simp_tac HOLCF_ss 1,
   148                                 simp_tac HOLCF_ss 1,
   149 				UNTIL_SOLVED(fast_tac HOL_cs 1)]) 1;
   149                                 UNTIL_SOLVED(fast_tac HOL_cs 1)]) 1;
   150 	    fun sum_tac [(_,args)]       [p]        = 
   150             fun sum_tac [(_,args)]       [p]        = 
   151 				prod_tac args THEN sum_one_tac p
   151                                 prod_tac args THEN sum_one_tac p
   152 	    |   sum_tac ((_,args)::cons') (p::prems) = DETERM(
   152             |   sum_tac ((_,args)::cons') (p::prems) = DETERM(
   153 				common_tac ssumE THEN
   153                                 common_tac ssumE THEN
   154 				kill_neq_tac 1 THEN kill_neq_tac 2 THEN
   154                                 kill_neq_tac 1 THEN kill_neq_tac 2 THEN
   155 				prod_tac args THEN sum_one_tac p) THEN
   155                                 prod_tac args THEN sum_one_tac p) THEN
   156 				sum_tac cons' prems
   156                                 sum_tac cons' prems
   157 	    |   sum_tac _ _ = Imposs "theorems:sum_tac";
   157             |   sum_tac _ _ = Imposs "theorems:sum_tac";
   158 	  in pg'' thy [] (exh (fn l => foldr (op ===>) (l,mk_trp(%"P")))
   158           in pg'' thy [] (exh (fn l => foldr (op ===>) (l,mk_trp(%"P")))
   159 			      (fn T => T ==> %"P") mk_All
   159                               (fn T => T ==> %"P") mk_All
   160 			      (fn l => foldr (op ===>) (map mk_trp l,mk_trp(%"P")))
   160                               (fn l => foldr (op ===>) (map mk_trp l,mk_trp(%"P")))
   161 			      bound_arg)
   161                               bound_arg)
   162 			     (fn prems => [
   162                              (fn prems => [
   163 				cut_facts_tac [excluded_middle] 1,
   163                                 cut_facts_tac [excluded_middle] 1,
   164 				etac disjE 1,
   164                                 etac disjE 1,
   165 				rtac (hd prems) 2,
   165                                 rtac (hd prems) 2,
   166 				etac rep_defin' 2,
   166                                 etac rep_defin' 2,
   167 				if is_one_con_one_arg (not o is_lazy) cons
   167                                 if is_one_con_one_arg (not o is_lazy) cons
   168 				then rtac (hd (tl prems)) 1 THEN atac 2 THEN
   168                                 then rtac (hd (tl prems)) 1 THEN atac 2 THEN
   169 				     rewrite_goals_tac axs_con_def THEN
   169                                      rewrite_goals_tac axs_con_def THEN
   170 				     simp_tac (HOLCF_ss addsimps [ax_rep_iso]) 1
   170                                      simp_tac (HOLCF_ss addsimps [ax_rep_iso]) 1
   171 				else sum_tac cons (tl prems)])end;
   171                                 else sum_tac cons (tl prems)])end;
   172 val exhaust = pg [] (mk_trp(exh (foldr' mk_disj) Id mk_ex (foldr' mk_conj) (K %))) [
   172 val exhaust = pg [] (mk_trp(exh (foldr' mk_disj) Id mk_ex (foldr' mk_conj) (K %))) [
   173 				rtac cases 1,
   173                                 rtac cases 1,
   174 				UNTIL_SOLVED(fast_tac HOL_cs 1)];
   174                                 UNTIL_SOLVED(fast_tac HOL_cs 1)];
   175 end;
   175 end;
   176 
   176 
   177 local 
   177 local 
   178 val when_app = foldl (op `) (%%(dname^"_when"), map % (when_funs cons));
   178 val when_app = foldl (op `) (%%(dname^"_when"), map % (when_funs cons));
   179 val when_appl = pg [ax_when_def] (mk_trp(when_app`%x_name===when_body cons 
   179 val when_appl = pg [ax_when_def] (mk_trp(when_app`%x_name===when_body cons 
   180 		(fn (_,n) => %(nth_elem(n-1,when_funs cons)))`(dc_rep`%x_name))) [
   180                 (fn (_,n) => %(nth_elem(n-1,when_funs cons)))`(dc_rep`%x_name))) [
   181 				simp_tac HOLCF_ss 1];
   181                                 simp_tac HOLCF_ss 1];
   182 in
   182 in
   183 val when_strict = pg [] ((if is_one_con_one_arg (K true) cons 
   183 val when_strict = pg [] ((if is_one_con_one_arg (K true) cons 
   184 	then fn t => mk_trp(strict(%"f")) ===> t else Id)(mk_trp(strict when_app))) [
   184         then fn t => mk_trp(strict(%"f")) ===> t else Id)(mk_trp(strict when_app))) [
   185 				simp_tac(HOLCF_ss addsimps [when_appl,rep_strict]) 1];
   185                                 simp_tac(HOLCF_ss addsimps [when_appl,rep_strict]) 1];
   186 val when_apps = let fun one_when n (con,args) = pg axs_con_def
   186 val when_apps = let fun one_when n (con,args) = pg axs_con_def
   187 		(lift_defined % (nonlazy args, mk_trp(when_app`(con_app con args) ===
   187                 (lift_defined % (nonlazy args, mk_trp(when_app`(con_app con args) ===
   188 		 mk_cfapp(%(nth_elem(n,when_funs cons)),map %# args))))[
   188                  mk_cfapp(%(nth_elem(n,when_funs cons)),map %# args))))[
   189 			asm_simp_tac (HOLCF_ss addsimps [when_appl,ax_abs_iso]) 1];
   189                         asm_simp_tac (HOLCF_ss addsimps [when_appl,ax_abs_iso]) 1];
   190 		in mapn one_when 0 cons end;
   190                 in mapn one_when 0 cons end;
   191 end;
   191 end;
   192 val when_rews = when_strict::when_apps;
   192 val when_rews = when_strict::when_apps;
   193 
   193 
   194 (* ----- theorems concerning the constructors, discriminators and selectors ------- *)
   194 (* ----- theorems concerning the constructors, discriminators and selectors ------- *)
   195 
   195 
   196 val dis_stricts = map (fn (con,_) => pg axs_dis_def (mk_trp(
   196 val dis_stricts = map (fn (con,_) => pg axs_dis_def (mk_trp(
   197 			(if is_one_con_one_arg (K true) cons then mk_not else Id)
   197                         (if is_one_con_one_arg (K true) cons then mk_not else Id)
   198 		         (strict(%%(dis_name con))))) [
   198                          (strict(%%(dis_name con))))) [
   199 		simp_tac (HOLCF_ss addsimps (if is_one_con_one_arg (K true) cons 
   199                 simp_tac (HOLCF_ss addsimps (if is_one_con_one_arg (K true) cons 
   200 					then [ax_when_def] else when_rews)) 1]) cons;
   200                                         then [ax_when_def] else when_rews)) 1]) cons;
   201 val dis_apps = let fun one_dis c (con,args)= pg (axs_dis_def)
   201 val dis_apps = let fun one_dis c (con,args)= pg (axs_dis_def)
   202 		   (lift_defined % (nonlazy args, (*(if is_one_con_one_arg is_lazy cons
   202                    (lift_defined % (nonlazy args, (*(if is_one_con_one_arg is_lazy cons
   203 			then curry (lift_defined %#) args else Id)
   203                         then curry (lift_defined %#) args else Id)
   204 #################*)
   204 #################*)
   205 			(mk_trp((%%(dis_name c))`(con_app con args) ===
   205                         (mk_trp((%%(dis_name c))`(con_app con args) ===
   206 			      %%(if con=c then "TT" else "FF"))))) [
   206                               %%(if con=c then "TT" else "FF"))))) [
   207 				asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
   207                                 asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
   208 	in flat(map (fn (c,_) => map (one_dis c) cons) cons) end;
   208         in flat(map (fn (c,_) => map (one_dis c) cons) cons) end;
   209 val dis_defins = map (fn (con,args) => pg [] (defined(%x_name)==> 
   209 val dis_defins = map (fn (con,args) => pg [] (defined(%x_name)==> 
   210 		      defined(%%(dis_name con)`%x_name)) [
   210                       defined(%%(dis_name con)`%x_name)) [
   211 				rtac cases 1,
   211                                 rtac cases 1,
   212 				contr_tac 1,
   212                                 contr_tac 1,
   213 				UNTIL_SOLVED (CHANGED(asm_simp_tac 
   213                                 UNTIL_SOLVED (CHANGED(asm_simp_tac 
   214 				              (HOLCF_ss addsimps dis_apps) 1))]) cons;
   214                                               (HOLCF_ss addsimps dis_apps) 1))]) cons;
   215 val dis_rews = dis_stricts @ dis_defins @ dis_apps;
   215 val dis_rews = dis_stricts @ dis_defins @ dis_apps;
   216 
   216 
   217 val con_stricts = flat(map (fn (con,args) => map (fn vn =>
   217 val con_stricts = flat(map (fn (con,args) => map (fn vn =>
   218 			pg (axs_con_def) 
   218                         pg (axs_con_def) 
   219 			   (mk_trp(con_app2 con (fn arg => if vname arg = vn 
   219                            (mk_trp(con_app2 con (fn arg => if vname arg = vn 
   220 					then UU else %# arg) args === UU))[
   220                                         then UU else %# arg) args === UU))[
   221 				asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1]
   221                                 asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1]
   222 			) (nonlazy args)) cons);
   222                         ) (nonlazy args)) cons);
   223 val con_defins = map (fn (con,args) => pg []
   223 val con_defins = map (fn (con,args) => pg []
   224 			(lift_defined % (nonlazy args,
   224                         (lift_defined % (nonlazy args,
   225 				mk_trp(defined(con_app con args)))) ([
   225                                 mk_trp(defined(con_app con args)))) ([
   226 				rtac swap3 1] @ (if is_one_con_one_arg (K true) cons 
   226                                 rtac swap3 1] @ (if is_one_con_one_arg (K true) cons 
   227 				then [
   227                                 then [
   228 				  if is_lazy (hd args) then rtac defined_up 2
   228                                   if is_lazy (hd args) then rtac defined_up 2
   229 						       else atac 2,
   229                                                        else atac 2,
   230 				  rtac abs_defin' 1,	
   230                                   rtac abs_defin' 1,    
   231 				  asm_full_simp_tac (HOLCF_ss addsimps axs_con_def) 1]
   231                                   asm_full_simp_tac (HOLCF_ss addsimps axs_con_def) 1]
   232 				else [
   232                                 else [
   233 				  eres_inst_tac [("f",dis_name con)] cfun_arg_cong 1,
   233                                   eres_inst_tac [("f",dis_name con)] cfun_arg_cong 1,
   234 				  asm_simp_tac (HOLCF_ss addsimps dis_rews) 1])))cons;
   234                                   asm_simp_tac (HOLCF_ss addsimps dis_rews) 1])))cons;
   235 val con_rews = con_stricts @ con_defins;
   235 val con_rews = con_stricts @ con_defins;
   236 
   236 
   237 val sel_stricts = let fun one_sel sel = pg axs_sel_def (mk_trp(strict(%%sel))) [
   237 val sel_stricts = let fun one_sel sel = pg axs_sel_def (mk_trp(strict(%%sel))) [
   238 				simp_tac (HOLCF_ss addsimps when_rews) 1];
   238                                 simp_tac (HOLCF_ss addsimps when_rews) 1];
   239 in flat(map (fn (_,args) => map (fn arg => one_sel (sel_of arg)) args) cons) end;
   239 in flat(map (fn (_,args) => map (fn arg => one_sel (sel_of arg)) args) cons) end;
   240 val sel_apps = let fun one_sel c n sel = map (fn (con,args) => 
   240 val sel_apps = let fun one_sel c n sel = map (fn (con,args) => 
   241 		let val nlas = nonlazy args;
   241                 let val nlas = nonlazy args;
   242 		    val vns  = map vname args;
   242                     val vns  = map vname args;
   243 		in pg axs_sel_def (lift_defined %
   243                 in pg axs_sel_def (lift_defined %
   244 		   (filter (fn v => con=c andalso (v<>nth_elem(n,vns))) nlas,
   244                    (filter (fn v => con=c andalso (v<>nth_elem(n,vns))) nlas,
   245    mk_trp((%%sel)`(con_app con args) === (if con=c then %(nth_elem(n,vns)) else UU))))
   245    mk_trp((%%sel)`(con_app con args) === (if con=c then %(nth_elem(n,vns)) else UU))))
   246 			    ( (if con=c then [] 
   246                             ( (if con=c then [] 
   247 			       else map(case_UU_tac(when_rews@con_stricts)1) nlas)
   247                                else map(case_UU_tac(when_rews@con_stricts)1) nlas)
   248 			     @(if con=c andalso ((nth_elem(n,vns)) mem nlas)
   248                              @(if con=c andalso ((nth_elem(n,vns)) mem nlas)
   249 					 then[case_UU_tac (when_rews @ con_stricts) 1 
   249                                          then[case_UU_tac (when_rews @ con_stricts) 1 
   250 							  (nth_elem(n,vns))] else [])
   250                                                           (nth_elem(n,vns))] else [])
   251 			     @ [asm_simp_tac(HOLCF_ss addsimps when_rews)1])end) cons;
   251                              @ [asm_simp_tac(HOLCF_ss addsimps when_rews)1])end) cons;
   252 in flat(map  (fn (c,args) => 
   252 in flat(map  (fn (c,args) => 
   253 	flat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end;
   253         flat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end;
   254 val sel_defins = if length cons = 1 then map (fn arg => pg [] (defined(%x_name) ==> 
   254 val sel_defins = if length cons = 1 then map (fn arg => pg [] (defined(%x_name) ==> 
   255 			defined(%%(sel_of arg)`%x_name)) [
   255                         defined(%%(sel_of arg)`%x_name)) [
   256 				rtac cases 1,
   256                                 rtac cases 1,
   257 				contr_tac 1,
   257                                 contr_tac 1,
   258 				UNTIL_SOLVED (CHANGED(asm_simp_tac 
   258                                 UNTIL_SOLVED (CHANGED(asm_simp_tac 
   259 				              (HOLCF_ss addsimps sel_apps) 1))]) 
   259                                               (HOLCF_ss addsimps sel_apps) 1))]) 
   260 		 (filter_out is_lazy (snd(hd cons))) else [];
   260                  (filter_out is_lazy (snd(hd cons))) else [];
   261 val sel_rews = sel_stricts @ sel_defins @ sel_apps;
   261 val sel_rews = sel_stricts @ sel_defins @ sel_apps;
   262 
   262 
   263 val distincts_le = let
   263 val distincts_le = let
   264     fun dist (con1, args1) (con2, args2) = pg []
   264     fun dist (con1, args1) (con2, args2) = pg []
   265 	      (lift_defined % ((nonlazy args1),
   265               (lift_defined % ((nonlazy args1),
   266 			     (mk_trp (con_app con1 args1 ~<< con_app con2 args2))))([
   266                              (mk_trp (con_app con1 args1 ~<< con_app con2 args2))))([
   267 			rtac swap3 1,
   267                         rtac swap3 1,
   268 			eres_inst_tac [("fo5",dis_name con1)] monofun_cfun_arg 1]
   268                         eres_inst_tac [("fo5",dis_name con1)] monofun_cfun_arg 1]
   269 		      @ map (case_UU_tac (con_stricts @ dis_rews) 1) (nonlazy args2)
   269                       @ map (case_UU_tac (con_stricts @ dis_rews) 1) (nonlazy args2)
   270 		      @[asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]);
   270                       @[asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]);
   271     fun distinct (con1,args1) (con2,args2) =
   271     fun distinct (con1,args1) (con2,args2) =
   272 	let val arg1 = (con1, args1);
   272         let val arg1 = (con1, args1);
   273 	    val arg2 = (con2, (map (fn (arg,vn) => upd_vname (K vn) arg)
   273             val arg2 = (con2, (map (fn (arg,vn) => upd_vname (K vn) arg)
   274 			      (args2~~variantlist(map vname args2,map vname args1))));
   274                               (args2~~variantlist(map vname args2,map vname args1))));
   275 	in [dist arg1 arg2, dist arg2 arg1] end;
   275         in [dist arg1 arg2, dist arg2 arg1] end;
   276     fun distincts []      = []
   276     fun distincts []      = []
   277     |   distincts (c::cs) = (map (distinct c) cs) :: distincts cs;
   277     |   distincts (c::cs) = (map (distinct c) cs) :: distincts cs;
   278 in distincts cons end;
   278 in distincts cons end;
   279 val dists_le = flat (flat distincts_le);
   279 val dists_le = flat (flat distincts_le);
   280 val dists_eq = let
   280 val dists_eq = let
   281     fun distinct (_,args1) ((_,args2),leqs) = let
   281     fun distinct (_,args1) ((_,args2),leqs) = let
   282 	val (le1,le2) = (hd leqs, hd(tl leqs));
   282         val (le1,le2) = (hd leqs, hd(tl leqs));
   283 	val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) in
   283         val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) in
   284 	if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else
   284         if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else
   285 	if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else
   285         if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else
   286 					[eq1, eq2] end;
   286                                         [eq1, eq2] end;
   287     fun distincts []      = []
   287     fun distincts []      = []
   288     |   distincts ((c,leqs)::cs) = flat(map (distinct c) ((map fst cs)~~leqs)) @
   288     |   distincts ((c,leqs)::cs) = flat(map (distinct c) ((map fst cs)~~leqs)) @
   289 				   distincts cs;
   289                                    distincts cs;
   290     in distincts (cons~~distincts_le) end;
   290     in distincts (cons~~distincts_le) end;
   291 
   291 
   292 local 
   292 local 
   293   fun pgterm rel con args = let
   293   fun pgterm rel con args = let
   294 		fun append s = upd_vname(fn v => v^s);
   294                 fun append s = upd_vname(fn v => v^s);
   295 		val (largs,rargs) = (args, map (append "'") args);
   295                 val (largs,rargs) = (args, map (append "'") args);
   296 		in pg [] (mk_trp (rel(con_app con largs,con_app con rargs)) ===>
   296                 in pg [] (mk_trp (rel(con_app con largs,con_app con rargs)) ===>
   297 		      lift_defined % ((nonlazy largs),lift_defined % ((nonlazy rargs),
   297                       lift_defined % ((nonlazy largs),lift_defined % ((nonlazy rargs),
   298 			    mk_trp (foldr' mk_conj 
   298                             mk_trp (foldr' mk_conj 
   299 				(map rel (map %# largs ~~ map %# rargs)))))) end;
   299                                 (map rel (map %# largs ~~ map %# rargs)))))) end;
   300   val cons' = filter (fn (_,args) => args<>[]) cons;
   300   val cons' = filter (fn (_,args) => args<>[]) cons;
   301 in
   301 in
   302 val inverts = map (fn (con,args) => 
   302 val inverts = map (fn (con,args) => 
   303 		pgterm (op <<) con args (flat(map (fn arg => [
   303                 pgterm (op <<) con args (flat(map (fn arg => [
   304 				TRY(rtac conjI 1),
   304                                 TRY(rtac conjI 1),
   305 				dres_inst_tac [("fo5",sel_of arg)] monofun_cfun_arg 1,
   305                                 dres_inst_tac [("fo5",sel_of arg)] monofun_cfun_arg 1,
   306 				asm_full_simp_tac (HOLCF_ss addsimps sel_apps) 1]
   306                                 asm_full_simp_tac (HOLCF_ss addsimps sel_apps) 1]
   307 			     			      ) args))) cons';
   307                                                       ) args))) cons';
   308 val injects = map (fn ((con,args),inv_thm) => 
   308 val injects = map (fn ((con,args),inv_thm) => 
   309 			   pgterm (op ===) con args [
   309                            pgterm (op ===) con args [
   310 				etac (antisym_less_inverse RS conjE) 1,
   310                                 etac (antisym_less_inverse RS conjE) 1,
   311 				dtac inv_thm 1, REPEAT(atac 1),
   311                                 dtac inv_thm 1, REPEAT(atac 1),
   312 				dtac inv_thm 1, REPEAT(atac 1),
   312                                 dtac inv_thm 1, REPEAT(atac 1),
   313 				TRY(safe_tac HOL_cs),
   313                                 TRY(safe_tac HOL_cs),
   314 				REPEAT(rtac antisym_less 1 ORELSE atac 1)] )
   314                                 REPEAT(rtac antisym_less 1 ORELSE atac 1)] )
   315 		  (cons'~~inverts);
   315                   (cons'~~inverts);
   316 end;
   316 end;
   317 
   317 
   318 (* ----- theorems concerning one induction step ----------------------------------- *)
   318 (* ----- theorems concerning one induction step ----------------------------------- *)
   319 
   319 
   320 val copy_strict = pg [ax_copy_def] ((if is_one_con_one_arg (K true) cons then fn t =>
   320 val copy_strict = pg [ax_copy_def] ((if is_one_con_one_arg (K true) cons then fn t =>
   321 	 mk_trp(strict(cproj (%"f") eqs (rec_of (hd(snd(hd cons)))))) ===> t
   321          mk_trp(strict(cproj (%"f") eqs (rec_of (hd(snd(hd cons)))))) ===> t
   322 	else Id) (mk_trp(strict(dc_copy`%"f")))) [
   322         else Id) (mk_trp(strict(dc_copy`%"f")))) [
   323 				asm_simp_tac(HOLCF_ss addsimps [abs_strict,rep_strict,
   323                                 asm_simp_tac(HOLCF_ss addsimps [abs_strict,rep_strict,
   324 							cfst_strict,csnd_strict]) 1];
   324                                                         cfst_strict,csnd_strict]) 1];
   325 val copy_apps = map (fn (con,args) => pg (ax_copy_def::axs_con_def)
   325 val copy_apps = map (fn (con,args) => pg (ax_copy_def::axs_con_def)
   326 		    (lift_defined %# (filter is_nonlazy_rec args,
   326                     (lift_defined %# (filter is_nonlazy_rec args,
   327 			mk_trp(dc_copy`%"f"`(con_app con args) ===
   327                         mk_trp(dc_copy`%"f"`(con_app con args) ===
   328 			   (con_app2 con (app_rec_arg (cproj (%"f") eqs)) args))))
   328                            (con_app2 con (app_rec_arg (cproj (%"f") eqs)) args))))
   329 				 (map (case_UU_tac [ax_abs_iso] 1 o vname)
   329                                  (map (case_UU_tac [ax_abs_iso] 1 o vname)
   330 				   (filter(fn a=>not(is_rec a orelse is_lazy a))args)@
   330                                    (filter(fn a=>not(is_rec a orelse is_lazy a))args)@
   331 				 [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1])
   331                                  [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1])
   332 		)cons;
   332                 )cons;
   333 val copy_stricts = map(fn(con,args)=>pg[](mk_trp(dc_copy`UU`(con_app con args) ===UU))
   333 val copy_stricts = map(fn(con,args)=>pg[](mk_trp(dc_copy`UU`(con_app con args) ===UU))
   334 	     (let val rews = cfst_strict::csnd_strict::copy_strict::copy_apps@con_rews
   334              (let val rews = cfst_strict::csnd_strict::copy_strict::copy_apps@con_rews
   335 			 in map (case_UU_tac rews 1) (nonlazy args) @ [
   335                          in map (case_UU_tac rews 1) (nonlazy args) @ [
   336 			     asm_simp_tac (HOLCF_ss addsimps rews) 1] end))
   336                              asm_simp_tac (HOLCF_ss addsimps rews) 1] end))
   337 		   (filter (fn (_,args)=>exists is_nonlazy_rec args) cons);
   337                    (filter (fn (_,args)=>exists is_nonlazy_rec args) cons);
   338 val copy_rews = copy_strict::copy_apps @ copy_stricts;
   338 val copy_rews = copy_strict::copy_apps @ copy_stricts;
   339 
   339 
   340 in     (iso_rews, exhaust, cases, when_rews,
   340 in     (iso_rews, exhaust, cases, when_rews,
   341 	con_rews, sel_rews, dis_rews, dists_eq, dists_le, inverts, injects,
   341         con_rews, sel_rews, dis_rews, dists_eq, dists_le, inverts, injects,
   342 	copy_rews)
   342         copy_rews)
   343 end; (* let *)
   343 end; (* let *)
   344 
   344 
   345 
   345 
   346 fun comp_theorems thy (comp_dname, eqs: eq list, casess, con_rews, copy_rews) =
   346 fun comp_theorems thy (comp_dname, eqs: eq list, casess, con_rews, copy_rews) =
   347 let
   347 let
   367 fun dc_take dn = %%(dn^"_take");
   367 fun dc_take dn = %%(dn^"_take");
   368 val x_name = idx_name dnames "x"; 
   368 val x_name = idx_name dnames "x"; 
   369 val P_name = idx_name dnames "P";
   369 val P_name = idx_name dnames "P";
   370 
   370 
   371 local
   371 local
   372   val iterate_ss = simpset_of "Fix";	
   372   val iterate_ss = simpset_of "Fix";    
   373   val iterate_Cprod_strict_ss = iterate_ss addsimps [cfst_strict, csnd_strict];
   373   val iterate_Cprod_strict_ss = iterate_ss addsimps [cfst_strict, csnd_strict];
   374   val iterate_Cprod_ss = iterate_ss addsimps [cfst2,csnd2,csplit2];
   374   val iterate_Cprod_ss = iterate_ss addsimps [cfst2,csnd2,csplit2];
   375   val copy_con_rews  = copy_rews @ con_rews;
   375   val copy_con_rews  = copy_rews @ con_rews;
   376   val copy_take_defs = (if length dnames=1 then [] else [ax_copy2_def]) @axs_take_def;
   376   val copy_take_defs = (if length dnames=1 then [] else [ax_copy2_def]) @axs_take_def;
   377   val take_stricts = pg copy_take_defs (mk_trp(foldr' mk_conj (map (fn ((dn,args),_)=>
   377   val take_stricts = pg copy_take_defs (mk_trp(foldr' mk_conj (map (fn ((dn,args),_)=>
   378 		  (dc_take dn $ %"n")`UU === mk_constrain(Type(dn,args),UU)) eqs)))([
   378                   (dc_take dn $ %"n")`UU === mk_constrain(Type(dn,args),UU)) eqs)))([
   379 				nat_ind_tac "n" 1,
   379                                 nat_ind_tac "n" 1,
   380 				simp_tac iterate_ss 1,
   380                                 simp_tac iterate_ss 1,
   381 				simp_tac iterate_Cprod_strict_ss 1,
   381                                 simp_tac iterate_Cprod_strict_ss 1,
   382 				asm_simp_tac iterate_Cprod_ss 1,
   382                                 asm_simp_tac iterate_Cprod_ss 1,
   383 				TRY(safe_tac HOL_cs)] @
   383                                 TRY(safe_tac HOL_cs)] @
   384 			map(K(asm_simp_tac (HOL_ss addsimps copy_rews)1))dnames);
   384                         map(K(asm_simp_tac (HOL_ss addsimps copy_rews)1))dnames);
   385   val take_stricts' = rewrite_rule copy_take_defs take_stricts;
   385   val take_stricts' = rewrite_rule copy_take_defs take_stricts;
   386   val take_0s = mapn (fn n => fn dn => pg axs_take_def(mk_trp((dc_take dn $ %%"0")
   386   val take_0s = mapn (fn n => fn dn => pg axs_take_def(mk_trp((dc_take dn $ %%"0")
   387 								`%x_name n === UU))[
   387                                                                 `%x_name n === UU))[
   388 				simp_tac iterate_Cprod_strict_ss 1]) 1 dnames;
   388                                 simp_tac iterate_Cprod_strict_ss 1]) 1 dnames;
   389   val take_apps = pg copy_take_defs (mk_trp(foldr' mk_conj 
   389   val take_apps = pg copy_take_defs (mk_trp(foldr' mk_conj 
   390 	    (flat(map (fn ((dn,_),cons) => map (fn (con,args) => foldr mk_all 
   390             (flat(map (fn ((dn,_),cons) => map (fn (con,args) => foldr mk_all 
   391 		(map vname args,(dc_take dn $ (%%"Suc" $ %"n"))`(con_app con args) ===
   391                 (map vname args,(dc_take dn $ (%%"Suc" $ %"n"))`(con_app con args) ===
   392   		 con_app2 con (app_rec_arg (fn n=>dc_take (nth_elem(n,dnames))$ %"n"))
   392                  con_app2 con (app_rec_arg (fn n=>dc_take (nth_elem(n,dnames))$ %"n"))
   393 			      args)) cons) eqs)))) ([
   393                               args)) cons) eqs)))) ([
   394 				nat_ind_tac "n" 1,
   394                                 nat_ind_tac "n" 1,
   395 				simp_tac iterate_Cprod_strict_ss 1,
   395                                 simp_tac iterate_Cprod_strict_ss 1,
   396 				simp_tac (HOLCF_ss addsimps copy_con_rews) 1,
   396                                 simp_tac (HOLCF_ss addsimps copy_con_rews) 1,
   397 				TRY(safe_tac HOL_cs)] @
   397                                 TRY(safe_tac HOL_cs)] @
   398 			(flat(map (fn ((dn,_),cons) => map (fn (con,args) => EVERY (
   398                         (flat(map (fn ((dn,_),cons) => map (fn (con,args) => EVERY (
   399 				asm_full_simp_tac iterate_Cprod_ss 1::
   399                                 asm_full_simp_tac iterate_Cprod_ss 1::
   400 				map (case_UU_tac (take_stricts'::copy_con_rews) 1)
   400                                 map (case_UU_tac (take_stricts'::copy_con_rews) 1)
   401 				    (nonlazy args) @[
   401                                     (nonlazy args) @[
   402 				asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1])
   402                                 asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1])
   403 		 	) cons) eqs)));
   403                         ) cons) eqs)));
   404 in
   404 in
   405 val take_rews = atomize take_stricts @ take_0s @ atomize take_apps;
   405 val take_rews = atomize take_stricts @ take_0s @ atomize take_apps;
   406 end; (* local *)
   406 end; (* local *)
   407 
   407 
   408 val take_lemmas = mapn (fn n => fn(dn,ax_reach) => pg'' thy axs_take_def (mk_All("n",
   408 val take_lemmas = mapn (fn n => fn(dn,ax_reach) => pg'' thy axs_take_def (mk_All("n",
   409 		mk_trp(dc_take dn $ Bound 0 `%(x_name n) === 
   409                 mk_trp(dc_take dn $ Bound 0 `%(x_name n) === 
   410 		       dc_take dn $ Bound 0 `%(x_name n^"'")))
   410                        dc_take dn $ Bound 0 `%(x_name n^"'")))
   411 	   ===> mk_trp(%(x_name n) === %(x_name n^"'"))) (fn prems => [
   411            ===> mk_trp(%(x_name n) === %(x_name n^"'"))) (fn prems => [
   412 				res_inst_tac[("t",x_name n    )](ax_reach RS subst) 1,
   412                                 res_inst_tac[("t",x_name n    )](ax_reach RS subst) 1,
   413 				res_inst_tac[("t",x_name n^"'")](ax_reach RS subst) 1,
   413                                 res_inst_tac[("t",x_name n^"'")](ax_reach RS subst) 1,
   414 				rtac (fix_def2 RS ssubst) 1,
   414                                 rtac (fix_def2 RS ssubst) 1,
   415 				REPEAT(CHANGED(rtac (contlub_cfun_arg RS ssubst) 1
   415                                 REPEAT(CHANGED(rtac (contlub_cfun_arg RS ssubst) 1
   416 					       THEN chain_tac 1)),
   416                                                THEN chain_tac 1)),
   417 				rtac (contlub_cfun_fun RS ssubst) 1,
   417                                 rtac (contlub_cfun_fun RS ssubst) 1,
   418 				rtac (contlub_cfun_fun RS ssubst) 2,
   418                                 rtac (contlub_cfun_fun RS ssubst) 2,
   419 				rtac lub_equal 3,
   419                                 rtac lub_equal 3,
   420 				chain_tac 1,
   420                                 chain_tac 1,
   421 				rtac allI 1,
   421                                 rtac allI 1,
   422 				resolve_tac prems 1])) 1 (dnames~~axs_reach);
   422                                 resolve_tac prems 1])) 1 (dnames~~axs_reach);
   423 
   423 
   424 local
   424 local
   425   fun one_con p (con,args) = foldr mk_All (map vname args,
   425   fun one_con p (con,args) = foldr mk_All (map vname args,
   426 	lift_defined (bound_arg (map vname args)) (nonlazy args,
   426         lift_defined (bound_arg (map vname args)) (nonlazy args,
   427 	lift (fn arg => %(P_name (1+rec_of arg)) $ bound_arg args arg)
   427         lift (fn arg => %(P_name (1+rec_of arg)) $ bound_arg args arg)
   428 	     (filter is_rec args,mk_trp(%p $ con_app2 con (bound_arg args) args))));
   428              (filter is_rec args,mk_trp(%p $ con_app2 con (bound_arg args) args))));
   429   fun one_eq ((p,cons),concl) = (mk_trp(%p $ UU) ===> 
   429   fun one_eq ((p,cons),concl) = (mk_trp(%p $ UU) ===> 
   430 			   foldr (op ===>) (map (one_con p) cons,concl));
   430                            foldr (op ===>) (map (one_con p) cons,concl));
   431   fun ind_term concf = foldr one_eq (mapn (fn n => fn x => (P_name n, x)) 1 conss,
   431   fun ind_term concf = foldr one_eq (mapn (fn n => fn x => (P_name n, x)) 1 conss,
   432 	mk_trp(foldr' mk_conj (mapn (fn n => concf (P_name n,x_name n)) 1 dnames)));
   432         mk_trp(foldr' mk_conj (mapn (fn n => concf (P_name n,x_name n)) 1 dnames)));
   433   val take_ss = HOL_ss addsimps take_rews;
   433   val take_ss = HOL_ss addsimps take_rews;
   434   fun ind_tacs tacsf thms1 thms2 prems = TRY(safe_tac HOL_cs)::
   434   fun ind_tacs tacsf thms1 thms2 prems = TRY(safe_tac HOL_cs)::
   435 				flat (mapn (fn n => fn (thm1,thm2) => 
   435                                 flat (mapn (fn n => fn (thm1,thm2) => 
   436 				  tacsf (n,prems) (thm1,thm2) @ 
   436                                   tacsf (n,prems) (thm1,thm2) @ 
   437 				  flat (map (fn cons =>
   437                                   flat (map (fn cons =>
   438 				    (resolve_tac prems 1 ::
   438                                     (resolve_tac prems 1 ::
   439 				     flat (map (fn (_,args) => 
   439                                      flat (map (fn (_,args) => 
   440 				       resolve_tac prems 1::
   440                                        resolve_tac prems 1::
   441 				       map (K(atac 1)) (nonlazy args) @
   441                                        map (K(atac 1)) (nonlazy args) @
   442 				       map (K(atac 1)) (filter is_rec args))
   442                                        map (K(atac 1)) (filter is_rec args))
   443 				     cons)))
   443                                      cons)))
   444 				   conss))
   444                                    conss))
   445 				0 (thms1~~thms2));
   445                                 0 (thms1~~thms2));
   446   local 
   446   local 
   447     fun all_rec_to ns lazy_rec (n,cons) = forall (exists (fn arg => 
   447     fun all_rec_to ns lazy_rec (n,cons) = forall (exists (fn arg => 
   448 		  is_rec arg andalso not(rec_of arg mem ns) andalso
   448                   is_rec arg andalso not(rec_of arg mem ns) andalso
   449 		  ((rec_of arg =  n andalso not(lazy_rec orelse is_lazy arg)) orelse 
   449                   ((rec_of arg =  n andalso not(lazy_rec orelse is_lazy arg)) orelse 
   450 		    rec_of arg <> n andalso all_rec_to (rec_of arg::ns) 
   450                     rec_of arg <> n andalso all_rec_to (rec_of arg::ns) 
   451 		      (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss))))
   451                       (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss))))
   452 		  ) o snd) cons;
   452                   ) o snd) cons;
   453     fun warn (n,cons) = if all_rec_to [] false (n,cons) then (writeln 
   453     fun warn (n,cons) = if all_rec_to [] false (n,cons) then (writeln 
   454 			   ("WARNING: domain "^nth_elem(n,dnames)^" is empty!"); true)
   454                            ("WARNING: domain "^nth_elem(n,dnames)^" is empty!"); true)
   455 			else false;
   455                         else false;
   456     fun lazy_rec_to ns lazy_rec (n,cons) = exists (exists (fn arg => 
   456     fun lazy_rec_to ns lazy_rec (n,cons) = exists (exists (fn arg => 
   457 		  is_rec arg andalso not(rec_of arg mem ns) andalso
   457                   is_rec arg andalso not(rec_of arg mem ns) andalso
   458 		  ((rec_of arg =  n andalso (lazy_rec orelse is_lazy arg)) orelse 
   458                   ((rec_of arg =  n andalso (lazy_rec orelse is_lazy arg)) orelse 
   459 		    rec_of arg <> n andalso lazy_rec_to (rec_of arg::ns)
   459                     rec_of arg <> n andalso lazy_rec_to (rec_of arg::ns)
   460 		     (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss))))
   460                      (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss))))
   461 		 ) o snd) cons;
   461                  ) o snd) cons;
   462   in val is_emptys = map warn (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs);
   462   in val is_emptys = map warn (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs);
   463      val is_finite = forall (not o lazy_rec_to [] false) 
   463      val is_finite = forall (not o lazy_rec_to [] false) 
   464 			    (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs)
   464                             (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs)
   465   end;
   465   end;
   466 in
   466 in
   467 val finite_ind = pg'' thy [] (ind_term (fn (P,x) => fn dn => 
   467 val finite_ind = pg'' thy [] (ind_term (fn (P,x) => fn dn => 
   468 			  mk_all(x,%P $ (dc_take dn $ %"n" `Bound 0)))) (fn prems=> [
   468                           mk_all(x,%P $ (dc_take dn $ %"n" `Bound 0)))) (fn prems=> [
   469 				nat_ind_tac "n" 1,
   469                                 nat_ind_tac "n" 1,
   470 				simp_tac (take_ss addsimps prems) 1,
   470                                 simp_tac (take_ss addsimps prems) 1,
   471 				TRY(safe_tac HOL_cs)]
   471                                 TRY(safe_tac HOL_cs)]
   472 				@ flat(mapn (fn n => fn (cons,cases) => [
   472                                 @ flat(mapn (fn n => fn (cons,cases) => [
   473 				 res_inst_tac [("x",x_name n)] cases 1,
   473                                  res_inst_tac [("x",x_name n)] cases 1,
   474 				 asm_simp_tac (take_ss addsimps prems) 1]
   474                                  asm_simp_tac (take_ss addsimps prems) 1]
   475 				 @ flat(map (fn (con,args) => 
   475                                  @ flat(map (fn (con,args) => 
   476 				  asm_simp_tac take_ss 1 ::
   476                                   asm_simp_tac take_ss 1 ::
   477 				  map (fn arg =>
   477                                   map (fn arg =>
   478 				   case_UU_tac (prems@con_rews) 1 (
   478                                    case_UU_tac (prems@con_rews) 1 (
   479 				   nth_elem(rec_of arg,dnames)^"_take n1`"^vname arg))
   479                                    nth_elem(rec_of arg,dnames)^"_take n1`"^vname arg))
   480 				  (filter is_nonlazy_rec args) @ [
   480                                   (filter is_nonlazy_rec args) @ [
   481 				  resolve_tac prems 1] @
   481                                   resolve_tac prems 1] @
   482 				  map (K (atac 1))      (nonlazy args) @
   482                                   map (K (atac 1))      (nonlazy args) @
   483 				  map (K (etac spec 1)) (filter is_rec args)) 
   483                                   map (K (etac spec 1)) (filter is_rec args)) 
   484 				 cons))
   484                                  cons))
   485 				1 (conss~~casess)));
   485                                 1 (conss~~casess)));
   486 
   486 
   487 val (finites,ind) = if is_finite then
   487 val (finites,ind) = if is_finite then
   488 let 
   488 let 
   489   fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %"x" === %"x");
   489   fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %"x" === %"x");
   490   val finite_lemmas1a = map (fn dn => pg [] (mk_trp(defined (%"x")) ===> 
   490   val finite_lemmas1a = map (fn dn => pg [] (mk_trp(defined (%"x")) ===> 
   491 	mk_trp(mk_disj(mk_all("n",dc_take dn $ Bound 0 ` %"x" === UU),
   491         mk_trp(mk_disj(mk_all("n",dc_take dn $ Bound 0 ` %"x" === UU),
   492 	take_enough dn)) ===> mk_trp(take_enough dn)) [
   492         take_enough dn)) ===> mk_trp(take_enough dn)) [
   493 				etac disjE 1,
   493                                 etac disjE 1,
   494 				etac notE 1,
   494                                 etac notE 1,
   495 				resolve_tac take_lemmas 1,
   495                                 resolve_tac take_lemmas 1,
   496 				asm_simp_tac take_ss 1,
   496                                 asm_simp_tac take_ss 1,
   497 				atac 1]) dnames;
   497                                 atac 1]) dnames;
   498   val finite_lemma1b = pg [] (mk_trp (mk_all("n",foldr' mk_conj (mapn 
   498   val finite_lemma1b = pg [] (mk_trp (mk_all("n",foldr' mk_conj (mapn 
   499 	(fn n => fn ((dn,args),_) => mk_constrainall(x_name n,Type(dn,args),
   499         (fn n => fn ((dn,args),_) => mk_constrainall(x_name n,Type(dn,args),
   500 	 mk_disj(dc_take dn $ Bound 1 ` Bound 0 === UU,
   500          mk_disj(dc_take dn $ Bound 1 ` Bound 0 === UU,
   501 		 dc_take dn $ Bound 1 ` Bound 0 === Bound 0))) 1 eqs)))) ([
   501                  dc_take dn $ Bound 1 ` Bound 0 === Bound 0))) 1 eqs)))) ([
   502 				rtac allI 1,
   502                                 rtac allI 1,
   503 				nat_ind_tac "n" 1,
   503                                 nat_ind_tac "n" 1,
   504 				simp_tac take_ss 1,
   504                                 simp_tac take_ss 1,
   505 				TRY(safe_tac(empty_cs addSEs[conjE] addSIs[conjI]))] @
   505                                 TRY(safe_tac(empty_cs addSEs[conjE] addSIs[conjI]))] @
   506 				flat(mapn (fn n => fn (cons,cases) => [
   506                                 flat(mapn (fn n => fn (cons,cases) => [
   507 				  simp_tac take_ss 1,
   507                                   simp_tac take_ss 1,
   508 				  rtac allI 1,
   508                                   rtac allI 1,
   509 				  res_inst_tac [("x",x_name n)] cases 1,
   509                                   res_inst_tac [("x",x_name n)] cases 1,
   510 				  asm_simp_tac take_ss 1] @ 
   510                                   asm_simp_tac take_ss 1] @ 
   511 				  flat(map (fn (con,args) => 
   511                                   flat(map (fn (con,args) => 
   512 				    asm_simp_tac take_ss 1 ::
   512                                     asm_simp_tac take_ss 1 ::
   513 				    flat(map (fn arg => [
   513                                     flat(map (fn arg => [
   514 				      eres_inst_tac [("x",vname arg)] all_dupE 1,
   514                                       eres_inst_tac [("x",vname arg)] all_dupE 1,
   515 				      etac disjE 1,
   515                                       etac disjE 1,
   516 				      asm_simp_tac (HOL_ss addsimps con_rews) 1,
   516                                       asm_simp_tac (HOL_ss addsimps con_rews) 1,
   517 				      asm_simp_tac take_ss 1])
   517                                       asm_simp_tac take_ss 1])
   518 				    (filter is_nonlazy_rec args)))
   518                                     (filter is_nonlazy_rec args)))
   519 				  cons))
   519                                   cons))
   520 				1 (conss~~casess))) handle ERROR => raise ERROR;
   520                                 1 (conss~~casess))) handle ERROR => raise ERROR;
   521   val all_finite=map (fn(dn,l1b)=>pg axs_finite_def (mk_trp(%%(dn^"_finite") $ %"x"))[
   521   val all_finite=map (fn(dn,l1b)=>pg axs_finite_def (mk_trp(%%(dn^"_finite") $ %"x"))[
   522 				case_UU_tac take_rews 1 "x",
   522                                 case_UU_tac take_rews 1 "x",
   523 				eresolve_tac finite_lemmas1a 1,
   523                                 eresolve_tac finite_lemmas1a 1,
   524 				step_tac HOL_cs 1,
   524                                 step_tac HOL_cs 1,
   525 				step_tac HOL_cs 1,
   525                                 step_tac HOL_cs 1,
   526 				cut_facts_tac [l1b] 1,
   526                                 cut_facts_tac [l1b] 1,
   527 				fast_tac HOL_cs 1]) (dnames~~atomize finite_lemma1b);
   527                                 fast_tac HOL_cs 1]) (dnames~~atomize finite_lemma1b);
   528 in
   528 in
   529 (all_finite,
   529 (all_finite,
   530  pg'' thy [] (ind_term (fn (P,x) => fn dn => %P $ %x))
   530  pg'' thy [] (ind_term (fn (P,x) => fn dn => %P $ %x))
   531 			       (ind_tacs (fn _ => fn (all_fin,finite_ind) => [
   531                                (ind_tacs (fn _ => fn (all_fin,finite_ind) => [
   532 				rtac (rewrite_rule axs_finite_def all_fin RS exE) 1,
   532                                 rtac (rewrite_rule axs_finite_def all_fin RS exE) 1,
   533 				etac subst 1,
   533                                 etac subst 1,
   534 				rtac finite_ind 1]) all_finite (atomize finite_ind))
   534                                 rtac finite_ind 1]) all_finite (atomize finite_ind))
   535 ) end (* let *) else
   535 ) end (* let *) else
   536 (mapn (fn n => fn dn => read_instantiate_sg (sign_of thy) 
   536 (mapn (fn n => fn dn => read_instantiate_sg (sign_of thy) 
   537 	  	    [("P",dn^"_finite "^x_name n)] excluded_middle) 1 dnames,
   537                     [("P",dn^"_finite "^x_name n)] excluded_middle) 1 dnames,
   538  pg'' thy [] (foldr (op ===>) (mapn (fn n =>K(mk_trp(%%"adm" $ %(P_name n))))1
   538  pg'' thy [] (foldr (op ===>) (mapn (fn n =>K(mk_trp(%%"adm" $ %(P_name n))))1
   539 				       dnames,ind_term (fn(P,x)=>fn dn=> %P $ %x)))
   539                                        dnames,ind_term (fn(P,x)=>fn dn=> %P $ %x)))
   540 			       (ind_tacs (fn (n,prems) => fn (ax_reach,finite_ind) =>[
   540                                (ind_tacs (fn (n,prems) => fn (ax_reach,finite_ind) =>[
   541 				rtac (ax_reach RS subst) 1,
   541                                 rtac (ax_reach RS subst) 1,
   542 				res_inst_tac [("x",x_name n)] spec 1,
   542                                 res_inst_tac [("x",x_name n)] spec 1,
   543 				rtac wfix_ind 1,
   543                                 rtac wfix_ind 1,
   544 				rtac adm_impl_admw 1,
   544                                 rtac adm_impl_admw 1,
   545 				resolve_tac adm_thms 1,
   545                                 resolve_tac adm_thms 1,
   546 				rtac adm_subst 1,
   546                                 rtac adm_subst 1,
   547 				cont_tacR 1,
   547                                 cont_tacR 1,
   548 				resolve_tac prems 1,
   548                                 resolve_tac prems 1,
   549 				strip_tac 1,
   549                                 strip_tac 1,
   550 			        rtac(rewrite_rule axs_take_def finite_ind) 1])
   550                                 rtac(rewrite_rule axs_take_def finite_ind) 1])
   551 				 axs_reach (atomize finite_ind))
   551                                  axs_reach (atomize finite_ind))
   552 )
   552 )
   553 end; (* local *)
   553 end; (* local *)
   554 
   554 
   555 local
   555 local
   556   val xs = mapn (fn n => K (x_name n)) 1 dnames;
   556   val xs = mapn (fn n => K (x_name n)) 1 dnames;
   557   fun bnd_arg n i = Bound(2*(length dnames - n)-i-1);
   557   fun bnd_arg n i = Bound(2*(length dnames - n)-i-1);
   558   val take_ss = HOL_ss addsimps take_rews;
   558   val take_ss = HOL_ss addsimps take_rews;
   559   val sproj   = bin_branchr (fn s => "fst("^s^")") (fn s => "snd("^s^")");
   559   val sproj   = bin_branchr (fn s => "fst("^s^")") (fn s => "snd("^s^")");
   560   val coind_lemma = pg [ax_bisim_def] (mk_trp(mk_imp(%%(comp_dname^"_bisim") $ %"R",
   560   val coind_lemma = pg [ax_bisim_def] (mk_trp(mk_imp(%%(comp_dname^"_bisim") $ %"R",
   561 		foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs,
   561                 foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs,
   562 		  foldr mk_imp (mapn (fn n => K(proj (%"R") dnames n $ 
   562                   foldr mk_imp (mapn (fn n => K(proj (%"R") dnames n $ 
   563 				      bnd_arg n 0 $ bnd_arg n 1)) 0 dnames,
   563                                       bnd_arg n 0 $ bnd_arg n 1)) 0 dnames,
   564 		    foldr' mk_conj (mapn (fn n => fn dn => 
   564                     foldr' mk_conj (mapn (fn n => fn dn => 
   565 				(dc_take dn $ %"n" `bnd_arg n 0 === 
   565                                 (dc_take dn $ %"n" `bnd_arg n 0 === 
   566 				(dc_take dn $ %"n" `bnd_arg n 1))) 0 dnames)))))) ([
   566                                 (dc_take dn $ %"n" `bnd_arg n 1))) 0 dnames)))))) ([
   567 				rtac impI 1,
   567                                 rtac impI 1,
   568 				nat_ind_tac "n" 1,
   568                                 nat_ind_tac "n" 1,
   569 				simp_tac take_ss 1,
   569                                 simp_tac take_ss 1,
   570 				safe_tac HOL_cs] @
   570                                 safe_tac HOL_cs] @
   571 				flat(mapn (fn n => fn x => [
   571                                 flat(mapn (fn n => fn x => [
   572 				  etac allE 1, etac allE 1, 
   572                                   etac allE 1, etac allE 1, 
   573 				  eres_inst_tac [("P1",sproj "R" dnames n^
   573                                   eres_inst_tac [("P1",sproj "R" dnames n^
   574 						  " "^x^" "^x^"'")](mp RS disjE) 1,
   574                                                   " "^x^" "^x^"'")](mp RS disjE) 1,
   575 				  TRY(safe_tac HOL_cs),
   575                                   TRY(safe_tac HOL_cs),
   576 				  REPEAT(CHANGED(asm_simp_tac take_ss 1))]) 
   576                                   REPEAT(CHANGED(asm_simp_tac take_ss 1))]) 
   577 				0 xs));
   577                                 0 xs));
   578 in
   578 in
   579 val coind = pg [] (mk_trp(%%(comp_dname^"_bisim") $ %"R") ===>
   579 val coind = pg [] (mk_trp(%%(comp_dname^"_bisim") $ %"R") ===>
   580 		foldr (op ===>) (mapn (fn n => fn x => 
   580                 foldr (op ===>) (mapn (fn n => fn x => 
   581 			mk_trp(proj (%"R") dnames n $ %x $ %(x^"'"))) 0 xs,
   581                         mk_trp(proj (%"R") dnames n $ %x $ %(x^"'"))) 0 xs,
   582 			mk_trp(foldr' mk_conj (map (fn x => %x === %(x^"'")) xs)))) ([
   582                         mk_trp(foldr' mk_conj (map (fn x => %x === %(x^"'")) xs)))) ([
   583 				TRY(safe_tac HOL_cs)] @
   583                                 TRY(safe_tac HOL_cs)] @
   584 				flat(map (fn take_lemma => [
   584                                 flat(map (fn take_lemma => [
   585 				  rtac take_lemma 1,
   585                                   rtac take_lemma 1,
   586 				  cut_facts_tac [coind_lemma] 1,
   586                                   cut_facts_tac [coind_lemma] 1,
   587 				  fast_tac HOL_cs 1])
   587                                   fast_tac HOL_cs 1])
   588 				take_lemmas));
   588                                 take_lemmas));
   589 end; (* local *)
   589 end; (* local *)
   590 
   590 
   591 
   591 
   592 in (take_rews, take_lemmas, finites, finite_ind, ind, coind)
   592 in (take_rews, take_lemmas, finites, finite_ind, ind, coind)
   593 
   593