src/HOL/Lambda/ParRed.ML
author paulson
Tue, 07 May 1996 18:19:13 +0200
changeset 1730 1c7f793fc374
parent 1486 7b95d7b49f7a
child 1974 b50f96543dec
permissions -rw-r--r--
Updated for new form of induction rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1288
6eb89a693e05 Improved a few proofs.
nipkow
parents: 1269
diff changeset
     1
(*  Title:      HOL/Lambda/ParRed.ML
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     4
    Copyright   1995 TU Muenchen
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     5
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     6
Properties of => and cd, in particular the diamond property of => and
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     7
confluence of beta.
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     8
*)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     9
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    10
open ParRed;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    11
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    12
Addsimps par_beta.intrs;
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    13
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    14
val par_beta_cases = map (par_beta.mk_cases db.simps)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    15
    ["Var n => t", "Fun s => Fun t",
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    16
     "(Fun s) @ t => u", "s @ t => u", "Fun s => t"];
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    17
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    18
val parred_cs = lambda_cs addSIs par_beta.intrs addSEs par_beta_cases;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    19
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    20
(*** beta <= par_beta <= beta^* ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    21
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    22
goal ParRed.thy "(Var n => t) = (t = Var n)";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    23
by(fast_tac parred_cs 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    24
qed "par_beta_varL";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    25
Addsimps [par_beta_varL];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    26
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    27
goal ParRed.thy "t => t";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    28
by(db.induct_tac "t" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    29
by(ALLGOALS Asm_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    30
qed"par_beta_refl";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    31
Addsimps [par_beta_refl];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    32
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    33
goal ParRed.thy "beta <= par_beta";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1431
diff changeset
    34
by (rtac subsetI 1);
1431
be7c6d77e19c Polished proofs.
nipkow
parents: 1288
diff changeset
    35
by (split_all_tac 1);
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    36
by (etac beta.induct 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    37
by (ALLGOALS(fast_tac (parred_cs addSIs [par_beta_refl])));
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    38
qed "beta_subset_par_beta";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    39
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    40
goal ParRed.thy "par_beta <= beta^*";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1431
diff changeset
    41
by (rtac subsetI 1);
1431
be7c6d77e19c Polished proofs.
nipkow
parents: 1288
diff changeset
    42
by (split_all_tac 1);
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    43
by (etac par_beta.induct 1);
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    44
by (ALLGOALS
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    45
    (fast_tac (parred_cs addIs
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    46
	                  [rtrancl_beta_Fun,rtrancl_beta_App,rtrancl_refl,
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    47
			   rtrancl_into_rtrancl] 
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    48
                         addEs [rtrancl_trans])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    49
qed "par_beta_subset_beta";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    50
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    51
(*** => ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    52
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1156
diff changeset
    53
goal ParRed.thy "!t' n. t => t' --> lift t n => lift t' n";
ab725b698cb2 Renamed some vars.
nipkow
parents: 1156
diff changeset
    54
by(db.induct_tac "t" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    55
by(ALLGOALS(fast_tac (parred_cs addss (!simpset))));
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    56
qed_spec_mp "par_beta_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    57
Addsimps [par_beta_lift];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    58
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    59
goal ParRed.thy
1124
a6233ea105a4 Polished the presentation making it completely definitional.
nipkow
parents: 1120
diff changeset
    60
  "!s s' t' n. s => s' --> t => t' --> t[s/n] => t'[s'/n]";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    61
by(db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    62
  by(asm_simp_tac (addsplit(!simpset)) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    63
 by(strip_tac 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    64
 bes par_beta_cases 1;
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    65
  by(Asm_simp_tac 1);
1288
6eb89a693e05 Improved a few proofs.
nipkow
parents: 1269
diff changeset
    66
 by(asm_simp_tac (!simpset addsimps [subst_subst RS sym]) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    67
 by(fast_tac (parred_cs addSIs [par_beta_lift] addss (!simpset)) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    68
by(fast_tac (parred_cs addss (!simpset)) 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    69
qed_spec_mp "par_beta_subst";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    70
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    71
(*** Confluence (directly) ***)
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    72
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    73
goalw ParRed.thy [diamond_def,commute_def,square_def] "diamond(par_beta)";
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    74
by (rtac (impI RS allI RS allI) 1);
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    75
by (etac par_beta.induct 1);
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    76
by(ALLGOALS(fast_tac (parred_cs addSIs [par_beta_subst])));
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    77
qed "diamond_par_beta";
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    78
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    79
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    80
(*** cd ***)
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    81
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    82
goal ParRed.thy "!t. s => t --> t => cd s";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    83
by(db.induct_tac "s" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    84
  by(Simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    85
 be rev_mp 1;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    86
 by(db.induct_tac "db1" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    87
 by(ALLGOALS(fast_tac (parred_cs addSIs [par_beta_subst] addss (!simpset))));
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    88
qed_spec_mp "par_beta_cd";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    89
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    90
(*** Confluence (via cd) ***)
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    91
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    92
goalw ParRed.thy [diamond_def,commute_def,square_def] "diamond(par_beta)";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    93
by(fast_tac (HOL_cs addIs [par_beta_cd]) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    94
qed "diamond_par_beta2";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    95
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    96
goal ParRed.thy "confluent(beta)";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    97
by(fast_tac (HOL_cs addIs [diamond_par_beta2,diamond_to_confluence,
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    98
                           par_beta_subset_beta,beta_subset_par_beta]) 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    99
qed"beta_confluent";