src/HOL/Lambda/ParRed.ML
author paulson
Thu, 26 Sep 1996 12:47:47 +0200
changeset 2031 03a843f0f447
parent 1974 b50f96543dec
child 2057 4d7a4b25a11f
permissions -rw-r--r--
Ran expandshort
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
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
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    20
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    21
(*** beta <= par_beta <= beta^* ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    22
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    23
goal ParRed.thy "(Var n => t) = (t = Var n)";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    24
by (Fast_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    25
qed "par_beta_varL";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    26
Addsimps [par_beta_varL];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    27
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    28
goal ParRed.thy "t => t";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    29
by (db.induct_tac "t" 1);
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    30
by (ALLGOALS Asm_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    31
qed"par_beta_refl";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    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
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    34
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    35
goal ParRed.thy "beta <= par_beta";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1431
diff changeset
    36
by (rtac subsetI 1);
1431
be7c6d77e19c Polished proofs.
nipkow
parents: 1288
diff changeset
    37
by (split_all_tac 1);
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    38
by (etac beta.induct 1);
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    39
by (ALLGOALS(fast_tac (!claset addSIs [par_beta_refl])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    40
qed "beta_subset_par_beta";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    41
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    42
goal ParRed.thy "par_beta <= beta^*";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1431
diff changeset
    43
by (rtac subsetI 1);
1431
be7c6d77e19c Polished proofs.
nipkow
parents: 1288
diff changeset
    44
by (split_all_tac 1);
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    45
by (etac par_beta.induct 1);
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    46
by (ALLGOALS (slow_best_tac (!claset addIs [rtrancl_into_rtrancl] 
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    47
                                     addEs [rtrancl_trans])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    48
qed "par_beta_subset_beta";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    49
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    50
(*** => ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    51
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1156
diff changeset
    52
goal ParRed.thy "!t' n. t => t' --> lift t n => lift t' n";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    53
by (db.induct_tac "t" 1);
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    54
by (ALLGOALS(fast_tac (!claset addss (!simpset))));
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    55
qed_spec_mp "par_beta_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    56
Addsimps [par_beta_lift];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    57
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    58
goal ParRed.thy
1124
a6233ea105a4 Polished the presentation making it completely definitional.
nipkow
parents: 1120
diff changeset
    59
  "!s s' t' n. s => s' --> t => t' --> t[s/n] => t'[s'/n]";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    60
by (db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    61
  by(asm_simp_tac (addsplit(!simpset)) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    62
 by(strip_tac 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    63
 bes par_beta_cases 1;
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    64
  by(Asm_simp_tac 1);
1288
6eb89a693e05 Improved a few proofs.
nipkow
parents: 1269
diff changeset
    65
 by(asm_simp_tac (!simpset addsimps [subst_subst RS sym]) 1);
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    66
 by(fast_tac (!claset addSIs [par_beta_lift] addss (!simpset)) 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    67
by (fast_tac (!claset addss (!simpset)) 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    68
qed_spec_mp "par_beta_subst";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    69
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    70
(*** Confluence (directly) ***)
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    71
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    72
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
    73
by (rtac (impI RS allI RS allI) 1);
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    74
by (etac par_beta.induct 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    75
by (ALLGOALS(fast_tac (!claset addSIs [par_beta_subst])));
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    76
qed "diamond_par_beta";
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    77
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
(*** cd ***)
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    80
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    81
goal ParRed.thy "!t. s => t --> t => cd s";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    82
by (db.induct_tac "s" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    83
  by(Simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    84
 be rev_mp 1;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    85
 by(db.induct_tac "db1" 1);
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    86
 by(ALLGOALS(fast_tac (!claset addSIs [par_beta_subst] addss !simpset)));
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    87
qed_spec_mp "par_beta_cd";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    88
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    89
(*** Confluence (via cd) ***)
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    90
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    91
goalw ParRed.thy [diamond_def,commute_def,square_def] "diamond(par_beta)";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    92
by (fast_tac (HOL_cs addIs [par_beta_cd]) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    93
qed "diamond_par_beta2";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    94
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    95
goal ParRed.thy "confluent(beta)";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    96
by (fast_tac (HOL_cs addIs [diamond_par_beta2,diamond_to_confluence,
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    97
                           par_beta_subset_beta,beta_subset_par_beta]) 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    98
qed"beta_confluent";