src/HOL/UNITY/FP.ML
author nipkow
Tue, 23 Jun 1998 18:07:45 +0200
changeset 5071 548f398d770b
parent 5069 3ea049f7979d
child 5232 e5a7cdd07ea5
permissions -rw-r--r--
Consequences of the change from [ := ] to ( := ) in theory Update.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/FP
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
Fixed Point of a Program
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming", 1994
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    11
Goal "Union(B) Int A = (UN C:B. C Int A)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
qed "Int_Union2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
open FP;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    17
Goalw [FP_Orig_def, stable_def] "stable Acts (FP_Orig Acts Int B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    18
by (stac Int_Union2 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
by (rtac ball_constrains_UN 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
by (Simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
qed "stable_FP_Orig_Int";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    24
val prems = goalw thy [FP_Orig_def, stable_def]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
    "(!!B. stable Acts (A Int B)) ==> A <= FP_Orig Acts";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    26
by (blast_tac (claset() addIs prems) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    27
qed "FP_Orig_weakest";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    29
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    30
Goal "stable Acts (FP Acts Int B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
by (subgoal_tac "FP Acts Int B = (UN x:B. FP Acts Int {x})" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
by (Blast_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
by (asm_simp_tac (simpset() addsimps [Int_insert_right]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
by (rewrite_goals_tac [FP_def, stable_def]);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
by (rtac ball_constrains_UN 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    36
by (Simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    37
qed "stable_FP_Int";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    38
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    39
Goal "FP Acts <= FP_Orig Acts";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    40
by (rtac (stable_FP_Int RS FP_Orig_weakest) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
val lemma1 = result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    43
Goalw [FP_Orig_def, FP_def] "FP_Orig Acts <= FP Acts";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    44
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    45
by (dres_inst_tac [("x", "{x}")] spec 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    46
by (asm_full_simp_tac (simpset() addsimps [Int_insert_right]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    47
val lemma2 = result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    48
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    49
Goal "FP Acts = FP_Orig Acts";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    50
by (rtac ([lemma1,lemma2] MRS equalityI) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    51
qed "FP_equivalence";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    52
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    53
val [prem] = goal thy
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
    "(!!B. stable Acts (A Int B)) ==> A <= FP Acts";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    55
by (simp_tac (simpset() addsimps [FP_equivalence, prem RS FP_Orig_weakest]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    56
qed "FP_weakest";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    57
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    58
Goalw [FP_def, stable_def, constrains_def]
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    59
    "Compl (FP Acts) = (UN act:Acts. Compl{s. act^^{s} <= {s}})";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    60
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    61
qed "Compl_FP";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    62
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    63
Goal "A - (FP Acts) = (UN act:Acts. A - {s. act^^{s} <= {s}})";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    64
by (simp_tac (simpset() addsimps [Diff_eq, Compl_FP]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    65
qed "Diff_FP";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    66