diff -r d09d0f160888 -r 360e3215f029 src/HOL/MicroJava/J/Conform.thy --- a/src/HOL/MicroJava/J/Conform.thy Sun Dec 16 00:17:44 2001 +0100 +++ b/src/HOL/MicroJava/J/Conform.thy Sun Dec 16 00:18:17 2001 +0100 @@ -8,14 +8,14 @@ theory Conform = State + WellType: -types 'c env_ = "'c prog \ (vname \ ty)" (* same as env of WellType.thy *) +types 'c env_ = "'c prog \ (vname \ ty)" -- "same as @{text env} of @{text WellType.thy}" constdefs hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50) "h<=|h' == \a C fs. h a = Some(C,fs) --> (\fs'. h' a = Some(C,fs'))" - conf :: "'c prog => aheap => val => ty => bool" + conf :: "'c prog => aheap => val => ty => bool" ("_,_ |- _ ::<= _" [51,51,51,51] 50) "G,h|-v::<=T == \T'. typeof (option_map obj_ty o h) v = Some T' \ G\T'\T" @@ -29,7 +29,7 @@ hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50) "G|-h h [ok] == \a obj. h a = Some obj --> G,h|-obj [ok]" - conforms :: "state => java_mb env_ => bool" ("_ ::<= _" [51,51] 50) + conforms :: "state => java_mb env_ => bool" ("_ ::<= _" [51,51] 50) "s::<=E == prg E|-h heap s [ok] \ prg E,heap s|-locals s[::<=]localT E" @@ -112,7 +112,8 @@ apply (simp) done -lemma defval_conf [rule_format (no_asm)]: "is_type G T --> G,h\default_val T::\T" +lemma defval_conf [rule_format (no_asm)]: + "is_type G T --> G,h\default_val T::\T" apply (unfold conf_def) apply (rule_tac "y" = "T" in ty.exhaust) apply (erule ssubst) @@ -127,7 +128,8 @@ apply auto done -lemma conf_widen [rule_format (no_asm)]: "wf_prog wf_mb G ==> G,h\x::\T --> G\T\T' --> G,h\x::\T'" +lemma conf_widen [rule_format (no_asm)]: + "wf_prog wf_mb G ==> G,h\x::\T --> G\T\T' --> G,h\x::\T'" apply (unfold conf_def) apply (rule val.induct) apply (auto intro: widen_trans) @@ -168,7 +170,8 @@ apply (fast dest: non_npD) done -lemma non_np_objD' [rule_format (no_asm)]: "a' \ Null ==> wf_prog wf_mb G ==> G,h\a'::\RefT t --> +lemma non_np_objD' [rule_format (no_asm)]: + "a' \ Null ==> wf_prog wf_mb G ==> G,h\a'::\RefT t --> (\C. t = ClassT C --> C \ Object) --> (\a C fs. a' = Addr a \ h a = Some (C,fs) \ G\Class C\RefT t)" apply(rule_tac "y" = "t" in ref_ty.exhaust) @@ -176,7 +179,9 @@ apply (fast dest: non_np_objD) done -lemma conf_list_gext_widen [rule_format (no_asm)]: "wf_prog wf_mb G ==> \Ts Ts'. list_all2 (conf G h) vs Ts --> list_all2 (\T T'. G\T\T') Ts Ts' --> list_all2 (conf G h) vs Ts'" +lemma conf_list_gext_widen [rule_format (no_asm)]: + "wf_prog wf_mb G ==> \Ts Ts'. list_all2 (conf G h) vs Ts --> + list_all2 (\T T'. G\T\T') Ts Ts' --> list_all2 (conf G h) vs Ts'" apply(induct_tac "vs") apply(clarsimp) apply(clarsimp) @@ -208,7 +213,8 @@ apply auto done -lemma lconf_init_vars_lemma [rule_format (no_asm)]: "\x. P x --> R (dv x) x ==> (\x. map_of fs f = Some x --> P x) --> +lemma lconf_init_vars_lemma [rule_format (no_asm)]: + "\x. P x --> R (dv x) x ==> (\x. map_of fs f = Some x --> P x) --> (\T. map_of fs f = Some T --> (\v. map_of (map (\(f,ft). (f, dv ft)) fs) f = Some v \ R v T))" apply( induct_tac "fs") @@ -231,7 +237,9 @@ apply auto done -lemma lconf_ext_list [rule_format (no_asm)]: "G,h\l[::\]L ==> \vs Ts. nodups vns --> length Ts = length vns --> list_all2 (\v T. G,h\v::\T) vs Ts --> G,h\l(vns[\]vs)[::\]L(vns[\]Ts)" +lemma lconf_ext_list [rule_format (no_asm)]: + "G,h\l[::\]L ==> \vs Ts. nodups vns --> length Ts = length vns --> + list_all2 (\v T. G,h\v::\T) vs Ts --> G,h\l(vns[\]vs)[::\]L(vns[\]Ts)" apply (unfold lconf_def) apply( induct_tac "vns") apply( clarsimp) @@ -291,12 +299,14 @@ apply( fast dest: conforms_localD elim!: conformsI lconf_hext) done -lemma conforms_upd_obj: "[|(h,l)::\(G, lT); G,h(a\obj)\obj\; h\|h(a\obj)|] ==> (h(a\obj),l)::\(G, lT)" -apply( rule conforms_hext) -apply auto -apply( rule hconfI) -apply( drule conforms_heapD) -apply( tactic {* auto_tac (HOL_cs addEs [thm "oconf_hext"] addDs [thm "hconfD"], simpset() delsimps [split_paired_All]) *}) +lemma conforms_upd_obj: + "[|(h,l)::\(G, lT); G,h(a\obj)\obj\; h\|h(a\obj)|] ==> (h(a\obj),l)::\(G, lT)" +apply(rule conforms_hext) +apply auto +apply(rule hconfI) +apply(drule conforms_heapD) +apply(tactic {* auto_tac (HOL_cs addEs [thm "oconf_hext"] + addDs [thm "hconfD"], simpset() delsimps [split_paired_All]) *}) done lemma conforms_upd_local: