src/HOLCF/domain/theorems.ML
author oheimb
Wed, 26 Jun 1996 17:38:34 +0200
changeset 1829 5a3687398716
parent 1781 cc5f55a0fbd7
child 1834 c780a4f39454
permissions -rw-r--r--
function names in when_rews now meta-quantified
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1638
69c094639823 *** empty log message ***
oheimb
parents: 1637
diff changeset
     1
 (* theorems.ML
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     2
   Author : David von Oheimb
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     3
   Created: 06-Jun-95
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     4
   Updated: 08-Jun-95 first proof from cterms
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     5
   Updated: 26-Jun-95 proofs for exhaustion thms
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     6
   Updated: 27-Jun-95 proofs for discriminators, constructors and selectors
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     7
   Updated: 06-Jul-95 proofs for distinctness, invertibility and injectivity
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     8
   Updated: 17-Jul-95 proofs for induction rules
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     9
   Updated: 19-Jul-95 proof for co-induction rule
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    10
   Updated: 28-Aug-95 definedness theorems for selectors (completion)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    11
   Updated: 05-Sep-95 simultaneous domain equations (main part)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    12
   Updated: 11-Sep-95 simultaneous domain equations (coding finished)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    13
   Updated: 13-Sep-95 simultaneous domain equations (debugging)
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    14
   Updated: 26-Oct-95 debugging and enhancement of proofs for take_apps, ind
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    15
   Updated: 16-Feb-96 bug concerning  domain Triv = triv  fixed
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    16
   Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    17
   Copyright 1995, 1996 TU Muenchen
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    18
*)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    19
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    20
structure Domain_Theorems = struct
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    21
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    22
local
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    23
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    24
open Domain_Library;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    25
infixr 0 ===>;infixr 0 ==>;infix 0 == ; 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    26
infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    27
infix 9 `   ; infix 9 `% ; infix 9 `%%; infixr 9 oo;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    28
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    29
(* ----- general proof facilities ------------------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    30
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    31
fun inferT sg pre_tm = #2 (Sign.infer_types sg (K None) (K None) [] true 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    32
			   ([pre_tm],propT));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    33
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    34
fun pg'' thy defs t = let val sg = sign_of thy;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    35
		          val ct = Thm.cterm_of sg (inferT sg t);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    36
		      in prove_goalw_cterm defs ct end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    37
fun pg'  thy defs t tacsf=pg'' thy defs t (fn []   => tacsf 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    38
				| prems=> (cut_facts_tac prems 1)::tacsf);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    39
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    40
fun REPEAT_DETERM_UNTIL p tac = 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    41
let fun drep st = if p st then Sequence.single st
1638
69c094639823 *** empty log message ***
oheimb
parents: 1637
diff changeset
    42
			  else (case Sequence.pull(tac st) of
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    43
		                  None        => Sequence.null
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    44
				| Some(st',_) => drep st')
1638
69c094639823 *** empty log message ***
oheimb
parents: 1637
diff changeset
    45
in drep end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    46
val UNTIL_SOLVED = REPEAT_DETERM_UNTIL (has_fewer_prems 1);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    47
1674
33aff4d854e4 *** empty log message ***
oheimb
parents: 1644
diff changeset
    48
local val trueI2 = prove_goal HOL.thy"f~=x ==> True"(fn _ => [rtac TrueI 1]) in
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    49
val kill_neq_tac = dtac trueI2 end;
1674
33aff4d854e4 *** empty log message ***
oheimb
parents: 1644
diff changeset
    50
fun case_UU_tac rews i v =	case_tac (v^"=UU") i THEN
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    51
				asm_simp_tac (HOLCF_ss addsimps rews) i;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    52
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    53
val chain_tac = REPEAT_DETERM o resolve_tac 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    54
		[is_chain_iterate, ch2ch_fappR, ch2ch_fappL];
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    55
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    56
(* ----- general proofs ----------------------------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    57
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    58
val quant_ss = HOL_ss addsimps (map (fn s => prove_goal HOL.thy s (fn _ =>[
1644
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    59
		fast_tac HOL_cs 1]))["(!x. P x & Q)=((!x. P x) & Q)",
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    60
			    	     "(!x. P & Q x) = (P & (!x. Q x))"]);
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    61
1644
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    62
val all2E = prove_goal HOL.thy "[| !x y . P x y; P x y ==> R |] ==> R"
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    63
 (fn prems =>[
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    64
				resolve_tac prems 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    65
				cut_facts_tac prems 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    66
				fast_tac HOL_cs 1]);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    67
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    68
val swap3 = prove_goal HOL.thy "[| Q ==> P; ~P |] ==> ~Q" (fn prems => [
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    69
                                cut_facts_tac prems 1,
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    70
                                etac swap 1,
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    71
                                dtac notnotD 1,
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    72
				etac (hd prems) 1]);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    73
1644
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    74
val dist_eqI = prove_goal Porder.thy "~ x << y ==> x ~= y" (fn prems => [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    75
                                rtac swap3 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    76
				etac (antisym_less_inverse RS conjunct1) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    77
				resolve_tac prems 1]);
1644
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    78
val cfst_strict  = prove_goal Cprod3.thy "cfst`UU = UU" (fn _ => [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    79
			(simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]);
