| author | haftmann | 
| Sat, 16 Sep 2023 06:38:44 +0000 | |
| changeset 78669 | 18ea58bdcf77 | 
| parent 74534 | 638301b86f0a | 
| child 79965 | 233d70cad0cf | 
| permissions | -rw-r--r-- | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Tools/Function/function_core.ML  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
2  | 
Author: Alexander Krauss, TU Muenchen  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
3  | 
|
| 41114 | 4  | 
Core of the function package.  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
7  | 
signature FUNCTION_CORE =  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
9  | 
val trace: bool Unsynchronized.ref  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
10  | 
val prepare_function : Function_Common.function_config  | 
| 63004 | 11  | 
-> binding (* defname *)  | 
| 63011 | 12  | 
-> ((binding * typ) * mixfix) list (* defined symbol *)  | 
13  | 
-> ((string * typ) list * term list * term * term) list (* specification *)  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
14  | 
-> local_theory  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
15  | 
-> (term (* f *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
16  | 
* thm (* goalstate *)  | 
| 60643 | 17  | 
* (Proof.context -> thm -> Function_Common.function_result) (* continuation *)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
18  | 
) * local_theory  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
20  | 
end  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
22  | 
structure Function_Core : FUNCTION_CORE =  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
23  | 
struct  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
24  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
25  | 
val trace = Unsynchronized.ref false  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
26  | 
fun trace_msg msg = if ! trace then tracing (msg ()) else ()  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
28  | 
val boolT = HOLogic.boolT  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
29  | 
val mk_eq = HOLogic.mk_eq  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
31  | 
open Function_Lib  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
32  | 
open Function_Common  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
33  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
34  | 
datatype globals = Globals of  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
35  | 
 {fvar: term,
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
36  | 
domT: typ,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
37  | 
ranT: typ,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
38  | 
h: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
39  | 
y: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
40  | 
x: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
41  | 
z: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
42  | 
a: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
43  | 
P: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
44  | 
D: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
45  | 
Pbool:term}  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
46  | 
|
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
47  | 
datatype rec_call_info = RCInfo of  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
48  | 
 {RIvs: (string * typ) list,  (* Call context: fixes and assumes *)
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
49  | 
CCas: thm list,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
50  | 
rcarg: term, (* The recursive argument *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
51  | 
llRI: thm,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
52  | 
h_assum: term}  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
54  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
55  | 
datatype clause_context = ClauseContext of  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
56  | 
 {ctxt : Proof.context,
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
57  | 
qs : term list,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
58  | 
gs : term list,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
59  | 
lhs: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
60  | 
rhs: term,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
61  | 
cqs: cterm list,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
62  | 
ags: thm list,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
63  | 
case_hyp : thm}  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
66  | 
fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
 | 
| 42361 | 67  | 
  ClauseContext { ctxt = Proof_Context.transfer thy ctxt,
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
68  | 
qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
70  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
71  | 
datatype clause_info = ClauseInfo of  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
72  | 
 {no: int,
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
73  | 
qglr : ((string * typ) list * term list * term * term),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
74  | 
cdata : clause_context,  | 
| 58816 | 75  | 
tree: Function_Context_Tree.ctx_tree,  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
76  | 
lGI: thm,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
77  | 
RCs: rec_call_info list}  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
78  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
79  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
80  | 
(* Theory dependencies. *)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
81  | 
val acc_induct_rule = @{thm accp_induct_rule}
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
82  | 
|
| 
55085
 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 
blanchet 
parents: 
53603 
diff
changeset
 | 
83  | 
val ex1_implies_ex = @{thm Fun_Def.fundef_ex1_existence}
 | 
| 
 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 
blanchet 
parents: 
53603 
diff
changeset
 | 
84  | 
val ex1_implies_un = @{thm Fun_Def.fundef_ex1_uniqueness}
 | 
| 
 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 
blanchet 
parents: 
53603 
diff
changeset
 | 
85  | 
val ex1_implies_iff = @{thm Fun_Def.fundef_ex1_iff}
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
86  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
87  | 
val acc_downward = @{thm accp_downward}
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
88  | 
val accI = @{thm accp.accI}
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
91  | 
fun find_calls tree =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
92  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
93  | 
fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
94  | 
([], (fixes, assumes, arg) :: xs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
95  | 
| add_Ri _ _ _ _ = raise Match  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
96  | 
in  | 
| 58816 | 97  | 
rev (Function_Context_Tree.traverse_tree add_Ri tree [])  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
98  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
101  | 
(** building proof obligations *)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
103  | 
fun mk_compat_proof_obligations domT ranT fvar f glrs =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
104  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
105  | 
fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
106  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
107  | 
val shift = incr_boundvars (length qs')  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
108  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
109  | 
Logic.mk_implies  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
110  | 
(HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
111  | 
HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
112  | 
|> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')  | 
| 
46217
 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 
wenzelm 
parents: 
42793 
diff
changeset
 | 
113  | 
|> fold_rev (fn (n,T) => fn b => Logic.all_const T $ Abs(n,T,b)) (qs @ qs')  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
114  | 
|> curry abstract_over fvar  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
115  | 
|> curry subst_bound f  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
116  | 
end  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
117  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
118  | 
map mk_impl (unordered_pairs glrs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
119  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
120  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
121  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
122  | 
fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
123  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
124  | 
    fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
125  | 
HOLogic.mk_Trueprop Pbool  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
126  | 
|> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
127  | 
|> fold_rev (curry Logic.mk_implies) gs  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
128  | 
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
129  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
130  | 
HOLogic.mk_Trueprop Pbool  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
131  | 
|> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
132  | 
    |> mk_forall_rename ("x", x)
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
133  | 
    |> mk_forall_rename ("P", Pbool)
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
134  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
135  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
136  | 
(** making a context with it's own local bindings **)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
138  | 
fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
139  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
140  | 
val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
141  | 
|>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
142  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
143  | 
fun inst t = subst_bounds (rev qs, t)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
144  | 
val gs = map inst pre_gs  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
145  | 
val lhs = inst pre_lhs  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
146  | 
val rhs = inst pre_rhs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
147  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
148  | 
val cqs = map (Thm.cterm_of ctxt') qs  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
149  | 
val ags = map (Thm.assume o Thm.cterm_of ctxt') gs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
150  | 
|
| 59618 | 151  | 
val case_hyp =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
152  | 
Thm.assume (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (mk_eq (x, lhs))))  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
153  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
154  | 
    ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
155  | 
cqs = cqs, ags = ags, case_hyp = case_hyp }  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
156  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
158  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
159  | 
(* lowlevel term function. FIXME: remove *)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
160  | 
fun abstract_over_list vs body =  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
161  | 
let  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
162  | 
fun abs lev v tm =  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
163  | 
if v aconv tm then Bound lev  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
164  | 
else  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
165  | 
(case tm of  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
166  | 
Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
167  | 
| t $ u => abs lev v t $ abs lev v u  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
168  | 
| t => t)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
169  | 
in  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
170  | 
fold_index (fn (i, v) => fn t => abs i v t) vs body  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
171  | 
end  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
172  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
174  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
175  | 
fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
176  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
177  | 
    val Globals {h, ...} = globals
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
178  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
179  | 
    val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
180  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
181  | 
(* Instantiate the GIntro thm with "f" and import into the clause context. *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
182  | 
val lGI = GIntro_thm  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
183  | 
|> Thm.forall_elim (Thm.cterm_of ctxt f)  | 
| 36945 | 184  | 
|> fold Thm.forall_elim cqs  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
185  | 
|> fold Thm.elim_implies ags  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
186  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
187  | 
fun mk_call_info (rcfix, rcassm, rcarg) RI =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
188  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
189  | 
val llRI = RI  | 
| 36945 | 190  | 
|> fold Thm.forall_elim cqs  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
191  | 
|> fold (Thm.forall_elim o Thm.cterm_of ctxt o Free) rcfix  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
192  | 
|> fold Thm.elim_implies ags  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
193  | 
|> fold Thm.elim_implies rcassm  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
194  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
195  | 
val h_assum =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
196  | 
HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))  | 
| 59582 | 197  | 
|> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
198  | 
|> fold_rev (Logic.all o Free) rcfix  | 
| 42361 | 199  | 
|> Pattern.rewrite_term (Proof_Context.theory_of ctxt) [(f, h)] []  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
200  | 
|> abstract_over_list (rev qs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
201  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
202  | 
        RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
203  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
204  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
205  | 
val RC_infos = map2 mk_call_info RCs RIntro_thms  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
206  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
207  | 
    ClauseInfo {no=no, cdata=cdata, qglr=qglr, lGI=lGI, RCs=RC_infos,
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
208  | 
tree=tree}  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
209  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
210  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
212  | 
fun store_compat_thms 0 thms = []  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
213  | 
| store_compat_thms n thms =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
214  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
215  | 
val (thms1, thms2) = chop n thms  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
216  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
217  | 
(thms1 :: store_compat_thms (n - 1) thms2)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
218  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
220  | 
(* expects i <= j *)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
221  | 
fun lookup_compat_thm i j cts =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
222  | 
nth (nth cts (i - 1)) (j - i)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
223  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
224  | 
(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
225  | 
(* if j < i, then turn around *)  | 
| 59627 | 226  | 
fun get_compat_thm ctxt cts i j ctxi ctxj =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
227  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
228  | 
    val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
229  | 
    val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
230  | 
|
| 59627 | 231  | 
val lhsi_eq_lhsj = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
232  | 
in if j < i then  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
233  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
234  | 
val compat = lookup_compat_thm j i cts  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
235  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
236  | 
compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)  | 
| 36945 | 237  | 
|> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
238  | 
|> fold Thm.elim_implies agsj  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
239  | 
|> fold Thm.elim_implies agsi  | 
| 36945 | 240  | 
|> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
241  | 
end  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
242  | 
else  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
243  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
244  | 
val compat = lookup_compat_thm i j cts  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
245  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
246  | 
compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)  | 
| 36945 | 247  | 
|> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
248  | 
|> fold Thm.elim_implies agsi  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
249  | 
|> fold Thm.elim_implies agsj  | 
| 36945 | 250  | 
|> Thm.elim_implies (Thm.assume lhsi_eq_lhsj)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
251  | 
|> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
252  | 
end  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
253  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
254  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
255  | 
(* Generates the replacement lemma in fully quantified form. *)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
256  | 
fun mk_replacement_lemma ctxt h ih_elim clause =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
257  | 
let  | 
| 61340 | 258  | 
    val ClauseInfo {cdata=ClauseContext {qs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
259  | 
local open Conv in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
260  | 
val ih_conv = arg1_conv o arg_conv o arg_conv  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
261  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
262  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
263  | 
val ih_elim_case =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
264  | 
Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
265  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
266  | 
    val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
267  | 
    val h_assums = map (fn RCInfo {h_assum, ...} =>
 | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
268  | 
Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qs, h_assum)))) RCs  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
269  | 
|
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
270  | 
val (eql, _) =  | 
| 58816 | 271  | 
Function_Context_Tree.rewrite_by_tree ctxt h ih_elim_case (Ris ~~ h_assums) tree  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
272  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67149 
diff
changeset
 | 
273  | 
val replace_lemma = HOLogic.mk_obj_eq eql  | 
| 59582 | 274  | 
|> Thm.implies_intr (Thm.cprop_of case_hyp)  | 
275  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) h_assums  | 
|
276  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) ags  | 
|
| 36945 | 277  | 
|> fold_rev Thm.forall_intr cqs  | 
| 70494 | 278  | 
|> Thm.close_derivation \<^here>  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
279  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
280  | 
replace_lemma  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
281  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
282  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
283  | 
|
| 59627 | 284  | 
fun mk_uniqueness_clause ctxt globals compat_store clausei clausej RLj =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
285  | 
let  | 
| 59627 | 286  | 
val thy = Proof_Context.theory_of ctxt  | 
287  | 
||
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
288  | 
    val Globals {h, y, x, fvar, ...} = globals
 | 
| 59582 | 289  | 
    val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} =
 | 
290  | 
clausei  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
291  | 
    val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
292  | 
|
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
293  | 
    val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} =
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
294  | 
mk_clause_context x ctxti cdescj  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
295  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
296  | 
val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'  | 
| 59627 | 297  | 
val compat = get_compat_thm ctxt compat_store i j cctxi cctxj  | 
| 59582 | 298  | 
val Ghsj' =  | 
299  | 
      map (fn RCInfo {h_assum, ...} =>
 | 
|
| 59627 | 300  | 
Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qsj', h_assum)))) RCsj  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
301  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
302  | 
val RLj_import = RLj  | 
| 36945 | 303  | 
|> fold Thm.forall_elim cqsj'  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
304  | 
|> fold Thm.elim_implies agsj'  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
305  | 
|> fold Thm.elim_implies Ghsj'  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
306  | 
|
| 59627 | 307  | 
val y_eq_rhsj'h = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))  | 
308  | 
val lhsi_eq_lhsj' = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
309  | 
(* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
310  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
311  | 
(trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)  | 
| 36945 | 312  | 
|> Thm.implies_elim RLj_import  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
313  | 
(* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
314  | 
|> (fn it => trans OF [it, compat])  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
315  | 
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
316  | 
|> (fn it => trans OF [y_eq_rhsj'h, it])  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
317  | 
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)  | 
| 59582 | 318  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) Ghsj'  | 
319  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) agsj'  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
320  | 
(* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)  | 
| 59582 | 321  | 
|> Thm.implies_intr (Thm.cprop_of y_eq_rhsj'h)  | 
322  | 
|> Thm.implies_intr (Thm.cprop_of lhsi_eq_lhsj')  | 
|
| 59627 | 323  | 
|> fold_rev Thm.forall_intr (Thm.cterm_of ctxt h :: cqsj')  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
324  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
325  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
326  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
327  | 
|
| 59627 | 328  | 
fun mk_uniqueness_case ctxt  | 
329  | 
globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
330  | 
let  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
331  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
332  | 
    val Globals {x, y, ranT, fvar, ...} = globals
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
333  | 
    val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
334  | 
val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
335  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
336  | 
val ih_intro_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ih_intro  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
337  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
338  | 
    fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
 | 
| 59582 | 339  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
340  | 
|> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
341  | 
|
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
342  | 
val existence = fold (curry op COMP o prep_RC) RCs lGI  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
343  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
344  | 
val P = Thm.cterm_of ctxt (mk_eq (y, rhsC))  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
345  | 
val G_lhs_y = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (G $ lhs $ y)))  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
346  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
347  | 
val unique_clauses =  | 
| 59627 | 348  | 
map2 (mk_uniqueness_clause ctxt globals compat_store clausei) clauses rep_lemmas  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
349  | 
|
| 
36270
 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 
krauss 
parents: 
34232 
diff
changeset
 | 
350  | 
fun elim_implies_eta A AB =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58816 
diff
changeset
 | 
351  | 
      Thm.bicompose (SOME ctxt) {flatten = false, match = true, incremented = false}
 | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58816 
diff
changeset
 | 
352  | 
(false, A, 0) 1 AB  | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
51717 
diff
changeset
 | 
353  | 
|> Seq.list_of |> the_single  | 
| 
36270
 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 
krauss 
parents: 
34232 
diff
changeset
 | 
354  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
355  | 
val uniqueness = G_cases  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
356  | 
|> Thm.forall_elim (Thm.cterm_of ctxt lhs)  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
357  | 
|> Thm.forall_elim (Thm.cterm_of ctxt y)  | 
| 36945 | 358  | 
|> Thm.forall_elim P  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
359  | 
|> Thm.elim_implies G_lhs_y  | 
| 
36270
 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 
krauss 
parents: 
34232 
diff
changeset
 | 
360  | 
|> fold elim_implies_eta unique_clauses  | 
| 59582 | 361  | 
|> Thm.implies_intr (Thm.cprop_of G_lhs_y)  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
362  | 
|> Thm.forall_intr (Thm.cterm_of ctxt y)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
363  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
364  | 
val P2 = Thm.cterm_of ctxt (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
365  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
366  | 
val exactly_one =  | 
| 59582 | 367  | 
      @{thm ex1I}
 | 
| 60801 | 368  | 
|> Thm.instantiate'  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
369  | 
[SOME (Thm.ctyp_of ctxt ranT)]  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
370  | 
[SOME P2, SOME (Thm.cterm_of ctxt rhsC)]  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
371  | 
|> curry (op COMP) existence  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
372  | 
|> curry (op COMP) uniqueness  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
373  | 
|> simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp RS sym])  | 
| 59582 | 374  | 
|> Thm.implies_intr (Thm.cprop_of case_hyp)  | 
375  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) ags  | 
|
| 36945 | 376  | 
|> fold_rev Thm.forall_intr cqs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
377  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
378  | 
val function_value =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
379  | 
existence  | 
| 36945 | 380  | 
|> Thm.implies_intr ihyp  | 
| 59582 | 381  | 
|> Thm.implies_intr (Thm.cprop_of case_hyp)  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
382  | 
|> Thm.forall_intr (Thm.cterm_of ctxt x)  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
383  | 
|> Thm.forall_elim (Thm.cterm_of ctxt lhs)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
384  | 
|> curry (op RS) refl  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
385  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
386  | 
(exactly_one, function_value)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
387  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
388  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
389  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
390  | 
fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
391  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
392  | 
    val Globals {h, domT, ranT, x, ...} = globals
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
393  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
394  | 
(* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)  | 
| 
46217
 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 
wenzelm 
parents: 
42793 
diff
changeset
 | 
395  | 
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
396  | 
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),  | 
| 67149 | 397  | 
HOLogic.mk_Trueprop (Const (\<^const_name>\<open>Ex1\<close>, (ranT --> boolT) --> boolT) $  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
398  | 
          Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
 | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
399  | 
|> Thm.cterm_of ctxt  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
400  | 
|
| 36945 | 401  | 
val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
402  | 
val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
403  | 
val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)  | 
| 60801 | 404  | 
|> Thm.instantiate' [] [NONE, SOME (Thm.cterm_of ctxt h)]  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
405  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
406  | 
val _ = trace_msg (K "Proving Replacement lemmas...")  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
407  | 
val repLemmas = map (mk_replacement_lemma ctxt h ih_elim) clauses  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
408  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
409  | 
val _ = trace_msg (K "Proving cases for unique existence...")  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
410  | 
val (ex1s, values) =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
411  | 
split_list  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
412  | 
(map  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
413  | 
(mk_uniqueness_case ctxt globals G f ihyp ih_intro G_elim compat_store clauses repLemmas)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
414  | 
clauses)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
415  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
416  | 
val _ = trace_msg (K "Proving: Graph is a function")  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
417  | 
val graph_is_function = complete  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
418  | 
|> Thm.forall_elim_vars 0  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
419  | 
|> fold (curry op COMP) ex1s  | 
| 36945 | 420  | 
|> Thm.implies_intr (ihyp)  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
421  | 
|> Thm.implies_intr (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ x)))  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
422  | 
|> Thm.forall_intr (Thm.cterm_of ctxt x)  | 
| 52467 | 423  | 
|> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)  | 
| 59582 | 424  | 
|> (fn it =>  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
425  | 
fold (Thm.forall_intr o Thm.cterm_of ctxt o Var)  | 
| 59618 | 426  | 
(Term.add_vars (Thm.prop_of it) []) it)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
427  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
428  | 
val goalstate = Conjunction.intr graph_is_function complete  | 
| 70494 | 429  | 
|> Thm.close_derivation \<^here>  | 
| 52456 | 430  | 
|> Goal.protect 0  | 
| 59582 | 431  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) compat  | 
432  | 
|> Thm.implies_intr (Thm.cprop_of complete)  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
433  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
434  | 
(goalstate, values)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
435  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
436  | 
|
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
437  | 
(* wrapper -- restores quantifiers in rule specifications *)  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
438  | 
fun inductive_def (binding as ((R, T), _)) intrs lthy =  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
439  | 
let  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
440  | 
    val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) =
 | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
