67 NoSyn); |
67 NoSyn); |
68 in (map (fn a => "'" ^ a) vs @ map (fst o fst) iTs, tname, NoSyn, |
68 in (map (fn a => "'" ^ a) vs @ map (fst o fst) iTs, tname, NoSyn, |
69 map constr_of_intr intrs) |
69 map constr_of_intr intrs) |
70 end; |
70 end; |
71 |
71 |
72 fun gen_realizes (Const ("realizes", Type ("fun", [T, _])) $ t $ |
|
73 (Const ("op :", Type ("fun", [U, _])) $ x $ Var (ixn, _))) = |
|
74 Var (ixn, [T, U] ---> HOLogic.boolT) $ t $ x |
|
75 | gen_realizes (Const ("op :", Type ("fun", [U, _])) $ x $ Var (ixn, _)) = |
|
76 Var (ixn, U --> HOLogic.boolT) $ x |
|
77 | gen_realizes (bla as Const ("realizes", Type ("fun", [T, _])) $ t $ P) = |
|
78 if T = Extraction.nullT then P |
|
79 else (case strip_comb P of |
|
80 (Var (ixn, U), ts) => list_comb (Var (ixn, T --> U), t :: ts) |
|
81 | _ => error "gen_realizes: variable expected") |
|
82 | gen_realizes (t $ u) = gen_realizes t $ gen_realizes u |
|
83 | gen_realizes (Abs (s, T, t)) = Abs (s, T, gen_realizes t) |
|
84 | gen_realizes t = t; |
|
85 |
|
86 fun mk_rlz T = Const ("realizes", [T, HOLogic.boolT] ---> HOLogic.boolT); |
72 fun mk_rlz T = Const ("realizes", [T, HOLogic.boolT] ---> HOLogic.boolT); |
87 fun mk_rlz' T = Const ("realizes", [T, propT] ---> propT); |
|
88 |
73 |
89 (** turn "P" into "%r x. realizes r (P x)" or "%r x. realizes r (x : P)" **) |
74 (** turn "P" into "%r x. realizes r (P x)" or "%r x. realizes r (x : P)" **) |
90 |
75 |
91 fun gen_rvar vs (t as Var ((a, 0), T)) = |
76 fun gen_rvar vs (t as Var ((a, 0), T)) = |
92 let val U = TVar (("'" ^ a, 0), HOLogic.typeS) |
77 let val U = TVar (("'" ^ a, 0), HOLogic.typeS) |
266 end |
251 end |
267 else error msg; |
252 else error msg; |
268 |
253 |
269 fun mk_realizer thy vs params ((rule, rrule), rt) = |
254 fun mk_realizer thy vs params ((rule, rrule), rt) = |
270 let |
255 let |
271 val prems = prems_of rule; |
256 val prems = prems_of rule ~~ prems_of rrule; |
|
257 val rvs = map fst (relevant_vars (prop_of rule)); |
272 val xs = rev (Term.add_vars ([], prop_of rule)); |
258 val xs = rev (Term.add_vars ([], prop_of rule)); |
273 val rs = gen_rems (op = o pairself fst) |
259 val vs1 = map Var (filter_out (fn ((a, _), _) => a mem rvs) xs); |
274 (rev (Term.add_vars ([], prop_of rrule)), xs); |
260 val rlzvs = rev (Term.add_vars ([], prop_of rrule)); |
|
261 val vs2 = map (fn (ixn, _) => Var (ixn, the (assoc (rlzvs, ixn)))) xs; |
|
262 val rs = gen_rems (op = o pairself fst) (rlzvs, xs); |
275 |
263 |
276 fun mk_prf _ [] prf = prf |
264 fun mk_prf _ [] prf = prf |
277 | mk_prf rs (prem :: prems) prf = |
265 | mk_prf rs ((prem, rprem) :: prems) prf = |
278 let val T = Extraction.etype_of thy vs [] prem |
266 if Extraction.etype_of thy vs [] prem = Extraction.nullT |
279 in if T = Extraction.nullT |
267 then AbsP ("H", Some rprem, mk_prf rs prems prf) |
280 then AbsP ("H", Some (mk_rlz' T $ Extraction.nullt $ prem), |
268 else forall_intr_prf (Var (hd rs), AbsP ("H", Some rprem, |
281 mk_prf rs prems prf) |
269 mk_prf (tl rs) prems prf)); |
282 else forall_intr_prf (Var (hd rs), AbsP ("H", Some (mk_rlz' T $ |
|
283 Var (hd rs) $ prem), mk_prf (tl rs) prems prf)) |
|
284 end; |
|
285 |
|
286 val subst = map (fn v as (ixn, _) => (ixn, gen_rvar vs (Var v))) xs; |
|
287 val prf = Proofterm.map_proof_terms |
|
288 (subst_vars ([], subst)) I (prf_of rrule); |
|
289 |
270 |
290 in (Thm.name_of_thm rule, (vs, |
271 in (Thm.name_of_thm rule, (vs, |
291 if rt = Extraction.nullt then rt else |
272 if rt = Extraction.nullt then rt else |
292 foldr (uncurry lambda) (map Var xs, rt), |
273 foldr (uncurry lambda) (vs1, rt), |
293 foldr forall_intr_prf (map Var xs, mk_prf rs prems (Proofterm.proof_combP |
274 foldr forall_intr_prf (vs2, mk_prf rs prems (Proofterm.proof_combP |
294 (prf, map PBound (length prems - 1 downto 0)))))) |
275 (prf_of rrule, map PBound (length prems - 1 downto 0)))))) |
295 end; |
276 end; |
296 |
277 |
297 fun add_rule (rss, r) = |
278 fun add_rule (rss, r) = |
298 let |
279 let |
299 val _ $ (_ $ _ $ S) = concl_of r; |
280 val _ $ (_ $ _ $ S) = concl_of r; |
346 in ((recs2, dummies'), map (head_of o hd o rev o snd o strip_comb o |
327 in ((recs2, dummies'), map (head_of o hd o rev o snd o strip_comb o |
347 fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) recs1) |
328 fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) recs1) |
348 end |
329 end |
349 else ((recs, dummies), replicate (length rs) Extraction.nullt)) |
330 else ((recs, dummies), replicate (length rs) Extraction.nullt)) |
350 ((get #rec_thms dt_info, dummies), rss); |
331 ((get #rec_thms dt_info, dummies), rss); |
351 val rintrs = map (fn (intr, c) => Pattern.eta_contract (gen_realizes |
332 val rintrs = map (fn (intr, c) => Pattern.eta_contract |
352 (Extraction.realizes_of thy2 vs |
333 (Extraction.realizes_of thy2 vs |
353 c (prop_of (forall_intr_list (map (cterm_of (sign_of thy2) o Var) |
334 c (prop_of (forall_intr_list (map (cterm_of (sign_of thy2) o Var) |
354 (rev (Term.add_vars ([], prop_of intr)) \\ params')) intr))))) |
335 (rev (Term.add_vars ([], prop_of intr)) \\ params')) intr)))) |
355 (intrs ~~ flat constrss); |
336 (intrs ~~ flat constrss); |
356 val rlzsets = distinct (map (fn rintr => snd (HOLogic.dest_mem |
337 val rlzsets = distinct (map (fn rintr => snd (HOLogic.dest_mem |
357 (HOLogic.dest_Trueprop (Logic.strip_assums_concl rintr)))) rintrs); |
338 (HOLogic.dest_Trueprop (Logic.strip_assums_concl rintr)))) rintrs); |
358 |
339 |
359 (** realizability predicate **) |
340 (** realizability predicate **) |
375 |
356 |
376 fun add_ind_realizer (thy, Ps) = |
357 fun add_ind_realizer (thy, Ps) = |
377 let |
358 let |
378 val r = indrule_realizer thy induct raw_induct rsets params' |
359 val r = indrule_realizer thy induct raw_induct rsets params' |
379 (vs @ Ps) rec_names rss intrs dummies; |
360 (vs @ Ps) rec_names rss intrs dummies; |
380 val rlz = strip_all (Logic.unvarify (gen_realizes |
361 val rlz = strip_all (Logic.unvarify |
381 (Extraction.realizes_of thy (vs @ Ps) r (prop_of induct)))); |
362 (Extraction.realizes_of thy (vs @ Ps) r (prop_of induct))); |
382 val rews = map mk_meta_eq |
363 val rews = map mk_meta_eq |
383 (fst_conv :: snd_conv :: get #rec_thms dt_info); |
364 (fst_conv :: snd_conv :: get #rec_thms dt_info); |
384 val thm = simple_prove_goal_cterm (cterm_of (sign_of thy) rlz) (fn prems => |
365 val thm = simple_prove_goal_cterm (cterm_of (sign_of thy) rlz) (fn prems => |
385 [if length rss = 1 then |
366 [if length rss = 1 then |
386 cut_facts_tac [hd prems] 1 THEN etac (#induct ind_info) 1 |
367 cut_facts_tac [hd prems] 1 THEN etac (#induct ind_info) 1 |
414 else list_abs (map (pair "x") Ts, list_comb (Const (case_name, T), |
395 else list_abs (map (pair "x") Ts, list_comb (Const (case_name, T), |
415 (if dummy then |
396 (if dummy then |
416 [Abs ("x", HOLogic.unitT, Const ("arbitrary", body_type T))] |
397 [Abs ("x", HOLogic.unitT, Const ("arbitrary", body_type T))] |
417 else []) @ |
398 else []) @ |
418 map Bound ((length prems - 1 downto 0) @ [length prems]))); |
399 map Bound ((length prems - 1 downto 0) @ [length prems]))); |
419 val rlz = strip_all (Logic.unvarify (gen_realizes |
400 val rlz = strip_all (Logic.unvarify |
420 (Extraction.realizes_of thy (vs @ Ps) r (prop_of elim)))); |
401 (Extraction.realizes_of thy (vs @ Ps) r (prop_of elim))); |
421 val rews = map mk_meta_eq case_thms; |
402 val rews = map mk_meta_eq case_thms; |
422 val thm = simple_prove_goal_cterm (cterm_of (sign_of thy) rlz) (fn prems => |
403 val thm = simple_prove_goal_cterm (cterm_of (sign_of thy) rlz) (fn prems => |
423 [cut_facts_tac [hd prems] 1, |
404 [cut_facts_tac [hd prems] 1, |
424 etac elimR 1, |
405 etac elimR 1, |
425 ALLGOALS (EVERY' [etac Pair_inject, asm_simp_tac HOL_basic_ss]), |
406 ALLGOALS (EVERY' [etac Pair_inject, asm_simp_tac HOL_basic_ss]), |