| author | wenzelm | 
| Thu, 01 Nov 2018 13:53:29 +0100 | |
| changeset 69222 | 8365124a86ae | 
| parent 67649 | 1e1782c1aedf | 
| child 69590 | e65314985426 | 
| 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  | 
| 59970 | 62  | 
val atomize_term: Proof.context -> 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  | 
| 59970 | 66  | 
val rulified_term: Proof.context -> thm -> 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  | 
| 61844 | 74  | 
val cases_context_tactic: bool -> term option list list -> thm option ->  | 
75  | 
thm list -> int -> context_tactic  | 
|
76  | 
val cases_tac: Proof.context -> bool -> term option list list -> thm option ->  | 
|
77  | 
thm list -> int -> tactic  | 
|
| 
27323
 
385c0ce33173
tuned get_inductT: *all* rules for missing instantiation;
 
wenzelm 
parents: 
27140 
diff
changeset
 | 
78  | 
val get_inductT: Proof.context -> term option list list -> thm list list  | 
| 61844 | 79  | 
val gen_induct_context_tactic: ((Rule_Cases.info * int) * thm -> (Rule_Cases.info * int) * thm) ->  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
80  | 
bool -> (binding option * (term * bool)) option list list ->  | 
| 
45014
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents: 
44942 
diff
changeset
 | 
81  | 
(string * typ) list list -> term option list -> thm list option ->  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
82  | 
thm list -> int -> context_tactic  | 
| 61844 | 83  | 
val gen_induct_tac: Proof.context ->  | 
84  | 
((Rule_Cases.info * int) * thm -> (Rule_Cases.info * int) * thm) ->  | 
|
85  | 
bool -> (binding option * (term * bool)) option list list ->  | 
|
86  | 
(string * typ) list list -> term option list -> thm list option ->  | 
|
87  | 
thm list -> int -> tactic  | 
|
88  | 
val induct_context_tactic: bool ->  | 
|
89  | 
(binding option * (term * bool)) option list list ->  | 
|
| 26924 | 90  | 
(string * typ) list list -> term option list -> thm list option ->  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
91  | 
thm list -> int -> context_tactic  | 
| 61844 | 92  | 
val induct_tac: Proof.context -> bool ->  | 
93  | 
(binding option * (term * bool)) option list list ->  | 
|
94  | 
(string * typ) list list -> term option list -> thm list option ->  | 
|
95  | 
thm list -> int -> tactic  | 
|
96  | 
val coinduct_context_tactic: term option list -> term option list -> thm option ->  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
97  | 
thm list -> int -> context_tactic  | 
| 61844 | 98  | 
val coinduct_tac: Proof.context -> term option list -> term option list -> thm option ->  | 
99  | 
thm list -> int -> tactic  | 
|
| 
45014
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents: 
44942 
diff
changeset
 | 
100  | 
val gen_induct_setup: binding ->  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
101  | 
(bool -> (binding option * (term * bool)) option list list ->  | 
| 
45014
 
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
 
nipkow 
parents: 
44942 
diff
changeset
 | 
102  | 
(string * typ) list list -> term option list -> thm list option ->  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
103  | 
thm list -> int -> context_tactic) -> local_theory -> local_theory  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
|
| 37524 | 106  | 
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
 | 
