28 |
28 |
29 goal UTLemmas.thy "t <: u & u <: v --> t <: v"; |
29 goal UTLemmas.thy "t <: u & u <: v --> t <: v"; |
30 by (uterm_ind_tac "v" 1); |
30 by (uterm_ind_tac "v" 1); |
31 by (ALLGOALS (simp_tac utlemmas_ss)); |
31 by (ALLGOALS (simp_tac utlemmas_ss)); |
32 by (fast_tac HOL_cs 1); |
32 by (fast_tac HOL_cs 1); |
33 val occs_trans = conjI RS (result() RS mp); |
33 val occs_trans = store_thm("occs_trans", conjI RS (result() RS mp)); |
34 |
34 |
35 goal UTLemmas.thy "~ t <: v --> ~ t <: u | ~ u <: v"; |
35 goal UTLemmas.thy "~ t <: v --> ~ t <: u | ~ u <: v"; |
36 by (fast_tac (HOL_cs addIs [occs_trans]) 1); |
36 by (fast_tac (HOL_cs addIs [occs_trans]) 1); |
37 val contr_occs_trans = result() RS mp; |
37 val contr_occs_trans = store_thm("contr_occs_trans", result() RS mp); |
38 |
38 |
39 goal UTLemmas.thy "t <: Comb(t,u)"; |
39 goal UTLemmas.thy "t <: Comb(t,u)"; |
40 by (simp_tac utlemmas_ss 1); |
40 by (simp_tac utlemmas_ss 1); |
41 val occs_Comb1 = result(); |
41 qed "occs_Comb1"; |
42 |
42 |
43 goal UTLemmas.thy "u <: Comb(t,u)"; |
43 goal UTLemmas.thy "u <: Comb(t,u)"; |
44 by (simp_tac utlemmas_ss 1); |
44 by (simp_tac utlemmas_ss 1); |
45 val occs_Comb2 = result(); |
45 qed "occs_Comb2"; |
46 |
46 |
47 goal HOL.thy "(~(P|Q)) = (~P & ~Q)"; |
47 goal HOL.thy "(~(P|Q)) = (~P & ~Q)"; |
48 by (fast_tac HOL_cs 1); |
48 by (fast_tac HOL_cs 1); |
49 val demorgan_disj = result(); |
49 qed "demorgan_disj"; |
50 |
50 |
51 goal UTLemmas.thy "~ t <: t"; |
51 goal UTLemmas.thy "~ t <: t"; |
52 by (uterm_ind_tac "t" 1); |
52 by (uterm_ind_tac "t" 1); |
53 by (ALLGOALS (simp_tac (utlemmas_ss addsimps [demorgan_disj]))); |
53 by (ALLGOALS (simp_tac (utlemmas_ss addsimps [demorgan_disj]))); |
54 by (REPEAT (resolve_tac [impI,conjI] 1 ORELSE |
54 by (REPEAT (resolve_tac [impI,conjI] 1 ORELSE |
55 (etac contrapos 1 THEN etac subst 1 THEN |
55 (etac contrapos 1 THEN etac subst 1 THEN |
56 resolve_tac [occs_Comb1,occs_Comb2] 1) ORELSE |
56 resolve_tac [occs_Comb1,occs_Comb2] 1) ORELSE |
57 (etac (contr_occs_trans RS disjE) 1 THEN assume_tac 2) ORELSE |
57 (etac (contr_occs_trans RS disjE) 1 THEN assume_tac 2) ORELSE |
58 eresolve_tac ([occs_Comb1,occs_Comb2] RLN(2,[notE])) 1)); |
58 eresolve_tac ([occs_Comb1,occs_Comb2] RLN(2,[notE])) 1)); |
59 val occs_irrefl = result(); |
59 qed "occs_irrefl"; |
60 |
60 |
61 goal UTLemmas.thy "t <: u --> ~t=u"; |
61 goal UTLemmas.thy "t <: u --> ~t=u"; |
62 by (fast_tac (HOL_cs addEs [occs_irrefl RS notE]) 1); |
62 by (fast_tac (HOL_cs addEs [occs_irrefl RS notE]) 1); |
63 val occs_irrefl2 = result() RS mp; |
63 val occs_irrefl2 = store_thm("occs_irrefl2", result() RS mp); |
64 |
64 |
65 |
65 |
66 (**** vars_of lemmas ****) |
66 (**** vars_of lemmas ****) |
67 |
67 |
68 goal UTLemmas.thy "(v : vars_of(Var(w))) = (w=v)"; |
68 goal UTLemmas.thy "(v : vars_of(Var(w))) = (w=v)"; |
69 by (simp_tac utlemmas_ss 1); |
69 by (simp_tac utlemmas_ss 1); |
70 by (fast_tac HOL_cs 1); |
70 by (fast_tac HOL_cs 1); |
71 val vars_var_iff = result(); |
71 qed "vars_var_iff"; |
72 |
72 |
73 goal UTLemmas.thy "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)"; |
73 goal UTLemmas.thy "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)"; |
74 by (uterm_ind_tac "t" 1); |
74 by (uterm_ind_tac "t" 1); |
75 by (ALLGOALS (simp_tac utlemmas_ss)); |
75 by (ALLGOALS (simp_tac utlemmas_ss)); |
76 by (fast_tac HOL_cs 1); |
76 by (fast_tac HOL_cs 1); |
77 val vars_iff_occseq = result(); |
77 qed "vars_iff_occseq"; |