441  | 
lthy  | 
| 
59880
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
442  | 
|> Proof_Context.concealed  | 
| 69709 | 443  | 
|> Inductive.add_inductive  | 
| 33350 | 444  | 
          {quiet_mode = true,
 | 
445  | 
verbose = ! trace,  | 
|
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
446  | 
alt_name = Binding.empty,  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
447  | 
coind = false,  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
448  | 
no_elim = false,  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
449  | 
no_ind = false,  | 
| 
49170
 
03bee3a6a1b7
discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
 
wenzelm 
parents: 
46217 
diff
changeset
 | 
450  | 
skip_mono = true}  | 
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
451  | 
[binding] (* relation *)  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
452  | 
[] (* no parameters *)  | 
| 63352 | 453  | 
(map (fn t => (Binding.empty_atts, t)) intrs) (* intro rules *)  | 
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
454  | 
[] (* no special monos *)  | 
| 
59880
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
455  | 
||> Proof_Context.restore_naming lthy  | 
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
456  | 
|
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
457  | 
fun requantify orig_intro thm =  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
458  | 
let  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
459  | 
val (qs, t) = dest_all_all orig_intro  | 
| 42483 | 460  | 
val frees = Variable.add_frees lthy t [] |> remove (op =) (Binding.name_of R, T)  | 
| 59582 | 461  | 
val vars = Term.add_vars (Thm.prop_of thm) []  | 
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
462  | 
val varmap = AList.lookup (op =) (frees ~~ map fst vars)  | 
| 42483 | 463  | 
          #> the_default ("", 0)
 | 
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
464  | 
in  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
465  | 
fold_rev (fn Free (n, T) =>  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
466  | 
forall_intr_rename (n, Thm.cterm_of lthy (Var (varmap (n, T), T)))) qs thm  | 
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
467  | 
end  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
468  | 
in  | 
| 74245 | 469  | 
((Rdef, map2 requantify intrs intrs_gen, Thm.forall_intr_vars elim_gen, induct), lthy)  | 
| 
33348
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
470  | 
end  | 
| 
 
