src/HOL/Tools/datatype_abs_proofs.ML
changeset 26531 96e82c7861fa
parent 26475 3cc1e48d0ce1
child 26711 3a478bfa1650
equal deleted inserted replaced
26530:777f334ff652 26531:96e82c7861fa
    68         val insts' = (map cert induct_Ps) ~~ (map cert insts);
    68         val insts' = (map cert induct_Ps) ~~ (map cert insts);
    69         val induct' = refl RS ((List.nth
    69         val induct' = refl RS ((List.nth
    70           (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
    70           (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
    71 
    71 
    72       in
    72       in
    73         Goal.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
    73         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
    74           (fn prems => EVERY
    74           (fn prems => EVERY
    75             [rtac induct' 1,
    75             [rtac induct' 1,
    76              REPEAT (rtac TrueI 1),
    76              REPEAT (rtac TrueI 1),
    77              REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
    77              REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
    78              REPEAT (rtac TrueI 1)])
    78              REPEAT (rtac TrueI 1)])
   154         (([], 0), descr' ~~ recTs ~~ rec_sets');
   154         (([], 0), descr' ~~ recTs ~~ rec_sets');
   155 
   155 
   156     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
   156     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
   157         InductivePackage.add_inductive_global (serial_string ())
   157         InductivePackage.add_inductive_global (serial_string ())
   158           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
   158           {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
   159             alt_name = big_rec_name', coind = false, no_elim = false, no_ind = true}
   159             alt_name = big_rec_name', coind = false, no_elim = false, no_ind = true,
       
   160             skip_mono = true}
   160           (map (fn (s, T) => ((s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
   161           (map (fn (s, T) => ((s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
   161           (map dest_Free rec_fns)
   162           (map dest_Free rec_fns)
   162           (map (fn x => (("", []), x)) rec_intr_ts) [] thy0;
   163           (map (fn x => (("", []), x)) rec_intr_ts) [] thy0;
   163 
   164 
   164     (* prove uniqueness and termination of primrec combinators *)
   165     (* prove uniqueness and termination of primrec combinators *)
   214         val (tac, _) = Library.foldl mk_unique_tac
   215         val (tac, _) = Library.foldl mk_unique_tac
   215           (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
   216           (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
   216               THEN rewtac (mk_meta_eq choice_eq), rec_intrs),
   217               THEN rewtac (mk_meta_eq choice_eq), rec_intrs),
   217             descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
   218             descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
   218 
   219 
   219       in split_conj_thm (Goal.prove_global thy1 [] []
   220       in split_conj_thm (SkipProof.prove_global thy1 [] []
   220         (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
   221         (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
   221       end;
   222       end;
   222 
   223 
   223     val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
   224     val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
   224 
   225 
   248 
   249 
   249     (* prove characteristic equations for primrec combinators *)
   250     (* prove characteristic equations for primrec combinators *)
   250 
   251 
   251     val _ = message "Proving characteristic theorems for primrec combinators ..."
   252     val _ = message "Proving characteristic theorems for primrec combinators ..."
   252 
   253 
   253     val rec_thms = map (fn t => Goal.prove_global thy2 [] [] t
   254     val rec_thms = map (fn t => SkipProof.prove_global thy2 [] [] t
   254       (fn _ => EVERY
   255       (fn _ => EVERY
   255         [rewrite_goals_tac reccomb_defs,
   256         [rewrite_goals_tac reccomb_defs,
   256          rtac the1_equality 1,
   257          rtac the1_equality 1,
   257          resolve_tac rec_unique_thms 1,
   258          resolve_tac rec_unique_thms 1,
   258          resolve_tac rec_intrs 1,
   259          resolve_tac rec_intrs 1,
   326 
   327 
   327         in (defs @ [def_thm], thy')
   328         in (defs @ [def_thm], thy')
   328         end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
   329         end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
   329           (Library.take (length newTs, reccomb_names)));
   330           (Library.take (length newTs, reccomb_names)));
   330 
   331 
   331     val case_thms = map (map (fn t => Goal.prove_global thy2 [] [] t
   332     val case_thms = map (map (fn t => SkipProof.prove_global thy2 [] [] t
   332       (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
   333       (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
   333           (DatatypeProp.make_cases new_type_names descr sorts thy2)
   334           (DatatypeProp.make_cases new_type_names descr sorts thy2)
   334 
   335 
   335   in
   336   in
   336     thy2
   337     thy2
   359         val exhaustion' = cterm_instantiate
   360         val exhaustion' = cterm_instantiate
   360           [(cert lhs, cert (Free ("x", T)))] exhaustion;
   361           [(cert lhs, cert (Free ("x", T)))] exhaustion;
   361         val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
   362         val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
   362           (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
   363           (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
   363       in
   364       in
   364         (Goal.prove_global thy [] [] t1 tacf,
   365         (SkipProof.prove_global thy [] [] t1 tacf,
   365          Goal.prove_global thy [] [] t2 tacf)
   366          SkipProof.prove_global thy [] [] t2 tacf)
   366       end;
   367       end;
   367 
   368 
   368     val split_thm_pairs = map prove_split_thms
   369     val split_thm_pairs = map prove_split_thms
   369       ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
   370       ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
   370         dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
   371         dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
   379   end;
   380   end;
   380 
   381 
   381 fun prove_weak_case_congs new_type_names descr sorts thy =
   382 fun prove_weak_case_congs new_type_names descr sorts thy =
   382   let
   383   let
   383     fun prove_weak_case_cong t =
   384     fun prove_weak_case_cong t =
   384        Goal.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   385        SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   385          (fn prems => EVERY [rtac ((hd prems) RS arg_cong) 1])
   386          (fn prems => EVERY [rtac ((hd prems) RS arg_cong) 1])
   386 
   387 
   387     val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
   388     val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
   388       new_type_names descr sorts thy)
   389       new_type_names descr sorts thy)
   389 
   390 
   400         (* For goal i, select the correct disjunct to attack, then prove it *)
   401         (* For goal i, select the correct disjunct to attack, then prove it *)
   401         fun tac i 0 = EVERY [TRY (rtac disjI1 i),
   402         fun tac i 0 = EVERY [TRY (rtac disjI1 i),
   402               hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
   403               hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
   403           | tac i n = rtac disjI2 i THEN tac i (n - 1)
   404           | tac i n = rtac disjI2 i THEN tac i (n - 1)
   404       in 
   405       in 
   405         Goal.prove_global thy [] [] t (fn _ =>
   406         SkipProof.prove_global thy [] [] t (fn _ =>
   406           EVERY [rtac allI 1,
   407           EVERY [rtac allI 1,
   407            exh_tac (K exhaustion) 1,
   408            exh_tac (K exhaustion) 1,
   408            ALLGOALS (fn i => tac i (i-1))])
   409            ALLGOALS (fn i => tac i (i-1))])
   409       end;
   410       end;
   410 
   411 
   422         val cert = cterm_of thy;
   423         val cert = cterm_of thy;
   423         val nchotomy' = nchotomy RS spec;
   424         val nchotomy' = nchotomy RS spec;
   424         val nchotomy'' = cterm_instantiate
   425         val nchotomy'' = cterm_instantiate
   425           [(cert (hd (add_term_vars (concl_of nchotomy', []))), cert Ma)] nchotomy'
   426           [(cert (hd (add_term_vars (concl_of nchotomy', []))), cert Ma)] nchotomy'
   426       in
   427       in
   427         Goal.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   428         SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   428           (fn prems => 
   429           (fn prems => 
   429             let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
   430             let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
   430             in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
   431             in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
   431                 cut_facts_tac [nchotomy''] 1,
   432                 cut_facts_tac [nchotomy''] 1,
   432                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
   433                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),