author | wenzelm |
Tue, 27 May 1997 15:45:07 +0200 | |
changeset 3362 | 0b268cff9344 |
parent 3207 | fe79ad367d77 |
child 4089 | 96fba19bcbe2 |
permissions | -rw-r--r-- |
1288 | 1 |
(* Title: HOL/Lambda/ParRed.ML |
1120 | 2 |
ID: $Id$ |
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1995 TU Muenchen |
|
5 |
||
6 |
Properties of => and cd, in particular the diamond property of => and |
|
7 |
confluence of beta. |
|
8 |
*) |
|
9 |
||
10 |
open ParRed; |
|
11 |
||
1266 | 12 |
Addsimps par_beta.intrs; |
1120 | 13 |
|
2159
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
14 |
val par_beta_cases = map (par_beta.mk_cases dB.simps) |
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
15 |
["Var n => t", "Abs s => Abs t", |
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
16 |
"(Abs s) @ t => u", "s @ t => u", "Abs s => t"]; |
1120 | 17 |
|
1974
b50f96543dec
Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents:
1730
diff
changeset
|
18 |
AddSIs par_beta.intrs; |
b50f96543dec
Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents:
1730
diff
changeset
|
19 |
AddSEs par_beta_cases; |
1120 | 20 |
|
21 |
(*** beta <= par_beta <= beta^* ***) |
|
22 |
||
23 |
goal ParRed.thy "(Var n => t) = (t = Var n)"; |
|
2891 | 24 |
by (Blast_tac 1); |
1120 | 25 |
qed "par_beta_varL"; |
1266 | 26 |
Addsimps [par_beta_varL]; |
1120 | 27 |
|
28 |
goal ParRed.thy "t => t"; |
|
2159
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
29 |
by (dB.induct_tac "t" 1); |
2031 | 30 |
by (ALLGOALS Asm_simp_tac); |
1120 | 31 |
qed"par_beta_refl"; |
1266 | 32 |
Addsimps [par_beta_refl]; |
1974
b50f96543dec
Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents:
1730
diff
changeset
|
33 |
(* AddSIs [par_beta_refl]; causes search to blow up *) |
1120 | 34 |
|
35 |
goal ParRed.thy "beta <= par_beta"; |
|
1465 | 36 |
by (rtac subsetI 1); |
1431 | 37 |
by (split_all_tac 1); |
1730 | 38 |
by (etac beta.induct 1); |
2891 | 39 |
by (ALLGOALS(blast_tac (!claset addSIs [par_beta_refl]))); |
1120 | 40 |
qed "beta_subset_par_beta"; |
41 |
||
42 |
goal ParRed.thy "par_beta <= beta^*"; |
|
1465 | 43 |
by (rtac subsetI 1); |
1431 | 44 |
by (split_all_tac 1); |
1730 | 45 |
by (etac par_beta.induct 1); |
2922 | 46 |
by (Blast_tac 1); |
47 |
(* rtrancl_refl complicates the proof by increasing the branching factor*) |
|
48 |
by (ALLGOALS (blast_tac (!claset delrules [rtrancl_refl] |
|
49 |
addIs [rtrancl_into_rtrancl]))); |
|
1120 | 50 |
qed "par_beta_subset_beta"; |
51 |
||
52 |
(*** => ***) |
|
53 |
||
1172 | 54 |
goal ParRed.thy "!t' n. t => t' --> lift t n => lift t' n"; |
2159
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
55 |
by (dB.induct_tac "t" 1); |
2031 | 56 |
by (ALLGOALS(fast_tac (!claset addss (!simpset)))); |
1486 | 57 |
qed_spec_mp "par_beta_lift"; |
1266 | 58 |
Addsimps [par_beta_lift]; |
1120 | 59 |
|
60 |
goal ParRed.thy |
|
1124
a6233ea105a4
Polished the presentation making it completely definitional.
nipkow
parents:
1120
diff
changeset
|
61 |
"!s s' t' n. s => s' --> t => t' --> t[s/n] => t'[s'/n]"; |
2159
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
62 |
by (dB.induct_tac "t" 1); |
2057 | 63 |
by (asm_simp_tac (addsplit(!simpset)) 1); |
64 |
by (strip_tac 1); |
|
65 |
by (eresolve_tac par_beta_cases 1); |
|
66 |
by (Asm_simp_tac 1); |
|
67 |
by (asm_simp_tac (!simpset addsimps [subst_subst RS sym]) 1); |
|
68 |
by (fast_tac (!claset addSIs [par_beta_lift] addss (!simpset)) 1); |
|
2031 | 69 |
by (fast_tac (!claset addss (!simpset)) 1); |
1486 | 70 |
qed_spec_mp "par_beta_subst"; |
1120 | 71 |
|
1156 | 72 |
(*** Confluence (directly) ***) |
73 |
||
1269 | 74 |
goalw ParRed.thy [diamond_def,commute_def,square_def] "diamond(par_beta)"; |
1730 | 75 |
by (rtac (impI RS allI RS allI) 1); |
76 |
by (etac par_beta.induct 1); |
|
2922 | 77 |
by (ALLGOALS(blast_tac (!claset addSIs [par_beta_subst]))); |
1156 | 78 |
qed "diamond_par_beta"; |
79 |
||
80 |
||
81 |
(*** cd ***) |
|
82 |
||
1120 | 83 |
goal ParRed.thy "!t. s => t --> t => cd s"; |
2159
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
84 |
by (dB.induct_tac "s" 1); |
2057 | 85 |
by (Simp_tac 1); |
86 |
by (etac rev_mp 1); |
|
2159
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents:
2057
diff
changeset
|
87 |
by (dB.induct_tac "dB1" 1); |
2891 | 88 |
by (ALLGOALS(fast_tac (!claset addSIs [par_beta_subst] |
3207 | 89 |
addss (!simpset)))); |
1486 | 90 |
qed_spec_mp "par_beta_cd"; |
1120 | 91 |
|
1156 | 92 |
(*** Confluence (via cd) ***) |
1120 | 93 |
|
1269 | 94 |
goalw ParRed.thy [diamond_def,commute_def,square_def] "diamond(par_beta)"; |
2922 | 95 |
by (blast_tac (!claset addIs [par_beta_cd]) 1); |
1266 | 96 |
qed "diamond_par_beta2"; |
1120 | 97 |
|
98 |
goal ParRed.thy "confluent(beta)"; |
|
2891 | 99 |
by (blast_tac (HOL_cs addIs [diamond_par_beta2, diamond_to_confluence, |
100 |
par_beta_subset_beta, beta_subset_par_beta]) 1); |
|
1120 | 101 |
qed"beta_confluent"; |