bb65583ab70d
absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
 
krauss 
parents: 
33278 
diff
changeset
 | 
471  | 
|
| 63004 | 472  | 
fun define_graph (G_binding, G_type) fvar clauses RCss lthy =  | 
| 33349 | 473  | 
let  | 
| 63004 | 474  | 
val G = Free (Binding.name_of G_binding, G_type)  | 
| 33349 | 475  | 
    fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
 | 
476  | 
let  | 
|
477  | 
fun mk_h_assm (rcfix, rcassm, rcarg) =  | 
|
| 63001 | 478  | 
HOLogic.mk_Trueprop (G $ rcarg $ (fvar $ rcarg))  | 
| 59582 | 479  | 
|> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm  | 
| 33349 | 480  | 
|> fold_rev (Logic.all o Free) rcfix  | 
481  | 
in  | 
|
| 63001 | 482  | 
HOLogic.mk_Trueprop (G $ lhs $ rhs)  | 
| 33349 | 483  | 
|> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs  | 
484  | 
|> fold_rev (curry Logic.mk_implies) gs  | 
|
485  | 
|> fold_rev Logic.all (fvar :: qs)  | 
|
486  | 
end  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
487  | 
|
| 33349 | 488  | 
val G_intros = map2 mk_GIntro clauses RCss  | 
| 63004 | 489  | 
in inductive_def ((G_binding, G_type), NoSyn) G_intros lthy end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
490  | 
|
| 63004 | 491  | 
fun define_function defname (fname, mixfix) domT ranT G default lthy =  | 
| 33349 | 492  | 
let  | 
| 63004 | 493  | 
val f_def_binding =  | 
494  | 
Thm.make_def_binding (Config.get lthy function_internals)  | 
|
| 63005 | 495  | 
(derived_name_suffix defname "_sumC")  | 
| 33349 | 496  | 
val f_def =  | 
| 67149 | 497  | 
      Abs ("x", domT, Const (\<^const_name>\<open>Fun_Def.THE_default\<close>, ranT --> (ranT --> boolT) --> ranT)
 | 
| 33349 | 498  | 
        $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
 | 
499  | 
|> Syntax.check_term lthy  | 
|
500  | 
in  | 
|
| 63005 | 501  | 
lthy |> Local_Theory.define  | 
| 63011 | 502  | 
((Binding.map_name (suffix "C") fname, mixfix), ((f_def_binding, []), f_def))  | 
| 33349 | 503  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
504  | 
|
| 63004 | 505  | 
fun define_recursion_relation (R_binding, R_type) qglrs clauses RCss lthy =  | 
| 33349 | 506  | 
let  | 
| 63004 | 507  | 
val R = Free (Binding.name_of R_binding, R_type)  | 
| 33349 | 508  | 
    fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
 | 
| 63004 | 509  | 
HOLogic.mk_Trueprop (R $ rcarg $ lhs)  | 
| 59582 | 510  | 
|> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm  | 
| 33349 | 511  | 
|> fold_rev (curry Logic.mk_implies) gs  | 
512  | 
|> fold_rev (Logic.all o Free) rcfix  | 
|
513  | 
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)  | 
|
514  | 
(* "!!qs xs. CS ==> G => (r, lhs) : R" *)  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
515  | 
|
| 33349 | 516  | 
val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
517  | 
|
| 33349 | 518  | 
val ((R, RIntro_thms, R_elim, _), lthy) =  | 
| 63004 | 519  | 
inductive_def ((R_binding, R_type), NoSyn) (flat R_intross) lthy  | 
| 33349 | 520  | 
in  | 
521  | 
((R, Library.unflat R_intross RIntro_thms, R_elim), lthy)  | 
|
522  | 
end  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
523  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
524  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
525  | 
fun fix_globals domT ranT fvar ctxt =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
526  | 
let  | 
| 59618 | 527  | 
val ([h, y, x, z, a, D, P, Pbool], ctxt') = Variable.variant_fixes  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
528  | 
["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
529  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
530  | 
    (Globals {h = Free (h, domT --> ranT),
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
531  | 
y = Free (y, ranT),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
532  | 
x = Free (x, domT),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
533  | 
z = Free (z, domT),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
534  | 
a = Free (a, domT),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
535  | 
D = Free (D, domT --> boolT),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
536  | 
P = Free (P, domT --> boolT),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
537  | 
Pbool = Free (Pbool, boolT),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
538  | 
fvar = fvar,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
539  | 
domT = domT,  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
540  | 
ranT = ranT},  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
541  | 
ctxt')  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
542  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
543  | 
|
| 59627 | 544  | 
fun inst_RC ctxt fvar f (rcfix, rcassm, rcarg) =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
545  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
546  | 
fun inst_term t = subst_bound(f, abstract_over (fvar, t))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
547  | 
in  | 
| 59627 | 548  | 
(rcfix, map (Thm.assume o Thm.cterm_of ctxt o inst_term o Thm.prop_of) rcassm, inst_term rcarg)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
549  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
550  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
551  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
552  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
553  | 
(**********************************************************  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
554  | 
* PROVING THE RULES  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
555  | 
**********************************************************)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
556  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
557  | 
fun mk_psimps ctxt globals R clauses valthms f_iff graph_is_function =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
558  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
559  | 
    val Globals {domT, z, ...} = globals
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
560  | 
|
| 59582 | 561  | 
fun mk_psimp  | 
562  | 
      (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
 | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
563  | 
let  | 
| 59618 | 564  | 
val lhs_acc =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
565  | 
Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)  | 
| 59618 | 566  | 
val z_smaller =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
567  | 
Thm.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
568  | 
in  | 
| 36945 | 569  | 
((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
570  | 
|> (fn it => it COMP graph_is_function)  | 
| 36945 | 571  | 
|> Thm.implies_intr z_smaller  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
572  | 
|> Thm.forall_intr (Thm.cterm_of ctxt z)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
573  | 
|> (fn it => it COMP valthm)  | 
| 36945 | 574  | 
|> Thm.implies_intr lhs_acc  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
575  | 
|> asm_simplify (put_simpset HOL_basic_ss ctxt addsimps [f_iff])  | 
| 59582 | 576  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) ags  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
577  | 
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
578  | 
end  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
579  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
580  | 
map2 mk_psimp clauses valthms  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
581  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
582  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
583  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
584  | 
(** Induction rule **)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
585  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
586  | 
|
| 
34065
 
6f8f9835e219
moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
 
haftmann 
parents: 
33855 
diff
changeset
 | 
587  | 
val acc_subset_induct = @{thm predicate1I} RS @{thm accp_subset_induct}
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
588  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
589  | 
|
| 59627 | 590  | 
fun mk_partial_induct_rule ctxt globals R complete_thm clauses =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
591  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
592  | 
    val Globals {domT, x, z, a, P, D, ...} = globals
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
593  | 
val acc_R = mk_acc domT R  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
594  | 
|
| 59627 | 595  | 
val x_D = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ x)))  | 
596  | 
val a_D = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ a))  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
597  | 
|
| 59627 | 598  | 
val D_subset =  | 
599  | 
Thm.cterm_of ctxt (Logic.all x  | 
|
600  | 
(Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
601  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
602  | 
val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
603  | 
Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
604  | 
Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
605  | 
HOLogic.mk_Trueprop (D $ z)))))  | 
| 59627 | 606  | 
|> Thm.cterm_of ctxt  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
607  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
608  | 
(* Inductive Hypothesis: !!z. (z,x):R ==> P z *)  | 
| 
46217
 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 
