src/HOLCF/Tools/Domain/domain_theorems.ML
author huffman
Sat, 16 Oct 2010 15:26:30 -0700
changeset 40025 876689e6bbdf
parent 40023 a868e9d73031
child 40026 8f8f18a88685
permissions -rw-r--r--
reimplement proof automation for coinduct rules
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
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    12
  val comp_theorems :
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
    13
      binding * Domain_Library.eq list ->
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
    14
      binding list ->
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    15
      Domain_Take_Proofs.take_induct_info ->
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    16
      Domain_Constructors.constr_info list ->
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    17
      theory -> thm list * theory
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    18
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    19
  val quiet_mode: bool Unsynchronized.ref;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    20
  val trace_domain: bool Unsynchronized.ref;
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    21
end;
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    22
31023
d027411c9a38 use opaque ascription for all HOLCF code
huffman
parents: 31005
diff changeset
    23
structure Domain_Theorems :> DOMAIN_THEOREMS =
31005
e55eed7d9b55 add module signature for domain_theorems.ML
huffman
parents: 31004
diff changeset
    24
struct
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    25
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    26
val quiet_mode = Unsynchronized.ref false;
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32172
diff changeset
    27
val trace_domain = Unsynchronized.ref false;
29402
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    28
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    29
fun message s = if !quiet_mode then () else writeln s;
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    30
fun trace s = if !trace_domain then tracing s else ();
9610f3870d64 add tracing for domain package proofs
huffman
parents: 29064
diff changeset
    31
40022
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
    32