1644
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    80
val csnd_strict  = prove_goal Cprod3.thy "csnd`UU = UU" (fn _ => [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    81
			(simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    82
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    83
in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    84
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    85
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    86
fun theorems thy (((dname,_),cons) : eq, eqs :eq list) =
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    87
let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    88
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    89
val dummy = writeln ("Proving isomorphism properties of domain "^dname^"...");
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    90
val pg = pg' thy;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    91
(*
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    92
infixr 0 y;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    93
val b = 0;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    94
fun _ y t = by t;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    95
fun  g  defs t = let val sg = sign_of thy;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    96
		     val ct = Thm.cterm_of sg (inferT sg t);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    97
		 in goalw_cterm defs ct end;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    98
*)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    99
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   100
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   101
(* ----- getting the axioms and definitions --------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   102
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   103
local val ga = get_axiom thy in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   104
val ax_abs_iso    = ga (dname^"_abs_iso"   );
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   105
val ax_rep_iso    = ga (dname^"_rep_iso"   );
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   106
val ax_when_def   = ga (dname^"_when_def"  );
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   107
val axs_con_def   = map (fn (con,_) => ga (extern_name con ^"_def")) cons;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   108
val axs_dis_def   = map (fn (con,_) => ga (   dis_name con ^"_def")) cons;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   109
val axs_sel_def   = flat(map (fn (_,args) => 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   110
		    map (fn     arg => ga (sel_of arg      ^"_def")) args)cons);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   111
val ax_copy_def   = ga (dname^"_copy_def"  );
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   112
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   113
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   114
(* ----- theorems concerning the isomorphism -------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   115
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   116
val dc_abs  = %%(dname^"_abs");
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   117
val dc_rep  = %%(dname^"_rep");
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   118
val dc_copy = %%(dname^"_copy");
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   119
val x_name = "x";
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   120
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   121
val (rep_strict, abs_strict) = let 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   122
	 val r = ax_rep_iso RS (ax_abs_iso RS (allI  RSN(2,allI RS iso_strict)))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   123
	       in (r RS conjunct1, r RS conjunct2) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   124
val abs_defin' = pg [] ((dc_abs`%x_name === UU) ==> (%x_name === UU)) [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   125
			   res_inst_tac [("t",x_name)] (ax_abs_iso RS subst) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   126
				etac ssubst 1, rtac rep_strict 1];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   127
val rep_defin' = pg [] ((dc_rep`%x_name === UU) ==> (%x_name === UU)) [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   128
			   res_inst_tac [("t",x_name)] (ax_rep_iso RS subst) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   129
				etac ssubst 1, rtac abs_strict 1];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   130
val iso_rews = [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   131
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   132
local 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   133
val iso_swap = pg [] (dc_rep`%"x" === %"y" ==> %"x" === dc_abs`%"y") [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   134
			    dres_inst_tac [("f",dname^"_abs")] cfun_arg_cong 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   135
			    etac (ax_rep_iso RS subst) 1];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   136
fun exh foldr1 cn quant foldr2 var = let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   137
  fun one_con (con,args) = let val vns = map vname args in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   138
    foldr quant (vns, foldr2 ((%x_name === con_app2 con (var vns) vns)::
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   139
			      map (defined o (var vns)) (nonlazy args))) end
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   140
  in foldr1 ((cn(%x_name===UU))::map one_con cons) end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   141
in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   142
val cases = let 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   143
	    fun common_tac thm = rtac thm 1 THEN contr_tac 1;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   144
	    fun unit_tac true = common_tac liftE1
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   145
	    |   unit_tac _    = all_tac;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   146
	    fun prod_tac []          = common_tac oneE
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   147
	    |   prod_tac [arg]       = unit_tac (is_lazy arg)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   148
	    |   prod_tac (arg::args) = 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   149
				common_tac sprodE THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   150
				kill_neq_tac 1 THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   151
				unit_tac (is_lazy arg) THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   152
				prod_tac args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   153
	    fun sum_rest_tac p = SELECT_GOAL(EVERY[
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   154
				rtac p 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   155
				rewrite_goals_tac axs_con_def,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   156
				dtac iso_swap 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   157
				simp_tac HOLCF_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   158
				UNTIL_SOLVED(fast_tac HOL_cs 1)]) 1;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   159
	    fun sum_tac [(_,args)]       [p]        = 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   160
				prod_tac args THEN sum_rest_tac p
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   161
	    |   sum_tac ((_,args)::cons') (p::prems) = DETERM(
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   162
				common_tac ssumE THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   163
				kill_neq_tac 1 THEN kill_neq_tac 2 THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   164
				prod_tac args THEN sum_rest_tac p) THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   165
				sum_tac cons' prems
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   166
	    |   sum_tac _ _ = Imposs "theorems:sum_tac";
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   167
	  in pg'' thy [] (exh (fn l => foldr (op ===>) (l,mk_trp(%"P")))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   168
			      (fn T => T ==> %"P") mk_All
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   169
			      (fn l => foldr (op ===>) (map mk_trp l,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   170
							    mk_trp(%"P")))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   171
			      bound_arg)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   172
			     (fn prems => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   173
				cut_facts_tac [excluded_middle] 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   174
				etac disjE 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   175
				rtac (hd prems) 2,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   176
				etac rep_defin' 2,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   177
				if length cons = 1 andalso 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   178
				   length (snd(hd cons)) = 1 andalso 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   179
				   not(is_lazy(hd(snd(hd cons))))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   180
				then rtac (hd (tl prems)) 1 THEN atac 2 THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   181
				     rewrite_goals_tac axs_con_def THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   182
				     simp_tac (HOLCF_ss addsimps [ax_rep_iso]) 1
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   183
				else sum_tac cons (tl prems)])end;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   184
val exhaust= pg[](mk_trp(exh (foldr' mk_disj) Id mk_ex (foldr' mk_conj) (K %)))[
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   185
				rtac cases 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   186
				UNTIL_SOLVED(fast_tac HOL_cs 1)];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   187
end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   188
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   189
local 
1829
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   190
  fun bind_fun t = foldr mk_All ((if length cons = 1 then ["f"] 
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   191
		  else mapn (fn n => K("f"^(string_of_int n))) 1 cons),t);
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   192
  fun bound_fun i _ = Bound (length cons - i);
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   193
  val when_app  = foldl (op `) (%%(dname^"_when"), mapn bound_fun 1 cons);
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   194
  val when_appl = pg [ax_when_def] (bind_fun(mk_trp(when_app`%x_name ===
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   195
	     when_body cons (fn (m,n)=> bound_fun (n-m) 0)`(dc_rep`%x_name))))[
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   196
				simp_tac HOLCF_ss 1];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   197
