author | kleing |
Mon, 19 Sep 2005 14:20:45 +0200 | |
changeset 17483 | c6005bfc1630 |
parent 17357 | ee2bdca144c7 |
child 17858 | bc4db8cfd92f |
permissions | -rw-r--r-- |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/Isar/obtain.ML |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
3 |
Author: Markus Wenzel, TU Muenchen |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
4 |
|
8543 | 5 |
The 'obtain' language element -- generalized existence at the level of |
6 |
proof texts. |
|
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
7 |
|
9468 | 8 |
<chain_facts> |
9 |
obtain x where "P x" <proof> == |
|
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
10 |
|
12970 | 11 |
have "!!thesis. (!!x. P x ==> thesis) ==> thesis" |
12 |
proof succeed |
|
9468 | 13 |
fix thesis |
12970 | 14 |
assume that [intro?]: "!!x. P x ==> thesis" |
15 |
<chain_facts> show thesis <proof (insert that)> |
|
16 |
qed |
|
10379
93630e0c5ae9
improved handling of "that": insert into goal, only declare as Pure "intro";
wenzelm
parents:
9481
diff
changeset
|
17 |
fix x assm (obtained) "P x" |
8094 | 18 |
*) |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
19 |
|
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
20 |
signature OBTAIN = |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
21 |
sig |
11890 | 22 |
val obtain: (string list * string option) list -> |
17111 | 23 |
((string * Attrib.src list) * (string * (string list * string list)) list) list |
17357 | 24 |
-> bool -> Proof.state -> Proof.state |
11890 | 25 |
val obtain_i: (string list * typ option) list -> |
26 |
((string * Proof.context attribute list) * (term * (term list * term list)) list) list |
|
17357 | 27 |
-> bool -> Proof.state -> Proof.state |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
28 |
end; |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
29 |
|
10379
93630e0c5ae9
improved handling of "that": insert into goal, only declare as Pure "intro";
wenzelm
parents:
9481
diff
changeset
|
30 |
structure Obtain: OBTAIN = |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
31 |
struct |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
32 |
|
8094 | 33 |
|
11816 | 34 |
(** export_obtain **) |
9468 | 35 |
|
11816 | 36 |
fun export_obtain state parms rule _ cprops thm = |
9468 | 37 |
let |
17111 | 38 |
val {thy, prop, maxidx, ...} = Thm.rep_thm thm; |
39 |
val cparms = map (Thm.cterm_of thy) parms; |
|
9468 | 40 |
|
41 |
val thm' = thm |
|
11816 | 42 |
|> Drule.implies_intr_goals cprops |
9468 | 43 |
|> Drule.forall_intr_list cparms |
44 |
|> Drule.forall_elim_vars (maxidx + 1); |
|
11816 | 45 |
val elim_tacs = replicate (length cprops) (Tactic.etac Drule.triv_goal); |
9468 | 46 |
|
47 |
val concl = Logic.strip_assums_concl prop; |
|
48 |
val bads = parms inter (Term.term_frees concl); |
|
49 |
in |
|
50 |
if not (null bads) then |
|
9481 | 51 |
raise Proof.STATE ("Conclusion contains obtained parameters: " ^ |
12055 | 52 |
space_implode " " (map (ProofContext.string_of_term (Proof.context_of state)) bads), state) |
17111 | 53 |
else if not (ObjectLogic.is_judgment thy (Logic.strip_assums_concl prop)) then |
10379
93630e0c5ae9
improved handling of "that": insert into goal, only declare as Pure "intro";
wenzelm
parents:
9481
diff
changeset
|
54 |
raise Proof.STATE ("Conclusions of 'obtain' context must be object-logic judgments", state) |
9468 | 55 |
else (Tactic.rtac thm' THEN' RANGE elim_tacs) 1 rule |
56 |
end; |
|
57 |
||
58 |
||
59 |
||
8094 | 60 |
(** obtain(_i) **) |
61 |
||
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
62 |
val thatN = "that"; |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
63 |
|
17357 | 64 |
fun gen_obtain prep_att prep_vars prep_propp raw_vars raw_asms int state = |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
65 |
let |
9468 | 66 |
val _ = Proof.assert_forward_or_chain state; |
17357 | 67 |
val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else []; |
10464 | 68 |
val thy = Proof.theory_of state; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
69 |
|
8543 | 70 |
(*obtain vars*) |
11890 | 71 |
val (vars_ctxt, vars) = foldl_map prep_vars (Proof.context_of state, raw_vars); |
15570 | 72 |
val xs = List.concat (map fst vars); |
10582 | 73 |
val fix_ctxt = vars_ctxt |> ProofContext.fix_i vars; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
74 |
|
8543 | 75 |
(*obtain asms*) |
11890 | 76 |
val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms); |
15570 | 77 |
val asm_props = List.concat (map (map fst) proppss); |
17111 | 78 |
val asms = map fst (Attrib.map_specs (prep_att thy) raw_asms) ~~ proppss; |
10464 | 79 |
|
10582 | 80 |
val _ = ProofContext.warn_extra_tfrees fix_ctxt asms_ctxt; |
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
81 |
|
12970 | 82 |
(*obtain statements*) |
16606 | 83 |
val thesisN = Term.variant xs AutoBind.thesisN; |
12970 | 84 |
val bind_thesis = ProofContext.bind_skolem fix_ctxt [thesisN]; |
17111 | 85 |
val bound_thesis = bind_thesis (ObjectLogic.fixed_judgment thy thesisN); |
12970 | 86 |
val bound_thesis_raw as (bound_thesis_name, _) = |
87 |
Term.dest_Free (bind_thesis (Free (thesisN, propT))); |
|
88 |
val bound_thesis_var = |
|
16787 | 89 |
fold_aterms (fn Free (x, T) => (fn v => if x = bound_thesis_name then (x, T) else v) |
90 |
| _ => I) bound_thesis bound_thesis_raw; |
|
9468 | 91 |
|
10582 | 92 |
fun occs_var x = Library.get_first (fn t => |
93 |
ProofContext.find_free t (ProofContext.get_skolem fix_ctxt x)) asm_props; |
|
94 |
val raw_parms = map occs_var xs; |
|
15570 | 95 |
val parms = List.mapPartial I raw_parms; |
10582 | 96 |
val parm_names = |
15570 | 97 |
List.mapPartial (fn (SOME (Free a), x) => SOME (a, x) | _ => NONE) (raw_parms ~~ xs); |
10582 | 98 |
|
99 |
val that_prop = |
|
100 |
Term.list_all_free (map #1 parm_names, Logic.list_implies (asm_props, bound_thesis)) |
|
101 |
|> Library.curry Logic.list_rename_params (map #2 parm_names); |
|
12970 | 102 |
val obtain_prop = |
103 |
Logic.list_rename_params ([AutoBind.thesisN], |
|
104 |
Term.list_all_free ([bound_thesis_var], Logic.mk_implies (that_prop, bound_thesis))); |
|
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
105 |
|
17357 | 106 |
fun after_qed _ _ = |
107 |
Proof.local_qed (NONE, false) |
|
108 |
#> Seq.map (`Proof.the_fact #-> (fn this => |
|
109 |
Proof.fix_i vars |
|
110 |
#> Proof.assm_i (export_obtain state parms this) asms)); |
|
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
111 |
in |
8094 | 112 |
state |
9468 | 113 |
|> Proof.enter_forward |
17357 | 114 |
|> Proof.have_i (K (K Seq.single)) [(("", []), [(obtain_prop, ([], []))])] int |
17111 | 115 |
|> Proof.proof (SOME (Method.Basic (K Method.succeed))) |> Seq.hd |
15531 | 116 |
|> Proof.fix_i [([thesisN], NONE)] |
117 |
|> Proof.assume_i [((thatN, [ContextRules.intro_query_local NONE]), [(that_prop, ([], []))])] |
|
16842 | 118 |
|> `Proof.the_facts |
17357 | 119 |
||> Proof.chain_facts chain_facts |
120 |
||> Proof.show_i after_qed [(("", []), [(bound_thesis, ([], []))])] false |
|
121 |
|-> (Proof.refine o Method.Basic o K o Method.insert) |> Seq.hd |
|
7674
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
122 |
end; |
99305245f6bd
The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff
changeset
|
123 |
|
17111 | 124 |
val obtain = gen_obtain Attrib.local_attribute ProofContext.read_vars ProofContext.read_propp; |
125 |
val obtain_i = gen_obtain (K I) ProofContext.cert_vars ProofContext.cert_propp; |
|
8094 | 126 |
|
127 |
end; |