| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Fri, 26 Mar 2010 10:07:26 +0100 | |
| changeset 1653 | a2142526bb01 | 
| child 1656 | c9d3dda79fe3 | 
| permissions | -rw-r--r-- | 
| 1653 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 1 | theory Tacs | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 2 | imports Main | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 4 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 5 | (* General not-nominal/quotient functionality useful for proving *) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 6 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 7 | (* A version of case_rule_tac that takes more exhaust rules *) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 8 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 9 | fun case_rules_tac ctxt0 s rules i st = | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 10 | let | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 11 | val (_, ctxt) = Variable.focus_subgoal i st ctxt0; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 12 | val ty = fastype_of (ProofContext.read_term_schematic ctxt s) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 13 | fun exhaust_ty thm = fastype_of (hd (Induct.vars_of (Thm.term_of (Thm.cprem_of thm 1)))); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 14 | val ty_rules = filter (fn x => exhaust_ty x = ty) rules; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 15 | in | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 16 | InductTacs.case_rule_tac ctxt0 s (hd ty_rules) i st | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 17 | end | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 18 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 19 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 20 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 21 | fun mk_conjl props = | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 22 | fold (fn a => fn b => | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 23 |     if a = @{term True} then b else
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 24 |     if b = @{term True} then a else
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 25 |     HOLogic.mk_conj (a, b)) (rev props) @{term True};
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 26 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 27 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 28 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 29 | val split_conj_tac = REPEAT o etac conjE THEN' TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 30 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 31 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 32 | (* Given function for buildng a goal for an input, prepares a | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 33 | one common goals for all the inputs and proves it by induction | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 34 | together *) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 35 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 36 | fun prove_by_induct tys build_goal ind utac inputs ctxt = | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 37 | let | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 38 | val names = Datatype_Prop.make_tnames tys; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 39 | val (names', ctxt') = Variable.variant_fixes names ctxt; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 40 | val frees = map Free (names' ~~ tys); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 41 | val (gls_lists, ctxt'') = fold_map (build_goal (tys ~~ frees)) inputs ctxt'; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 42 | val gls = flat gls_lists; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 43 | fun trm_gls_map t = filter (exists_subterm (fn s => s = t)) gls; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 44 | val trm_gl_lists = map trm_gls_map frees; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 45 | val trm_gl_insts = map2 (fn n => fn l => [NONE, if l = [] then NONE else SOME n]) names' trm_gl_lists | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 46 | val trm_gls = map mk_conjl trm_gl_lists; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 47 | val gl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj trm_gls); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 48 |   fun tac {context,...} = (
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 49 | InductTacs.induct_rules_tac context [(flat trm_gl_insts)] [ind] | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | THEN_ALL_NEW split_conj_tac THEN_ALL_NEW utac) 1 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 51 | val th_loc = Goal.prove ctxt'' [] [] gl tac | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 52 | val ths_loc = HOLogic.conj_elims th_loc | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 53 | val ths = Variable.export ctxt'' ctxt ths_loc | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 54 | in | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 55 |   filter (fn x => not (prop_of x = prop_of @{thm TrueI})) ths
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 56 | end | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 57 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 58 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 59 | (* An induction for a single relation is "R x y \<Longrightarrow> P x y" | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 60 | but for multiple relations is "(R1 x y \<longrightarrow> P x y) \<and> (R2 a b \<longrightarrow> P2 a b)" *) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 61 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 62 | fun rel_indtac induct = (rtac impI THEN' etac induct) ORELSE' rtac induct | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 63 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 64 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 65 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 66 | fun prove_by_rel_induct alphas build_goal ind utac inputs ctxt = | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 67 | let | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 68 | val tys = map (domain_type o fastype_of) alphas; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 69 | val names = Datatype_Prop.make_tnames tys; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 70 | val (namesl, ctxt') = Variable.variant_fixes names ctxt; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 71 | val (namesr, ctxt'') = Variable.variant_fixes names ctxt'; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 72 | val freesl = map Free (namesl ~~ tys); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 73 | val freesr = map Free (namesr ~~ tys); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 74 | val (gls_lists, ctxt'') = fold_map (build_goal (tys ~~ (freesl ~~ freesr))) inputs ctxt''; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 75 | val gls = flat gls_lists; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 76 | fun trm_gls_map t = filter (exists_subterm (fn s => s = t)) gls; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 77 | val trm_gl_lists = map trm_gls_map freesl; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 78 | val trm_gls = map mk_conjl trm_gl_lists; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 79 | val pgls = map | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 80 | (fn ((alpha, gl), (l, r)) => HOLogic.mk_imp (alpha $ l $ r, gl)) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 81 | ((alphas ~~ trm_gls) ~~ (freesl ~~ freesr)) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 82 | val gl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj pgls); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 83 |   fun tac {context,...} = (rel_indtac ind THEN_ALL_NEW split_conj_tac THEN_ALL_NEW
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 84 |     TRY o rtac @{thm TrueI} THEN_ALL_NEW utac context) 1
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 85 | val th_loc = Goal.prove ctxt'' [] [] gl tac | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 86 | val ths_loc = HOLogic.conj_elims th_loc | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 87 | val ths = Variable.export ctxt'' ctxt ths_loc | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 88 | in | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 89 |   filter (fn x => not (prop_of x = prop_of @{thm TrueI})) ths
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 90 | end | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 91 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 92 | (* Code for transforming an inductive relation to a function *) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 93 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 94 | fun rel_inj_tac dist_inj intrs elims = | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 95 | SOLVED' (asm_full_simp_tac (HOL_ss addsimps intrs)) ORELSE' | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 96 |   (rtac @{thm iffI} THEN' RANGE [
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 97 | (eresolve_tac elims THEN_ALL_NEW | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 98 | asm_full_simp_tac (HOL_ss addsimps dist_inj) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 99 | ), | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 100 | asm_full_simp_tac (HOL_ss addsimps intrs)]) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 101 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 102 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 103 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 104 | fun build_rel_inj_gl thm = | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 105 | let | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 106 | val prop = prop_of thm; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 107 | val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 108 | val hyps = map HOLogic.dest_Trueprop (Logic.strip_imp_prems prop); | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 109 | fun list_conj l = foldr1 HOLogic.mk_conj l; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 110 | in | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 111 | if hyps = [] then concl | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 112 | else HOLogic.mk_eq (concl, list_conj hyps) | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 113 | end; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 114 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 115 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 116 | ML {*
 | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 117 | fun build_rel_inj intrs dist_inj elims ctxt = | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 118 | let | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 119 | val ((_, thms_imp), ctxt') = Variable.import false intrs ctxt; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 120 | val gls = map (HOLogic.mk_Trueprop o build_rel_inj_gl) thms_imp; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 121 | fun tac _ = rel_inj_tac dist_inj intrs elims 1; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 122 | val thms = map (fn gl => Goal.prove ctxt' [] [] gl tac) gls; | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 123 | in | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 124 | Variable.export ctxt' ctxt thms | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 125 | end | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 126 | *} | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 127 | |
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 128 | end | 
| 
a2142526bb01
Removed another cheat and cleaned the code a bit.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 129 |