wenzelm 
parents: 
42793 
diff
changeset
 | 
609  | 
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
610  | 
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
611  | 
HOLogic.mk_Trueprop (P $ Bound 0)))  | 
| 59627 | 612  | 
|> Thm.cterm_of ctxt  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
613  | 
|
| 36945 | 614  | 
val aihyp = Thm.assume ihyp  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
615  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
616  | 
fun prove_case clause =  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
617  | 
let  | 
| 59627 | 618  | 
        val ClauseInfo {cdata = ClauseContext {ctxt = ctxt1, qs, cqs, ags, gs, lhs, case_hyp, ...},
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
619  | 
RCs, qglr = (oqs, _, _, _), ...} = clause  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
620  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
621  | 
val case_hyp_conv = K (case_hyp RS eq_reflection)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
622  | 
local open Conv in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
623  | 
val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
624  | 
val sih =  | 
| 74534 | 625  | 
fconv_rule  | 
626  | 
(binder_conv (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt1) aihyp  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
627  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
628  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
629  | 
        fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
 | 
| 59627 | 630  | 
|> Thm.forall_elim (Thm.cterm_of ctxt rcarg)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
631  | 
|> Thm.elim_implies llRI  | 
| 59582 | 632  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas  | 
| 59627 | 633  | 
|> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
634  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
635  | 
val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
636  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
637  | 
val step = HOLogic.mk_Trueprop (P $ lhs)  | 
| 59582 | 638  | 
|> fold_rev (curry Logic.mk_implies o Thm.prop_of) P_recs  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
639  | 
|> fold_rev (curry Logic.mk_implies) gs  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
640  | 
|> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
641  | 
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)  | 
| 59627 | 642  | 
|> Thm.cterm_of ctxt  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
643  | 
|
| 36945 | 644  | 
val P_lhs = Thm.assume step  | 
645  | 
|> fold Thm.forall_elim cqs  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
646  | 
|> Thm.elim_implies lhs_D  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
647  | 
|> fold Thm.elim_implies ags  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
648  | 
|> fold Thm.elim_implies P_recs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
649  | 
|
| 59627 | 650  | 
val res = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (P $ x))  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
651  | 
|> Conv.arg_conv (Conv.arg_conv case_hyp_conv)  | 
| 36945 | 652  | 
|> Thm.symmetric (* P lhs == P x *)  | 
653  | 
|> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)  | 
|
| 59582 | 654  | 
|> Thm.implies_intr (Thm.cprop_of case_hyp)  | 
655  | 
|> fold_rev (Thm.implies_intr o Thm.cprop_of) ags  | 
|
| 36945 | 656  | 
|> fold_rev Thm.forall_intr cqs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
657  | 
in  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
658  | 
(res, step)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
659  | 
end  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
660  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
661  | 
val (cases, steps) = split_list (map prove_case clauses)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
662  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
663  | 
val istep = complete_thm  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
664  | 
|> Thm.forall_elim_vars 0  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
665  | 
|> fold (curry op COMP) cases (* P x *)  | 
| 36945 | 666  | 
|> Thm.implies_intr ihyp  | 
| 59582 | 667  | 
|> Thm.implies_intr (Thm.cprop_of x_D)  | 
| 59627 | 668  | 
|> Thm.forall_intr (Thm.cterm_of ctxt x)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
669  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
670  | 
val subset_induct_rule =  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
671  | 
acc_subset_induct  | 
| 36945 | 672  | 
|> (curry op COMP) (Thm.assume D_subset)  | 
673  | 
|> (curry op COMP) (Thm.assume D_dcl)  | 
|
674  | 
|> (curry op COMP) (Thm.assume a_D)  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
675  | 
|> (curry op COMP) istep  | 
| 36945 | 676  | 
|> fold_rev Thm.implies_intr steps  | 
677  | 
|> Thm.implies_intr a_D  | 
|
678  | 
|> Thm.implies_intr D_dcl  | 
|
679  | 
|> Thm.implies_intr D_subset  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
680  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
681  | 
val simple_induct_rule =  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
682  | 
subset_induct_rule  | 
| 59627 | 683  | 
|> Thm.forall_intr (Thm.cterm_of ctxt D)  | 
684  | 
|> Thm.forall_elim (Thm.cterm_of ctxt acc_R)  | 
|
| 60752 | 685  | 
|> assume_tac ctxt 1 |> Seq.hd  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
686  | 
|> (curry op COMP) (acc_downward  | 
| 60801 | 687  | 
|> (Thm.instantiate' [SOME (Thm.ctyp_of ctxt domT)] (map (SOME o Thm.cterm_of ctxt) [R, x, z]))  | 
| 59627 | 688  | 
|> Thm.forall_intr (Thm.cterm_of ctxt z)  | 
689  | 
|> Thm.forall_intr (Thm.cterm_of ctxt x))  | 
|
690  | 
|> Thm.forall_intr (Thm.cterm_of ctxt a)  | 
|
691  | 
|> Thm.forall_intr (Thm.cterm_of ctxt P)  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
692  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
693  | 
simple_induct_rule  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
694  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
695  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
696  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
697  | 
(* FIXME: broken by design *)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
698  | 
fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
699  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
700  | 
    val ClauseInfo {cdata = ClauseContext {gs, lhs, cqs, ...},
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
701  | 
qglr = (oqs, _, _, _), ...} = clause  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
702  | 
val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
703  | 
|> fold_rev (curry Logic.mk_implies) gs  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
704  | 
|> Thm.cterm_of ctxt  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
705  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
706  | 
Goal.init goal  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
707  | 
|> (SINGLE (resolve_tac ctxt [accI] 1)) |> the  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58963 
diff
changeset
 | 
708  | 
|> (SINGLE (eresolve_tac ctxt [Thm.forall_elim_vars 0 R_cases] 1)) |> the  | 
| 42793 | 709  | 
|> (SINGLE (auto_tac ctxt)) |> the  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
710  | 
|> Goal.conclude  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
711  | 
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
712  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
713  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
714  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
715  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
716  | 
(** Termination rule **)  | 
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
717  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
718  | 
val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule}
 | 
