src/HOLCF/Tools/domain/domain_theorems.ML
author huffman
Fri, 08 May 2009 16:19:51 -0700
changeset 31076 99fe356cbbc2
parent 31023 d027411c9a38
child 31160 2823f1b6b860
permissions -rw-r--r--
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/domain/domain_theorems.ML
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
    Author:     David von Oheimb
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
                New proofs/tactics by Brian Huffman
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
Proof generator for domain command.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 26336
diff changeset
     9
val HOLCF_ss = @{simpset};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    10
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    11
signature DOMAIN_THEOREMS =
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    12
sig
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    13
  val theorems: Domain_Library.eq * Domain_Library.eq list -> theory -> thm list * theory;
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    14
  val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory;
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    15
end;
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    16
31023
d027411c9a38 use opaque ascription for all HOLCF code
huffman
parents: 31005
diff changeset
    17
structure Domain_Theorems :> DOMAIN_THEOREMS =
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    18
struct
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    20
val quiet_mode = ref false;
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    21
val trace_domain = ref false;
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    22
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    23
fun message s = if !quiet_mode then () else writeln s;
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    24
fun trace s = if !trace_domain then tracing s else ();
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    25
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    26
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    28
val adm_impl_admw = @{thm adm_impl_admw};
25895
0eaadfa8889e converted adm_all and adm_ball to rule_format; cleaned up
huffman
parents: 25805
diff changeset
    29
val adm_all = @{thm adm_all};
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    30
val adm_conj = @{thm adm_conj};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    31
val adm_subst = @{thm adm_subst};
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31023
diff changeset
    32
val antisym_less_inverse = @{thm below_antisym_inverse};
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    33
val beta_cfun = @{thm beta_cfun};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    34
val cfun_arg_cong = @{thm cfun_arg_cong};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    35
val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    36
val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    37
val chain_iterate = @{thm chain_iterate};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    38
val compact_ONE = @{thm compact_ONE};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    39
val compact_sinl = @{thm compact_sinl};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    40
val compact_sinr = @{thm compact_sinr};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    41
val compact_spair = @{thm compact_spair};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    42
val compact_up = @{thm compact_up};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    43
val contlub_cfun_arg = @{thm contlub_cfun_arg};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    44
val contlub_cfun_fun = @{thm contlub_cfun_fun};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    45
val fix_def2 = @{thm fix_def2};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    46
val injection_eq = @{thm injection_eq};
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31023
diff changeset
    47
val injection_less = @{thm injection_below};
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    48
val lub_equal = @{thm lub_equal};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    49
val monofun_cfun_arg = @{thm monofun_cfun_arg};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    50
val retraction_strict = @{thm retraction_strict};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    51
val spair_eq = @{thm spair_eq};
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31023
diff changeset
    52
