src/HOLCF/Tools/Domain/domain_theorems.ML
author huffman
Sat, 27 Feb 2010 18:45:06 -0800
changeset 35464 2ae3362ba8ee
parent 35462 f5461b02d754
child 35466 9fcfd5763181
permissions -rw-r--r--
removed dead code
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32126
a5042f260440 obey captialized directory names convention
haftmann
parents: 31288
diff changeset
     1
(*  Title:      HOLCF/Tools/Domain/domain_theorems.ML
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     David von Oheimb
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
     3
    Author:     Brian Huffman
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
Proof generator for domain command.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 26336
diff changeset
     8
val HOLCF_ss = @{simpset};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    10
signature DOMAIN_THEOREMS =
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    11
sig
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
    12
  val theorems:
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
    13
    Domain_Library.eq * Domain_Library.eq list
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
    14
    -> typ * (binding * (bool * binding option * typ) list * mixfix) list
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
    15
    -> theory -> thm list * theory;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
    16
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    17
  val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory;
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    18
  val quiet_mode: bool Unsynchronized.ref;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    19
  val trace_domain: bool Unsynchronized.ref;
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    20
end;
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    21
31023
d027411c9a38 use opaque ascription for all HOLCF code
huffman
parents: 31005
diff changeset
    22
structure Domain_Theorems :> DOMAIN_THEOREMS =
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    23
struct
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    24
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    25
val quiet_mode = Unsynchronized.ref false;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    26
val trace_domain = Unsynchronized.ref false;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    27
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    28
fun message s = if !quiet_mode then () else writeln s;
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    29
fun trace s = if !trace_domain then tracing s else ();
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    30
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    31
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
    32
val adm_all = @{thm adm_all};
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    33
val adm_conj = @{thm adm_conj};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    34
val adm_subst = @{thm adm_subst};
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    35
val ch2ch_fst = @{thm ch2ch_fst};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    36
val ch2ch_snd = @{thm ch2ch_snd};
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    37
val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    38
val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    39
val chain_iterate = @{thm chain_iterate};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    40
val contlub_cfun_fun = @{thm contlub_cfun_fun};
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    41
val contlub_fst = @{thm contlub_fst};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    42
val contlub_snd = @{thm contlub_snd};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    43
val contlubE = @{thm contlubE};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    44
val cont_const = @{thm cont_const};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    45
val cont_id = @{thm cont_id};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    46
val cont2cont_fst = @{thm cont2cont_fst};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    47
val cont2cont_snd = @{thm cont2cont_snd};
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
    48
val cont2cont_Rep_CFun = @{thm cont2cont_Rep_CFun};
25805
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    49
val fix_def2 = @{thm fix_def2};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    50
val lub_equal = @{thm lub_equal};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    51
val retraction_strict = @{thm retraction_strict};
5df82bb5b982 new-style theorem references
huffman
parents: 25132
diff changeset
    52
val wfix_ind = @{thm wfix_ind};
35464
2ae3362ba8ee removed dead code
huffman
parents: 35462
diff changeset
    53
val iso_intro = @{thm iso.intro};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    54
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    55
open Domain_Library;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
infixr 0 ===>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    57
infixr 0 ==>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
infix 0 == ; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
infix 1 ===;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    60
infix 1 ~= ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
infix 1 <<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
infix 1 ~<<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    63
infix 9 `   ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    64
infix 9 `% ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
infix 9 `%%;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    66
infixr 9 oo;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    67
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    68
(* ----- general proof facilities ------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    69
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    70
fun legacy_infer_term thy t =
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    71
  let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy)
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    72
  in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end;
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    73
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    74
fun pg'' thy defs t tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
  let
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    76
    val t' = legacy_infer_term thy t;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
    val asms = Logic.strip_imp_prems t';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    78
    val prop = Logic.strip_imp_concl t';
26711
3a478bfa1650 prove_global: pass context;
wenzelm
parents: 26343
diff changeset
    79
    fun tac {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
      rewrite_goals_tac defs THEN
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    81
      EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
  in Goal.prove_global thy [] asms prop tac end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    84
fun pg' thy defs t tacsf =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
  let
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    86
    fun tacs {prems, context} =
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    87
      if null prems then tacsf context
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    88
      else cut_facts_tac prems 1 :: tacsf context;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
  in pg'' thy defs t tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
    91
(* FIXME!!!!!!!!! *)
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
    92
(* We should NEVER re-parse variable names as strings! *)
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
    93
(* The names can conflict with existing constants or other syntax! *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    94
fun case_UU_tac ctxt rews i v =
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    95
  InductTacs.case_tac ctxt (v^"=UU") i THEN
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
  asm_simp_tac (HOLCF_ss addsimps rews) i;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    97
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    98
val chain_tac =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
  REPEAT_DETERM o resolve_tac 
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   100
    [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL, ch2ch_fst, ch2ch_snd];
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   101
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   102
(* ----- general proofs ----------------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   103
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
   104
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
   105
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   106
fun theorems
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   107
    (((dname, _), cons) : eq, eqs : eq list)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   108
    (dom_eqn : typ * (binding * (bool * binding option * typ) list * mixfix) list)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   109
    (thy : theory) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   110
let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   112
val _ = message ("Proving isomorphism properties of domain "^dname^" ...");
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33504
diff changeset
   113
val map_tab = Domain_Isomorphism.get_map_tab thy;
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33504
diff changeset
   114
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   115
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   116
(* ----- getting the axioms and definitions --------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   117
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   119
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   120
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   121
  val ax_abs_iso  = ga "abs_iso"  dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   122
  val ax_rep_iso  = ga "rep_iso"  dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   123
  val ax_when_def = ga "when_def" dname;
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   124
  fun get_def mk_name (con, _, _) = ga (mk_name con^"_def") dname;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   125
  val axs_pat_def = map (get_def pat_name) cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   126
  val ax_copy_def = ga "copy_def" dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   127
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   129
(* ----- define constructors ------------------------------------------------ *)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   130
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   131
val lhsT = fst dom_eqn;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   132
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   133
val rhsT =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   134
  let
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   135
    fun mk_arg_typ (lazy, sel, T) = if lazy then mk_uT T else T;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   136
    fun mk_con_typ (bind, args, mx) =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   137
        if null args then oneT else foldr1 mk_sprodT (map mk_arg_typ args);
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   138
    fun mk_eq_typ (_, cons) = foldr1 mk_ssumT (map mk_con_typ cons);
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   139
  in
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   140
    mk_eq_typ dom_eqn
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   141
  end;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   142
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   143
val rep_const = Const(dname^"_rep", lhsT ->> rhsT);
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   144
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   145
val abs_const = Const(dname^"_abs", rhsT ->> lhsT);
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   146
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   147
val (result, thy) =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   148
  Domain_Constructors.add_domain_constructors
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   149
    (Long_Name.base_name dname) dom_eqn
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   150
    (rep_const, abs_const) (ax_rep_iso, ax_abs_iso) ax_when_def thy;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   151
35451
a726a033b313 don't bother returning con_defs
huffman
parents: 35448
diff changeset
   152
val con_appls = #con_betas result;
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   153
val {exhaust, casedist, ...} = result;
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   154
val {con_compacts, con_rews, inverts, injects, dist_les, dist_eqs, ...} = result;
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   155
val {sel_rews, ...} = result;
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   156
val when_rews = #cases result;
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   157
val when_strict = hd when_rews;
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   158
val dis_rews = #dis_rews result;
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   159
val axs_mat_def = #match_rews result;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   160
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   161
(* ----- theorems concerning the isomorphism -------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   162
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   163
val pg = pg' thy;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   164
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   165
val dc_copy = %%:(dname^"_copy");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   166
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   167
val abs_strict = ax_rep_iso RS (allI RS retraction_strict);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   168
val rep_strict = ax_abs_iso RS (allI RS retraction_strict);
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34974
diff changeset
   169
val iso_rews = map Drule.export_without_context [ax_abs_iso, ax_rep_iso, abs_strict, rep_strict];
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   170
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   171
(* ----- theorems concerning the constructors, discriminators and selectors - *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   172
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   173
local
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   174
  fun mat_strict (con, _, _) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   175
    let
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   176
      val goal = mk_trp (%%:(mat_name con) ` UU ` %:"rhs" === UU);
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   177
      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
   178
    in pg axs_mat_def goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   179
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   180
  val _ = trace " Proving mat_stricts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   181
  val mat_stricts = map mat_strict cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   182
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   183
  fun one_mat c (con, _, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   184
    let
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   185
      val lhs = %%:(mat_name c) ` con_app con args ` %:"rhs";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   186
      val rhs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   187
        if con = c
30912
4022298c1a86 change definition of match combinators for fixrec package
huffman
parents: 30911
diff changeset
   188
        then list_ccomb (%:"rhs", map %# args)
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   189
        else mk_fail;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   190
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   191
      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
   192
    in pg axs_mat_def goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   193
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   194
  val _ = trace " Proving mat_apps...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   195
  val mat_apps =
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   196
    maps (fn (c,_,_) => map (one_mat c) cons) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   197
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   198
  val mat_rews = mat_stricts @ mat_apps;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   199
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   200
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   201
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   202
  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
   203
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   204
  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
   205
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   206
  fun pat_rhs (con,_,[]) = mk_return ((%:"rhs") ` HOLogic.unit)
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   207
    | pat_rhs (con,_,args) =
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   208
        (mk_branch (mk_ctuple_pat (ps args)))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   209
          `(%:"rhs")`(mk_ctuple (map %# args));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   210
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   211
  fun pat_strict c =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   212
    let
25132
dffe405b090d removed obsolete ML bindings;
wenzelm
parents: 24712
diff changeset
   213
      val axs = @{thm branch_def} :: axs_pat_def;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   214
      val goal = mk_trp (strict (pat_lhs c ` (%:"rhs")));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   215
      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
   216
    in pg axs goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   217
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   218
  fun pat_app c (con, _, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   219
    let
25132
dffe405b090d removed obsolete ML bindings;
wenzelm
parents: 24712
diff changeset
   220
      val axs = @{thm branch_def} :: axs_pat_def;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   221
      val lhs = (pat_lhs c)`(%:"rhs")`(con_app con args);
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   222
      val rhs = if con = first c then pat_rhs c else mk_fail;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   223
      val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   224
      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
   225
    in pg axs goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   226
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   227
  val _ = trace " Proving pat_stricts...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   228
  val pat_stricts = map pat_strict cons;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   229
  val _ = trace " Proving pat_apps...";
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   230
  val pat_apps = maps (fn c => map (pat_app c) cons) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   231
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   232
  val pat_rews = pat_stricts @ pat_apps;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   233
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   234
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   235
(* ----- theorems concerning one induction step ----------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   236
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   237
val copy_strict =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   238
  let
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   239
    val _ = trace " Proving copy_strict...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   240
    val goal = mk_trp (strict (dc_copy `% "f"));
33504
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33427
diff changeset
   241
    val rules = [abs_strict, rep_strict] @ @{thms domain_map_stricts};
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   242
    val tacs = [asm_simp_tac (HOLCF_ss addsimps rules) 1];
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   243
  in
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   244
    SOME (pg [ax_copy_def] goal (K tacs))
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   245
    handle
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   246
      THM (s, _, _) => (trace s; NONE)
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   247
    | ERROR s => (trace s; NONE)
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   248
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   249
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   250
local
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   251
  fun copy_app (con, _, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   252
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   253
      val lhs = dc_copy`%"f"`(con_app con args);
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   254
      fun one_rhs arg =
33971
9c7fa7f76950 modernized structure Datatype_Aux
haftmann
parents: 33810
diff changeset
   255
          if Datatype_Aux.is_rec_type (dtyp_of arg)
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33504
diff changeset
   256
          then Domain_Axioms.copy_of_dtyp map_tab
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33504
diff changeset
   257
                 (proj (%:"f") eqs) (dtyp_of arg) ` (%# arg)
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   258
          else (%# arg);
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   259
      val rhs = con_app2 con one_rhs args;
35059
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   260
      fun is_rec arg = Datatype_Aux.is_rec_type (dtyp_of arg);
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   261
      fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   262
      fun nonlazy_rec args = map vname (filter is_nonlazy_rec args);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   263
      val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs));
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   264
      val args' = filter_out (fn a => is_rec a orelse is_lazy a) args;
33504
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33427
diff changeset
   265
      val stricts = abs_strict :: rep_strict :: @{thms domain_map_stricts};
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   266
                        (* FIXME! case_UU_tac *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   267
      fun tacs1 ctxt = map (case_UU_tac ctxt stricts 1 o vname) args';
33504
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33427
diff changeset
   268
      val rules = [ax_abs_iso] @ @{thms domain_map_simps};
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   269
      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps rules) 1];
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   270
    in pg (ax_copy_def::con_appls) goal (fn ctxt => (tacs1 ctxt @ tacs2)) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   271
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   272
  val _ = trace " Proving copy_apps...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   273
  val copy_apps = map copy_app cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   274
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   275
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   276
local
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   277
  fun one_strict (con, _, args) = 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   278
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   279
      val goal = mk_trp (dc_copy`UU`(con_app con args) === UU);
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   280
      val rews = the_list copy_strict @ copy_apps @ con_rews;
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   281
                        (* FIXME! case_UU_tac *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   282
      fun tacs ctxt = map (case_UU_tac ctxt rews 1) (nonlazy args) @
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   283
        [asm_simp_tac (HOLCF_ss addsimps rews) 1];
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   284
    in
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   285
      SOME (pg [] goal tacs)
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   286
      handle
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   287
        THM (s, _, _) => (trace s; NONE)
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   288
      | ERROR s => (trace s; NONE)
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   289
    end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   290
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   291
  fun has_nonlazy_rec (_, _, args) = exists is_nonlazy_rec args;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   292
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   293
  val _ = trace " Proving copy_stricts...";
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   294
  val copy_stricts = map_filter one_strict (filter has_nonlazy_rec cons);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   295
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   296
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   297
val copy_rews = the_list copy_strict @ copy_apps @ copy_stricts;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   298
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   299
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   300
  thy
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   301
    |> Sign.add_path (Long_Name.base_name dname)
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   302
    |> snd o PureThy.add_thmss [
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   303
        ((Binding.name "iso_rews"  , iso_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   304
        ((Binding.name "exhaust"   , [exhaust]   ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   305
        ((Binding.name "casedist"  , [casedist]  ), [Induct.cases_type dname]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   306
        ((Binding.name "when_rews" , when_rews   ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   307
        ((Binding.name "compacts"  , con_compacts), [Simplifier.simp_add]),
33427
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   308
        ((Binding.name "con_rews"  , con_rews    ),
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   309
         [Simplifier.simp_add, Fixrec.fixrec_simp_add]),
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   310
        ((Binding.name "sel_rews"  , sel_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   311
        ((Binding.name "dis_rews"  , dis_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   312
        ((Binding.name "pat_rews"  , pat_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   313
        ((Binding.name "dist_les"  , dist_les    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   314
        ((Binding.name "dist_eqs"  , dist_eqs    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   315
        ((Binding.name "inverts"   , inverts     ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   316
        ((Binding.name "injects"   , injects     ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   317
        ((Binding.name "copy_rews" , copy_rews   ), [Simplifier.simp_add]),
33427
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   318
        ((Binding.name "match_rews", mat_rews    ),
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   319
         [Simplifier.simp_add, Fixrec.fixrec_simp_add])]
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24503
diff changeset
   320
    |> Sign.parent_path
28536
8dccb6035d0f established canonical argument order
haftmann
parents: 27239
diff changeset
   321
    |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   322
        pat_rews @ dist_les @ dist_eqs @ copy_rews)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   323
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   324
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   325
fun comp_theorems (comp_dnam, eqs: eq list) thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   326
let
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   327
val global_ctxt = ProofContext.init thy;
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33504
diff changeset
   328
val map_tab = Domain_Isomorphism.get_map_tab thy;
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   329
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   330
val dnames = map (fst o fst) eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   331
val conss  = map  snd        eqs;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28536
diff changeset
   332
val comp_dname = Sign.full_bname thy comp_dnam;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   333
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   334
val _ = message ("Proving induction properties of domain "^comp_dname^" ...");
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   335
val pg = pg' thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   336
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   337
(* ----- getting the composite axiom and definitions ------------------------ *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   338
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   339
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   340
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   341
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   342
  val axs_reach      = map (ga "reach"     ) dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   343
  val axs_take_def   = map (ga "take_def"  ) dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   344
  val axs_finite_def = map (ga "finite_def") dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   345
  val ax_copy2_def   =      ga "copy_def"  comp_dnam;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   346
(* TEMPORARILY DISABLED
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   347
  val ax_bisim_def   =      ga "bisim_def" comp_dnam;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   348
TEMPORARILY DISABLED *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   349
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   350
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   351
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   352
  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
   353
  fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   354
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   355
  val cases = map (gt  "casedist" ) dnames;
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   356
  val con_rews  = maps (gts "con_rews" ) dnames;
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   357
  val copy_rews = maps (gts "copy_rews") dnames;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   358
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   359
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   360
fun dc_take dn = %%:(dn^"_take");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   361
val x_name = idx_name dnames "x"; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   362
val P_name = idx_name dnames "P";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   363
val n_eqs = length eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   364
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   365
(* ----- theorems concerning finite approximation and finite induction ------ *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   366
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   367
local
32149
ef59550a55d3 renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents: 32126
diff changeset
   368
  val iterate_Cprod_ss = global_simpset_of @{theory Fix};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   369
  val copy_con_rews  = copy_rews @ con_rews;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   370
  val copy_take_defs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   371
    (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
   372
  val _ = trace " Proving take_stricts...";
35057
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   373
  fun one_take_strict ((dn, args), _) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   374
    let
35057
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   375
      val goal = mk_trp (strict (dc_take dn $ %:"n"));
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   376
      val rules = [
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   377
        @{thm monofun_fst [THEN monofunE]},
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   378
        @{thm monofun_snd [THEN monofunE]}];
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   379
      val tacs = [
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   380
        rtac @{thm UU_I} 1,
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   381
        rtac @{thm below_eq_trans} 1,
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   382
        resolve_tac axs_reach 2,
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   383
        rtac @{thm monofun_cfun_fun} 1,
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   384
        REPEAT (resolve_tac rules 1),
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   385
        rtac @{thm iterate_below_fix} 1];
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   386
    in pg axs_take_def goal (K tacs) end;
03d023236fcd rewrite proof script for take_stricts
huffman
parents: 34974
diff changeset
   387
  val take_stricts = map one_take_strict eqs;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   388
  fun take_0 n dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   389
    let
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   390
      val goal = mk_trp ((dc_take dn $ @{term "0::nat"}) `% x_name n === UU);
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   391
    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
   392
  val take_0s = mapn take_0 1 dnames;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   393
  val _ = trace " Proving take_apps...";
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   394
  fun one_take_app dn (con, _, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   395
    let
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   396
      fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n";
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   397
      fun one_rhs arg =
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   398
          if Datatype_Aux.is_rec_type (dtyp_of arg)
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   399
          then Domain_Axioms.copy_of_dtyp map_tab
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   400
                 mk_take (dtyp_of arg) ` (%# arg)
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   401
          else (%# arg);
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   402
      val lhs = (dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args);
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   403
      val rhs = con_app2 con one_rhs args;
35059
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   404
      fun is_rec arg = Datatype_Aux.is_rec_type (dtyp_of arg);
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   405
      fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   406
      fun nonlazy_rec args = map vname (filter is_nonlazy_rec args);
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   407
      val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs));
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   408
      val tacs = [asm_simp_tac (HOLCF_ss addsimps copy_con_rews) 1];
acbc346e5310 correct definedness side conditions for copy_apps and take_apps
huffman
parents: 35058
diff changeset
   409
    in pg copy_take_defs goal (K tacs) end;
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   410
  fun one_take_apps ((dn, _), cons) = map (one_take_app dn) cons;
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   411
  val take_apps = maps one_take_apps eqs;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   412
in
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34974
diff changeset
   413
  val take_rews = map Drule.export_without_context
35058
d0cc1650b378 handle case where copy_stricts cannot be proven; rewrite proof script for take_apps
huffman
parents: 35057
diff changeset
   414
    (take_stricts @ take_0s @ take_apps);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   415
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   416
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   417
local
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   418
  fun one_con p (con, _, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   419
    let
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   420
      val P_names = map P_name (1 upto (length dnames));
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   421
      val vns = Name.variant_list P_names (map vname args);
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   422
      val nonlazy_vns = map snd (filter_out (is_lazy o fst) (args ~~ vns));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   423
      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
   424
      val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args);
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   425
      val t2 = lift ind_hyp (filter is_rec args, t1);
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   426
      val t3 = lift_defined (bound_arg vns) (nonlazy_vns, t2);
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   427
    in Library.foldr mk_All (vns, t3) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   428
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   429
  fun one_eq ((p, cons), concl) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   430
    mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   431
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   432
  fun ind_term concf = Library.foldr one_eq
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   433
    (mapn (fn n => fn x => (P_name n, x)) 1 conss,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   434
     mk_trp (foldr1 mk_conj (mapn concf 1 dnames)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   435
  val take_ss = HOL_ss addsimps take_rews;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   436
  fun quant_tac ctxt i = EVERY
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   437
    (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
   438
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   439
  fun ind_prems_tac prems = EVERY
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   440
    (maps (fn cons =>
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   441
      (resolve_tac prems 1 ::
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   442
        maps (fn (_,_,args) => 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   443
          resolve_tac prems 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   444
          map (K(atac 1)) (nonlazy args) @
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   445
          map (K(atac 1)) (filter is_rec args))
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   446
        cons))
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   447
      conss);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   448
  local 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   449
    (* check whether every/exists constructor of the n-th part of the equation:
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   450
       it has a possibly indirectly recursive argument that isn't/is possibly 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   451
       indirectly lazy *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   452
    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
   453
          is_rec arg andalso not(rec_of arg mem ns) andalso
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   454
          ((rec_of arg =  n andalso nfn(lazy_rec orelse is_lazy arg)) orelse 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   455
            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
   456
              (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg))))
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   457
          ) o third) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   458
    fun all_rec_to ns  = rec_to forall not all_rec_to  ns;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   459
    fun warn (n,cons) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   460
      if all_rec_to [] false (n,cons)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   461
      then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   462
      else false;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   463
    fun lazy_rec_to ns = rec_to exists I  lazy_rec_to ns;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   464
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   465
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   466
    val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   467
    val is_emptys = map warn n__eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   468
    val is_finite = forall (not o lazy_rec_to [] false) n__eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   469
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   470
in (* local *)
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   471
  val _ = trace " Proving finite_ind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   472
  val finite_ind =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   473
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   474
      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
   475
      val goal = ind_term concf;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   476
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   477
      fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   478
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   479
          val tacs1 = [
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   480
            quant_tac context 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   481
            simp_tac HOL_ss 1,
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   482
            InductTacs.induct_tac context [[SOME "n"]] 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   483
            simp_tac (take_ss addsimps prems) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   484
            TRY (safe_tac HOL_cs)];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   485
          fun arg_tac arg =
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   486
                        (* FIXME! case_UU_tac *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   487
            case_UU_tac context (prems @ con_rews) 1
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   488
              (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg);
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   489
          fun con_tacs (con, _, args) = 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   490
            asm_simp_tac take_ss 1 ::
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   491
            map arg_tac (filter is_nonlazy_rec args) @
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   492
            [resolve_tac prems 1] @
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   493
            map (K (atac 1)) (nonlazy args) @
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   494
            map (K (etac spec 1)) (filter is_rec args);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   495
          fun cases_tacs (cons, cases) =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   496
            res_inst_tac context [(("x", 0), "x")] cases 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   497
            asm_simp_tac (take_ss addsimps prems) 1 ::
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   498
            maps con_tacs cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   499
        in
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   500
          tacs1 @ maps cases_tacs (conss ~~ cases)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   501
        end;
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   502
    in pg'' thy [] goal tacf
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   503
       handle ERROR _ => (warning "Proof of finite_ind failed."; TrueI)
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   504
    end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   505
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   506
  val _ = trace " Proving take_lemmas...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   507
  val take_lemmas =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   508
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   509
      fun take_lemma n (dn, ax_reach) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   510
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   511
          val lhs = dc_take dn $ Bound 0 `%(x_name n);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   512
          val rhs = dc_take dn $ Bound 0 `%(x_name n^"'");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   513
          val concl = mk_trp (%:(x_name n) === %:(x_name n^"'"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   514
          val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl;
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   515
          val rules = [contlub_fst RS contlubE RS ssubst,
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   516
                       contlub_snd RS contlubE RS ssubst];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   517
          fun tacf {prems, context} = [
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   518
            res_inst_tac context [(("t", 0), x_name n    )] (ax_reach RS subst) 1,
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   519
            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
   520
            stac fix_def2 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   521
            REPEAT (CHANGED
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   522
              (resolve_tac rules 1 THEN chain_tac 1)),
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   523
            stac contlub_cfun_fun 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   524
            stac contlub_cfun_fun 2,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   525
            rtac lub_equal 3,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   526
            chain_tac 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   527
            rtac allI 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   528
            resolve_tac prems 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   529
        in pg'' thy axs_take_def goal tacf end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   530
    in mapn take_lemma 1 (dnames ~~ axs_reach) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   531
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   532
(* ----- theorems concerning finiteness and induction ----------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   533
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   534
  val _ = trace " Proving finites, ind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   535
  val (finites, ind) =
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   536
  (
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   537
    if is_finite
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   538
    then (* finite case *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   539
      let 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   540
        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
   541
        fun dname_lemma dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   542
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   543
            val prem1 = mk_trp (defined (%:"x"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   544
            val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   545
            val prem2 = mk_trp (mk_disj (disj1, take_enough dn));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   546
            val concl = mk_trp (take_enough dn);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   547
            val goal = prem1 ===> prem2 ===> concl;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   548
            val tacs = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   549
              etac disjE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   550
              etac notE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   551
              resolve_tac take_lemmas 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   552
              asm_simp_tac take_ss 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   553
              atac 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   554
          in pg [] goal (K tacs) end;
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   555
        val _ = trace " Proving finite_lemmas1a";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   556
        val finite_lemmas1a = map dname_lemma dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   557
 
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   558
        val _ = trace " Proving finite_lemma1b";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   559
        val finite_lemma1b =
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
            fun mk_eqn n ((dn, args), _) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   562
              let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   563
                val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   564
                val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   565
              in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   566
                mk_constrainall
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   567
                  (x_name n, Type (dn,args), mk_disj (disj1, disj2))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   568
              end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   569
            val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   570
              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
   571
            fun arg_tacs ctxt vn = [
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   572
              eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   573
              etac disjE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   574
              asm_simp_tac (HOL_ss addsimps con_rews) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   575
              asm_simp_tac take_ss 1];
35288
aa7da51ae1ef add mixfix field to type Domain_Library.cons
huffman
parents: 35287
diff changeset
   576
            fun con_tacs ctxt (con, _, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   577
              asm_simp_tac take_ss 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   578
              maps (arg_tacs ctxt) (nonlazy_rec args);
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   579
            fun foo_tacs ctxt n (cons, cases) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   580
              simp_tac take_ss 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   581
              rtac allI 1 ::
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   582
              res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   583
              asm_simp_tac take_ss 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   584
              maps (con_tacs ctxt) cons;
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   585
            fun tacs ctxt =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   586
              rtac allI 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   587
              InductTacs.induct_tac ctxt [[SOME "n"]] 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   588
              simp_tac take_ss 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   589
              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
   590
              flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   591
          in pg [] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   592
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   593
        fun one_finite (dn, l1b) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   594
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   595
            val goal = mk_trp (%%:(dn^"_finite") $ %:"x");
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   596
            fun tacs ctxt = [
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   597
                        (* FIXME! case_UU_tac *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   598
              case_UU_tac ctxt take_rews 1 "x",
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   599
              eresolve_tac finite_lemmas1a 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   600
              step_tac HOL_cs 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   601
              step_tac HOL_cs 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   602
              cut_facts_tac [l1b] 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   603
              fast_tac HOL_cs 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   604
          in pg axs_finite_def goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   605
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   606
        val _ = trace " Proving finites";
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   607
        val finites = map one_finite (dnames ~~ atomize global_ctxt finite_lemma1b);
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   608
        val _ = trace " Proving ind";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   609
        val ind =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   610
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   611
            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
   612
            fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   613
              let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   614
                fun finite_tacs (finite, fin_ind) = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   615
                  rtac(rewrite_rule axs_finite_def finite RS exE)1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   616
                  etac subst 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   617
                  rtac fin_ind 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   618
                  ind_prems_tac prems];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   619
              in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   620
                TRY (safe_tac HOL_cs) ::
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   621
                maps finite_tacs (finites ~~ atomize global_ctxt finite_ind)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   622
              end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   623
          in pg'' thy [] (ind_term concf) tacf end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   624
      in (finites, ind) end (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   625
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   626
    else (* infinite case *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   627
      let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   628
        fun one_finite n dn =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   629
          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
   630
        val finites = mapn one_finite 1 dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   631
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   632
        val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   633
          let
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   634
            fun one_adm n _ = mk_trp (mk_adm (%:(P_name n)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   635
            fun concf n dn = %:(P_name n) $ %:(x_name n);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   636
          in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end;
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   637
        val cont_rules =
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   638
            [cont_id, cont_const, cont2cont_Rep_CFun,
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   639
             cont2cont_fst, cont2cont_snd];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   640
        fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   641
          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
   642
          quant_tac context 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   643
          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
   644
          REPEAT_DETERM (rtac adm_all 1),
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   645
          REPEAT_DETERM (
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   646
            TRY (rtac adm_conj 1) THEN 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   647
            rtac adm_subst 1 THEN 
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   648
            REPEAT (resolve_tac cont_rules 1) THEN
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   649
            resolve_tac prems 1),
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   650
          strip_tac 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   651
          rtac (rewrite_rule axs_take_def finite_ind) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   652
          ind_prems_tac prems];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   653
        val ind = (pg'' thy [] goal tacf
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   654
          handle ERROR _ =>
33396
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   655
            (warning "Cannot prove infinite induction rule"; TrueI));
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   656
      in (finites, ind) end
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   657
  )
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   658
      handle THM _ =>
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   659
             (warning "Induction proofs failed (THM raised)."; ([], TrueI))
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   660
           | ERROR _ =>
33810
38375b16ffd9 nicer warning message for indirect-recursive domain definitions
huffman
parents: 33801
diff changeset
   661
             (warning "Cannot prove induction rule"; ([], TrueI));
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   662
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   663
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   664
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   665
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   666
(* ----- theorem concerning coinduction ------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   667
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   668
(* COINDUCTION TEMPORARILY DISABLED
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   669
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   670
  val xs = mapn (fn n => K (x_name n)) 1 dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   671
  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   672
  val take_ss = HOL_ss addsimps take_rews;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   673
  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
   674
  val _ = trace " Proving coind_lemma...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   675
  val coind_lemma =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   676
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   677
      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
   678
      fun mk_eqn n dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   679
        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   680
        (dc_take dn $ %:"n" ` bnd_arg n 1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   681
      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   682
      val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   683
        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   684
          Library.foldr mk_all2 (xs,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   685
            Library.foldr mk_imp (mapn mk_prj 0 dnames,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   686
              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   687
      fun x_tacs ctxt n x = [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   688
        rotate_tac (n+1) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   689
        etac all2E 1,
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   690
        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
   691
        TRY (safe_tac HOL_cs),
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   692
        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   693
      fun tacs ctxt = [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   694
        rtac impI 1,
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   695
        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   696
        simp_tac take_ss 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   697
        safe_tac HOL_cs] @
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   698
        flat (mapn (x_tacs ctxt) 0 xs);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   699
    in pg [ax_bisim_def] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   700
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   701
  val _ = trace " Proving coind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   702
  val coind = 
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_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   705
      fun mk_eqn x = %:x === %:(x^"'");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   706
      val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   707
        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   708
          Logic.list_implies (mapn mk_prj 0 xs,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   709
            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   710
      val tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   711
        TRY (safe_tac HOL_cs) ::
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   712
        maps (fn take_lemma => [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   713
          rtac take_lemma 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   714
          cut_facts_tac [coind_lemma] 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   715
          fast_tac HOL_cs 1])
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   716
        take_lemmas;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   717
    in pg [] goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   718
end; (* local *)
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   719
COINDUCTION TEMPORARILY DISABLED *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   720
32172
c4e55f30d527 renamed functor ProjectRuleFun to Project_Rule;
wenzelm
parents: 32149
diff changeset
   721
val inducts = Project_Rule.projections (ProofContext.init thy) ind;
30829
d64a293f23ba domain package registers induction rules
huffman
parents: 30807
diff changeset
   722
fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]);
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   723
val induct_failed = (Thm.prop_of ind = Thm.prop_of TrueI);
30829
d64a293f23ba domain package registers induction rules
huffman
parents: 30807
diff changeset
   724
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24503
diff changeset
   725
in thy |> Sign.add_path comp_dnam
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   726
       |> snd o PureThy.add_thmss [
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   727
           ((Binding.name "take_rews"  , take_rews   ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   728
           ((Binding.name "take_lemmas", take_lemmas ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   729
           ((Binding.name "finites"    , finites     ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   730
           ((Binding.name "finite_ind" , [finite_ind]), []),
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   731
           ((Binding.name "ind"        , [ind]       ), [])(*,
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   732
           ((Binding.name "coind"      , [coind]     ), [])*)]
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   733
       |> (if induct_failed then I
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   734
           else snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts)))
28536
8dccb6035d0f established canonical argument order
haftmann
parents: 27239
diff changeset
   735
       |> Sign.parent_path |> pair take_rews
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   736
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   737
end; (* struct *)