open HOLCF_Library;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
infixr 0 ===>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
infix 0 == ; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    35
infix 1 ===;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    36
infix 9 `   ;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    37
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
(* ----- general proof facilities ------------------------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    39
35800
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    40
local
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    41
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    42
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
    43
  | 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
    44
  | 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
    45
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    46
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
    47
  | 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
    48
  | 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
    49
  | 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
    50
  | 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
    51
  | 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
    52
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    53
in
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 intern_term thy =
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    56
  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
    57
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    58
end;
76b2a53a199d moved old Sign.intern_term to the place where it is still used;
wenzelm
parents: 35782
diff changeset
    59
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    60
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
    61
  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
    62
  in singleton (Syntax.check_terms ctxt) (intern_term thy t) end;
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    63
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    64
fun pg'' thy defs t tacs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
  let
24503
2439587f516b legacy_infer_term: ProofContext.mode_schematic;
wenzelm
parents: 23894
diff changeset
    66
    val t' = legacy_infer_term thy t;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    67
    val asms = Logic.strip_imp_prems t';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    68
    val prop = Logic.strip_imp_concl t';
26711
3a478bfa1650 prove_global: pass context;
wenzelm
parents: 26343
diff changeset
    69
    fun tac {prems, context} =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
      rewrite_goals_tac defs THEN
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    71
      EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    72
  in Goal.prove_global thy [] asms prop tac end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    73
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    74
fun pg' thy defs t tacsf =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
  let
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    76
    fun tacs {prems, context} =
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    77
      if null prems then tacsf context
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27153
diff changeset
    78
      else cut_facts_tac prems 1 :: tacsf context;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    79
  in pg'' thy defs t tacs end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
40013
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
    81
(******************************************************************************)
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
    82
(***************************** proofs about take ******************************)
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
    83
(******************************************************************************)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    84
40013
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
    85
fun take_theorems
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
    86
    (dbinds : binding list)
35775
9b7e2e17be69 pass take_info as argument to Domain_Theorems.theorems
huffman
parents: 35774
diff changeset
    87
    (take_info : Domain_Take_Proofs.take_induct_info)
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    88
    (constr_infos : Domain_Constructors.constr_info list)
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    89
    (thy : theory) : thm list list * theory =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
let
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    91
  val {take_consts, take_Suc_thms, deflation_take_thms, ...} = take_info;
35523
cc57f4a274a3 fix proof script for take_apps so it works with indirect recursion
huffman
parents: 35521
diff changeset
    92
  val deflation_thms = Domain_Take_Proofs.get_deflation_thms thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    94
  val n = Free ("n", @{typ nat});
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    95
  val n' = @{const Suc} $ n;
35559
119653afcd6e remove copy_of_dtyp from domain_axioms.ML
huffman
parents: 35558
diff changeset
    96
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    97
  local
40017
575d3aa1f3c5 include iso_info as part of constr_info type
huffman
parents: 40016
diff changeset
    98
    val newTs = map (#absT o #iso_info) constr_infos;
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
    99
    val subs = newTs ~~ map (fn t => t $ n) take_consts;
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   100
    fun is_ID (Const (c, _)) = (c = @{const_name ID})
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   101
      | is_ID _              = false;
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   102
  in
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   103
    fun map_of_arg v T =
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   104
      let val m = Domain_Take_Proofs.map_of_typ thy subs T;
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   105
      in if is_ID m then v else mk_capply (m, v) end;
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   106
  end
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   107
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   108
  fun prove_take_apps
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   109
      ((dbind, take_const), constr_info) thy =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   110
    let
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   111
      val {iso_info, con_specs, con_betas, ...} = constr_info;
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   112
      val {abs_inverse, ...} = iso_info;
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   113
      fun prove_take_app (con_const, args) =
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   114
        let
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   115
          val Ts = map snd args;
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   116
          val ns = Name.variant_list ["n"] (Datatype_Prop.make_tnames Ts);
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   117
          val vs = map Free (ns ~~ Ts);
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   118
          val lhs = mk_capply (take_const $ n', list_ccomb (con_const, vs));
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   119
          val rhs = list_ccomb (con_const, map2 map_of_arg vs Ts);
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   120
          val goal = mk_trp (mk_eq (lhs, rhs));
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   121
          val rules =
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   122
              [abs_inverse] @ con_betas @ @{thms take_con_rules}
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   123
              @ take_Suc_thms @ deflation_thms @ deflation_take_thms;
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   124
          val tac = simp_tac (HOL_basic_ss addsimps rules) 1;
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   125
        in
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   126
          Goal.prove_global thy [] [] goal (K tac)
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   127
        end;
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   128
      val take_apps = map prove_take_app con_specs;
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   129
    in
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   130
      yield_singleton Global_Theory.add_thmss
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   131
        ((Binding.qualified true "take_rews" dbind, take_apps),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   132
        [Simplifier.simp_add]) thy
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   133
    end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   134
in
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   135
  fold_map prove_take_apps
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   136
    (dbinds ~~ take_consts ~~ constr_infos) thy
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   137
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   138
40013
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
   139
(* ----- general proofs ----------------------------------------------------- *)
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
   140
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
   141
val all2E = @{lemma "!x y . P x y ==> (P x y ==> R) ==> R" by simp}
9db8fb58fddc add function take_theorems
huffman
parents: 39557
diff changeset
   142
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   143
val case_UU_allI =
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   144
    @{lemma "(!!x. x ~= UU ==> P x) ==> P UU ==> ALL x. P x" by metis};
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   145
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   146
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   147
(****************************** induction rules *******************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   148
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   149
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   150
fun prove_induction
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   151
    (comp_dbind : binding)
40018
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   152
    (constr_infos : Domain_Constructors.constr_info list)
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   153
    (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
   154
    (take_rews : thm list)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   155
    (thy : theory) =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   156
let
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   157
  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
   158
40018
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   159
  val iso_infos = map #iso_info constr_infos;
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   160
  val exhausts = map #exhaust constr_infos;
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   161
  val con_rews = maps #con_rews constr_infos;
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   162
  val {take_consts, take_induct_thms, ...} = take_info;
35658
3d8da9fac424 pass take_info as an argument to comp_theorems
huffman
parents: 35657
diff changeset
   163
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   164
  val newTs = map #absT iso_infos;
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   165
  val P_names = Datatype_Prop.indexify_names (map (K "P") newTs);
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   166
  val x_names = Datatype_Prop.indexify_names (map (K "x") newTs);
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   167
  val P_types = map (fn T => T --> HOLogic.boolT) newTs;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   168
  val Ps = map Free (P_names ~~ P_types);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   169
  val xs = map Free (x_names ~~ newTs);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   170
  val n = Free ("n", HOLogic.natT);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   171
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   172
  fun con_assm defined p (con, args) =
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   173
    let
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   174
      val Ts = map snd args;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   175
      val ns = Name.variant_list P_names (Datatype_Prop.make_tnames Ts);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   176
      val vs = map Free (ns ~~ Ts);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   177
      val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs));
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   178
      fun ind_hyp (v, T) t =
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   179
          case AList.lookup (op =) (newTs ~~ Ps) T of NONE => t
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   180
          | SOME p' => Logic.mk_implies (mk_trp (p' $ v), t);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   181
      val t1 = mk_trp (p $ list_ccomb (con, vs));
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   182
      val t2 = fold_rev ind_hyp (vs ~~ Ts) t1;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   183
      val t3 = Logic.list_implies (map (mk_trp o mk_defined) nonlazy, t2);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   184
    in fold_rev Logic.all vs (if defined then t3 else t2) end;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   185
  fun eq_assms ((p, T), cons) =
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   186
      mk_trp (p $ HOLCF_Library.mk_bottom T) :: map (con_assm true p) cons;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   187
  val assms = maps eq_assms (Ps ~~ newTs ~~ map #con_specs constr_infos);
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   188
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   189
  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
   190
  fun quant_tac ctxt i = EVERY
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   191
    (map (fn name => res_inst_tac ctxt [(("x", 0), name)] spec i) x_names);
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   192
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   193
  (* FIXME: move this message to domain_take_proofs.ML *)
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   194
  val is_finite = #is_finite take_info;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   195
  val _ = if is_finite
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   196
          then message ("Proving finiteness rule for domain "^comp_dname^" ...")
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   197
          else ();
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   198
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   199
  val _ = trace " Proving finite_ind...";
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   200
  val finite_ind =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   201
    let
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   202
      val concls =
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   203
          map (fn ((P, t), x) => P $ mk_capply (t $ n, x))
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   204
              (Ps ~~ take_consts ~~ xs);
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   205
      val goal = mk_trp (foldr1 mk_conj concls);
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   206
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   207
      fun tacf {prems, context} =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   208
        let
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   209
          (* Prove stronger prems, without definedness side conditions *)
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   210
          fun con_thm p (con, args) =
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   211
            let
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   212
              val subgoal = con_assm false p (con, args);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   213
              val rules = prems @ con_rews @ simp_thms;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   214
              val simplify = asm_simp_tac (HOL_basic_ss addsimps rules);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   215
              fun arg_tac (lazy, _) =
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   216
                  rtac (if lazy then allI else case_UU_allI) 1;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   217
              val tacs =
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   218
                  rewrite_goals_tac @{thms atomize_all atomize_imp} ::
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   219
                  map arg_tac args @
40022
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   220
                  [REPEAT (rtac impI 1), ALLGOALS simplify];
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   221
            in
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   222
              Goal.prove context [] [] subgoal (K (EVERY tacs))
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   223
            end;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   224
          fun eq_thms (p, cons) = map (con_thm p) cons;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   225
          val conss = map #con_specs constr_infos;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   226
          val prems' = maps eq_thms (Ps ~~ conss);
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   227
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   228
          val tacs1 = [
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   229
            quant_tac context 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   230
            simp_tac HOL_ss 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   231
            InductTacs.induct_tac context [[SOME "n"]] 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   232
            simp_tac (take_ss addsimps prems) 1,
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   233
            TRY (safe_tac HOL_cs)];
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   234
          fun con_tac _ = 
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   235
            asm_simp_tac take_ss 1 THEN
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   236
            (resolve_tac prems' THEN_ALL_NEW etac spec) 1;
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   237
          fun cases_tacs (cons, exhaust) =
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   238
            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
   239
            asm_simp_tac (take_ss addsimps prems) 1 ::
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   240
            map con_tac cons;
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   241
          val tacs = tacs1 @ maps cases_tacs (conss ~~ exhausts)
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   242
        in
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   243
          EVERY (map DETERM tacs)
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   244
        end;
40020
0cbb08bf18df rewrite proof automation for finite_ind; get rid of case_UU_tac
huffman
parents: 40019
diff changeset
   245
    in Goal.prove_global thy [] assms goal tacf end;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   246
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
   247
  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
   248
  val ind =
40022
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   249
    let
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   250
      val concls = map (op $) (Ps ~~ xs);
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   251
      val goal = mk_trp (foldr1 mk_conj concls);
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   252
      val adms = if is_finite then [] else map (mk_trp o mk_adm) Ps;
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   253
      fun tacf {prems, context} =
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   254
        let
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   255
          fun finite_tac (take_induct, fin_ind) =
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   256
              rtac take_induct 1 THEN
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   257
              (if is_finite then all_tac else resolve_tac prems 1) THEN
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   258
              (rtac fin_ind THEN_ALL_NEW solve_tac prems) 1;
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   259
          val fin_inds = Project_Rule.projections context finite_ind;
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   260
        in
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   261
          TRY (safe_tac HOL_cs) THEN
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   262
          EVERY (map finite_tac (take_induct_thms ~~ fin_inds))
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   263
        end;
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   264
    in Goal.prove_global thy [] (adms @ assms) goal tacf end
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   265
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   266
  (* case names for induction rules *)
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   267
  val dnames = map (fst o dest_Type) newTs;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   268
  val case_ns =
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   269
    let
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   270
      val adms =
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   271
          if is_finite then [] else
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   272
          if length dnames = 1 then ["adm"] else
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   273
          map (fn s => "adm_" ^ Long_Name.base_name s) dnames;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   274
      val bottoms =
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   275
          if length dnames = 1 then ["bottom"] else
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   276
          map (fn s => "bottom_" ^ Long_Name.base_name s) dnames;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   277
      fun one_eq bot constr_info =
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   278
        let fun name_of (c, args) = Long_Name.base_name (fst (dest_Const c));
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   279
        in bot :: map name_of (#con_specs constr_info) end;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   280
    in adms @ flat (map2 one_eq bottoms constr_infos) end;
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   281
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   282
  val inducts = Project_Rule.projections (ProofContext.init_global thy) ind;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   283
  fun ind_rule (dname, rule) =
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   284
      ((Binding.empty, [rule]),
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   285
       [Rule_Cases.case_names case_ns, Induct.induct_type dname]);
35630
8e562d56d404 add case_names attribute to casedist and ind rules
huffman
parents: 35601
diff changeset
   286
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   287
in
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   288
  thy
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39288
diff changeset
   289
  |> snd o Global_Theory.add_thmss [
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   290
     ((Binding.qualified true "finite_induct" comp_dbind, [finite_ind]), []),
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   291
     ((Binding.qualified true "induct"        comp_dbind, [ind]       ), [])]
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39288
diff changeset
   292
  |> (snd o Global_Theory.add_thmss (map ind_rule (dnames ~~ inducts)))
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   293
end; (* prove_induction *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   294
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   295
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   296
(************************ bisimulation and coinduction ************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   297
(******************************************************************************)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   298
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   299
fun prove_coinduction
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   300
    (comp_dbind : binding, dbinds : binding list)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   301
    (constr_infos : Domain_Constructors.constr_info list)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   302
    (take_info : Domain_Take_Proofs.take_induct_info)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   303
    (take_rews : thm list list)
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   304
    (thy : theory) : theory =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   305
let
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   306
  val comp_dname = Sign.full_name thy comp_dbind;
27232
7cd256da0a36 atomize: proper context;
wenzelm
parents: 27208
diff changeset
   307
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   308
  val iso_infos = map #iso_info constr_infos;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   309
  val newTs = map #absT iso_infos;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   310
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   311
  val {take_consts, take_0_thms, take_lemma_thms, ...} = take_info;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   312
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   313
  val R_names = Datatype_Prop.indexify_names (map (K "R") newTs);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   314
  val R_types = map (fn T => T --> T --> boolT) newTs;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   315
  val Rs = map Free (R_names ~~ R_types);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   316
  val n = Free ("n", natT);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   317
  val reserved = "x" :: "y" :: R_names;
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   318
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   319
  (* declare bisimulation predicate *)
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   320
  val bisim_bind = Binding.suffix_name "_bisim" comp_dbind;
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   321
  val bisim_type = R_types ---> boolT;
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   322
  val (bisim_const, thy) =
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   323
      Sign.declare_const ((bisim_bind, bisim_type), NoSyn) thy;
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   324
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   325
  (* define bisimulation predicate *)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   326
  local
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   327
    fun one_con T (con, args) =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   328
      let
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   329
        val Ts = map snd args;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   330
        val ns1 = Name.variant_list reserved (Datatype_Prop.make_tnames Ts);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   331
        val ns2 = map (fn n => n^"'") ns1;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   332
        val vs1 = map Free (ns1 ~~ Ts);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   333
        val vs2 = map Free (ns2 ~~ Ts);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   334
        val eq1 = mk_eq (Free ("x", T), list_ccomb (con, vs1));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   335
        val eq2 = mk_eq (Free ("y", T), list_ccomb (con, vs2));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   336
        fun rel ((v1, v2), T) =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   337
            case AList.lookup (op =) (newTs ~~ Rs) T of
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   338
              NONE => mk_eq (v1, v2) | SOME r => r $ v1 $ v2;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   339
        val eqs = foldr1 mk_conj (map rel (vs1 ~~ vs2 ~~ Ts) @ [eq1, eq2]);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   340
      in
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   341
        Library.foldr mk_ex (vs1 @ vs2, eqs)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   342
      end;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   343
    fun one_eq ((T, R), cons) =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   344
      let
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   345
        val x = Free ("x", T);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   346
        val y = Free ("y", T);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   347
        val disj1 = mk_conj (mk_eq (x, mk_bottom T), mk_eq (y, mk_bottom T));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   348
        val disjs = disj1 :: map (one_con T) cons;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   349
      in
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   350
        mk_all (x, mk_all (y, mk_imp (R $ x $ y, foldr1 mk_disj disjs)))
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   351
      end;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   352
    val conjs = map one_eq (newTs ~~ Rs ~~ map #con_specs constr_infos);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   353
    val bisim_rhs = lambdas Rs (Library.foldr1 mk_conj conjs);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   354
    val bisim_eqn = Logic.mk_equals (bisim_const, bisim_rhs);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   355
  in
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   356
    val (bisim_def_thm, thy) = thy |>
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   357
        yield_singleton (Global_Theory.add_defs false)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   358
         ((Binding.qualified true "bisim_def" comp_dbind, bisim_eqn), []);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   359
  end (* local *)
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   360
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   361
  (* prove coinduction lemma *)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   362
  val coind_lemma =
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   363
    let
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   364
      val assm = mk_trp (list_comb (bisim_const, Rs));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   365
      fun one ((T, R), take_const) =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   366
        let
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   367
          val x = Free ("x", T);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   368
          val y = Free ("y", T);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   369
          val lhs = mk_capply (take_const $ n, x);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   370
          val rhs = mk_capply (take_const $ n, y);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   371
        in
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   372
          mk_all (x, mk_all (y, mk_imp (R $ x $ y, mk_eq (lhs, rhs))))
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   373
        end;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   374
      val goal =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   375
          mk_trp (foldr1 mk_conj (map one (newTs ~~ Rs ~~ take_consts)));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   376
      val rules = @{thm Rep_CFun_strict1} :: take_0_thms;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   377
      fun tacf {prems, context} =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   378
        let
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   379
          val prem' = rewrite_rule [bisim_def_thm] (hd prems);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   380
          val prems' = Project_Rule.projections context prem';
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   381
          val dests = map (fn th => th RS spec RS spec RS mp) prems';
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   382
          fun one_tac (dest, rews) =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   383
              dtac dest 1 THEN safe_tac HOL_cs THEN
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   384
              ALLGOALS (asm_simp_tac (HOL_basic_ss addsimps rews));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   385
        in
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   386
          rtac @{thm nat.induct} 1 THEN
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   387
          simp_tac (HOL_ss addsimps rules) 1 THEN
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   388
          safe_tac HOL_cs THEN
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   389
          EVERY (map one_tac (dests ~~ take_rews))
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   390
        end
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   391
    in
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   392
      Goal.prove_global thy [] [assm] goal tacf
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   393
    end;
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   394
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   395
  (* prove individual coinduction rules *)
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   396
  fun prove_coind ((T, R), take_lemma) =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   397
    let
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   398
      val x = Free ("x", T);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   399
      val y = Free ("y", T);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   400
      val assm1 = mk_trp (list_comb (bisim_const, Rs));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   401
      val assm2 = mk_trp (R $ x $ y);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   402
      val goal = mk_trp (mk_eq (x, y));
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   403
      fun tacf {prems, context} =
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   404
        let
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   405
          val rule = hd prems RS coind_lemma;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   406
        in
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   407
          rtac take_lemma 1 THEN
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   408
          asm_simp_tac (HOL_basic_ss addsimps (rule :: prems)) 1
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   409
        end;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   410
    in
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   411
      Goal.prove_global thy [] [assm1, assm2] goal tacf
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   412
    end;
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   413
  val coinds = map prove_coind (newTs ~~ Rs ~~ take_lemma_thms);
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   414
  val coind_binds = map (Binding.qualified true "coinduct") dbinds;
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   415
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35494
diff changeset
   416
in
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   417
  thy |> snd o Global_Theory.add_thms
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   418
    (map Thm.no_attributes (coind_binds ~~ coinds))
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   419
end; (* let *)
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   420
40018
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   421
(******************************************************************************)
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   422
(******************************* main function ********************************)
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   423
(******************************************************************************)
bf85fef3cce4 avoid using Global_Theory.get_thm
huffman
parents: 40017
diff changeset
   424
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
   425
fun comp_theorems
40022
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   426
    (comp_dbind : binding, eqs : Domain_Library.eq list)
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   427
    (dbinds : binding list)
35659
a78bc1930a7a include take_info within take_induct_info type
huffman
parents: 35658
diff changeset
   428
    (take_info : Domain_Take_Proofs.take_induct_info)
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   429
    (constr_infos : Domain_Constructors.constr_info list)
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
   430
    (thy : theory) =
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   431
let
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   432
val dnames = map (fst o fst) eqs;
35774
218e9766a848 replace some string arguments with bindings
huffman
parents: 35663
diff changeset
   433
val comp_dname = Sign.full_name thy comp_dbind;
35574
ee5df989b7c4 move coinduction-related stuff into function prove_coindunction
huffman
parents: 35560
diff changeset
   434
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   435
(* Test for emptiness *)
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   436
local
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   437
  open Domain_Library;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   438
  val conss = map snd eqs;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   439
  fun rec_to ns lazy_rec (n,cons) = forall (exists (fn arg => 
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   440
        is_rec arg andalso not (member (op =) ns (rec_of arg)) andalso
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   441
        ((rec_of arg =  n andalso not (lazy_rec orelse is_lazy arg)) orelse 
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   442
          rec_of arg <> n andalso rec_to (rec_of arg::ns) 
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   443
            (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg))))
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   444
        ) o snd) cons;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   445
  fun warn (n,cons) =
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   446
    if rec_to [] false (n,cons)
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   447
    then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   448
    else false;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   449
in
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   450
  val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   451
  val is_emptys = map warn n__eqs;
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   452
end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   453
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   454
(* Test for indirect recursion *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   455
local
40022
3a4a24b714f3 simplify automation of induct proof
huffman
parents: 40020
diff changeset
   456
  open Domain_Library;
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   457
  fun indirect_arg arg =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   458
      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
   459
  fun indirect_con (_, args) = exists indirect_arg args;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   460
  fun indirect_eq (_, cons) = exists indirect_con cons;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   461
in
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   462
  val is_indirect = exists indirect_eq eqs;
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   463
  val _ =
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   464
      if is_indirect
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   465
      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
   466
      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
   467
end;
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   468
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   469
(* theorems about take *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   470
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   471
val (take_rewss, thy) =
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40018
diff changeset
   472
    take_theorems dbinds take_info constr_infos thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   473
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   474
val {take_lemma_thms, take_0_thms, take_strict_thms, ...} = take_info;
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   475
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   476
val take_rews = take_0_thms @ take_strict_thms @ flat take_rewss;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   477
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   478
(* prove induction rules, unless definition is indirect recursive *)
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   479
val thy =
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35574
diff changeset
   480
    if is_indirect then thy else
40023
a868e9d73031 move emptiness check to comp_theorems; remove eqs parameter to prove_induction; remove dead code
huffman
parents: 40022
diff changeset
   481
    prove_induction comp_dbind constr_infos take_info take_rews thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   482
35599
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   483
val thy =
20670f5564e9 skip coinduction proofs for indirect-recursive domain definitions
huffman
parents: 35597
diff changeset
   484
    if is_indirect then thy else
40025
876689e6bbdf reimplement proof automation for coinduct rules
huffman
parents: 40023
diff changeset
   485
    prove_coinduction (comp_dbind, dbinds) constr_infos take_info take_rewss thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   486
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35630
diff changeset
   487
in
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35630
diff changeset
   488
  (take_rews, thy)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   489
end; (* let *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   490
end; (* struct *)