| 4714 |      1 | (*  Title:      HOL/Lex/MaxChop.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Tobias Nipkow
 | 
|  |      4 |     Copyright   1998 TUM
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | (* Termination of chop *)
 | 
|  |      8 | 
 | 
| 5069 |      9 | Goalw [reducing_def] "reducing(%qs. maxsplit P ([],qs) [] qs)";
 | 
| 5132 |     10 | by (asm_full_simp_tac (simpset() addsimps [maxsplit_eq]) 1);
 | 
| 4714 |     11 | qed "reducing_maxsplit";
 | 
|  |     12 | 
 | 
|  |     13 | val [tc] = chopr.tcs;
 | 
|  |     14 | goalw_cterm [reducing_def] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop tc));
 | 
| 5132 |     15 | by (blast_tac (claset() addDs [sym]) 1);
 | 
| 4714 |     16 | val lemma = result();
 | 
|  |     17 | 
 | 
| 8624 |     18 | val chopr_rule = let val [chopr_rule] = chopr.simps in lemma RS chopr_rule end;
 | 
| 4714 |     19 | 
 | 
| 5118 |     20 | Goalw [chop_def] "reducing splitf ==> \
 | 
| 4714 |     21 | \ chop splitf xs = (let (pre,post) = splitf xs \
 | 
|  |     22 | \                   in if pre=[] then ([],xs) \
 | 
|  |     23 | \                      else let (xss,zs) = chop splitf post \
 | 
|  |     24 | \                             in (pre#xss,zs))";
 | 
| 6918 |     25 | by (asm_simp_tac (simpset() addsimps [chopr_rule]) 1);
 | 
| 5132 |     26 | by (simp_tac (simpset() addsimps [Let_def] addsplits [split_split]) 1);
 | 
| 4714 |     27 | qed "chop_rule";
 | 
|  |     28 | 
 | 
| 5069 |     29 | Goalw [is_maxsplitter_def,reducing_def]
 | 
| 5118 |     30 |  "is_maxsplitter P splitf ==> reducing splitf";
 | 
| 5132 |     31 | by (Asm_full_simp_tac 1);
 | 
| 4714 |     32 | qed "is_maxsplitter_reducing";
 | 
|  |     33 | 
 | 
| 5118 |     34 | Goal "is_maxsplitter P splitf ==> \
 | 
| 4714 |     35 | \ !yss zs. chop splitf xs = (yss,zs) --> xs = concat yss @ zs";
 | 
| 9747 |     36 | by (induct_thm_tac length_induct "xs" 1);
 | 
| 5132 |     37 | by (asm_simp_tac (simpset() delsplits [split_if]
 | 
| 6918 |     38 |                            addsimps [chop_rule,is_maxsplitter_reducing]) 1);
 | 
| 5132 |     39 | by (asm_full_simp_tac (simpset() addsimps [Let_def,is_maxsplitter_def]
 | 
| 4832 |     40 |                                 addsplits [split_split]) 1);
 | 
| 4714 |     41 | qed_spec_mp "chop_concat";
 | 
|  |     42 | 
 | 
| 5118 |     43 | Goal "is_maxsplitter P splitf ==> \
 | 
| 4714 |     44 | \ !yss zs. chop splitf xs = (yss,zs) --> (!ys : set yss. ys ~= [])";
 | 
| 9747 |     45 | by (induct_thm_tac length_induct "xs" 1);
 | 
| 6918 |     46 | by (asm_simp_tac (simpset() addsimps [chop_rule,is_maxsplitter_reducing]) 1);
 | 
| 5132 |     47 | by (asm_full_simp_tac (simpset() addsimps [Let_def,is_maxsplitter_def]
 | 
| 6918 |     48 |                                  addsplits [split_split]) 1);
 | 
| 5132 |     49 | by (simp_tac (simpset() addsimps [Let_def,maxsplit_eq]
 | 
| 6918 |     50 |                         addsplits [split_split]) 1);
 | 
| 5132 |     51 | by (etac thin_rl 1);
 | 
|  |     52 | by (strip_tac 1);
 | 
|  |     53 | by (rtac ballI 1);
 | 
|  |     54 | by (etac exE 1);
 | 
|  |     55 | by (etac allE 1);
 | 
|  |     56 | by Auto_tac;
 | 
| 9896 |     57 | qed "chop_nonempty";
 | 
| 4714 |     58 | 
 | 
|  |     59 | val [prem] = goalw thy [is_maxchopper_def]
 | 
|  |     60 |  "is_maxsplitter P splitf ==> is_maxchopper P (chop splitf)";
 | 
| 5132 |     61 | by (Clarify_tac 1);
 | 
|  |     62 | by (rtac iffI 1);
 | 
|  |     63 |  by (rtac conjI 1);
 | 
|  |     64 |   by (etac (prem RS chop_concat) 1);
 | 
|  |     65 |  by (rtac conjI 1);
 | 
| 9896 |     66 |   by (etac (prem RS (chop_nonempty RS spec RS spec RS mp)) 1);
 | 
| 5132 |     67 |  by (etac rev_mp 1);
 | 
|  |     68 |  by (stac (prem RS is_maxsplitter_reducing RS chop_rule) 1);
 | 
|  |     69 |  by (simp_tac (simpset() addsimps [Let_def,rewrite_rule[is_maxsplitter_def]prem]
 | 
| 4832 |     70 |                         addsplits [split_split]) 1);
 | 
| 5132 |     71 |  by (Clarify_tac 1);
 | 
|  |     72 |  by (rtac conjI 1);
 | 
|  |     73 |   by (Clarify_tac 1);
 | 
|  |     74 |  by (Clarify_tac 1);
 | 
|  |     75 |  by (Asm_full_simp_tac 1);
 | 
|  |     76 |  by (forward_tac [prem RS chop_concat] 1);
 | 
|  |     77 |  by (Clarify_tac 1);
 | 
|  |     78 | by (stac (prem RS is_maxsplitter_reducing RS chop_rule) 1);
 | 
| 5161 |     79 | by (simp_tac (simpset() addsimps [Let_def,rewrite_rule[is_maxsplitter_def]prem]
 | 
| 4832 |     80 |                         addsplits [split_split]) 1);
 | 
| 5132 |     81 | by (Clarify_tac 1);
 | 
| 5168 |     82 | by (rename_tac "xs1 ys1 xss1 ys" 1);
 | 
| 5184 |     83 | by (split_asm_tac [list.split_asm] 1);
 | 
| 5132 |     84 |  by (Asm_full_simp_tac 1);
 | 
|  |     85 |  by (full_simp_tac (simpset() addsimps [is_maxpref_def]) 1);
 | 
| 10338 |     86 |  by (blast_tac (claset() addIs [thm "prefix_append" RS iffD2]) 1);
 | 
| 5161 |     87 | by (rtac conjI 1);
 | 
|  |     88 |  by (Clarify_tac 1);
 | 
| 5132 |     89 |  by (full_simp_tac (simpset() addsimps [is_maxpref_def]) 1);
 | 
| 10338 |     90 |  by (blast_tac (claset() addIs [thm "prefix_append" RS iffD2]) 1);
 | 
| 5132 |     91 | by (Clarify_tac 1);
 | 
| 5168 |     92 | by (rename_tac "us uss" 1);
 | 
| 5161 |     93 | by (subgoal_tac "xs1=us" 1);
 | 
| 5132 |     94 |  by (Asm_full_simp_tac 1);
 | 
|  |     95 | by (Asm_full_simp_tac 1);
 | 
|  |     96 | by (full_simp_tac (simpset() addsimps [is_maxpref_def]) 1);
 | 
| 10338 |     97 | by (blast_tac (claset() addIs [thm "prefix_append" RS iffD2, order_antisym]) 1);
 | 
| 4714 |     98 | qed "is_maxchopper_chop";
 |