src/HOL/Tools/inductive.ML
author haftmann
Fri, 27 Jun 2025 08:09:26 +0200
changeset 82775 61c39a9e5415
parent 82643 f1c14af17591
permissions -rw-r--r--
typo
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
     1
(*  Title:      HOL/Tools/inductive.ML
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
     3
    Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     4
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
     5
(Co)Inductive Definition module for HOL.
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     6
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     7
Features:
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
     8
  * least or greatest fixedpoints
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
     9
  * mutually recursive definitions
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    10
  * definitions involving arbitrary monotone operators
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    11
  * automatically proves introduction and elimination rules
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    12
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    13
  Introduction rules have the form
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
    14
  [| M Pj ti, ..., Q x, ... |] ==> Pk t
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    15
  where M is some monotone operator (usually the identity)
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
    16
  Q x is any side condition on the free variables
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    17
  ti, t are any terms
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
    18
  Pj, Pk are two of the predicates being defined in mutual recursion
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    19
*)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    20
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    21
signature INDUCTIVE =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    22
sig
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    23
  type result =
81810
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
    24
    {preds: term list, elims: thm list, raw_induct: thm, def: thm, mono: thm,
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
    25
     induct: thm, inducts: thm list, intrs: thm list, eqs: thm list}
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    26
  val transform_result: morphism -> result -> result
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    27
  type info = {names: string list, coind: bool} * result
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    28
  val the_inductive: Proof.context -> term -> info
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    29
  val the_inductive_global: Proof.context -> string -> info
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59880
diff changeset
    30
  val print_inductives: bool -> Proof.context -> unit
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
    31
  val get_monos: Proof.context -> thm list
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    32
  val mono_add: attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    33
  val mono_del: attribute
53994
4237859c186d eliminated clone of Inductive.mk_cases_tac;
wenzelm
parents: 52732
diff changeset
    34
  val mk_cases_tac: Proof.context -> tactic
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
    35
  val mk_cases: Proof.context -> term -> thm
10910
058775a575db export inductive_forall_name, inductive_forall_def, rulify;
wenzelm
parents: 10804
diff changeset
    36
  val inductive_forall_def: thm
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
    37
  val rulify: Proof.context -> thm -> thm
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
    38
  val inductive_cases: (Attrib.binding * term list) list -> bool -> local_theory ->
53995
1d457fc83f5c tuned signature;
wenzelm
parents: 53994
diff changeset
    39
    (string * thm list) list * local_theory
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
    40
  val inductive_cases_cmd: (Attrib.binding * string list) list -> bool -> local_theory ->
53995
1d457fc83f5c tuned signature;
wenzelm
parents: 53994
diff changeset
    41
    (string * thm list) list * local_theory
59845
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
    42
  val ind_cases_rules: Proof.context ->
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
    43
    string list -> (binding * string option * mixfix) list -> thm list
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
    44
  val inductive_simps: (Attrib.binding * term list) list -> bool -> local_theory ->
53995
1d457fc83f5c tuned signature;
wenzelm
parents: 53994
diff changeset
    45
    (string * thm list) list * local_theory
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
    46
  val inductive_simps_cmd: (Attrib.binding * string list) list -> bool -> local_theory ->
