src/HOL/Tools/Datatype/datatype_abs_proofs.ML
changeset 32970 fbd2bb2489a8
parent 32952 aeb1e44fbc19
child 33056 791a4655cae3
equal deleted inserted replaced
32969:15489e162b21 32970:fbd2bb2489a8
    67         val insts' = (map cert induct_Ps) ~~ (map cert insts);
    67         val insts' = (map cert induct_Ps) ~~ (map cert insts);
    68         val induct' = refl RS ((nth
    68         val induct' = refl RS ((nth
    69           (split_conj_thm (cterm_instantiate insts' induct)) i) RSN (2, rev_mp))
    69           (split_conj_thm (cterm_instantiate insts' induct)) i) RSN (2, rev_mp))
    70 
    70 
    71       in
    71       in
    72         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
    72         Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
    73           (fn {prems, ...} => EVERY
    73           (fn {prems, ...} => EVERY
    74             [rtac induct' 1,
    74             [rtac induct' 1,
    75              REPEAT (rtac TrueI 1),
    75              REPEAT (rtac TrueI 1),
    76              REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
    76              REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
    77              REPEAT (rtac TrueI 1)])
    77              REPEAT (rtac TrueI 1)])
   213           (map cert insts)) induct;
   213           (map cert insts)) induct;
   214         val (tac, _) = fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
   214         val (tac, _) = fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
   215            (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
   215            (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
   216               THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs));
   216               THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs));
   217 
   217 
   218       in split_conj_thm (SkipProof.prove_global thy1 [] []
   218       in split_conj_thm (Skip_Proof.prove_global thy1 [] []
   219         (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
   219         (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
   220       end;
   220       end;
   221 
   221 
   222     val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
   222     val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
   223 
   223 
   248 
   248 
   249     (* prove characteristic equations for primrec combinators *)
   249     (* prove characteristic equations for primrec combinators *)
   250 
   250 
   251     val _ = message config "Proving characteristic theorems for primrec combinators ..."
   251     val _ = message config "Proving characteristic theorems for primrec combinators ..."
   252 
   252 
   253     val rec_thms = map (fn t => SkipProof.prove_global thy2 [] [] t
   253     val rec_thms = map (fn t => Skip_Proof.prove_global thy2 [] [] t
   254       (fn _ => EVERY
   254       (fn _ => EVERY
   255         [rewrite_goals_tac reccomb_defs,
   255         [rewrite_goals_tac reccomb_defs,
   256          rtac the1_equality 1,
   256          rtac the1_equality 1,
   257          resolve_tac rec_unique_thms 1,
   257          resolve_tac rec_unique_thms 1,
   258          resolve_tac rec_intrs 1,
   258          resolve_tac rec_intrs 1,
   328         in (defs @ [def_thm], thy')
   328         in (defs @ [def_thm], thy')
   329         end) (hd descr ~~ newTs ~~ case_names ~~
   329         end) (hd descr ~~ newTs ~~ case_names ~~
   330           Library.take (length newTs, reccomb_names)) ([], thy1)
   330           Library.take (length newTs, reccomb_names)) ([], thy1)
   331       ||> Theory.checkpoint;
   331       ||> Theory.checkpoint;
   332 
   332 
   333     val case_thms = map (map (fn t => SkipProof.prove_global thy2 [] [] t
   333     val case_thms = map (map (fn t => Skip_Proof.prove_global thy2 [] [] t
   334       (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
   334       (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
   335           (DatatypeProp.make_cases new_type_names descr sorts thy2)
   335           (DatatypeProp.make_cases new_type_names descr sorts thy2)
   336   in
   336   in
   337     thy2
   337     thy2
   338     |> Context.the_theory o fold (fold Nitpick_Const_Simps.add_thm) case_thms
   338     |> Context.the_theory o fold (fold Nitpick_Const_Simps.add_thm) case_thms
   362         val exhaustion' = cterm_instantiate
   362         val exhaustion' = cterm_instantiate
   363           [(cert lhs, cert (Free ("x", T)))] exhaustion;
   363           [(cert lhs, cert (Free ("x", T)))] exhaustion;
   364         val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
   364         val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
   365           (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
   365           (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
   366       in
   366       in
   367         (SkipProof.prove_global thy [] [] t1 tacf,
   367         (Skip_Proof.prove_global thy [] [] t1 tacf,
   368          SkipProof.prove_global thy [] [] t2 tacf)
   368          Skip_Proof.prove_global thy [] [] t2 tacf)
   369       end;
   369       end;
   370 
   370 
   371     val split_thm_pairs = map prove_split_thms
   371     val split_thm_pairs = map prove_split_thms
   372       ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
   372       ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
   373         dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
   373         dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
   382   end;
   382   end;
   383 
   383 
   384 fun prove_weak_case_congs new_type_names descr sorts thy =
   384 fun prove_weak_case_congs new_type_names descr sorts thy =
   385   let
   385   let
   386     fun prove_weak_case_cong t =
   386     fun prove_weak_case_cong t =
   387        SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   387        Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   388          (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
   388          (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
   389 
   389 
   390     val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
   390     val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
   391       new_type_names descr sorts thy)
   391       new_type_names descr sorts thy)
   392 
   392 
   403         (* For goal i, select the correct disjunct to attack, then prove it *)
   403         (* For goal i, select the correct disjunct to attack, then prove it *)
   404         fun tac i 0 = EVERY [TRY (rtac disjI1 i),
   404         fun tac i 0 = EVERY [TRY (rtac disjI1 i),
   405               hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
   405               hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
   406           | tac i n = rtac disjI2 i THEN tac i (n - 1)
   406           | tac i n = rtac disjI2 i THEN tac i (n - 1)
   407       in 
   407       in 
   408         SkipProof.prove_global thy [] [] t (fn _ =>
   408         Skip_Proof.prove_global thy [] [] t (fn _ =>
   409           EVERY [rtac allI 1,
   409           EVERY [rtac allI 1,
   410            exh_tac (K exhaustion) 1,
   410            exh_tac (K exhaustion) 1,
   411            ALLGOALS (fn i => tac i (i-1))])
   411            ALLGOALS (fn i => tac i (i-1))])
   412       end;
   412       end;
   413 
   413 
   425         val cert = cterm_of thy;
   425         val cert = cterm_of thy;
   426         val nchotomy' = nchotomy RS spec;
   426         val nchotomy' = nchotomy RS spec;
   427         val [v] = Term.add_vars (concl_of nchotomy') [];
   427         val [v] = Term.add_vars (concl_of nchotomy') [];
   428         val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
   428         val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
   429       in
   429       in
   430         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   430         Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   431           (fn {prems, ...} => 
   431           (fn {prems, ...} => 
   432             let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
   432             let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
   433             in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
   433             in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
   434                 cut_facts_tac [nchotomy''] 1,
   434                 cut_facts_tac [nchotomy''] 1,
   435                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
   435                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),