76 #> snd |
76 #> snd |
77 in |
77 in |
78 (saved_simps, fold2 add_for_f fnames simps_by_f lthy) |
78 (saved_simps, fold2 add_for_f fnames simps_by_f lthy) |
79 end |
79 end |
80 |
80 |
81 fun prepare_function do_print prep default_constraint fixspec eqns config lthy = |
81 fun prepare_function do_print prep fixspec eqns config lthy = |
82 let |
82 let |
83 val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx)) |
83 val ((fixes0, spec0), ctxt') = |
84 val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy |
84 lthy |
|
85 |> Proof_Context.set_object_logic_constraint |
|
86 |> prep fixspec eqns |
|
87 ||> Proof_Context.restore_object_logic_constraint lthy; |
85 val fixes = map (apfst (apfst Binding.name_of)) fixes0; |
88 val fixes = map (apfst (apfst Binding.name_of)) fixes0; |
86 val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0; |
89 val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0; |
87 val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec |
90 val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec |
88 |
91 |
89 val defname = mk_defname fixes |
92 val defname = mk_defname fixes |
141 end |
144 end |
142 in |
145 in |
143 ((goal_state, afterqed), lthy') |
146 ((goal_state, afterqed), lthy') |
144 end |
147 end |
145 |
148 |
146 fun gen_add_function do_print prep default_constraint fixspec eqns config tac lthy = |
149 fun gen_add_function do_print prep fixspec eqns config tac lthy = |
147 let |
150 let |
148 val ((goal_state, afterqed), lthy') = |
151 val ((goal_state, afterqed), lthy') = |
149 prepare_function do_print prep default_constraint fixspec eqns config lthy |
152 prepare_function do_print prep fixspec eqns config lthy |
150 val pattern_thm = |
153 val pattern_thm = |
151 case SINGLE (tac lthy') goal_state of |
154 case SINGLE (tac lthy') goal_state of |
152 NONE => error "pattern completeness and compatibility proof failed" |
155 NONE => error "pattern completeness and compatibility proof failed" |
153 | SOME st => Goal.finish lthy' st |
156 | SOME st => Goal.finish lthy' st |
154 in |
157 in |
155 lthy' |
158 lthy' |
156 |> afterqed [[pattern_thm]] |
159 |> afterqed [[pattern_thm]] |
157 end |
160 end |
158 |
161 |
159 val default_constraint_any = Type_Infer.anyT @{sort type}; |
162 val add_function = gen_add_function false Specification.check_spec |
160 val default_constraint_any' = YXML.string_of_body (Term_XML.Encode.typ default_constraint_any); |
163 fun add_function_cmd a b c d int = gen_add_function int Specification.read_spec a b c d |
161 |
164 |
162 val add_function = |
165 fun gen_function do_print prep fixspec eqns config lthy = |
163 gen_add_function false Specification.check_spec default_constraint_any |
|
164 fun add_function_cmd a b c d int = |
|
165 gen_add_function int Specification.read_spec default_constraint_any' a b c d |
|
166 |
|
167 fun gen_function do_print prep default_constraint fixspec eqns config lthy = |
|
168 let |
166 let |
169 val ((goal_state, afterqed), lthy') = |
167 val ((goal_state, afterqed), lthy') = |
170 prepare_function do_print prep default_constraint fixspec eqns config lthy |
168 prepare_function do_print prep fixspec eqns config lthy |
171 in |
169 in |
172 lthy' |
170 lthy' |
173 |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (Thm.concl_of goal_state), [])]] |
171 |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (Thm.concl_of goal_state), [])]] |
174 |> Proof.refine_singleton (Method.primitive_text (K (K goal_state))) |
172 |> Proof.refine_singleton (Method.primitive_text (K (K goal_state))) |
175 end |
173 end |
176 |
174 |
177 val function = |
175 val function = gen_function false Specification.check_spec |
178 gen_function false Specification.check_spec default_constraint_any |
176 fun function_cmd a b c int = gen_function int Specification.read_spec a b c |
179 fun function_cmd a b c int = |
|
180 gen_function int Specification.read_spec default_constraint_any' a b c |
|
181 |
177 |
182 fun prepare_termination_proof prep_term raw_term_opt lthy = |
178 fun prepare_termination_proof prep_term raw_term_opt lthy = |
183 let |
179 let |
184 val term_opt = Option.map (prep_term lthy) raw_term_opt |
180 val term_opt = Option.map (prep_term lthy) raw_term_opt |
185 val info = |
181 val info = |