| author | blanchet | 
| Wed, 17 Feb 2016 11:54:34 +0100 | |
| changeset 62326 | 3cf7a067599c | 
| parent 61585 | a9599d3d7610 | 
| child 81706 | 7beb0cf38292 | 
| permissions | -rw-r--r-- | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
2  | 
|
| 60500 | 3  | 
section \<open>Debugging facilities for code generated towards Isabelle/ML\<close>  | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
5  | 
theory Debug  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
6  | 
imports Main  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
7  | 
begin  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
8  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
9  | 
context  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
10  | 
begin  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
11  | 
|
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
12  | 
qualified definition trace :: "String.literal \<Rightarrow> unit" where  | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
13  | 
[simp]: "trace s = ()"  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
14  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
15  | 
qualified definition tracing :: "String.literal \<Rightarrow> 'a \<Rightarrow> 'a" where  | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
16  | 
[simp]: "tracing s = id"  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
18  | 
lemma [code]:  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
19  | 
"tracing s = (let u = trace s in id)"  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
20  | 
by simp  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
21  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
22  | 
qualified definition flush :: "'a \<Rightarrow> unit" where  | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
23  | 
[simp]: "flush x = ()"  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
24  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
25  | 
qualified definition flushing :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" where  | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
26  | 
[simp]: "flushing x = id"  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
28  | 
lemma [code, code_unfold]:  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
29  | 
"flushing x = (let u = flush x in id)"  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
30  | 
by simp  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
31  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
32  | 
qualified definition timing :: "String.literal \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
 | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
33  | 
[simp]: "timing s f x = f x"  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
34  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
35  | 
end  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
36  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51929 
diff
changeset
 | 
37  | 
code_printing  | 
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
38  | 
constant Debug.trace \<rightharpoonup> (Eval) "Output.tracing"  | 
| 61585 | 39  | 
| constant Debug.flush \<rightharpoonup> (Eval) "Output.tracing/ (@{make'_string} _)" \<comment> \<open>note indirection via antiquotation\<close>
 | 
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60500 
diff
changeset
 | 
40  | 
| constant Debug.timing \<rightharpoonup> (Eval) "Timing.timeap'_msg"  | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
41  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51929 
diff
changeset
 | 
42  | 
code_reserved Eval Output Timing  | 
| 
48427
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
44  | 
end  | 
| 
 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 
haftmann 
parents:  
diff
changeset
 | 
45  |