src/HOLCF/Tools/Domain/domain_theorems.ML
author huffman
Tue, 02 Mar 2010 15:53:07 -0800
changeset 35521 47eec4da124a
parent 35514 a2cfa413eaab
child 35523 cc57f4a274a3
permissions -rw-r--r--
remove unused mixfix component from type cons
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
(* ----- general proofs ----------------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
   100
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
   101
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   102
fun theorems
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   103
    (((dname, _), cons) : eq, eqs : eq list)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   104
    (dom_eqn : typ * (binding * (bool * binding option * typ) list * mixfix) list)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   105
    (thy : theory) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   106
let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   107
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   108
val _ = message ("Proving isomorphism properties of domain "^dname^" ...");
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35512
diff changeset
   109
val map_tab = Domain_Take_Proofs.get_map_tab thy;
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33504
diff changeset
   110
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   112
(* ----- getting the axioms and definitions --------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   113
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   114
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   115
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   116
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   117
  val ax_abs_iso  = ga "abs_iso"  dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
  val ax_rep_iso  = ga "rep_iso"  dname;
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   119
  val ax_take_0      = ga "take_0" dname;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   120
  val ax_take_Suc    = ga "take_Suc" dname;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   121
  val ax_take_strict = ga "take_strict" dname;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   122
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   123
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   124
(* ----- define constructors ------------------------------------------------ *)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   125
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   126
val lhsT = fst dom_eqn;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   127
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   128
val rhsT =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   129
  let
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   130
    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
   131
    fun mk_con_typ (bind, args, mx) =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   132
        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
   133
    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
   134
  in
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   135
    mk_eq_typ dom_eqn
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   136
  end;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   137
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   138
val rep_const = Const(dname^"_rep", lhsT ->> rhsT);
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   139
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   140
val abs_const = Const(dname^"_abs", rhsT ->> lhsT);
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   141
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35512
diff changeset
   142
val iso_info : Domain_Take_Proofs.iso_info =
35481
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   143
  {
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   144
    absT = lhsT,
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   145
    repT = rhsT,
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   146
    abs_const = abs_const,
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   147
    rep_const = rep_const,
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   148
    abs_inverse = ax_abs_iso,
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   149
    rep_inverse = ax_rep_iso
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   150
  };
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35468
diff changeset
   151
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   152
val (result, thy) =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   153
  Domain_Constructors.add_domain_constructors
35486
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35482
diff changeset
   154
    (Long_Name.base_name dname) (snd dom_eqn) iso_info thy;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   155
35451
a726a033b313 don't bother returning con_defs
huffman
parents: 35448
diff changeset
   156
val con_appls = #con_betas result;
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   157
val {exhaust, casedist, ...} = result;
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   158
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
   159
val {sel_rews, ...} = result;
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   160
val when_rews = #cases result;
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   161
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
   162
val dis_rews = #dis_rews result;
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35464
diff changeset
   163
val mat_rews = #match_rews result;
35482
d756837b708d move proofs of pat_rews to domain_constructors.ML
huffman
parents: 35481
diff changeset
   164
val pat_rews = #pat_rews result;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   165
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   166
(* ----- theorems concerning the isomorphism -------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   167
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   168
val pg = pg' thy;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   169
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   170
val abs_strict = ax_rep_iso RS (allI RS retraction_strict);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   171
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
   172
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
   173
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   174
(* ----- theorems concerning one induction step ----------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   175
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   176
local
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   177
  fun dc_take dn = %%:(dn^"_take");
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   178
  val dnames = map (fst o fst) eqs;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   179
  fun get_take_strict dn = PureThy.get_thm thy (dn ^ ".take_strict");
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   180
  val axs_take_strict = map get_take_strict dnames;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   181
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   182
  fun one_take_app (con, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   183
    let
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   184
      fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n";
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   185
      fun one_rhs arg =
33971
9c7fa7f76950 modernized structure Datatype_Aux
haftmann
parents: 33810
diff changeset
   186
          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
   187
          then Domain_Axioms.copy_of_dtyp map_tab
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   188
                 mk_take (dtyp_of arg) ` (%# arg)
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   189
          else (%# arg);
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   190
      val lhs = (dc_take dname $ (%%:"Suc" $ %:"n"))`(con_app con args);
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   191
      val rhs = con_app2 con one_rhs args;
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   192
      val goal = mk_trp (lhs === rhs);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   193
      val rules = [ax_take_Suc, ax_abs_iso, @{thm cfcomp2}];
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   194
      val rules2 = @{thms take_con_rules ID1} @ axs_take_strict;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   195
      val tacs =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   196
          [simp_tac (HOL_basic_ss addsimps rules) 1,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   197
           asm_simp_tac (HOL_basic_ss addsimps rules2) 1];
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   198
    in pg con_appls goal (K tacs) end;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   199
  val take_apps = map (Drule.export_without_context o one_take_app) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   200
in
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   201
  val take_rews = ax_take_0 :: ax_take_strict :: take_apps;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   202
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   203
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   204
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   205
  thy
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   206
    |> Sign.add_path (Long_Name.base_name dname)
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   207
    |> snd o PureThy.add_thmss [
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   208
        ((Binding.name "iso_rews"  , iso_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   209
        ((Binding.name "exhaust"   , [exhaust]   ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   210
        ((Binding.name "casedist"  , [casedist]  ), [Induct.cases_type dname]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   211
        ((Binding.name "when_rews" , when_rews   ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   212
        ((Binding.name "compacts"  , con_compacts), [Simplifier.simp_add]),
33427
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   213
        ((Binding.name "con_rews"  , con_rews    ),
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   214
         [Simplifier.simp_add, Fixrec.fixrec_simp_add]),
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   215
        ((Binding.name "sel_rews"  , sel_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   216
        ((Binding.name "dis_rews"  , dis_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   217
        ((Binding.name "pat_rews"  , pat_rews    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   218
        ((Binding.name "dist_les"  , dist_les    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   219
        ((Binding.name "dist_eqs"  , dist_eqs    ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   220
        ((Binding.name "inverts"   , inverts     ), [Simplifier.simp_add]),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   221
        ((Binding.name "injects"   , injects     ), [Simplifier.simp_add]),
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   222
        ((Binding.name "take_rews" , take_rews   ), [Simplifier.simp_add]),
33427
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   223
        ((Binding.name "match_rews", mat_rews    ),
3182812d33ed domain package registers fixrec_simp lemmas
huffman
parents: 33396
diff changeset
   224
         [Simplifier.simp_add, Fixrec.fixrec_simp_add])]
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24503
diff changeset
   225
    |> Sign.parent_path
28536
8dccb6035d0f established canonical argument order
haftmann
parents: 27239
diff changeset
   226
    |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   227
        pat_rews @ dist_les @ dist_eqs)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   228
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   229
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   230
fun comp_theorems (comp_dnam, eqs: eq list) thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   231
let
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35512
diff changeset
   232
val map_tab = Domain_Take_Proofs.get_map_tab thy;
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   233
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   234
val dnames = map (fst o fst) eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   235
val conss  = map  snd        eqs;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28536
diff changeset
   236
val comp_dname = Sign.full_bname thy comp_dnam;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   237
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   238
val _ = message ("Proving induction properties of domain "^comp_dname^" ...");
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   239
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   240
(* ----- define bisimulation predicate -------------------------------------- *)
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   241
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   242
local
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   243
  open HOLCF_Library
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   244
  val dtypes  = map (Type o fst) eqs;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   245
  val relprod = mk_tupleT (map (fn tp => tp --> tp --> boolT) dtypes);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   246
  val bisim_bind = Binding.name (comp_dnam ^ "_bisim");
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   247
  val bisim_type = relprod --> boolT;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   248
in
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   249
  val (bisim_const, thy) =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   250
      Sign.declare_const ((bisim_bind, bisim_type), NoSyn) thy;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   251
end;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   252
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   253
local
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   254
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   255
  fun legacy_infer_term thy t =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   256
      singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   257
  fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain propT t);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   258
  fun infer_props thy = map (apsnd (legacy_infer_prop thy));
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   259
  fun add_defs_i x = PureThy.add_defs false (map Thm.no_attributes x);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   260
  fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   261
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   262
  val comp_dname = Sign.full_bname thy comp_dnam;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   263
  val dnames = map (fst o fst) eqs;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   264
  val x_name = idx_name dnames "x"; 
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   265
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   266
  fun one_con (con, args) =
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   267
    let
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   268
      val nonrec_args = filter_out is_rec args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   269
      val    rec_args = filter is_rec args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   270
      val    recs_cnt = length rec_args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   271
      val allargs     = nonrec_args @ rec_args
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   272
                        @ map (upd_vname (fn s=> s^"'")) rec_args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   273
      val allvns      = map vname allargs;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   274
      fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   275
      val vns1        = map (vname_arg "" ) args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   276
      val vns2        = map (vname_arg "'") args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   277
      val allargs_cnt = length nonrec_args + 2*recs_cnt;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   278
      val rec_idxs    = (recs_cnt-1) downto 0;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   279
      val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg)
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   280
                                             (allargs~~((allargs_cnt-1) downto 0)));
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   281
      fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ 
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   282
                              Bound (2*recs_cnt-i) $ Bound (recs_cnt-i);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   283
      val capps =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   284
          List.foldr
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   285
            mk_conj
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   286
            (mk_conj(
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   287
             Bound(allargs_cnt+1)===list_ccomb(%%:con,map (bound_arg allvns) vns1),
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   288
             Bound(allargs_cnt+0)===list_ccomb(%%:con,map (bound_arg allvns) vns2)))
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   289
            (mapn rel_app 1 rec_args);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   290
    in
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   291
      List.foldr
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   292
        mk_ex
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   293
        (Library.foldr mk_conj
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   294
                       (map (defined o Bound) nonlazy_idxs,capps)) allvns
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   295
    end;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   296
  fun one_comp n (_,cons) =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   297
      mk_all (x_name(n+1),
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   298
      mk_all (x_name(n+1)^"'",
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   299
      mk_imp (proj (Bound 2) eqs n $ Bound 1 $ Bound 0,
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   300
      foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU)
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   301
                      ::map one_con cons))));
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   302
  val bisim_eqn =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   303
      %%:(comp_dname^"_bisim") ==
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   304
         mk_lam("R", foldr1 mk_conj (mapn one_comp 0 eqs));
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   305
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   306
in
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   307
  val ([ax_bisim_def], thy) =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   308
      thy
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   309
        |> Sign.add_path comp_dnam
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   310
        |> add_defs_infer [(Binding.name "bisim_def", bisim_eqn)]
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   311
        ||> Sign.parent_path;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   312
end; (* local *)
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   313
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   314
val pg = pg' thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   315
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   316
(* ----- getting the composite axiom and definitions ------------------------ *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   317
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   318
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   319
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   320
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   321
  val axs_take_def   = map (ga "take_def"  ) dnames;
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   322
  val axs_chain_take = map (ga "chain_take") dnames;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   323
  val axs_lub_take   = map (ga "lub_take"  ) dnames;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   324
  val axs_finite_def = map (ga "finite_def") dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   325
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   326
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   327
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   328
  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
   329
  fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   330
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   331
  val cases = map (gt  "casedist" ) dnames;
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   332
  val con_rews  = maps (gts "con_rews" ) dnames;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   333
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   334
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   335
fun dc_take dn = %%:(dn^"_take");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   336
val x_name = idx_name dnames "x"; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   337
val P_name = idx_name dnames "P";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   338
val n_eqs = length eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   339
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   340
(* ----- theorems concerning finite approximation and finite induction ------ *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   341
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   342
val take_rews =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   343
    maps (fn dn => PureThy.get_thms thy (dn ^ ".take_rews")) dnames;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   344
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   345
local
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   346
  fun one_con p (con, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   347
    let
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
   348
      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
   349
      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
   350
      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
   351
      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
   352
      val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args);
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   353
      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
   354
      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
   355
    in Library.foldr mk_All (vns, t3) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   356
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   357
  fun one_eq ((p, cons), concl) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   358
    mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl);
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 ind_term concf = Library.foldr one_eq
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   361
    (mapn (fn n => fn x => (P_name n, x)) 1 conss,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   362
     mk_trp (foldr1 mk_conj (mapn concf 1 dnames)));
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   363
  val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   364
  fun quant_tac ctxt i = EVERY
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   365
    (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
   366
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   367
  fun ind_prems_tac prems = EVERY
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   368
    (maps (fn cons =>
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   369
      (resolve_tac prems 1 ::
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   370
        maps (fn (_,args) => 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   371
          resolve_tac prems 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   372
          map (K(atac 1)) (nonlazy args) @
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   373
          map (K(atac 1)) (filter is_rec args))
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   374
        cons))
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   375
      conss);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   376
  local 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   377
    (* check whether every/exists constructor of the n-th part of the equation:
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   378
       it has a possibly indirectly recursive argument that isn't/is possibly 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   379
       indirectly lazy *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   380
    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
   381
          is_rec arg andalso not(rec_of arg mem ns) andalso
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   382
          ((rec_of arg =  n andalso nfn(lazy_rec orelse is_lazy arg)) orelse 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   383
            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
   384
              (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg))))
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   385
          ) o snd) cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   386
    fun all_rec_to ns  = rec_to forall not all_rec_to  ns;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   387
    fun warn (n,cons) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   388
      if all_rec_to [] false (n,cons)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   389
      then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   390
      else false;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   391
    fun lazy_rec_to ns = rec_to exists I  lazy_rec_to ns;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   392
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   393
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   394
    val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   395
    val is_emptys = map warn n__eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   396
    val is_finite = forall (not o lazy_rec_to [] false) n__eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   397
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   398
in (* local *)
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   399
  val _ = trace " Proving finite_ind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   400
  val finite_ind =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   401
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   402
      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
   403
      val goal = ind_term concf;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   404
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   405
      fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   406
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   407
          val tacs1 = [
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   408
            quant_tac context 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   409
            simp_tac HOL_ss 1,
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   410
            InductTacs.induct_tac context [[SOME "n"]] 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   411
            simp_tac (take_ss addsimps prems) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   412
            TRY (safe_tac HOL_cs)];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   413
          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
   414
                        (* FIXME! case_UU_tac *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   415
            case_UU_tac context (prems @ con_rews) 1
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   416
              (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg);
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   417
          fun con_tacs (con, args) = 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   418
            asm_simp_tac take_ss 1 ::
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   419
            map arg_tac (filter is_nonlazy_rec args) @
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   420
            [resolve_tac prems 1] @
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   421
            map (K (atac 1)) (nonlazy args) @
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 32957
diff changeset
   422
            map (K (etac spec 1)) (filter is_rec args);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   423
          fun cases_tacs (cons, cases) =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   424
            res_inst_tac context [(("x", 0), "x")] cases 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   425
            asm_simp_tac (take_ss addsimps prems) 1 ::
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   426
            maps con_tacs cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   427
        in
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   428
          tacs1 @ maps cases_tacs (conss ~~ cases)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   429
        end;
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   430
    in pg'' thy [] goal tacf
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   431
       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
   432
    end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   433
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   434
  val _ = trace " Proving take_lemmas...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   435
  val take_lemmas =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   436
    let
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   437
      fun take_lemma (ax_chain_take, ax_lub_take) =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   438
        @{thm lub_ID_take_lemma} OF [ax_chain_take, ax_lub_take];
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   439
    in map take_lemma (axs_chain_take ~~ axs_lub_take) end;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   440
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   441
  val axs_reach =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   442
    let
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   443
      fun reach (ax_chain_take, ax_lub_take) =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   444
        @{thm lub_ID_reach} OF [ax_chain_take, ax_lub_take];
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   445
    in map reach (axs_chain_take ~~ axs_lub_take) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   446
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   447
(* ----- theorems concerning finiteness and induction ----------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   448
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   449
  val global_ctxt = ProofContext.init thy;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   450
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   451
  val _ = trace " Proving finites, ind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   452
  val (finites, ind) =
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   453
  (
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   454
    if is_finite
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   455
    then (* finite case *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   456
      let 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   457
        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
   458
        fun dname_lemma dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   459
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   460
            val prem1 = mk_trp (defined (%:"x"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   461
            val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   462
            val prem2 = mk_trp (mk_disj (disj1, take_enough dn));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   463
            val concl = mk_trp (take_enough dn);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   464
            val goal = prem1 ===> prem2 ===> concl;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   465
            val tacs = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   466
              etac disjE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   467
              etac notE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   468
              resolve_tac take_lemmas 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   469
              asm_simp_tac take_ss 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   470
              atac 1];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   471
          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
   472
        val _ = trace " Proving finite_lemmas1a";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   473
        val finite_lemmas1a = map dname_lemma dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   474
 
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   475
        val _ = trace " Proving finite_lemma1b";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   476
        val finite_lemma1b =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   477
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   478
            fun mk_eqn n ((dn, args), _) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   479
              let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   480
                val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   481
                val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   482
              in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   483
                mk_constrainall
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   484
                  (x_name n, Type (dn,args), mk_disj (disj1, disj2))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   485
              end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   486
            val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   487
              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
   488
            fun arg_tacs ctxt vn = [
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   489
              eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   490
              etac disjE 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   491
              asm_simp_tac (HOL_ss addsimps con_rews) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   492
              asm_simp_tac take_ss 1];
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   493
            fun con_tacs ctxt (con, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   494
              asm_simp_tac take_ss 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   495
              maps (arg_tacs ctxt) (nonlazy_rec args);
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   496
            fun foo_tacs ctxt n (cons, cases) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   497
              simp_tac take_ss 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   498
              rtac allI 1 ::
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   499
              res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   500
              asm_simp_tac take_ss 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   501
              maps (con_tacs ctxt) cons;
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   502
            fun tacs ctxt =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   503
              rtac allI 1 ::
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   504
              InductTacs.induct_tac ctxt [[SOME "n"]] 1 ::
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   505
              simp_tac take_ss 1 ::
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   506
              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
   507
              flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   508
          in pg [] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   509
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   510
        fun one_finite (dn, l1b) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   511
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   512
            val goal = mk_trp (%%:(dn^"_finite") $ %:"x");
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   513
            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
   514
                        (* FIXME! case_UU_tac *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   515
              case_UU_tac ctxt take_rews 1 "x",
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   516
              eresolve_tac finite_lemmas1a 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   517
              step_tac HOL_cs 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   518
              step_tac HOL_cs 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   519
              cut_facts_tac [l1b] 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   520
              fast_tac HOL_cs 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   521
          in pg axs_finite_def goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   522
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   523
        val _ = trace " Proving finites";
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   524
        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
   525
        val _ = trace " Proving ind";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   526
        val ind =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   527
          let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   528
            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
   529
            fun tacf {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   530
              let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   531
                fun finite_tacs (finite, fin_ind) = [
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   532
                  rtac(rewrite_rule axs_finite_def finite RS exE)1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   533
                  etac subst 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   534
                  rtac fin_ind 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   535
                  ind_prems_tac prems];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   536
              in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   537
                TRY (safe_tac HOL_cs) ::
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   538
                maps finite_tacs (finites ~~ atomize global_ctxt finite_ind)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   539
              end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   540
          in pg'' thy [] (ind_term concf) tacf end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   541
      in (finites, ind) end (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   542
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   543
    else (* infinite case *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   544
      let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   545
        fun one_finite n dn =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   546
          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
   547
        val finites = mapn one_finite 1 dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   548
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   549
        val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   550
          let
26012
f6917792f8a4 new term-building combinators
huffman
parents: 25895
diff changeset
   551
            fun one_adm n _ = mk_trp (mk_adm (%:(P_name n)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   552
            fun concf n dn = %:(P_name n) $ %:(x_name n);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   553
          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
   554
        val cont_rules =
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   555
            [cont_id, cont_const, cont2cont_Rep_CFun,
45c5c3c51918 domain package no longer uses cfst/csnd/cpair
huffman
parents: 33317
diff changeset
   556
             cont2cont_fst, cont2cont_snd];
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   557
        val subgoal =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   558
          let fun p n dn = %:(P_name n) $ (dc_take dn $ Bound 0 `%(x_name n));
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   559
          in mk_trp (mk_all ("n", foldr1 mk_conj (mapn p 1 dnames))) end;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   560
        val subgoal' = legacy_infer_term thy subgoal;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   561
        fun tacf {prems, context} =
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   562
          let
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   563
            val subtac =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   564
                EVERY [rtac allI 1, rtac finite_ind 1, ind_prems_tac prems];
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   565
            val subthm = Goal.prove context [] [] subgoal' (K subtac);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   566
          in
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   567
            map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   568
            cut_facts_tac (subthm :: take (length dnames) prems) 1,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   569
            REPEAT (rtac @{thm conjI} 1 ORELSE
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   570
                    EVERY [etac @{thm admD [OF _ ch2ch_Rep_CFunL]} 1,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   571
                           resolve_tac axs_chain_take 1,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   572
                           asm_simp_tac HOL_basic_ss 1])
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   573
            ]
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   574
          end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   575
        val ind = (pg'' thy [] goal tacf
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   576
          handle ERROR _ =>
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   577
            (warning "Cannot prove infinite induction rule"; TrueI)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   578
                  );
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   579
      in (finites, ind) end
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   580
  )
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   581
      handle THM _ =>
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   582
             (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
   583
           | ERROR _ =>
33810
38375b16ffd9 nicer warning message for indirect-recursive domain definitions
huffman
parents: 33801
diff changeset
   584
             (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
   585
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   586
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   587
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   588
(* ----- theorem concerning coinduction ------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   589
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   590
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   591
  val xs = mapn (fn n => K (x_name n)) 1 dnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   592
  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   593
  val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   594
  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
   595
  val _ = trace " Proving coind_lemma...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   596
  val coind_lemma =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   597
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   598
      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
   599
      fun mk_eqn n dn =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   600
        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   601
        (dc_take dn $ %:"n" ` bnd_arg n 1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   602
      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   603
      val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   604
        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   605
          Library.foldr mk_all2 (xs,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   606
            Library.foldr mk_imp (mapn mk_prj 0 dnames,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   607
              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   608
      fun x_tacs ctxt n x = [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   609
        rotate_tac (n+1) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   610
        etac all2E 1,
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27232
diff changeset
   611
        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
   612
        TRY (safe_tac HOL_cs),
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   613
        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   614
      fun tacs ctxt = [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   615
        rtac impI 1,
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   616
        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   617
        simp_tac take_ss 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   618
        safe_tac HOL_cs] @
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   619
        flat (mapn (x_tacs ctxt) 0 xs);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   620
    in pg [ax_bisim_def] goal tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   621
in
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   622
  val _ = trace " Proving coind...";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   623
  val coind = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   624
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   625
      fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   626
      fun mk_eqn x = %:x === %:(x^"'");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   627
      val goal =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   628
        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   629
          Logic.list_implies (mapn mk_prj 0 xs,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   630
            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   631
      val tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   632
        TRY (safe_tac HOL_cs) ::
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   633
        maps (fn take_lemma => [
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   634
          rtac take_lemma 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   635
          cut_facts_tac [coind_lemma] 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   636
          fast_tac HOL_cs 1])
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26012
diff changeset
   637
        take_lemmas;
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
   638
    in pg [] goal (K tacs) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   639
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   640
32172
c4e55f30d527 renamed functor ProjectRuleFun to Project_Rule;
wenzelm
parents: 32149
diff changeset
   641
val inducts = Project_Rule.projections (ProofContext.init thy) ind;
30829
d64a293f23ba domain package registers induction rules
huffman
parents: 30807
diff changeset
   642
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
   643
val induct_failed = (Thm.prop_of ind = Thm.prop_of TrueI);
30829
d64a293f23ba domain package registers induction rules
huffman
parents: 30807
diff changeset
   644
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24503
diff changeset
   645
in thy |> Sign.add_path comp_dnam
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   646
       |> snd o PureThy.add_thmss [
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   647
           ((Binding.name "take_lemmas", take_lemmas ), []),
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   648
           ((Binding.name "reach"      , axs_reach   ), []),
31004
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   649
           ((Binding.name "finites"    , finites     ), []),
ac7e90792089 declare take_rews as simp rules
huffman
parents: 30913
diff changeset
   650
           ((Binding.name "finite_ind" , [finite_ind]), []),
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   651
           ((Binding.name "ind"        , [ind]       ), []),
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   652
           ((Binding.name "coind"      , [coind]     ), [])]
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   653
       |> (if induct_failed then I
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31160
diff changeset
   654
           else snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts)))
28536
8dccb6035d0f established canonical argument order
haftmann
parents: 27239
diff changeset
   655
       |> Sign.parent_path |> pair take_rews
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   656
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   657
end; (* struct *)