| author | wenzelm | 
| Tue, 13 Jun 2000 18:33:34 +0200 | |
| changeset 9058 | 7856a01119fb | 
| parent 8442 | 96023903c2df | 
| child 9111 | 33b32680669a | 
| permissions | -rw-r--r-- | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 1 | (* Title: HOL/UNITY/GenPrefix.thy | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 4 | Copyright 1999 University of Cambridge | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 5 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 6 | Charpentier's Generalized Prefix Relation | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 7 | (xs,ys) : genPrefix r | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 8 | if ys = xs' @ zs where length xs = length xs' | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 9 | and corresponding elements of xs, xs' are pairwise related by r | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 10 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 11 | Based on Lex/Prefix | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 12 | *) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 13 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 14 | (*** preliminary lemmas ***) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 15 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 16 | Goal "([], xs) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 17 | by (cut_facts_tac [genPrefix.Nil RS genPrefix.append] 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 18 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 19 | qed "Nil_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 20 | AddIffs[Nil_genPrefix]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 21 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 22 | Goal "(xs,ys) : genPrefix r ==> length xs <= length ys"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 23 | by (etac genPrefix.induct 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 24 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 25 | qed "genPrefix_length_le"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 26 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 27 | Goal "[| (xs', ys'): genPrefix r |] \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 28 | \ ==> (ALL x xs. xs' = x#xs --> (EX y ys. ys' = y#ys & (x,y) : r & (xs, ys) : genPrefix r))"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 29 | by (etac genPrefix.induct 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 30 | by (Blast_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 31 | by (Blast_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 32 | by (force_tac (claset() addIs [genPrefix.append], | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 33 | simpset()) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 34 | val lemma = result(); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 35 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 36 | (*As usual converting it to an elimination rule is tiresome*) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 37 | val major::prems = | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 38 | Goal "[| (x#xs, zs): genPrefix r; \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 39 | \ !!y ys. [| zs = y#ys; (x,y) : r; (xs, ys) : genPrefix r |] ==> P \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 40 | \ |] ==> P"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 41 | by (cut_facts_tac [major RS lemma] 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 42 | by (Full_simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 43 | by (REPEAT (eresolve_tac [exE, conjE] 1)); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 44 | by (REPEAT (ares_tac prems 1)); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 45 | qed "cons_genPrefixE"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 46 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 47 | AddSEs [cons_genPrefixE]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 48 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 49 | Goal "((x#xs,y#ys) : genPrefix r) = ((x,y) : r & (xs,ys) : genPrefix r)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 50 | by (blast_tac (claset() addIs [genPrefix.prepend]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 51 | qed"Cons_genPrefix_Cons"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 52 | AddIffs [Cons_genPrefix_Cons]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 53 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 54 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 55 | (*** genPrefix is a partial order ***) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 56 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 57 | Goalw [refl_def] "reflexive r ==> reflexive (genPrefix r)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 58 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 59 | by (induct_tac "x" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 60 | by (blast_tac (claset() addIs [genPrefix.prepend]) 2); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 61 | by (blast_tac (claset() addIs [genPrefix.Nil]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 62 | qed "refl_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 63 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 64 | Goal "reflexive r ==> (l,l) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 65 | by (etac ([refl_genPrefix,UNIV_I] MRS reflD) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 66 | qed "genPrefix_refl"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 67 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 68 | Addsimps [genPrefix_refl]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 69 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 70 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 71 | (** Transitivity **) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 72 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 73 | (*Merely a lemma for proving transitivity*) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 74 | Goal "ALL zs. (xs @ ys, zs) : genPrefix r --> (xs, zs) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 75 | by (induct_tac "xs" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 76 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 77 | qed_spec_mp "append_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 78 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 79 | Goalw [trans_def] | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 80 | "[| (x, y) : genPrefix r; trans r |] \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 81 | \ ==> ALL z. (y, z) : genPrefix r --> (x, z) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 82 | by (etac genPrefix.induct 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 83 | by (blast_tac (claset() addDs [append_genPrefix]) 3); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 84 | by (blast_tac (claset() addIs [genPrefix.prepend]) 2); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 85 | by (Blast_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 86 | qed_spec_mp "genPrefix_trans"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 87 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 88 | Goal "trans r ==> trans (genPrefix r)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 89 | by (blast_tac (claset() addIs [transI, genPrefix_trans]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 90 | qed "trans_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 91 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 92 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 93 | (** Antisymmetry **) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 94 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 95 | Goal "[| (xs,ys) : genPrefix r; antisym r |] \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 96 | \ ==> (ys,xs) : genPrefix r --> xs = ys"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 97 | by (etac genPrefix.induct 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 98 | by (full_simp_tac (simpset() addsimps [antisym_def]) 2); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 99 | by (Blast_tac 2); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 100 | by (Blast_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 101 | (*append case is hardest*) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 102 | by (Clarify_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 103 | by (subgoal_tac "length zs = 0" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 104 | by (Force_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 105 | by (REPEAT (dtac genPrefix_length_le 1)); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 106 | by (full_simp_tac (simpset() delsimps [length_0_conv]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 107 | qed_spec_mp "genPrefix_antisym"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 108 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 109 | Goal "antisym r ==> antisym (genPrefix r)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 110 | by (blast_tac (claset() addIs [antisymI, genPrefix_antisym]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 111 | qed "antisym_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 112 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 113 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 114 | (*** recursion equations ***) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 115 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 116 | Goal "((xs, []) : genPrefix r) = (xs = [])"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 117 | by (induct_tac "xs" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 118 | by (Blast_tac 2); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 119 | by (Simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 120 | qed "genPrefix_Nil"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 121 | Addsimps [genPrefix_Nil]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 122 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 123 | Goalw [refl_def] | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 124 | "reflexive r ==> ((xs@ys, xs@zs) : genPrefix r) = ((ys,zs) : genPrefix r)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 125 | by (induct_tac "xs" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 126 | by (ALLGOALS Asm_simp_tac); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 127 | qed "same_genPrefix_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 128 | Addsimps [same_genPrefix_genPrefix]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 129 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 130 | Goal "((xs, y#ys) : genPrefix r) = \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 131 | \ (xs=[] | (EX z zs. xs=z#zs & (z,y) : r & (zs,ys) : genPrefix r))"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 132 | by (case_tac "xs" 1); | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 133 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 134 | qed "genPrefix_Cons"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 135 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 136 | Goal "[| reflexive r; (xs,ys) : genPrefix r |] \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 137 | \ ==> (xs@zs, take (length xs) ys @ zs) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 138 | by (etac genPrefix.induct 1); | 
| 7499 | 139 | by (ftac genPrefix_length_le 3); | 
| 7053 
8f246bc87ab2
tweaked proof after removal of diff_is_0_eq RS iffD2
 paulson parents: 
6824diff
changeset | 140 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_is_0_eq RS iffD2]))); | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 141 | qed "genPrefix_take_append"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 142 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 143 | Goal "[| reflexive r; (xs,ys) : genPrefix r; length xs = length ys |] \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 144 | \ ==> (xs@zs, ys @ zs) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 145 | by (dtac genPrefix_take_append 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 146 | by (assume_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 147 | by (asm_full_simp_tac (simpset() addsimps [take_all]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 148 | qed "genPrefix_append_both"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 149 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 150 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 151 | (*NOT suitable for rewriting since [y] has the form y#ys*) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 152 | Goal "xs @ y # ys = (xs @ [y]) @ ys"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 153 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 154 | qed "append_cons_eq"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 155 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 156 | Goal "[| (xs,ys) : genPrefix r; reflexive r |] \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 157 | \ ==> length xs < length ys --> (xs @ [ys ! length xs], ys) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 158 | by (etac genPrefix.induct 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 159 | by (Blast_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 160 | by (Asm_simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 161 | (*Append case is hardest*) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 162 | by (Asm_simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 163 | by (forward_tac [genPrefix_length_le RS le_imp_less_or_eq] 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 164 | by (etac disjE 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 165 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [neq_Nil_conv, nth_append]))); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 166 | by (blast_tac (claset() addIs [genPrefix.append]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 167 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 168 | by (stac append_cons_eq 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 169 | by (blast_tac (claset() addIs [genPrefix_append_both, genPrefix.append]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 170 | val lemma = result() RS mp; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 171 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 172 | Goal "[| (xs,ys) : genPrefix r; length xs < length ys; reflexive r |] \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 173 | \ ==> (xs @ [ys ! length xs], ys) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 174 | by (blast_tac (claset() addIs [lemma]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 175 | qed "append_one_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 176 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 177 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 178 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 179 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 180 | (** Proving the equivalence with Charpentier's definition **) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 181 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 182 | Goal "ALL i ys. i < length xs \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 183 | \ --> (xs, ys) : genPrefix r --> (xs ! i, ys ! i) : r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 184 | by (induct_tac "xs" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 185 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 186 | by (case_tac "i" 1); | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 187 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 188 | qed_spec_mp "genPrefix_imp_nth"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 189 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 190 | Goal "ALL ys. length xs <= length ys \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 191 | \ --> (ALL i. i < length xs --> (xs ! i, ys ! i) : r) \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 192 | \ --> (xs, ys) : genPrefix r"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 193 | by (induct_tac "xs" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 194 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq_0_disj, | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 195 | all_conj_distrib]))); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 196 | by (Clarify_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 197 | by (case_tac "ys" 1); | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 198 | by (ALLGOALS Force_tac); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 199 | qed_spec_mp "nth_imp_genPrefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 200 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 201 | Goal "((xs,ys) : genPrefix r) = \ | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 202 | \ (length xs <= length ys & (ALL i. i < length xs --> (xs!i, ys!i) : r))"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 203 | by (blast_tac (claset() addIs [genPrefix_length_le, genPrefix_imp_nth, | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 204 | nth_imp_genPrefix]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 205 | qed "genPrefix_iff_nth"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 206 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 207 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 208 | (** <= is a partial order: **) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 209 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 210 | AddIffs [reflexive_Id, antisym_Id, trans_Id]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 211 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 212 | Goalw [prefix_def] "xs <= (xs::'a list)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 213 | by (Simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 214 | qed "prefix_refl"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 215 | AddIffs[prefix_refl]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 216 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 217 | Goalw [prefix_def] "!!xs::'a list. [| xs <= ys; ys <= zs |] ==> xs <= zs"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 218 | by (blast_tac (claset() addIs [genPrefix_trans]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 219 | qed "prefix_trans"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 220 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 221 | Goalw [prefix_def] "!!xs::'a list. [| xs <= ys; ys <= xs |] ==> xs = ys"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 222 | by (blast_tac (claset() addIs [genPrefix_antisym]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 223 | qed "prefix_antisym"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 224 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 225 | Goalw [strict_prefix_def] "!!xs::'a list. (xs < zs) = (xs <= zs & xs ~= zs)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 226 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 227 | qed "prefix_less_le"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 228 | |
| 7839 | 229 | (*Monotonicity of "set" operator WRT prefix*) | 
| 230 | Goalw [prefix_def] "xs <= ys ==> set xs <= set ys"; | |
| 231 | by (etac genPrefix.induct 1); | |
| 232 | by Auto_tac; | |
| 233 | qed "set_mono"; | |
| 234 | ||
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 235 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 236 | (** recursion equations **) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 237 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 238 | Goalw [prefix_def] "[] <= xs"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 239 | by (simp_tac (simpset() addsimps [Nil_genPrefix]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 240 | qed "Nil_prefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 241 | AddIffs[Nil_prefix]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 242 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 243 | Goalw [prefix_def] "(xs <= []) = (xs = [])"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 244 | by (simp_tac (simpset() addsimps [genPrefix_Nil]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 245 | qed "prefix_Nil"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 246 | Addsimps [prefix_Nil]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 247 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 248 | Goalw [prefix_def] "(x#xs <= y#ys) = (x=y & xs<=ys)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 249 | by (Simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 250 | qed"Cons_prefix_Cons"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 251 | Addsimps [Cons_prefix_Cons]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 252 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 253 | Goalw [prefix_def] "(xs@ys <= xs@zs) = (ys <= zs)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 254 | by (Simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 255 | qed "same_prefix_prefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 256 | Addsimps [same_prefix_prefix]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 257 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 258 | AddIffs (* (xs@ys <= xs) = (ys <= []) *) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 259 |  [simplify (simpset()) (read_instantiate [("zs","[]")] same_prefix_prefix)];
 | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 260 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 261 | Goalw [prefix_def] "xs <= ys ==> xs <= ys@zs"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 262 | by (etac genPrefix.append 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 263 | qed "prefix_appendI"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 264 | Addsimps [prefix_appendI]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 265 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 266 | Goalw [prefix_def] | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 267 | "(xs <= y#ys) = (xs=[] | (? zs. xs=y#zs & zs <= ys))"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 268 | by (simp_tac (simpset() addsimps [genPrefix_Cons]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 269 | by Auto_tac; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 270 | qed "prefix_Cons"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 271 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 272 | Goalw [prefix_def] | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 273 | "[| xs <= ys; length xs < length ys |] ==> xs @ [ys ! length xs] <= ys"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 274 | by (asm_simp_tac (simpset() addsimps [append_one_genPrefix]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 275 | qed "append_one_prefix"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 276 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 277 | Goalw [prefix_def] "xs <= ys ==> length xs <= length ys"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 278 | by (etac genPrefix_length_le 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 279 | qed "prefix_length_le"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 280 | |
| 8216 | 281 | Goalw [prefix_def] "xs<=ys ==> xs~=ys --> length xs < length ys"; | 
| 282 | by (etac genPrefix.induct 1); | |
| 283 | by Auto_tac; | |
| 284 | val lemma = result(); | |
| 285 | ||
| 286 | Goalw [strict_prefix_def] "xs < ys ==> length xs < length ys"; | |
| 287 | by (blast_tac (claset() addIs [lemma RS mp]) 1); | |
| 288 | qed "strict_prefix_length_less"; | |
| 289 | ||
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 290 | Goal "mono length"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 291 | by (blast_tac (claset() addIs [monoI, prefix_length_le]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 292 | qed "mono_length"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 293 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 294 | (*Equivalence to the definition used in Lex/Prefix.thy*) | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 295 | Goalw [prefix_def] "(xs <= zs) = (EX ys. zs = xs@ys)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 296 | by (auto_tac (claset(), | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 297 | simpset() addsimps [genPrefix_iff_nth, nth_append])); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 298 | by (res_inst_tac [("x", "drop (length xs) zs")] exI 1);
 | 
| 6824 | 299 | by (rtac nth_equalityI 1); | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 300 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [nth_append]))); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 301 | qed "prefix_iff"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 302 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 303 | Goal "(xs <= ys@[y]) = (xs = ys@[y] | xs <= ys)"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 304 | by (simp_tac (simpset() addsimps [prefix_iff]) 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 305 | by (rtac iffI 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 306 | by (etac exE 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 307 | by (rename_tac "zs" 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 308 |  by (res_inst_tac [("xs","zs")] rev_exhaust 1);
 | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 309 | by (Asm_full_simp_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 310 | by (hyp_subst_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 311 | by (asm_full_simp_tac (simpset() delsimps [append_assoc] | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 312 | addsimps [append_assoc RS sym])1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 313 | by (Force_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 314 | qed "prefix_snoc"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 315 | Addsimps [prefix_snoc]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 316 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 317 | Goal "(xs <= ys@zs) = (xs <= ys | (? us. xs = ys@us & us <= zs))"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 318 | by (res_inst_tac [("xs","zs")] rev_induct 1);
 | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 319 | by (Force_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 320 | by (asm_full_simp_tac (simpset() delsimps [append_assoc] | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 321 | addsimps [append_assoc RS sym])1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 322 | by (Force_tac 1); | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 323 | qed "prefix_append_iff"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 324 | |
| 8216 | 325 | (*Although the prefix ordering is not linear, the prefixes of a list | 
| 326 | are linearly ordered.*) | |
| 327 | Goal "!!zs::'a list. xs <= zs --> ys <= zs --> xs <= ys | ys <= xs"; | |
| 328 | by (rev_induct_tac "zs" 1); | |
| 329 | by Auto_tac; | |
| 330 | qed_spec_mp "common_prefix_linear"; | |
| 331 | ||
| 8128 | 332 | Goal "r<=s ==> genPrefix r <= genPrefix s"; | 
| 333 | by (Clarify_tac 1); | |
| 334 | by (etac genPrefix.induct 1); | |
| 335 | by (auto_tac (claset() addIs [genPrefix.append], simpset())); | |
| 336 | qed "genPrefix_mono"; | |
| 337 | ||
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 338 | |
| 6824 | 339 | (*** pfixLe, pfixGe: properties inherited from the translations ***) | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 340 | |
| 8128 | 341 | (** pfixLe **) | 
| 342 | ||
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 343 | Goalw [refl_def, Le_def] "reflexive Le"; | 
| 6824 | 344 | by Auto_tac; | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 345 | qed "reflexive_Le"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 346 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 347 | Goalw [antisym_def, Le_def] "antisym Le"; | 
| 6824 | 348 | by Auto_tac; | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 349 | qed "antisym_Le"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 350 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 351 | Goalw [trans_def, Le_def] "trans Le"; | 
| 6824 | 352 | by Auto_tac; | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 353 | qed "trans_Le"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 354 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 355 | AddIffs [reflexive_Le, antisym_Le, trans_Le]; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 356 | |
| 8128 | 357 | Goal "x pfixLe x"; | 
| 358 | by (Simp_tac 1); | |
| 359 | qed "pfixLe_refl"; | |
| 360 | AddIffs[pfixLe_refl]; | |
| 361 | ||
| 362 | Goal "[| x pfixLe y; y pfixLe z |] ==> x pfixLe z"; | |
| 363 | by (blast_tac (claset() addIs [genPrefix_trans]) 1); | |
| 364 | qed "pfixLe_trans"; | |
| 365 | ||
| 366 | Goal "[| x pfixLe y; y pfixLe x |] ==> x = y"; | |
| 367 | by (blast_tac (claset() addIs [genPrefix_antisym]) 1); | |
| 368 | qed "pfixLe_antisym"; | |
| 369 | ||
| 370 | Goalw [prefix_def, Le_def] "xs<=ys ==> xs pfixLe ys"; | |
| 371 | by (blast_tac (claset() addIs [impOfSubs genPrefix_mono]) 1); | |
| 372 | qed "prefix_imp_pfixLe"; | |
| 373 | ||
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 374 | Goalw [refl_def, Ge_def] "reflexive Ge"; | 
| 6824 | 375 | by Auto_tac; | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 376 | qed "reflexive_Ge"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 377 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 378 | Goalw [antisym_def, Ge_def] "antisym Ge"; | 
| 6824 | 379 | by Auto_tac; | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 380 | qed "antisym_Ge"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 381 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 382 | Goalw [trans_def, Ge_def] "trans Ge"; | 
| 6824 | 383 | by Auto_tac; | 
| 6804 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 384 | qed "trans_Ge"; | 
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 385 | |
| 
66dc8e62a987
Generalized prefix theory, replacing the reference to directory Lex.
 paulson parents: diff
changeset | 386 | AddIffs [reflexive_Ge, antisym_Ge, trans_Ge]; | 
| 8067 | 387 | |
| 8128 | 388 | Goal "x pfixGe x"; | 
| 389 | by (Simp_tac 1); | |
| 390 | qed "pfixGe_refl"; | |
| 391 | AddIffs[pfixGe_refl]; | |
| 8067 | 392 | |
| 8128 | 393 | Goal "[| x pfixGe y; y pfixGe z |] ==> x pfixGe z"; | 
| 394 | by (blast_tac (claset() addIs [genPrefix_trans]) 1); | |
| 395 | qed "pfixGe_trans"; | |
| 396 | ||
| 397 | Goal "[| x pfixGe y; y pfixGe x |] ==> x = y"; | |
| 398 | by (blast_tac (claset() addIs [genPrefix_antisym]) 1); | |
| 399 | qed "pfixGe_antisym"; | |
| 8067 | 400 | |
| 401 | Goalw [prefix_def, Ge_def] "xs<=ys ==> xs pfixGe ys"; | |
| 402 | by (blast_tac (claset() addIs [impOfSubs genPrefix_mono]) 1); | |
| 403 | qed "prefix_imp_pfixGe"; | |
| 404 |