src/HOL/Lambda/ParRed.ML
author paulson
Tue, 19 Jan 1999 11:18:11 +0100
changeset 6141 a6922171b396
parent 5184 9b8547a9496a
child 8624 69619f870939
permissions -rw-r--r--
removal of the (thm list) argument of mk_cases
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
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    10
Addsimps par_beta.intrs;
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    11
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5184
diff changeset
    12
val par_beta_cases = 
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5184
diff changeset
    13
    map par_beta.mk_cases
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5184
diff changeset
    14
       ["Var n => t", 
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5184
diff changeset
    15
	"Abs s => Abs t",
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5184
diff changeset
    16
	"(Abs s) $ t => u", 
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5184
diff changeset
    17
	"s $ t => u", 
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 5184
diff changeset
    18
	"Abs s => t"];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    19
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    20
AddSIs par_beta.intrs;
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    21
AddSEs par_beta_cases;
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    22
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    23
(*** beta <= par_beta <= beta^* ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    24
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    25
Goal "(Var n => t) = (t = Var n)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
    26
by (Blast_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    27
qed "par_beta_varL";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    28
Addsimps [par_beta_varL];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    29
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    30
Goal "t => t";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5168
diff changeset
    31
by (induct_tac "t" 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
    32
by (ALLGOALS Asm_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    33
qed"par_beta_refl";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    34
Addsimps [par_beta_refl];
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1730
diff changeset
    35
(* AddSIs [par_beta_refl]; causes search to blow up *)
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    36
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    37
Goal "beta <= par_beta";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1431
diff changeset
    38
by (rtac subsetI 1);
1431
be7c6d77e19c Polished proofs.
nipkow
parents: 1288
diff changeset
    39
by (split_all_tac 1);
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    40
by (etac beta.induct 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    41
by (ALLGOALS(blast_tac (claset() addSIs [par_beta_refl])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    42
qed "beta_subset_par_beta";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    43
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    44
Goal "par_beta <= beta^*";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1431
diff changeset
    45
by (rtac subsetI 1);
1431
be7c6d77e19c Polished proofs.
nipkow
parents: 1288
diff changeset
    46
by (split_all_tac 1);
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    47
by (etac par_beta.induct 1);
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
    48
by (Blast_tac 1);
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
    49
(* rtrancl_refl complicates the proof by increasing the branching factor*)
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    50
by (ALLGOALS (blast_tac (claset() delrules [rtrancl_refl]
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
    51
				 addIs [rtrancl_into_rtrancl])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    52
qed "par_beta_subset_beta";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    53
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    54
(*** => ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    55
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    56
Goal "!t' n. t => t' --> lift t n => lift t' n";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5168
diff changeset
    57
by (induct_tac "t" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    58
by (ALLGOALS(fast_tac (claset() addss (simpset()))));
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    59
qed_spec_mp "par_beta_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    60
Addsimps [par_beta_lift];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    61
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    62
Goal
1124
a6233ea105a4 Polished the presentation making it completely definitional.
nipkow
parents: 1120
diff changeset
    63
  "!s s' t' n. s => s' --> t => t' --> t[s/n] => t'[s'/n]";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5168
diff changeset
    64
by (induct_tac "t" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    65
  by (asm_simp_tac (addsplit(simpset())) 1);
2057
4d7a4b25a11f Ran expandshort
paulson
parents: 2031
diff changeset
    66
 by (strip_tac 1);
4d7a4b25a11f Ran expandshort
paulson
parents: 2031
diff changeset
    67
 by (eresolve_tac par_beta_cases 1);
4d7a4b25a11f Ran expandshort
paulson
parents: 2031
diff changeset
    68
  by (Asm_simp_tac 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    69
 by (asm_simp_tac (simpset() addsimps [subst_subst RS sym]) 1);
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    70
 by (fast_tac (claset() addSIs [par_beta_lift] addss (simpset())) 1);
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    71
by (fast_tac (claset() addss (simpset())) 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    72
qed_spec_mp "par_beta_subst";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    73
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    74
(*** Confluence (directly) ***)
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    75
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    76
Goalw [diamond_def,commute_def,square_def] "diamond(par_beta)";
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    77
by (rtac (impI RS allI RS allI) 1);
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    78
by (etac par_beta.induct 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    79
by (ALLGOALS(blast_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
    80
qed "diamond_par_beta";
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    81
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    82
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    83
(*** cd ***)
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    84
5134
51f81581e3d9 Replace awkward primrec by recdef.
nipkow
parents: 5069
diff changeset
    85
Addsimps cd.rules;
51f81581e3d9 Replace awkward primrec by recdef.
nipkow
parents: 5069
diff changeset
    86
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    87
Goal "!t. s => t --> t => cd s";
5168
adafef6eb295 isatool expandshort;
wenzelm
parents: 5146
diff changeset
    88
by (res_inst_tac[("u","s")] cd.induct 1);
adafef6eb295 isatool expandshort;
wenzelm
parents: 5146
diff changeset
    89
by (Auto_tac);
adafef6eb295 isatool expandshort;
wenzelm
parents: 5146
diff changeset
    90
by (fast_tac (claset() addSIs [par_beta_subst]) 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    91
qed_spec_mp "par_beta_cd";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    92
1156
b373cb33352f Put in direct proof of C-R w/o detour via cd.
nipkow
parents: 1124
diff changeset
    93
(*** Confluence (via cd) ***)
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    94
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    95
Goalw [diamond_def,commute_def,square_def] "diamond(par_beta)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3207
diff changeset
    96
by (blast_tac (claset() addIs [par_beta_cd]) 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    97
qed "diamond_par_beta2";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    98
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4089
diff changeset
    99
Goal "confluent(beta)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   100
by (blast_tac (HOL_cs addIs [diamond_par_beta2, diamond_to_confluence,
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2637
diff changeset
   101
			     par_beta_subset_beta, beta_subset_par_beta]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   102
qed"beta_confluent";