src/HOLCF/domain/theorems.ML
author wenzelm
Sat, 30 Dec 2006 16:08:10 +0100
changeset 21969 a8bf1106cb7c
parent 21619 dea0914773f7
child 22578 b0eb5652f210
permissions -rw-r--r--
removed conditional combinator; avoid handle _; showctxt: print_context (cf. local theory context); searchtheorems: proper find_theorems; refrain from setting ml_prompts again; tuned init_pgip;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2276
3eb9a113029e modified file headers
oheimb
parents: 2267
diff changeset
     1
(*  Title:      HOLCF/domain/theorems.ML
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
     2
    ID:         $Id$
12030
wenzelm
parents: 11531
diff changeset
     3
    Author:     David von Oheimb
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
     4
                New proofs/tactics by Brian Huffman
2445
51993fea433f removed Holcfb.thy and Holcfb.ML, moving classical3 to HOL.ML as classical2
oheimb
parents: 2276
diff changeset
     5
12030
wenzelm
parents: 11531
diff changeset
     6
Proof generator for domain section.
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     7
*)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     8
16842
wenzelm
parents: 16778
diff changeset
     9
val HOLCF_ss = simpset();
2446
c2a9bf6c0948 The previous log message was wrong. The correct one is:
oheimb
parents: 2445
diff changeset
    10
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    11
structure Domain_Theorems = struct
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    12
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    13
local
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    14
19764
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    15
val adm_impl_admw = thm "adm_impl_admw";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    16
val antisym_less_inverse = thm "antisym_less_inverse";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    17
val beta_cfun = thm "beta_cfun";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    18
val cfun_arg_cong = thm "cfun_arg_cong";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    19
val ch2ch_Rep_CFunL = thm "ch2ch_Rep_CFunL";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    20
val ch2ch_Rep_CFunR = thm "ch2ch_Rep_CFunR";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    21
val chain_iterate = thm "chain_iterate";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    22
val compact_ONE = thm "compact_ONE";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    23
val compact_sinl = thm "compact_sinl";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    24
val compact_sinr = thm "compact_sinr";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    25
val compact_spair = thm "compact_spair";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    26
val compact_up = thm "compact_up";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    27
val contlub_cfun_arg = thm "contlub_cfun_arg";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    28
val contlub_cfun_fun = thm "contlub_cfun_fun";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    29
val fix_def2 = thm "fix_def2";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    30
val injection_eq = thm "injection_eq";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    31
val injection_less = thm "injection_less";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    32
val lub_equal = thm "lub_equal";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    33
val monofun_cfun_arg = thm "monofun_cfun_arg";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    34
val retraction_strict = thm "retraction_strict";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    35
val spair_eq = thm "spair_eq";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    36
val spair_less = thm "spair_less";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    37
val sscase1 = thm "sscase1";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    38
val ssplit1 = thm "ssplit1";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    39
val strictify1 = thm "strictify1";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    40
val wfix_ind = thm "wfix_ind";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
    41
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    42
open Domain_Library;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    43
infixr 0 ===>;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    44
infixr 0 ==>;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    45
infix 0 == ; 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    46
infix 1 ===;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    47
infix 1 ~= ;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    48
infix 1 <<;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    49
infix 1 ~<<;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    50
infix 9 `   ;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    51
infix 9 `% ;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    52
infix 9 `%%;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    53
infixr 9 oo;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    54
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    55
(* ----- general proof facilities ------------------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    56
18972
2905d1805e1e adapted Sign.infer_types;
wenzelm
parents: 18728
diff changeset
    57
(* FIXME better avoid this low-level stuff *)
14820
3f80d6510ee9 Sign.infer_types: Sign.pp;
wenzelm
parents: 13454
diff changeset
    58
fun inferT sg pre_tm =
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    59
  let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    60
    val pp = Sign.pp sg;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    61
    val consts = Sign.consts_of sg;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    62
    val (t, _) =
20155
da0505518e69 Sign.infer_types: Name.context;
wenzelm
parents: 20071
diff changeset
    63
      Sign.infer_types pp sg consts (K NONE) (K NONE) Name.context true
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    64
        ([pre_tm],propT);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    65
  in t end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    66
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
    67
fun pg'' thy defs t tacs =
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    68
  let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    69
    val t' = inferT thy t;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    70
    val asms = Logic.strip_imp_prems t';
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    71
    val prop = Logic.strip_imp_concl t';
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    72
    fun tac prems =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    73
      rewrite_goals_tac defs THEN
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    74
      EVERY (tacs (map (rewrite_rule defs) prems));
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19764
diff changeset
    75
  in Goal.prove_global thy [] asms prop tac end;
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
    76
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    77
fun pg' thy defs t tacsf =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    78
  let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    79
    fun tacs [] = tacsf
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    80
      | tacs prems = cut_facts_tac prems 1 :: tacsf;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    81
  in pg'' thy defs t tacs end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    82
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    83
fun case_UU_tac rews i v =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    84
  case_tac (v^"=UU") i THEN
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    85
  asm_simp_tac (HOLCF_ss addsimps rews) i;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    86
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    87
val chain_tac =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    88
  REPEAT_DETERM o resolve_tac 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    89
    [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL];
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    90
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
    91
(* ----- general proofs ----------------------------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    92
1644
681f70ca3cf7 Removed 8bit characters used by mistake
oheimb
parents: 1638
diff changeset
    93
val all2E = prove_goal HOL.thy "[| !x y . P x y; P x y ==> R |] ==> R"
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    94
  (fn prems =>[
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    95
    resolve_tac prems 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    96
    cut_facts_tac prems 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
    97
    fast_tac HOL_cs 1]);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    98
18057
ad97e231bf8a dist_eqI: the_context();
wenzelm
parents: 17985
diff changeset
    99
val dist_eqI = prove_goal (the_context ()) "!!x::'a::po. ~ x << y ==> x ~= y" 
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   100
  (fn prems =>
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   101
    [blast_tac (claset () addDs [antisym_less_inverse]) 1]);
4755
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   102
(*
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   103
infixr 0 y;
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   104
val b = 0;
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   105
fun _ y t = by t;
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   106
fun g defs t = let val sg = sign_of thy;
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   107
                     val ct = Thm.cterm_of sg (inferT sg t);
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   108
                 in goalw_cterm defs ct end;
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4721
diff changeset
   109
*)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   110
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   111
in
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   112
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   113
fun theorems (((dname, _), cons) : eq, eqs : eq list) thy =
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   114
let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   115
4030
ca44afcc259c debugging concerning sort variables
oheimb
parents: 4008
diff changeset
   116
val dummy = writeln ("Proving isomorphism properties of domain "^dname^" ...");
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   117
val pg = pg' thy;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   118
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   119
(* ----- getting the axioms and definitions --------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   120
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   121
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   122
  fun ga s dn = get_thm thy (Name (dn ^ "." ^ s));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   123
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   124
  val ax_abs_iso  = ga "abs_iso"  dname;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   125
  val ax_rep_iso  = ga "rep_iso"  dname;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   126
  val ax_when_def = ga "when_def" dname;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   127
  fun get_def mk_name (con,_) = ga (mk_name con^"_def") dname;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   128
  val axs_con_def = map (get_def extern_name) cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   129
  val axs_dis_def = map (get_def dis_name) cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   130
  val axs_mat_def = map (get_def mat_name) cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   131
  val axs_pat_def = map (get_def pat_name) cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   132
  val axs_sel_def =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   133
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   134
      fun def_of_sel sel = ga (sel^"_def") dname;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   135
      fun def_of_arg arg = Option.map def_of_sel (sel_of arg);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   136
      fun defs_of_con (_, args) = List.mapPartial def_of_arg args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   137
    in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   138
      List.concat (map defs_of_con cons)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   139
    end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   140
  val ax_copy_def = ga "copy_def" dname;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   141
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   142
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   143
(* ----- theorems concerning the isomorphism -------------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   144
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 10835
diff changeset
   145
val dc_abs  = %%:(dname^"_abs");
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 10835
diff changeset
   146
val dc_rep  = %%:(dname^"_rep");
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 10835
diff changeset
   147
val dc_copy = %%:(dname^"_copy");
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   148
val x_name = "x";
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   149
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   150
val iso_locale = iso_intro OF [ax_abs_iso, ax_rep_iso];
16224
57094b83774e Domain package generates match functions for new datatypes, for use with the fixrec package
huffman
parents: 15742
diff changeset
   151
val abs_strict = ax_rep_iso RS (allI RS retraction_strict);
57094b83774e Domain package generates match functions for new datatypes, for use with the fixrec package
huffman
parents: 15742
diff changeset
   152
val rep_strict = ax_abs_iso RS (allI RS retraction_strict);
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   153
val abs_defin' = iso_locale RS iso_abs_defin';
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   154
val rep_defin' = iso_locale RS iso_rep_defin';
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   155
val iso_rews = map standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   156
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   157
(* ----- generating beta reduction rules from definitions-------------------- *)
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   158
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   159
local
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   160
  fun arglist (Const _ $ Abs (s, _, t)) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   161
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   162
      val (vars,body) = arglist t;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   163
    in (s :: vars, body) end
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   164
    | arglist t = ([], t);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   165
  fun bind_fun vars t = Library.foldr mk_All (vars, t);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   166
  fun bound_vars 0 = []
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   167
    | bound_vars i = Bound (i-1) :: bound_vars (i - 1);
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   168
in
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   169
  fun appl_of_def def =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   170
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   171
      val (_ $ con $ lam) = concl_of def;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   172
      val (vars, rhs) = arglist lam;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   173
      val lhs = list_ccomb (con, bound_vars (length vars));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   174
      val appl = bind_fun vars (lhs == rhs);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   175
      val cs = ContProc.cont_thms lam;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   176
      val betas = map (fn c => mk_meta_eq (c RS beta_cfun)) cs;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   177
    in pg (def::betas) appl [rtac reflexive_thm 1] end;
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   178
end;
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   179
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   180
val when_appl = appl_of_def ax_when_def;
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   181
val con_appls = map appl_of_def axs_con_def;
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   182
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   183
local
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   184
  fun arg2typ n arg =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   185
    let val t = TVar (("'a", n), pcpoS)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   186
    in (n + 1, if is_lazy arg then mk_uT t else t) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   187
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   188
  fun args2typ n [] = (n, oneT)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   189
    | args2typ n [arg] = arg2typ n arg
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   190
    | args2typ n (arg::args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   191
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   192
      val (n1, t1) = arg2typ n arg;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   193
      val (n2, t2) = args2typ n1 args
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   194
    in (n2, mk_sprodT (t1, t2)) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   195
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   196
  fun cons2typ n [] = (n,oneT)
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   197
    | cons2typ n [con] = args2typ n (snd con)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   198
    | cons2typ n (con::cons) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   199
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   200
      val (n1, t1) = args2typ n (snd con);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   201
      val (n2, t2) = cons2typ n1 cons
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   202
    in (n2, mk_ssumT (t1, t2)) end;
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   203
in
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   204
  fun cons2ctyp cons = ctyp_of (sign_of thy) (snd (cons2typ 1 cons));
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   205
end;
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   206
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   207
local 
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   208
  val iso_swap = iso_locale RS iso_iso_swap;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   209
  fun one_con (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   210
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   211
      val vns = map vname args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   212
      val eqn = %:x_name === con_app2 con %: vns;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   213
      val conj = foldr1 mk_conj (eqn :: map (defined o %:) (nonlazy args));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   214
    in Library.foldr mk_ex (vns, conj) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   215
21619
dea0914773f7 stripped some legacy bindings
haftmann
parents: 20731
diff changeset
   216
  val conj_assoc = thm "conj_assoc";
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   217
  val exh = foldr1 mk_disj ((%:x_name === UU) :: map one_con cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   218
  val thm1 = instantiate' [SOME (cons2ctyp cons)] [] exh_start;
15742
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   219
  val thm2 = rewrite_rule (map mk_meta_eq ex_defined_iffs) thm1;
64eae3513064 speed improvements for the domain package
huffman
parents: 15570
diff changeset
   220
  val thm3 = rewrite_rule [mk_meta_eq conj_assoc] thm2;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   221
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   222
  (* first 3 rules replace "x = UU \/ P" with "rep$x = UU \/ P" *)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   223
  val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   224
    rtac disjE 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   225
    etac (rep_defin' RS disjI1) 2,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   226
    etac disjI2 2,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   227
    rewrite_goals_tac [mk_meta_eq iso_swap],
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   228
    rtac thm3 1];
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   229
in
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   230
  val exhaust = pg con_appls (mk_trp exh) tacs;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   231
  val casedist =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   232
    standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   233
end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   234
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   235
local 
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   236
  fun bind_fun t = Library.foldr mk_All (when_funs cons, t);
1829
5a3687398716 function names in when_rews now meta-quantified
oheimb
parents: 1781
diff changeset
   237
  fun bound_fun i _ = Bound (length cons - i);
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   238
  val when_app = list_ccomb (%%:(dname^"_when"), mapn bound_fun 1 cons);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   239
in
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   240
  val when_strict =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   241
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   242
      val axs = [when_appl, mk_meta_eq rep_strict];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   243
      val goal = bind_fun (mk_trp (strict when_app));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   244
      val tacs = [resolve_tac [sscase1, ssplit1, strictify1] 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   245
    in pg axs goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   246
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   247
  val when_apps =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   248
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   249
      fun one_when n (con,args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   250
        let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   251
          val axs = when_appl :: con_appls;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   252
          val goal = bind_fun (lift_defined %: (nonlazy args, 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   253
                mk_trp (when_app`(con_app con args) ===
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   254
                       list_ccomb (bound_fun n 0, map %# args))));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   255
          val tacs = [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   256
        in pg axs goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   257
    in mapn one_when 1 cons end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   258
end;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   259
val when_rews = when_strict :: when_apps;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   260
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   261
(* ----- theorems concerning the constructors, discriminators and selectors - *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   262
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   263
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   264
  fun dis_strict (con, _) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   265
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   266
      val goal = mk_trp (strict (%%:(dis_name con)));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   267
    in pg axs_dis_def goal [rtac when_strict 1] end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   268
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   269
  fun dis_app c (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   270
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   271
      val lhs = %%:(dis_name c) ` con_app con args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   272
      val rhs = %%:(if con = c then TT_N else FF_N);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   273
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   274
      val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   275
    in pg axs_dis_def goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   276
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   277
  val dis_apps = List.concat (map (fn (c,_) => map (dis_app c) cons) cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   278
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   279
  fun dis_defin (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   280
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   281
      val goal = defined (%:x_name) ==> defined (%%:(dis_name con) `% x_name);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   282
      val tacs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   283
        [rtac casedist 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   284
         contr_tac 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   285
         DETERM_UNTIL_SOLVED (CHANGED
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   286
          (asm_simp_tac (HOLCF_ss addsimps dis_apps) 1))];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   287
    in pg [] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   288
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   289
  val dis_stricts = map dis_strict cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   290
  val dis_defins = map dis_defin cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   291
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   292
  val dis_rews = dis_stricts @ dis_defins @ dis_apps;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   293
end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   294
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   295
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   296
  fun mat_strict (con, _) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   297
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   298
      val goal = mk_trp (strict (%%:(mat_name con)));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   299
      val tacs = [rtac when_strict 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   300
    in pg axs_mat_def goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   301
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   302
  val mat_stricts = map mat_strict cons;
16224
57094b83774e Domain package generates match functions for new datatypes, for use with the fixrec package
huffman
parents: 15742
diff changeset
   303
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   304
  fun one_mat c (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   305
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   306
      val lhs = %%:(mat_name c) ` con_app con args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   307
      val rhs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   308
        if con = c
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   309
        then %%:returnN ` mk_ctuple (map %# args)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   310
        else %%:failN;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   311
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   312
      val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   313
    in pg axs_mat_def goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   314
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   315
  val mat_apps =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   316
    List.concat (map (fn (c,_) => map (one_mat c) cons) cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   317
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   318
  val mat_rews = mat_stricts @ mat_apps;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   319
end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   320
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   321
local
18113
fb76eea85835 generate pattern combinators for new datatypes
huffman
parents: 18084
diff changeset
   322
  fun ps args = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   323
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18113
diff changeset
   324
  fun pat_lhs (con,args) = %%:branchN $ list_comb (%%:(pat_name con), ps args);
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   325
18293
4eaa654c92f2 reimplement Case expression pattern matching to support lazy patterns
huffman
parents: 18113
diff changeset
   326
  fun pat_rhs (con,[]) = %%:returnN ` ((%:"rhs") ` HOLogic.unit)
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   327
    | pat_rhs (con,args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   328
        (%%:branchN $ foldr1 cpair_pat (ps args))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   329
          `(%:"rhs")`(mk_ctuple (map %# args));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   330
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   331
  fun pat_strict c =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   332
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   333
      val axs = branch_def :: axs_pat_def;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   334
      val goal = mk_trp (strict (pat_lhs c ` (%:"rhs")));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   335
      val tacs = [simp_tac (HOLCF_ss addsimps [when_strict]) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   336
    in pg axs goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   337
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   338
  fun pat_app c (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   339
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   340
      val axs = branch_def :: axs_pat_def;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   341
      val lhs = (pat_lhs c)`(%:"rhs")`(con_app con args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   342
      val rhs = if con = fst c then pat_rhs c else %%:failN;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   343
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   344
      val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   345
    in pg axs goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   346
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   347
  val pat_stricts = map pat_strict cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   348
  val pat_apps = List.concat (map (fn c => map (pat_app c) cons) cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   349
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   350
  val pat_rews = pat_stricts @ pat_apps;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   351
end;
18113
fb76eea85835 generate pattern combinators for new datatypes
huffman
parents: 18084
diff changeset
   352
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   353
local
21619
dea0914773f7 stripped some legacy bindings
haftmann
parents: 20731
diff changeset
   354
  val rev_contrapos = thm "rev_contrapos";
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   355
  fun con_strict (con, args) = 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   356
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   357
      fun one_strict vn =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   358
        let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   359
          fun f arg = if vname arg = vn then UU else %# arg;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   360
          val goal = mk_trp (con_app2 con f args === UU);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   361
          val tacs = [asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   362
        in pg con_appls goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   363
    in map one_strict (nonlazy args) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   364
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   365
  fun con_defin (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   366
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   367
      val concl = mk_trp (defined (con_app con args));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   368
      val goal = lift_defined %: (nonlazy args, concl);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   369
      val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   370
        rtac rev_contrapos 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   371
        eres_inst_tac [("f",dis_name con)] cfun_arg_cong 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   372
        asm_simp_tac (HOLCF_ss addsimps dis_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   373
    in pg [] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   374
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   375
  val con_stricts = List.concat (map con_strict cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   376
  val con_defins = map con_defin cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   377
  val con_rews = con_stricts @ con_defins;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   378
end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   379
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   380
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   381
  val rules =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   382
    [compact_sinl, compact_sinr, compact_spair, compact_up, compact_ONE];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   383
  fun con_compact (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   384
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   385
      val concl = mk_trp (%%:compactN $ con_app con args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   386
      val goal = lift (fn x => %%:compactN $ %#x) (args, concl);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   387
      val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   388
        rtac (iso_locale RS iso_compact_abs) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   389
        REPEAT (resolve_tac rules 1 ORELSE atac 1)];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   390
    in pg con_appls goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   391
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   392
  val con_compacts = map con_compact cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   393
end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   394
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   395
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   396
  fun one_sel sel =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   397
    pg axs_sel_def (mk_trp (strict (%%:sel)))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   398
      [simp_tac (HOLCF_ss addsimps when_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   399
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   400
  fun sel_strict (_, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   401
    List.mapPartial (Option.map one_sel o sel_of) args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   402
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   403
  val sel_stricts = List.concat (map sel_strict cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   404
end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   405
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   406
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   407
  fun sel_app_same c n sel (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   408
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   409
      val nlas = nonlazy args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   410
      val vns = map vname args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   411
      val vnn = List.nth (vns, n);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   412
      val nlas' = List.filter (fn v => v <> vnn) nlas;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   413
      val lhs = (%%:sel)`(con_app con args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   414
      val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   415
      val tacs1 =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   416
        if vnn mem nlas
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   417
        then [case_UU_tac (when_rews @ con_stricts) 1 vnn]
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   418
        else [];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   419
      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   420
    in pg axs_sel_def goal (tacs1 @ tacs2) end;
17840
11bcd77cfa22 domain package generates compactness lemmas for new constructors
huffman
parents: 17811
diff changeset
   421
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   422
  fun sel_app_diff c n sel (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   423
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   424
      val nlas = nonlazy args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   425
      val goal = mk_trp (%%:sel ` con_app con args === UU);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   426
      val tacs1 = map (case_UU_tac (when_rews @ con_stricts) 1) nlas;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   427
      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   428
    in pg axs_sel_def goal (tacs1 @ tacs2) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   429
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   430
  fun sel_app c n sel (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   431
    if con = c
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   432
    then sel_app_same c n sel (con, args)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   433
    else sel_app_diff c n sel (con, args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   434
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   435
  fun one_sel c n sel = map (sel_app c n sel) cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   436
  fun one_sel' c n arg = Option.map (one_sel c n) (sel_of arg);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   437
  fun one_con (c, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   438
    List.concat (List.mapPartial I (mapn (one_sel' c) 0 args));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   439
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   440
  val sel_apps = List.concat (map one_con cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   441
end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   442
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   443
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   444
  fun sel_defin sel =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   445
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   446
      val goal = defined (%:x_name) ==> defined (%%:sel`%x_name);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   447
      val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   448
        rtac casedist 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   449
        contr_tac 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   450
        DETERM_UNTIL_SOLVED (CHANGED
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   451
          (asm_simp_tac (HOLCF_ss addsimps sel_apps) 1))];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   452
    in pg [] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   453
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   454
  val sel_defins =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   455
    if length cons = 1
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   456
    then List.mapPartial (fn arg => Option.map sel_defin (sel_of arg))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   457
                 (filter_out is_lazy (snd (hd cons)))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   458
    else [];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   459
end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   460
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   461
val sel_rews = sel_stricts @ sel_defins @ sel_apps;
21619
dea0914773f7 stripped some legacy bindings
haftmann
parents: 20731
diff changeset
   462
val rev_contrapos = thm "rev_contrapos";
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   463
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   464
val distincts_le =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   465
  let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   466
    fun dist (con1, args1) (con2, args2) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   467
      let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   468
        val goal = lift_defined %: (nonlazy args1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   469
                        mk_trp (con_app con1 args1 ~<< con_app con2 args2));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   470
        val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   471
          rtac rev_contrapos 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   472
          eres_inst_tac [("f", dis_name con1)] monofun_cfun_arg 1]
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   473
          @ map (case_UU_tac (con_stricts @ dis_rews) 1) (nonlazy args2)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   474
          @ [asm_simp_tac (HOLCF_ss addsimps dis_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   475
      in pg [] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   476
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   477
    fun distinct (con1, args1) (con2, args2) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   478
        let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   479
          val arg1 = (con1, args1);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   480
          val arg2 =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   481
            (con2, ListPair.map (fn (arg,vn) => upd_vname (K vn) arg)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20046
diff changeset
   482
              (args2, Name.variant_list (map vname args1) (map vname args2)));
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   483
        in [dist arg1 arg2, dist arg2 arg1] end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   484
    fun distincts []      = []
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   485
      | distincts (c::cs) = (map (distinct c) cs) :: distincts cs;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   486
  in distincts cons end;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   487
val dist_les = List.concat (List.concat distincts_le);
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   488
val dist_eqs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   489
  let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   490
    fun distinct (_,args1) ((_,args2), leqs) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   491
      let
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   492
        val (le1,le2) = (hd leqs, hd(tl leqs));
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   493
        val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   494
      in
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   495
        if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   496
        if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   497
          [eq1, eq2]
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   498
      end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   499
    fun distincts []      = []
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   500
      | distincts ((c,leqs)::cs) = List.concat
2267
b2326aefecbc Replaced map...~~ by ListPair.map
paulson
parents: 2033
diff changeset
   501
	            (ListPair.map (distinct c) ((map #1 cs),leqs)) @
b2326aefecbc Replaced map...~~ by ListPair.map
paulson
parents: 2033
diff changeset
   502
		    distincts cs;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   503
  in map standard (distincts (cons ~~ distincts_le)) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   504
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   505
local 
16321
ef32a42f4079 faster proofs of inverts and injects lemmas, with fewer strictness hypotheses
huffman
parents: 16224
diff changeset
   506
  fun pgterm rel con args =
ef32a42f4079 faster proofs of inverts and injects lemmas, with fewer strictness hypotheses
huffman
parents: 16224
diff changeset
   507
    let
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   508
      fun append s = upd_vname (fn v => v^s);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   509
      val (largs, rargs) = (args, map (append "'") args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   510
      val concl =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   511
        foldr1 mk_conj (ListPair.map rel (map %# largs, map %# rargs));
19549
a8ed346f8635 inverts and injects generated lemmas now take the form of rewrite rules, instead of dest rules
huffman
parents: 18972
diff changeset
   512
      val prem = rel (con_app con largs, con_app con rargs);
a8ed346f8635 inverts and injects generated lemmas now take the form of rewrite rules, instead of dest rules
huffman
parents: 18972
diff changeset
   513
      val sargs = case largs of [_] => [] | _ => nonlazy args;
a8ed346f8635 inverts and injects generated lemmas now take the form of rewrite rules, instead of dest rules
huffman
parents: 18972
diff changeset
   514
      val prop = lift_defined %: (sargs, mk_trp (prem === concl));
16321
ef32a42f4079 faster proofs of inverts and injects lemmas, with fewer strictness hypotheses
huffman
parents: 16224
diff changeset
   515
    in pg con_appls prop end;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   516
  val cons' = List.filter (fn (_,args) => args<>[]) cons;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   517
in
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   518
  val inverts =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   519
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   520
      val abs_less = ax_abs_iso RS (allI RS injection_less);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   521
      val tacs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   522
        [asm_full_simp_tac (HOLCF_ss addsimps [abs_less, spair_less]) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   523
    in map (fn (con, args) => pgterm (op <<) con args tacs) cons' end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   524
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   525
  val injects =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   526
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   527
      val abs_eq = ax_abs_iso RS (allI RS injection_eq);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   528
      val tacs = [asm_full_simp_tac (HOLCF_ss addsimps [abs_eq, spair_eq]) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   529
    in map (fn (con, args) => pgterm (op ===) con args tacs) cons' end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   530
end;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   531
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   532
(* ----- theorems concerning one induction step ----------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   533
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   534
val copy_strict =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   535
  let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   536
    val goal = mk_trp (strict (dc_copy `% "f"));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   537
    val tacs = [asm_simp_tac (HOLCF_ss addsimps [abs_strict, when_strict]) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   538
  in pg [ax_copy_def] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   539
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   540
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   541
  fun copy_app (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   542
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   543
      val lhs = dc_copy`%"f"`(con_app con args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   544
      val rhs = con_app2 con (app_rec_arg (cproj (%:"f") eqs)) args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   545
      val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   546
      val args' = List.filter (fn a => not (is_rec a orelse is_lazy a)) args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   547
      val stricts = abs_strict::when_strict::con_stricts;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   548
      val tacs1 = map (case_UU_tac stricts 1 o vname) args';
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   549
      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_apps) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   550
    in pg [ax_copy_def] goal (tacs1 @ tacs2) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   551
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   552
  val copy_apps = map copy_app cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   553
end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   554
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   555
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   556
  fun one_strict (con, args) = 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   557
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   558
      val goal = mk_trp (dc_copy`UU`(con_app con args) === UU);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   559
      val rews = copy_strict :: copy_apps @ con_rews;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   560
      val tacs = map (case_UU_tac rews 1) (nonlazy args) @
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   561
        [asm_simp_tac (HOLCF_ss addsimps rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   562
    in pg [] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   563
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   564
  fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   565
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   566
  val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   567
end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   568
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   569
val copy_rews = copy_strict :: copy_apps @ copy_stricts;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   570
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   571
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   572
  thy
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   573
    |> Theory.add_path (Sign.base_name dname)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   574
    |> (snd o (PureThy.add_thmss (map Thm.no_attributes [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   575
        ("iso_rews" , iso_rews  ),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   576
        ("exhaust"  , [exhaust] ),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   577
        ("casedist" , [casedist]),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   578
        ("when_rews", when_rews ),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   579
        ("compacts", con_compacts),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   580
        ("con_rews", con_rews),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   581
        ("sel_rews", sel_rews),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   582
        ("dis_rews", dis_rews),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   583
        ("pat_rews", pat_rews),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   584
        ("dist_les", dist_les),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   585
        ("dist_eqs", dist_eqs),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   586
        ("inverts" , inverts ),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   587
        ("injects" , injects ),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   588
        ("copy_rews", copy_rews)])))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   589
    |> (snd o PureThy.add_thmss
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   590
        [(("match_rews", mat_rews), [Simplifier.simp_add])])
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   591
    |> Theory.parent_path
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   592
    |> rpair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   593
        pat_rews @ dist_les @ dist_eqs @ copy_rews)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   594
end; (* let *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   595
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   596
fun comp_theorems (comp_dnam, eqs: eq list) thy =
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   597
let
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3323
diff changeset
   598
val dnames = map (fst o fst) eqs;
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3323
diff changeset
   599
val conss  = map  snd        eqs;
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3323
diff changeset
   600
val comp_dname = Sign.full_name (sign_of thy) comp_dnam;
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3323
diff changeset
   601
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   602
val d = writeln("Proving induction properties of domain "^comp_dname^" ...");
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   603
val pg = pg' thy;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   604
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   605
(* ----- getting the composite axiom and definitions ------------------------ *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   606
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   607
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   608
  fun ga s dn = get_thm thy (Name (dn ^ "." ^ s));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   609
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   610
  val axs_reach      = map (ga "reach"     ) dnames;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   611
  val axs_take_def   = map (ga "take_def"  ) dnames;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   612
  val axs_finite_def = map (ga "finite_def") dnames;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   613
  val ax_copy2_def   =      ga "copy_def"  comp_dnam;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   614
  val ax_bisim_def   =      ga "bisim_def" comp_dnam;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   615
end;
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   616
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   617
local
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   618
  fun gt  s dn = get_thm  thy (Name (dn ^ "." ^ s));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   619
  fun gts s dn = get_thms thy (Name (dn ^ "." ^ s));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   620
in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   621
  val cases = map (gt  "casedist" ) dnames;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   622
  val con_rews  = List.concat (map (gts "con_rews" ) dnames);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   623
  val copy_rews = List.concat (map (gts "copy_rews") dnames);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   624
end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   625
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 10835
diff changeset
   626
fun dc_take dn = %%:(dn^"_take");
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   627
val x_name = idx_name dnames "x"; 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   628
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
   629
val n_eqs = length eqs;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   630
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   631
(* ----- theorems concerning finite approximation and finite induction ------ *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   632
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   633
local
19764
372065f34795 removed obsolete ML files;
wenzelm
parents: 19561
diff changeset
   634
  val iterate_Cprod_ss = simpset_of (theory "Fix");
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   635
  val copy_con_rews  = copy_rews @ con_rews;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   636
  val copy_take_defs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   637
    (if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   638
  val take_stricts =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   639
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   640
      fun one_eq ((dn, args), _) = strict (dc_take dn $ %:"n");
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   641
      val goal = mk_trp (foldr1 mk_conj (map one_eq eqs));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   642
      val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   643
        induct_tac "n" 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   644
        simp_tac iterate_Cprod_ss 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   645
        asm_simp_tac (iterate_Cprod_ss addsimps copy_rews) 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   646
    in pg copy_take_defs goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   647
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   648
  val take_stricts' = rewrite_rule copy_take_defs take_stricts;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   649
  fun take_0 n dn =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   650
    let
20731
2ef8b7332b4f replaced constant 0 by HOL.zero
haftmann
parents: 20155
diff changeset
   651
      val goal = mk_trp ((dc_take dn $ %%:"HOL.zero") `% x_name n === UU);
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   652
    in pg axs_take_def goal [simp_tac iterate_Cprod_ss 1] end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   653
  val take_0s = mapn take_0 1 dnames;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   654
  val c_UU_tac = case_UU_tac (take_stricts'::copy_con_rews) 1;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   655
  val take_apps =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   656
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   657
      fun mk_eqn dn (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   658
        let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   659
          fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n";
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   660
          val lhs = (dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   661
          val rhs = con_app2 con (app_rec_arg mk_take) args;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   662
        in Library.foldr mk_all (map vname args, lhs === rhs) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   663
      fun mk_eqns ((dn, _), cons) = map (mk_eqn dn) cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   664
      val goal = mk_trp (foldr1 mk_conj (List.concat (map mk_eqns eqs)));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   665
      val simps = List.filter (has_fewer_prems 1) copy_rews;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   666
      fun con_tac (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   667
        if nonlazy_rec args = []
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   668
        then all_tac
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   669
        else EVERY (map c_UU_tac (nonlazy_rec args)) THEN
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   670
          asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   671
      fun eq_tacs ((dn, _), cons) = map con_tac cons;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   672
      val tacs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   673
        simp_tac iterate_Cprod_ss 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   674
        induct_tac "n" 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   675
        simp_tac (iterate_Cprod_ss addsimps copy_con_rews) 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   676
        asm_full_simp_tac (HOLCF_ss addsimps simps) 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   677
        TRY (safe_tac HOL_cs) ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   678
        List.concat (map eq_tacs eqs);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   679
    in pg copy_take_defs goal tacs end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   680
in
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   681
  val take_rews = map standard
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   682
    (atomize take_stricts @ take_0s @ atomize take_apps);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   683
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   684
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   685
local
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   686
  fun one_con p (con,args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   687
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   688
      fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   689
      val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   690
      val t2 = lift ind_hyp (List.filter is_rec args, t1);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   691
      val t3 = lift_defined (bound_arg (map vname args)) (nonlazy args, t2);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   692
    in Library.foldr mk_All (map vname args, t3) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   693
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   694
  fun one_eq ((p, cons), concl) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   695
    mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   696
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   697
  fun ind_term concf = Library.foldr one_eq
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   698
    (mapn (fn n => fn x => (P_name n, x)) 1 conss,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   699
     mk_trp (foldr1 mk_conj (mapn concf 1 dnames)));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   700
  val take_ss = HOL_ss addsimps take_rews;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   701
  fun quant_tac i = EVERY
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   702
    (mapn (fn n => fn _ => res_inst_tac [("x", x_name n)] spec i) 1 dnames);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   703
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   704
  fun ind_prems_tac prems = EVERY
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   705
    (List.concat (map (fn cons =>
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   706
      (resolve_tac prems 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   707
        List.concat (map (fn (_,args) => 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   708
          resolve_tac prems 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   709
          map (K(atac 1)) (nonlazy args) @
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   710
          map (K(atac 1)) (List.filter is_rec args))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   711
        cons)))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   712
      conss));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   713
  local 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   714
    (* 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
   715
       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
   716
       indirectly lazy *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   717
    fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => 
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   718
          is_rec arg andalso not(rec_of arg mem ns) andalso
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   719
          ((rec_of arg =  n andalso nfn(lazy_rec orelse is_lazy arg)) orelse 
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   720
            rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   721
              (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg))))
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1834
diff changeset
   722
          ) o snd) cons;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   723
    fun all_rec_to ns  = rec_to forall not all_rec_to  ns;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   724
    fun warn (n,cons) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   725
      if all_rec_to [] false (n,cons)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   726
      then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   727
      else false;
16842
wenzelm
parents: 16778
diff changeset
   728
    fun lazy_rec_to ns = rec_to exists I  lazy_rec_to ns;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   729
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   730
  in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   731
    val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   732
    val is_emptys = map warn n__eqs;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   733
    val is_finite = forall (not o lazy_rec_to [] false) n__eqs;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   734
  end;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   735
in (* local *)
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   736
  val finite_ind =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   737
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   738
      fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   739
      val goal = ind_term concf;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   740
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   741
      fun tacf prems =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   742
        let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   743
          val tacs1 = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   744
            quant_tac 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   745
            simp_tac HOL_ss 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   746
            induct_tac "n" 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   747
            simp_tac (take_ss addsimps prems) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   748
            TRY (safe_tac HOL_cs)];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   749
          fun arg_tac arg =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   750
            case_UU_tac (prems @ con_rews) 1
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   751
              (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   752
          fun con_tacs (con, args) = 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   753
            asm_simp_tac take_ss 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   754
            map arg_tac (List.filter is_nonlazy_rec args) @
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   755
            [resolve_tac prems 1] @
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   756
            map (K (atac 1))      (nonlazy args) @
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   757
            map (K (etac spec 1)) (List.filter is_rec args);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   758
          fun cases_tacs (cons, cases) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   759
            res_inst_tac [("x","x")] cases 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   760
            asm_simp_tac (take_ss addsimps prems) 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   761
            List.concat (map con_tacs cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   762
        in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   763
          tacs1 @ List.concat (map cases_tacs (conss ~~ cases))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   764
        end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   765
    in pg'' thy [] goal tacf end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   766
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   767
  val take_lemmas =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   768
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   769
      fun take_lemma n (dn, ax_reach) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   770
        let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   771
          val lhs = dc_take dn $ Bound 0 `%(x_name n);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   772
          val rhs = dc_take dn $ Bound 0 `%(x_name n^"'");
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   773
          val concl = mk_trp (%:(x_name n) === %:(x_name n^"'"));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   774
          val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   775
          fun tacf prems = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   776
            res_inst_tac [("t", x_name n    )] (ax_reach RS subst) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   777
            res_inst_tac [("t", x_name n^"'")] (ax_reach RS subst) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   778
            stac fix_def2 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   779
            REPEAT (CHANGED
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   780
              (rtac (contlub_cfun_arg RS ssubst) 1 THEN chain_tac 1)),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   781
            stac contlub_cfun_fun 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   782
            stac contlub_cfun_fun 2,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   783
            rtac lub_equal 3,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   784
            chain_tac 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   785
            rtac allI 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   786
            resolve_tac prems 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   787
        in pg'' thy axs_take_def goal tacf end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   788
    in mapn take_lemma 1 (dnames ~~ axs_reach) end;
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   789
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   790
(* ----- theorems concerning finiteness and induction ----------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   791
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   792
  val (finites, ind) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   793
    if is_finite
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   794
    then (* finite case *)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   795
      let 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   796
        fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %:"x" === %:"x");
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   797
        fun dname_lemma dn =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   798
          let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   799
            val prem1 = mk_trp (defined (%:"x"));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   800
            val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   801
            val prem2 = mk_trp (mk_disj (disj1, take_enough dn));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   802
            val concl = mk_trp (take_enough dn);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   803
            val goal = prem1 ===> prem2 ===> concl;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   804
            val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   805
              etac disjE 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   806
              etac notE 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   807
              resolve_tac take_lemmas 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   808
              asm_simp_tac take_ss 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   809
              atac 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   810
          in pg [] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   811
        val finite_lemmas1a = map dname_lemma dnames;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   812
 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   813
        val finite_lemma1b =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   814
          let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   815
            fun mk_eqn n ((dn, args), _) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   816
              let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   817
                val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   818
                val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   819
              in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   820
                mk_constrainall
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   821
                  (x_name n, Type (dn,args), mk_disj (disj1, disj2))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   822
              end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   823
            val goal =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   824
              mk_trp (mk_all ("n", foldr1 mk_conj (mapn mk_eqn 1 eqs)));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   825
            fun arg_tacs vn = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   826
              eres_inst_tac [("x", vn)] all_dupE 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   827
              etac disjE 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   828
              asm_simp_tac (HOL_ss addsimps con_rews) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   829
              asm_simp_tac take_ss 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   830
            fun con_tacs (con, args) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   831
              asm_simp_tac take_ss 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   832
              List.concat (map arg_tacs (nonlazy_rec args));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   833
            fun foo_tacs n (cons, cases) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   834
              simp_tac take_ss 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   835
              rtac allI 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   836
              res_inst_tac [("x",x_name n)] cases 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   837
              asm_simp_tac take_ss 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   838
              List.concat (map con_tacs cons);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   839
            val tacs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   840
              rtac allI 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   841
              induct_tac "n" 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   842
              simp_tac take_ss 1 ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   843
              TRY (safe_tac (empty_cs addSEs [conjE] addSIs [conjI])) ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   844
              List.concat (mapn foo_tacs 1 (conss ~~ cases));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   845
          in pg [] goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   846
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   847
        fun one_finite (dn, l1b) =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   848
          let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   849
            val goal = mk_trp (%%:(dn^"_finite") $ %:"x");
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   850
            val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   851
              case_UU_tac take_rews 1 "x",
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   852
              eresolve_tac finite_lemmas1a 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   853
              step_tac HOL_cs 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   854
              step_tac HOL_cs 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   855
              cut_facts_tac [l1b] 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   856
              fast_tac HOL_cs 1];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   857
          in pg axs_finite_def goal tacs end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   858
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   859
        val finites = map one_finite (dnames ~~ atomize finite_lemma1b);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   860
        val ind =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   861
          let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   862
            fun concf n dn = %:(P_name n) $ %:(x_name n);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   863
            fun tacf prems =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   864
              let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   865
                fun finite_tacs (finite, fin_ind) = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   866
                  rtac(rewrite_rule axs_finite_def finite RS exE)1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   867
                  etac subst 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   868
                  rtac fin_ind 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   869
                  ind_prems_tac prems];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   870
              in
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   871
                TRY (safe_tac HOL_cs) ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   872
                List.concat (map finite_tacs (finites ~~ atomize finite_ind))
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   873
              end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   874
          in pg'' thy [] (ind_term concf) tacf end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   875
      in (finites, ind) end (* let *)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   876
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   877
    else (* infinite case *)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   878
      let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   879
        fun one_finite n dn =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   880
          read_instantiate_sg (sign_of thy)
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   881
            [("P",dn^"_finite "^x_name n)] excluded_middle;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   882
        val finites = mapn one_finite 1 dnames;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   883
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   884
        val goal =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   885
          let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   886
            fun one_adm n _ = mk_trp (%%:admN $ %:(P_name n));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   887
            fun concf n dn = %:(P_name n) $ %:(x_name n);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   888
          in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   889
        fun tacf prems =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   890
          map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   891
          quant_tac 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   892
          rtac (adm_impl_admw RS wfix_ind) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   893
          REPEAT_DETERM (rtac adm_all2 1),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   894
          REPEAT_DETERM (
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   895
            TRY (rtac adm_conj 1) THEN 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   896
            rtac adm_subst 1 THEN 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   897
            cont_tacR 1 THEN resolve_tac prems 1),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   898
          strip_tac 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   899
          rtac (rewrite_rule axs_take_def finite_ind) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   900
          ind_prems_tac prems];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   901
        val ind = (pg'' thy [] goal tacf
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   902
          handle ERROR _ =>
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   903
            (warning "Cannot prove infinite induction rule"; refl));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   904
      in (finites, ind) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   905
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   906
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   907
(* ----- theorem concerning coinduction ------------------------------------- *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1512
diff changeset
   908
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   909
local
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   910
  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
   911
  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   912
  val take_ss = HOL_ss addsimps take_rews;
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   913
  val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   914
  val coind_lemma =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   915
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   916
      fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1;
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   917
      fun mk_eqn n dn =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   918
        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   919
        (dc_take dn $ %:"n" ` bnd_arg n 1);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   920
      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   921
      val goal =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   922
        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   923
          Library.foldr mk_all2 (xs,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   924
            Library.foldr mk_imp (mapn mk_prj 0 dnames,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   925
              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   926
      fun x_tacs n x = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   927
        rotate_tac (n+1) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   928
        etac all2E 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   929
        eres_inst_tac [("P1", sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   930
        TRY (safe_tac HOL_cs),
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   931
        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   932
      val tacs = [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   933
        rtac impI 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   934
        induct_tac "n" 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   935
        simp_tac take_ss 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   936
        safe_tac HOL_cs] @
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   937
        List.concat (mapn x_tacs 0 xs);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   938
    in pg [ax_bisim_def] goal tacs end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   939
in
19561
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   940
  val coind = 
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   941
    let
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   942
      fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   943
      fun mk_eqn x = %:x === %:(x^"'");
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   944
      val goal =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   945
        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   946
          Logic.list_implies (mapn mk_prj 0 xs,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   947
            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   948
      val tacs =
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   949
        TRY (safe_tac HOL_cs) ::
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   950
        List.concat (map (fn take_lemma => [
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   951
          rtac take_lemma 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   952
          cut_facts_tac [coind_lemma] 1,
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   953
          fast_tac HOL_cs 1])
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   954
        take_lemmas);
2a4983dc963d extensive reformatting and cleaning up code (no changes in functionality)
huffman
parents: 19549
diff changeset
   955
    in pg [] goal tacs end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   956
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   957
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   958
in thy |> Theory.add_path comp_dnam
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18293
diff changeset
   959
       |> (snd o (PureThy.add_thmss (map Thm.no_attributes [
4043
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   960
		("take_rews"  , take_rews  ),
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   961
		("take_lemmas", take_lemmas),
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   962
		("finites"    , finites    ),
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   963
		("finite_ind", [finite_ind]),
35766855f344 domain package:
oheimb
parents: 4030
diff changeset
   964
		("ind"       , [ind       ]),
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8149
diff changeset
   965
		("coind"     , [coind     ])])))
12037
0282eacef4e7 adapted to new-style theories;
wenzelm
parents: 12030
diff changeset
   966
       |> Theory.parent_path |> rpair take_rews
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   967
end; (* let *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   968
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   969
end; (* struct *)