| author | wenzelm | 
| Thu, 13 Feb 2014 11:23:55 +0100 | |
| changeset 55445 | a76c679c0218 | 
| parent 55440 | 721b4561007a | 
| child 55642 | 63beb38e9258 | 
| permissions | -rw-r--r-- | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 1 | (* Title: HOL/Tools/Predicate_Compile/predicate_compile_proof.ML | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 3 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 4 | Proof procedure for the compiler from predicates specified by intro/elim rules to equations. | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 5 | *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 6 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 7 | signature PREDICATE_COMPILE_PROOF = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 8 | sig | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 9 | type indprem = Predicate_Compile_Aux.indprem; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 10 | type mode = Predicate_Compile_Aux.mode | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 11 | val prove_pred : Predicate_Compile_Aux.options -> theory | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 12 | -> (string * (term list * indprem list) list) list | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 13 | -> (string * typ) list -> string -> bool * mode | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 14 | -> (term list * (indprem * Mode_Inference.mode_derivation) list) list * term | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 15 | -> Thm.thm | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 16 | end; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 17 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 18 | structure Predicate_Compile_Proof : PREDICATE_COMPILE_PROOF = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 19 | struct | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 20 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 21 | open Predicate_Compile_Aux; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 22 | open Core_Data; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 23 | open Mode_Inference; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 24 | |
| 55437 | 25 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 26 | (* debug stuff *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 27 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 28 | fun print_tac options s = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 29 | if show_proof_trace options then Tactical.print_tac s else Seq.single; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 30 | |
| 55437 | 31 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 32 | (** auxiliary **) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 33 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 34 | datatype assertion = Max_number_of_subgoals of int | 
| 55437 | 35 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 36 | fun assert_tac (Max_number_of_subgoals i) st = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 37 | if (nprems_of st <= i) then Seq.single st | 
| 55437 | 38 | else | 
| 39 |     raise Fail ("assert_tac: Numbers of subgoals mismatch at goal state :\n" ^
 | |
| 40 | Pretty.string_of (Pretty.chunks | |
| 41 | (Goal_Display.pretty_goals_without_context st))) | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 42 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 43 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 44 | (** special setup for simpset **) | 
| 55437 | 45 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 46 | val HOL_basic_ss' = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 47 |   simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 48 |     addsimps @{thms simp_thms Pair_eq}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 49 | setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 50 |     setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI})))
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 51 | |
| 55437 | 52 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 53 | (* auxillary functions *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 54 | |
| 55399 | 55 | (* returns true if t is an application of a datatype constructor *) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 56 | (* which then consequently would be splitted *) | 
| 55399 | 57 | fun is_constructor ctxt t = | 
| 58 | (case fastype_of t of | |
| 59 |     Type (s, _) => s <> @{type_name fun} andalso can (Ctr_Sugar.dest_ctr ctxt s) t
 | |
| 55440 | 60 | | _ => false) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 61 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 62 | (* MAJOR FIXME: prove_params should be simple | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 63 | - different form of introrule for parameters ? *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 64 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 65 | fun prove_param options ctxt nargs t deriv = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 66 | let | 
| 55399 | 67 | val (f, args) = strip_comb (Envir.eta_contract t) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 68 | val mode = head_mode_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 69 | val param_derivations = param_derivations_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 70 | val ho_args = ho_args_of mode args | 
| 55437 | 71 | val f_tac = | 
| 72 | (case f of | |
| 73 | Const (name, _) => simp_tac (put_simpset HOL_basic_ss ctxt addsimps | |
| 74 |            [@{thm eval_pred}, predfun_definition_of ctxt name mode,
 | |
| 75 |            @{thm split_eta}, @{thm split_beta}, @{thm fst_conv},
 | |
| 76 |            @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
 | |
| 77 | | Free _ => | |
| 78 |         Subgoal.FOCUS_PREMS (fn {context = ctxt', params = params, prems, asms, concl, schematics} =>
 | |
| 79 | let | |
| 80 | val prems' = maps dest_conjunct_prem (take nargs prems) | |
| 81 | in | |
| 82 |             rewrite_goal_tac ctxt' (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
 | |
| 83 | end) ctxt 1 | |
| 84 | | Abs _ => raise Fail "prove_param: No valid parameter term") | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 85 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 86 |     REPEAT_DETERM (rtac @{thm ext} 1)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 87 | THEN print_tac options "prove_param" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 88 | THEN f_tac | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 89 | THEN print_tac options "after prove_param" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 90 | THEN (REPEAT_DETERM (atac 1)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 91 | THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 92 |     THEN REPEAT_DETERM (rtac @{thm refl} 1)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 93 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 94 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 95 | fun prove_expr options ctxt nargs (premposition : int) (t, deriv) = | 
| 55437 | 96 | (case strip_comb t of | 
| 50056 | 97 | (Const (name, _), args) => | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 98 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 99 | val mode = head_mode_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 100 | val introrule = predfun_intro_of ctxt name mode | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 101 | val param_derivations = param_derivations_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 102 | val ho_args = ho_args_of mode args | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 103 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 104 | print_tac options "before intro rule:" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 105 | THEN rtac introrule 1 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 106 | THEN print_tac options "after intro rule" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 107 | (* for the right assumption in first position *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 108 | THEN rotate_tac premposition 1 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 109 | THEN atac 1 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 110 | THEN print_tac options "parameter goal" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 111 | (* work with parameter arguments *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 112 | THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 113 | THEN (REPEAT_DETERM (atac 1)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 114 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 115 | | (Free _, _) => | 
| 55437 | 116 | print_tac options "proving parameter call.." | 
| 117 |       THEN Subgoal.FOCUS_PREMS (fn {context = ctxt', params, prems, asms, concl, schematics} =>
 | |
| 118 | let | |
| 119 | val param_prem = nth prems premposition | |
| 120 | val (param, _) = strip_comb (HOLogic.dest_Trueprop (prop_of param_prem)) | |
| 121 | val prems' = maps dest_conjunct_prem (take nargs prems) | |
| 122 | fun param_rewrite prem = | |
| 123 | param = snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of prem))) | |
| 124 | val SOME rew_eq = find_first param_rewrite prems' | |
| 125 | val param_prem' = rewrite_rule ctxt' | |
| 126 |               (map (fn th => th RS @{thm eq_reflection})
 | |
| 127 |                 [rew_eq RS @{thm sym}, @{thm split_beta}, @{thm fst_conv}, @{thm snd_conv}])
 | |
| 128 | param_prem | |
| 129 | in | |
| 130 | rtac param_prem' 1 | |
| 131 | end) ctxt 1 | |
| 132 | THEN print_tac options "after prove parameter call") | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 133 | |
| 55437 | 134 | fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 135 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 136 | fun prove_match options ctxt nargs out_ts = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 137 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 138 | val eval_if_P = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 139 |       @{lemma "P ==> Predicate.eval x z ==> Predicate.eval (if P then x else y) z" by simp} 
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 140 | fun get_case_rewrite t = | 
| 55399 | 141 | if is_constructor ctxt t then | 
| 45701 | 142 | let | 
| 55399 | 143 |           val SOME {case_thms, ...} = Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type (fastype_of t)))
 | 
| 45701 | 144 | in | 
| 55399 | 145 | fold (union Thm.eq_thm) (case_thms :: map get_case_rewrite (snd (strip_comb t))) [] | 
| 45701 | 146 | end | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 147 | else [] | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 148 |     val simprules = insert Thm.eq_thm @{thm "unit.cases"} (insert Thm.eq_thm @{thm "prod.cases"}
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 149 | (fold (union Thm.eq_thm) (map get_case_rewrite out_ts) [])) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 150 | (* replace TRY by determining if it necessary - are there equations when calling compile match? *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 151 | in | 
| 55437 | 152 | (* make this simpset better! *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 153 | asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt addsimps simprules) 1 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 154 | THEN print_tac options "after prove_match:" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 155 | THEN (DETERM (TRY | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 156 | (rtac eval_if_P 1 | 
| 50056 | 157 |            THEN (SUBPROOF (fn {context, params, prems, asms, concl, schematics} =>
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 158 |              (REPEAT_DETERM (rtac @{thm conjI} 1
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 159 | THEN (SOLVED (asm_simp_tac (put_simpset HOL_basic_ss' ctxt) 1)))) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 160 | THEN print_tac options "if condition to be solved:" | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 161 | THEN asm_simp_tac (put_simpset HOL_basic_ss' ctxt) 1 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 162 | THEN TRY ( | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 163 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 164 | val prems' = maps dest_conjunct_prem (take nargs prems) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 165 | in | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52732diff
changeset | 166 | rewrite_goal_tac ctxt | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 167 |                     (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 168 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 169 |              THEN REPEAT_DETERM (rtac @{thm refl} 1))
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 170 | THEN print_tac options "after if simp; in SUBPROOF") ctxt 1)))) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 171 | THEN print_tac options "after if simplification" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 172 | end; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 173 | |
| 55437 | 174 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 175 | (* corresponds to compile_fun -- maybe call that also compile_sidecond? *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 176 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 177 | fun prove_sidecond ctxt t = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 178 | let | 
| 55437 | 179 | fun preds_of t nameTs = | 
| 180 | (case strip_comb t of | |
| 181 | (Const (name, T), args) => | |
| 182 | if is_registered ctxt name then (name, T) :: nameTs | |
| 183 | else fold preds_of args nameTs | |
| 184 | | _ => nameTs) | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 185 | val preds = preds_of t [] | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 186 | val defs = map | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 187 | (fn (pred, T) => predfun_definition_of ctxt pred | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 188 | (all_input_of T)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 189 | preds | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 190 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 191 | simp_tac (put_simpset HOL_basic_ss ctxt | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 192 |       addsimps (@{thms HOL.simp_thms eval_pred} @ defs)) 1
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 193 | (* need better control here! *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 194 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 195 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 196 | fun prove_clause options ctxt nargs mode (_, clauses) (ts, moded_ps) = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 197 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 198 | val (in_ts, clause_out_ts) = split_mode mode ts; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 199 | fun prove_prems out_ts [] = | 
| 55437 | 200 | (prove_match options ctxt nargs out_ts) | 
| 201 | THEN print_tac options "before simplifying assumptions" | |
| 202 | THEN asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1 | |
| 203 | THEN print_tac options "before single intro rule" | |
| 204 | THEN Subgoal.FOCUS_PREMS | |
| 205 |            (fn {context = ctxt', params, prems, asms, concl, schematics} =>
 | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 206 | let | 
| 55437 | 207 | val prems' = maps dest_conjunct_prem (take nargs prems) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 208 | in | 
| 55437 | 209 |               rewrite_goal_tac ctxt' (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
 | 
| 210 | end) ctxt 1 | |
| 211 |         THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
 | |
| 212 | | prove_prems out_ts ((p, deriv) :: ps) = | |
| 213 | let | |
| 214 | val premposition = (find_index (equal p) clauses) + nargs | |
| 215 | val mode = head_mode_of deriv | |
| 216 | val rest_tac = | |
| 217 |             rtac @{thm bindI} 1
 | |
| 218 | THEN (case p of Prem t => | |
| 219 | let | |
| 220 | val (_, us) = strip_comb t | |
| 221 | val (_, out_ts''') = split_mode mode us | |
| 222 | val rec_tac = prove_prems out_ts''' ps | |
| 223 | in | |
| 224 | print_tac options "before clause:" | |
| 225 | (*THEN asm_simp_tac (put_simpset HOL_basic_ss ctxt) 1*) | |
| 226 | THEN print_tac options "before prove_expr:" | |
| 227 | THEN prove_expr options ctxt nargs premposition (t, deriv) | |
| 228 | THEN print_tac options "after prove_expr:" | |
| 229 | THEN rec_tac | |
| 230 | end | |
| 231 | | Negprem t => | |
| 232 | let | |
| 233 | val (t, args) = strip_comb t | |
| 234 | val (_, out_ts''') = split_mode mode args | |
| 235 | val rec_tac = prove_prems out_ts''' ps | |
| 236 | val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE) | |
| 237 | val neg_intro_rule = | |
| 238 | Option.map (fn name => | |
| 239 | the (predfun_neg_intro_of ctxt name mode)) name | |
| 240 | val param_derivations = param_derivations_of deriv | |
| 241 | val params = ho_args_of mode args | |
| 242 | in | |
| 243 | print_tac options "before prove_neg_expr:" | |
| 244 | THEN full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps | |
| 245 |                   [@{thm split_eta}, @{thm split_beta}, @{thm fst_conv},
 | |
| 246 |                    @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
 | |
| 247 | THEN (if (is_some name) then | |
| 248 | print_tac options "before applying not introduction rule" | |
| 249 | THEN Subgoal.FOCUS_PREMS | |
| 250 |                       (fn {context, params = params, prems, asms, concl, schematics} =>
 | |
| 251 | rtac (the neg_intro_rule) 1 | |
| 252 | THEN rtac (nth prems premposition) 1) ctxt 1 | |
| 253 | THEN print_tac options "after applying not introduction rule" | |
| 254 | THEN (EVERY (map2 (prove_param options ctxt nargs) params param_derivations)) | |
| 255 | THEN (REPEAT_DETERM (atac 1)) | |
| 256 | else | |
| 257 |                     rtac @{thm not_predI'} 1
 | |
| 258 | (* test: *) | |
| 259 |                     THEN dtac @{thm sym} 1
 | |
| 260 | THEN asm_full_simp_tac | |
| 261 |                       (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1)
 | |
| 262 | THEN simp_tac | |
| 263 |                       (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1
 | |
| 264 | THEN rec_tac | |
| 265 | end | |
| 266 | | Sidecond t => | |
| 267 |              rtac @{thm if_predI} 1
 | |
| 268 | THEN print_tac options "before sidecond:" | |
| 269 | THEN prove_sidecond ctxt t | |
| 270 | THEN print_tac options "after sidecond:" | |
| 271 | THEN prove_prems [] ps) | |
| 272 | in (prove_match options ctxt nargs out_ts) | |
| 273 | THEN rest_tac | |
| 274 | end | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 275 | val prems_tac = prove_prems in_ts moded_ps | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 276 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 277 | print_tac options "Proving clause..." | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 278 |     THEN rtac @{thm bindI} 1
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 279 |     THEN rtac @{thm singleI} 1
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 280 | THEN prems_tac | 
| 55437 | 281 | end | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 282 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 283 | fun select_sup 1 1 = [] | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 284 |   | select_sup _ 1 = [rtac @{thm supI1}]
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 285 |   | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 286 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 287 | fun prove_one_direction options ctxt clauses preds pred mode moded_clauses = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 288 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 289 | val T = the (AList.lookup (op =) preds pred) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 290 | val nargs = length (binder_types T) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 291 | val pred_case_rule = the_elim_of ctxt pred | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 292 | in | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52732diff
changeset | 293 |     REPEAT_DETERM (CHANGED (rewrite_goals_tac ctxt @{thms split_paired_all}))
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 294 | THEN print_tac options "before applying elim rule" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 295 | THEN etac (predfun_elim_of ctxt pred mode) 1 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 296 | THEN etac pred_case_rule 1 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 297 | THEN print_tac options "after applying elim rule" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 298 | THEN (EVERY (map | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 299 | (fn i => EVERY' (select_sup (length moded_clauses) i) i) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 300 | (1 upto (length moded_clauses)))) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 301 | THEN (EVERY (map2 (prove_clause options ctxt nargs mode) clauses moded_clauses)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 302 | THEN print_tac options "proved one direction" | 
| 55437 | 303 | end | 
| 304 | ||
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 305 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 306 | (** Proof in the other direction **) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 307 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 308 | fun prove_match2 options ctxt out_ts = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 309 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 310 | fun split_term_tac (Free _) = all_tac | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 311 | | split_term_tac t = | 
| 55399 | 312 | if is_constructor ctxt t then | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 313 | let | 
| 55399 | 314 |             val SOME {case_thms, split_asm, ...} =
 | 
| 315 | Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type (fastype_of t))) | |
| 316 | val num_of_constrs = length case_thms | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 317 | val (_, ts) = strip_comb t | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 318 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 319 |             print_tac options ("Term " ^ (Syntax.string_of_term ctxt t) ^ 
 | 
| 55420 | 320 | " splitting with rules \n" ^ Display.string_of_thm ctxt split_asm) | 
| 45701 | 321 | THEN TRY (Splitter.split_asm_tac [split_asm] 1 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 322 | THEN (print_tac options "after splitting with split_asm rules") | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 323 | (* THEN (Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt) 1) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 324 |               THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 325 | THEN (REPEAT_DETERM_N (num_of_constrs - 1) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 326 |                 (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)))
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 327 | THEN (assert_tac (Max_number_of_subgoals 2)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 328 | THEN (EVERY (map split_term_tac ts)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 329 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 330 | else all_tac | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 331 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 332 | split_term_tac (HOLogic.mk_tuple out_ts) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 333 |     THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 334 |     THEN (etac @{thm botE} 2))))
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 335 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 336 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 337 | (* VERY LARGE SIMILIRATIY to function prove_param | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 338 | -- join both functions | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 339 | *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 340 | (* TODO: remove function *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 341 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 342 | fun prove_param2 options ctxt t deriv = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 343 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 344 | val (f, args) = strip_comb (Envir.eta_contract t) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 345 | val mode = head_mode_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 346 | val param_derivations = param_derivations_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 347 | val ho_args = ho_args_of mode args | 
| 55437 | 348 | val f_tac = | 
| 349 | (case f of | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 350 | Const (name, _) => full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 351 |            (@{thm eval_pred}::(predfun_definition_of ctxt name mode)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 352 |            :: @{thm "Product_Type.split_conv"}::[])) 1
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 353 | | Free _ => all_tac | 
| 55437 | 354 | | _ => error "prove_param2: illegal parameter term") | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 355 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 356 | print_tac options "before simplification in prove_args:" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 357 | THEN f_tac | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 358 | THEN print_tac options "after simplification in prove_args" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 359 | THEN EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 360 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 361 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 362 | fun prove_expr2 options ctxt (t, deriv) = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 363 | (case strip_comb t of | 
| 50056 | 364 | (Const (name, _), args) => | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 365 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 366 | val mode = head_mode_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 367 | val param_derivations = param_derivations_of deriv | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 368 | val ho_args = ho_args_of mode args | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 369 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 370 |           etac @{thm bindE} 1
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52732diff
changeset | 371 |           THEN (REPEAT_DETERM (CHANGED (rewrite_goals_tac ctxt @{thms split_paired_all})))
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 372 | THEN print_tac options "prove_expr2-before" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 373 | THEN etac (predfun_elim_of ctxt name mode) 1 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 374 | THEN print_tac options "prove_expr2" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 375 | THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 376 | THEN print_tac options "finished prove_expr2" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 377 | end | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 378 |       | _ => etac @{thm bindE} 1)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 379 | |
| 55437 | 380 | fun prove_sidecond2 options ctxt t = | 
| 381 | let | |
| 382 | fun preds_of t nameTs = | |
| 383 | (case strip_comb t of | |
| 384 | (Const (name, T), args) => | |
| 385 | if is_registered ctxt name then (name, T) :: nameTs | |
| 386 | else fold preds_of args nameTs | |
| 387 | | _ => nameTs) | |
| 388 | val preds = preds_of t [] | |
| 389 | val defs = map | |
| 390 | (fn (pred, T) => predfun_definition_of ctxt pred | |
| 391 | (all_input_of T)) | |
| 392 | preds | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 393 | in | 
| 55437 | 394 | (* only simplify the one assumption *) | 
| 395 |     full_simp_tac (put_simpset HOL_basic_ss' ctxt addsimps @{thm eval_pred} :: defs) 1 
 | |
| 396 | (* need better control here! *) | |
| 397 | THEN print_tac options "after sidecond2 simplification" | |
| 398 | end | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 399 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 400 | fun prove_clause2 options ctxt pred mode (ts, ps) i = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 401 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 402 | val pred_intro_rule = nth (intros_of ctxt pred) (i - 1) | 
| 50056 | 403 | val (in_ts, _) = split_mode mode ts; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 404 | val split_simpset = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 405 | put_simpset HOL_basic_ss' ctxt | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 406 |       addsimps [@{thm split_eta}, @{thm split_beta},
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 407 |         @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}]
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 408 | fun prove_prems2 out_ts [] = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 409 | print_tac options "before prove_match2 - last call:" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 410 | THEN prove_match2 options ctxt out_ts | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 411 | THEN print_tac options "after prove_match2 - last call:" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 412 |       THEN (etac @{thm singleE} 1)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 413 |       THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 414 | THEN (asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 415 | THEN TRY ( | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 416 |         (REPEAT_DETERM (etac @{thm Pair_inject} 1))
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 417 | THEN (asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 418 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 419 | THEN SOLVED (print_tac options "state before applying intro rule:" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 420 | THEN (rtac pred_intro_rule | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 421 | (* How to handle equality correctly? *) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 422 | THEN_ALL_NEW (K (print_tac options "state before assumption matching") | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51552diff
changeset | 423 | THEN' (atac ORELSE' ((CHANGED o asm_full_simp_tac split_simpset) THEN' (TRY o atac))) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 424 | THEN' (K (print_tac options "state after pre-simplification:")) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 425 | THEN' (K (print_tac options "state after assumption matching:")))) 1)) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 426 | | prove_prems2 out_ts ((p, deriv) :: ps) = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 427 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 428 | val mode = head_mode_of deriv | 
| 55437 | 429 | val rest_tac = | 
| 430 | (case p of | |
| 431 | Prem t => | |
| 432 | let | |
| 433 | val (_, us) = strip_comb t | |
| 434 | val (_, out_ts''') = split_mode mode us | |
| 435 | val rec_tac = prove_prems2 out_ts''' ps | |
| 436 | in | |
| 437 | (prove_expr2 options ctxt (t, deriv)) THEN rec_tac | |
| 438 | end | |
| 439 | | Negprem t => | |
| 440 | let | |
| 441 | val (_, args) = strip_comb t | |
| 442 | val (_, out_ts''') = split_mode mode args | |
| 443 | val rec_tac = prove_prems2 out_ts''' ps | |
| 444 | val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE) | |
| 445 | val param_derivations = param_derivations_of deriv | |
| 446 | val ho_args = ho_args_of mode args | |
| 447 | in | |
| 448 | print_tac options "before neg prem 2" | |
| 449 |                 THEN etac @{thm bindE} 1
 | |
| 450 | THEN (if is_some name then | |
| 451 | full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps | |
| 452 | [predfun_definition_of ctxt (the name) mode]) 1 | |
| 453 |                     THEN etac @{thm not_predE} 1
 | |
| 454 |                     THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1
 | |
| 455 | THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations)) | |
| 456 | else | |
| 457 |                     etac @{thm not_predE'} 1)
 | |
| 458 | THEN rec_tac | |
| 459 | end | |
| 460 | | Sidecond t => | |
| 461 |               etac @{thm bindE} 1
 | |
| 462 |               THEN etac @{thm if_predE} 1
 | |
| 463 | THEN prove_sidecond2 options ctxt t | |
| 464 | THEN prove_prems2 [] ps) | |
| 465 | in | |
| 466 | print_tac options "before prove_match2:" | |
| 467 | THEN prove_match2 options ctxt out_ts | |
| 468 | THEN print_tac options "after prove_match2:" | |
| 469 | THEN rest_tac | |
| 470 | end | |
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 471 | val prems_tac = prove_prems2 in_ts ps | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 472 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 473 | print_tac options "starting prove_clause2" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 474 |     THEN etac @{thm bindE} 1
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 475 |     THEN (etac @{thm singleE'} 1)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 476 |     THEN (TRY (etac @{thm Pair_inject} 1))
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 477 | THEN print_tac options "after singleE':" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 478 | THEN prems_tac | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 479 | end; | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 480 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 481 | fun prove_other_direction options ctxt pred mode moded_clauses = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 482 | let | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 483 | fun prove_clause clause i = | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 484 |       (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 485 | THEN (prove_clause2 options ctxt pred mode clause i) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 486 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 487 |     (DETERM (TRY (rtac @{thm unit.induct} 1)))
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52732diff
changeset | 488 |      THEN (REPEAT_DETERM (CHANGED (rewrite_goals_tac ctxt @{thms split_paired_all})))
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 489 | THEN (rtac (predfun_intro_of ctxt pred mode) 1) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 490 |      THEN (REPEAT_DETERM (rtac @{thm refl} 2))
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52732diff
changeset | 491 | THEN ( | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52732diff
changeset | 492 |        if null moded_clauses then etac @{thm botE} 1
 | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 493 | else EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses)))) | 
| 55437 | 494 | end | 
| 495 | ||
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 496 | |
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 497 | (** proof procedure **) | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 498 | |
| 50056 | 499 | fun prove_pred options thy clauses preds pred (_, mode) (moded_clauses, compiled_term) = | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 500 | let | 
| 51552 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
50056diff
changeset | 501 | val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) | 
| 55437 | 502 | val clauses = (case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []) | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 503 | in | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 504 | Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 505 | (if not (skip_proof options) then | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 506 | (fn _ => | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 507 |         rtac @{thm pred_iffI} 1
 | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 508 | THEN print_tac options "after pred_iffI" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 509 | THEN prove_one_direction options ctxt clauses preds pred mode moded_clauses | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 510 | THEN print_tac options "proved one direction" | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 511 | THEN prove_other_direction options ctxt pred mode moded_clauses | 
| 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 512 | THEN print_tac options "proved other direction") | 
| 51552 
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
 wenzelm parents: 
50056diff
changeset | 513 | else (fn _ => ALLGOALS Skip_Proof.cheat_tac)) | 
| 55437 | 514 | end | 
| 40052 
ea46574ca815
splitting large core file into core_data, mode_inference and predicate_compile_proof
 bulwahn parents: diff
changeset | 515 | |
| 55437 | 516 | end |