in
1829
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   198
val when_strict = pg [] (bind_fun(mk_trp(strict when_app))) [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   199
			simp_tac(HOLCF_ss addsimps [when_appl,rep_strict]) 1];
1829
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   200
val when_apps = let fun one_when n (con,args) = pg axs_con_def 
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   201
		(bind_fun (lift_defined % (nonlazy args, 
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   202
		mk_trp(when_app`(con_app con args) ===
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   203
                       mk_cfapp(bound_fun n 0,map %# args)))))[
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   204
		asm_simp_tac (HOLCF_ss addsimps [when_appl,ax_abs_iso]) 1];
1829
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   205
	in mapn one_when 1 cons end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   206
end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   207
val when_rews = when_strict::when_apps;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   208
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   209
(* ----- theorems concerning the constructors, discriminators and selectors - *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   210
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   211
val dis_rews = let
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   212
  val dis_stricts = map (fn (con,_) => pg axs_dis_def (mk_trp(
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   213
		      	     strict(%%(dis_name con)))) [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   214
				simp_tac (HOLCF_ss addsimps when_rews) 1]) cons;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   215
  val dis_apps = let fun one_dis c (con,args)= pg axs_dis_def
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   216
		   (lift_defined % (nonlazy args,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   217
			(mk_trp((%%(dis_name c))`(con_app con args) ===
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   218
			      %%(if con=c then "TT" else "FF"))))) [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   219
				asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   220
	in flat(map (fn (c,_) => map (one_dis c) cons) cons) end;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   221
  val dis_defins = map (fn (con,args) => pg [] (defined(%x_name) ==> 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   222
		      defined(%%(dis_name con)`%x_name)) [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   223
				rtac cases 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   224
				contr_tac 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   225
				UNTIL_SOLVED (CHANGED(asm_simp_tac 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   226
				        (HOLCF_ss addsimps dis_apps) 1))]) cons;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   227
in dis_stricts @ dis_defins @ dis_apps end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   228
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   229
val con_stricts = flat(map (fn (con,args) => map (fn vn =>
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   230
			pg (axs_con_def) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   231
			   (mk_trp(con_app2 con (fn arg => if vname arg = vn 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   232
					then UU else %# arg) args === UU))[
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   233
				asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1]
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   234
			) (nonlazy args)) cons);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   235
val con_defins = map (fn (con,args) => pg []
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   236
			(lift_defined % (nonlazy args,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   237
				mk_trp(defined(con_app con args)))) ([
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   238
			  rtac swap3 1, 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   239
			  eres_inst_tac [("f",dis_name con)] cfun_arg_cong 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   240
			  asm_simp_tac (HOLCF_ss addsimps dis_rews) 1] )) cons;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   241
val con_rews = con_stricts @ con_defins;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   242
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   243
val sel_stricts = let fun one_sel sel = pg axs_sel_def (mk_trp(strict(%%sel))) [
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   244
				simp_tac (HOLCF_ss addsimps when_rews) 1];
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   245
in flat(map (fn (_,args) =>map (fn arg => one_sel (sel_of arg)) args) cons) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   246
val sel_apps = let fun one_sel c n sel = map (fn (con,args) => 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   247
		let val nlas = nonlazy args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   248
		    val vns  = map vname args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   249
		in pg axs_sel_def (lift_defined %
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   250
		   (filter (fn v => con=c andalso (v<>nth_elem(n,vns))) nlas,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   251
				mk_trp((%%sel)`(con_app con args) === 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   252
				(if con=c then %(nth_elem(n,vns)) else UU))))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   253
			    ( (if con=c then [] 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   254
		       else map(case_UU_tac(when_rews@con_stricts)1) nlas)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   255
		     @(if con=c andalso ((nth_elem(n,vns)) mem nlas)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   256
				 then[case_UU_tac (when_rews @ con_stricts) 1 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   257
						  (nth_elem(n,vns))] else [])
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   258
		     @ [asm_simp_tac(HOLCF_ss addsimps when_rews)1])end) cons;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   259
