src/HOL/Lex/MaxChop.ML
author nipkow
Wed, 30 Aug 2000 16:29:21 +0200
changeset 9747 043098ba5098
parent 8624 69619f870939
child 9896 1319676eb2db
permissions -rw-r--r--
introduced induct_thm_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Lex/MaxChop.ML
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     4
    Copyright   1998 TUM
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     5
*)
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     6
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     7
(* Termination of chop *)
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
     8
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4910
diff changeset
     9
Goalw [reducing_def] "reducing(%qs. maxsplit P ([],qs) [] qs)";
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    10
by (asm_full_simp_tac (simpset() addsimps [maxsplit_eq]) 1);
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    11
qed "reducing_maxsplit";
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    12
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    13
val [tc] = chopr.tcs;
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    14
goalw_cterm [reducing_def] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop tc));
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    15
by (blast_tac (claset() addDs [sym]) 1);
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    16
val lemma = result();
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    17
8624
69619f870939 recdef.rules -> recdef.simps
nipkow
parents: 6918
diff changeset
    18
val chopr_rule = let val [chopr_rule] = chopr.simps in lemma RS chopr_rule end;
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    19
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    20
Goalw [chop_def] "reducing splitf ==> \
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    21
\ chop splitf xs = (let (pre,post) = splitf xs \
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    22
\                   in if pre=[] then ([],xs) \
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    23
\                      else let (xss,zs) = chop splitf post \
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    24
\                             in (pre#xss,zs))";
6918
63c9df6b5c4b Now if_weak_cong is a standard congruence rule
paulson
parents: 6540
diff changeset
    25
by (asm_simp_tac (simpset() addsimps [chopr_rule]) 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    26
by (simp_tac (simpset() addsimps [Let_def] addsplits [split_split]) 1);
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    27
qed "chop_rule";
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    28
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4910
diff changeset
    29
Goalw [is_maxsplitter_def,reducing_def]
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    30
 "is_maxsplitter P splitf ==> reducing splitf";
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    31
by (Asm_full_simp_tac 1);
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    32
qed "is_maxsplitter_reducing";
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    33
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    34
Goal "is_maxsplitter P splitf ==> \
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    35
\ !yss zs. chop splitf xs = (yss,zs) --> xs = concat yss @ zs";
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 8624
diff changeset
    36
by (induct_thm_tac length_induct "xs" 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    37
by (asm_simp_tac (simpset() delsplits [split_if]
6918
63c9df6b5c4b Now if_weak_cong is a standard congruence rule
paulson
parents: 6540
diff changeset
    38
                           addsimps [chop_rule,is_maxsplitter_reducing]) 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    39
by (asm_full_simp_tac (simpset() addsimps [Let_def,is_maxsplitter_def]
4832
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4714
diff changeset
    40
                                addsplits [split_split]) 1);
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    41
qed_spec_mp "chop_concat";
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    42
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    43
Goal "is_maxsplitter P splitf ==> \
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    44
\ !yss zs. chop splitf xs = (yss,zs) --> (!ys : set yss. ys ~= [])";
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 8624
diff changeset
    45
by (induct_thm_tac length_induct "xs" 1);
6918
63c9df6b5c4b Now if_weak_cong is a standard congruence rule
paulson
parents: 6540
diff changeset
    46
by (asm_simp_tac (simpset() addsimps [chop_rule,is_maxsplitter_reducing]) 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    47
by (asm_full_simp_tac (simpset() addsimps [Let_def,is_maxsplitter_def]
6918
63c9df6b5c4b Now if_weak_cong is a standard congruence rule
paulson
parents: 6540
diff changeset
    48
                                 addsplits [split_split]) 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    49
by (simp_tac (simpset() addsimps [Let_def,maxsplit_eq]
6918
63c9df6b5c4b Now if_weak_cong is a standard congruence rule
paulson
parents: 6540
diff changeset
    50
                        addsplits [split_split]) 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    51
by (etac thin_rl 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    52
by (strip_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    53
by (rtac ballI 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    54
by (etac exE 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    55
by (etac allE 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    56
by Auto_tac;
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    57
qed_spec_mp "chop_nonempty";
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    58
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    59
val [prem] = goalw thy [is_maxchopper_def]
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    60
 "is_maxsplitter P splitf ==> is_maxchopper P (chop splitf)";
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    61
by (Clarify_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    62
by (rtac iffI 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    63
 by (rtac conjI 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    64
  by (etac (prem RS chop_concat) 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    65
 by (rtac conjI 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    66
  by (etac (prem RS chop_nonempty) 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    67
 by (etac rev_mp 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    68
 by (stac (prem RS is_maxsplitter_reducing RS chop_rule) 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    69
 by (simp_tac (simpset() addsimps [Let_def,rewrite_rule[is_maxsplitter_def]prem]
4832
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4714
diff changeset
    70
                        addsplits [split_split]) 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    71
 by (Clarify_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    72
 by (rtac conjI 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    73
  by (Clarify_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    74
 by (Clarify_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    75
 by (Asm_full_simp_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    76
 by (forward_tac [prem RS chop_concat] 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    77
 by (Clarify_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    78
by (stac (prem RS is_maxsplitter_reducing RS chop_rule) 1);
5161
e7457679e26d Simplified last proof.
nipkow
parents: 5132
diff changeset
    79
by (simp_tac (simpset() addsimps [Let_def,rewrite_rule[is_maxsplitter_def]prem]
4832
bc11b5b06f87 Added conversion of reg.expr. to automata.
nipkow
parents: 4714
diff changeset
    80
                        addsplits [split_split]) 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    81
by (Clarify_tac 1);
5168
adafef6eb295 isatool expandshort;
wenzelm
parents: 5161
diff changeset
    82
by (rename_tac "xs1 ys1 xss1 ys" 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5168
diff changeset
    83
by (split_asm_tac [list.split_asm] 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    84
 by (Asm_full_simp_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    85
 by (full_simp_tac (simpset() addsimps [is_maxpref_def]) 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    86
 by (blast_tac (claset() addIs [prefix_append RS iffD2]) 1);
5161
e7457679e26d Simplified last proof.
nipkow
parents: 5132
diff changeset
    87
by (rtac conjI 1);
e7457679e26d Simplified last proof.
nipkow
parents: 5132
diff changeset
    88
 by (Clarify_tac 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    89
 by (full_simp_tac (simpset() addsimps [is_maxpref_def]) 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    90
 by (blast_tac (claset() addIs [prefix_append RS iffD2]) 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    91
by (Clarify_tac 1);
5168
adafef6eb295 isatool expandshort;
wenzelm
parents: 5161
diff changeset
    92
by (rename_tac "us uss" 1);
5161
e7457679e26d Simplified last proof.
nipkow
parents: 5132
diff changeset
    93
by (subgoal_tac "xs1=us" 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    94
 by (Asm_full_simp_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    95
by (Asm_full_simp_tac 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    96
by (full_simp_tac (simpset() addsimps [is_maxpref_def]) 1);
24f992a25adc isatool expandshort;
wenzelm
parents: 5118
diff changeset
    97
by (blast_tac (claset() addIs [prefix_append RS iffD2,prefix_antisym]) 1);
4714
bcdf68c78e18 New scanner in abstract form.
nipkow
parents:
diff changeset
    98
qed "is_maxchopper_chop";