| author | Andreas Lochbihler | 
| Wed, 11 Feb 2015 15:03:21 +0100 | |
| changeset 59523 | 860fb1c65553 | 
| parent 58886 | 8a6cac7c7247 | 
| child 61361 | 8b5f00202e1a | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/J/JBasis.thy | 
| 41589 | 2 | Author: David von Oheimb, TU Muenchen | 
| 11070 | 3 | *) | 
| 8011 | 4 | |
| 58886 | 5 | chapter {* Java Source Language \label{cha:j} *}
 | 
| 6 | ||
| 7 | section {* Some Auxiliary Definitions *}
 | |
| 8011 | 8 | |
| 44035 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 Andreas Lochbihler parents: 
41589diff
changeset | 9 | theory JBasis imports Main "~~/src/HOL/Library/Transitive_Closure_Table" begin | 
| 8011 | 10 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 11 | lemmas [simp] = Let_def | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 12 | |
| 58886 | 13 | subsection "unique" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 14 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
24783diff
changeset | 15 | definition unique :: "('a \<times> 'b) list => bool" where
 | 
| 45634 | 16 | "unique == distinct \<circ> map fst" | 
| 8011 | 17 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 18 | |
| 45634 | 19 | lemma fst_in_set_lemma: "(x, y) : set xys ==> x : fst ` set xys" | 
| 20 | by (induct xys) auto | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 21 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 22 | lemma unique_Nil [simp]: "unique []" | 
| 45634 | 23 | by (simp add: unique_def) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 24 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 25 | lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))" | 
| 45634 | 26 | by (auto simp: unique_def dest: fst_in_set_lemma) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 27 | |
| 45634 | 28 | lemma unique_append: "unique l' ==> unique l ==> | 
| 29 | (!(x,y):set l. !(x',y'):set l'. x' ~= x) ==> unique (l @ l')" | |
| 30 | by (induct l) (auto dest: fst_in_set_lemma) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 31 | |
| 45634 | 32 | lemma unique_map_inj: "unique l ==> inj f ==> unique (map (%(k,x). (f k, g k x)) l)" | 
| 33 | by (induct l) (auto dest: fst_in_set_lemma simp add: inj_eq) | |
| 34 | ||
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 35 | |
| 58886 | 36 | subsection "More about Maps" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 37 | |
| 45634 | 38 | lemma map_of_SomeI: "unique l ==> (k, x) : set l ==> map_of l k = Some x" | 
| 39 | by (induct l) auto | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 40 | |
| 45634 | 41 | lemma Ball_set_table: "(\<forall>(x,y)\<in>set l. P x y) ==> (\<forall>x. \<forall>y. map_of l x = Some y --> P x y)" | 
| 42 | by (induct l) auto | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 43 | |
| 45634 | 44 | lemma table_of_remap_SomeD: | 
| 45 | "map_of (map (\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) ==> | |
| 46 | map_of t (k, k') = Some x" | |
| 47 | by (atomize (full), induct t) auto | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10042diff
changeset | 48 | |
| 8011 | 49 | end |