| 
55085
 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 
blanchet 
parents: 
53603 
diff
changeset
 | 
719  | 
val wf_in_rel = @{thm Fun_Def.wf_in_rel}
 | 
| 
 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 
blanchet 
parents: 
53603 
diff
changeset
 | 
720  | 
val in_rel_def = @{thm Fun_Def.in_rel_def}
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
721  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
722  | 
fun mk_nest_term_case ctxt globals R' ihyp clause =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
723  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
724  | 
    val Globals {z, ...} = globals
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
725  | 
    val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree,
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
726  | 
qglr=(oqs, _, _, _), ...} = clause  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
727  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
728  | 
val ih_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ihyp  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
729  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
730  | 
fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
731  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
732  | 
val used = (u @ sub)  | 
| 59618 | 733  | 
|> map (fn (ctx, thm) => Function_Context_Tree.export_thm ctxt ctx thm)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
734  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
735  | 
val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)  | 
| 59582 | 736  | 
|> fold_rev (curry Logic.mk_implies o Thm.prop_of) used (* additional hyps *)  | 
| 58816 | 737  | 
|> Function_Context_Tree.export_term (fixes, assumes)  | 
| 59582 | 738  | 
|> fold_rev (curry Logic.mk_implies o Thm.prop_of) ags  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
739  | 
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
740  | 
|> Thm.cterm_of ctxt  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
741  | 
|
| 36945 | 742  | 
val thm = Thm.assume hyp  | 
743  | 
|> fold Thm.forall_elim cqs  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
744  | 
|> fold Thm.elim_implies ags  | 
| 59618 | 745  | 
|> Function_Context_Tree.import_thm ctxt (fixes, assumes)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
746  | 
|> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
747  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
748  | 
val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
749  | 
|> Thm.cterm_of ctxt |> Thm.assume  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
750  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
751  | 
val acc = thm COMP ih_case  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
752  | 
val z_acc_local = acc  | 
| 36945 | 753  | 
|> Conv.fconv_rule  | 
754  | 
(Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection)))))  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
755  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
756  | 
val ethm = z_acc_local  | 
| 59618 | 757  | 
|> Function_Context_Tree.export_thm ctxt (fixes,  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
758  | 
z_eq_arg :: case_hyp :: ags @ assumes)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
759  | 
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
760  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
761  | 
val sub' = sub @ [(([],[]), acc)]  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
762  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
763  | 
(sub', (hyp :: hyps, ethm :: thms))  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
764  | 
end  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
765  | 
| step _ _ _ _ = raise Match  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
766  | 
in  | 
| 58816 | 767  | 
Function_Context_Tree.traverse_tree step tree  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
768  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
769  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
770  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
771  | 
fun mk_nest_term_rule ctxt globals R R_cases clauses =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
772  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
773  | 
    val Globals { domT, x, z, ... } = globals
 | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
774  | 
val acc_R = mk_acc domT R  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
775  | 
|
| 61340 | 776  | 
val ([Rn], ctxt') = Variable.variant_fixes ["R"] ctxt  | 
777  | 
val R' = Free (Rn, fastype_of R)  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
778  | 
|
| 61340 | 779  | 
val Rrel = Free (Rn, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))  | 
| 67149 | 780  | 
val inrel_R = Const (\<^const_name>\<open>Fun_Def.in_rel\<close>,  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
781  | 
HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
782  | 
|
| 67149 | 783  | 
val wfR' = HOLogic.mk_Trueprop (Const (\<^const_name>\<open>Wellfounded.wfP\<close>,  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
784  | 
(domT --> domT --> boolT) --> boolT) $ R')  | 
| 61340 | 785  | 
|> Thm.cterm_of ctxt' (* "wf R'" *)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
786  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
787  | 
(* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)  | 
| 
46217
 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 
wenzelm 
parents: 
42793 
diff
changeset
 | 
788  | 
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
789  | 
Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
790  | 
HOLogic.mk_Trueprop (acc_R $ Bound 0)))  | 
| 61340 | 791  | 
|> Thm.cterm_of ctxt'  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
792  | 
|
| 36945 | 793  | 
val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
794  | 
|
| 61340 | 795  | 
val R_z_x = Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (R $ z $ x))  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
796  | 
|
| 61340 | 797  | 
val (hyps, cases) = fold (mk_nest_term_case ctxt' globals R' ihyp_a) clauses ([], [])  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
798  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
799  | 
R_cases  | 
| 61340 | 800  | 
|> Thm.forall_elim (Thm.cterm_of ctxt' z)  | 
801  | 
|> Thm.forall_elim (Thm.cterm_of ctxt' x)  | 
|
802  | 
|> Thm.forall_elim (Thm.cterm_of ctxt' (acc_R $ z))  | 
|
| 36945 | 803  | 
|> curry op COMP (Thm.assume R_z_x)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
804  | 
|> fold_rev (curry op COMP) cases  | 
| 36945 | 805  | 
|> Thm.implies_intr R_z_x  | 
| 61340 | 806  | 
|> Thm.forall_intr (Thm.cterm_of ctxt' z)  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
807  | 
|> (fn it => it COMP accI)  | 
| 36945 | 808  | 
|> Thm.implies_intr ihyp  | 
| 61340 | 809  | 
|> Thm.forall_intr (Thm.cterm_of ctxt' x)  | 
| 52467 | 810  | 
|> (fn it => Drule.compose (it, 2, wf_induct_rule))  | 
| 36945 | 811  | 
|> curry op RS (Thm.assume wfR')  | 
| 74245 | 812  | 
|> Thm.forall_intr_vars  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
813  | 
|> (fn it => it COMP allI)  | 
| 36945 | 814  | 
|> fold Thm.implies_intr hyps  | 
815  | 
|> Thm.implies_intr wfR'  | 
|
| 61340 | 816  | 
|> Thm.forall_intr (Thm.cterm_of ctxt' R')  | 
817  | 
|> Thm.forall_elim (Thm.cterm_of ctxt' inrel_R)  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
818  | 
|> curry op RS wf_in_rel  | 
| 61340 | 819  | 
|> full_simplify (put_simpset HOL_basic_ss ctxt' addsimps [in_rel_def])  | 
820  | 
    |> Thm.forall_intr_name ("R", Thm.cterm_of ctxt' Rrel)
 | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
821  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
822  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
823  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
824  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
825  | 
fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
826  | 
let  | 
| 
41846
 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 
krauss 
parents: 
41417 
diff
changeset
 | 
827  | 
    val FunctionConfig {domintros, default=default_opt, ...} = config
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
828  | 
|
| 63011 | 829  | 
val default_str = the_default "%x. HOL.undefined" (* FIXME proper term!? *) default_opt  | 
830  | 
val fvar = (Binding.name_of fname, fT)  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
831  | 
val domT = domain_type fT  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
832  | 
val ranT = range_type fT  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
833  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
834  | 
val default = Syntax.parse_term lthy default_str  | 
| 39288 | 835  | 
|> Type.constraint fT |> Syntax.check_term lthy  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
836  | 
|
| 63011 | 837  | 
val (globals, ctxt') = fix_globals domT ranT (Free fvar) lthy  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
838  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
839  | 
    val Globals { x, h, ... } = globals
 | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
840  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
841  | 
val clauses = map (mk_clause_context x ctxt') abstract_qglrs  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
842  | 
|
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
843  | 
val n = length abstract_qglrs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
844  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
845  | 
    fun build_tree (ClauseContext { ctxt, rhs, ...}) =
 | 
| 
65418
 
c821f1f3d92d
more general signature; works for all terms, not just frees
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
63352 
diff
changeset
 | 
846  | 
Function_Context_Tree.mk_tree (Free fvar) h ctxt rhs  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
847  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
848  | 
val trees = map build_tree clauses  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
849  | 
val RCss = map find_calls trees  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
850  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
851  | 
val ((G, GIntro_thms, G_elim, G_induct), lthy) =  | 
| 63001 | 852  | 
PROFILE "def_graph"  | 
| 63004 | 853  | 
(define_graph  | 
| 63011 | 854  | 
(derived_name_suffix defname "_graph", domT --> ranT --> boolT) (Free fvar) clauses RCss)  | 
855  | 
lthy  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
856  | 
|
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
857  | 
val ((f, (_, f_defthm)), lthy) =  | 
| 63004 | 858  | 
PROFILE "def_fun" (define_function defname (fname, mixfix) domT ranT G default) lthy  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
859  | 
|
| 63011 | 860  | 
val RCss = map (map (inst_RC lthy (Free fvar) f)) RCss  | 
861  | 
val trees = map (Function_Context_Tree.inst_tree lthy (Free fvar) f) trees  | 
|
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
862  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
863  | 
val ((R, RIntro_thmss, R_elim), lthy) =  | 
| 63001 | 864  | 
PROFILE "def_rel"  | 
| 63005 | 865  | 
(define_recursion_relation (derived_name_suffix defname "_rel", domT --> domT --> boolT)  | 
| 63001 | 866  | 
abstract_qglrs clauses RCss) lthy  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
867  | 
|
| 52384 | 868  | 
val dom = mk_acc domT R  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
869  | 
val (_, lthy) =  | 
| 63004 | 870  | 
Local_Theory.abbrev Syntax.mode_default  | 
| 63011 | 871  | 
((derived_name_suffix defname "_dom", NoSyn), dom) lthy  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
872  | 
|
| 42361 | 873  | 
val newthy = Proof_Context.theory_of lthy  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
874  | 
val clauses = map (transfer_clause_ctx newthy) clauses  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
875  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
876  | 
val xclauses = PROFILE "xclauses"  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
55990 
diff
changeset
 | 
877  | 
      (@{map 7} (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees
 | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
878  | 
RCss GIntro_thms) RIntro_thmss  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
879  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
880  | 
val complete =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
881  | 
mk_completeness globals clauses abstract_qglrs |> Thm.cterm_of lthy |> Thm.assume  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
882  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
883  | 
val compat =  | 
| 63011 | 884  | 
mk_compat_proof_obligations domT ranT (Free fvar) f abstract_qglrs  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59618 
diff
changeset
 | 
885  | 
|> map (Thm.cterm_of lthy #> Thm.assume)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
886  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
887  | 
val compat_store = store_compat_thms n compat  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
888  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
889  | 
val (goalstate, values) = PROFILE "prove_stuff"  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
890  | 
(prove_stuff lthy globals G f R xclauses complete compat  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
891  | 
compat_store G_elim) f_defthm  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
892  | 
|
| 60643 | 893  | 
fun mk_partial_rules newctxt provedgoal =  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
894  | 
let  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
895  | 
val (graph_is_function, complete_thm) =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
896  | 
provedgoal  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
897  | 
|> Conjunction.elim  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
898  | 
|> apfst (Thm.forall_elim_vars 0)  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
899  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
900  | 
val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
901  | 
|
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
902  | 
val psimps = PROFILE "Proving simplification rules"  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
903  | 
(mk_psimps newctxt globals R xclauses values f_iff) graph_is_function  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
904  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
905  | 
val simple_pinduct = PROFILE "Proving partial induction rule"  | 
| 59627 | 906  | 
(mk_partial_induct_rule newctxt globals R complete_thm) xclauses  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
907  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
908  | 
val total_intro = PROFILE "Proving nested termination rule"  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49170 
diff
changeset
 | 
909  | 
(mk_nest_term_rule newctxt globals R R_elim) xclauses  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
910  | 
|
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
911  | 
val dom_intros =  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
912  | 
if domintros then SOME (PROFILE "Proving domain introduction rules"  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
913  | 
(map (mk_domain_intro lthy globals R R_elim)) xclauses)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
914  | 
else NONE  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
915  | 
in  | 
| 
53603
 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 
Manuel Eberl 
parents: 
52467 
diff
changeset
 | 
916  | 
        FunctionResult {fs=[f], G=G, R=R, dom=dom,
 | 
| 
 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 
Manuel Eberl 
parents: 
52467 
diff
changeset
 | 
917  | 
cases=[complete_thm], psimps=psimps, pelims=[],  | 
| 
 
59ef06cda7b9
generate elim rules for elimination of function equalities;
 
Manuel Eberl 
parents: 
52467 
diff
changeset
 | 
918  | 
simple_pinducts=[simple_pinduct],  | 
| 
41846
 
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
 
krauss 
parents: 
41417 
diff
changeset
 | 
919  | 
termination=total_intro, domintros=dom_intros}  | 
| 
34232
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
920  | 
end  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
921  | 
in  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
922  | 
((f, goalstate, mk_partial_rules), lthy)  | 
| 
 
36a2a3029fd3
new year's resolution: reindented code in function package
 
krauss 
parents: 
34065 
diff
changeset
 | 
923  | 
end  | 
| 
33099
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
924  | 
|
| 
 
b8cdd3d73022
function package: more standard names for structures and files
 
krauss 
parents:  
diff
changeset
 | 
925  | 
end  |