src/HOLCF/Tools/Domain/domain_theorems.ML
author haftmann
Wed, 05 May 2010 18:25:34 +0200
changeset 36692 54b64d4ad524
parent 36610 bafd82950e24
child 36837 4d1dd57103b9
permissions -rw-r--r--
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
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:
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
    13
      Domain_Library.eq * Domain_Library.eq list ->
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
    14
      binding ->
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
    15
      (binding * (bool * binding option * typ) list * mixfix) list ->
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
    16
      Domain_Take_Proofs.iso_info ->
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
    17
      Domain_Take_Proofs.take_induct_info ->
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
    18
      theory -> thm list * theory;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
    19
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    20
  val comp_theorems :
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
    21
      binding * Domain_Library.eq list ->
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    22
      Domain_Take_Proofs.take_induct_info ->
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    23
      theory -> thm list * theory
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    24
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    25
  val quiet_mode: bool Unsynchronized.ref;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    26
  val trace_domain: bool Unsynchronized.ref;
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    27
end;
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    28
31023
d027411c9a38 use opaque ascription for all HOLCF code
huffman
parents: 31005
diff changeset
    29
structure Domain_Theorems :> DOMAIN_THEOREMS =
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    30
struct
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    31
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    32
val quiet_mode = Unsynchronized.ref false;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    33
val trace_domain = Unsynchronized.ref false;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    34
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    35
fun message s = if !quiet_mode then () else writeln s;
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    36
fun trace s = if !trace_domain then tracing s else ();
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    37
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
open Domain_Library;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    39
infixr 0 ===>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
infixr 0 ==>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    41
infix 0 == ; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    42
infix 1 ===;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    43
infix 1 ~= ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
infix 1 <<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    45
infix 1 ~<<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    46
infix 9 `   ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    47
infix 9 `% ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    48
infix 9 `%%;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    49
infixr 9 oo;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    50
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    51
(* ----- general proof facilities ------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    52
35800
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    53
local
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    54
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    55
fun map_typ f g (Type (c, Ts)) = Type (g c, map (map_typ f g) Ts)
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    56
  | map_typ f _ (TFree (x, S)) = TFree (x, map f S)
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    57
  | map_typ f _ (TVar (xi, S)) = TVar (xi, map f S);
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    58
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    59
fun map_term f g h (Const (c, T)) = Const (h c, map_typ f g T)
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    60
  | map_term f g _ (Free (x, T)) = Free (x, map_typ f g T)
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    61
  | map_term f g _ (Var (xi, T)) = Var (xi, map_typ f g T)
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    62
  | map_term _ _ _ (t as Bound _) = t
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    63
  | map_term f g h (Abs (x, T, t)) = Abs (x, map_typ f g T, map_term f g h t)
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    64
  | map_term f g h (t $ u) = map_term f g h t $ map_term f g h u;
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    65
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    66
in
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    67
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    68
fun intern_term thy =
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    69
  map_term (Sign.intern_class thy) (Sign.intern_type thy) (Sign.intern_const thy);
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    70
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    71
end;
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    72
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    73
fun legacy_infer_term thy t =
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36160
diff changeset
    74
  let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init_global thy)
35800
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    75
  in singleton (Syntax.check_terms ctxt) (intern_term thy t) end;
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    76
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
fun pg'' thy defs t tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    78
  let
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    79
    val t' = legacy_infer_term thy t;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
    val asms = Logic.strip_imp_prems t';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    81
    val prop = Logic.strip_imp_concl t';
26711
3a478bfa1650 prove_global: pass context;
wenzelm
parents: 26343
diff changeset
    82
    fun tac {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
      rewrite_goals_tac defs THEN
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    84
      EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
  in Goal.prove_global thy [] asms prop tac end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
fun pg' thy defs t tacsf =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
  let
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    89
    fun tacs {prems, context} =
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    90
      if null prems then tacsf context
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    91
      else cut_facts_tac prems 1 :: tacsf context;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    92
  in pg'' thy defs t tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
    94
(* FIXME!!!!!!!!! *)
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 35288
diff changeset
    95
(* 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
    96
(* 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
    97
fun case_UU_tac ctxt rews i v =
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    98
  InductTacs.case_tac ctxt (v^"=UU") i THEN
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
  asm_simp_tac (HOLCF_ss addsimps rews) i;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   100
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   101
(* ----- general proofs ----------------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   102
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
   103
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
   104
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   105
fun theorems
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   106
    (((dname, _), cons) : eq, eqs : eq list)
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   107
    (dbind : binding)
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   108
    (spec : (binding * (bool * binding option * typ) list * mixfix) list)
35558
bb088a6fafbc add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
huffman
parents: 35557
diff changeset
   109
    (iso_info : Domain_Take_Proofs.iso_info)
35775
9b7e2e17be69 pass take_info as argument to Domain_Theorems.theorems
huffman
parents: 35774
diff changeset
   110
    (take_info : Domain_Take_Proofs.take_induct_info)
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   111
    (thy : theory) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   112
let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   113
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
   114
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
   115
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
   116
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   117
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
(* ----- getting the axioms and definitions --------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   119
35558
bb088a6fafbc add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
huffman
parents: 35557
diff changeset
   120
val ax_abs_iso = #abs_inverse iso_info;
bb088a6fafbc add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
huffman
parents: 35557
diff changeset
   121
val ax_rep_iso = #rep_inverse iso_info;
bb088a6fafbc add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
huffman
parents: 35557
diff changeset
   122
bb088a6fafbc add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
huffman
parents: 35557
diff changeset
   123
val abs_const = #abs_const iso_info;
bb088a6fafbc add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
huffman
parents: 35557
diff changeset
   124
val rep_const = #rep_const iso_info;
bb088a6fafbc add_axioms returns an iso_info; add_theorems takes an iso_info as an argument
huffman
parents: 35557
diff changeset
   125
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   126
local
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26342
diff changeset
   127
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
in
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   129
  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
   130
  val ax_take_strict = ga "take_strict" dname;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   131
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   132
35775
9b7e2e17be69 pass take_info as argument to Domain_Theorems.theorems
huffman
parents: 35774
diff changeset
   133
val {take_Suc_thms, deflation_take_thms, ...} = take_info;
9b7e2e17be69 pass take_info as argument to Domain_Theorems.theorems
huffman
parents: 35774
diff changeset
   134
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   135
(* ----- define constructors ------------------------------------------------ *)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   136
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   137
val (result, thy) =
35777
bcc77916b7b9 pass binding as argument to add_domain_constructors; proper binding for case combinator
huffman
parents: 35776
diff changeset
   138
    Domain_Constructors.add_domain_constructors dbind spec iso_info thy;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   139
35451
a726a033b313 don't bother returning con_defs
huffman
parents: 35448
diff changeset
   140
val con_appls = #con_betas result;
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   141
val {nchotomy, exhaust, ...} = result;
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   142
val {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
   143
val {sel_rews, ...} = result;
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   144
val when_rews = #cases result;
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   145
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
   146
val dis_rews = #dis_rews result;
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35464
diff changeset
   147
val mat_rews = #match_rews result;
35482
d756837b708d move proofs of pat_rews to domain_constructors.ML
huffman
parents: 35481
diff changeset
   148
val pat_rews = #pat_rews result;
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   149
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   150
(* ----- theorems concerning the isomorphism -------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   151
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   152
val pg = pg' thy;
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   153
35560
d607ea103dcb remove unnecessary theorem references
huffman
parents: 35559
diff changeset
   154
val retraction_strict = @{thm retraction_strict};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   155
val abs_strict = ax_rep_iso RS (allI RS retraction_strict);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   156
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
   157
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
   158
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   159
(* ----- theorems concerning one induction step ----------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   160
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   161
local
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   162
  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
   163
  val dnames = map (fst o fst) eqs;
35523
cc57f4a274a3 fix proof script for take_apps so it works with indirect recursion
huffman
parents: 35521
diff changeset
   164
  val deflation_thms = Domain_Take_Proofs.get_deflation_thms thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   165
35559
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   166
  fun copy_of_dtyp tab r dt =
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   167
      if Datatype_Aux.is_rec_type dt then copy tab r dt else ID
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   168
  and copy tab r (Datatype_Aux.DtRec i) = r i
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   169
    | copy tab r (Datatype_Aux.DtTFree a) = ID
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   170
    | copy tab r (Datatype_Aux.DtType (c, ds)) =
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   171
      case Symtab.lookup tab c of
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   172
        SOME f => list_ccomb (%%:f, map (copy_of_dtyp tab r) ds)
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   173
      | NONE => (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID);
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   174
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   175
  fun one_take_app (con, args) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   176
    let
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   177
      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
   178
      fun one_rhs arg =
33971
9c7fa7f76950 modernized structure Datatype_Aux
haftmann
parents: 33810
diff changeset
   179
          if Datatype_Aux.is_rec_type (dtyp_of arg)
35559
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
   180
          then 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
   181
                 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
   182
          else (%# arg);
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   183
      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
   184
      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
   185
      val goal = mk_trp (lhs === rhs);
35590
f638444c9667 fix proof script so 'domain foo = Foo foo' works
huffman
parents: 35585
diff changeset
   186
      val rules =
36160
f84fa49a0b69 add rule deflation_ID to proof script for take + constructor rules
huffman
parents: 35800
diff changeset
   187
          [ax_abs_iso] @ @{thms take_con_rules}
35775
9b7e2e17be69 pass take_info as argument to Domain_Theorems.theorems
huffman
parents: 35774
diff changeset
   188
          @ take_Suc_thms @ deflation_thms @ deflation_take_thms;
35590
f638444c9667 fix proof script so 'domain foo = Foo foo' works
huffman
parents: 35585
diff changeset
   189
      val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1];
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   190
    in pg con_appls goal (K tacs) end;
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35528
diff changeset
   191
  val take_apps = map one_take_app cons;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   192
in
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   193
  val take_rews = ax_take_0 :: ax_take_strict :: take_apps;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   194
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   195
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   196
val case_ns =
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   197
    "bottom" :: map (fn (b,_,_) => Binding.name_of b) spec;
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   198
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   199
fun qualified name = Binding.qualified true name dbind;
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   200
val simp = Simplifier.simp_add;
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   201
val fixrec_simp = Fixrec.fixrec_simp_add;
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   202
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   203
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   204
  thy
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   205
  |> PureThy.add_thmss [
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   206
     ((qualified "iso_rews"  , iso_rews    ), [simp]),
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   207
     ((qualified "nchotomy"  , [nchotomy]  ), []),
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   208
     ((qualified "exhaust"   , [exhaust]   ),
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   209
      [Rule_Cases.case_names case_ns, Induct.cases_type dname]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   210
     ((qualified "when_rews" , when_rews   ), [simp]),
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   211
     ((qualified "compacts"  , compacts    ), [simp]),
35776
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   212
     ((qualified "con_rews"  , con_rews    ), [simp, fixrec_simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   213
     ((qualified "sel_rews"  , sel_rews    ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   214
     ((qualified "dis_rews"  , dis_rews    ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   215
     ((qualified "pat_rews"  , pat_rews    ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   216
     ((qualified "dist_les"  , dist_les    ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   217
     ((qualified "dist_eqs"  , dist_eqs    ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   218
     ((qualified "inverts"   , inverts     ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   219
     ((qualified "injects"   , injects     ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   220
     ((qualified "take_rews" , take_rews   ), [simp]),
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   221
     ((qualified "match_rews", mat_rews    ), [simp, fixrec_simp])]
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   222
  |> snd
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   223
  |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
b0bc15d8ad3b pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents: 35775
diff changeset
   224
      pat_rews @ dist_les @ dist_eqs)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   225
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   226
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   227
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   228
(****************************** induction rules *******************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   229
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   230
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   231
fun prove_induction
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   232
    (comp_dbind : binding, eqs : eq list)
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   233
    (take_rews : thm list)
35659
a78bc1930a7a include take_info within take_induct_info type
huffman
parents: 35658
diff changeset
   234
    (take_info : Domain_Take_Proofs.take_induct_info)
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   235
    (thy : theory) =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   236
let
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   237
  val comp_dname = Sign.full_name thy comp_dbind;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   238
  val dnames = map (fst o fst) eqs;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   239
  val conss  = map  snd        eqs;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   240
  fun dc_take dn = %%:(dn^"_take");
35662
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   241
  val x_name = idx_name dnames "x";
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   242
  val P_name = idx_name dnames "P";
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   243
  val pg = pg' thy;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   244
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   245
  local
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   246
    fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   247
    fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   248
  in
35597
e4331b99b03f introduce notion of 'decisive' deflations; use them to simplify proof script for rule 'finites'
huffman
parents: 35590
diff changeset
   249
    val axs_rep_iso = map (ga "rep_iso") dnames;
e4331b99b03f introduce notion of 'decisive' deflations; use them to simplify proof script for rule 'finites'
huffman
parents: 35590
diff changeset
   250
    val axs_abs_iso = map (ga "abs_iso") dnames;
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   251
    val exhausts = map (ga  "exhaust" ) dnames;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   252
    val con_rews  = maps (gts "con_rews" ) dnames;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   253
  end;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   254
35662
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   255
  val {take_consts, ...} = take_info;
35659
a78bc1930a7a include take_info within take_induct_info type
huffman
parents: 35658
diff changeset
   256
  val {take_0_thms, take_Suc_thms, chain_take_thms, ...} = take_info;
35660
8169419cd824 remove redundant function arguments
huffman
parents: 35659
diff changeset
   257
  val {lub_take_thms, finite_defs, reach_thms, ...} = take_info;
35661
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   258
  val {take_induct_thms, ...} = take_info;
35658
3d8da9fac424 pass take_info as an argument to comp_theorems
huffman
parents: 35657
diff changeset
   259
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   260
  fun one_con p (con, args) =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   261
    let
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   262
      val P_names = map P_name (1 upto (length dnames));
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   263
      val vns = Name.variant_list P_names (map vname args);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   264
      val nonlazy_vns = map snd (filter_out (is_lazy o fst) (args ~~ vns));
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   265
      fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   266
      val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   267
      val t2 = lift ind_hyp (filter is_rec args, t1);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   268
      val t3 = lift_defined (bound_arg vns) (nonlazy_vns, t2);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   269
    in Library.foldr mk_All (vns, t3) end;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   270
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   271
  fun one_eq ((p, cons), concl) =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   272
    mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   273
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   274
  fun ind_term concf = Library.foldr one_eq
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   275
    (mapn (fn n => fn x => (P_name n, x)) 1 conss,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   276
     mk_trp (foldr1 mk_conj (mapn concf 1 dnames)));
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   277
  val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   278
  fun quant_tac ctxt i = EVERY
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   279
    (mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   280
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   281
  fun ind_prems_tac prems = EVERY
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   282
    (maps (fn cons =>
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   283
      (resolve_tac prems 1 ::
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   284
        maps (fn (_,args) => 
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   285
          resolve_tac prems 1 ::
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   286
          map (K(atac 1)) (nonlazy args) @
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   287
          map (K(atac 1)) (filter is_rec args))
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   288
        cons))
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   289
      conss);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   290
  local 
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   291
    (* check whether every/exists constructor of the n-th part of the equation:
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   292
       it has a possibly indirectly recursive argument that isn't/is possibly 
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   293
       indirectly lazy *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   294
    fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => 
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   295
          is_rec arg andalso not (member (op =) ns (rec_of arg)) andalso
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   296
          ((rec_of arg =  n andalso nfn(lazy_rec orelse is_lazy arg)) orelse 
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   297
            rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) 
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   298
              (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg))))
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   299
          ) o snd) cons;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   300
    fun all_rec_to ns  = rec_to forall not all_rec_to  ns;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   301
    fun warn (n,cons) =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   302
      if all_rec_to [] false (n,cons)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   303
      then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   304
      else false;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   305
    fun lazy_rec_to ns = rec_to exists I  lazy_rec_to ns;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   306
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   307
  in
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   308
    val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   309
    val is_emptys = map warn n__eqs;
35661
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   310
    val is_finite = #is_finite take_info;
35601
50ba5010b876 print message when finiteness of domain definition is detected
huffman
parents: 35599
diff changeset
   311
    val _ = if is_finite
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   312
            then message ("Proving finiteness rule for domain "^comp_dname^" ...")
35601
50ba5010b876 print message when finiteness of domain definition is detected
huffman
parents: 35599
diff changeset
   313
            else ();
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   314
  end;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   315
  val _ = trace " Proving finite_ind...";
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   316
  val finite_ind =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   317
    let
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   318
      fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n));
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   319
      val goal = ind_term concf;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   320
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   321
      fun tacf {prems, context} =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   322
        let
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   323
          val tacs1 = [
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   324
            quant_tac context 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   325
            simp_tac HOL_ss 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   326
            InductTacs.induct_tac context [[SOME "n"]] 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   327
            simp_tac (take_ss addsimps prems) 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   328
            TRY (safe_tac HOL_cs)];
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   329
          fun arg_tac arg =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   330
                        (* FIXME! case_UU_tac *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   331
            case_UU_tac context (prems @ con_rews) 1
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   332
              (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   333
          fun con_tacs (con, args) = 
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   334
            asm_simp_tac take_ss 1 ::
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   335
            map arg_tac (filter is_nonlazy_rec args) @
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   336
            [resolve_tac prems 1] @
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   337
            map (K (atac 1)) (nonlazy args) @
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   338
            map (K (etac spec 1)) (filter is_rec args);
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   339
          fun cases_tacs (cons, exhaust) =
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   340
            res_inst_tac context [(("y", 0), "x")] exhaust 1 ::
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   341
            asm_simp_tac (take_ss addsimps prems) 1 ::
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   342
            maps con_tacs cons;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   343
        in
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   344
          tacs1 @ maps cases_tacs (conss ~~ exhausts)
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   345
        end;
35663
ada7bc39c6b1 remove unnecessary error handling code
huffman
parents: 35662
diff changeset
   346
    in pg'' thy [] goal tacf end;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   347
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   348
(* ----- theorems concerning finiteness and induction ----------------------- *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   349
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36160
diff changeset
   350
  val global_ctxt = ProofContext.init_global thy;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   351
35661
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   352
  val _ = trace " Proving ind...";
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   353
  val ind =
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   354
    if is_finite
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   355
    then (* finite case *)
35597
e4331b99b03f introduce notion of 'decisive' deflations; use them to simplify proof script for rule 'finites'
huffman
parents: 35590
diff changeset
   356
      let
35661
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   357
        fun concf n dn = %:(P_name n) $ %:(x_name n);
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   358
        fun tacf {prems, context} =
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   359
          let
35661
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   360
            fun finite_tacs (take_induct, fin_ind) = [
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   361
                rtac take_induct 1,
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   362
                rtac fin_ind 1,
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   363
                ind_prems_tac prems];
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   364
          in
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   365
            TRY (safe_tac HOL_cs) ::
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   366
            maps finite_tacs (take_induct_thms ~~ atomize global_ctxt finite_ind)
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   367
          end;
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35660
diff changeset
   368
      in pg'' thy [] (ind_term concf) tacf end
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   369
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   370
    else (* infinite case *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   371
      let
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   372
        val goal =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   373
          let
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   374
            fun one_adm n _ = mk_trp (mk_adm (%:(P_name n)));
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   375
            fun concf n dn = %:(P_name n) $ %:(x_name n);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   376
          in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   377
        val cont_rules =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   378
            @{thms cont_id cont_const cont2cont_Rep_CFun
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   379
                   cont2cont_fst cont2cont_snd};
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   380
        val subgoal =
35662
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   381
          let
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   382
            val Ts = map (Type o fst) eqs;
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   383
            val P_names = Datatype_Prop.indexify_names (map (K "P") dnames);
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   384
            val x_names = Datatype_Prop.indexify_names (map (K "x") dnames);
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   385
            val P_types = map (fn T => T --> HOLogic.boolT) Ts;
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   386
            val Ps = map Free (P_names ~~ P_types);
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   387
            val xs = map Free (x_names ~~ Ts);
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   388
            val n = Free ("n", HOLogic.natT);
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   389
            val goals =
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   390
                map (fn ((P,t),x) => P $ HOLCF_Library.mk_capply (t $ n, x))
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   391
                  (Ps ~~ take_consts ~~ xs);
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   392
          in
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   393
            HOLogic.mk_Trueprop
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   394
            (HOLogic.mk_all ("n", HOLogic.natT, foldr1 HOLogic.mk_conj goals))
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   395
          end;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   396
        fun tacf {prems, context} =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   397
          let
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   398
            val subtac =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   399
                EVERY [rtac allI 1, rtac finite_ind 1, ind_prems_tac prems];
35662
44d7aafdddb9 construct fully typed goal in proof of induction rule
huffman
parents: 35661
diff changeset
   400
            val subthm = Goal.prove context [] [] subgoal (K subtac);
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   401
          in
35660
8169419cd824 remove redundant function arguments
huffman
parents: 35659
diff changeset
   402
            map (fn ax_reach => rtac (ax_reach RS subst) 1) reach_thms @ [
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   403
            cut_facts_tac (subthm :: take (length dnames) prems) 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   404
            REPEAT (rtac @{thm conjI} 1 ORELSE
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   405
                    EVERY [etac @{thm admD [OF _ ch2ch_Rep_CFunL]} 1,
35659
a78bc1930a7a include take_info within take_induct_info type
huffman
parents: 35658
diff changeset
   406
                           resolve_tac chain_take_thms 1,
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   407
                           asm_simp_tac HOL_basic_ss 1])
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   408
            ]
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   409
          end;
35663
ada7bc39c6b1 remove unnecessary error handling code
huffman
parents: 35662
diff changeset
   410
      in pg'' thy [] goal tacf end;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   411
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   412
val case_ns =
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   413
  let
35782
8a314dd86714 add case name 'adm' for infinite induction rules
huffman
parents: 35781
diff changeset
   414
    val adms =
8a314dd86714 add case name 'adm' for infinite induction rules
huffman
parents: 35781
diff changeset
   415
        if is_finite then [] else
8a314dd86714 add case name 'adm' for infinite induction rules
huffman
parents: 35781
diff changeset
   416
        if length dnames = 1 then ["adm"] else
8a314dd86714 add case name 'adm' for infinite induction rules
huffman
parents: 35781
diff changeset
   417
        map (fn s => "adm_" ^ Long_Name.base_name s) dnames;
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   418
    val bottoms =
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   419
        if length dnames = 1 then ["bottom"] else
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   420
        map (fn s => "bottom_" ^ Long_Name.base_name s) dnames;
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   421
    fun one_eq bot (_,cons) =
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   422
          bot :: map (fn (c,_) => Long_Name.base_name c) cons;
35782
8a314dd86714 add case name 'adm' for infinite induction rules
huffman
parents: 35781
diff changeset
   423
  in adms @ flat (map2 one_eq bottoms eqs) end;
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   424
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36160
diff changeset
   425
val inducts = Project_Rule.projections (ProofContext.init_global thy) ind;
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   426
fun ind_rule (dname, rule) =
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   427
    ((Binding.empty, [rule]),
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   428
     [Rule_Cases.case_names case_ns, Induct.induct_type dname]);
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   429
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   430
in
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   431
  thy
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   432
  |> snd o PureThy.add_thmss [
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   433
     ((Binding.qualified true "finite_induct" comp_dbind, [finite_ind]), []),
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   434
     ((Binding.qualified true "induct"        comp_dbind, [ind]       ), [])]
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   435
  |> (snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts)))
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   436
end; (* prove_induction *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   437
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   438
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   439
(************************ bisimulation and coinduction ************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   440
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   441
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   442
fun prove_coinduction
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   443
    (comp_dbind : binding, eqs : eq list)
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   444
    (take_lemmas : thm list)
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   445
    (thy : theory) : theory =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   446
let
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   447
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   448
val dnames = map (fst o fst) eqs;
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   449
val comp_dname = Sign.full_name thy comp_dbind;
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   450
fun dc_take dn = %%:(dn^"_take");
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   451
val x_name = idx_name dnames "x"; 
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   452
val n_eqs = length eqs;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   453
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   454
val take_rews =
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   455
    maps (fn dn => PureThy.get_thms thy (dn ^ ".take_rews")) dnames;
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   456
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   457
(* ----- define bisimulation predicate -------------------------------------- *)
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   458
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   459
local
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   460
  open HOLCF_Library
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   461
  val dtypes  = map (Type o fst) eqs;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   462
  val relprod = mk_tupleT (map (fn tp => tp --> tp --> boolT) dtypes);
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   463
  val bisim_bind = Binding.suffix_name "_bisim" comp_dbind;
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   464
  val bisim_type = relprod --> boolT;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   465
in
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   466
  val (bisim_const, thy) =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   467
      Sign.declare_const ((bisim_bind, bisim_type), NoSyn) thy;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   468
end;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   469
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   470
local
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   471
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   472
  fun legacy_infer_term thy t =
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36160
diff changeset
   473
      singleton (Syntax.check_terms (ProofContext.init_global thy)) (intern_term thy t);
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   474
  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
   475
  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
   476
  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
   477
  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
   478
35521
47eec4da124a remove unused mixfix component from type cons
huffman
parents: 35514
diff changeset
   479
  fun one_con (con, args) =
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   480
    let
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   481
      val nonrec_args = filter_out is_rec args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   482
      val    rec_args = filter is_rec args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   483
      val    recs_cnt = length rec_args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   484
      val allargs     = nonrec_args @ rec_args
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   485
                        @ map (upd_vname (fn s=> s^"'")) rec_args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   486
      val allvns      = map vname allargs;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   487
      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
   488
      val vns1        = map (vname_arg "" ) args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   489
      val vns2        = map (vname_arg "'") args;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   490
      val allargs_cnt = length nonrec_args + 2*recs_cnt;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   491
      val rec_idxs    = (recs_cnt-1) downto 0;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   492
      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
   493
                                             (allargs~~((allargs_cnt-1) downto 0)));
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   494
      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
   495
                              Bound (2*recs_cnt-i) $ Bound (recs_cnt-i);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   496
      val capps =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   497
          List.foldr
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   498
            mk_conj
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   499
            (mk_conj(
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   500
             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
   501
             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
   502
            (mapn rel_app 1 rec_args);
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   503
    in
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   504
      List.foldr
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   505
        mk_ex
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   506
        (Library.foldr mk_conj
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   507
                       (map (defined o Bound) nonlazy_idxs,capps)) allvns
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   508
    end;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   509
  fun one_comp n (_,cons) =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   510
      mk_all (x_name(n+1),
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   511
      mk_all (x_name(n+1)^"'",
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   512
      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
   513
      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
   514
                      ::map one_con cons))));
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   515
  val bisim_eqn =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   516
      %%:(comp_dname^"_bisim") ==
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   517
         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
   518
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   519
in
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   520
  val (ax_bisim_def, thy) =
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   521
      yield_singleton add_defs_infer
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   522
        (Binding.qualified true "bisim_def" comp_dbind, bisim_eqn) thy;
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   523
end; (* local *)
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   524
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   525
(* ----- theorem concerning coinduction ------------------------------------- *)
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   526
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   527
local
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   528
  val pg = pg' thy;
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   529
  val xs = mapn (fn n => K (x_name n)) 1 dnames;
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   530
  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   531
  val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   532
  val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   533
  val _ = trace " Proving coind_lemma...";
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   534
  val coind_lemma =
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   535
    let
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   536
      fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1;
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   537
      fun mk_eqn n dn =
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   538
        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   539
        (dc_take dn $ %:"n" ` bnd_arg n 1);
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   540
      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   541
      val goal =
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   542
        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   543
          Library.foldr mk_all2 (xs,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   544
            Library.foldr mk_imp (mapn mk_prj 0 dnames,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   545
              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   546
      fun x_tacs ctxt n x = [
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   547
        rotate_tac (n+1) 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   548
        etac all2E 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   549
        eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   550
        TRY (safe_tac HOL_cs),
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   551
        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   552
      fun tacs ctxt = [
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   553
        rtac impI 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   554
        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   555
        simp_tac take_ss 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   556
        safe_tac HOL_cs] @
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   557
        flat (mapn (x_tacs ctxt) 0 xs);
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   558
    in pg [ax_bisim_def] goal tacs end;
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   559
in
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   560
  val _ = trace " Proving coind...";
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   561
  val coind = 
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   562
    let
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   563
      fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   564
      fun mk_eqn x = %:x === %:(x^"'");
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   565
      val goal =
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   566
        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   567
          Logic.list_implies (mapn mk_prj 0 xs,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   568
            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   569
      val tacs =
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   570
        TRY (safe_tac HOL_cs) ::
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   571
        maps (fn take_lemma => [
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   572
          rtac take_lemma 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   573
          cut_facts_tac [coind_lemma] 1,
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   574
          fast_tac HOL_cs 1])
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   575
        take_lemmas;
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   576
    in pg [] goal (K tacs) end;
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   577
end; (* local *)
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   578
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   579
in thy |> snd o PureThy.add_thmss
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   580
    [((Binding.qualified true "coinduct" comp_dbind, [coind]), [])]
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   581
end; (* let *)
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   582
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
   583
fun comp_theorems
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   584
    (comp_dbind : binding, eqs : eq list)
35659
a78bc1930a7a include take_info within take_induct_info type
huffman
parents: 35658
diff changeset
   585
    (take_info : Domain_Take_Proofs.take_induct_info)
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
   586
    (thy : theory) =
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   587
let
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   588
val map_tab = Domain_Take_Proofs.get_map_tab thy;
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   589
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   590
val dnames = map (fst o fst) eqs;
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   591
val comp_dname = Sign.full_name thy comp_dbind;
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   592
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   593
(* ----- getting the composite axiom and definitions ------------------------ *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   594
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   595
(* Test for indirect recursion *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   596
local
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   597
  fun indirect_arg arg =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   598
      rec_of arg = ~1 andalso Datatype_Aux.is_rec_type (dtyp_of arg);
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   599
  fun indirect_con (_, args) = exists indirect_arg args;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   600
  fun indirect_eq (_, cons) = exists indirect_con cons;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   601
in
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   602
  val is_indirect = exists indirect_eq eqs;
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   603
  val _ =
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   604
      if is_indirect
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   605
      then message "Indirect recursion detected, skipping proofs of (co)induction rules"
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   606
      else message ("Proving induction properties of domain "^comp_dname^" ...");
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   607
end;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   608
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   609
(* theorems about take *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   610
35659
a78bc1930a7a include take_info within take_induct_info type
huffman
parents: 35658
diff changeset
   611
val take_lemmas = #take_lemma_thms take_info;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   612
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   613
val take_rews =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   614
    maps (fn dn => PureThy.get_thms thy (dn ^ ".take_rews")) dnames;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   615
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   616
(* prove induction rules, unless definition is indirect recursive *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   617
val thy =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   618
    if is_indirect then thy else
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   619
    prove_induction (comp_dbind, eqs) take_rews take_info thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   620
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   621
val thy =
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   622
    if is_indirect then thy else
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   623
    prove_coinduction (comp_dbind, eqs) take_lemmas thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   624
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35630
diff changeset
   625
in
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35630
diff changeset
   626
  (take_rews, thy)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   627
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   628
end; (* struct *)