in flat(map  (fn (c,args) => 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   260
     flat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   261
val sel_defins = if length cons=1 then map (fn arg => pg [](defined(%x_name)==> 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   262
			defined(%%(sel_of arg)`%x_name)) [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   263
				rtac cases 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   264
				contr_tac 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   265
				UNTIL_SOLVED (CHANGED(asm_simp_tac 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   266
				             (HOLCF_ss addsimps sel_apps) 1))]) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   267
		 (filter_out is_lazy (snd(hd cons))) else [];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   268
val sel_rews = sel_stricts @ sel_defins @ sel_apps;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   269
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   270
val distincts_le = let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   271
    fun dist (con1, args1) (con2, args2) = pg []
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   272
	      (lift_defined % ((nonlazy args1),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   273
			(mk_trp (con_app con1 args1 ~<< con_app con2 args2))))([
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   274
			rtac swap3 1,
1781
cc5f55a0fbd7 adapted use of monofun_cfun_arg
oheimb
parents: 1674
diff changeset
   275
			eres_inst_tac[("fo",dis_name con1)] monofun_cfun_arg 1]
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   276
		      @map(case_UU_tac (con_stricts @ dis_rews)1)(nonlazy args2)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   277
		      @[asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   278
    fun distinct (con1,args1) (con2,args2) =
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   279
	let val arg1 = (con1, args1);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   280
	    val arg2 = (con2, (map (fn (arg,vn) => upd_vname (K vn) arg)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   281
			(args2~~variantlist(map vname args2,map vname args1))));
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   282
	in [dist arg1 arg2, dist arg2 arg1] end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   283
    fun distincts []      = []
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   284
    |   distincts (c::cs) = (map (distinct c) cs) :: distincts cs;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   285
in distincts cons end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   286
val dists_le = flat (flat distincts_le);
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   287
val dists_eq = let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   288
    fun distinct (_,args1) ((_,args2),leqs) = let
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   289
	val (le1,le2) = (hd leqs, hd(tl leqs));
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   290
	val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) in
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   291
	if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   292
	if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   293
					[eq1, eq2] end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   294
    fun distincts []      = []
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   295
    |   distincts ((c,leqs)::cs) = flat(map (distinct c) ((map fst cs)~~leqs)) @
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   296
				   distincts cs;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   297
    in distincts (cons~~distincts_le) end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   298
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   299
local 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   300
  fun pgterm rel con args = let
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   301
		fun append s = upd_vname(fn v => v^s);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   302
		val (largs,rargs) = (args, map (append "'") args);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   303
		in pg [] (mk_trp (rel(con_app con largs,con_app con rargs)) ===>
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   304
		      lift_defined % ((nonlazy largs),lift_defined % ((nonlazy rargs),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   305
			    mk_trp (foldr' mk_conj 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   306
				(map rel (map %# largs ~~ map %# rargs)))))) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   307
  val cons' = filter (fn (_,args) => args<>[]) cons;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   308
in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   309
val inverts = map (fn (con,args) => 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   310
		pgterm (op <<) con args (flat(map (fn arg => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   311
				TRY(rtac conjI 1),
1781
cc5f55a0fbd7 adapted use of monofun_cfun_arg
oheimb
parents: 1674
diff changeset
   312
				dres_inst_tac [("fo",sel_of arg)] monofun_cfun_arg 1,
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   313
				asm_full_simp_tac (HOLCF_ss addsimps sel_apps) 1]
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   314
			     			      ) args))) cons';
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   315
val injects = map (fn ((con,args),inv_thm) => 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   316
			   pgterm (op ===) con args [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   317
				etac (antisym_less_inverse RS conjE) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   318
				dtac inv_thm 1, REPEAT(atac 1),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   319
				dtac inv_thm 1, REPEAT(atac 1),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   320
				TRY(safe_tac HOL_cs),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   321
				REPEAT(rtac antisym_less 1 ORELSE atac 1)] )
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   322
		  (cons'~~inverts);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   323
end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   324
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   325
(* ----- theorems concerning one induction step ----------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   326
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   327
val copy_strict = pg[ax_copy_def](mk_trp(strict(dc_copy`%"f"))) [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   328
		   asm_simp_tac(HOLCF_ss addsimps [abs_strict, when_strict,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   329
						   cfst_strict,csnd_strict]) 1];
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   330
val copy_apps = map (fn (con,args) => pg [ax_copy_def]
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   331
		    (lift_defined % (nonlazy_rec args,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   332
			mk_trp(dc_copy`%"f"`(con_app con args) ===
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   333
		(con_app2 con (app_rec_arg (cproj (%"f") (length eqs))) args))))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   334
			(map (case_UU_tac (abs_strict::when_strict::con_stricts)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   335
				 1 o vname)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   336
			 (filter (fn a => not (is_rec a orelse is_lazy a)) args)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   337
			@[asm_simp_tac (HOLCF_ss addsimps when_apps) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   338
		          simp_tac (HOLCF_ss addsimps axs_con_def) 1]))cons;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   339
val copy_stricts = map (fn (con,args) => pg [] (mk_trp(dc_copy`UU`
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   340
					(con_app con args) ===UU))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   341
     (let val rews = cfst_strict::csnd_strict::copy_strict::copy_apps@con_rews
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   342
			 in map (case_UU_tac rews 1) (nonlazy args) @ [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   343
			     asm_simp_tac (HOLCF_ss addsimps rews) 1] end))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   344
  		        (filter (fn (_,args)=>exists is_nonlazy_rec args) cons);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   345
val copy_rews = copy_strict::copy_apps @ copy_stricts;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   346
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   347
in     (iso_rews, exhaust, cases, when_rews,
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   348
	con_rews, sel_rews, dis_rews, dists_le, dists_eq, inverts, injects,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   349
	copy_rews)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   350
end; (* let *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   351
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   352
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   353
fun comp_theorems thy (comp_dname, eqs: eq list, casess, con_rews, copy_rews) =
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   354
let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   355
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   356
val dummy = writeln("Proving induction properties of domain "^comp_dname^"...");
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   357
val pg = pg' thy;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   358
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   359
val dnames = map (fst o fst) eqs;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   360
val conss  = map  snd        eqs;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   361
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   362
(* ----- getting the composite axiom and definitions ------------------------ *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   363
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   364
local val ga = get_axiom thy in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   365
val axs_reach      = map (fn dn => ga (dn ^  "_reach"   )) dnames;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   366
val axs_take_def   = map (fn dn => ga (dn ^  "_take_def")) dnames;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   367
val axs_finite_def = map (fn dn => ga (dn ^"_finite_def")) dnames;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   368
val ax_copy2_def   = ga (comp_dname^ "_copy_def");
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   369
val ax_bisim_def   = ga (comp_dname^"_bisim_def");
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   370
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   371
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   372
fun dc_take dn = %%(dn^"_take");
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   373
val x_name = idx_name dnames "x"; 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   374
val P_name = idx_name dnames "P";
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   375
val n_eqs = length eqs;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   376
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   377
(* ----- theorems concerning finite approximation and finite induction ------ *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   378
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   379
local
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   380
  val iterate_Cprod_ss = simpset_of "Fix"
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   381
			 addsimps [cfst_strict, csnd_strict]addsimps Cprod_rews;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   382
  val copy_con_rews  = copy_rews @ con_rews;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   383
  val copy_take_defs =(if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   384
  val take_stricts=pg copy_take_defs(mk_trp(foldr' mk_conj(map(fn((dn,args),_)=>
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   385
	    (dc_take dn $ %"n")`UU === mk_constrain(Type(dn,args),UU)) eqs)))([
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   386
			nat_ind_tac "n" 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   387
			simp_tac iterate_Cprod_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   388
			asm_simp_tac (iterate_Cprod_ss addsimps copy_rews)1]);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   389
  val take_stricts' = rewrite_rule copy_take_defs take_stricts;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   390
  val take_0s = mapn(fn n=> fn dn => pg axs_take_def(mk_trp((dc_take dn $ %%"0")
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   391
							`%x_name n === UU))[
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   392
				simp_tac iterate_Cprod_ss 1]) 1 dnames;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   393
  val c_UU_tac = case_UU_tac (take_stricts'::copy_con_rews) 1;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   394
  val take_apps = pg copy_take_defs (mk_trp(foldr' mk_conj 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   395
	    (flat(map (fn ((dn,_),cons) => map (fn (con,args) => foldr mk_all 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   396
	(map vname args,(dc_take dn $ (%%"Suc" $ %"n"))`(con_app con args) ===
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   397
  	 con_app2 con (app_rec_arg (fn n=>dc_take (nth_elem(n,dnames))$ %"n"))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   398
			      args)) cons) eqs)))) ([
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   399
				simp_tac iterate_Cprod_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   400
				nat_ind_tac "n" 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   401
			    simp_tac(iterate_Cprod_ss addsimps copy_con_rews) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   402
				asm_full_simp_tac (HOLCF_ss addsimps 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   403
				      (filter (has_fewer_prems 1) copy_rews)) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   404
				TRY(safe_tac HOL_cs)] @
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   405
			(flat(map (fn ((dn,_),cons) => map (fn (con,args) => 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   406
				if nonlazy_rec args = [] then all_tac else
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   407
				EVERY(map c_UU_tac (nonlazy_rec args)) THEN
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   408
				asm_full_simp_tac (HOLCF_ss addsimps copy_rews)1
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   409
		 					   ) cons) eqs)));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   410
in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   411
val take_rews = atomize take_stricts @ take_0s @ atomize take_apps;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   412
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   413
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   414
local
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   415
  fun one_con p (con,args) = foldr mk_All (map vname args,
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   416
	lift_defined (bound_arg (map vname args)) (nonlazy args,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   417
	lift (fn arg => %(P_name (1+rec_of arg)) $ bound_arg args arg)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   418
         (filter is_rec args,mk_trp(%p $ con_app2 con (bound_arg args) args))));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   419
  fun one_eq ((p,cons),concl) = (mk_trp(%p $ UU) ===> 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   420
			   foldr (op ===>) (map (one_con p) cons,concl));
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   421
  fun ind_term concf = foldr one_eq (mapn (fn n => fn x => (P_name n, x))1conss,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   422
			mk_trp(foldr' mk_conj (mapn concf 1 dnames)));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   423
  val take_ss = HOL_ss addsimps take_rews;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   424
  fun quant_tac i = EVERY(mapn(fn n=> fn _=> res_inst_tac[("x",x_name n)]spec i)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   425
			       1 dnames);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   426
  fun ind_prems_tac prems = EVERY(flat (map (fn cons => (
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   427
				     resolve_tac prems 1 ::
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   428
				     flat (map (fn (_,args) => 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   429
				       resolve_tac prems 1 ::
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   430
				       map (K(atac 1)) (nonlazy args) @
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   431
				       map (K(atac 1)) (filter is_rec args))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   432
				     cons))) conss));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   433
  local 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   434
    (* check whether every/exists constructor of the n-th part of the equation:
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   435
       it has a possibly indirectly recursive argument that isn't/is possibly 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   436
       indirectly lazy *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   437
    fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   438
	  is_rec arg andalso not(rec_of arg mem ns) andalso
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   439
	  ((rec_of arg =  n andalso nfn(lazy_rec orelse is_lazy arg)) orelse 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   440
	    rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   441
	      (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss))))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   442
	  ) o snd) cons;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   443
    fun all_rec_to ns  = rec_to forall not all_rec_to  ns;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   444
    fun warn (n,cons)  = if all_rec_to [] false (n,cons) then (writeln 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   445
        ("WARNING: domain "^nth_elem(n,dnames)^" is empty!"); true) else false;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   446
    fun lazy_rec_to ns = rec_to exists Id  lazy_rec_to ns;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   447
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   448
  in val n__eqs     = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   449
     val is_emptys = map warn n__eqs;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   450
     val is_finite = forall (not o lazy_rec_to [] false) n__eqs;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   451
  end;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   452
in (* local *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   453
val finite_ind = pg'' thy [] (ind_term (fn n => fn dn => %(P_name n)$
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   454
			     (dc_take dn $ %"n" `%(x_name n)))) (fn prems => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   455
				quant_tac 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   456
				simp_tac quant_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   457
				nat_ind_tac "n" 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   458
				simp_tac (take_ss addsimps prems) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   459
				TRY(safe_tac HOL_cs)]
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   460
				@ flat(map (fn (cons,cases) => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   461
				 res_inst_tac [("x","x")] cases 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   462
				 asm_simp_tac (take_ss addsimps prems) 1]
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   463
				 @ flat(map (fn (con,args) => 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   464
				  asm_simp_tac take_ss 1 ::
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   465
				  map (fn arg =>
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   466
				   case_UU_tac (prems@con_rews) 1 (
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   467
			   nth_elem(rec_of arg,dnames)^"_take n1`"^vname arg))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   468
				  (filter is_nonlazy_rec args) @ [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   469
				  resolve_tac prems 1] @
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   470
				  map (K (atac 1))      (nonlazy args) @
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   471
				  map (K (etac spec 1)) (filter is_rec args)) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   472
				 cons))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   473
				(conss~~casess)));
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   474
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   475
val take_lemmas =mapn(fn n=> fn(dn,ax_reach)=> pg'' thy axs_take_def(mk_All("n",
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   476
		mk_trp(dc_take dn $ Bound 0 `%(x_name n) === 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   477
		       dc_take dn $ Bound 0 `%(x_name n^"'")))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   478
	   ===> mk_trp(%(x_name n) === %(x_name n^"'"))) (fn prems => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   479
			res_inst_tac[("t",x_name n    )](ax_reach RS subst) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   480
			res_inst_tac[("t",x_name n^"'")](ax_reach RS subst) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   481
				rtac (fix_def2 RS ssubst) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   482
				REPEAT(CHANGED(rtac(contlub_cfun_arg RS ssubst)1
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   483
					       THEN chain_tac 1)),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   484
				rtac (contlub_cfun_fun RS ssubst) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   485
				rtac (contlub_cfun_fun RS ssubst) 2,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   486
				rtac lub_equal 3,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   487
				chain_tac 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   488
				rtac allI 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   489
				resolve_tac prems 1])) 1 (dnames~~axs_reach);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   490
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   491
(* ----- theorems concerning finiteness and induction ----------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   492
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   493
val (finites,ind) = if is_finite then
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   494
  let 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   495
    fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %"x" === %"x");
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   496
    val finite_lemmas1a = map (fn dn => pg [] (mk_trp(defined (%"x")) ===> 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   497
	mk_trp(mk_disj(mk_all("n",dc_take dn $ Bound 0 ` %"x" === UU),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   498
	take_enough dn)) ===> mk_trp(take_enough dn)) [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   499
				etac disjE 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   500
				etac notE 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   501
				resolve_tac take_lemmas 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   502
				asm_simp_tac take_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   503
				atac 1]) dnames;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   504
    val finite_lemma1b = pg [] (mk_trp (mk_all("n",foldr' mk_conj (mapn 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   505
	(fn n => fn ((dn,args),_) => mk_constrainall(x_name n,Type(dn,args),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   506
	 mk_disj(dc_take dn $ Bound 1 ` Bound 0 === UU,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   507
		 dc_take dn $ Bound 1 ` Bound 0 === Bound 0))) 1 eqs)))) ([
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   508
				rtac allI 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   509
				nat_ind_tac "n" 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   510
				simp_tac take_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   511
			TRY(safe_tac(empty_cs addSEs[conjE] addSIs[conjI]))] @
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   512
				flat(mapn (fn n => fn (cons,cases) => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   513
				  simp_tac take_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   514
				  rtac allI 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   515
				  res_inst_tac [("x",x_name n)] cases 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   516
				  asm_simp_tac take_ss 1] @ 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   517
				  flat(map (fn (con,args) => 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   518
				    asm_simp_tac take_ss 1 ::
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   519
				    flat(map (fn vn => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   520
				      eres_inst_tac [("x",vn)] all_dupE 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   521
				      etac disjE 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   522
				      asm_simp_tac (HOL_ss addsimps con_rews) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   523
				      asm_simp_tac take_ss 1])
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   524
				    (nonlazy_rec args)))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   525
				  cons))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   526
				1 (conss~~casess))) handle ERROR => raise ERROR;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   527
    val finites = map (fn (dn,l1b) => pg axs_finite_def (mk_trp(
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   528
						%%(dn^"_finite") $ %"x"))[
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   529
				case_UU_tac take_rews 1 "x",
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   530
				eresolve_tac finite_lemmas1a 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   531
				step_tac HOL_cs 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   532
				step_tac HOL_cs 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   533
				cut_facts_tac [l1b] 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   534
			fast_tac HOL_cs 1]) (dnames~~atomize finite_lemma1b);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   535
  in
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   536
  (finites,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   537
   pg'' thy[](ind_term (fn n => fn dn => %(P_name n) $ %(x_name n)))(fn prems =>
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   538
				TRY(safe_tac HOL_cs) ::
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   539
			 flat (map (fn (finite,fin_ind) => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   540
			       rtac(rewrite_rule axs_finite_def finite RS exE)1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   541
				etac subst 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   542
				rtac fin_ind 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   543
				ind_prems_tac prems]) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   544
			           (finites~~(atomize finite_ind)) ))
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   545
) end (* let *) else
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   546
  (mapn (fn n => fn dn => read_instantiate_sg (sign_of thy) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   547
	  	    [("P",dn^"_finite "^x_name n)] excluded_middle) 1 dnames,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   548
   pg'' thy [] (foldr (op ===>) (mapn (fn n => K(mk_trp(%%"adm" $ %(P_name n))))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   549
	       1 dnames, ind_term (fn n => fn dn => %(P_name n) $ %(x_name n))))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   550
		   (fn prems => map (fn ax_reach => rtac (ax_reach RS subst) 1) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   551
				    axs_reach @ [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   552
				quant_tac 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   553
				rtac (adm_impl_admw RS wfix_ind) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   554
				REPEAT_DETERM(rtac adm_all2 1),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   555
				REPEAT_DETERM(TRY(rtac adm_conj 1) THEN 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   556
						  rtac adm_subst 1 THEN 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   557
					cont_tacR 1 THEN resolve_tac prems 1),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   558
				strip_tac 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   559
				rtac (rewrite_rule axs_take_def finite_ind) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   560
				ind_prems_tac prems])
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   561
)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   562
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   563
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   564
(* ----- theorem concerning coinduction ------------------------------------- *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   565
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   566
local
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   567
  val xs = mapn (fn n => K (x_name n)) 1 dnames;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   568
  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   569
  val take_ss = HOL_ss addsimps take_rews;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   570
  val sproj   = prj (fn s => "fst("^s^")") (fn s => "snd("^s^")");
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   571
  val coind_lemma=pg[ax_bisim_def](mk_trp(mk_imp(%%(comp_dname^"_bisim") $ %"R",
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   572
		foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   573
		  foldr mk_imp (mapn (fn n => K(proj (%"R") n_eqs n $ 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   574
				      bnd_arg n 0 $ bnd_arg n 1)) 0 dnames,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   575
		    foldr' mk_conj (mapn (fn n => fn dn => 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   576
				(dc_take dn $ %"n" `bnd_arg n 0 === 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   577
				(dc_take dn $ %"n" `bnd_arg n 1)))0 dnames))))))
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   578
			     ([ rtac impI 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   579
				nat_ind_tac "n" 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   580
				simp_tac take_ss 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   581
				safe_tac HOL_cs] @
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   582
				flat(mapn (fn n => fn x => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   583
				  rotate_tac (n+1) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   584
				  etac all2E 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   585
				  eres_inst_tac [("P1", sproj "R" n_eqs n^
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   586
					" "^x^" "^x^"'")](mp RS disjE) 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   587
				  TRY(safe_tac HOL_cs),
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   588
				  REPEAT(CHANGED(asm_simp_tac take_ss 1))]) 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   589
				0 xs));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   590
in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   591
val coind = pg [] (mk_trp(%%(comp_dname^"_bisim") $ %"R") ===>
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   592
		foldr (op ===>) (mapn (fn n => fn x => 
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   593
		  mk_trp(proj (%"R") n_eqs n $ %x $ %(x^"'"))) 0 xs,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   594
		  mk_trp(foldr' mk_conj (map (fn x => %x === %(x^"'")) xs)))) ([
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   595
				TRY(safe_tac HOL_cs)] @
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   596
				flat(map (fn take_lemma => [
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   597
				  rtac take_lemma 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   598
				  cut_facts_tac [coind_lemma] 1,
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   599
				  fast_tac HOL_cs 1])
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   600
				take_lemmas));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   601
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   602
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   603
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   604
in (take_rews, take_lemmas, finites, finite_ind, ind, coind)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   605
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   606
end; (* let *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   607
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   608
end; (* struct *)