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