val spair_less = @{thm spair_below};
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    53
val sscase1 = @{thm sscase1};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    54
val ssplit1 = @{thm ssplit1};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    55
val strictify1 = @{thm strictify1};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    56
val wfix_ind = @{thm wfix_ind};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    57
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    58
val iso_intro       = @{thm iso.intro};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    59
val iso_abs_iso     = @{thm iso.abs_iso};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    60
val iso_rep_iso     = @{thm iso.rep_iso};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    61
val iso_abs_strict  = @{thm iso.abs_strict};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    62
val iso_rep_strict  = @{thm iso.rep_strict};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    63
val iso_abs_defin'  = @{thm iso.abs_defin'};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    64
val iso_rep_defin'  = @{thm iso.rep_defin'};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    65
val iso_abs_defined = @{thm iso.abs_defined};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    66
val iso_rep_defined = @{thm iso.rep_defined};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    67
val iso_compact_abs = @{thm iso.compact_abs};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    68
val iso_compact_rep = @{thm iso.compact_rep};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    69
val iso_iso_swap    = @{thm iso.iso_swap};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    70
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    71
val exh_start = @{thm exh_start};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    72
val ex_defined_iffs = @{thms ex_defined_iffs};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    73
val exh_casedist0 = @{thm exh_casedist0};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    74
val exh_casedists = @{thms exh_casedists};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    76
open Domain_Library;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
infixr 0 ===>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    78
infixr 0 ==>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    79
infix 0 == ; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
infix 1 ===;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    81
infix 1 ~= ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
infix 1 <<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
infix 1 ~<<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    84
infix 9 `   ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
infix 9 `% ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
infix 9 `%%;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
infixr 9 oo;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
(* ----- general proof facilities ------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    91
fun legacy_infer_term thy t =
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    92
  let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy)
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    93
  in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end;
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    94
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    95
fun pg'' thy defs t tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
  let
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    97
    val t' = legacy_infer_term thy t;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    98
    val asms = Logic.strip_imp_prems t';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
    val prop = Logic.strip_imp_concl t';
26711
3a478bfa1650 prove_global: pass context;
wenzelm
parents: 26343
diff changeset
   100
    fun tac {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   101
      rewrite_goals_tac defs THEN
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   102
      EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   103
  in Goal.prove_global thy [] asms prop tac end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   104
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   105
fun pg' thy defs t tacsf =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   106
  let
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   107
    fun tacs {prems, context} =
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   108
      if null prems then tacsf context
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   109
      else cut_facts_tac prems 1 :: tacsf context;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   110
  in pg'' thy defs t tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   112
fun case_UU_tac ctxt rews i v =
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   113
  InductTacs.case_tac ctxt (v^"=UU") i THEN
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   114
  asm_simp_tac (HOLCF_ss addsimps rews) i;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   115
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   116
val chain_tac =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   117
  REPEAT_DETERM o resolve_tac 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
    [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   119
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   120
(* ----- general proofs ----------------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   121
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
   122
val all2E = @{lemma "!x y . P x y ==> (P x y ==> R) ==> R" by simp}
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   123
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31023
diff changeset
   124
val dist_eqI = @{lemma "!!x::'a::po. ~ x << y ==> x ~= y" by (blast dest!: below_antisym_inverse)}
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   125
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   126
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   127
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
fun theorems (((dname, _), cons) : eq, eqs : eq list) thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   129
let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   130
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   131
val _ = message ("Proving isomorphism properties of domain "^dname^" ...");
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   132
val pg = pg' thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   133
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   134
(* ----- getting the axioms and definitions --------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   135
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   136
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   137
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   138
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   139
  val ax_abs_iso  = ga "abs_iso"  dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   140
  val ax_rep_iso  = ga "rep_iso"  dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   141
  val ax_when_def = ga "when_def" dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   142
  fun get_def mk_name (con,_) = ga (mk_name con^"_def") dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   143
  val axs_con_def = map (get_def extern_name) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   144
  val axs_dis_def = map (get_def dis_name) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   145
  val axs_mat_def = map (get_def mat_name) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   146
  val axs_pat_def = map (get_def pat_name) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   147
  val axs_sel_def =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   148
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   149
      fun def_of_sel sel = ga (sel^"_def") dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   150
      fun def_of_arg arg = Option.map def_of_sel (sel_of arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   151
      fun defs_of_con (_, args) = List.mapPartial def_of_arg args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   152
    in
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   153
      maps defs_of_con cons
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   154
    end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   155
  val ax_copy_def = ga "copy_def" dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   156
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   157
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   158
(* ----- theorems concerning the isomorphism -------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   159
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   160
val dc_abs  = %%:(dname^"_abs");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   161
val dc_rep  = %%:(dname^"_rep");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   162
val dc_copy = %%:(dname^"_copy");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   163
val x_name = "x";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   164
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   165
val iso_locale = iso_intro OF [ax_abs_iso, ax_rep_iso];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   166
val abs_strict = ax_rep_iso RS (allI RS retraction_strict);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   167
val rep_strict = ax_abs_iso RS (allI RS retraction_strict);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   168
val abs_defin' = iso_locale RS iso_abs_defin';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   169
val rep_defin' = iso_locale RS iso_rep_defin';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   170
val iso_rews = map standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   171
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   172
(* ----- generating beta reduction rules from definitions-------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   173
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   174
val _ = trace " Proving beta reduction rules...";
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   175
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   176
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   177
  fun arglist (Const _ $ Abs (s, _, t)) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   178
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   179
      val (vars,body) = arglist t;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   180
    in (s :: vars, body) end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   181
    | arglist t = ([], t);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   182
  fun bind_fun vars t = Library.foldr mk_All (vars, t);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   183
  fun bound_vars 0 = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   184
    | bound_vars i = Bound (i-1) :: bound_vars (i - 1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   185
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   186
  fun appl_of_def def =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   187
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   188
      val (_ $ con $ lam) = concl_of def;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   189
      val (vars, rhs) = arglist lam;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   190
      val lhs = list_ccomb (con, bound_vars (length vars));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   191
      val appl = bind_fun vars (lhs == rhs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   192
      val cs = ContProc.cont_thms lam;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   193
      val betas = map (fn c => mk_meta_eq (c RS beta_cfun)) cs;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   194
    in pg (def::betas) appl (K [rtac reflexive_thm 1]) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   195
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   196
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   197
val _ = trace "Proving when_appl...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   198
val when_appl = appl_of_def ax_when_def;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   199
val _ = trace "Proving con_appls...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   200
val con_appls = map appl_of_def axs_con_def;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   201
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   202
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   203
  fun arg2typ n arg =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   204
    let val t = TVar (("'a", n), pcpoS)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   205
    in (n + 1, if is_lazy arg then mk_uT t else t) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   206
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   207
  fun args2typ n [] = (n, oneT)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   208
    | args2typ n [arg] = arg2typ n arg
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   209
    | args2typ n (arg::args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   210
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   211
      val (n1, t1) = arg2typ n arg;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   212
      val (n2, t2) = args2typ n1 args
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   213
    in (n2, mk_sprodT (t1, t2)) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   214
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   215
  fun cons2typ n [] = (n,oneT)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   216
    | cons2typ n [con] = args2typ n (snd con)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   217
    | cons2typ n (con::cons) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   218
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   219
      val (n1, t1) = args2typ n (snd con);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   220
      val (n2, t2) = cons2typ n1 cons
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   221
    in (n2, mk_ssumT (t1, t2)) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   222
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   223
  fun cons2ctyp cons = ctyp_of thy (snd (cons2typ 1 cons));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   224
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   225
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   226
local 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   227
  val iso_swap = iso_locale RS iso_iso_swap;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   228
  fun one_con (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   229
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   230
      val vns = map vname args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   231
      val eqn = %:x_name === con_app2 con %: vns;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   232
      val conj = foldr1 mk_conj (eqn :: map (defined o %:) (nonlazy args));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   233
    in Library.foldr mk_ex (vns, conj) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   234
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23152
diff changeset
   235
  val conj_assoc = @{thm conj_assoc};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   236
  val exh = foldr1 mk_disj ((%:x_name === UU) :: map one_con cons);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   237
  val thm1 = instantiate' [SOME (cons2ctyp cons)] [] exh_start;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   238
  val thm2 = rewrite_rule (map mk_meta_eq ex_defined_iffs) thm1;
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
   239
  val thm3 = rewrite_rule [mk_meta_eq @{thm conj_assoc}] thm2;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   240
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   241
  (* first 3 rules replace "x = UU \/ P" with "rep$x = UU \/ P" *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   242
  val tacs = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   243
    rtac disjE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   244
    etac (rep_defin' RS disjI1) 2,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   245
    etac disjI2 2,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   246
    rewrite_goals_tac [mk_meta_eq iso_swap],
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   247
    rtac thm3 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   248
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   249
  val _ = trace " Proving exhaust...";
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   250
  val exhaust = pg con_appls (mk_trp exh) (K tacs);
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   251
  val _ = trace " Proving casedist...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   252
  val casedist =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   253
    standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   254
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   255
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   256
local 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   257
  fun bind_fun t = Library.foldr mk_All (when_funs cons, t);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   258
  fun bound_fun i _ = Bound (length cons - i);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   259
  val when_app = list_ccomb (%%:(dname^"_when"), mapn bound_fun 1 cons);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   260
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   261
  val _ = trace " Proving when_strict...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   262
  val when_strict =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   263
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   264
      val axs = [when_appl, mk_meta_eq rep_strict];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   265
      val goal = bind_fun (mk_trp (strict when_app));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   266
      val tacs = [resolve_tac [sscase1, ssplit1, strictify1] 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   267
    in pg axs goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   268
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   269
  val _ = trace " Proving when_apps...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   270
  val when_apps =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   271
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   272
      fun one_when n (con,args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   273
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   274
          val axs = when_appl :: con_appls;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   275
          val goal = bind_fun (lift_defined %: (nonlazy args, 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   276
                mk_trp (when_app`(con_app con args) ===
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   277
                       list_ccomb (bound_fun n 0, map %# args))));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   278
          val tacs = [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   279
        in pg axs goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   280
    in mapn one_when 1 cons end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   281
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   282
val when_rews = when_strict :: when_apps;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   283
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   284
(* ----- theorems concerning the constructors, discriminators and selectors - *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   285
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   286
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   287
  fun dis_strict (con, _) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   288
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   289
      val goal = mk_trp (strict (%%:(dis_name con)));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   290
    in pg axs_dis_def goal (K [rtac when_strict 1]) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   291
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   292
  fun dis_app c (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   293
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   294
      val lhs = %%:(dis_name c) ` con_app con args;
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   295
      val rhs = if con = c then TT else FF;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   296
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   297
      val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   298
    in pg axs_dis_def goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   299
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   300
  val _ = trace " Proving dis_apps...";
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   301
  val dis_apps = maps (fn (c,_) => map (dis_app c) cons) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   302
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   303
  fun dis_defin (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   304
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   305
      val goal = defined (%:x_name) ==> defined (%%:(dis_name con) `% x_name);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   306
      val tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   307
        [rtac casedist 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   308
         contr_tac 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   309
         DETERM_UNTIL_SOLVED (CHANGED
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   310
          (asm_simp_tac (HOLCF_ss addsimps dis_apps) 1))];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   311
    in pg [] goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   312
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   313
  val _ = trace " Proving dis_stricts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   314
  val dis_stricts = map dis_strict cons;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   315
  val _ = trace " Proving dis_defins...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   316
  val dis_defins = map dis_defin cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   317
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   318
  val dis_rews = dis_stricts @ dis_defins @ dis_apps;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   319
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   320
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   321
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   322
  fun mat_strict (con, _) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   323
    let
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   324
      val goal = mk_trp (%%:(mat_name con) ` UU ` %:"rhs" === UU);
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   325
      val tacs = [asm_simp_tac (HOLCF_ss addsimps [when_strict]) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   326
    in pg axs_mat_def goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   327
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   328
  val _ = trace " Proving mat_stricts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   329
  val mat_stricts = map mat_strict cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   330
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   331
  fun one_mat c (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   332
    let
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   333
      val lhs = %%:(mat_name c) ` con_app con args ` %:"rhs";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   334
      val rhs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   335
        if con = c
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   336
        then list_ccomb (%:"rhs", map %# args)
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   337
        else mk_fail;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   338
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   339
      val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   340
    in pg axs_mat_def goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   341
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   342
  val _ = trace " Proving mat_apps...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   343
  val mat_apps =
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   344
    maps (fn (c,_) => map (one_mat c) cons) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   345
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   346
  val mat_rews = mat_stricts @ mat_apps;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   347
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   348
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   349
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   350
  fun ps args = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   351
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   352
  fun pat_lhs (con,args) = mk_branch (list_comb (%%:(pat_name con), ps args));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   353
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   354
  fun pat_rhs (con,[]) = mk_return ((%:"rhs") ` HOLogic.unit)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   355
    | pat_rhs (con,args) =
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   356
        (mk_branch (mk_ctuple_pat (ps args)))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   357
          `(%:"rhs")`(mk_ctuple (map %# args));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   358
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   359
  fun pat_strict c =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   360
    let
25132
dffe405b090d removed obsolete ML bindings;
wenzelm
parents: 24712
diff changeset
   361
      val axs = @{thm branch_def} :: axs_pat_def;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   362
      val goal = mk_trp (strict (pat_lhs c ` (%:"rhs")));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   363
      val tacs = [simp_tac (HOLCF_ss addsimps [when_strict]) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   364
    in pg axs goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   365
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   366
  fun pat_app c (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   367
    let
25132
dffe405b090d removed obsolete ML bindings;
wenzelm
parents: 24712
diff changeset
   368
      val axs = @{thm branch_def} :: axs_pat_def;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   369
      val lhs = (pat_lhs c)`(%:"rhs")`(con_app con args);
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   370
      val rhs = if con = fst c then pat_rhs c else mk_fail;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   371
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   372
      val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   373
    in pg axs goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   374
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   375
  val _ = trace " Proving pat_stricts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   376
  val pat_stricts = map pat_strict cons;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   377
  val _ = trace " Proving pat_apps...";
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   378
  val pat_apps = maps (fn c => map (pat_app c) cons) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   379
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   380
  val pat_rews = pat_stricts @ pat_apps;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   381
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   382
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   383
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   384
  fun con_strict (con, args) = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   385
    let
30911
7809cbaa1b61 domain package: simplify internal proofs of con_rews
huffman
parents: 30829
diff changeset
   386
      val rules = abs_strict :: @{thms con_strict_rules};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   387
      fun one_strict vn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   388
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   389
          fun f arg = if vname arg = vn then UU else %# arg;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   390
          val goal = mk_trp (con_app2 con f args === UU);
30911
7809cbaa1b61 domain package: simplify internal proofs of con_rews
huffman
parents: 30829
diff changeset
   391
          val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   392
        in pg con_appls goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   393
    in map one_strict (nonlazy args) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   394
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   395
  fun con_defin (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   396
    let
30913
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   397
      fun iff_disj (t, []) = HOLogic.mk_not t
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   398
        | iff_disj (t, ts) = t === foldr1 HOLogic.mk_disj ts;
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   399
      val lhs = con_app con args === UU;
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   400
      val rhss = map (fn x => %:x === UU) (nonlazy args);
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   401
      val goal = mk_trp (iff_disj (lhs, rhss));
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   402
      val rule1 = iso_locale RS @{thm iso.abs_defined_iff};
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   403
      val rules = rule1 :: @{thms con_defined_iff_rules};
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30912
diff changeset
   404
      val tacs = [simp_tac (HOL_ss addsimps rules) 1];
30911
7809cbaa1b61 domain package: simplify internal proofs of con_rews
huffman
parents: 30829
diff changeset
   405
    in pg con_appls goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   406
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   407
  val _ = trace " Proving con_stricts...";
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   408
  val con_stricts = maps con_strict cons;
30911
7809cbaa1b61 domain package: simplify internal proofs of con_rews
huffman
parents: 30829
diff changeset
   409
  val _ = trace " Proving con_defins...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   410
  val con_defins = map con_defin cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   411
  val con_rews = con_stricts @ con_defins;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   412
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   413
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   414
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   415
  val rules =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   416
    [compact_sinl, compact_sinr, compact_spair, compact_up, compact_ONE];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   417
  fun con_compact (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   418
    let
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   419
      val concl = mk_trp (mk_compact (con_app con args));
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   420
      val goal = lift (fn x => mk_compact (%#x)) (args, concl);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   421
      val tacs = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   422
        rtac (iso_locale RS iso_compact_abs) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   423
        REPEAT (resolve_tac rules 1 ORELSE atac 1)];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   424
    in pg con_appls goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   425
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   426
  val _ = trace " Proving con_compacts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   427
  val con_compacts = map con_compact cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   428
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   429
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   430
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   431
  fun one_sel sel =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   432
    pg axs_sel_def (mk_trp (strict (%%:sel)))
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   433
      (K [simp_tac (HOLCF_ss addsimps when_rews) 1]);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   434
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   435
  fun sel_strict (_, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   436
    List.mapPartial (Option.map one_sel o sel_of) args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   437
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   438
  val _ = trace " Proving sel_stricts...";
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   439
  val sel_stricts = maps sel_strict cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   440
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   441
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   442
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   443
  fun sel_app_same c n sel (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   444
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   445
      val nlas = nonlazy args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   446
      val vns = map vname args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   447
      val vnn = List.nth (vns, n);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   448
      val nlas' = List.filter (fn v => v <> vnn) nlas;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   449
      val lhs = (%%:sel)`(con_app con args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   450
      val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   451
      fun tacs1 ctxt =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   452
        if vnn mem nlas
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   453
        then [case_UU_tac ctxt (when_rews @ con_stricts) 1 vnn]
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   454
        else [];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   455
      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   456
    in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   457
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   458
  fun sel_app_diff c n sel (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   459
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   460
      val nlas = nonlazy args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   461
      val goal = mk_trp (%%:sel ` con_app con args === UU);
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   462
      fun tacs1 ctxt = map (case_UU_tac ctxt (when_rews @ con_stricts) 1) nlas;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   463
      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   464
    in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   465
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   466
  fun sel_app c n sel (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   467
    if con = c
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   468
    then sel_app_same c n sel (con, args)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   469
    else sel_app_diff c n sel (con, args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   470
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   471
  fun one_sel c n sel = map (sel_app c n sel) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   472
  fun one_sel' c n arg = Option.map (one_sel c n) (sel_of arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   473
  fun one_con (c, args) =
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   474
    flat (map_filter I (mapn (one_sel' c) 0 args));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   475
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   476
  val _ = trace " Proving sel_apps...";
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   477
  val sel_apps = maps one_con cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   478
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   479
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   480
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   481
  fun sel_defin sel =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   482
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   483
      val goal = defined (%:x_name) ==> defined (%%:sel`%x_name);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   484
      val tacs = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   485
        rtac casedist 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   486
        contr_tac 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   487
        DETERM_UNTIL_SOLVED (CHANGED
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   488
          (asm_simp_tac (HOLCF_ss addsimps sel_apps) 1))];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   489
    in pg [] goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   490
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   491
  val _ = trace " Proving sel_defins...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   492
  val sel_defins =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   493
    if length cons = 1
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   494
    then List.mapPartial (fn arg => Option.map sel_defin (sel_of arg))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   495
                 (filter_out is_lazy (snd (hd cons)))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   496
    else [];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   497
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   498
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   499
val sel_rews = sel_stricts @ sel_defins @ sel_apps;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   500
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   501
val _ = trace " Proving dist_les...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   502
val distincts_le =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   503
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   504
    fun dist (con1, args1) (con2, args2) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   505
      let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   506
        val goal = lift_defined %: (nonlazy args1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   507
                        mk_trp (con_app con1 args1 ~<< con_app con2 args2));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   508
        fun tacs ctxt = [
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
   509
          rtac @{thm rev_contrapos} 1,
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   510
          eres_inst_tac ctxt [(("f", 0), dis_name con1)] monofun_cfun_arg 1]
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   511
          @ map (case_UU_tac ctxt (con_stricts @ dis_rews) 1) (nonlazy args2)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   512
          @ [asm_simp_tac (HOLCF_ss addsimps dis_rews) 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   513
      in pg [] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   514
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   515
    fun distinct (con1, args1) (con2, args2) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   516
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   517
          val arg1 = (con1, args1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   518
          val arg2 =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   519
            (con2, ListPair.map (fn (arg,vn) => upd_vname (K vn) arg)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   520
              (args2, Name.variant_list (map vname args1) (map vname args2)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   521
        in [dist arg1 arg2, dist arg2 arg1] end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   522
    fun distincts []      = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   523
      | distincts (c::cs) = (map (distinct c) cs) :: distincts cs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   524
  in distincts cons end;
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   525
val dist_les = flat (flat distincts_le);
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   526
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   527
val _ = trace " Proving dist_eqs...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   528
val dist_eqs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   529
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   530
    fun distinct (_,args1) ((_,args2), leqs) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   531
      let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   532
        val (le1,le2) = (hd leqs, hd(tl leqs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   533
        val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   534
      in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   535
        if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   536
        if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   537
          [eq1, eq2]
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   538
      end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   539
    fun distincts []      = []
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   540
      | distincts ((c,leqs)::cs) = flat
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   541
	            (ListPair.map (distinct c) ((map #1 cs),leqs)) @
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   542
		    distincts cs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   543
  in map standard (distincts (cons ~~ distincts_le)) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   544
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   545
local 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   546
  fun pgterm rel con args =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   547
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   548
      fun append s = upd_vname (fn v => v^s);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   549
      val (largs, rargs) = (args, map (append "'") args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   550
      val concl =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   551
        foldr1 mk_conj (ListPair.map rel (map %# largs, map %# rargs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   552
      val prem = rel (con_app con largs, con_app con rargs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   553
      val sargs = case largs of [_] => [] | _ => nonlazy args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   554
      val prop = lift_defined %: (sargs, mk_trp (prem === concl));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   555
    in pg con_appls prop end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   556
  val cons' = List.filter (fn (_,args) => args<>[]) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   557
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   558
  val _ = trace " Proving inverts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   559
  val inverts =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   560
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   561
      val abs_less = ax_abs_iso RS (allI RS injection_less);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   562
      val tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   563
        [asm_full_simp_tac (HOLCF_ss addsimps [abs_less, spair_less]) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   564
    in map (fn (con, args) => pgterm (op <<) con args (K tacs)) cons' end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   565
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   566
  val _ = trace " Proving injects...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   567
  val injects =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   568
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   569
      val abs_eq = ax_abs_iso RS (allI RS injection_eq);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   570
      val tacs = [asm_full_simp_tac (HOLCF_ss addsimps [abs_eq, spair_eq]) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   571
    in map (fn (con, args) => pgterm (op ===) con args (K tacs)) cons' end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   572
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   573
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   574
(* ----- theorems concerning one induction step ----------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   575
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   576
val copy_strict =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   577
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   578
    val goal = mk_trp (strict (dc_copy `% "f"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   579
    val tacs = [asm_simp_tac (HOLCF_ss addsimps [abs_strict, when_strict]) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   580
  in pg [ax_copy_def] goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   581
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   582
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   583
  fun copy_app (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   584
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   585
      val lhs = dc_copy`%"f"`(con_app con args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   586
      val rhs = con_app2 con (app_rec_arg (cproj (%:"f") eqs)) args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   587
      val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   588
      val args' = List.filter (fn a => not (is_rec a orelse is_lazy a)) args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   589
      val stricts = abs_strict::when_strict::con_stricts;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   590
      fun tacs1 ctxt = map (case_UU_tac ctxt stricts 1 o vname) args';
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   591
      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_apps) 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   592
    in pg [ax_copy_def] goal (fn ctxt => (tacs1 ctxt @ tacs2)) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   593
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   594
  val _ = trace " Proving copy_apps...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   595
  val copy_apps = map copy_app cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   596
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   597
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   598
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   599
  fun one_strict (con, args) = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   600
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   601
      val goal = mk_trp (dc_copy`UU`(con_app con args) === UU);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   602
      val rews = copy_strict :: copy_apps @ con_rews;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   603
      fun tacs ctxt = map (case_UU_tac ctxt rews 1) (nonlazy args) @
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   604
        [asm_simp_tac (HOLCF_ss addsimps rews) 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   605
    in pg [] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   606
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   607
  fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   608
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   609
  val _ = trace " Proving copy_stricts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   610
  val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   611
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   612
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   613
val copy_rews = copy_strict :: copy_apps @ copy_stricts;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   614
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   615
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   616
  thy
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   617
    |> Sign.add_path (Long_Name.base_name dname)
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   618
    |> snd o PureThy.add_thmss [
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   619
        ((Binding.name "iso_rews"  , iso_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   620
        ((Binding.name "exhaust"   , [exhaust]   ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   621
        ((Binding.name "casedist"  , [casedist]  ), [Induct.cases_type dname]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   622
        ((Binding.name "when_rews" , when_rews   ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   623
        ((Binding.name "compacts"  , con_compacts), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   624
        ((Binding.name "con_rews"  , con_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   625
        ((Binding.name "sel_rews"  , sel_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   626
        ((Binding.name "dis_rews"  , dis_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   627
        ((Binding.name "pat_rews"  , pat_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   628
        ((Binding.name "dist_les"  , dist_les    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   629
        ((Binding.name "dist_eqs"  , dist_eqs    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   630
        ((Binding.name "inverts"   , inverts     ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   631
        ((Binding.name "injects"   , injects     ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   632
        ((Binding.name "copy_rews" , copy_rews   ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   633
        ((Binding.name "match_rews", mat_rews    ), [Simplifier.simp_add])]
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24503
diff changeset
   634
    |> Sign.parent_path
28536
8dccb6035d0f established canonical argument order
haftmann
parents: 27239
diff changeset
   635
    |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   636
        pat_rews @ dist_les @ dist_eqs @ copy_rews)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   637
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   638
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   639
fun comp_theorems (comp_dnam, eqs: eq list) thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   640
let
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   641
val global_ctxt = ProofContext.init thy;
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   642
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   643
val dnames = map (fst o fst) eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   644
val conss  = map  snd        eqs;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28536
diff changeset
   645
val comp_dname = Sign.full_bname thy comp_dnam;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   646
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   647
val _ = message ("Proving induction properties of domain "^comp_dname^" ...");
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   648
val pg = pg' thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   649
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   650
(* ----- getting the composite axiom and definitions ------------------------ *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   651
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   652
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   653
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   654
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   655
  val axs_reach      = map (ga "reach"     ) dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   656
  val axs_take_def   = map (ga "take_def"  ) dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   657
  val axs_finite_def = map (ga "finite_def") dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   658
  val ax_copy2_def   =      ga "copy_def"  comp_dnam;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   659
  val ax_bisim_def   =      ga "bisim_def" comp_dnam;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   660
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   661
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   662
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   663
  fun gt  s dn = PureThy.get_thm  thy (dn ^ "." ^ s);
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   664
  fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   665
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   666
  val cases = map (gt  "casedist" ) dnames;
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   667
  val con_rews  = maps (gts "con_rews" ) dnames;
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   668
  val copy_rews = maps (gts "copy_rews") dnames;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   669
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   670
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   671
fun dc_take dn = %%:(dn^"_take");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   672
val x_name = idx_name dnames "x"; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   673
val P_name = idx_name dnames "P";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   674
val n_eqs = length eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   675
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   676
(* ----- theorems concerning finite approximation and finite induction ------ *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   677
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   678
local
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 26336
diff changeset
   679
  val iterate_Cprod_ss = simpset_of @{theory Fix};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   680
  val copy_con_rews  = copy_rews @ con_rews;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   681
  val copy_take_defs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   682
    (if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   683
  val _ = trace " Proving take_stricts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   684
  val take_stricts =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   685
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   686
      fun one_eq ((dn, args), _) = strict (dc_take dn $ %:"n");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   687
      val goal = mk_trp (foldr1 mk_conj (map one_eq eqs));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   688
      fun tacs ctxt = [
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   689
        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   690
        simp_tac iterate_Cprod_ss 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   691
        asm_simp_tac (iterate_Cprod_ss addsimps copy_rews) 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   692
    in pg copy_take_defs goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   693
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   694
  val take_stricts' = rewrite_rule copy_take_defs take_stricts;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   695
  fun take_0 n dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   696
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   697
      val goal = mk_trp ((dc_take dn $ %%:"HOL.zero") `% x_name n === UU);
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   698
    in pg axs_take_def goal (K [simp_tac iterate_Cprod_ss 1]) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   699
  val take_0s = mapn take_0 1 dnames;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   700
  fun c_UU_tac ctxt = case_UU_tac ctxt (take_stricts'::copy_con_rews) 1;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   701
  val _ = trace " Proving take_apps...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   702
  val take_apps =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   703
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   704
      fun mk_eqn dn (con, args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   705
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   706
          fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   707
          val lhs = (dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   708
          val rhs = con_app2 con (app_rec_arg mk_take) args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   709
        in Library.foldr mk_all (map vname args, lhs === rhs) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   710
      fun mk_eqns ((dn, _), cons) = map (mk_eqn dn) cons;
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   711
      val goal = mk_trp (foldr1 mk_conj (maps mk_eqns eqs));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   712
      val simps = List.filter (has_fewer_prems 1) copy_rews;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   713
      fun con_tac ctxt (con, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   714
        if nonlazy_rec args = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   715
        then all_tac
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   716
        else EVERY (map (c_UU_tac ctxt) (nonlazy_rec args)) THEN
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   717
          asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   718
      fun eq_tacs ctxt ((dn, _), cons) = map (con_tac ctxt) cons;
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   719
      fun tacs ctxt =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   720
        simp_tac iterate_Cprod_ss 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   721
        InductTacs.induct_tac ctxt [[SOME "n"]] 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   722
        simp_tac (iterate_Cprod_ss addsimps copy_con_rews) 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   723
        asm_full_simp_tac (HOLCF_ss addsimps simps) 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   724
        TRY (safe_tac HOL_cs) ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   725
        maps (eq_tacs ctxt) eqs;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   726
    in pg copy_take_defs goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   727
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   728
  val take_rews = map standard
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   729
    (atomize global_ctxt take_stricts @ take_0s @ atomize global_ctxt take_apps);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   730
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   731
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   732
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   733
  fun one_con p (con,args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   734
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   735
      fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   736
      val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   737
      val t2 = lift ind_hyp (List.filter is_rec args, t1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   738
      val t3 = lift_defined (bound_arg (map vname args)) (nonlazy args, t2);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   739
    in Library.foldr mk_All (map vname args, t3) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   740
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   741
  fun one_eq ((p, cons), concl) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   742
    mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   743
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   744
  fun ind_term concf = Library.foldr one_eq
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   745
    (mapn (fn n => fn x => (P_name n, x)) 1 conss,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   746
     mk_trp (foldr1 mk_conj (mapn concf 1 dnames)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   747
  val take_ss = HOL_ss addsimps take_rews;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   748
  fun quant_tac ctxt i = EVERY
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   749
    (mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   750
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   751
  fun ind_prems_tac prems = EVERY
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   752
    (maps (fn cons =>
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   753
      (resolve_tac prems 1 ::
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   754
        maps (fn (_,args) => 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   755
          resolve_tac prems 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   756
          map (K(atac 1)) (nonlazy args) @
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   757
          map (K(atac 1)) (List.filter is_rec args))
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   758
        cons))
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   759
      conss);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   760
  local 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   761
    (* check whether every/exists constructor of the n-th part of the equation:
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   762
       it has a possibly indirectly recursive argument that isn't/is possibly 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   763
       indirectly lazy *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   764
    fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   765
          is_rec arg andalso not(rec_of arg mem ns) andalso
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   766
          ((rec_of arg =  n andalso nfn(lazy_rec orelse is_lazy arg)) orelse 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   767
            rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   768
              (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg))))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   769
          ) o snd) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   770
    fun all_rec_to ns  = rec_to forall not all_rec_to  ns;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   771
    fun warn (n,cons) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   772
      if all_rec_to [] false (n,cons)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   773
      then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   774
      else false;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   775
    fun lazy_rec_to ns = rec_to exists I  lazy_rec_to ns;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   776
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   777
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   778
    val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   779
    val is_emptys = map warn n__eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   780
    val is_finite = forall (not o lazy_rec_to [] false) n__eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   781
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   782
in (* local *)
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   783
  val _ = trace " Proving finite_ind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   784
  val finite_ind =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   785
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   786
      fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   787
      val goal = ind_term concf;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   788
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   789
      fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   790
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   791
          val tacs1 = [
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   792
            quant_tac context 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   793
            simp_tac HOL_ss 1,
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   794
            InductTacs.induct_tac context [[SOME "n"]] 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   795
            simp_tac (take_ss addsimps prems) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   796
            TRY (safe_tac HOL_cs)];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   797
          fun arg_tac arg =
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   798
            case_UU_tac context (prems @ con_rews) 1
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   799
              (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   800
          fun con_tacs (con, args) = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   801
            asm_simp_tac take_ss 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   802
            map arg_tac (List.filter is_nonlazy_rec args) @
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   803
            [resolve_tac prems 1] @
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   804
            map (K (atac 1))      (nonlazy args) @
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   805
            map (K (etac spec 1)) (List.filter is_rec args);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   806
          fun cases_tacs (cons, cases) =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   807
            res_inst_tac context [(("x", 0), "x")] cases 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   808
            asm_simp_tac (take_ss addsimps prems) 1 ::
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   809
            maps con_tacs cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   810
        in
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   811
          tacs1 @ maps cases_tacs (conss ~~ cases)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   812
        end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   813
    in pg'' thy [] goal tacf end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   814
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   815
  val _ = trace " Proving take_lemmas...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   816
  val take_lemmas =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   817
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   818
      fun take_lemma n (dn, ax_reach) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   819
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   820
          val lhs = dc_take dn $ Bound 0 `%(x_name n);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   821
          val rhs = dc_take dn $ Bound 0 `%(x_name n^"'");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   822
          val concl = mk_trp (%:(x_name n) === %:(x_name n^"'"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   823
          val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   824
          fun tacf {prems, context} = [
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   825
            res_inst_tac context [(("t", 0), x_name n    )] (ax_reach RS subst) 1,
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   826
            res_inst_tac context [(("t", 0), x_name n^"'")] (ax_reach RS subst) 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   827
            stac fix_def2 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   828
            REPEAT (CHANGED
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   829
              (rtac (contlub_cfun_arg RS ssubst) 1 THEN chain_tac 1)),
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   830
            stac contlub_cfun_fun 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   831
            stac contlub_cfun_fun 2,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   832
            rtac lub_equal 3,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   833
            chain_tac 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   834
            rtac allI 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   835
            resolve_tac prems 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   836
        in pg'' thy axs_take_def goal tacf end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   837
    in mapn take_lemma 1 (dnames ~~ axs_reach) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   838
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   839
(* ----- theorems concerning finiteness and induction ----------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   840
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   841
  val _ = trace " Proving finites, ind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   842
  val (finites, ind) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   843
    if is_finite
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   844
    then (* finite case *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   845
      let 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   846
        fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %:"x" === %:"x");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   847
        fun dname_lemma dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   848
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   849
            val prem1 = mk_trp (defined (%:"x"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   850
            val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   851
            val prem2 = mk_trp (mk_disj (disj1, take_enough dn));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   852
            val concl = mk_trp (take_enough dn);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   853
            val goal = prem1 ===> prem2 ===> concl;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   854
            val tacs = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   855
              etac disjE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   856
              etac notE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   857
              resolve_tac take_lemmas 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   858
              asm_simp_tac take_ss 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   859
              atac 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   860
          in pg [] goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   861
        val finite_lemmas1a = map dname_lemma dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   862
 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   863
        val finite_lemma1b =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   864
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   865
            fun mk_eqn n ((dn, args), _) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   866
              let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   867
                val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   868
                val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   869
              in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   870
                mk_constrainall
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   871
                  (x_name n, Type (dn,args), mk_disj (disj1, disj2))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   872
              end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   873
            val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   874
              mk_trp (mk_all ("n", foldr1 mk_conj (mapn mk_eqn 1 eqs)));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   875
            fun arg_tacs ctxt vn = [
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   876
              eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   877
              etac disjE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   878
              asm_simp_tac (HOL_ss addsimps con_rews) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   879
              asm_simp_tac take_ss 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   880
            fun con_tacs ctxt (con, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   881
              asm_simp_tac take_ss 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   882
              maps (arg_tacs ctxt) (nonlazy_rec args);
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   883
            fun foo_tacs ctxt n (cons, cases) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   884
              simp_tac take_ss 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   885
              rtac allI 1 ::
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   886
              res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   887
              asm_simp_tac take_ss 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   888
              maps (con_tacs ctxt) cons;
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   889
            fun tacs ctxt =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   890
              rtac allI 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   891
              InductTacs.induct_tac ctxt [[SOME "n"]] 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   892
              simp_tac take_ss 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   893
              TRY (safe_tac (empty_cs addSEs [conjE] addSIs [conjI])) ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   894
              flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   895
          in pg [] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   896
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   897
        fun one_finite (dn, l1b) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   898
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   899
            val goal = mk_trp (%%:(dn^"_finite") $ %:"x");
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   900
            fun tacs ctxt = [
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   901
              case_UU_tac ctxt take_rews 1 "x",
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   902
              eresolve_tac finite_lemmas1a 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   903
              step_tac HOL_cs 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   904
              step_tac HOL_cs 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   905
              cut_facts_tac [l1b] 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   906
              fast_tac HOL_cs 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   907
          in pg axs_finite_def goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   908
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   909
        val finites = map one_finite (dnames ~~ atomize global_ctxt finite_lemma1b);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   910
        val ind =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   911
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   912
            fun concf n dn = %:(P_name n) $ %:(x_name n);
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   913
            fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   914
              let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   915
                fun finite_tacs (finite, fin_ind) = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   916
                  rtac(rewrite_rule axs_finite_def finite RS exE)1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   917
                  etac subst 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   918
                  rtac fin_ind 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   919
                  ind_prems_tac prems];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   920
              in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   921
                TRY (safe_tac HOL_cs) ::
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   922
                maps finite_tacs (finites ~~ atomize global_ctxt finite_ind)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   923
              end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   924
          in pg'' thy [] (ind_term concf) tacf end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   925
      in (finites, ind) end (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   926
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   927
    else (* infinite case *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   928
      let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   929
        fun one_finite n dn =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   930
          read_instantiate global_ctxt [(("P", 0), dn ^ "_finite " ^ x_name n)] excluded_middle;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   931
        val finites = mapn one_finite 1 dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   932
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   933
        val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   934
          let
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   935
            fun one_adm n _ = mk_trp (mk_adm (%:(P_name n)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   936
            fun concf n dn = %:(P_name n) $ %:(x_name n);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   937
          in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   938
        fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   939
          map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   940
          quant_tac context 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   941
          rtac (adm_impl_admw RS wfix_ind) 1,
25895
0eaadfa8889e converted adm_all and adm_ball to rule_format; cleaned up
huffman
parents: 25805
diff changeset
   942
          REPEAT_DETERM (rtac adm_all 1),
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   943
          REPEAT_DETERM (
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   944
            TRY (rtac adm_conj 1) THEN 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   945
            rtac adm_subst 1 THEN 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   946
            cont_tacR 1 THEN resolve_tac prems 1),
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   947
          strip_tac 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   948
          rtac (rewrite_rule axs_take_def finite_ind) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   949
          ind_prems_tac prems];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   950
        val ind = (pg'' thy [] goal tacf
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   951
          handle ERROR _ =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   952
            (warning "Cannot prove infinite induction rule"; refl));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   953
      in (finites, ind) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   954
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   955
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   956
(* ----- theorem concerning coinduction ------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   957
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   958
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   959
  val xs = mapn (fn n => K (x_name n)) 1 dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   960
  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   961
  val take_ss = HOL_ss addsimps take_rews;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   962
  val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   963
  val _ = trace " Proving coind_lemma...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   964
  val coind_lemma =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   965
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   966
      fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   967
      fun mk_eqn n dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   968
        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   969
        (dc_take dn $ %:"n" ` bnd_arg n 1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   970
      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   971
      val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   972
        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   973
          Library.foldr mk_all2 (xs,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   974
            Library.foldr mk_imp (mapn mk_prj 0 dnames,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   975
              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   976
      fun x_tacs ctxt n x = [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   977
        rotate_tac (n+1) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   978
        etac all2E 1,
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   979
        eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   980
        TRY (safe_tac HOL_cs),
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   981
        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   982
      fun tacs ctxt = [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   983
        rtac impI 1,
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   984
        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   985
        simp_tac take_ss 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   986
        safe_tac HOL_cs] @
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   987
        flat (mapn (x_tacs ctxt) 0 xs);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   988
    in pg [ax_bisim_def] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   989
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   990
  val _ = trace " Proving coind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   991
  val coind = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   992
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   993
      fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   994
      fun mk_eqn x = %:x === %:(x^"'");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   995
      val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   996
        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   997
          Logic.list_implies (mapn mk_prj 0 xs,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   998
            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   999
      val tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1000
        TRY (safe_tac HOL_cs) ::
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
  1001
        maps (fn take_lemma => [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1002
          rtac take_lemma 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1003
          cut_facts_tac [coind_lemma] 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1004
          fast_tac HOL_cs 1])
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
  1005
        take_lemmas;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
  1006
    in pg [] goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1007
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1008
30829
d64a293f23ba domain package registers induction rules
huffman
parents: 30807
diff changeset
  1009
val inducts = ProjectRule.projections (ProofContext.init thy) ind;
d64a293f23ba domain package registers induction rules
huffman
parents: 30807
diff changeset
  1010
fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]);
d64a293f23ba domain package registers induction rules
huffman
parents: 30807
diff changeset
  1011
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24503
diff changeset
  1012
in thy |> Sign.add_path comp_dnam
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1013
       |> snd o PureThy.add_thmss [
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1014
           ((Binding.name "take_rews"  , take_rews   ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1015
           ((Binding.name "take_lemmas", take_lemmas ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1016
           ((Binding.name "finites"    , finites     ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1017
           ((Binding.name "finite_ind" , [finite_ind]), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1018
           ((Binding.name "ind"        , [ind]       ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1019
           ((Binding.name "coind"      , [coind]     ), [])]
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
  1020
       |> snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts))
28536
8dccb6035d0f established canonical argument order
haftmann
parents: 27239
diff changeset
  1021
       |> Sign.parent_path |> pair take_rews
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1022
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1023
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
  1024
end; (* struct *)