Nominal/Tacs.thy
changeset 2300 9fb315392493
parent 2108 c5b7be27f105
child 2302 c6db12ddb60c
equal deleted inserted replaced
2299:09bbed4f21d6 2300:9fb315392493
    15 in
    15 in
    16   InductTacs.case_rule_tac ctxt0 s (hd ty_rules) i st
    16   InductTacs.case_rule_tac ctxt0 s (hd ty_rules) i st
    17 end
    17 end
    18 *}
    18 *}
    19 
    19 
    20 ML {*
       
    21 fun mk_conjl props =
       
    22   fold (fn a => fn b =>
       
    23     if a = @{term True} then b else
       
    24     if b = @{term True} then a else
       
    25     HOLogic.mk_conj (a, b)) (rev props) @{term True};
       
    26 *}
       
    27 
    20 
    28 ML {*
       
    29 val split_conj_tac = REPEAT o etac conjE THEN' TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI)
       
    30 *}
       
    31 
       
    32 (* Given function for buildng a goal for an input, prepares a
       
    33    one common goals for all the inputs and proves it by induction
       
    34    together *)
       
    35 ML {*
       
    36 fun prove_by_induct tys build_goal ind utac inputs ctxt =
       
    37 let
       
    38   val names = Datatype_Prop.make_tnames tys;
       
    39   val (names', ctxt') = Variable.variant_fixes names ctxt;
       
    40   val frees = map Free (names' ~~ tys);
       
    41   val (gls_lists, ctxt'') = fold_map (build_goal (tys ~~ frees)) inputs ctxt';
       
    42   val gls = flat gls_lists;
       
    43   fun trm_gls_map t = filter (exists_subterm (fn s => s = t)) gls;
       
    44   val trm_gl_lists = map trm_gls_map frees;
       
    45   val trm_gl_insts = map2 (fn n => fn l => [NONE, if l = [] then NONE else SOME n]) names' trm_gl_lists
       
    46   val trm_gls = map mk_conjl trm_gl_lists;
       
    47   val gl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj trm_gls);
       
    48   fun tac {context,...} = (
       
    49     InductTacs.induct_rules_tac context [(flat trm_gl_insts)] [ind]
       
    50     THEN_ALL_NEW split_conj_tac THEN_ALL_NEW utac) 1
       
    51   val th_loc = Goal.prove ctxt'' [] [] gl tac
       
    52   val ths_loc = HOLogic.conj_elims th_loc
       
    53   val ths = Variable.export ctxt'' ctxt ths_loc
       
    54 in
       
    55   filter (fn x => not (prop_of x = prop_of @{thm TrueI})) ths
       
    56 end
       
    57 *}
       
    58 
    21 
    59 ML {*
    22 ML {*
    60 fun prove_by_rel_induct alphas build_goal ind utac inputs ctxt =
    23 fun prove_by_rel_induct alphas build_goal ind utac inputs ctxt =
    61 let
    24 let
    62   val tys = map (domain_type o fastype_of) alphas;
    25   val tys = map (domain_type o fastype_of) alphas;
    79   val th_loc = Goal.prove ctxt'' [] [] gl tac
    42   val th_loc = Goal.prove ctxt'' [] [] gl tac
    80   val ths_loc = HOLogic.conj_elims th_loc
    43   val ths_loc = HOLogic.conj_elims th_loc
    81   val ths = Variable.export ctxt'' ctxt ths_loc
    44   val ths = Variable.export ctxt'' ctxt ths_loc
    82 in
    45 in
    83   filter (fn x => not (prop_of x = prop_of @{thm TrueI})) ths
    46   filter (fn x => not (prop_of x = prop_of @{thm TrueI})) ths
    84 end
       
    85 *}
       
    86 (* Code for transforming an inductive relation to a function *)
       
    87 ML {*
       
    88 fun rel_inj_tac dist_inj intrs elims =
       
    89   SOLVED' (asm_full_simp_tac (HOL_ss addsimps intrs)) ORELSE'
       
    90   (rtac @{thm iffI} THEN' RANGE [
       
    91      (eresolve_tac elims THEN_ALL_NEW
       
    92        asm_full_simp_tac (HOL_ss addsimps dist_inj)
       
    93      ),
       
    94      asm_full_simp_tac (HOL_ss addsimps intrs)])
       
    95 *}
       
    96 
       
    97 ML {*
       
    98 fun build_rel_inj_gl thm =
       
    99   let
       
   100     val prop = prop_of thm;
       
   101     val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop);
       
   102     val hyps = map HOLogic.dest_Trueprop (Logic.strip_imp_prems prop);
       
   103     fun list_conj l = foldr1 HOLogic.mk_conj l;
       
   104   in
       
   105     if hyps = [] then concl
       
   106     else HOLogic.mk_eq (concl, list_conj hyps)
       
   107   end;
       
   108 *}
       
   109 
       
   110 ML {*
       
   111 fun build_rel_inj intrs dist_inj elims ctxt =
       
   112 let
       
   113   val ((_, thms_imp), ctxt') = Variable.import false intrs ctxt;
       
   114   val gls = map (HOLogic.mk_Trueprop o build_rel_inj_gl) thms_imp;
       
   115   fun tac _ = rel_inj_tac dist_inj intrs elims 1;
       
   116   val thms = map (fn gl => Goal.prove ctxt' [] [] gl tac) gls;
       
   117 in
       
   118   Variable.export ctxt' ctxt thms
       
   119 end
    47 end
   120 *}
    48 *}
   121 
    49 
   122 ML {*
    50 ML {*
   123 fun repeat_mp thm = repeat_mp (mp OF [thm]) handle THM _ => thm
    51 fun repeat_mp thm = repeat_mp (mp OF [thm]) handle THM _ => thm