author | wenzelm |
Fri, 20 Mar 2015 14:48:04 +0100 | |
changeset 59763 | 56d2c357e6b5 |
parent 59621 | 291934bac95e |
child 59843 | b640b5e6b023 |
permissions | -rw-r--r-- |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Tools/induct.ML |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
2 |
Author: Markus Wenzel, TU Muenchen |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
3 |
|
26924 | 4 |
Proof by cases, induction, and coinduction. |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
5 |
*) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
6 |
|
37524 | 7 |
signature INDUCT_ARGS = |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
8 |
sig |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
9 |
val cases_default: thm |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
10 |
val atomize: thm list |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
11 |
val rulify: thm list |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
12 |
val rulify_fallback: thm list |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
13 |
val equal_def: thm |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
14 |
val dest_def: term -> (term * term) option |
58957 | 15 |
val trivial_tac: Proof.context -> int -> tactic |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
16 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
17 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
18 |
signature INDUCT = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
19 |
sig |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
20 |
(*rule declarations*) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
21 |
val vars_of: term -> term list |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
22 |
val dest_rules: Proof.context -> |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
23 |
{type_cases: (string * thm) list, pred_cases: (string * thm) list, |
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
24 |
type_induct: (string * thm) list, pred_induct: (string * thm) list, |
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
25 |
type_coinduct: (string * thm) list, pred_coinduct: (string * thm) list} |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
26 |
val print_rules: Proof.context -> unit |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
27 |
val lookup_casesT: Proof.context -> string -> thm option |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
28 |
val lookup_casesP: Proof.context -> string -> thm option |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
29 |
val lookup_inductT: Proof.context -> string -> thm option |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
30 |
val lookup_inductP: Proof.context -> string -> thm option |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
31 |
val lookup_coinductT: Proof.context -> string -> thm option |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
32 |
val lookup_coinductP: Proof.context -> string -> thm option |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
33 |
val find_casesT: Proof.context -> typ -> thm list |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
34 |
val find_casesP: Proof.context -> term -> thm list |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
35 |
val find_inductT: Proof.context -> typ -> thm list |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
36 |
val find_inductP: Proof.context -> term -> thm list |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
37 |
val find_coinductT: Proof.context -> typ -> thm list |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
38 |
val find_coinductP: Proof.context -> term -> thm list |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
39 |
val cases_type: string -> attribute |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
40 |
val cases_pred: string -> attribute |
27140 | 41 |
val cases_del: attribute |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
42 |
val induct_type: string -> attribute |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
43 |
val induct_pred: string -> attribute |
27140 | 44 |
val induct_del: attribute |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
45 |
val coinduct_type: string -> attribute |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
46 |
val coinduct_pred: string -> attribute |
27140 | 47 |
val coinduct_del: attribute |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51658
diff
changeset
|
48 |
val map_simpset: (Proof.context -> Proof.context) -> Context.generic -> Context.generic |
36602
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
wenzelm
parents:
35625
diff
changeset
|
49 |
val induct_simp_add: attribute |
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
wenzelm
parents:
35625
diff
changeset
|
50 |
val induct_simp_del: attribute |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
51 |
val no_simpN: string |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
52 |
val casesN: string |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
53 |
val inductN: string |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
54 |
val coinductN: string |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
55 |
val typeN: string |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
56 |
val predN: string |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
57 |
val setN: string |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
58 |
(*proof methods*) |
45132 | 59 |
val arbitrary_tac: Proof.context -> int -> (string * typ) list -> int -> tactic |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
60 |
val add_defs: (binding option * (term * bool)) option list -> Proof.context -> |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
61 |
(term option list * thm list) * Proof.context |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
62 |
val atomize_term: theory -> term -> term |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
63 |
val atomize_cterm: Proof.context -> conv |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
64 |
val atomize_tac: Proof.context -> int -> tactic |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
65 |
val inner_atomize_tac: Proof.context -> int -> tactic |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
66 |
val rulified_term: thm -> theory * term |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
67 |
val rulify_tac: Proof.context -> int -> tactic |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
68 |
val simplified_rule: Proof.context -> thm -> thm |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
69 |
val simplify_tac: Proof.context -> int -> tactic |
58957 | 70 |
val trivial_tac: Proof.context -> int -> tactic |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
71 |
val rotate_tac: int -> int -> int -> tactic |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
72 |
val internalize: Proof.context -> int -> thm -> thm |
26940 | 73 |
val guess_instance: Proof.context -> thm -> int -> thm -> thm Seq.seq |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
74 |
val cases_tac: Proof.context -> bool -> term option list list -> thm option -> |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
75 |
thm list -> int -> cases_tactic |
27323
385c0ce33173
tuned get_inductT: *all* rules for missing instantiation;
wenzelm
parents:
27140
diff
changeset
|
76 |
val get_inductT: Proof.context -> term option list list -> thm list list |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
77 |
type case_data = (((string * string list) * string list) list * int) (* FIXME -> rule_cases.ML *) |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
78 |
val gen_induct_tac: (theory -> case_data * thm -> case_data * thm) -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
79 |
Proof.context -> bool -> (binding option * (term * bool)) option list list -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
80 |
(string * typ) list list -> term option list -> thm list option -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
81 |
thm list -> int -> cases_tactic |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
82 |
val induct_tac: Proof.context -> bool -> (binding option * (term * bool)) option list list -> |
26924 | 83 |
(string * typ) list list -> term option list -> thm list option -> |
84 |
thm list -> int -> cases_tactic |
|
85 |
val coinduct_tac: Proof.context -> term option list -> term option list -> thm option -> |
|
86 |
thm list -> int -> cases_tactic |
|
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
87 |
val gen_induct_setup: binding -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
88 |
(Proof.context -> bool -> (binding option * (term * bool)) option list list -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
89 |
(string * typ) list list -> term option list -> thm list option -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
90 |
thm list -> int -> cases_tactic) -> |
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
91 |
theory -> theory |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
92 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
93 |
|
37524 | 94 |
functor Induct(Induct_Args: INDUCT_ARGS): INDUCT = |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
95 |
struct |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
96 |
|
37523 | 97 |
(** variables -- ordered left-to-right, preferring right **) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
98 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
99 |
fun vars_of tm = |
45131 | 100 |
rev (distinct (op =) (Term.fold_aterms (fn t as Var _ => cons t | _ => I) tm [])); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
101 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
102 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
103 |
|
37523 | 104 |
val mk_var = Net.encode_type o #2 o Term.dest_Var; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
105 |
|
47060
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
wenzelm
parents:
46961
diff
changeset
|
106 |
fun concl_var which thm = mk_var (which (vars_of (Thm.concl_of thm))) handle List.Empty => |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
107 |
raise THM ("No variables in conclusion of rule", 0, [thm]); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
108 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
109 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
110 |
|
47060
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
wenzelm
parents:
46961
diff
changeset
|
111 |
fun left_var_prem thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle List.Empty => |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
112 |
raise THM ("No variables in major premise of rule", 0, [thm]); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
113 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
114 |
val left_var_concl = concl_var hd; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
115 |
val right_var_concl = concl_var List.last; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
116 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
117 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
118 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
119 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
120 |
|
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
121 |
(** constraint simplification **) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
122 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
123 |
(* rearrange parameters and premises to allow application of one-point-rules *) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
124 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
125 |
fun swap_params_conv ctxt i j cv = |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
126 |
let |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
127 |
fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
128 |
| conv1 k ctxt = |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
129 |
Conv.rewr_conv @{thm swap_params} then_conv |
37525 | 130 |
Conv.forall_conv (conv1 (k - 1) o snd) ctxt |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
131 |
fun conv2 0 ctxt = conv1 j ctxt |
37525 | 132 |
| conv2 k ctxt = Conv.forall_conv (conv2 (k - 1) o snd) ctxt |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
133 |
in conv2 i ctxt end; |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
134 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
135 |
fun swap_prems_conv 0 = Conv.all_conv |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
136 |
| swap_prems_conv i = |
37525 | 137 |
Conv.implies_concl_conv (swap_prems_conv (i - 1)) then_conv |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
138 |
Conv.rewr_conv Drule.swap_prems_eq |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
139 |
|
42361 | 140 |
fun drop_judgment ctxt = Object_Logic.drop_judgment (Proof_Context.theory_of ctxt); |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
141 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
142 |
fun find_eq ctxt t = |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
143 |
let |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
144 |
val l = length (Logic.strip_params t); |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
145 |
val Hs = Logic.strip_assums_hyp t; |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
146 |
fun find (i, t) = |
37525 | 147 |
(case Induct_Args.dest_def (drop_judgment ctxt t) of |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
148 |
SOME (Bound j, _) => SOME (i, j) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
149 |
| SOME (_, Bound j) => SOME (i, j) |
37525 | 150 |
| _ => NONE); |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
151 |
in |
37525 | 152 |
(case get_first find (map_index I Hs) of |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
153 |
NONE => NONE |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
154 |
| SOME (0, 0) => NONE |
37525 | 155 |
| SOME (i, j) => SOME (i, l - j - 1, j)) |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
156 |
end; |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
157 |
|
37525 | 158 |
fun mk_swap_rrule ctxt ct = |
59582 | 159 |
(case find_eq ctxt (Thm.term_of ct) of |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
160 |
NONE => NONE |
37525 | 161 |
| SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct)); |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
162 |
|
37525 | 163 |
val rearrange_eqs_simproc = |
56245 | 164 |
Simplifier.simproc_global Pure.thy "rearrange_eqs" ["Pure.all t"] |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59586
diff
changeset
|
165 |
(fn ctxt => fn t => mk_swap_rrule ctxt (Thm.global_cterm_of (Proof_Context.theory_of ctxt) t)); |
37525 | 166 |
|
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
167 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
168 |
(* rotate k premises to the left by j, skipping over first j premises *) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
169 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
170 |
fun rotate_conv 0 j 0 = Conv.all_conv |
37525 | 171 |
| rotate_conv 0 j k = swap_prems_conv j then_conv rotate_conv 1 j (k - 1) |
172 |
| rotate_conv i j k = Conv.implies_concl_conv (rotate_conv (i - 1) j k); |
|
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
173 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
174 |
fun rotate_tac j 0 = K all_tac |
37525 | 175 |
| rotate_tac j k = SUBGOAL (fn (goal, i) => |
176 |
CONVERSION (rotate_conv |
|
177 |
j (length (Logic.strip_assums_hyp goal) - j - k) k) i); |
|
178 |
||
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
179 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
180 |
(* rulify operators around definition *) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
181 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
182 |
fun rulify_defs_conv ctxt ct = |
59582 | 183 |
if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) andalso |
184 |
not (is_some (Induct_Args.dest_def (drop_judgment ctxt (Thm.term_of ct)))) |
|
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
185 |
then |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
186 |
(Conv.forall_conv (rulify_defs_conv o snd) ctxt else_conv |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
187 |
Conv.implies_conv (Conv.try_conv (rulify_defs_conv ctxt)) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
188 |
(Conv.try_conv (rulify_defs_conv ctxt)) else_conv |
37524 | 189 |
Conv.first_conv (map Conv.rewr_conv Induct_Args.rulify) then_conv |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
190 |
Conv.try_conv (rulify_defs_conv ctxt)) ct |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
191 |
else Conv.no_conv ct; |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
192 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
193 |
|
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
194 |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
195 |
(** induct data **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
196 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
197 |
(* rules *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
198 |
|
30560 | 199 |
type rules = (string * thm) Item_Net.T; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
200 |
|
33373 | 201 |
fun init_rules index : rules = |
202 |
Item_Net.init |
|
203 |
(fn ((s1, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm_prop (th1, th2)) |
|
204 |
(single o index); |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
205 |
|
27140 | 206 |
fun filter_rules (rs: rules) th = |
30560 | 207 |
filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs); |
27140 | 208 |
|
30560 | 209 |
fun lookup_rule (rs: rules) = AList.lookup (op =) (Item_Net.content rs); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
210 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
211 |
fun pretty_rules ctxt kind rs = |
30560 | 212 |
let val thms = map snd (Item_Net.content rs) |
51584 | 213 |
in Pretty.big_list kind (map (Display.pretty_thm_item ctxt) thms) end; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
214 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
215 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
216 |
(* context data *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
217 |
|
37524 | 218 |
structure Data = Generic_Data |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
219 |
( |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
220 |
type T = (rules * rules) * (rules * rules) * (rules * rules) * simpset; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
221 |
val empty = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
222 |
((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)), |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
223 |
(init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)), |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
224 |
(init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2)), |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51658
diff
changeset
|
225 |
simpset_of (empty_simpset @{context} |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51658
diff
changeset
|
226 |
addsimprocs [rearrange_eqs_simproc] addsimps [Drule.norm_hhf_eq])); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
227 |
val extend = I; |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
228 |
fun merge (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1), simpset1), |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
229 |
((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2), simpset2)) = |
30560 | 230 |
((Item_Net.merge (casesT1, casesT2), Item_Net.merge (casesP1, casesP2)), |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
231 |
(Item_Net.merge (inductT1, inductT2), Item_Net.merge (inductP1, inductP2)), |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
232 |
(Item_Net.merge (coinductT1, coinductT2), Item_Net.merge (coinductP1, coinductP2)), |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
233 |
merge_ss (simpset1, simpset2)); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
234 |
); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
235 |
|
37524 | 236 |
val get_local = Data.get o Context.Proof; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
237 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
238 |
fun dest_rules ctxt = |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
239 |
let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in |
30560 | 240 |
{type_cases = Item_Net.content casesT, |
241 |
pred_cases = Item_Net.content casesP, |
|
242 |
type_induct = Item_Net.content inductT, |
|
243 |
pred_induct = Item_Net.content inductP, |
|
244 |
type_coinduct = Item_Net.content coinductT, |
|
245 |
pred_coinduct = Item_Net.content coinductP} |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
246 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
247 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
248 |
fun print_rules ctxt = |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
249 |
let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
250 |
[pretty_rules ctxt "coinduct type:" coinductT, |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
251 |
pretty_rules ctxt "coinduct pred:" coinductP, |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
252 |
pretty_rules ctxt "induct type:" inductT, |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
253 |
pretty_rules ctxt "induct pred:" inductP, |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
254 |
pretty_rules ctxt "cases type:" casesT, |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
255 |
pretty_rules ctxt "cases pred:" casesP] |
56334
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
56245
diff
changeset
|
256 |
|> Pretty.writeln_chunks |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
257 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
258 |
|
24867 | 259 |
val _ = |
58893
9e0ecb66d6a7
eliminated unused int_only flag (see also c12484a27367);
wenzelm
parents:
58838
diff
changeset
|
260 |
Outer_Syntax.command @{command_spec "print_induct_rules"} |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
45375
diff
changeset
|
261 |
"print induction and cases rules" |
51658
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
wenzelm
parents:
51584
diff
changeset
|
262 |
(Scan.succeed (Toplevel.unknown_context o |
24867 | 263 |
Toplevel.keep (print_rules o Toplevel.context_of))); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
264 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
265 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
266 |
(* access rules *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
267 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
268 |
val lookup_casesT = lookup_rule o #1 o #1 o get_local; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
269 |
val lookup_casesP = lookup_rule o #2 o #1 o get_local; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
270 |
val lookup_inductT = lookup_rule o #1 o #2 o get_local; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
271 |
val lookup_inductP = lookup_rule o #2 o #2 o get_local; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
272 |
val lookup_coinductT = lookup_rule o #1 o #3 o get_local; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
273 |
val lookup_coinductP = lookup_rule o #2 o #3 o get_local; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
274 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
275 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
276 |
fun find_rules which how ctxt x = |
30560 | 277 |
map snd (Item_Net.retrieve (which (get_local ctxt)) (how x)); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
278 |
|
37523 | 279 |
val find_casesT = find_rules (#1 o #1) Net.encode_type; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
280 |
val find_casesP = find_rules (#2 o #1) I; |
37523 | 281 |
val find_inductT = find_rules (#1 o #2) Net.encode_type; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
282 |
val find_inductP = find_rules (#2 o #2) I; |
37523 | 283 |
val find_coinductT = find_rules (#1 o #3) Net.encode_type; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
284 |
val find_coinductP = find_rules (#2 o #3) I; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
285 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
286 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
287 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
288 |
(** attributes **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
289 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
290 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
291 |
|
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
292 |
fun mk_att f g name = |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
293 |
Thm.mixed_attribute (fn (context, thm) => |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
294 |
let |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
295 |
val thm' = g thm; |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
296 |
val context' = Data.map (f (name, thm')) context; |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
297 |
in (context', thm') end); |
27140 | 298 |
|
37525 | 299 |
fun del_att which = |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
59057
diff
changeset
|
300 |
Thm.declaration_attribute (fn th => Data.map (which (apply2 (fn rs => |
37525 | 301 |
fold Item_Net.remove (filter_rules rs th) rs)))); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
302 |
|
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
303 |
fun add_casesT rule x = @{apply 4(1)} (apfst (Item_Net.update rule)) x; |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
304 |
fun add_casesP rule x = @{apply 4(1)} (apsnd (Item_Net.update rule)) x; |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
305 |
fun add_inductT rule x = @{apply 4(2)} (apfst (Item_Net.update rule)) x; |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
306 |
fun add_inductP rule x = @{apply 4(2)} (apsnd (Item_Net.update rule)) x; |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
307 |
fun add_coinductT rule x = @{apply 4(3)} (apfst (Item_Net.update rule)) x; |
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
308 |
fun add_coinductP rule x = @{apply 4(3)} (apsnd (Item_Net.update rule)) x; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
309 |
|
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
310 |
val consumes0 = Rule_Cases.default_consumes 0; |
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
45328
diff
changeset
|
311 |
val consumes1 = Rule_Cases.default_consumes 1; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
312 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
313 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
314 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
315 |
val cases_type = mk_att add_casesT consumes0; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
316 |
val cases_pred = mk_att add_casesP consumes1; |
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
317 |
val cases_del = del_att @{apply 4(1)}; |
27140 | 318 |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
319 |
val induct_type = mk_att add_inductT consumes0; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
320 |
val induct_pred = mk_att add_inductP consumes1; |
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
321 |
val induct_del = del_att @{apply 4(2)}; |
27140 | 322 |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
323 |
val coinduct_type = mk_att add_coinductT consumes0; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
324 |
val coinduct_pred = mk_att add_coinductP consumes1; |
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
325 |
val coinduct_del = del_att @{apply 4(3)}; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
326 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51658
diff
changeset
|
327 |
fun map_simpset f context = |
59057
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
wenzelm
parents:
58957
diff
changeset
|
328 |
context |> (Data.map o @{apply 4(4)} o Simplifier.simpset_map (Context.proof_of context)) f; |
36602
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
wenzelm
parents:
35625
diff
changeset
|
329 |
|
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
wenzelm
parents:
35625
diff
changeset
|
330 |
fun induct_simp f = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51658
diff
changeset
|
331 |
Thm.declaration_attribute (fn thm => map_simpset (fn ctxt => f (ctxt, [thm]))); |
36602
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
wenzelm
parents:
35625
diff
changeset
|
332 |
|
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
wenzelm
parents:
35625
diff
changeset
|
333 |
val induct_simp_add = induct_simp (op addsimps); |
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
wenzelm
parents:
35625
diff
changeset
|
334 |
val induct_simp_del = induct_simp (op delsimps); |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
335 |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
336 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
337 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
338 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
339 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
340 |
(** attribute syntax **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
341 |
|
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
342 |
val no_simpN = "no_simp"; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
343 |
val casesN = "cases"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
344 |
val inductN = "induct"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
345 |
val coinductN = "coinduct"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
346 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
347 |
val typeN = "type"; |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
348 |
val predN = "pred"; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
349 |
val setN = "set"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
350 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
351 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
352 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
353 |
fun spec k arg = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
354 |
Scan.lift (Args.$$$ k -- Args.colon) |-- arg || |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
355 |
Scan.lift (Args.$$$ k) >> K ""; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
356 |
|
30528 | 357 |
fun attrib add_type add_pred del = |
55951
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents:
54742
diff
changeset
|
358 |
spec typeN (Args.type_name {proper = false, strict = false}) >> add_type || |
55954 | 359 |
spec predN (Args.const {proper = false, strict = false}) >> add_pred || |
360 |
spec setN (Args.const {proper = false, strict = false}) >> add_pred || |
|
30528 | 361 |
Scan.lift Args.del >> K del; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
362 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
363 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
364 |
|
58826 | 365 |
val _ = |
366 |
Theory.setup |
|
367 |
(Attrib.setup @{binding cases} (attrib cases_type cases_pred cases_del) |
|
368 |
"declaration of cases rule" #> |
|
369 |
Attrib.setup @{binding induct} (attrib induct_type induct_pred induct_del) |
|
370 |
"declaration of induction rule" #> |
|
371 |
Attrib.setup @{binding coinduct} (attrib coinduct_type coinduct_pred coinduct_del) |
|
372 |
"declaration of coinduction rule" #> |
|
373 |
Attrib.setup @{binding induct_simp} (Attrib.add_del induct_simp_add induct_simp_del) |
|
374 |
"declaration of rules for simplifying induction or cases rules"); |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
375 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
376 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
377 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
378 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
379 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
380 |
(** method utils **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
381 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
382 |
(* alignment *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
383 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
384 |
fun align_left msg xs ys = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
385 |
let val m = length xs and n = length ys |
33957 | 386 |
in if m < n then error msg else (take n xs ~~ ys) end; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
387 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
388 |
fun align_right msg xs ys = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
389 |
let val m = length xs and n = length ys |
33957 | 390 |
in if m < n then error msg else (drop (m - n) xs ~~ ys) end; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
391 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
392 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
393 |
(* prep_inst *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
394 |
|
32432
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
395 |
fun prep_inst ctxt align tune (tm, ts) = |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
396 |
let |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59586
diff
changeset
|
397 |
val cert = Thm.global_cterm_of (Proof_Context.theory_of ctxt); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
398 |
fun prep_var (x, SOME t) = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
399 |
let |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
400 |
val cx = cert x; |
59586 | 401 |
val xT = Thm.typ_of_cterm cx; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
402 |
val ct = cert (tune t); |
59586 | 403 |
val tT = Thm.typ_of_cterm ct; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
404 |
in |
32432
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
405 |
if Type.could_unify (tT, xT) then SOME (cx, ct) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
406 |
else error (Pretty.string_of (Pretty.block |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
407 |
[Pretty.str "Ill-typed instantiation:", Pretty.fbrk, |
32432
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
408 |
Syntax.pretty_term ctxt (Thm.term_of ct), Pretty.str " ::", Pretty.brk 1, |
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
409 |
Syntax.pretty_typ ctxt tT])) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
410 |
end |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
411 |
| prep_var (_, NONE) = NONE; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
412 |
val xs = vars_of tm; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
413 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
414 |
align "Rule has fewer variables than instantiations given" xs ts |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
415 |
|> map_filter prep_var |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
416 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
417 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
418 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
419 |
(* trace_rules *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
420 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
421 |
fun trace_rules _ kind [] = error ("Unable to figure out " ^ kind ^ " rule") |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
422 |
| trace_rules ctxt _ rules = Method.trace ctxt rules; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
423 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
424 |
|
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
425 |
(* mark equality constraints in cases rule *) |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
426 |
|
37524 | 427 |
val equal_def' = Thm.symmetric Induct_Args.equal_def; |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
428 |
|
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
429 |
fun mark_constraints n ctxt = Conv.fconv_rule |
45130 | 430 |
(Conv.prems_conv ~1 (Conv.params_conv ~1 (K (Conv.prems_conv n |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
431 |
(Raw_Simplifier.rewrite ctxt false [equal_def']))) ctxt)); |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
432 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
433 |
fun unmark_constraints ctxt = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
434 |
Conv.fconv_rule (Raw_Simplifier.rewrite ctxt true [Induct_Args.equal_def]); |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
435 |
|
37525 | 436 |
|
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
437 |
(* simplify *) |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
438 |
|
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
439 |
fun simplify_conv' ctxt = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51658
diff
changeset
|
440 |
Simplifier.full_rewrite (put_simpset (#4 (get_local ctxt)) ctxt); |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
441 |
|
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
442 |
fun simplify_conv ctxt ct = |
59582 | 443 |
if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) then |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
444 |
(Conv.try_conv (rulify_defs_conv ctxt) then_conv simplify_conv' ctxt) ct |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
445 |
else Conv.all_conv ct; |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
446 |
|
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
447 |
fun gen_simplified_rule cv ctxt = |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
448 |
Conv.fconv_rule (Conv.prems_conv ~1 (cv ctxt)); |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
449 |
|
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
450 |
val simplified_rule' = gen_simplified_rule simplify_conv'; |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
451 |
val simplified_rule = gen_simplified_rule simplify_conv; |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
452 |
|
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
453 |
fun simplify_tac ctxt = CONVERSION (simplify_conv ctxt); |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
454 |
|
37524 | 455 |
val trivial_tac = Induct_Args.trivial_tac; |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
456 |
|
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
457 |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
458 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
459 |
(** cases method **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
460 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
461 |
(* |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
462 |
rule selection scheme: |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
463 |
cases - default case split |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
464 |
`A t` cases ... - predicate/set cases |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
465 |
cases t - type cases |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
466 |
... cases ... r - explicit rule |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
467 |
*) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
468 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
469 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
470 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
471 |
fun get_casesT ctxt ((SOME t :: _) :: _) = find_casesT ctxt (Term.fastype_of t) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
472 |
| get_casesT _ _ = []; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
473 |
|
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
474 |
fun get_casesP ctxt (fact :: _) = find_casesP ctxt (Thm.concl_of fact) |
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
475 |
| get_casesP _ _ = []; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
476 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
477 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
478 |
|
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
479 |
fun cases_tac ctxt simp insts opt_rule facts = |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
480 |
let |
42361 | 481 |
val thy = Proof_Context.theory_of ctxt; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
482 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
483 |
fun inst_rule r = |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
484 |
(if null insts then r |
45131 | 485 |
else |
486 |
(align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts |
|
487 |
|> maps (prep_inst ctxt align_left I) |
|
488 |
|> Drule.cterm_instantiate) r) |
|
489 |
|> simp ? mark_constraints (Rule_Cases.get_constraints r) ctxt |
|
490 |
|> pair (Rule_Cases.get r); |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
491 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
492 |
val ruleq = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
493 |
(case opt_rule of |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
494 |
SOME r => Seq.single (inst_rule r) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
495 |
| NONE => |
37524 | 496 |
(get_casesP ctxt facts @ get_casesT ctxt insts @ [Induct_Args.cases_default]) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
497 |
|> tap (trace_rules ctxt casesN) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
498 |
|> Seq.of_list |> Seq.maps (Seq.try inst_rule)); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
499 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
500 |
fn i => fn st => |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
501 |
ruleq |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
502 |
|> Seq.maps (Rule_Cases.consume ctxt [] facts) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
503 |
|> Seq.maps (fn ((cases, (_, more_facts)), rule) => |
45131 | 504 |
let |
505 |
val rule' = rule |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
506 |
|> simp ? (simplified_rule' ctxt #> unmark_constraints ctxt); |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
507 |
in |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
508 |
CASES (Rule_Cases.make_common (thy, |
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
509 |
Thm.prop_of (Rule_Cases.internalize_params rule')) cases) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59058
diff
changeset
|
510 |
((Method.insert_tac more_facts THEN' resolve_tac ctxt [rule'] THEN_ALL_NEW |
58957 | 511 |
(if simp then TRY o trivial_tac ctxt else K all_tac)) i) st |
34987
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
berghofe
parents:
34907
diff
changeset
|
512 |
end) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
513 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
514 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
515 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
516 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
517 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
518 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
519 |
(** induct method **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
520 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
521 |
val conjunction_congs = [@{thm Pure.all_conjunction}, @{thm imp_conjunction}]; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
522 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
523 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
524 |
(* atomize *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
525 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
526 |
fun atomize_term thy = |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
38715
diff
changeset
|
527 |
Raw_Simplifier.rewrite_term thy Induct_Args.atomize [] |
35625 | 528 |
#> Object_Logic.drop_judgment thy; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
529 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
530 |
fun atomize_cterm ctxt = Raw_Simplifier.rewrite ctxt true Induct_Args.atomize; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
531 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
532 |
fun atomize_tac ctxt = rewrite_goal_tac ctxt Induct_Args.atomize; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
533 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
534 |
fun inner_atomize_tac ctxt = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
535 |
rewrite_goal_tac ctxt (map Thm.symmetric conjunction_congs) THEN' atomize_tac ctxt; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
536 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
537 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
538 |
(* rulify *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
539 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
540 |
fun rulify_term thy = |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
38715
diff
changeset
|
541 |
Raw_Simplifier.rewrite_term thy (Induct_Args.rulify @ conjunction_congs) [] #> |
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
38715
diff
changeset
|
542 |
Raw_Simplifier.rewrite_term thy Induct_Args.rulify_fallback []; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
543 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
544 |
fun rulified_term thm = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
545 |
let |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
546 |
val thy = Thm.theory_of_thm thm; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
547 |
val rulify = rulify_term thy; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
548 |
val (As, B) = Logic.strip_horn (Thm.prop_of thm); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
549 |
in (thy, Logic.list_implies (map rulify As, rulify B)) end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
550 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
551 |
fun rulify_tac ctxt = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
552 |
rewrite_goal_tac ctxt (Induct_Args.rulify @ conjunction_congs) THEN' |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
553 |
rewrite_goal_tac ctxt Induct_Args.rulify_fallback THEN' |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
554 |
Goal.conjunction_tac THEN_ALL_NEW |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
555 |
(rewrite_goal_tac ctxt [@{thm Pure.conjunction_imp}] THEN' Goal.norm_hhf_tac ctxt); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
556 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
557 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
558 |
(* prepare rule *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
559 |
|
32432
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
560 |
fun rule_instance ctxt inst rule = |
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
561 |
Drule.cterm_instantiate (prep_inst ctxt align_left I (Thm.prop_of rule, inst)) rule; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
562 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
563 |
fun internalize ctxt k th = |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
564 |
th |> Thm.permute_prems 0 k |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
565 |
|> Conv.fconv_rule (Conv.concl_conv (Thm.nprems_of th - k) (atomize_cterm ctxt)); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
566 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
567 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
568 |
(* guess rule instantiation -- cannot handle pending goal parameters *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
569 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
570 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
571 |
|
32032 | 572 |
fun dest_env thy env = |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
573 |
let |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59586
diff
changeset
|
574 |
val cert = Thm.global_cterm_of thy; |
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59586
diff
changeset
|
575 |
val certT = Thm.global_ctyp_of thy; |
32032 | 576 |
val pairs = Vartab.dest (Envir.term_env env); |
577 |
val types = Vartab.dest (Envir.type_env env); |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
578 |
val ts = map (cert o Envir.norm_term env o #2 o #2) pairs; |
59586 | 579 |
val xs = map2 (curry (cert o Var)) (map #1 pairs) (map Thm.typ_of_cterm ts); |
32032 | 580 |
in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) types, xs ~~ ts) end; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
581 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
582 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
583 |
|
26940 | 584 |
fun guess_instance ctxt rule i st = |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
585 |
let |
42361 | 586 |
val thy = Proof_Context.theory_of ctxt; |
26626
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents:
26568
diff
changeset
|
587 |
val maxidx = Thm.maxidx_of st; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
588 |
val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*) |
29276 | 589 |
val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal)); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
590 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
591 |
if not (null params) then |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
592 |
(warning ("Cannot determine rule instantiation due to pending parameter(s): " ^ |
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
47060
diff
changeset
|
593 |
commas_quote (map (Syntax.string_of_term ctxt o Syntax_Trans.mark_bound_abs) params)); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
594 |
Seq.single rule) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
595 |
else |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
596 |
let |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
597 |
val rule' = Thm.incr_indexes (maxidx + 1) rule; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
598 |
val concl = Logic.strip_assums_concl goal; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
599 |
in |
58950
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents:
58893
diff
changeset
|
600 |
Unify.smash_unifiers (Context.Proof ctxt) |
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents:
58893
diff
changeset
|
601 |
[(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule')) |
43333
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents:
43326
diff
changeset
|
602 |
|> Seq.map (fn env => Drule.instantiate_normalize (dest_env thy env) rule') |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
603 |
end |
43333
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents:
43326
diff
changeset
|
604 |
end |
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents:
43326
diff
changeset
|
605 |
handle General.Subscript => Seq.empty; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
606 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
607 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
608 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
609 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
610 |
(* special renaming of rule parameters *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
611 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
612 |
fun special_rename_params ctxt [[SOME (Free (z, Type (T, _)))]] [thm] = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
613 |
let |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42361
diff
changeset
|
614 |
val x = Name.clean (Variable.revert_fixed ctxt z); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
615 |
fun index i [] = [] |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
616 |
| index i (y :: ys) = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
617 |
if x = y then x ^ string_of_int i :: index (i + 1) ys |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
618 |
else y :: index i ys; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
619 |
fun rename_params [] = [] |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
620 |
| rename_params ((y, Type (U, _)) :: ys) = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
621 |
(if U = T then x else y) :: rename_params ys |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
622 |
| rename_params ((y, _) :: ys) = y :: rename_params ys; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
623 |
fun rename_asm A = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
624 |
let |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
625 |
val xs = rename_params (Logic.strip_params A); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
626 |
val xs' = |
28375 | 627 |
(case filter (fn x' => x' = x) xs of |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
628 |
[] => xs | [_] => xs | _ => index 1 xs); |
45328 | 629 |
in Logic.list_rename_params xs' A end; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
630 |
fun rename_prop p = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
631 |
let val (As, C) = Logic.strip_horn p |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
632 |
in Logic.list_implies (map rename_asm As, C) end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
633 |
val cp' = cterm_fun rename_prop (Thm.cprop_of thm); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
634 |
val thm' = Thm.equal_elim (Thm.reflexive cp') thm; |
33368 | 635 |
in [Rule_Cases.save thm thm'] end |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
636 |
| special_rename_params _ _ ths = ths; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
637 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
638 |
|
45132 | 639 |
(* arbitrary_tac *) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
640 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
641 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
642 |
|
56245 | 643 |
fun goal_prefix k ((c as Const (@{const_name Pure.all}, _)) $ Abs (a, T, B)) = |
644 |
c $ Abs (a, T, goal_prefix k B) |
|
45156 | 645 |
| goal_prefix 0 _ = Term.dummy_prop |
56245 | 646 |
| goal_prefix k ((c as Const (@{const_name Pure.imp}, _)) $ A $ B) = |
647 |
c $ A $ goal_prefix (k - 1) B |
|
45156 | 648 |
| goal_prefix _ _ = Term.dummy_prop; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
649 |
|
56245 | 650 |
fun goal_params k (Const (@{const_name Pure.all}, _) $ Abs (_, _, B)) = goal_params k B + 1 |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
651 |
| goal_params 0 _ = 0 |
56245 | 652 |
| goal_params k (Const (@{const_name Pure.imp}, _) $ _ $ B) = goal_params (k - 1) B |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
653 |
| goal_params _ _ = 0; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
654 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
655 |
fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) => |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
656 |
let |
42361 | 657 |
val thy = Proof_Context.theory_of ctxt; |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59586
diff
changeset
|
658 |
val cert = Thm.global_cterm_of thy; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
659 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
660 |
val v = Free (x, T); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
661 |
fun spec_rule prfx (xs, body) = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
662 |
@{thm Pure.meta_spec} |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42361
diff
changeset
|
663 |
|> Thm.rename_params_rule ([Name.clean (Variable.revert_fixed ctxt x)], 1) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
664 |
|> Thm.lift_rule (cert prfx) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
665 |
|> `(Thm.prop_of #> Logic.strip_assums_concl) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
666 |
|-> (fn pred $ arg => |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
667 |
Drule.cterm_instantiate |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
668 |
[(cert (Term.head_of pred), cert (Logic.rlist_abs (xs, body))), |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
669 |
(cert (Term.head_of arg), cert (Logic.rlist_abs (xs, v)))]); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
670 |
|
56245 | 671 |
fun goal_concl k xs (Const (@{const_name Pure.all}, _) $ Abs (a, T, B)) = |
672 |
goal_concl k ((a, T) :: xs) B |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
673 |
| goal_concl 0 xs B = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
674 |
if not (Term.exists_subterm (fn t => t aconv v) B) then NONE |
44241 | 675 |
else SOME (xs, absfree (x, T) (Term.incr_boundvars 1 B)) |
56245 | 676 |
| goal_concl k xs (Const (@{const_name Pure.imp}, _) $ _ $ B) = |
677 |
goal_concl (k - 1) xs B |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
678 |
| goal_concl _ _ _ = NONE; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
679 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
680 |
(case goal_concl n [] goal of |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
681 |
SOME concl => |
58956
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents:
58950
diff
changeset
|
682 |
(compose_tac ctxt (false, spec_rule (goal_prefix n goal) concl, 1) THEN' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59058
diff
changeset
|
683 |
resolve_tac ctxt [asm_rl]) i |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
684 |
| NONE => all_tac) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
685 |
end); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
686 |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
687 |
fun miniscope_tac p = |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
688 |
CONVERSION o |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
689 |
Conv.params_conv p (fn ctxt => |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
690 |
Raw_Simplifier.rewrite ctxt true [Thm.symmetric Drule.norm_hhf_eq]); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
691 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
692 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
693 |
|
45132 | 694 |
fun arbitrary_tac _ _ [] = K all_tac |
695 |
| arbitrary_tac ctxt n xs = SUBGOAL (fn (goal, i) => |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
696 |
(EVERY' (map (meta_spec_tac ctxt n) xs) THEN' |
24832 | 697 |
(miniscope_tac (goal_params n goal) ctxt)) i); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
698 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
699 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
700 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
701 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
702 |
(* add_defs *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
703 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
704 |
fun add_defs def_insts = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
705 |
let |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
706 |
fun add (SOME (_, (t, true))) ctxt = ((SOME t, []), ctxt) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
707 |
| add (SOME (SOME x, (t, _))) ctxt = |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27865
diff
changeset
|
708 |
let val ([(lhs, (_, th))], ctxt') = |
49748
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
wenzelm
parents:
49660
diff
changeset
|
709 |
Local_Defs.add_defs [((x, NoSyn), ((Thm.def_binding x, []), t))] ctxt |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
710 |
in ((SOME lhs, [th]), ctxt') end |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
711 |
| add (SOME (NONE, (t as Free _, _))) ctxt = ((SOME t, []), ctxt) |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
712 |
| add (SOME (NONE, (t, _))) ctxt = |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
713 |
let |
43326
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents:
43278
diff
changeset
|
714 |
val (s, _) = Name.variant "x" (Variable.names_of ctxt); |
49748
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
wenzelm
parents:
49660
diff
changeset
|
715 |
val x = Binding.name s; |
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
wenzelm
parents:
49660
diff
changeset
|
716 |
val ([(lhs, (_, th))], ctxt') = ctxt |
a346daa8a1f4
clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
wenzelm
parents:
49660
diff
changeset
|
717 |
|> Local_Defs.add_defs [((x, NoSyn), ((Thm.def_binding x, []), t))]; |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
718 |
in ((SOME lhs, [th]), ctxt') end |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
719 |
| add NONE ctxt = ((NONE, []), ctxt); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
720 |
in fold_map add def_insts #> apfst (split_list #> apsnd flat) end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
721 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
722 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
723 |
(* induct_tac *) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
724 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
725 |
(* |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
726 |
rule selection scheme: |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
727 |
`A x` induct ... - predicate/set induction |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
728 |
induct x - type induction |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
729 |
... induct ... r - explicit rule |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
730 |
*) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
731 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
732 |
fun get_inductT ctxt insts = |
32188 | 733 |
fold_rev (map_product cons) (insts |> map |
27323
385c0ce33173
tuned get_inductT: *all* rules for missing instantiation;
wenzelm
parents:
27140
diff
changeset
|
734 |
((fn [] => NONE | ts => List.last ts) #> |
385c0ce33173
tuned get_inductT: *all* rules for missing instantiation;
wenzelm
parents:
27140
diff
changeset
|
735 |
(fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #> |
385c0ce33173
tuned get_inductT: *all* rules for missing instantiation;
wenzelm
parents:
27140
diff
changeset
|
736 |
find_inductT ctxt)) [[]] |
33368 | 737 |
|> filter_out (forall Rule_Cases.is_inner_rule); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
738 |
|
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
739 |
fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact)) |
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
740 |
| get_inductP _ _ = []; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
741 |
|
45130 | 742 |
type case_data = (((string * string list) * string list) list * int); |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
743 |
|
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
744 |
fun gen_induct_tac mod_cases ctxt simp def_insts arbitrary taking opt_rule facts = |
56231
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
745 |
SUBGOAL_CASES (fn (_, i, st) => |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
746 |
let |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
747 |
val thy = Proof_Context.theory_of ctxt; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
748 |
|
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
749 |
val ((insts, defs), defs_ctxt) = fold_map add_defs def_insts ctxt |>> split_list; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
750 |
val atomized_defs = map (map (Conv.fconv_rule (atomize_cterm defs_ctxt))) defs; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
751 |
|
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
752 |
fun inst_rule (concls, r) = |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
753 |
(if null insts then `Rule_Cases.get r |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
754 |
else (align_left "Rule has fewer conclusions than arguments given" |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
755 |
(map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
756 |
|> maps (prep_inst ctxt align_right (atomize_term thy)) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
757 |
|> Drule.cterm_instantiate) r |> pair (Rule_Cases.get r)) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
758 |
|> mod_cases thy |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
759 |
|> (fn ((cases, consumes), th) => (((cases, concls), consumes), th)); |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
760 |
|
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
761 |
val ruleq = |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
762 |
(case opt_rule of |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
763 |
SOME rs => Seq.single (inst_rule (Rule_Cases.strict_mutual_rule ctxt rs)) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
764 |
| NONE => |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
765 |
(get_inductP ctxt facts @ |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
766 |
map (special_rename_params defs_ctxt insts) (get_inductT ctxt insts)) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
767 |
|> map_filter (Rule_Cases.mutual_rule ctxt) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
768 |
|> tap (trace_rules ctxt inductN o map #2) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
769 |
|> Seq.of_list |> Seq.maps (Seq.try inst_rule)); |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
770 |
|
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
771 |
fun rule_cases ctxt rule cases = |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
772 |
let |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
773 |
val rule' = Rule_Cases.internalize_params rule; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
774 |
val rule'' = rule' |> simp ? simplified_rule ctxt; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
775 |
val nonames = map (fn ((cn, _), cls) => ((cn, []), cls)); |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
776 |
val cases' = if Thm.eq_thm_prop (rule', rule'') then cases else nonames cases; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
777 |
in Rule_Cases.make_nested (Thm.prop_of rule'') (rulified_term rule'') cases' end; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
778 |
in |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
779 |
fn st => |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
780 |
ruleq |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
781 |
|> Seq.maps (Rule_Cases.consume defs_ctxt (flat defs) facts) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
782 |
|> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
783 |
(PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
784 |
(CONJUNCTS (ALLGOALS |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
785 |
let |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
786 |
val adefs = nth_list atomized_defs (j - 1); |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
787 |
val frees = fold (Term.add_frees o Thm.prop_of) adefs []; |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
788 |
val xs = nth_list arbitrary (j - 1); |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
789 |
val k = nth concls (j - 1) + more_consumes |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
790 |
in |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
791 |
Method.insert_tac (more_facts @ adefs) THEN' |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
792 |
(if simp then |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
793 |
rotate_tac k (length adefs) THEN' |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
794 |
arbitrary_tac defs_ctxt k (List.partition (member op = frees) xs |> op @) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
795 |
else |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
796 |
arbitrary_tac defs_ctxt k xs) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
797 |
end) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
798 |
THEN' inner_atomize_tac defs_ctxt) j)) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
799 |
THEN' atomize_tac defs_ctxt) i st |> Seq.maps (fn st' => |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
800 |
guess_instance ctxt (internalize ctxt more_consumes rule) i st' |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
801 |
|> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking)) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
802 |
|> Seq.maps (fn rule' => |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
803 |
CASES (rule_cases ctxt rule' cases) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59058
diff
changeset
|
804 |
(resolve_tac ctxt [rule'] i THEN |
56231
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
805 |
PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) st'))) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
806 |
end) |
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
807 |
THEN_ALL_NEW_CASES |
58957 | 808 |
((if simp then simplify_tac ctxt THEN' (TRY o trivial_tac ctxt) else K all_tac) |
56231
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
809 |
THEN_ALL_NEW rulify_tac ctxt); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
810 |
|
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
811 |
val induct_tac = gen_induct_tac (K I); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
812 |
|
45130 | 813 |
|
814 |
||
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
815 |
(** coinduct method **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
816 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
817 |
(* |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
818 |
rule selection scheme: |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
819 |
goal "A x" coinduct ... - predicate/set coinduction |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
820 |
coinduct x - type coinduction |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
821 |
coinduct ... r - explicit rule |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
822 |
*) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
823 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
824 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
825 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
826 |
fun get_coinductT ctxt (SOME t :: _) = find_coinductT ctxt (Term.fastype_of t) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
827 |
| get_coinductT _ _ = []; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
828 |
|
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
829 |
fun get_coinductP ctxt goal = find_coinductP ctxt (Logic.strip_assums_concl goal); |
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
830 |
|
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
831 |
fun main_prop_of th = |
33368 | 832 |
if Rule_Cases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
833 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
834 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
835 |
|
56231
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
836 |
fun coinduct_tac ctxt inst taking opt_rule facts = SUBGOAL_CASES (fn (goal, i, st) => |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
837 |
let |
42361 | 838 |
val thy = Proof_Context.theory_of ctxt; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
839 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
840 |
fun inst_rule r = |
33368 | 841 |
if null inst then `Rule_Cases.get r |
32432
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
842 |
else Drule.cterm_instantiate (prep_inst ctxt align_right I (main_prop_of r, inst)) r |
33368 | 843 |
|> pair (Rule_Cases.get r); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
844 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
845 |
fun ruleq goal = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
846 |
(case opt_rule of |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
847 |
SOME r => Seq.single (inst_rule r) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
848 |
| NONE => |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
849 |
(get_coinductP ctxt goal @ get_coinductT ctxt inst) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
850 |
|> tap (trace_rules ctxt coinductN) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
851 |
|> Seq.of_list |> Seq.maps (Seq.try inst_rule)); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
852 |
in |
56231
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
853 |
fn st => |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
854 |
ruleq goal |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53168
diff
changeset
|
855 |
|> Seq.maps (Rule_Cases.consume ctxt [] facts) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
856 |
|> Seq.maps (fn ((cases, (_, more_facts)), rule) => |
26940 | 857 |
guess_instance ctxt rule i st |
32432
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents:
32188
diff
changeset
|
858 |
|> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking)) |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
859 |
|> Seq.maps (fn rule' => |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
860 |
CASES (Rule_Cases.make_common (thy, |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
861 |
Thm.prop_of (Rule_Cases.internalize_params rule')) cases) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59058
diff
changeset
|
862 |
(Method.insert_tac more_facts i THEN resolve_tac ctxt [rule'] i) st)) |
56231
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents:
55954
diff
changeset
|
863 |
end); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
864 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
865 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
866 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
867 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
868 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
869 |
(** concrete syntax **) |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
870 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
871 |
val arbitraryN = "arbitrary"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
872 |
val takingN = "taking"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
873 |
val ruleN = "rule"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
874 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
875 |
local |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
876 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
877 |
fun single_rule [rule] = rule |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
878 |
| single_rule _ = error "Single rule expected"; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
879 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
880 |
fun named_rule k arg get = |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
881 |
Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|-- |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
882 |
(fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name => |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
883 |
(case get (Context.proof_of context) name of SOME x => x |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
884 |
| NONE => error ("No rule for " ^ k ^ " " ^ quote name)))))); |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
885 |
|
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
886 |
fun rule get_type get_pred = |
55951
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents:
54742
diff
changeset
|
887 |
named_rule typeN (Args.type_name {proper = false, strict = false}) get_type || |
55954 | 888 |
named_rule predN (Args.const {proper = false, strict = false}) get_pred || |
889 |
named_rule setN (Args.const {proper = false, strict = false}) get_pred || |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
890 |
Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
891 |
|
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
892 |
val cases_rule = rule lookup_casesT lookup_casesP >> single_rule; |
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
893 |
val induct_rule = rule lookup_inductT lookup_inductP; |
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
894 |
val coinduct_rule = rule lookup_coinductT lookup_coinductP >> single_rule; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
895 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
896 |
val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
897 |
|
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
898 |
val inst' = Scan.lift (Args.$$$ "_") >> K NONE || |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
899 |
Args.term >> (SOME o rpair false) || |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
900 |
Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --| |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
901 |
Scan.lift (Args.$$$ ")"); |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
902 |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
903 |
val def_inst = |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
27865
diff
changeset
|
904 |
((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME) |
34907
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
905 |
-- (Args.term >> rpair false)) >> SOME || |
b0aaec87751c
Added infrastructure for simplifying equality constraints.
berghofe
parents:
33957
diff
changeset
|
906 |
inst' >> Option.map (pair NONE); |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
907 |
|
27370 | 908 |
val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => |
909 |
error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
910 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
911 |
fun unless_more_args scan = Scan.unless (Scan.lift |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
912 |
((Args.$$$ arbitraryN || Args.$$$ takingN || Args.$$$ typeN || |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
913 |
Args.$$$ predN || Args.$$$ setN || Args.$$$ ruleN) -- Args.colon)) scan; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
914 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
915 |
val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |-- |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36602
diff
changeset
|
916 |
Parse.and_list1' (Scan.repeat (unless_more_args free))) []; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
917 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
918 |
val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |-- |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
919 |
Scan.repeat1 (unless_more_args inst)) []; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
920 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
921 |
in |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
922 |
|
58826 | 923 |
val _ = |
924 |
Theory.setup |
|
925 |
(Method.setup @{binding cases} |
|
926 |
(Scan.lift (Args.mode no_simpN) -- |
|
927 |
(Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >> |
|
928 |
(fn (no_simp, (insts, opt_rule)) => fn ctxt => |
|
929 |
METHOD_CASES (fn facts => Seq.DETERM (HEADGOAL |
|
930 |
(cases_tac ctxt (not no_simp) insts opt_rule facts))))) |
|
931 |
"case analysis on types or predicates/sets"); |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
932 |
|
58002 | 933 |
fun gen_induct_setup binding tac = |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
934 |
Method.setup binding |
53168 | 935 |
(Scan.lift (Args.mode no_simpN) -- |
936 |
(Parse.and_list' (Scan.repeat (unless_more_args def_inst)) -- |
|
937 |
(arbitrary -- taking -- Scan.option induct_rule)) >> |
|
58002 | 938 |
(fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn ctxt => fn facts => |
939 |
Seq.DETERM (HEADGOAL (tac ctxt (not no_simp) insts arbitrary taking opt_rule facts)))) |
|
30722
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents:
30560
diff
changeset
|
940 |
"induction on types or predicates/sets"; |
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
941 |
|
58826 | 942 |
val _ = Theory.setup (gen_induct_setup @{binding induct} induct_tac); |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44942
diff
changeset
|
943 |
|
58826 | 944 |
val _ = |
945 |
Theory.setup |
|
946 |
(Method.setup @{binding coinduct} |
|
947 |
(Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >> |
|
948 |
(fn ((insts, taking), opt_rule) => fn ctxt => fn facts => |
|
949 |
Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts)))) |
|
950 |
"coinduction on types or predicates/sets"); |
|
24830
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
951 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
952 |
end; |
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
953 |
|
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents:
diff
changeset
|
954 |
end; |