Theory Type
chapter ‹Concepts for all JinjaThreads Languages \label{cha:j}›
section ‹JinjaThreads types›
theory Type
imports
"../Basic/Auxiliary"
begin
type_synonym cname = String.literal
type_synonym mname = String.literal
type_synonym vname = String.literal
definition Object :: cname
where "Object ≡ STR ''java/lang/Object''"
definition Thread :: cname
where "Thread ≡ STR ''java/lang/Thread''"
definition Throwable :: cname
where "Throwable ≡ STR ''java/lang/Throwable''"
definition this :: vname
where "this ≡ STR ''this''"
definition run :: mname
where "run ≡ STR ''run()V''"
definition start :: mname
where "start ≡ STR ''start()V''"
definition wait :: mname
where "wait ≡ STR ''wait()V''"
definition notify :: mname
where "notify ≡ STR ''notify()V''"
definition notifyAll :: mname
where "notifyAll ≡ STR ''notifyAll()V''"
definition join :: mname
where "join ≡ STR ''join()V''"
definition interrupt :: mname
where "interrupt ≡ STR ''interrupt()V''"
definition isInterrupted :: mname
where "isInterrupted ≡ STR ''isInterrupted()Z''"
definition hashcode :: mname
where "hashcode = STR ''hashCode()I''"
definition clone :: mname
where "clone = STR ''clone()Ljava/lang/Object;''"
definition print :: mname
where "print = STR ''~print(I)V''"
definition currentThread :: mname
where "currentThread = STR ''~Thread.currentThread()Ljava/lang/Thread;''"
definition interrupted :: mname
where "interrupted = STR ''~Thread.interrupted()Z''"
definition yield :: mname
where "yield = STR ''~Thread.yield()V''"
lemmas identifier_name_defs [code_unfold] =
this_def run_def start_def wait_def notify_def notifyAll_def join_def interrupt_def isInterrupted_def
hashcode_def clone_def print_def currentThread_def interrupted_def yield_def
lemma Object_Thread_Throwable_neq [simp]:
"Thread ≠ Object" "Object ≠ Thread"
"Object ≠ Throwable" "Throwable ≠ Object"
"Thread ≠ Throwable" "Throwable ≠ Thread"
by(auto simp add: Thread_def Object_def Throwable_def)
lemma synth_method_names_neq_aux:
"start ≠ wait" "start ≠ notify" "start ≠ notifyAll" "start ≠ join" "start ≠ interrupt" "start ≠ isInterrupted"
"start ≠ hashcode" "start ≠ clone" "start ≠ print" "start ≠ currentThread"
"start ≠ interrupted" "start ≠ yield" "start ≠ run"
"wait ≠ notify" "wait ≠ notifyAll" "wait ≠ join" "wait ≠ interrupt" "wait ≠ isInterrupted"
"wait ≠ hashcode" "wait ≠ clone" "wait ≠ print" "wait ≠ currentThread"
"wait ≠ interrupted" "wait ≠ yield" "wait ≠ run"
"notify ≠ notifyAll" "notify ≠ join" "notify ≠ interrupt" "notify ≠ isInterrupted"
"notify ≠ hashcode" "notify ≠ clone" "notify ≠ print" "notify ≠ currentThread"
"notify ≠ interrupted" "notify ≠ yield" "notify ≠ run"
"notifyAll ≠ join" "notifyAll ≠ interrupt" "notifyAll ≠ isInterrupted"
"notifyAll ≠ hashcode" "notifyAll ≠ clone" "notifyAll ≠ print" "notifyAll ≠ currentThread"
"notifyAll ≠ interrupted" "notifyAll ≠ yield" "notifyAll ≠ run"
"join ≠ interrupt" "join ≠ isInterrupted"
"join ≠ hashcode" "join ≠ clone" "join ≠ print" "join ≠ currentThread"
"join ≠ interrupted" "join ≠ yield" "join ≠ run"
"interrupt ≠ isInterrupted"
"interrupt ≠ hashcode" "interrupt ≠ clone" "interrupt ≠ print" "interrupt ≠ currentThread"
"interrupt ≠ interrupted" "interrupt ≠ yield" "interrupt ≠ run"
"isInterrupted ≠ hashcode" "isInterrupted ≠ clone" "isInterrupted ≠ print" "isInterrupted ≠ currentThread"
"isInterrupted ≠ interrupted" "isInterrupted ≠ yield" "isInterrupted ≠ run"
"hashcode ≠ clone" "hashcode ≠ print" "hashcode ≠ currentThread"
"hashcode ≠ interrupted" "hashcode ≠ yield" "hashcode ≠ run"
"clone ≠ print" "clone ≠ currentThread"
"clone ≠ interrupted" "clone ≠ yield" "clone ≠ run"
"print ≠ currentThread"
"print ≠ interrupted" "print ≠ yield" "print ≠ run"
"currentThread ≠ interrupted" "currentThread ≠ yield" "currentThread ≠ run"
"interrupted ≠ yield" "interrupted ≠ run"
"yield ≠ run"
by(simp_all add: identifier_name_defs)
lemmas synth_method_names_neq [simp] = synth_method_names_neq_aux synth_method_names_neq_aux[symmetric]