53995
1d457fc83f5c tuned signature;
wenzelm
parents: 53994
diff changeset
    47
    (string * thm list) list * local_theory
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    48
  type flags =
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
    49
    {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool,
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
    50
      no_elim: bool, no_ind: bool, skip_mono: bool}
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    51
  val add_inductive:
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    52
    flags -> ((binding * typ) * mixfix) list ->
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    53
    (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory ->
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    54
    result * local_theory
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    55
  val add_inductive_cmd: bool -> bool ->
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29388
diff changeset
    56
    (binding * string option * mixfix) list ->
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29388
diff changeset
    57
    (binding * string option * mixfix) list ->
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
    58
    Specification.multi_specs_cmd ->
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 56334
diff changeset
    59
    (Facts.ref * Token.src list) list ->
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    60
    local_theory -> result * local_theory
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
    61
  val arities_of: thm -> (string * int) list
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
    62
  val params_of: thm -> term list
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
    63
  val partition_rules: thm -> thm list -> (string * thm list) list
25822
05756950011c Added function partition_rules'.
berghofe
parents: 25510
diff changeset
    64
  val partition_rules': thm -> (thm * 'a) list -> (string * (thm * 'a) list) list
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
    65
  val unpartition_rules: thm list -> (string * 'a list) list -> 'a list
60362
befdc10ebb42 clarified context;
wenzelm
parents: 60097
diff changeset
    66
  val infer_intro_vars: theory -> thm -> int -> thm list -> term list list
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61951
diff changeset
    67
  val inductive_internals: bool Config.T
59532
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
    68
  val select_disj_tac: Proof.context -> int -> int -> int -> tactic
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
    69
  type add_ind_def =
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    70
    flags ->
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
    71
    term list -> (Attrib.binding * term) list -> thm list ->
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
    72
    term list -> (binding * mixfix) list ->
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    73
    local_theory -> result * local_theory
71214
5727bcc3c47c proper spec_rule name via naming/binding/Morphism.binding;
wenzelm
parents: 71179
diff changeset
    74
  val declare_rules: binding -> bool -> bool -> binding -> string list -> term list ->
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 56334
diff changeset
    75
    thm list -> binding list -> Token.src list list -> (thm * string list * int) list ->
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
    76
    thm list -> thm -> local_theory -> thm list * thm list * thm list * thm * thm list * local_theory
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
    77
  val add_ind_def: add_ind_def
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    78
  val gen_add_inductive: add_ind_def -> flags ->
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29388
diff changeset
    79
    ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    80
    thm list -> local_theory -> result * local_theory
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    81
  val gen_add_inductive_cmd: add_ind_def -> bool -> bool ->
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29388
diff changeset
    82
    (binding * string option * mixfix) list ->
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29388
diff changeset
    83
    (binding * string option * mixfix) list ->
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
    84
    Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list ->
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
    85
    local_theory -> result * local_theory
49324
4f28543ae7fa removed obsolete argument "int" and thus made SML/NJ happy (cf. 03bee3a6a1b7);
wenzelm
parents: 49170
diff changeset
    86
  val gen_ind_decl: add_ind_def -> bool -> (local_theory -> local_theory) parser
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
    87
end;
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
    88
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31177
diff changeset
    89
structure Inductive: INDUCTIVE =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    90
struct
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    91
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
    92
(** theory context references **)
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
    93
59940
087d81f5213e local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents: 59936
diff changeset
    94
val inductive_forall_def = @{thm HOL.induct_forall_def};
087d81f5213e local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents: 59936
diff changeset
    95
val inductive_conj_def = @{thm HOL.induct_conj_def};
32602
f2b741473860 more antiquotations
haftmann
parents: 32181
diff changeset
    96
val inductive_conj = @{thms induct_conj};
f2b741473860 more antiquotations
haftmann
parents: 32181
diff changeset
    97
val inductive_atomize = @{thms induct_atomize};
f2b741473860 more antiquotations
haftmann
parents: 32181
diff changeset
    98
val inductive_rulify = @{thms induct_rulify};
f2b741473860 more antiquotations
haftmann
parents: 32181
diff changeset
    99
val inductive_rulify_fallback = @{thms induct_rulify_fallback};
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   100
45649
wenzelm
parents: 45648
diff changeset
   101
val simp_thms1 =
wenzelm
parents: 45648
diff changeset
   102
  map mk_meta_eq
67091
1393c2340eec more symbols;
wenzelm
parents: 65411
diff changeset
   103
    @{lemma "(\<not> True) = False" "(\<not> False) = True"
1393c2340eec more symbols;
wenzelm
parents: 65411
diff changeset
   104
        "(True \<longrightarrow> P) = P" "(False \<longrightarrow> P) = True"
1393c2340eec more symbols;
wenzelm
parents: 65411
diff changeset
   105
        "(P \<and> True) = P" "(True \<and> P) = P"
45649
wenzelm
parents: 45648
diff changeset
   106
      by (fact simp_thms)+};
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   107
45649
wenzelm
parents: 45648
diff changeset
   108
val simp_thms2 =
wenzelm
parents: 45648
diff changeset
   109
  map mk_meta_eq [@{thm inf_fun_def}, @{thm inf_bool_def}] @ simp_thms1;
32652
3175e23b79f3 stripped legacy ML bindings
haftmann
parents: 32610
diff changeset
   110
45649
wenzelm
parents: 45648
diff changeset
   111
val simp_thms3 =
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   112
  @{thms le_rel_bool_arg_iff if_False if_True conj_ac
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   113
    le_fun_def le_bool_def sup_fun_def sup_bool_def simp_thms
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   114
    if_bool_eq_disj all_simps ex_simps imp_conjL};
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   115
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   116
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   117
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   118
(** misc utilities **)
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   119
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   120
val inductive_internals = Attrib.setup_config_bool \<^binding>\<open>inductive_internals\<close> (K false);
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
   121
26477
ecf06644f6cb eliminated quiete_mode ref (turned into proper argument);
wenzelm
parents: 26336
diff changeset
   122
fun message quiet_mode s = if quiet_mode then () else writeln s;
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   123
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   124
fun clean_message ctxt quiet_mode s =
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   125
  if Config.get ctxt quick_and_dirty then () else message quiet_mode s;
5662
72a2e33d3b9e Added quiet_mode flag.
berghofe
parents: 5553
diff changeset
   126
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   127
fun coind_prefix true = "co"
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   128
  | coind_prefix false = "";
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   129
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   130
fun log (b: int) m n = if m >= n then 0 else 1 + log b (b * m) n;
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   131
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   132
fun make_bool_args f g [] i = []
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   133
  | make_bool_args f g (x :: xs) i =
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   134
      (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2);
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   135
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   136
fun make_bool_args' xs =
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   137
  make_bool_args (K \<^Const>\<open>False\<close>) (K \<^Const>\<open>True\<close>) xs;
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   138
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   139
fun arg_types_of k c = drop k (binder_types (fastype_of c));
33077
3c9cf88ec841 arg_types_of auxiliary function; using multiset operations
haftmann
parents: 33056
diff changeset
   140
40316
665862241968 replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
wenzelm
parents: 39248
diff changeset
   141
fun find_arg T x [] = raise Fail "find_arg"
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   142
  | find_arg T x ((p as (_, (SOME _, _))) :: ps) =
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   143
      apsnd (cons p) (find_arg T x ps)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   144
  | find_arg T x ((p as (U, (NONE, y))) :: ps) =
23577
c5b93c69afd3 avoid polymorphic equality;
wenzelm
parents: 22997
diff changeset
   145
      if (T: typ) = U then (y, (U, (SOME x, y)) :: ps)
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   146
      else apsnd (cons p) (find_arg T x ps);
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   147
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   148
fun make_args Ts xs =
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   149
  map (fn (T, (NONE, ())) => \<^Const>\<open>undefined T\<close> | (_, (SOME t, ())) => t)
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   150
    (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts));
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   151
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   152
fun make_args' Ts xs Us =
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   153
  fst (fold_map (fn T => find_arg T ()) Us (Ts ~~ map (pair NONE) xs));
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   154
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   155
fun dest_predicate cs params t =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   156
  let
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   157
    val k = length params;
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   158
    val (c, ts) = strip_comb t;
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   159
    val (xs, ys) = chop k ts;
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31902
diff changeset
   160
    val i = find_index (fn c' => c' = c) cs;
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   161
  in
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   162
    if xs = params andalso i >= 0 then
33077
3c9cf88ec841 arg_types_of auxiliary function; using multiset operations
haftmann
parents: 33056
diff changeset
   163
      SOME (c, i, ys, chop (length ys) (arg_types_of k c))
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   164
    else NONE
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   165
  end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   166
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   167
fun mk_names a 0 = []
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   168
  | mk_names a 1 = [a]
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   169
  | mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n);
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   170
59532
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   171
fun select_disj_tac ctxt =
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   172
  let
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   173
    fun tacs 1 1 = []
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   174
      | tacs _ 1 = [resolve_tac ctxt @{thms disjI1}]
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   175
      | tacs n i = resolve_tac ctxt @{thms disjI2} :: tacs (n - 1) (i - 1);
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   176
  in fn n => fn i => EVERY' (tacs n i) end;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   177
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   178
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   179
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   180
(** context data **)
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   181
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
   182
type result =
81810
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
   183
  {preds: term list, elims: thm list, raw_induct: thm, def: thm, mono: thm,
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   184
   induct: thm, inducts: thm list, intrs: thm list, eqs: thm list};
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   185
81810
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
   186
fun transform_result phi {preds, elims, raw_induct: thm, induct, inducts, intrs, eqs, def, mono} =
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   187
  let
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   188
    val term = Morphism.term phi;
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   189
    val thm = Morphism.thm phi;
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   190
    val fact = Morphism.fact phi;
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   191
  in
81810
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
   192
   {preds = map term preds, elims = fact elims, raw_induct = thm raw_induct, def = thm def,
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
   193
    induct = thm induct, inducts = fact inducts, intrs = fact intrs, eqs = fact eqs,
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
   194
    mono = thm mono}
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   195
  end;
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   196
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
   197
type info = {names: string list, coind: bool} * result;
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   198
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   199
val empty_infos =
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   200
  Item_Net.init (op = o apply2 (#names o fst)) (#preds o snd)
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   201
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   202
val empty_equations =
45652
18214436e1d3 permissive update for improved "tool compliance";
wenzelm
parents: 45651
diff changeset
   203
  Item_Net.init Thm.eq_thm_prop
18214436e1d3 permissive update for improved "tool compliance";
wenzelm
parents: 45651
diff changeset
   204
    (single o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of);
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   205
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   206
datatype data = Data of
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
   207
 {infos: info Item_Net.T,
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   208
  monos: thm list,
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   209
  equations: thm Item_Net.T};
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   210
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   211
fun make_data (infos, monos, equations) =
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   212
  Data {infos = infos, monos = monos, equations = equations};
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   213
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   214
structure Data = Generic_Data
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   215
(
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   216
  type T = data;
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   217
  val empty = make_data (empty_infos, [], empty_equations);
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   218
  fun merge (Data {infos = infos1, monos = monos1, equations = equations1},
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   219
      Data {infos = infos2, monos = monos2, equations = equations2}) =
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   220
    make_data (Item_Net.merge (infos1, infos2),
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   221
      Thm.merge_thms (monos1, monos2),
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   222
      Item_Net.merge (equations1, equations2));
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   223
);
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   224
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   225
fun map_data f =
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   226
  Data.map (fn Data {infos, monos, equations} => make_data (f (infos, monos, equations)));
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   227
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   228
fun rep_data ctxt = Data.get (Context.Proof ctxt) |> (fn Data rep => rep);
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   229
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59880
diff changeset
   230
fun print_inductives verbose ctxt =
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   231
  let
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   232
    val {infos, monos, ...} = rep_data ctxt;
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   233
    val space = Consts.space_of (Proof_Context.consts_of ctxt);
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   234
    val consts =
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   235
      Item_Net.content infos
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   236
      |> maps (fn ({names, ...}, result) => map (rpair result) names)
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   237
  in
50301
56b4c9afd7be print formal entities with markup;
wenzelm
parents: 50214
diff changeset
   238
    [Pretty.block
56b4c9afd7be print formal entities with markup;
wenzelm
parents: 50214
diff changeset
   239
      (Pretty.breaks
56b4c9afd7be print formal entities with markup;
wenzelm
parents: 50214
diff changeset
   240
        (Pretty.str "(co)inductives:" ::
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59880
diff changeset
   241
          map (Pretty.mark_str o #1)
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   242
            (Name_Space.markup_entries verbose ctxt space consts))),
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61063
diff changeset
   243
     Pretty.big_list "monotonicity rules:" (map (Thm.pretty_thm_item ctxt) monos)]
82587
7415414bd9d8 more scalable: discontinue odd shortcuts from 6b3739fee456, which produce bulky strings internally;
wenzelm
parents: 82317
diff changeset
   244
  end |> Pretty.chunks |> Pretty.writeln;
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   245
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   246
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   247
(* inductive info *)
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   248
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   249
fun the_inductive ctxt term =
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   250
  Item_Net.retrieve (#infos (rep_data ctxt)) term
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   251
  |> the_single
67664
ad2b3e330c27 tuned signature;
wenzelm
parents: 67649
diff changeset
   252
  |> apsnd (transform_result (Morphism.transfer_morphism' ctxt))
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   253
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   254
fun the_inductive_global ctxt name =
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   255
  #infos (rep_data ctxt)
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   256
  |> Item_Net.content
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   257
  |> filter (fn ({names, ...}, _) => member op = names name)
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   258
  |> the_single
67664
ad2b3e330c27 tuned signature;
wenzelm
parents: 67649
diff changeset
   259
  |> apsnd (transform_result (Morphism.transfer_morphism' ctxt))
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   260
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
   261
fun put_inductives info =
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   262
  map_data (fn (infos, monos, equations) =>
67637
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   263
    (Item_Net.update (apsnd (transform_result Morphism.trim_context_morphism) info) infos,
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   264
      monos, equations));
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   265
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   266
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   267
(* monotonicity rules *)
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   268
67637
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   269
fun get_monos ctxt =
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   270
  #monos (rep_data ctxt)
67649
1e1782c1aedf tuned signature;
wenzelm
parents: 67637
diff changeset
   271
  |> map (Thm.transfer' ctxt);
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   272
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   273
fun mk_mono ctxt thm =
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   274
  let
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   275
    fun eq_to_mono thm' = thm' RS (thm' RS @{thm eq_to_mono});
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   276
    fun dest_less_concl thm = dest_less_concl (thm RS @{thm le_funD})
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   277
      handle THM _ => thm RS @{thm le_boolD}
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   278
  in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   279
    (case Thm.concl_of thm of
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   280
      \<^Const_>\<open>Pure.eq _ for _ _\<close> => eq_to_mono (HOLogic.mk_obj_eq thm)
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   281
    | _ $ \<^Const_>\<open>HOL.eq _ for _ _\<close> => eq_to_mono thm
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   282
    | _ $ \<^Const_>\<open>less_eq _ for _ _\<close> =>
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   283
      dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   284
        (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}])) thm))
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   285
    | _ => thm)
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61063
diff changeset
   286
  end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Thm.string_of_thm ctxt thm);
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   287
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   288
val mono_add =
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   289
  Thm.declaration_attribute (fn thm => fn context =>
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   290
    map_data (fn (infos, monos, equations) =>
67637
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   291
      (infos, Thm.add_thm (Thm.trim_context (mk_mono (Context.proof_of context) thm)) monos,
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   292
        equations)) context);
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   293
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   294
val mono_del =
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   295
  Thm.declaration_attribute (fn thm => fn context =>
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   296
    map_data (fn (infos, monos, equations) =>
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   297
      (infos, Thm.del_thm (mk_mono (Context.proof_of context) thm) monos, equations)) context);
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   298
58815
fd3f893a40ea modernized setup;
wenzelm
parents: 58028
diff changeset
   299
val _ =
fd3f893a40ea modernized setup;
wenzelm
parents: 58028
diff changeset
   300
  Theory.setup
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   301
    (Attrib.setup \<^binding>\<open>mono\<close> (Attrib.add_del mono_add mono_del)
58815
fd3f893a40ea modernized setup;
wenzelm
parents: 58028
diff changeset
   302
      "declaration of monotonicity rule");
fd3f893a40ea modernized setup;
wenzelm
parents: 58028
diff changeset
   303
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   304
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   305
(* equations *)
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   306
67637
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   307
fun retrieve_equations ctxt =
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   308
  Item_Net.retrieve (#equations (rep_data ctxt))
67649
1e1782c1aedf tuned signature;
wenzelm
parents: 67637
diff changeset
   309
  #> map (Thm.transfer' ctxt);
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   310
45652
18214436e1d3 permissive update for improved "tool compliance";
wenzelm
parents: 45651
diff changeset
   311
val equation_add_permissive =
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   312
  Thm.declaration_attribute (fn thm =>
45652
18214436e1d3 permissive update for improved "tool compliance";
wenzelm
parents: 45651
diff changeset
   313
    map_data (fn (infos, monos, equations) =>
67637
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   314
      (infos, monos, perhaps (try (Item_Net.update (Thm.trim_context thm))) equations)));
45651
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   315
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   316
172aa230ce69 just one data slot per module;
wenzelm
parents: 45649
diff changeset
   317
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   318
(** process rules **)
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   319
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   320
local
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   321
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
   322
fun err_in_rule ctxt name t msg =
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42364
diff changeset
   323
  error (cat_lines ["Ill-formed introduction rule " ^ Binding.print name,
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24915
diff changeset
   324
    Syntax.string_of_term ctxt t, msg]);
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   325
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
   326
fun err_in_prem ctxt name t p msg =
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24915
diff changeset
   327
  error (cat_lines ["Ill-formed premise", Syntax.string_of_term ctxt p,
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42364
diff changeset
   328
    "in introduction rule " ^ Binding.print name, Syntax.string_of_term ctxt t, msg]);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   329
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   330
val bad_concl = "Conclusion of introduction rule must be an inductive predicate";
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   331
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   332
val bad_ind_occ = "Inductive predicate occurs in argument of inductive predicate";
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   333
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   334
val bad_app = "Inductive predicate must be applied to parameter(s) ";
11358
416ea5c009f5 now checks for leading meta-quantifiers and complains, instead of
paulson
parents: 11036
diff changeset
   335
82643
f1c14af17591 prefer Simplifier over bootstrap-only Raw_Simplifier
haftmann
parents: 82587
diff changeset
   336
fun atomize_term thy = Simplifier.rewrite_term thy inductive_atomize [];
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   337
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   338
in
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   339
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27882
diff changeset
   340
fun check_rule ctxt cs params ((binding, att), rule) =
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   341
  let
81525
3e55334ef5af clarified names context: proper context, without consts;
wenzelm
parents: 81519
diff changeset
   342
    val params' = Variable.variant_names (Variable.declare_names rule ctxt) (Logic.strip_params rule);
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   343
    val frees = rev (map Free params');
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   344
    val concl = subst_bounds (frees, Logic.strip_assums_concl rule);
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   345
    val prems = map (curry subst_bounds frees) (Logic.strip_assums_hyp rule);
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
   346
    val rule' = Logic.list_implies (prems, concl);
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   347
    val aprems = map (atomize_term (Proof_Context.theory_of ctxt)) prems;
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
   348
    val arule = fold_rev (Logic.all o Free) params' (Logic.list_implies (aprems, concl));
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   349
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   350
    fun check_ind err t =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   351
      (case dest_predicate cs params t of
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   352
        NONE => err (bad_app ^
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24915
diff changeset
   353
          commas (map (Syntax.string_of_term ctxt) params))
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   354
      | SOME (_, _, ys, _) =>
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   355
          if exists (fn c => exists (fn t => Logic.occs (c, t)) ys) cs
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   356
          then err bad_ind_occ else ());
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   357
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   358
    fun check_prem' prem t =
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: 36642
diff changeset
   359
      if member (op =) cs (head_of t) then
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42364
diff changeset
   360
        check_ind (err_in_prem ctxt binding rule prem) t
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   361
      else
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   362
        (case t of
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   363
          Abs (_, _, t) => check_prem' prem t
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   364
        | t $ u => (check_prem' prem t; check_prem' prem u)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   365
        | _ => ());
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   366
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   367
    fun check_prem (prem, aprem) =
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   368
      if can HOLogic.dest_Trueprop aprem then check_prem' prem prem
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42364
diff changeset
   369
      else err_in_prem ctxt binding rule prem "Non-atomic premise";
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   370
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   371
    val _ =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   372
      (case concl of
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   373
        \<^Const_>\<open>Trueprop for t\<close> =>
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   374
          if member (op =) cs (head_of t) then
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42364
diff changeset
   375
           (check_ind (err_in_rule ctxt binding rule') t;
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   376
            List.app check_prem (prems ~~ aprems))
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   377
          else err_in_rule ctxt binding rule' bad_concl
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   378
       | _ => err_in_rule ctxt binding rule' bad_concl);
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   379
  in
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27882
diff changeset
   380
    ((binding, att), arule)
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   381
  end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   382
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   383
fun rulify ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   384
  hol_simplify ctxt inductive_conj
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   385
  #> hol_simplify ctxt inductive_rulify
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   386
  #> hol_simplify ctxt inductive_rulify_fallback
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54742
diff changeset
   387
  #> Simplifier.norm_hhf ctxt;
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   388
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   389
end;
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   390
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   391
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   392
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   393
(** proofs for (co)inductive predicates **)
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   394
26534
a2cb4de2a1aa Added skip_mono flag and inductive_flags type.
berghofe
parents: 26477
diff changeset
   395
(* prove monotonicity *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   396
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
   397
fun prove_mono quiet_mode skip_mono predT fp_fun monos ctxt =
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   398
 (message (quiet_mode orelse skip_mono andalso Config.get ctxt quick_and_dirty)
26534
a2cb4de2a1aa Added skip_mono flag and inductive_flags type.
berghofe
parents: 26477
diff changeset
   399
    "  Proving monotonicity ...";
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 50771
diff changeset
   400
  (if skip_mono then Goal.prove_sorry else Goal.prove_future) ctxt
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
   401
    [] []
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   402
    (HOLogic.mk_Trueprop
76054
a4b47c684445 moved mono and strict_mono to Fun and redefined them as abbreviations
desharna
parents: 74561
diff changeset
   403
      (\<^Const>\<open>monotone_on predT predT for
a4b47c684445 moved mono and strict_mono to Fun and redefined them as abbreviations
desharna
parents: 74561
diff changeset
   404
          \<^Const>\<open>top \<^Type>\<open>set predT\<close>\<close> \<^Const>\<open>less_eq predT\<close> \<^Const>\<open>less_eq predT\<close> fp_fun\<close>))
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   405
    (fn _ => EVERY [resolve_tac ctxt @{thms monoI} 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   406
      REPEAT (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}] 1),
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   407
      REPEAT (FIRST
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58936
diff changeset
   408
        [assume_tac ctxt 1,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   409
         resolve_tac ctxt (map (mk_mono ctxt) monos @ get_monos ctxt) 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   410
         eresolve_tac ctxt @{thms le_funE} 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   411
         dresolve_tac ctxt @{thms le_boolD} 1])]));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   412
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   413
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   414
(* prove introduction rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   415
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
   416
fun prove_intrs quiet_mode coind mono fp_def k intr_ts rec_preds_defs ctxt ctxt' =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   417
  let
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   418
    val _ = clean_message ctxt quiet_mode "  Proving the introduction rules ...";
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   419
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   420
    val unfold = funpow k (fn th => th RS fun_cong)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   421
      (mono RS (fp_def RS
32652
3175e23b79f3 stripped legacy ML bindings
haftmann
parents: 32610
diff changeset
   422
        (if coind then @{thm def_gfp_unfold} else @{thm def_lfp_unfold})));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   423
67091
1393c2340eec more symbols;
wenzelm
parents: 65411
diff changeset
   424
    val rules = [refl, TrueI, @{lemma "\<not> False" by (rule notI)}, exI, conjI];
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   425
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
   426
    val intrs = map_index (fn (i, intr) =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 50771
diff changeset
   427
      Goal.prove_sorry ctxt [] [] intr (fn _ => EVERY
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   428
       [rewrite_goals_tac ctxt rec_preds_defs,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   429
        resolve_tac ctxt [unfold RS iffD2] 1,
59532
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   430
        select_disj_tac ctxt (length intr_ts) (i + 1) 1,
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   431
        (*Not ares_tac, since refl must be tried before any equality assumptions;
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   432
          backtracking may occur if the premises have extra variables!*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   433
        DEPTH_SOLVE_1 (resolve_tac ctxt rules 1 APPEND assume_tac ctxt 1)])
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   434
       |> singleton (Proof_Context.export ctxt ctxt')) intr_ts
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   435
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   436
  in (intrs, unfold) end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   437
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   438
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   439
(* prove elimination rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   440
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
   441
fun prove_elims quiet_mode cs params intr_ts intr_names unfold rec_preds_defs ctxt ctxt''' =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   442
  let
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   443
    val _ = clean_message ctxt quiet_mode "  Proving the elimination rules ...";
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   444
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
   445
    val ([pname], ctxt') = Variable.variant_fixes ["P"] ctxt;
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   446
    val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT));
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   447
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   448
    fun dest_intr r =
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   449
      (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))),
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   450
       Logic.strip_assums_hyp r, Logic.strip_params r);
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   451
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   452
    val intrs = map dest_intr intr_ts ~~ intr_names;
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   453
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   454
    val rules1 = [disjE, exE, FalseE];
67091
1393c2340eec more symbols;
wenzelm
parents: 65411
diff changeset
   455
    val rules2 = [conjE, FalseE, @{lemma "\<not> True \<Longrightarrow> R" by (rule notE [OF _ TrueI])}];
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   456
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   457
    fun prove_elim c =
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   458
      let
33077
3c9cf88ec841 arg_types_of auxiliary function; using multiset operations
haftmann
parents: 33056
diff changeset
   459
        val Ts = arg_types_of (length params) c;
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   460
        val (anames, ctxt'') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt';
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   461
        val frees = map Free (anames ~~ Ts);
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   462
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   463
        fun mk_elim_prem ((_, _, us, _), ts, params') =
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 46215
diff changeset
   464
          Logic.list_all (params',
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   465
            Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   466
              (frees ~~ us) @ ts, P));
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33278
diff changeset
   467
        val c_intrs = filter (equal c o #1 o #1 o #1) intrs;
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   468
        val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) ::
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   469
           map mk_elim_prem (map #1 c_intrs)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   470
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 50771
diff changeset
   471
        (Goal.prove_sorry ctxt'' [] prems P
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   472
          (fn {context = ctxt4, prems} => EVERY
46708
b138dee7bed3 prefer cut_tac, where it is clear that the special variants cut_rules_tac or cut_facts_tac are not required;
wenzelm
parents: 46219
diff changeset
   473
            [cut_tac (hd prems) 1,
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   474
             rewrite_goals_tac ctxt4 rec_preds_defs,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   475
             dresolve_tac ctxt4 [unfold RS iffD1] 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   476
             REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules1)),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   477
             REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules2)),
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   478
             EVERY (map (fn prem =>
59499
14095f771781 misc tuning;
wenzelm
parents: 59498
diff changeset
   479
               DEPTH_SOLVE_1 (assume_tac ctxt4 1 ORELSE
14095f771781 misc tuning;
wenzelm
parents: 59498
diff changeset
   480
                resolve_tac ctxt [rewrite_rule ctxt4 rec_preds_defs prem, conjI] 1))
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   481
                (tl prems))])
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   482
          |> singleton (Proof_Context.export ctxt'' ctxt'''),
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 33966
diff changeset
   483
         map #2 c_intrs, length Ts)
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   484
      end
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   485
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   486
   in map prove_elim cs end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   487
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   488
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   489
(* prove simplification equations *)
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   490
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   491
fun prove_eqs quiet_mode cs params intr_ts intrs
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   492
    (elims: (thm * bstring list * int) list) ctxt ctxt'' =  (* FIXME ctxt'' ?? *)
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   493
  let
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   494
    val _ = clean_message ctxt quiet_mode "  Proving the simplification rules ...";
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   495
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   496
    fun dest_intr r =
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   497
      (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))),
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   498
       Logic.strip_assums_hyp r, Logic.strip_params r);
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   499
    val intr_ts' = map dest_intr intr_ts;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   500
37901
ea7d4423cb5b make SML/NJ happy, by adhoc type-constraints;
wenzelm
parents: 37734
diff changeset
   501
    fun prove_eq c (elim: thm * 'a * 'b) =
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   502
      let
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   503
        val Ts = arg_types_of (length params) c;
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   504
        val (anames, ctxt') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt;
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   505
        val frees = map Free (anames ~~ Ts);
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   506
        val c_intrs = filter (equal c o #1 o #1 o #1) (intr_ts' ~~ intrs);
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   507
        fun mk_intr_conj (((_, _, us, _), ts, params'), _) =
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   508
          let
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   509
            fun list_ex ([], t) = t
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   510
              | list_ex ((a, T) :: vars, t) =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   511
                  HOLogic.exists_const T $ Abs (a, T, list_ex (vars, t));
47876
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   512
            val conjs = map2 (curry HOLogic.mk_eq) frees us @ map HOLogic.dest_Trueprop ts;
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   513
          in
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   514
            list_ex (params', if null conjs then \<^Const>\<open>True\<close> else foldr1 HOLogic.mk_conj conjs)
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   515
          end;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   516
        val lhs = list_comb (c, params @ frees);
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   517
        val rhs =
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   518
          if null c_intrs then \<^Const>\<open>False\<close>
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   519
          else foldr1 HOLogic.mk_disj (map mk_intr_conj c_intrs);
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   520
        val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   521
        fun prove_intr1 (i, _) = Subgoal.FOCUS_PREMS (fn {context = ctxt'', params, prems, ...} =>
59532
82ab8168d940 proper context;
wenzelm
parents: 59499
diff changeset
   522
            select_disj_tac ctxt'' (length c_intrs) (i + 1) 1 THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   523
            EVERY (replicate (length params) (resolve_tac ctxt'' @{thms exI} 1)) THEN
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   524
            (if null prems then resolve_tac ctxt'' @{thms TrueI} 1
47876
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   525
             else
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   526
              let
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   527
                val (prems', last_prem) = split_last prems;
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   528
              in
58839
ccda99401bc8 eliminated aliases;
wenzelm
parents: 58815
diff changeset
   529
                EVERY (map (fn prem =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   530
                  (resolve_tac ctxt'' @{thms conjI} 1 THEN resolve_tac ctxt'' [prem] 1)) prems')
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   531
                THEN resolve_tac ctxt'' [last_prem] 1
47876
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   532
              end)) ctxt' 1;
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   533
        fun prove_intr2 (((_, _, us, _), ts, params'), intr) =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   534
          EVERY (replicate (length params') (eresolve_tac ctxt' @{thms exE} 1)) THEN
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   535
          (if null ts andalso null us then resolve_tac ctxt' [intr] 1
47876
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   536
           else
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   537
            EVERY (replicate (length ts + length us - 1) (eresolve_tac ctxt' @{thms conjE} 1)) THEN
59059
haftmann
parents: 58993
diff changeset
   538
            Subgoal.FOCUS_PREMS (fn {context = ctxt'', prems, ...} =>
47876
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   539
              let
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   540
                val (eqs, prems') = chop (length us) prems;
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   541
                val rew_thms = map (fn th => th RS @{thm eq_reflection}) eqs;
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   542
              in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   543
                rewrite_goal_tac ctxt'' rew_thms 1 THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   544
                resolve_tac ctxt'' [intr] 1 THEN
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   545
                EVERY (map (fn p => resolve_tac ctxt'' [p] 1) prems')
47876
0521ee2e504d tweaked Inductive.prove_eqs to allow degenerate definition like "inductive TRUE where TRUE";
wenzelm
parents: 46961
diff changeset
   546
              end) ctxt' 1);
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   547
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 50771
diff changeset
   548
        Goal.prove_sorry ctxt' [] [] eq (fn _ =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   549
          resolve_tac ctxt' @{thms iffI} 1 THEN
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   550
          eresolve_tac ctxt' [#1 elim] 1 THEN
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   551
          EVERY (map_index prove_intr1 c_intrs) THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   552
          (if null c_intrs then eresolve_tac ctxt' @{thms FalseE} 1
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   553
           else
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   554
            let val (c_intrs', last_c_intr) = split_last c_intrs in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   555
              EVERY (map (fn ci => eresolve_tac ctxt' @{thms disjE} 1 THEN prove_intr2 ci) c_intrs')
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   556
              THEN prove_intr2 last_c_intr
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   557
            end))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   558
        |> rulify ctxt'
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   559
        |> singleton (Proof_Context.export ctxt' ctxt'')
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   560
      end;
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   561
  in
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   562
    map2 prove_eq cs elims
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   563
  end;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   564
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   565
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   566
(* derivation of simplified elimination rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   567
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   568
local
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   569
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   570
(*delete needless equality assumptions*)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67768
diff changeset
   571
val refl_thin = Goal.prove_global \<^theory>\<open>HOL\<close> [] [] \<^prop>\<open>\<And>P. a = a \<Longrightarrow> P \<Longrightarrow> P\<close>
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58936
diff changeset
   572
  (fn {context = ctxt, ...} => assume_tac ctxt 1);
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   573
val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE];
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   574
fun elim_tac ctxt = REPEAT o eresolve_tac ctxt elim_rls;
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   575
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   576
fun simp_case_tac ctxt i =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   577
  EVERY' [elim_tac ctxt,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   578
    asm_full_simp_tac ctxt,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   579
    elim_tac ctxt,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   580
    REPEAT o bound_hyp_subst_tac ctxt] i;
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
   581
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   582
in
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   583
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   584
fun mk_cases_tac ctxt = ALLGOALS (simp_case_tac ctxt) THEN prune_params_tac ctxt;
53994
4237859c186d eliminated clone of Inductive.mk_cases_tac;
wenzelm
parents: 52732
diff changeset
   585
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
   586
fun mk_cases ctxt prop =
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   587
  let
21526
1e6bd5ed7abc added morh_result, the_inductive, add_inductive_global;
wenzelm
parents: 21508
diff changeset
   588
    fun err msg =
1e6bd5ed7abc added morh_result, the_inductive, add_inductive_global;
wenzelm
parents: 21508
diff changeset
   589
      error (Pretty.string_of (Pretty.block
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24915
diff changeset
   590
        [Pretty.str msg, Pretty.fbrk, Syntax.pretty_term ctxt prop]));
21526
1e6bd5ed7abc added morh_result, the_inductive, add_inductive_global;
wenzelm
parents: 21508
diff changeset
   591
24861
cc669ca5f382 tuned Induct interface: prefer pred'' over set'';
wenzelm
parents: 24830
diff changeset
   592
    val elims = Induct.find_casesP ctxt prop;
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
   593
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   594
    val cprop = Thm.cterm_of ctxt prop;
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
   595
    fun mk_elim rl =
53994
4237859c186d eliminated clone of Inductive.mk_cases_tac;
wenzelm
parents: 52732
diff changeset
   596
      Thm.implies_intr cprop
4237859c186d eliminated clone of Inductive.mk_cases_tac;
wenzelm
parents: 52732
diff changeset
   597
        (Tactic.rule_by_tactic ctxt (mk_cases_tac ctxt) (Thm.assume cprop RS rl))
70318
9eff9e2dc177 proper Proof_Context.export_morphism corresponding to Proof_Context.augment (see 7f568724d67e);
wenzelm
parents: 70308
diff changeset
   598
      |> singleton (Proof_Context.export (Proof_Context.augment prop ctxt) ctxt);
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   599
  in
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   600
    (case get_first (try mk_elim) elims of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   601
      SOME r => r
21526
1e6bd5ed7abc added morh_result, the_inductive, add_inductive_global;
wenzelm
parents: 21508
diff changeset
   602
    | NONE => err "Proposition not an inductive predicate:")
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   603
  end;
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   604
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   605
end;
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   606
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   607
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
   608
(* inductive_cases *)
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   609
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   610
fun gen_inductive_cases prep_att prep_prop args int lthy =
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   611
  let
46915
4b2eccaec3f6 more parallel inductive_cases;
wenzelm
parents: 46893
diff changeset
   612
    val thmss =
4b2eccaec3f6 more parallel inductive_cases;
wenzelm
parents: 46893
diff changeset
   613
      map snd args
58993
302104d8366b prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
wenzelm
parents: 58963
diff changeset
   614
      |> burrow (grouped 10 Par_List.map_independent (mk_cases lthy o prep_prop lthy));
46915
4b2eccaec3f6 more parallel inductive_cases;
wenzelm
parents: 46893
diff changeset
   615
    val facts =
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55111
diff changeset
   616
      map2 (fn ((a, atts), _) => fn thms => ((a, map (prep_att lthy) atts), [(thms, [])]))
46915
4b2eccaec3f6 more parallel inductive_cases;
wenzelm
parents: 46893
diff changeset
   617
        args thmss;
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   618
    val (res, lthy') = lthy |> Local_Theory.notes facts
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   619
    val _ =
82317
231b6d8231c6 more uniform Proof_Display.print_results for theory and proof output --- avoid loss of information seen in src/Doc/JEdit/document/output-and-state.png (the first bad changeset is f8c412a45af8, see also 53b59fa42696);
wenzelm
parents: 81810
diff changeset
   620
      Proof_Display.print_results {interactive = int, pos = Position.thread_data ()}
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   621
        lthy' ((Thm.theoremK, ""), res); 
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   622
  in (res, lthy') end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   623
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
   624
val inductive_cases = gen_inductive_cases (K I) Syntax.check_prop;
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
   625
val inductive_cases_cmd = gen_inductive_cases Attrib.check_src Syntax.read_prop;
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   626
59845
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   627
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   628
(* ind_cases *)
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   629
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   630
fun ind_cases_rules ctxt raw_props raw_fixes =
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   631
  let
63180
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63064
diff changeset
   632
    val (props, ctxt') = Specification.read_props raw_props raw_fixes ctxt;
59845
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   633
    val rules = Proof_Context.export ctxt' ctxt (map (mk_cases ctxt') props);
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   634
  in rules end;
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   635
58815
fd3f893a40ea modernized setup;
wenzelm
parents: 58028
diff changeset
   636
val _ =
fd3f893a40ea modernized setup;
wenzelm
parents: 58028
diff changeset
   637
  Theory.setup
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   638
    (Method.setup \<^binding>\<open>ind_cases\<close>
59845
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   639
      (Scan.lift (Scan.repeat1 Parse.prop -- Parse.for_fixes) >>
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   640
        (fn (props, fixes) => fn ctxt =>
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   641
          Method.erule ctxt 0 (ind_cases_rules ctxt props fixes)))
fafb4d12c307 ind_cases: clarified preparation of arguments;
wenzelm
parents: 59642
diff changeset
   642
      "case analysis for inductive definitions, based on simplified elimination rule");
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   643
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   644
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   645
(* derivation of simplified equation *)
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   646
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   647
fun mk_simp_eq ctxt prop =
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   648
  let
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   649
    val thy = Proof_Context.theory_of ctxt;
70308
7f568724d67e clarified signature;
wenzelm
parents: 69829
diff changeset
   650
    val ctxt' = Proof_Context.augment prop ctxt;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   651
    val lhs_of = fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of;
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   652
    val substs =
67637
e6bcd14139fc trim context of persistent data;
wenzelm
parents: 67149
diff changeset
   653
      retrieve_equations ctxt (HOLogic.dest_Trueprop prop)
38665
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 38388
diff changeset
   654
      |> map_filter
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 38388
diff changeset
   655
        (fn eq => SOME (Pattern.match thy (lhs_of eq, HOLogic.dest_Trueprop prop)
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 38388
diff changeset
   656
            (Vartab.empty, Vartab.empty), eq)
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   657
          handle Pattern.MATCH => NONE);
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   658
    val (subst, eq) =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   659
      (case substs of
38665
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 38388
diff changeset
   660
        [s] => s
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 38388
diff changeset
   661
      | _ => error
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   662
        ("equations matching pattern " ^ Syntax.string_of_term ctxt prop ^ " is not unique"));
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   663
    val inst =
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60362
diff changeset
   664
      map (fn v => (fst v, Thm.cterm_of ctxt' (Envir.subst_term subst (Var v))))
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   665
        (Term.add_vars (lhs_of eq) []);
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   666
  in
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60362
diff changeset
   667
    infer_instantiate ctxt' inst eq
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60362
diff changeset
   668
    |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.full_rewrite ctxt')))
70318
9eff9e2dc177 proper Proof_Context.export_morphism corresponding to Proof_Context.augment (see 7f568724d67e);
wenzelm
parents: 70308
diff changeset
   669
    |> singleton (Proof_Context.export ctxt' ctxt)
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   670
  end
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   671
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   672
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   673
(* inductive simps *)
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   674
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   675
fun gen_inductive_simps prep_att prep_prop args int lthy =
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   676
  let
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   677
    val facts = args |> map (fn ((a, atts), props) =>
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55111
diff changeset
   678
      ((a, map (prep_att lthy) atts),
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   679
        map (Thm.no_attributes o single o mk_simp_eq lthy o prep_prop lthy) props));
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   680
    val (res, lthy') = lthy |> Local_Theory.notes facts
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   681
    val _ =
82317
231b6d8231c6 more uniform Proof_Display.print_results for theory and proof output --- avoid loss of information seen in src/Doc/JEdit/document/output-and-state.png (the first bad changeset is f8c412a45af8, see also 53b59fa42696);
wenzelm
parents: 81810
diff changeset
   682
      Proof_Display.print_results {interactive = int, pos = Position.thread_data ()}
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   683
        lthy' ((Thm.theoremK, ""), res)
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
   684
  in (res, lthy') end;
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
   685
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
   686
val inductive_simps = gen_inductive_simps (K I) Syntax.check_prop;
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
   687
val inductive_simps_cmd = gen_inductive_simps Attrib.check_src Syntax.read_prop;
40902
bulwahn
parents: 40316
diff changeset
   688
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   689
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   690
(* prove induction rule *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   691
26477
ecf06644f6cb eliminated quiete_mode ref (turned into proper argument);
wenzelm
parents: 26336
diff changeset
   692
fun prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   693
    fp_def rec_preds_defs ctxt ctxt''' =  (* FIXME ctxt''' ?? *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   694
  let
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51798
diff changeset
   695
    val _ = clean_message ctxt quiet_mode "  Proving the induction rule ...";
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   696
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   697
    (* predicates for induction rule *)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   698
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
   699
    val (pnames, ctxt') = Variable.variant_fixes (mk_names "P" (length cs)) ctxt;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   700
    val preds =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   701
      map2 (curry Free) pnames
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   702
        (map (fn c => arg_types_of (length params) c ---> HOLogic.boolT) cs);
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   703
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   704
    (* transform an introduction rule into a premise for induction rule *)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   705
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   706
    fun mk_ind_prem r =
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   707
      let
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   708
        fun subst s =
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   709
          (case dest_predicate cs params s of
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   710
            SOME (_, i, ys, (_, Ts)) =>
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   711
              let
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   712
                val k = length Ts;
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   713
                val bs = map Bound (k - 1 downto 0);
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
   714
                val P = list_comb (nth preds i, map (incr_boundvars k) ys @ bs);
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   715
                val Q =
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   716
                  fold_rev Term.abs (mk_names "x" k ~~ Ts)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   717
                    (HOLogic.mk_binop \<^const_name>\<open>HOL.induct_conj\<close>
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   718
                      (list_comb (incr_boundvars k s, bs), P));
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   719
              in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   720
          | NONE =>
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   721
              (case s of
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   722
                t $ u => (fst (subst t) $ fst (subst u), NONE)
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   723
              | Abs (a, T, t) => (Abs (a, T, fst (subst t)), NONE)
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   724
              | _ => (s, NONE)));
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
   725
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   726
        fun mk_prem s prems =
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   727
          (case subst s of
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   728
            (_, SOME (t, u)) => t :: u :: prems
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   729
          | (t, _) => t :: prems);
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   730
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   731
        val SOME (_, i, ys, _) =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   732
          dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r));
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
   733
      in
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
   734
        fold_rev (Logic.all o Free) (Logic.strip_params r)
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
   735
          (Logic.list_implies (map HOLogic.mk_Trueprop (fold_rev mk_prem
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
   736
            (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r)) []),
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
   737
              HOLogic.mk_Trueprop (list_comb (nth preds i, ys))))
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   738
      end;
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   739
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   740
    val ind_prems = map mk_ind_prem intr_ts;
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   741
21526
1e6bd5ed7abc added morh_result, the_inductive, add_inductive_global;
wenzelm
parents: 21508
diff changeset
   742
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   743
    (* make conclusions for induction rules *)
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   744
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   745
    val Tss = map (binder_types o fastype_of) preds;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   746
    val (xnames, ctxt'') = Variable.variant_fixes (mk_names "x" (length (flat Tss))) ctxt';
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   747
    val mutual_ind_concl =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   748
      HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   749
        (map (fn (((xnames, Ts), c), P) =>
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   750
          let val frees = map Free (xnames ~~ Ts)
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   751
          in HOLogic.mk_imp (list_comb (c, params @ frees), list_comb (P, frees)) end)
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   752
        (unflat Tss xnames ~~ Tss ~~ cs ~~ preds)));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   753
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   754
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   755
    (* make predicate for instantiation of abstract induction rule *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   756
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   757
    val ind_pred =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   758
      fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   759
        (map_index (fn (i, P) => fold_rev (curry HOLogic.mk_imp)
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   760
           (make_bool_args HOLogic.mk_not I bs i)
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   761
           (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))) preds));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   762
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   763
    val ind_concl =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   764
      HOLogic.mk_Trueprop
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   765
        (HOLogic.mk_binrel \<^const_name>\<open>Orderings.less_eq\<close> (rec_const, ind_pred));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   766
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   767
    val raw_fp_induct = mono RS (fp_def RS @{thm def_lfp_induct});
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   768
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 50771
diff changeset
   769
    val induct = Goal.prove_sorry ctxt'' [] ind_prems ind_concl
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   770
      (fn {context = ctxt3, prems} => EVERY
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   771
        [rewrite_goals_tac ctxt3 [inductive_conj_def],
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   772
         DETERM (resolve_tac ctxt3 [raw_fp_induct] 1),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   773
         REPEAT (resolve_tac ctxt3 [@{thm le_funI}, @{thm le_boolI}] 1),
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   774
         rewrite_goals_tac ctxt3 simp_thms2,
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   775
         (*This disjE separates out the introduction rules*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   776
         REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [disjE, exE, FalseE])),
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   777
         (*Now break down the individual cases.  No disjE here in case
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   778
           some premise involves disjunction.*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   779
         REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [conjE] ORELSE' bound_hyp_subst_tac ctxt3)),
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   780
         REPEAT (FIRSTGOAL
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   781
           (resolve_tac ctxt3 [conjI, impI] ORELSE'
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   782
           (eresolve_tac ctxt3 [notE] THEN' assume_tac ctxt3))),
59499
14095f771781 misc tuning;
wenzelm
parents: 59498
diff changeset
   783
         EVERY (map (fn prem =>
14095f771781 misc tuning;
wenzelm
parents: 59498
diff changeset
   784
            DEPTH_SOLVE_1 (assume_tac ctxt3 1 ORELSE
14095f771781 misc tuning;
wenzelm
parents: 59498
diff changeset
   785
              resolve_tac ctxt3
14095f771781 misc tuning;
wenzelm
parents: 59498
diff changeset
   786
                [rewrite_rule ctxt3 (inductive_conj_def :: rec_preds_defs @ simp_thms2) prem,
14095f771781 misc tuning;
wenzelm
parents: 59498
diff changeset
   787
                  conjI, refl] 1)) prems)]);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   788
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 50771
diff changeset
   789
    val lemma = Goal.prove_sorry ctxt'' [] []
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   790
      (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn {context = ctxt3, ...} => EVERY
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   791
        [rewrite_goals_tac ctxt3 rec_preds_defs,
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   792
         REPEAT (EVERY
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   793
           [REPEAT (resolve_tac ctxt3 [conjI, impI] 1),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59060
diff changeset
   794
            REPEAT (eresolve_tac ctxt3 [@{thm le_funE}, @{thm le_boolE}] 1),
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58936
diff changeset
   795
            assume_tac ctxt3 1,
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53995
diff changeset
   796
            rewrite_goals_tac ctxt3 simp_thms1,
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58936
diff changeset
   797
            assume_tac ctxt3 1])]);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   798
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   799
  in singleton (Proof_Context.export ctxt'' ctxt''') (induct RS lemma) end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   800
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   801
(* prove coinduction rule *)
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   802
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   803
fun mk_If p t f = let val T = fastype_of t in \<^Const>\<open>If T\<close> $ p $ t $ f end;
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   804
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   805
fun prove_coindrule quiet_mode preds cs argTs bs xs params intr_ts mono
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   806
    fp_def rec_preds_defs ctxt ctxt''' =  (* FIXME ctxt''' ?? *)
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   807
  let
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   808
    val _ = clean_message ctxt quiet_mode "  Proving the coinduction rule ...";
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   809
    val n = length cs;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   810
    val (ns, xss) = map_split (fn pred =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   811
      make_args' argTs xs (arg_types_of (length params) pred) |> `length) preds;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   812
    val xTss = map (map fastype_of) xss;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   813
    val (Rs_names, names_ctxt) = Variable.variant_fixes (mk_names "X" n) ctxt;
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   814
    val Rs = map2 (fn name => fn Ts => Free (name, Ts ---> \<^Type>\<open>bool\<close>)) Rs_names xTss;
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   815
    val Rs_applied = map2 (curry list_comb) Rs xss;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   816
    val preds_applied = map2 (curry list_comb) (map (fn p => list_comb (p, params)) preds) xss;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   817
    val abstract_list = fold_rev (absfree o dest_Free);
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   818
    val bss = map (make_bool_args
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   819
      (fn b => HOLogic.mk_eq (b, \<^Const>\<open>False\<close>))
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   820
      (fn b => HOLogic.mk_eq (b, \<^Const>\<open>True\<close>)) bs) (0 upto n - 1);
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   821
    val eq_undefinedss = map (fn ys => map (fn x =>
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   822
        let val T = fastype_of x
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   823
        in \<^Const>\<open>HOL.eq T for x \<^Const>\<open>undefined T\<close>\<close> end)
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   824
      (subtract (op =) ys xs)) xss;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   825
    val R =
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   826
      @{fold 3} (fn bs => fn eqs => fn R => fn t => if null bs andalso null eqs then R else
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   827
        mk_If (Library.foldr1 HOLogic.mk_conj (bs @ eqs)) R t)
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   828
      bss eq_undefinedss Rs_applied \<^Const>\<open>False\<close>
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   829
      |> abstract_list (bs @ xs);
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   830
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   831
    fun subst t =
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   832
      (case dest_predicate cs params t of
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   833
        SOME (_, i, ts, (_, Us)) =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   834
          let
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   835
            val l = length Us;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   836
            val bs = map Bound (l - 1 downto 0);
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   837
            val args = map (incr_boundvars l) ts @ bs
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   838
          in
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   839
            HOLogic.mk_disj (list_comb (nth Rs i, args),
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   840
              list_comb (nth preds i, params @ args))
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   841
            |> fold_rev absdummy Us
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   842
          end
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   843
      | NONE =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   844
          (case t of
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   845
            t1 $ t2 => subst t1 $ subst t2
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   846
          | Abs (x, T, u) => Abs (x, T, subst u)
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   847
          | _ => t));
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   848
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   849
    fun mk_coind_prem r =
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   850
      let
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   851
        val SOME (_, i, ts, (Ts, _)) =
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   852
          dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r));
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   853
        val ps =
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   854
          map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   855
          map (subst o HOLogic.dest_Trueprop) (Logic.strip_assums_hyp r);
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   856
      in
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   857
        (i, fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P))
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   858
          (Logic.strip_params r)
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   859
          (if null ps then \<^Const>\<open>True\<close> else foldr1 HOLogic.mk_conj ps))
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   860
      end;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   861
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   862
    fun mk_prem i Ps = Logic.mk_implies
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   863
        ((nth Rs_applied i, Library.foldr1 HOLogic.mk_disj Ps) |> @{apply 2} HOLogic.mk_Trueprop)
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   864
      |> fold_rev Logic.all (nth xss i);
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   865
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   866
    val prems = map mk_coind_prem intr_ts |> AList.group (op =) |> sort (int_ord o apply2 fst)
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   867
      |> map (uncurry mk_prem);
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   868
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   869
    val concl = @{map 3} (fn xs =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   870
        Ctr_Sugar_Util.list_all_free xs oo curry HOLogic.mk_imp) xss Rs_applied preds_applied
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   871
      |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   872
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   873
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   874
    val pred_defs_sym = if null rec_preds_defs then [] else map2 (fn n => fn thm =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   875
        funpow n (fn thm => thm RS @{thm meta_fun_cong}) thm RS @{thm Pure.symmetric})
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   876
      ns rec_preds_defs;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   877
    val simps = simp_thms3 @ pred_defs_sym;
80701
39cd50407f79 tuned: prefer canonical argument order;
wenzelm
parents: 80636
diff changeset
   878
    val simproc = Simplifier.the_simproc ctxt "HOL.defined_All";
39cd50407f79 tuned: prefer canonical argument order;
wenzelm
parents: 80636
diff changeset
   879
    val simplify =
39cd50407f79 tuned: prefer canonical argument order;
wenzelm
parents: 80636
diff changeset
   880
      asm_full_simplify (Ctr_Sugar_Util.ss_only simps ctxt |> Simplifier.add_proc simproc);
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   881
    val coind = (mono RS (fp_def RS @{thm def_coinduct}))
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   882
      |> infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt R)]
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   883
      |> simplify;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   884
    fun idx_of t = find_index (fn R =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   885
      R = the_single (subtract (op =) (preds @ params) (map Free (Term.add_frees t [])))) Rs;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   886
    val coind_concls = HOLogic.dest_Trueprop (Thm.concl_of coind) |> HOLogic.dest_conj
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   887
      |> map (fn t => (idx_of t, t)) |> sort (int_ord o @{apply 2} fst) |> map snd;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   888
    val reorder_bound_goals = map_filter (fn (t, u) => if t aconv u then NONE else
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   889
      SOME (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u))))
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   890
      ((HOLogic.dest_Trueprop concl |> HOLogic.dest_conj) ~~ coind_concls);
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   891
    val reorder_bound_thms = map (fn goal => Goal.prove_sorry ctxt [] [] goal
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   892
      (fn {context = ctxt, prems = _} =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   893
        HEADGOAL (EVERY' [resolve_tac ctxt [iffI],
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   894
          REPEAT_DETERM o resolve_tac ctxt [allI, impI],
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   895
          REPEAT_DETERM o dresolve_tac ctxt [spec], eresolve_tac ctxt [mp], assume_tac ctxt,
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   896
          REPEAT_DETERM o resolve_tac ctxt [allI, impI],
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   897
          REPEAT_DETERM o dresolve_tac ctxt [spec], eresolve_tac ctxt [mp], assume_tac ctxt])))
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   898
      reorder_bound_goals;
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   899
    val coinduction = Goal.prove_sorry ctxt [] prems concl (fn {context = ctxt, prems = CIH} =>
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   900
      HEADGOAL (full_simp_tac
80701
39cd50407f79 tuned: prefer canonical argument order;
wenzelm
parents: 80636
diff changeset
   901
        (Ctr_Sugar_Util.ss_only (simps @ reorder_bound_thms) ctxt
39cd50407f79 tuned: prefer canonical argument order;
wenzelm
parents: 80636
diff changeset
   902
          |> Simplifier.add_proc simproc) THEN'
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   903
        resolve_tac ctxt [coind]) THEN
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   904
      ALLGOALS (REPEAT_ALL_NEW (REPEAT_DETERM o resolve_tac ctxt [allI, impI, conjI] THEN'
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   905
        REPEAT_DETERM o eresolve_tac ctxt [exE, conjE] THEN'
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   906
        dresolve_tac ctxt (map simplify CIH) THEN'
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   907
        REPEAT_DETERM o (assume_tac ctxt ORELSE'
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   908
          eresolve_tac ctxt [conjE] ORELSE' dresolve_tac ctxt [spec, mp]))))
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   909
  in
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   910
    coinduction
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   911
    |> length cs = 1 ? (Object_Logic.rulify ctxt #> rotate_prems ~1)
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   912
    |> singleton (Proof_Context.export names_ctxt ctxt''')
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   913
  end
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   914
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
   915
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   916
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   917
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   918
(** specification of (co)inductive predicates **)
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   919
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
   920
fun mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts monos params cnames_syn lthy =
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
   921
  let
33077
3c9cf88ec841 arg_types_of auxiliary function; using multiset operations
haftmann
parents: 33056
diff changeset
   922
    val argTs = fold (combine (op =) o arg_types_of (length params)) cs [];
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   923
    val k = log 2 1 (length cs);
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   924
    val predT = replicate k HOLogic.boolT ---> argTs ---> HOLogic.boolT;
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   925
    val fp_const = if coind then \<^Const>\<open>gfp predT\<close> else \<^Const>\<open>lfp predT\<close>;
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   926
81519
cdc43c0fdbfc clarified signature;
wenzelm
parents: 80727
diff changeset
   927
    val ctxt1 = fold Variable.declare_names intr_ts lthy
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   928
    val p :: xs =
81519
cdc43c0fdbfc clarified signature;
wenzelm
parents: 80727
diff changeset
   929
      map Free (Variable.variant_names ctxt1
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   930
        (("p", predT) :: (mk_names "x" (length argTs) ~~ argTs)));
81519
cdc43c0fdbfc clarified signature;
wenzelm
parents: 80727
diff changeset
   931
    val ctxt2 = fold Variable.declare_names (p :: xs) ctxt1
cdc43c0fdbfc clarified signature;
wenzelm
parents: 80727
diff changeset
   932
    val bs = map Free (Variable.variant_names ctxt2 (map (rpair HOLogic.boolT) (mk_names "b" k)));
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   933
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
   934
    fun subst t =
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
   935
      (case dest_predicate cs params t of
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   936
        SOME (_, i, ts, (Ts, Us)) =>
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
   937
          let
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
   938
            val l = length Us;
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   939
            val zs = map Bound (l - 1 downto 0);
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   940
          in
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   941
            fold_rev (Term.abs o pair "z") Us
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   942
              (list_comb (p,
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   943
                make_bool_args' bs i @ make_args argTs
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 46218
diff changeset
   944
                  ((map (incr_boundvars l) ts ~~ Ts) @ (zs ~~ Us))))
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   945
          end
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   946
      | NONE =>
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   947
          (case t of
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   948
            t1 $ t2 => subst t1 $ subst t2
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   949
          | Abs (x, T, u) => Abs (x, T, subst u)
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   950
          | _ => t));
5149
10f0be29c0d1 Fixed bug in transform_rule.
berghofe
parents: 5120
diff changeset
   951
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   952
    (* transform an introduction rule into a conjunction  *)
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   953
    (*   [| p_i t; ... |] ==> p_j u                       *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   954
    (* is transformed into                                *)
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
   955
    (*   b_j & x_j = u & p b_j t & ...                    *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   956
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   957
    fun transform_rule r =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   958
      let
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   959
        val SOME (_, i, ts, (Ts, _)) =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   960
          dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r));
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   961
        val ps =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   962
          make_bool_args HOLogic.mk_not I bs i @
21048
e57e91f72831 Restructured and repaired code dealing with case names
berghofe
parents: 21024
diff changeset
   963
          map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   964
          map (subst o HOLogic.dest_Trueprop) (Logic.strip_assums_hyp r);
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   965
      in
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   966
        fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P))
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   967
          (Logic.strip_params r)
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   968
          (if null ps then \<^Const>\<open>True\<close> else foldr1 HOLogic.mk_conj ps)
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   969
      end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   970
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   971
    (* make a disjunction of all introduction rules *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   972
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   973
    val fp_fun =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   974
      fold_rev lambda (p :: bs @ xs)
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   975
        (if null intr_ts then \<^Const>\<open>False\<close>
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   976
         else foldr1 HOLogic.mk_disj (map transform_rule intr_ts));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   977
61308
bb0596c7f921 avoid useless empty case_names;
wenzelm
parents: 61268
diff changeset
   978
    (* add definition of recursive predicates to theory *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   979
61948
wenzelm
parents: 61681
diff changeset
   980
    val is_auxiliary = length cs > 1;
wenzelm
parents: 61681
diff changeset
   981
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
   982
    val rec_binding =
63006
89d19aa73081 clarified bindings;
wenzelm
parents: 62969
diff changeset
   983
      if Binding.is_empty alt_name then Binding.conglomerate (map #1 cnames_syn) else alt_name;
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
   984
    val rec_name = Binding.name_of rec_binding;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   985
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61951
diff changeset
   986
    val internals = Config.get lthy inductive_internals;
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
   987
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
   988
    val ((rec_const, (_, fp_def)), lthy') = lthy
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
   989
      |> is_auxiliary ? Proof_Context.concealed
33766
c679f05600cd adapted Local_Theory.define -- eliminated odd thm kind;
wenzelm
parents: 33726
diff changeset
   990
      |> Local_Theory.define
61951
cbd310584cff clarified position information;
wenzelm
parents: 61948
diff changeset
   991
        ((rec_binding, case cnames_syn of [(_, mx)] => mx | _ => NoSyn),
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61951
diff changeset
   992
         ((Thm.make_def_binding internals rec_binding, @{attributes [nitpick_unfold]}),
80727
49067bf1cf92 tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
   993
           fold_rev lambda params (fp_const $ fp_fun)))
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
   994
      ||> Proof_Context.restore_naming lthy;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
   995
    val fp_def' =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
   996
      Simplifier.rewrite (put_simpset HOL_basic_ss lthy' addsimps [fp_def])
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   997
        (Thm.cterm_of lthy' (list_comb (rec_const, params)));
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33171
diff changeset
   998
    val specs =
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
   999
      if is_auxiliary then
61951
cbd310584cff clarified position information;
wenzelm
parents: 61948
diff changeset
  1000
        map_index (fn (i, ((b, mx), c)) =>
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33171
diff changeset
  1001
          let
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33171
diff changeset
  1002
            val Ts = arg_types_of (length params) c;
81519
cdc43c0fdbfc clarified signature;
wenzelm
parents: 80727
diff changeset
  1003
            val ctxt = fold Variable.declare_names intr_ts lthy';
cdc43c0fdbfc clarified signature;
wenzelm
parents: 80727
diff changeset
  1004
            val xs = map Free (Variable.variant_names ctxt (mk_names "x" (length Ts) ~~ Ts));
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33171
diff changeset
  1005
          in
61951
cbd310584cff clarified position information;
wenzelm
parents: 61948
diff changeset
  1006
            ((b, mx),
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61951
diff changeset
  1007
              ((Thm.make_def_binding internals b, []), fold_rev lambda (params @ xs)
61951
cbd310584cff clarified position information;
wenzelm
parents: 61948
diff changeset
  1008
                (list_comb (rec_const, params @ make_bool_args' bs i @
cbd310584cff clarified position information;
wenzelm
parents: 61948
diff changeset
  1009
                  make_args argTs (xs ~~ Ts)))))
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1010
          end) (cnames_syn ~~ cs)
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1011
      else [];
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1012
    val (consts_defs, lthy'') = lthy'
39248
4a3747517552 only conceal primitive definition theorems, not predicate names
haftmann
parents: 38864
diff changeset
  1013
      |> fold_map Local_Theory.define specs;
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
  1014
    val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1015
61063
d0c21a68d9c6 clarified context;
wenzelm
parents: 60784
diff changeset
  1016
    val (_, ctxt'') = Variable.add_fixes (map (fst o dest_Free) params) lthy'';
d0c21a68d9c6 clarified context;
wenzelm
parents: 60784
diff changeset
  1017
    val mono = prove_mono quiet_mode skip_mono predT fp_fun monos ctxt'';
d0c21a68d9c6 clarified context;
wenzelm
parents: 60784
diff changeset
  1018
    val (_, lthy''') = lthy''
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1019
      |> Local_Theory.note
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61951
diff changeset
  1020
        ((if internals
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1021
          then Binding.qualify true rec_name (Binding.name "mono")
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1022
          else Binding.empty, []),
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1023
          Proof_Context.export ctxt'' lthy'' [mono]);
61063
d0c21a68d9c6 clarified context;
wenzelm
parents: 60784
diff changeset
  1024
  in
d0c21a68d9c6 clarified context;
wenzelm
parents: 60784
diff changeset
  1025
    (lthy''', Proof_Context.transfer (Proof_Context.theory_of lthy''') ctxt'',
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1026
      rec_binding, mono, fp_def', map (#2 o #2) consts_defs,
61063
d0c21a68d9c6 clarified context;
wenzelm
parents: 60784
diff changeset
  1027
      list_comb (rec_const, params), preds, argTs, bs, xs)
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
  1028
  end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1029
71179
592e2afdd50c more informative spec rules: optional name;
wenzelm
parents: 70318
diff changeset
  1030
fun declare_rules rec_binding coind no_ind spec_name cnames
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1031
    preds intrs intr_bindings intr_atts elims eqs raw_induct lthy =
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1032
  let
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1033
    val rec_name = Binding.name_of rec_binding;
32773
f6fd4ccd8eea mandatory prefix where appropriate
haftmann
parents: 32652
diff changeset
  1034
    fun rec_qualified qualified = Binding.qualify qualified rec_name;
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1035
    val intr_names = map Binding.name_of intr_bindings;
61308
bb0596c7f921 avoid useless empty case_names;
wenzelm
parents: 61268
diff changeset
  1036
    val ind_case_names =
bb0596c7f921 avoid useless empty case_names;
wenzelm
parents: 61268
diff changeset
  1037
      if forall (equal "") intr_names then []
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1038
      else [Attrib.case_names intr_names];
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1039
    val induct =
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1040
      if coind then
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 50302
diff changeset
  1041
        (raw_induct,
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1042
          [Attrib.case_names [rec_name],
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1043
           Attrib.case_conclusion (rec_name, intr_names),
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1044
           Attrib.consumes (1 - Thm.nprems_of raw_induct),
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 76054
diff changeset
  1045
           Attrib.internal \<^here> (K (Induct.coinduct_pred (hd cnames)))])
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1046
      else if no_ind orelse length cnames > 1 then
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1047
        (raw_induct, ind_case_names @ [Attrib.consumes (~ (Thm.nprems_of raw_induct))])
50771
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 50302
diff changeset
  1048
      else
2852f997bfb5 prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
wenzelm
parents: 50302
diff changeset
  1049
        (raw_induct RSN (2, rev_mp),
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1050
          ind_case_names @ [Attrib.consumes (~ (Thm.nprems_of raw_induct))]);
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1051
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1052
    val (intrs', lthy1) =
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1053
      lthy |>
71179
592e2afdd50c more informative spec rules: optional name;
wenzelm
parents: 70318
diff changeset
  1054
      Spec_Rules.add spec_name
592e2afdd50c more informative spec rules: optional name;
wenzelm
parents: 70318
diff changeset
  1055
        (if coind then Spec_Rules.Co_Inductive else Spec_Rules.Inductive) preds intrs |>
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
  1056
      Local_Theory.notes
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33171
diff changeset
  1057
        (map (rec_qualified false) intr_bindings ~~ intr_atts ~~
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1058
          map (fn th => [([th], @{attributes [Pure.intro?]})]) intrs) |>>
24744
dcb8cf5fc99c - add_inductive_i now takes typ instead of typ option as argument
berghofe
parents: 24721
diff changeset
  1059
      map (hd o snd);
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1060
    val (((_, elims'), (_, [induct'])), lthy2) =
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1061
      lthy1 |>
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
  1062
      Local_Theory.note ((rec_qualified true (Binding.name "intros"), []), intrs') ||>>
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 33966
diff changeset
  1063
      fold_map (fn (name, (elim, cases, k)) =>
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
  1064
        Local_Theory.note
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1065
          ((Binding.qualify true (Long_Name.base_name name) (Binding.name "cases"),
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1066
            ((if forall (equal "") cases then [] else [Attrib.case_names cases]) @
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1067
              [Attrib.consumes (1 - Thm.nprems_of elim), Attrib.constraints k,
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 76054
diff changeset
  1068
               Attrib.internal \<^here> (K (Induct.cases_pred name))] @ @{attributes [Pure.elim?]})),
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1069
            [elim]) #>
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1070
        apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>>
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
  1071
      Local_Theory.note
61308
bb0596c7f921 avoid useless empty case_names;
wenzelm
parents: 61268
diff changeset
  1072
        ((rec_qualified true (Binding.name (coind_prefix coind ^ "induct")), #2 induct),
bb0596c7f921 avoid useless empty case_names;
wenzelm
parents: 61268
diff changeset
  1073
          [rulify lthy1 (#1 induct)]);
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1074
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1075
    val (eqs', lthy3) = lthy2 |>
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1076
      fold_map (fn (name, eq) => Local_Theory.note
38665
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 38388
diff changeset
  1077
          ((Binding.qualify true (Long_Name.base_name name) (Binding.name "simps"),
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 76054
diff changeset
  1078
            [Attrib.internal \<^here> (K equation_add_permissive)]), [eq])
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1079
          #> apfst (hd o snd))
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1080
        (if null eqs then [] else (cnames ~~ eqs))
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1081
    val (inducts, lthy4) =
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1082
      if no_ind orelse coind then ([], lthy3)
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1083
      else
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1084
        let val inducts = cnames ~~ Project_Rule.projects lthy3 (1 upto length cnames) induct' in
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1085
          lthy3 |>
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
  1086
          Local_Theory.notes [((rec_qualified true (Binding.name "inducts"), []),
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1087
            inducts |> map (fn (name, th) => ([th],
61308
bb0596c7f921 avoid useless empty case_names;
wenzelm
parents: 61268
diff changeset
  1088
              ind_case_names @
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 63006
diff changeset
  1089
                [Attrib.consumes (1 - Thm.nprems_of th),
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 76054
diff changeset
  1090
                 Attrib.internal \<^here> (K (Induct.induct_pred name))])))] |>> snd o hd
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1091
        end;
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1092
  in (intrs', elims', eqs', induct', inducts, lthy4) end;
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1093
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1094
type flags =
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
  1095
  {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool,
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
  1096
    no_elim: bool, no_ind: bool, skip_mono: bool};
26534
a2cb4de2a1aa Added skip_mono flag and inductive_flags type.
berghofe
parents: 26477
diff changeset
  1097
a2cb4de2a1aa Added skip_mono flag and inductive_flags type.
berghofe
parents: 26477
diff changeset
  1098
type add_ind_def =
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1099
  flags ->
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
  1100
  term list -> (Attrib.binding * term) list -> thm list ->
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29388
diff changeset
  1101
  term list -> (binding * mixfix) list ->
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1102
  local_theory -> result * local_theory;
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1103
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
  1104
fun add_ind_def {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono}
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1105
    cs intros monos params cnames_syn lthy =
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
  1106
  let
25288
6e0c8dd213df improved error message for missing predicates;
wenzelm
parents: 25143
diff changeset
  1107
    val _ = null cnames_syn andalso error "No inductive predicates given";
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1108
    val names = map (Binding.name_of o fst) cnames_syn;
26477
ecf06644f6cb eliminated quiete_mode ref (turned into proper argument);
wenzelm
parents: 26336
diff changeset
  1109
    val _ = message (quiet_mode andalso not verbose)
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27882
diff changeset
  1110
      ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names);
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
  1111
71214
5727bcc3c47c proper spec_rule name via naming/binding/Morphism.binding;
wenzelm
parents: 71179
diff changeset
  1112
    val spec_name = Binding.conglomerate (map #1 cnames_syn);
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
  1113
    val cnames = map (Local_Theory.full_name lthy o #1) cnames_syn;  (* FIXME *)
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1114
    val ((intr_names, intr_atts), intr_ts) =
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1115
      apfst split_list (split_list (map (check_rule lthy cs params) intros));
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
  1116
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1117
    val (lthy1, lthy2, rec_binding, mono, fp_def, rec_preds_defs, rec_const, preds,
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
  1118
      argTs, bs, xs) = mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1119
        monos params cnames_syn lthy;
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
  1120
26477
ecf06644f6cb eliminated quiete_mode ref (turned into proper argument);
wenzelm
parents: 26336
diff changeset
  1121
    val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs)
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
  1122
      intr_ts rec_preds_defs lthy2 lthy1;
33459
wenzelm
parents: 33458
diff changeset
  1123
    val elims =
wenzelm
parents: 33458
diff changeset
  1124
      if no_elim then []
wenzelm
parents: 33458
diff changeset
  1125
      else
wenzelm
parents: 33458
diff changeset
  1126
        prove_elims quiet_mode cs params intr_ts (map Binding.name_of intr_names)
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
  1127
          unfold rec_preds_defs lthy2 lthy1;
22605
41b092e7d89a - Removed occurrences of ProofContext.export in add_ind_def that
berghofe
parents: 22460
diff changeset
  1128
    val raw_induct = zero_var_indexes
33459
wenzelm
parents: 33458
diff changeset
  1129
      (if no_ind then Drule.asm_rl
wenzelm
parents: 33458
diff changeset
  1130
       else if coind then
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
  1131
         prove_coindrule quiet_mode preds cs argTs bs xs params intr_ts mono fp_def
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
  1132
           rec_preds_defs lthy2 lthy1
21024
63ab84bb64d1 Completely rewrote inductive definition package. Now allows to
berghofe
parents: 20901
diff changeset
  1133
       else
26477
ecf06644f6cb eliminated quiete_mode ref (turned into proper argument);
wenzelm
parents: 26336
diff changeset
  1134
         prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
  1135
           rec_preds_defs lthy2 lthy1);
63863
d14e580c3b8f don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents: 63395
diff changeset
  1136
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1137
    val eqs =
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1138
      if no_elim then [] else prove_eqs quiet_mode cs params intr_ts intrs elims lthy2 lthy1;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1139
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
  1140
    val elims' = map (fn (th, ns, i) => (rulify lthy1 th, ns, i)) elims;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
  1141
    val intrs' = map (rulify lthy1) intrs;
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1142
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1143
    val (intrs'', elims'', eqs', induct, inducts, lthy3) =
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61308
diff changeset
  1144
      declare_rules rec_binding coind no_ind
71179
592e2afdd50c more informative spec rules: optional name;
wenzelm
parents: 70318
diff changeset
  1145
        spec_name cnames preds intrs' intr_names intr_atts elims' eqs raw_induct lthy1;
21048
e57e91f72831 Restructured and repaired code dealing with case names
berghofe
parents: 21024
diff changeset
  1146
e57e91f72831 Restructured and repaired code dealing with case names
berghofe
parents: 21024
diff changeset
  1147
    val result =
e57e91f72831 Restructured and repaired code dealing with case names
berghofe
parents: 21024
diff changeset
  1148
      {preds = preds,
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1149
       intrs = intrs'',
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1150
       elims = elims'',
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51658
diff changeset
  1151
       raw_induct = rulify lthy3 raw_induct,
35646
b32d6c1bdb4d Added inducts field to inductive_result.
berghofe
parents: 35625
diff changeset
  1152
       induct = induct,
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1153
       inducts = inducts,
81810
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
  1154
       def = fp_def,
6cd42e67c0a8 store the {l,g}fp-definition and the monotonicity theorem for inductive predicates (by Jan van Brügge)
traytel
parents: 81525
diff changeset
  1155
       mono = mono,
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1156
       eqs = eqs'};
21367
7a0cc1bb4dcc inductive: canonical specification syntax (flattened result only);
wenzelm
parents: 21350
diff changeset
  1157
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
  1158
    val lthy4 = lthy3
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 76054
diff changeset
  1159
      |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi =>
45290
f599ac41e7f5 tuned signature -- refined terminology;
wenzelm
parents: 44868
diff changeset
  1160
        let val result' = transform_result phi result;
65411
3c628937899d use Item_Net to store inductive info
Lars Hupel <lars.hupel@mytum.de>
parents: 63863
diff changeset
  1161
        in put_inductives ({names = cnames, coind = coind}, result') end);
36642
084470c3cea2 Corrected handling of "for" parameters.
berghofe
parents: 36546
diff changeset
  1162
  in (result, lthy4) end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1163
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
  1164
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
  1165
(* external interfaces *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1166
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1167
fun gen_add_inductive mk_def
59059
haftmann
parents: 58993
diff changeset
  1168
    flags cnames_syn pnames spec monos lthy =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1169
  let
21766
3eb48154388e Abbreviations can now be specified simultaneously
berghofe
parents: 21658
diff changeset
  1170
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1171
    (* abbrevs *)
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1172
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1173
    val (_, ctxt1) = Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn) lthy;
21766
3eb48154388e Abbreviations can now be specified simultaneously
berghofe
parents: 21658
diff changeset
  1174
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1175
    fun get_abbrev ((name, atts), t) =
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1176
      if can (Logic.strip_assums_concl #> Logic.dest_equals) t then
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1177
        let
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28965
diff changeset
  1178
          val _ = Binding.is_empty name andalso null atts orelse
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1179
            error "Abbreviations may not have names or attributes";
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63285
diff changeset
  1180
          val ((x, T), rhs) = Local_Defs.abs_def (snd (Local_Defs.cert_def ctxt1 (K []) t));
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27882
diff changeset
  1181
          val var =
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1182
            (case find_first (fn ((c, _), _) => Binding.name_of c = x) cnames_syn of
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1183
              NONE => error ("Undeclared head of abbreviation " ^ quote x)
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27882
diff changeset
  1184
            | SOME ((b, T'), mx) =>
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1185
                if T <> T' then error ("Bad type specification for abbreviation " ^ quote x)
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27882
diff changeset
  1186
                else (b, mx));
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27882
diff changeset
  1187
        in SOME (var, rhs) end
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1188
      else NONE;
21766
3eb48154388e Abbreviations can now be specified simultaneously
berghofe
parents: 21658
diff changeset
  1189
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1190
    val abbrevs = map_filter get_abbrev spec;
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1191
    val bs = map (Binding.name_of o fst o fst) abbrevs;
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1192
21766
3eb48154388e Abbreviations can now be specified simultaneously
berghofe
parents: 21658
diff changeset
  1193
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1194
    (* predicates *)
21766
3eb48154388e Abbreviations can now be specified simultaneously
berghofe
parents: 21658
diff changeset
  1195
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1196
    val pre_intros = filter_out (is_some o get_abbrev) spec;
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1197
    val cnames_syn' = filter_out (member (op =) bs o Binding.name_of o fst o fst) cnames_syn;
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1198
    val cs = map (Free o apfst Binding.name_of o fst) cnames_syn';
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1199
    val ps = map Free pnames;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1200
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30218
diff changeset
  1201
    val (_, ctxt2) = lthy |> Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn');
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 35364
diff changeset
  1202
    val ctxt3 = ctxt2 |> fold (snd oo Local_Defs.fixed_abbrev) abbrevs;
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
  1203
    val expand = Assumption.export_term ctxt3 lthy #> Proof_Context.cert_term lthy;
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1204
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
  1205
    fun close_rule r =
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
  1206
      fold (Logic.all o Free) (fold_aterms
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
  1207
        (fn t as Free (v as (s, _)) =>
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
  1208
            if Variable.is_fixed ctxt1 s orelse
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
  1209
              member (op =) ps t then I else insert (op =) v
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45740
diff changeset
  1210
          | _ => I) r []) r;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1211
26736
e6091328718f added explicit check phase after reading of specification
haftmann
parents: 26534
diff changeset
  1212
    val intros = map (apsnd (Syntax.check_term lthy #> close_rule #> expand)) pre_intros;
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1213
    val preds = map (fn ((c, _), mx) => (c, mx)) cnames_syn';
21048
e57e91f72831 Restructured and repaired code dealing with case names
berghofe
parents: 21024
diff changeset
  1214
  in
25029
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1215
    lthy
3a72718c5ddd gen_add_inductive_i: treat abbrevs as local defs, expand by export;
wenzelm
parents: 25016
diff changeset
  1216
    |> mk_def flags cs intros monos ps preds
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
  1217
    ||> fold (snd oo Local_Theory.abbrev Syntax.mode_default) abbrevs
21048
e57e91f72831 Restructured and repaired code dealing with case names
berghofe
parents: 21024
diff changeset
  1218
  end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1219
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1220
fun gen_add_inductive_cmd mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos lthy =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1221
  let
30486
9cdc7ce0e389 simplified preparation and outer parsing of specification;
wenzelm
parents: 30435
diff changeset
  1222
    val ((vars, intrs), _) = lthy
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
  1223
      |> Proof_Context.set_mode Proof_Context.mode_abbrev
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
  1224
      |> Specification.read_multi_specs (cnames_syn @ pnames_syn) intro_srcs;
24721
2a029b78db60 proper Specification.read_specification;
wenzelm
parents: 24516
diff changeset
  1225
    val (cs, ps) = chop (length cnames_syn) vars;
2a029b78db60 proper Specification.read_specification;
wenzelm
parents: 24516
diff changeset
  1226
    val monos = Attrib.eval_thms lthy raw_monos;
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
  1227
    val flags =
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
  1228
     {quiet_mode = false, verbose = verbose, alt_name = Binding.empty,
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 47876
diff changeset
  1229
      coind = coind, no_elim = false, no_ind = false, skip_mono = false};
26128
fe2d24c26e0c inductive package: simplified group handling;
wenzelm
parents: 25978
diff changeset
  1230
  in
fe2d24c26e0c inductive package: simplified group handling;
wenzelm
parents: 25978
diff changeset
  1231
    lthy
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1232
    |> gen_add_inductive mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos
26128
fe2d24c26e0c inductive package: simplified group handling;
wenzelm
parents: 25978
diff changeset
  1233
  end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1234
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1235
val add_inductive = gen_add_inductive add_ind_def;
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1236
val add_inductive_cmd = gen_add_inductive_cmd add_ind_def;
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1237
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
  1238
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1239
(* read off arities of inductive predicates from raw induction rule *)
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1240
fun arities_of induct =
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1241
  map (fn (_ $ t $ u) =>
80636
4041e7c8059d tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents: 79732
diff changeset
  1242
      (dest_Const_name (head_of t), length (snd (strip_comb u))))
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
  1243
    (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct)));
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1244
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1245
(* read off parameters of inductive predicate from raw induction rule *)
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1246
fun params_of induct =
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1247
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
  1248
    val (_ $ t $ u :: _) = HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct));
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1249
    val (_, ts) = strip_comb t;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1250
    val (_, us) = strip_comb u;
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1251
  in
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1252
    List.take (ts, length ts - length us)
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1253
  end;
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1254
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1255
val pname_of_intr =
80636
4041e7c8059d tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents: 79732
diff changeset
  1256
  Thm.concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const_name;
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1257
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1258
(* partition introduction rules according to predicate name *)
25822
05756950011c Added function partition_rules'.
berghofe
parents: 25510
diff changeset
  1259
fun gen_partition_rules f induct intros =
05756950011c Added function partition_rules'.
berghofe
parents: 25510
diff changeset
  1260
  fold_rev (fn r => AList.map_entry op = (pname_of_intr (f r)) (cons r)) intros
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1261
    (map (rpair [] o fst) (arities_of induct));
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1262
25822
05756950011c Added function partition_rules'.
berghofe
parents: 25510
diff changeset
  1263
val partition_rules = gen_partition_rules I;
05756950011c Added function partition_rules'.
berghofe
parents: 25510
diff changeset
  1264
fun partition_rules' induct = gen_partition_rules fst induct;
05756950011c Added function partition_rules'.
berghofe
parents: 25510
diff changeset
  1265
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1266
fun unpartition_rules intros xs =
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1267
  fold_map (fn r => AList.map_entry_yield op = (pname_of_intr r)
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1268
    (fn x :: xs => (x, xs)) #>> the) intros xs |> fst;
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1269
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1270
(* infer order of variables in intro rules from order of quantifiers in elim rule *)
60362
befdc10ebb42 clarified context;
wenzelm
parents: 60097
diff changeset
  1271
fun infer_intro_vars thy elim arity intros =
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1272
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
  1273
    val _ :: cases = Thm.prems_of elim;
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
  1274
    val used = map (fst o fst) (Term.add_vars (Thm.prop_of elim) []);
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1275
    fun mtch (t, u) =
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1276
      let
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1277
        val params = Logic.strip_params t;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1278
        val vars =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1279
          map (Var o apfst (rpair 0))
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1280
            (Name.variant_list used (map fst params) ~~ map snd params);
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1281
        val ts =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1282
          map (curry subst_bounds (rev vars))
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1283
            (List.drop (Logic.strip_assums_hyp t, arity));
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1284
        val us = Logic.strip_imp_prems u;
45647
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1285
        val tab =
96af0578571c misc tuning;
wenzelm
parents: 45592
diff changeset
  1286
          fold (Pattern.first_order_match thy) (ts ~~ us) (Vartab.empty, Vartab.empty);
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1287
      in
32035
8e77b6a250d5 tuned/modernized Envir.subst_XXX;
wenzelm
parents: 31986
diff changeset
  1288
        map (Envir.subst_term tab) vars
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1289
      end
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1290
  in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
  1291
    map (mtch o apsnd Thm.prop_of) (cases ~~ intros)
22789
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1292
  end;
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1293
4d03dc1cad04 Added functions arities_of, params_of, partition_rules, and
berghofe
parents: 22675
diff changeset
  1294
25978
8ba1eba8d058 added theorem group;
wenzelm
parents: 25822
diff changeset
  1295
58815
fd3f893a40ea modernized setup;
wenzelm
parents: 58028
diff changeset
  1296
(** outer syntax **)
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
  1297
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1298
fun gen_ind_decl mk_def coind =
63285
e9c777bfd78c clarified syntax;
wenzelm
parents: 63180
diff changeset
  1299
  Parse.vars -- Parse.for_fixes --
63064
2f18172214c8 support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents: 63019
diff changeset
  1300
  Scan.optional Parse_Spec.where_multi_specs [] --
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
  1301
  Scan.optional (\<^keyword>\<open>monos\<close> |-- Parse.!!! Parse.thms1) []
26988
742e26213212 more uniform treatment of OuterSyntax.local_theory commands;
wenzelm
parents: 26928
diff changeset
  1302
  >> (fn (((preds, params), specs), monos) =>
69709
7263b59219c1 slightly more conventional naming schema
haftmann
parents: 69593
diff changeset
  1303
      (snd o gen_add_inductive_cmd mk_def true coind preds params specs monos));
23762
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1304
24eef53a9ad3 Reorganization due to introduction of inductive_set wrapper.
berghofe
parents: 23577
diff changeset
  1305
val ind_decl = gen_ind_decl add_ind_def;
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
  1306
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1307
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
  1308
  Outer_Syntax.local_theory \<^command_keyword>\<open>inductive\<close> "define inductive predicates"
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1309
    (ind_decl false);
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1310
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1311
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
  1312
  Outer_Syntax.local_theory \<^command_keyword>\<open>coinductive\<close> "define coinductive predicates"
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33457
diff changeset
  1313
    (ind_decl true);
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1314
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24861
diff changeset
  1315
val _ =
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
  1316
  Outer_Syntax.local_theory' \<^command_keyword>\<open>inductive_cases\<close>
50214
67fb9a168d10 tuned command descriptions;
wenzelm
parents: 49324
diff changeset
  1317
    "create simplified instances of elimination rules"
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
  1318
    (Parse.and_list1 Parse_Spec.simple_specs >> (#2 ooo inductive_cases_cmd));
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
  1319
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1320
val _ =
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
  1321
  Outer_Syntax.local_theory' \<^command_keyword>\<open>inductive_simps\<close>
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
  1322
    "create simplification rules for inductive predicates"
79732
a53287d9add3 improved output in inductive module;
Fabian Huch <huch@in.tum.de>
parents: 78095
diff changeset
  1323
    (Parse.and_list1 Parse_Spec.simple_specs >> (#2 ooo inductive_simps_cmd));
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37264
diff changeset
  1324
50302
9149a07a6c67 added 'print_inductives' command;
wenzelm
parents: 50301
diff changeset
  1325
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
  1326
  Outer_Syntax.command \<^command_keyword>\<open>print_inductives\<close>
50302
9149a07a6c67 added 'print_inductives' command;
wenzelm
parents: 50301
diff changeset
  1327
    "print (co)inductive definitions and monotonicity rules"
60097
d20ca79d50e4 discontinued pointless warnings: commands are only defined inside a theory context;
wenzelm
parents: 59940
diff changeset
  1328
    (Parse.opt_bang >> (fn b => Toplevel.keep (print_inductives b o Toplevel.context_of)));
50302
9149a07a6c67 added 'print_inductives' command;
wenzelm
parents: 50301
diff changeset
  1329
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
  1330
end;