author | nipkow |
Thu, 13 Apr 2000 15:01:50 +0200 | |
changeset 8703 | 816d8f6513be |
parent 8442 | 96023903c2df |
permissions | -rw-r--r-- |
5261 | 1 |
(* Title: HOL/Lambda/ListOrder.ML |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1998 TU Muenchen |
|
5 |
*) |
|
6 |
||
7 |
Goalw [step1_def] "step1(r^-1) = (step1 r)^-1"; |
|
5318 | 8 |
by (Blast_tac 1); |
5261 | 9 |
qed "step1_converse"; |
10 |
Addsimps [step1_converse]; |
|
11 |
||
12 |
Goal "(p : step1(r^-1)) = (p : (step1 r)^-1)"; |
|
5318 | 13 |
by (Auto_tac); |
5261 | 14 |
qed "in_step1_converse"; |
15 |
AddIffs [in_step1_converse]; |
|
16 |
||
17 |
Goalw [step1_def] "([],xs) ~: step1 r"; |
|
5318 | 18 |
by (Blast_tac 1); |
5261 | 19 |
qed "not_Nil_step1"; |
20 |
AddIffs [not_Nil_step1]; |
|
21 |
||
22 |
Goalw [step1_def] "(xs,[]) ~: step1 r"; |
|
5318 | 23 |
by (Blast_tac 1); |
5261 | 24 |
qed "not_step1_Nil"; |
25 |
AddIffs [not_step1_Nil]; |
|
26 |
||
27 |
Goalw [step1_def] |
|
28 |
"((y#ys,x#xs) : step1 r) = ((y,x):r & xs=ys | x=y & (ys,xs) : step1 r)"; |
|
5318 | 29 |
by (Asm_full_simp_tac 1); |
30 |
by (rtac iffI 1); |
|
31 |
by (etac exE 1); |
|
32 |
by (rename_tac "ts" 1); |
|
8442
96023903c2df
case_tac now subsumes both boolean and datatype cases;
wenzelm
parents:
8423
diff
changeset
|
33 |
by (case_tac "ts" 1); |
5318 | 34 |
by (Force_tac 1); |
35 |
by (Force_tac 1); |
|
36 |
by (etac disjE 1); |
|
37 |
by (Blast_tac 1); |
|
38 |
by (blast_tac (claset() addIs [Cons_eq_appendI]) 1); |
|
5261 | 39 |
qed "Cons_step1_Cons"; |
40 |
AddIffs [Cons_step1_Cons]; |
|
41 |
||
42 |
Goalw [step1_def] |
|
43 |
"(ys,xs) : step1 r & vs=us | ys=xs & (vs,us) : step1 r \ |
|
44 |
\ ==> (ys@vs,xs@us) : step1 r"; |
|
5318 | 45 |
by (Auto_tac); |
5758
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5525
diff
changeset
|
46 |
by (Blast_tac 1); |
5318 | 47 |
by (blast_tac (claset() addIs [append_eq_appendI]) 1); |
5261 | 48 |
qed "append_step1I"; |
49 |
||
50 |
Goal "[| (ys,x#xs) : step1 r; \ |
|
51 |
\ !y. ys = y#xs --> (y,x) : r --> R; \ |
|
52 |
\ !zs. ys = x#zs --> (zs,xs) : step1 r --> R \ |
|
53 |
\ |] ==> R"; |
|
8442
96023903c2df
case_tac now subsumes both boolean and datatype cases;
wenzelm
parents:
8423
diff
changeset
|
54 |
by (case_tac "ys" 1); |
5318 | 55 |
by (asm_full_simp_tac (simpset() addsimps [step1_def]) 1); |
56 |
by (Blast_tac 1); |
|
5261 | 57 |
val lemma = result(); |
58 |
bind_thm("Cons_step1E", |
|
59 |
impI RSN (3,impI RSN (3,allI RSN (3,impI RSN (2, |
|
60 |
impI RSN (2,allI RSN (2,lemma))))))); |
|
61 |
AddSEs [Cons_step1E]; |
|
62 |
||
63 |
Goalw [step1_def] |
|
64 |
"(ys@[y],xs@[x]) : step1 r ==> ((ys,xs) : step1 r & y=x | ys=xs & (y,x) : r)"; |
|
5318 | 65 |
by (Asm_full_simp_tac 1); |
66 |
by (clarify_tac (claset() delrules [disjCI]) 1); |
|
67 |
by (rename_tac "vs" 1); |
|
68 |
by (res_inst_tac [("xs","vs")]rev_exhaust 1); |
|
69 |
by (Force_tac 1); |
|
70 |
by (Asm_full_simp_tac 1); |
|
71 |
by (Blast_tac 1); |
|
5261 | 72 |
qed "Snoc_step1_SnocD"; |
73 |
||
74 |
Goal "x : acc r ==> !xs. xs : acc(step1 r) --> x#xs : acc(step1 r)"; |
|
5318 | 75 |
by (etac acc_induct 1); |
76 |
by (etac thin_rl 1); |
|
77 |
by (Clarify_tac 1); |
|
78 |
by (etac acc_induct 1); |
|
79 |
by (rtac accI 1); |
|
80 |
by (Blast_tac 1); |
|
5261 | 81 |
val lemma = result(); |
82 |
qed_spec_mp "Cons_acc_step1I"; |
|
83 |
AddSIs [Cons_acc_step1I]; |
|
84 |
||
85 |
Goal "xs : lists(acc r) ==> xs : acc(step1 r)"; |
|
5326 | 86 |
by (etac lists.induct 1); |
5318 | 87 |
by (rtac accI 1); |
88 |
by (asm_full_simp_tac (simpset() addsimps [step1_def]) 1); |
|
89 |
by (rtac accI 1); |
|
90 |
by (fast_tac (claset() addDs [acc_downward]) 1); |
|
5261 | 91 |
qed "lists_accD"; |
92 |
||
93 |
Goalw [step1_def] |
|
94 |
"[| x : set xs; (y,x) : r |] ==> ? ys. (ys,xs) : step1 r & y : set ys"; |
|
5318 | 95 |
by (dtac (in_set_conv_decomp RS iffD1) 1); |
96 |
by (Force_tac 1); |
|
5261 | 97 |
qed "ex_step1I"; |
98 |
||
99 |
Goal "xs : acc(step1 r) ==> xs : lists(acc r)"; |
|
5318 | 100 |
by (etac acc_induct 1); |
101 |
by (Clarify_tac 1); |
|
102 |
by (rtac accI 1); |
|
103 |
by (EVERY1[dtac ex_step1I, atac]); |
|
104 |
by (Blast_tac 1); |
|
5261 | 105 |
qed "lists_accI"; |