| author | haftmann | 
| Sat, 11 Nov 2017 18:33:35 +0000 | |
| changeset 67050 | 1e29e2666a15 | 
| parent 63146 | f1ecba0272f9 | 
| permissions | -rw-r--r-- | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
20625 
diff
changeset
 | 
1  | 
(* Title: HOL/UNITY/Comp/Counterc.thy  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
2  | 
Author: Sidi O Ehmety, Cambridge University Computer Laboratory  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Copyright 2001 University of Cambridge  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
4  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
20625 
diff
changeset
 | 
5  | 
A family of similar counters, version with a full use of  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
20625 
diff
changeset
 | 
6  | 
"compatibility ".  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
8  | 
From Charpentier and Chandy,  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
9  | 
Examples of Program Composition Illustrating the Use of Universal Properties  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
10  | 
In J. Rolim (editor), Parallel and Distributed Processing,  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
11  | 
Spriner LNCS 1586 (1999), pages 1215-1227.  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
12  | 
*)  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
13  | 
|
| 63146 | 14  | 
section\<open>A Family of Similar Counters: Version with Compatibility\<close>  | 
| 14094 | 15  | 
|
| 18556 | 16  | 
theory Counterc imports "../UNITY_Main" begin  | 
| 13792 | 17  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
18  | 
typedecl state  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
20  | 
consts  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
21  | 
C :: "state=>int"  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
22  | 
c :: "state=>nat=>int"  | 
| 
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
23  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
24  | 
primrec sum :: "[nat,state]=>int" where  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
25  | 
  (* sum I s = sigma_{i<I}. c s i *)
 | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11701 
diff
changeset
 | 
26  | 
"sum 0 s = 0"  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
27  | 
| "sum (Suc i) s = (c s) i + sum i s"  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
28  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
29  | 
primrec sumj :: "[nat, nat, state]=>int" where  | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11701 
diff
changeset
 | 
30  | 
"sumj 0 i s = 0"  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
31  | 
| "sumj (Suc n) i s = (if n=i then sum n s else (c s) n + sumj n i s)"  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
32  | 
|
| 42463 | 33  | 
type_synonym command = "(state*state)set"  | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
34  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
35  | 
definition a :: "nat=>command" where  | 
| 36866 | 36  | 
 "a i = {(s, s'). (c s') i = (c s) i + 1 & (C s') = (C s) + 1}"
 | 
| 
11194
 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
37  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
38  | 
definition Component :: "nat => state program" where  | 
| 36866 | 39  | 
  "Component i = mk_total_program({s. C s = 0 & (c s) i = 0},
 | 
40  | 
                                  {a i},
 | 
|
41  | 
\<Union>G \<in> preserves (%s. (c s) i). Acts G)"  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
42  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
43  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
44  | 
declare Component_def [THEN def_prg_Init, simp]  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
45  | 
declare Component_def [THEN def_prg_AllowedActs, simp]  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
46  | 
declare a_def [THEN def_act_simp, simp]  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
47  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
48  | 
(* Theorems about sum and sumj *)  | 
| 46911 | 49  | 
lemma sum_sumj_eq1: "I<i ==> sum I s = sumj I i s"  | 
50  | 
by (induct I) auto  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
51  | 
|
| 46911 | 52  | 
lemma sum_sumj_eq2: "i<I ==> sum I s = c s i + sumj I i s"  | 
53  | 
by (induct I) (auto simp add: linorder_neq_iff sum_sumj_eq1)  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
54  | 
|
| 46911 | 55  | 
lemma sum_ext: "(\<And>i. i<I \<Longrightarrow> c s' i = c s i) ==> sum I s' = sum I s"  | 
56  | 
by (induct I) auto  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
57  | 
|
| 46911 | 58  | 
lemma sumj_ext: "(\<And>j. j<I ==> j\<noteq>i ==> c s' j = c s j) ==> sumj I i s' = sumj I i s"  | 
59  | 
by (induct I) (auto intro!: sum_ext)  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
60  | 
|
| 46911 | 61  | 
lemma sum0: "(\<And>i. i<I ==> c s i = 0) ==> sum I s = 0"  | 
62  | 
by (induct I) auto  | 
|
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
63  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
64  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
65  | 
(* Safety properties for Components *)  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
66  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
67  | 
lemma Component_ok_iff:  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
68  | 
"(Component i ok G) =  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
69  | 
(G \<in> preserves (%s. c s i) & Component i \<in> Allowed G)"  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
70  | 
apply (auto simp add: ok_iff_Allowed Component_def [THEN def_total_prg_Allowed])  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
71  | 
done  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
72  | 
declare Component_ok_iff [iff]  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
73  | 
declare OK_iff_ok [iff]  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
74  | 
declare preserves_def [simp]  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
75  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
76  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
77  | 
lemma p2: "Component i \<in> stable {s. C s = (c s) i + k}"
 | 
| 
16184
 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 
paulson 
parents: 
14094 
diff
changeset
 | 
78  | 
by (simp add: Component_def, safety)  | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
79  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
80  | 
lemma p3:  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
81  | 
"[| OK I Component; i\<in>I |]  | 
| 14094 | 82  | 
      ==> Component i \<in> stable {s. \<forall>j\<in>I. j\<noteq>i --> c s j = c k j}"
 | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
83  | 
apply simp  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
84  | 
apply (unfold Component_def mk_total_program_def)  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
85  | 
apply (simp (no_asm_use) add: stable_def constrains_def)  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
86  | 
apply blast  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
87  | 
done  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
88  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
89  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
90  | 
lemma p2_p3_lemma1:  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
91  | 
     "[| OK {i. i<I} Component; i<I |] ==>  
 | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
92  | 
      \<forall>k. Component i \<in> stable ({s. C s = c s i + sumj I i k} Int  
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
20625 
diff
changeset
 | 
93  | 
                                {s. \<forall>j\<in>{i. i<I}. j\<noteq>i --> c s j = c k j})"
 | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
94  | 
by (blast intro: stable_Int [OF p2 p3])  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
95  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
96  | 
lemma p2_p3_lemma2:  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
97  | 
     "(\<forall>k. F \<in> stable ({s. C s = (c s) i + sumj I i k} Int  
 | 
| 14094 | 98  | 
                        {s. \<forall>j\<in>{i. i<I}. j\<noteq>i --> c s j = c k j}))   
 | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
99  | 
      ==> (F \<in> stable {s. C s = c s i + sumj I i s})"
 | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
100  | 
apply (simp add: constrains_def stable_def)  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
101  | 
apply (force intro!: sumj_ext)  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
102  | 
done  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
103  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
104  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
105  | 
lemma p2_p3:  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
106  | 
     "[| OK {i. i<I} Component; i<I |]  
 | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
107  | 
      ==> Component i \<in> stable {s. C s = c s i + sumj I i s}"
 | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
108  | 
by (blast intro: p2_p3_lemma1 [THEN p2_p3_lemma2])  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
109  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
110  | 
|
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
111  | 
(* Compositional correctness *)  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
112  | 
lemma safety:  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
113  | 
     "[| 0<I; OK {i. i<I} Component |]   
 | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
114  | 
      ==> (\<Squnion>i\<in>{i. i<I}. (Component i)) \<in> invariant {s. C s = sum I s}"
 | 
| 14094 | 115  | 
apply (simp (no_asm) add: invariant_def JN_stable sum_sumj_eq2)  | 
| 
14088
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
116  | 
apply (auto intro!: sum0 p2_p3)  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
117  | 
done  | 
| 
 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 
paulson 
parents: 
13812 
diff
changeset
 | 
118  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32960 
diff
changeset
 | 
119  | 
end  |