src/HOL/MicroJava/J/JBasis.thy
author haftmann
Tue, 03 Jan 2006 11:33:18 +0100
changeset 18552 30911da9fb27
parent 18447 da548623916a
child 18576 8d98b7711e47
permissions -rw-r--r--
class now a keyword
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/J/JBasis.thy
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     2
    ID:         $Id$
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     3
    Author:     David von Oheimb
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     4
    Copyright   1999 TU Muenchen
11070
cc421547e744 improved document (added headers etc)
oheimb
parents: 11026
diff changeset
     5
*)
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     6
12911
704713ca07ea new document
kleing
parents: 12888
diff changeset
     7
header {* 
704713ca07ea new document
kleing
parents: 12888
diff changeset
     8
  \chapter{Java Source Language}\label{cha:j}
704713ca07ea new document
kleing
parents: 12888
diff changeset
     9
  \isaheader{Some Auxiliary Definitions}
704713ca07ea new document
kleing
parents: 12888
diff changeset
    10
*}
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    11
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12911
diff changeset
    12
theory JBasis imports Main begin 
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    13
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    14
lemmas [simp] = Let_def
18447
da548623916a removed or modified some instances of [iff]
paulson
parents: 16417
diff changeset
    15
declare not_None_eq [iff]
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    16
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    17
section "unique"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    18
 
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    19
constdefs
11070
cc421547e744 improved document (added headers etc)
oheimb
parents: 11026
diff changeset
    20
  unique  :: "('a \<times> 'b) list => bool"
12888
f6c1e7306c40 nodups -> distinct
nipkow
parents: 12517
diff changeset
    21
  "unique  == distinct \<circ> map fst"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    22
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    23
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    24
lemma fst_in_set_lemma [rule_format (no_asm)]: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    25
      "(x, y) : set xys --> x : fst ` set xys"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    26
apply (induct_tac "xys")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    27
apply  auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    28
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    29
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    30
lemma unique_Nil [simp]: "unique []"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    31
apply (unfold unique_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    32
apply (simp (no_asm))
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    33
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    34
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    35
lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    36
apply (unfold unique_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    37
apply (auto dest: fst_in_set_lemma)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    38
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    39
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    40
lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l --> 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    41
 (!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    42
apply (induct_tac "l")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    43
apply  (auto dest: fst_in_set_lemma)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    44
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    45
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    46
lemma unique_map_inj [rule_format (no_asm)]: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    47
  "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    48
apply (induct_tac "l")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    49
apply  (auto dest: fst_in_set_lemma simp add: inj_eq)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    50
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    51
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    52
section "More about Maps"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    53
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    54
lemma map_of_SomeI [rule_format (no_asm)]: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    55
  "unique l --> (k, x) : set l --> map_of l k = Some x"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    56
apply (induct_tac "l")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    57
apply  auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    58
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    59
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    60
lemma Ball_set_table_: 
11070
cc421547e744 improved document (added headers etc)
oheimb
parents: 11026
diff changeset
    61
  "(\<forall>(x,y)\<in>set l. P x y) --> (\<forall>x. \<forall>y. map_of l x = Some y --> P x y)"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    62
apply(induct_tac "l")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    63
apply(simp_all (no_asm))
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    64
apply safe
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    65
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    66
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    67
lemmas Ball_set_table = Ball_set_table_ [THEN mp];
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    68
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    69
lemma table_of_remap_SomeD [rule_format (no_asm)]: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    70
"map_of (map (\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) --> 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    71
 map_of t (k, k') = Some x"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    72
apply (induct_tac "t")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    73
apply  auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    74
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10042
diff changeset
    75
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    76
end