107  | 
struct  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
|
| 37523 | 109  | 
(** 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
 | 
110  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
fun vars_of tm =  | 
| 45131 | 112  | 
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
 | 
113  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
local  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
|
| 37523 | 116  | 
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
 | 
117  | 
|
| 
47060
 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
118  | 
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
 | 
119  | 
  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
 | 
120  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
|
| 
47060
 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
123  | 
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
 | 
124  | 
  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
 | 
125  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
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
 | 
127  | 
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
 | 
128  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
|
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
133  | 
(** constraint simplification **)  | 
| 
 
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  | 
(* rearrange parameters and premises to allow application of one-point-rules *)  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
136  | 
|
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
137  | 
fun swap_params_conv ctxt i j cv =  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
138  | 
let  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
139  | 
fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
140  | 
| conv1 k ctxt =  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
141  | 
          Conv.rewr_conv @{thm swap_params} then_conv
 | 
| 60575 | 142  | 
Conv.forall_conv (conv1 (k - 1) o snd) ctxt;  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
143  | 
fun conv2 0 ctxt = conv1 j ctxt  | 
| 60575 | 144  | 
| 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
 | 
145  | 
in conv2 i ctxt end;  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
146  | 
|
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
147  | 
fun swap_prems_conv 0 = Conv.all_conv  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
148  | 
| swap_prems_conv i =  | 
| 37525 | 149  | 
Conv.implies_concl_conv (swap_prems_conv (i - 1)) then_conv  | 
| 60575 | 150  | 
Conv.rewr_conv Drule.swap_prems_eq;  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
151  | 
|
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
152  | 
fun find_eq ctxt t =  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
153  | 
let  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
154  | 
val l = length (Logic.strip_params t);  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
155  | 
val Hs = Logic.strip_assums_hyp t;  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
156  | 
fun find (i, t) =  | 
| 59970 | 157  | 
(case Induct_Args.dest_def (Object_Logic.drop_judgment ctxt t) of  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
158  | 
SOME (Bound j, _) => SOME (i, j)  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
159  | 
| SOME (_, Bound j) => SOME (i, j)  | 
| 37525 | 160  | 
| _ => NONE);  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
161  | 
in  | 
| 37525 | 162  | 
(case get_first find (map_index I Hs) of  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
163  | 
NONE => NONE  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
164  | 
| SOME (0, 0) => NONE  | 
| 37525 | 165  | 
| SOME (i, j) => SOME (i, l - j - 1, j))  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
166  | 
end;  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
167  | 
|
| 37525 | 168  | 
fun mk_swap_rrule ctxt ct =  | 
| 59582 | 169  | 
(case find_eq ctxt (Thm.term_of ct) of  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
170  | 
NONE => NONE  | 
| 37525 | 171  | 
| 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
 | 
172  | 
|
| 37525 | 173  | 
val rearrange_eqs_simproc =  | 
| 61144 | 174  | 
  Simplifier.make_simproc @{context} "rearrange_eqs"
 | 
| 
62958
 
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
 
wenzelm 
parents: 
62913 
diff
changeset
 | 
175  | 
   {lhss = [@{term "Pure.all (t :: 'a::{} \<Rightarrow> prop)"}],
 | 
| 62913 | 176  | 
proc = fn _ => fn ctxt => fn ct => mk_swap_rrule ctxt ct};  | 
| 37525 | 177  | 
|
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
178  | 
|
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
179  | 
(* 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
 | 
180  | 
|
| 59972 | 181  | 
fun rotate_conv 0 _ 0 = Conv.all_conv  | 
| 37525 | 182  | 
| rotate_conv 0 j k = swap_prems_conv j then_conv rotate_conv 1 j (k - 1)  | 
183  | 
| 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
 | 
184  | 
|
| 59972 | 185  | 
fun rotate_tac _ 0 = K all_tac  | 
| 37525 | 186  | 
| rotate_tac j k = SUBGOAL (fn (goal, i) =>  | 
187  | 
CONVERSION (rotate_conv  | 
|
188  | 
j (length (Logic.strip_assums_hyp goal) - j - k) k) i);  | 
|
189  | 
||
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
190  | 
|
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
191  | 
(* rulify operators around definition *)  | 
| 
 
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  | 
fun rulify_defs_conv ctxt ct =  | 
| 59582 | 194  | 
if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) andalso  | 
| 59970 | 195  | 
not (is_some (Induct_Args.dest_def (Object_Logic.drop_judgment ctxt (Thm.term_of ct))))  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
196  | 
then  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
197  | 
(Conv.forall_conv (rulify_defs_conv o snd) ctxt else_conv  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
198  | 
Conv.implies_conv (Conv.try_conv (rulify_defs_conv ctxt))  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
199  | 
(Conv.try_conv (rulify_defs_conv ctxt)) else_conv  | 
| 37524 | 200  | 
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
 | 
201  | 
Conv.try_conv (rulify_defs_conv ctxt)) ct  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
202  | 
else Conv.no_conv ct;  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
203  | 
|
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
204  | 
|
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
205  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
206  | 
(** induct data **)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
(* rules *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
|
| 30560 | 210  | 
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
 | 
211  | 
|
| 33373 | 212  | 
fun init_rules index : rules =  | 
213  | 
Item_Net.init  | 
|
214  | 
(fn ((s1, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm_prop (th1, th2))  | 
|
215  | 
(single o index);  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
216  | 
|
| 27140 | 217  | 
fun filter_rules (rs: rules) th =  | 
| 30560 | 218  | 
filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs);  | 
| 27140 | 219  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
fun pretty_rules ctxt kind rs =  | 
| 30560 | 221  | 
let val thms = map snd (Item_Net.content rs)  | 
| 61268 | 222  | 
in Pretty.big_list kind (map (Thm.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
 | 
223  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
224  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
(* context data *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
226  | 
|
| 37524 | 227  | 
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
 | 
228  | 
(  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
229  | 
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
 | 
230  | 
val empty =  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
231  | 
((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
 | 
232  | 
(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
 | 
233  | 
(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
 | 
234  | 
     simpset_of (empty_simpset @{context}
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51658 
diff
changeset
 | 
235  | 
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
 | 
236  | 
val extend = I;  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
237  | 
fun merge (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1), simpset1),  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
238  | 
((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2), simpset2)) =  | 
| 30560 | 239  | 
((Item_Net.merge (casesT1, casesT2), Item_Net.merge (casesP1, casesP2)),  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
240  | 
(Item_Net.merge (inductT1, inductT2), Item_Net.merge (inductP1, inductP2)),  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
241  | 
(Item_Net.merge (coinductT1, coinductT2), Item_Net.merge (coinductP1, coinductP2)),  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
242  | 
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
 | 
243  | 
);  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
244  | 
|
| 37524 | 245  | 
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
 | 
246  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
247  | 
fun dest_rules ctxt =  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
248  | 
let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in  | 
| 30560 | 249  | 
    {type_cases = Item_Net.content casesT,
 | 
250  | 
pred_cases = Item_Net.content casesP,  | 
|
251  | 
type_induct = Item_Net.content inductT,  | 
|
252  | 
pred_induct = Item_Net.content inductP,  | 
|
253  | 
type_coinduct = Item_Net.content coinductT,  | 
|
254  | 
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
 | 
255  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
256  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
257  | 
fun print_rules ctxt =  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
258  | 
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
 | 
259  | 
[pretty_rules ctxt "coinduct type:" coinductT,  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
260  | 
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
 | 
261  | 
pretty_rules ctxt "induct type:" inductT,  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
262  | 
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
 | 
263  | 
pretty_rules ctxt "cases type:" casesT,  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
264  | 
pretty_rules ctxt "cases pred:" casesP]  | 
| 
56334
 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 
wenzelm 
parents: 
56245 
diff
changeset
 | 
265  | 
|> Pretty.writeln_chunks  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
266  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
267  | 
|
| 24867 | 268  | 
val _ =  | 
| 67149 | 269  | 
Outer_Syntax.command \<^command_keyword>\<open>print_induct_rules\<close>  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
45375 
diff
changeset
 | 
270  | 
"print induction and cases rules"  | 
| 
60097
 
d20ca79d50e4
discontinued pointless warnings: commands are only defined inside a theory context;
 
wenzelm 
parents: 
59972 
diff
changeset
 | 
271  | 
(Scan.succeed (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
 | 
272  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
273  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
274  | 
(* access rules *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
275  | 
|
| 61058 | 276  | 
local  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
277  | 
|
| 61058 | 278  | 
fun lookup_rule which ctxt =  | 
279  | 
AList.lookup (op =) (Item_Net.content (which (get_local ctxt)))  | 
|
| 67649 | 280  | 
#> Option.map (Thm.transfer' ctxt);  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
281  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
282  | 
fun find_rules which how ctxt x =  | 
| 61058 | 283  | 
Item_Net.retrieve (which (get_local ctxt)) (how x)  | 
| 67649 | 284  | 
|> map (Thm.transfer' ctxt o snd);  | 
| 61058 | 285  | 
|
286  | 
in  | 
|
287  | 
||
288  | 
val lookup_casesT = lookup_rule (#1 o #1);  | 
|
289  | 
val lookup_casesP = lookup_rule (#2 o #1);  | 
|
290  | 
val lookup_inductT = lookup_rule (#1 o #2);  | 
|
291  | 
val lookup_inductP = lookup_rule (#2 o #2);  | 
|
292  | 
val lookup_coinductT = lookup_rule (#1 o #3);  | 
|
293  | 
val lookup_coinductP = lookup_rule (#2 o #3);  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
294  | 
|
| 37523 | 295  | 
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
 | 
296  | 
val find_casesP = find_rules (#2 o #1) I;  | 
| 37523 | 297  | 
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
 | 
298  | 
val find_inductP = find_rules (#2 o #2) I;  | 
| 37523 | 299  | 
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
 | 
300  | 
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
 | 
301  | 
|
| 61058 | 302  | 
end;  | 
303  | 
||
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
304  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
305  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
306  | 
(** attributes **)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
307  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
308  | 
local  | 
| 
 
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  | 
fun mk_att f g name =  | 
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
45328 
diff
changeset
 | 
311  | 
Thm.mixed_attribute (fn (context, thm) =>  | 
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
45328 
diff
changeset
 | 
312  | 
let  | 
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
45328 
diff
changeset
 | 
313  | 
val thm' = g thm;  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61844 
diff
changeset
 | 
314  | 
val context' =  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61844 
diff
changeset
 | 
315  | 
if Thm.is_free_dummy thm then context  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61844 
diff
changeset
 | 
316  | 
else Data.map (f (name, Thm.trim_context thm')) context;  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
45328 
diff
changeset
 | 
317  | 
in (context', thm') end);  | 
| 27140 | 318  | 
|
| 37525 | 319  | 
fun del_att which =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
59057 
diff
changeset
 | 
320  | 
Thm.declaration_attribute (fn th => Data.map (which (apply2 (fn rs =>  | 
| 37525 | 321  | 
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
 | 
322  | 
|
| 
59057
 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 
wenzelm 
parents: 
58957 
diff
changeset
 | 
323  | 
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
 | 
324  | 
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
 | 
325  | 
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
 | 
326  | 
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
 | 
327  | 
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
 | 
328  | 
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
 | 
329  | 
|
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
45328 
diff
changeset
 | 
330  | 
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
 | 
331  | 
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
 | 
332  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
333  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
334  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
335  | 
val cases_type = mk_att add_casesT consumes0;  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
336  | 
val cases_pred = mk_att add_casesP consumes1;  | 
| 
59057
 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 
wenzelm 
parents: 
58957 
diff
changeset
 | 
337  | 
val cases_del = del_att @{apply 4(1)};
 | 
| 27140 | 338  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
339  | 
val induct_type = mk_att add_inductT consumes0;  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
340  | 
val induct_pred = mk_att add_inductP consumes1;  | 
| 
59057
 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 
wenzelm 
parents: 
58957 
diff
changeset
 | 
341  | 
val induct_del = del_att @{apply 4(2)};
 | 
| 27140 | 342  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
343  | 
val coinduct_type = mk_att add_coinductT consumes0;  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
344  | 
val coinduct_pred = mk_att add_coinductP consumes1;  | 
| 
59057
 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 
wenzelm 
parents: 
58957 
diff
changeset
 | 
345  | 
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
 | 
346  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51658 
diff
changeset
 | 
347  | 
fun map_simpset f context =  | 
| 
59057
 
5b649fb2f2e1
added ML antiquotation @{apply n} or @{apply n(k)};
 
wenzelm 
parents: 
58957 
diff
changeset
 | 
348  | 
  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
 | 
349  | 
|
| 
 
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
 
wenzelm 
parents: 
35625 
diff
changeset
 | 
350  | 
fun induct_simp f =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51658 
diff
changeset
 | 
351  | 
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
 | 
352  | 
|
| 
 
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
 
wenzelm 
parents: 
35625 
diff
changeset
 | 
353  | 
val induct_simp_add = induct_simp (op addsimps);  | 
| 
 
0cbcdfd9e527
slightly more standard induct_simp_add/del attributes;
 
wenzelm 
parents: 
35625 
diff
changeset
 | 
354  | 
val induct_simp_del = induct_simp (op delsimps);  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
355  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
356  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
357  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
358  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
359  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
360  | 
(** attribute syntax **)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
361  | 
|
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
362  | 
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
 | 
363  | 
val casesN = "cases";  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
364  | 
val inductN = "induct";  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
365  | 
val coinductN = "coinduct";  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
366  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
367  | 
val typeN = "type";  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
368  | 
val predN = "pred";  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
369  | 
val setN = "set";  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
370  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
371  | 
local  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
372  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
373  | 
fun spec k arg =  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
374  | 
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
 | 
375  | 
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
 | 
376  | 
|
| 30528 | 377  | 
fun attrib add_type add_pred del =  | 
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
378  | 
  spec typeN (Args.type_name {proper = false, strict = false}) >> add_type ||
 | 
| 55954 | 379  | 
  spec predN (Args.const {proper = false, strict = false}) >> add_pred ||
 | 
380  | 
  spec setN (Args.const {proper = false, strict = false}) >> add_pred ||
 | 
|
| 30528 | 381  | 
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
 | 
382  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
383  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
384  | 
|
| 58826 | 385  | 
val _ =  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
386  | 
Theory.local_setup  | 
| 67149 | 387  | 
(Attrib.local_setup \<^binding>\<open>cases\<close> (attrib cases_type cases_pred cases_del)  | 
| 58826 | 388  | 
"declaration of cases rule" #>  | 
| 67149 | 389  | 
Attrib.local_setup \<^binding>\<open>induct\<close> (attrib induct_type induct_pred induct_del)  | 
| 58826 | 390  | 
"declaration of induction rule" #>  | 
| 67149 | 391  | 
Attrib.local_setup \<^binding>\<open>coinduct\<close> (attrib coinduct_type coinduct_pred coinduct_del)  | 
| 58826 | 392  | 
"declaration of coinduction rule" #>  | 
| 67149 | 393  | 
Attrib.local_setup \<^binding>\<open>induct_simp\<close> (Attrib.add_del induct_simp_add induct_simp_del)  | 
| 58826 | 394  | 
"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
 | 
395  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
396  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
397  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
398  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
399  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
400  | 
(** method utils **)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
401  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
402  | 
(* alignment *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
403  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
404  | 
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
 | 
405  | 
let val m = length xs and n = length ys  | 
| 33957 | 406  | 
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
 | 
407  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
408  | 
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
 | 
409  | 
let val m = length xs and n = length ys  | 
| 33957 | 410  | 
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
 | 
411  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
412  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
413  | 
(* prep_inst *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
414  | 
|
| 
32432
 
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
 
wenzelm 
parents: 
32188 
diff
changeset
 | 
415  | 
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
 | 
416  | 
let  | 
| 60784 | 417  | 
fun prep_var (Var (x, xT), SOME t) =  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
418  | 
let  | 
| 59843 | 419  | 
val ct = Thm.cterm_of ctxt (tune t);  | 
| 59586 | 420  | 
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
 | 
421  | 
in  | 
| 60784 | 422  | 
if Type.could_unify (tT, xT) then SOME (x, ct)  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
423  | 
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
 | 
424  | 
[Pretty.str "Ill-typed instantiation:", Pretty.fbrk,  | 
| 
32432
 
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
 
wenzelm 
parents: 
32188 
diff
changeset
 | 
425  | 
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
 | 
426  | 
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
 | 
427  | 
end  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
428  | 
| prep_var (_, NONE) = NONE;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
429  | 
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
 | 
430  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
431  | 
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
 | 
432  | 
|> map_filter prep_var  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
433  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
434  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
435  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
436  | 
(* trace_rules *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
437  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
438  | 
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
 | 
439  | 
| 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
 | 
440  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
441  | 
|
| 
34987
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
442  | 
(* mark equality constraints in cases rule *)  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
443  | 
|
| 37524 | 444  | 
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
 | 
445  | 
|
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
446  | 
fun mark_constraints n ctxt = Conv.fconv_rule  | 
| 45130 | 447  | 
(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
 | 
448  | 
(Raw_Simplifier.rewrite ctxt false [equal_def']))) ctxt));  | 
| 
34987
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
449  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
450  | 
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
 | 
451  | 
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
 | 
452  | 
|
| 37525 | 453  | 
|
| 
34987
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
454  | 
(* simplify *)  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
455  | 
|
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
456  | 
fun simplify_conv' ctxt =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51658 
diff
changeset
 | 
457  | 
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
 | 
458  | 
|
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
459  | 
fun simplify_conv ctxt ct =  | 
| 59582 | 460  | 
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
 | 
461  | 
(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
 | 
462  | 
else Conv.all_conv ct;  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
463  | 
|
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
464  | 
fun gen_simplified_rule cv ctxt =  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
465  | 
Conv.fconv_rule (Conv.prems_conv ~1 (cv ctxt));  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
466  | 
|
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
467  | 
val simplified_rule' = gen_simplified_rule simplify_conv';  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
468  | 
val simplified_rule = gen_simplified_rule simplify_conv;  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
469  | 
|
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
470  | 
fun simplify_tac ctxt = CONVERSION (simplify_conv ctxt);  | 
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
471  | 
|
| 37524 | 472  | 
val trivial_tac = Induct_Args.trivial_tac;  | 
| 
34987
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
473  | 
|
| 
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
474  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
475  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
476  | 
(** cases method **)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
477  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
478  | 
(*  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
479  | 
rule selection scheme:  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
480  | 
cases - default case split  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
481  | 
`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
 | 
482  | 
cases t - type cases  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
483  | 
... cases ... r - explicit rule  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
484  | 
*)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
485  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
486  | 
local  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
487  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
488  | 
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
 | 
489  | 
| get_casesT _ _ = [];  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
490  | 
|
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
491  | 
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
 | 
492  | 
| get_casesP _ _ = [];  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
493  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
494  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
495  | 
|
| 61844 | 496  | 
fun cases_context_tactic simp insts opt_rule facts i : context_tactic =  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
497  | 
fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
498  | 
let  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
499  | 
fun inst_rule r =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
500  | 
(if null insts then r  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
501  | 
else  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
502  | 
(align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
503  | 
|> maps (prep_inst ctxt align_left I)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
504  | 
|> infer_instantiate ctxt) r)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
505  | 
|> simp ? mark_constraints (Rule_Cases.get_constraints r) ctxt  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
506  | 
|> pair (Rule_Cases.get r);  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
507  | 
|
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
508  | 
val (opt_rule', facts') =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
509  | 
(case (opt_rule, facts) of  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
510  | 
(NONE, th :: ths) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
511  | 
if is_some (Object_Logic.elim_concl ctxt th) then (SOME th, ths)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
512  | 
else (opt_rule, facts)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
513  | 
| _ => (opt_rule, facts));  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
514  | 
|
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
515  | 
val ruleq =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
516  | 
(case opt_rule' of  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
517  | 
SOME r => Seq.single (inst_rule r)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
518  | 
| NONE =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
519  | 
(get_casesP ctxt facts' @ get_casesT ctxt insts @ [Induct_Args.cases_default])  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
520  | 
|> tap (trace_rules ctxt casesN)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
521  | 
|> Seq.of_list |> Seq.maps (Seq.try inst_rule));  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
522  | 
in  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
523  | 
ruleq  | 
| 60455 | 524  | 
|> Seq.maps (Rule_Cases.consume ctxt [] facts')  | 
525  | 
|> Seq.maps (fn ((cases, (_, facts'')), rule) =>  | 
|
| 45131 | 526  | 
let  | 
527  | 
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
 | 
528  | 
|> simp ? (simplified_rule' ctxt #> unmark_constraints ctxt);  | 
| 
34987
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
529  | 
in  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
530  | 
CONTEXT_CASES (Rule_Cases.make_common ctxt  | 
| 59970 | 531  | 
(Thm.prop_of (Rule_Cases.internalize_params rule')) cases)  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
532  | 
((Method.insert_tac ctxt facts'' THEN' resolve_tac ctxt [rule'] THEN_ALL_NEW  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
533  | 
(if simp then TRY o trivial_tac ctxt else K all_tac)) i) (ctxt, st)  | 
| 
34987
 
c1e8af37ee75
Added infrastructure for simplifying equality constraints in cases rules.
 
berghofe 
parents: 
34907 
diff
changeset
 | 
534  | 
end)  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
535  | 
end;  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
536  | 
|
| 61844 | 537  | 
fun cases_tac ctxt x1 x2 x3 x4 x5 =  | 
538  | 
Method.NO_CONTEXT_TACTIC ctxt (cases_context_tactic x1 x2 x3 x4 x5);  | 
|
539  | 
||
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
540  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
541  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
542  | 
|
| 
 
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  | 
(** induct method **)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
545  | 
|
| 59929 | 546  | 
val conjunction_congs = @{thms Pure.all_conjunction imp_conjunction};
 | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
547  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
549  | 
(* atomize *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
550  | 
|
| 59970 | 551  | 
fun atomize_term ctxt =  | 
552  | 
Raw_Simplifier.rewrite_term (Proof_Context.theory_of ctxt) Induct_Args.atomize []  | 
|
553  | 
#> Object_Logic.drop_judgment ctxt;  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
554  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
555  | 
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
 | 
556  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
557  | 
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
 | 
558  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
559  | 
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
 | 
560  | 
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
 | 
561  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
562  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
563  | 
(* rulify *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
564  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
565  | 
fun rulify_term thy =  | 
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
566  | 
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
 | 
567  | 
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
 | 
568  | 
|
| 59970 | 569  | 
fun rulified_term ctxt thm =  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
570  | 
let  | 
| 59970 | 571  | 
val rulify = rulify_term (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
 | 
572  | 
val (As, B) = Logic.strip_horn (Thm.prop_of thm);  | 
| 59970 | 573  | 
in Logic.list_implies (map rulify As, rulify B) end;  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
574  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
575  | 
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
 | 
576  | 
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
 | 
577  | 
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
 | 
578  | 
Goal.conjunction_tac THEN_ALL_NEW  | 
| 59929 | 579  | 
  (rewrite_goal_tac ctxt @{thms 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
 | 
580  | 
|
| 
 
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  | 
(* prepare rule *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
583  | 
|
| 
32432
 
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
 
wenzelm 
parents: 
32188 
diff
changeset
 | 
584  | 
fun rule_instance ctxt inst rule =  | 
| 60784 | 585  | 
infer_instantiate ctxt (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
 | 
586  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
587  | 
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
 | 
588  | 
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
 | 
589  | 
|> 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
 | 
590  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
591  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
592  | 
(* 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
 | 
593  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
594  | 
local  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
595  | 
|
| 59843 | 596  | 
fun dest_env ctxt env =  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
597  | 
let  | 
| 32032 | 598  | 
val pairs = Vartab.dest (Envir.term_env env);  | 
599  | 
val types = Vartab.dest (Envir.type_env env);  | 
|
| 59843 | 600  | 
val ts = map (Thm.cterm_of ctxt o Envir.norm_term env o #2 o #2) pairs;  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60576 
diff
changeset
 | 
601  | 
val xs = map #1 pairs ~~ map Thm.typ_of_cterm ts;  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60576 
diff
changeset
 | 
602  | 
in (map (fn (ai, (S, T)) => ((ai, S), Thm.ctyp_of ctxt 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
 | 
603  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
604  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
605  | 
|
| 26940 | 606  | 
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
 | 
607  | 
let  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26568 
diff
changeset
 | 
608  | 
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
 | 
609  | 
val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*)  | 
| 29276 | 610  | 
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
 | 
611  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
612  | 
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
 | 
613  | 
      (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
 | 
614  | 
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
 | 
615  | 
Seq.single rule)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
616  | 
else  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
617  | 
let  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
618  | 
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
 | 
619  | 
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
 | 
620  | 
in  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
621  | 
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
 | 
622  | 
[(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule'))  | 
| 59843 | 623  | 
|> Seq.map (fn env => Drule.instantiate_normalize (dest_env ctxt env) rule')  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
624  | 
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
 | 
625  | 
end  | 
| 
 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 
wenzelm 
parents: 
43326 
diff
changeset
 | 
626  | 
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
 | 
627  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
628  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
629  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
630  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
631  | 
(* 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
 | 
632  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
633  | 
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
 | 
634  | 
let  | 
| 
42488
 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
635  | 
val x = Name.clean (Variable.revert_fixed ctxt z);  | 
| 59972 | 636  | 
fun index _ [] = []  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
637  | 
| index i (y :: ys) =  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
638  | 
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
 | 
639  | 
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
 | 
640  | 
fun rename_params [] = []  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
641  | 
| 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
 | 
642  | 
(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
 | 
643  | 
| 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
 | 
644  | 
fun rename_asm A =  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
645  | 
let  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
646  | 
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
 | 
647  | 
val xs' =  | 
| 28375 | 648  | 
(case filter (fn x' => x' = x) xs of  | 
| 60575 | 649  | 
[] => xs  | 
650  | 
| [_] => xs  | 
|
651  | 
| _ => index 1 xs);  | 
|
| 45328 | 652  | 
in Logic.list_rename_params xs' A end;  | 
| 60313 | 653  | 
fun rename_prop prop =  | 
654  | 
let val (As, C) = Logic.strip_horn prop  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
655  | 
in Logic.list_implies (map rename_asm As, C) end;  | 
| 60313 | 656  | 
val thm' = Thm.renamed_prop (rename_prop (Thm.prop_of thm)) thm;  | 
| 33368 | 657  | 
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
 | 
658  | 
| 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
 | 
659  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
660  | 
|
| 45132 | 661  | 
(* arbitrary_tac *)  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
662  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
663  | 
local  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
664  | 
|
| 56245 | 665  | 
fun goal_prefix k ((c as Const (@{const_name Pure.all}, _)) $ Abs (a, T, B)) =
 | 
666  | 
c $ Abs (a, T, goal_prefix k B)  | 
|
| 45156 | 667  | 
| goal_prefix 0 _ = Term.dummy_prop  | 
| 56245 | 668  | 
  | goal_prefix k ((c as Const (@{const_name Pure.imp}, _)) $ A $ B) =
 | 
669  | 
c $ A $ goal_prefix (k - 1) B  | 
|
| 45156 | 670  | 
| 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
 | 
671  | 
|
| 56245 | 672  | 
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
 | 
673  | 
| goal_params 0 _ = 0  | 
| 56245 | 674  | 
  | 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
 | 
675  | 
| goal_params _ _ = 0;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
676  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
677  | 
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
 | 
678  | 
let  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
679  | 
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
 | 
680  | 
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
 | 
681  | 
      @{thm Pure.meta_spec}
 | 
| 
42488
 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
682  | 
|> Thm.rename_params_rule ([Name.clean (Variable.revert_fixed ctxt x)], 1)  | 
| 59843 | 683  | 
|> Thm.lift_rule (Thm.cterm_of ctxt prfx)  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
684  | 
|> `(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
 | 
685  | 
|-> (fn pred $ arg =>  | 
| 60784 | 686  | 
infer_instantiate ctxt  | 
687  | 
[(#1 (dest_Var (head_of pred)), Thm.cterm_of ctxt (Logic.rlist_abs (xs, body))),  | 
|
688  | 
(#1 (dest_Var (head_of arg)), Thm.cterm_of ctxt (Logic.rlist_abs (xs, v)))]);  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
689  | 
|
| 56245 | 690  | 
    fun goal_concl k xs (Const (@{const_name Pure.all}, _) $ Abs (a, T, B)) =
 | 
691  | 
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
 | 
692  | 
| 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
 | 
693  | 
if not (Term.exists_subterm (fn t => t aconv v) B) then NONE  | 
| 44241 | 694  | 
else SOME (xs, absfree (x, T) (Term.incr_boundvars 1 B))  | 
| 56245 | 695  | 
      | goal_concl k xs (Const (@{const_name Pure.imp}, _) $ _ $ B) =
 | 
696  | 
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
 | 
697  | 
| goal_concl _ _ _ = NONE;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
698  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
699  | 
(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
 | 
700  | 
SOME concl =>  | 
| 
58956
 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
701  | 
(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
 | 
702  | 
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
 | 
703  | 
| NONE => all_tac)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
704  | 
end);  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
705  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
706  | 
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
 | 
707  | 
CONVERSION o  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
708  | 
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
 | 
709  | 
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
 | 
710  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
711  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
712  | 
|
| 45132 | 713  | 
fun arbitrary_tac _ _ [] = K all_tac  | 
714  | 
| 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
 | 
715  | 
(EVERY' (map (meta_spec_tac ctxt n) xs) THEN'  | 
| 24832 | 716  | 
(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
 | 
717  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
718  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
719  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
720  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
721  | 
(* add_defs *)  | 
| 
 
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  | 
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
 | 
724  | 
let  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
725  | 
fun add (SOME (_, (t, true))) ctxt = ((SOME t, []), ctxt)  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
726  | 
| add (SOME (SOME x, (t, _))) ctxt =  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27865 
diff
changeset
 | 
727  | 
let val ([(lhs, (_, th))], ctxt') =  | 
| 63344 | 728  | 
Local_Defs.define [((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
 | 
729  | 
in ((SOME lhs, [th]), ctxt') end  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
730  | 
| add (SOME (NONE, (t as Free _, _))) ctxt = ((SOME t, []), ctxt)  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
731  | 
| add (SOME (NONE, (t, _))) ctxt =  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
732  | 
let  | 
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43278 
diff
changeset
 | 
733  | 
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
 | 
734  | 
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
 | 
735  | 
val ([(lhs, (_, th))], ctxt') = ctxt  | 
| 63344 | 736  | 
|> Local_Defs.define [((x, NoSyn), ((Thm.def_binding x, []), t))];  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
737  | 
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
 | 
738  | 
| 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
 | 
739  | 
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
 | 
740  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
741  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
742  | 
(* induct_tac *)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
743  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
744  | 
(*  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
745  | 
rule selection scheme:  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
746  | 
`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
 | 
747  | 
induct x - type induction  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
748  | 
... induct ... r - explicit rule  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
749  | 
*)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
750  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
751  | 
fun get_inductT ctxt insts =  | 
| 32188 | 752  | 
fold_rev (map_product cons) (insts |> map  | 
| 
27323
 
385c0ce33173
tuned get_inductT: *all* rules for missing instantiation;
 
wenzelm 
parents: 
27140 
diff
changeset
 | 
753  | 
((fn [] => NONE | ts => List.last ts) #>  | 
| 
 
385c0ce33173
tuned get_inductT: *all* rules for missing instantiation;
 
wenzelm 
parents: 
27140 
diff
changeset
 | 
754  | 
        (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
 | 
755  | 
find_inductT ctxt)) [[]]  | 
| 33368 | 756  | 
|> 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
 | 
757  | 
|
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
758  | 
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
 | 
759  | 
| get_inductP _ _ = [];  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
760  | 
|
| 61844 | 761  | 
fun gen_induct_context_tactic mod_cases simp def_insts arbitrary taking opt_rule facts =  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
762  | 
CONTEXT_SUBGOAL (fn (_, i) => fn (ctxt, st) =>  | 
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
763  | 
let  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
764  | 
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
 | 
765  | 
val atomized_defs = map (map (Conv.fconv_rule (atomize_cterm defs_ctxt))) defs;  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
766  | 
|
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
767  | 
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
 | 
768  | 
(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
 | 
769  | 
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
 | 
770  | 
(map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts  | 
| 59970 | 771  | 
|> maps (prep_inst ctxt align_right (atomize_term ctxt))  | 
| 60784 | 772  | 
|> infer_instantiate ctxt) r |> pair (Rule_Cases.get r))  | 
| 59970 | 773  | 
|> mod_cases  | 
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
774  | 
|> (fn ((cases, consumes), th) => (((cases, concls), consumes), th));  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
775  | 
|
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
776  | 
val ruleq =  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
777  | 
(case opt_rule of  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
778  | 
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
 | 
779  | 
| NONE =>  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
780  | 
(get_inductP ctxt facts @  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
781  | 
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
 | 
782  | 
|> 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
 | 
783  | 
|> 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
 | 
784  | 
|> Seq.of_list |> Seq.maps (Seq.try inst_rule));  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
785  | 
|
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
786  | 
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
 | 
787  | 
let  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
788  | 
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
 | 
789  | 
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
 | 
790  | 
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
 | 
791  | 
val cases' = if Thm.eq_thm_prop (rule', rule'') then cases else nonames cases;  | 
| 59970 | 792  | 
in Rule_Cases.make_nested ctxt (Thm.prop_of rule'') (rulified_term ctxt rule'') cases' end;  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
793  | 
|
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
794  | 
fun context_tac _ _ =  | 
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
795  | 
ruleq  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
796  | 
|> 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
 | 
797  | 
|> 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
 | 
798  | 
(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
 | 
799  | 
(CONJUNCTS (ALLGOALS  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
800  | 
let  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
801  | 
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
 | 
802  | 
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
 | 
803  | 
val xs = nth_list arbitrary (j - 1);  | 
| 60575 | 804  | 
val k = nth concls (j - 1) + more_consumes;  | 
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
805  | 
in  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
806  | 
Method.insert_tac defs_ctxt (more_facts @ adefs) THEN'  | 
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
807  | 
(if simp then  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
808  | 
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
 | 
809  | 
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
 | 
810  | 
else  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
811  | 
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
 | 
812  | 
end)  | 
| 
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
813  | 
THEN' inner_atomize_tac defs_ctxt) j))  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
814  | 
THEN' atomize_tac defs_ctxt) i st  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
815  | 
|> Seq.maps (fn st' =>  | 
| 
56231
 
b98813774a63
enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
816  | 
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
 | 
817  | 
|> 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
 | 
818  | 
|> Seq.maps (fn rule' =>  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
819  | 
CONTEXT_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
 | 
820  | 
(resolve_tac ctxt [rule'] i THEN  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
821  | 
PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) (ctxt, st'))));  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
822  | 
in  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
823  | 
(context_tac CONTEXT_THEN_ALL_NEW  | 
| 58957 | 824  | 
((if simp then simplify_tac ctxt THEN' (TRY o trivial_tac ctxt) else K all_tac)  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
825  | 
THEN_ALL_NEW rulify_tac ctxt)) i (ctxt, st)  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
826  | 
end)  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
827  | 
|
| 61844 | 828  | 
val induct_context_tactic = gen_induct_context_tactic I;  | 
829  | 
||
830  | 
fun gen_induct_tac ctxt x1 x2 x3 x4 x5 x6 x7 x8 =  | 
|
831  | 
Method.NO_CONTEXT_TACTIC ctxt (gen_induct_context_tactic x1 x2 x3 x4 x5 x6 x7 x8);  | 
|
832  | 
||
833  | 
fun induct_tac ctxt x1 x2 x3 x4 x5 x6 x7 =  | 
|
834  | 
Method.NO_CONTEXT_TACTIC ctxt (induct_context_tactic x1 x2 x3 x4 x5 x6 x7);  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
835  | 
|
| 45130 | 836  | 
|
837  | 
||
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
838  | 
(** coinduct method **)  | 
| 
 
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  | 
(*  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
841  | 
rule selection scheme:  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
842  | 
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
 | 
843  | 
coinduct x - type coinduction  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
844  | 
coinduct ... r - explicit rule  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
845  | 
*)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
846  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
847  | 
local  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
848  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
849  | 
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
 | 
850  | 
| get_coinductT _ _ = [];  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
851  | 
|
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
852  | 
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
 | 
853  | 
|
| 
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
854  | 
fun main_prop_of th =  | 
| 33368 | 855  | 
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
 | 
856  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
857  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
858  | 
|
| 61844 | 859  | 
fun coinduct_context_tactic inst taking opt_rule facts =  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
860  | 
CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
861  | 
let  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
862  | 
fun inst_rule r =  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
863  | 
if null inst then `Rule_Cases.get r  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
864  | 
else  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
865  | 
infer_instantiate ctxt (prep_inst ctxt align_right I (main_prop_of r, inst)) r  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
866  | 
|> pair (Rule_Cases.get r);  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
867  | 
in  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
868  | 
(case opt_rule of  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
869  | 
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
 | 
870  | 
| NONE =>  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
871  | 
(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
 | 
872  | 
|> tap (trace_rules ctxt coinductN)  | 
| 59972 | 873  | 
|> Seq.of_list |> Seq.maps (Seq.try inst_rule))  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53168 
diff
changeset
 | 
874  | 
|> 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
 | 
875  | 
|> Seq.maps (fn ((cases, (_, more_facts)), rule) =>  | 
| 26940 | 876  | 
guess_instance ctxt rule i st  | 
| 
32432
 
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
 
wenzelm 
parents: 
32188 
diff
changeset
 | 
877  | 
|> 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
 | 
878  | 
|> Seq.maps (fn rule' =>  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
879  | 
CONTEXT_CASES (Rule_Cases.make_common ctxt  | 
| 59970 | 880  | 
(Thm.prop_of (Rule_Cases.internalize_params rule')) cases)  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
881  | 
(Method.insert_tac ctxt more_facts i THEN resolve_tac ctxt [rule'] i) (ctxt, st)))  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
882  | 
end);  | 
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
883  | 
|
| 61844 | 884  | 
fun coinduct_tac ctxt x1 x2 x3 x4 x5 =  | 
885  | 
Method.NO_CONTEXT_TACTIC ctxt (coinduct_context_tactic x1 x2 x3 x4 x5);  | 
|
886  | 
||
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
887  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
888  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
889  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
890  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
891  | 
(** concrete syntax **)  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
892  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
893  | 
val arbitraryN = "arbitrary";  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
894  | 
val takingN = "taking";  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
895  | 
val ruleN = "rule";  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
896  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
897  | 
local  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
898  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
899  | 
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
 | 
900  | 
| 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
 | 
901  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
902  | 
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
 | 
903  | 
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
 | 
904  | 
(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
 | 
905  | 
(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
 | 
906  | 
      | 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
 | 
907  | 
|
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
908  | 
fun rule get_type get_pred =  | 
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
909  | 
  named_rule typeN (Args.type_name {proper = false, strict = false}) get_type ||
 | 
| 55954 | 910  | 
  named_rule predN (Args.const {proper = false, strict = false}) get_pred ||
 | 
911  | 
  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
 | 
912  | 
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
 | 
913  | 
|
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
914  | 
val cases_rule = rule lookup_casesT lookup_casesP >> single_rule;  | 
| 
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
915  | 
val induct_rule = rule lookup_inductT lookup_inductP;  | 
| 
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
916  | 
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
 | 
917  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
918  | 
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
 | 
919  | 
|
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
920  | 
val inst' = Scan.lift (Args.$$$ "_") >> K NONE ||  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
921  | 
Args.term >> (SOME o rpair false) ||  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
922  | 
  Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --|
 | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
923  | 
Scan.lift (Args.$$$ ")");  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
924  | 
|
| 
24830
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
925  | 
val def_inst =  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27865 
diff
changeset
 | 
926  | 
((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME)  | 
| 
34907
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
927  | 
-- (Args.term >> rpair false)) >> SOME ||  | 
| 
 
b0aaec87751c
Added infrastructure for simplifying equality constraints.
 
berghofe 
parents: 
33957 
diff
changeset
 | 
928  | 
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
 | 
929  | 
|
| 27370 | 930  | 
val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) =>  | 
931  | 
  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
 | 
932  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
933  | 
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
 | 
934  | 
((Args.$$$ arbitraryN || Args.$$$ takingN || Args.$$$ typeN ||  | 
| 
24861
 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 
wenzelm 
parents: 
24832 
diff
changeset
 | 
935  | 
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
 | 
936  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
937  | 
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
 | 
938  | 
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
 | 
939  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
940  | 
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
 | 
941  | 
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
 | 
942  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
943  | 
in  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
944  | 
|
| 58002 | 945  | 
fun gen_induct_setup binding tac =  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
946  | 
Method.local_setup binding  | 
| 53168 | 947  | 
(Scan.lift (Args.mode no_simpN) --  | 
948  | 
(Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --  | 
|
949  | 
(arbitrary -- taking -- Scan.option induct_rule)) >>  | 
|
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
950  | 
(fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn _ => fn facts =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
951  | 
Seq.DETERM (tac (not no_simp) insts arbitrary taking opt_rule facts 1)))  | 
| 
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
 | 
952  | 
"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
 | 
953  | 
|
| 58826 | 954  | 
val _ =  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
955  | 
Theory.local_setup  | 
| 67149 | 956  | 
(Method.local_setup \<^binding>\<open>cases\<close>  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
957  | 
(Scan.lift (Args.mode no_simpN) --  | 
| 
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
958  | 
(Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >>  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
959  | 
(fn (no_simp, (insts, opt_rule)) => fn _ =>  | 
| 
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
960  | 
CONTEXT_METHOD (fn facts =>  | 
| 61844 | 961  | 
Seq.DETERM (cases_context_tactic (not no_simp) insts opt_rule facts 1))))  | 
| 
59940
 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 
wenzelm 
parents: 
59936 
diff
changeset
 | 
962  | 
"case analysis on types or predicates/sets" #>  | 
| 67149 | 963  | 
gen_induct_setup \<^binding>\<open>induct\<close> induct_context_tactic #>  | 
964  | 
Method.local_setup \<^binding>\<open>coinduct\<close>  | 
|
| 58826 | 965  | 
(Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >>  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
966  | 
(fn ((insts, taking), opt_rule) => fn _ => fn facts =>  | 
| 61844 | 967  | 
Seq.DETERM (coinduct_context_tactic insts taking opt_rule facts 1)))  | 
| 58826 | 968  | 
"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
 | 
969  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
970  | 
end;  | 
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
971  | 
|
| 
 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
972  | 
end;  |