| author | panny | 
| Tue, 03 Dec 2013 02:51:20 +0100 | |
| changeset 54629 | a692901ecdc2 | 
| parent 53707 | d1c6bff9ff58 | 
| child 54730 | de2d99b459b3 | 
| permissions | -rw-r--r-- | 
| 48929 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 1 | (* Title: Pure/Pure.thy | 
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 2 | Author: Makarius | 
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 3 | |
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 4 | Final stage of bootstrapping Pure, based on implicit background theory. | 
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 5 | *) | 
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 6 | |
| 48638 | 7 | theory Pure | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 8 | keywords | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 9 |     "!!" "!" "%" "(" ")" "+" "," "--" ":" "::" ";" "<" "<=" "=" "=="
 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 10 | "=>" "?" "[" "\<equiv>" "\<leftharpoondown>" "\<rightharpoonup>" | 
| 52143 | 11 | "\<rightleftharpoons>" "\<subseteq>" "]" "and" "assumes" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 12 | "attach" "begin" "binder" "constrains" "defines" "fixes" "for" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 13 | "identifier" "if" "imports" "in" "includes" "infix" "infixl" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 14 | "infixr" "is" "keywords" "notes" "obtains" "open" "output" | 
| 51293 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 wenzelm parents: 
51274diff
changeset | 15 | "overloaded" "pervasive" "shows" "structure" "unchecked" "where" "|" | 
| 52449 
79e7fd57acc4
recover command tags from 4cf3f6153eb8 -- important for document preparation;
 wenzelm parents: 
52439diff
changeset | 16 | and "theory" :: thy_begin % "theory" | 
| 
79e7fd57acc4
recover command tags from 4cf3f6153eb8 -- important for document preparation;
 wenzelm parents: 
52439diff
changeset | 17 | and "ML_file" :: thy_load % "ML" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 18 | and "header" :: diag | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 19 | and "chapter" :: thy_heading1 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 20 | and "section" :: thy_heading2 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 21 | and "subsection" :: thy_heading3 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 22 | and "subsubsection" :: thy_heading4 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 23 | and "text" "text_raw" :: thy_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 24 | and "sect" :: prf_heading2 % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 25 | and "subsect" :: prf_heading3 % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 26 | and "subsubsect" :: prf_heading4 % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 27 | and "txt" "txt_raw" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 28 | and "classes" "classrel" "default_sort" "typedecl" "type_synonym" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 29 | "nonterminal" "arities" "judgment" "consts" "syntax" "no_syntax" | 
| 51313 | 30 | "translations" "no_translations" "defs" "definition" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 31 | "abbreviation" "type_notation" "no_type_notation" "notation" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 32 | "no_notation" "axiomatization" "theorems" "lemmas" "declare" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 33 | "hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl | 
| 51295 | 34 | and "ML" :: thy_decl % "ML" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 35 | and "ML_prf" :: prf_decl % "proof" (* FIXME % "ML" ?? *) | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 36 | and "ML_val" "ML_command" :: diag % "ML" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 37 | and "setup" "local_setup" "attribute_setup" "method_setup" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 38 | "declaration" "syntax_declaration" "simproc_setup" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 39 | "parse_ast_translation" "parse_translation" "print_translation" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 40 | "typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 41 | and "bundle" :: thy_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 42 | and "include" "including" :: prf_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 43 | and "print_bundles" :: diag | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 44 | and "context" "locale" :: thy_decl | 
| 51224 
c3e99efacb67
back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
 wenzelm parents: 
50603diff
changeset | 45 | and "sublocale" "interpretation" :: thy_goal | 
| 
c3e99efacb67
back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
 wenzelm parents: 
50603diff
changeset | 46 | and "interpret" :: prf_goal % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 47 | and "class" :: thy_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 48 | and "subclass" :: thy_goal | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 49 | and "instantiation" :: thy_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 50 | and "instance" :: thy_goal | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 51 | and "overloading" :: thy_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 52 | and "code_datatype" :: thy_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 53 | and "theorem" "lemma" "corollary" :: thy_goal | 
| 51274 
cfc83ad52571
discontinued pointless command category "thy_schematic_goal" -- this is checked dynamically;
 wenzelm parents: 
51270diff
changeset | 54 | and "schematic_theorem" "schematic_lemma" "schematic_corollary" :: thy_goal | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 55 | and "notepad" :: thy_decl | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 56 | and "have" :: prf_goal % "proof" | 
| 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 57 | and "hence" :: prf_goal % "proof" == "then have" | 
| 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 58 | and "show" :: prf_asm_goal % "proof" | 
| 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 59 | and "thus" :: prf_asm_goal % "proof" == "then show" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 60 | and "then" "from" "with" :: prf_chain % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 61 | and "note" "using" "unfolding" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 62 | and "fix" "assume" "presume" "def" :: prf_asm % "proof" | 
| 53371 
47b23c582127
more explicit indication of 'guess' as improper Isar (aka "script") element;
 wenzelm parents: 
52549diff
changeset | 63 | and "obtain" :: prf_asm_goal % "proof" | 
| 
47b23c582127
more explicit indication of 'guess' as improper Isar (aka "script") element;
 wenzelm parents: 
52549diff
changeset | 64 | and "guess" :: prf_asm_goal_script % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 65 | and "let" "write" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 66 | and "case" :: prf_asm % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 67 |   and "{" :: prf_open % "proof"
 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 68 | and "}" :: prf_close % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 69 | and "next" :: prf_block % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 70 | and "qed" :: qed_block % "proof" | 
| 53571 
e58ca0311c0f
more explicit indication of 'done' as proof script element;
 wenzelm parents: 
53371diff
changeset | 71 | and "by" ".." "." "sorry" :: "qed" % "proof" | 
| 
e58ca0311c0f
more explicit indication of 'done' as proof script element;
 wenzelm parents: 
53371diff
changeset | 72 | and "done" :: "qed_script" % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 73 | and "oops" :: qed_global % "proof" | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 74 | and "defer" "prefer" "apply" :: prf_script % "proof" | 
| 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 75 | and "apply_end" :: prf_script % "proof" == "" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 76 | and "proof" :: prf_block % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 77 | and "also" "moreover" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 78 | and "finally" "ultimately" :: prf_chain % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 79 | and "back" :: prf_script % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 80 | and "Isabelle.command" :: control | 
| 52060 | 81 | and "help" "print_commands" "print_options" | 
| 51585 | 82 | "print_context" "print_theory" "print_syntax" "print_abbrevs" "print_defn_rules" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 83 | "print_theorems" "print_locales" "print_classes" "print_locale" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 84 | "print_interps" "print_dependencies" "print_attributes" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 85 | "print_simpset" "print_rules" "print_trans_rules" "print_methods" | 
| 49569 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
48929diff
changeset | 86 | "print_antiquotations" "thy_deps" "locale_deps" "class_deps" "thm_deps" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 87 | "print_binds" "print_facts" "print_cases" "print_statement" "thm" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 88 | "prf" "full_prf" "prop" "term" "typ" "print_codesetup" "unused_thms" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 89 | :: diag | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 90 | and "cd" :: control | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 91 | and "pwd" :: diag | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 92 | and "use_thy" "remove_thy" "kill_thy" :: control | 
| 52549 | 93 | and "display_drafts" "print_state" "pr" :: diag | 
| 52438 
7b5a5116f3af
back to keyword 'pr' :: diag as required for ProofGeneral command line -- reject this TTY command in Isabelle/jEdit by other means;
 wenzelm parents: 
52437diff
changeset | 94 | and "pretty_setmargin" "disable_pr" "enable_pr" "commit" "quit" "exit" :: control | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 95 | and "welcome" :: diag | 
| 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 96 | and "init_toplevel" "linear_undo" "undo" "undos_proof" "cannot_undo" "kill" :: control | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 97 | and "end" :: thy_end % "theory" | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 98 | and "realizers" "realizability" "extract_type" "extract" :: thy_decl | 
| 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 99 | and "find_theorems" "find_consts" :: diag | 
| 52437 
c88354589b43
more formal ProofGeneral command setup within theory Pure;
 wenzelm parents: 
52430diff
changeset | 100 | and "ProofGeneral.process_pgip" "ProofGeneral.pr" "ProofGeneral.undo" | 
| 
c88354589b43
more formal ProofGeneral command setup within theory Pure;
 wenzelm parents: 
52430diff
changeset | 101 | "ProofGeneral.restart" "ProofGeneral.kill_proof" "ProofGeneral.inform_file_processed" | 
| 
c88354589b43
more formal ProofGeneral command setup within theory Pure;
 wenzelm parents: 
52430diff
changeset | 102 | "ProofGeneral.inform_file_retracted" :: control | 
| 48638 | 103 | begin | 
| 15803 | 104 | |
| 48891 | 105 | ML_file "Isar/isar_syn.ML" | 
| 53707 | 106 | ML_file "Tools/rule_insts.ML"; | 
| 48891 | 107 | ML_file "Tools/find_theorems.ML" | 
| 108 | ML_file "Tools/find_consts.ML" | |
| 52009 | 109 | ML_file "Tools/proof_general_pure.ML" | 
| 48891 | 110 | |
| 111 | ||
| 26435 | 112 | section {* Further content for the Pure theory *}
 | 
| 20627 | 113 | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 114 | subsection {* Meta-level connectives in assumptions *}
 | 
| 15803 | 115 | |
| 116 | lemma meta_mp: | |
| 18019 | 117 | assumes "PROP P ==> PROP Q" and "PROP P" | 
| 15803 | 118 | shows "PROP Q" | 
| 18019 | 119 | by (rule `PROP P ==> PROP Q` [OF `PROP P`]) | 
| 15803 | 120 | |
| 23432 | 121 | lemmas meta_impE = meta_mp [elim_format] | 
| 122 | ||
| 15803 | 123 | lemma meta_spec: | 
| 26958 | 124 | assumes "!!x. PROP P x" | 
| 125 | shows "PROP P x" | |
| 126 | by (rule `!!x. PROP P x`) | |
| 15803 | 127 | |
| 128 | lemmas meta_allE = meta_spec [elim_format] | |
| 129 | ||
| 26570 | 130 | lemma swap_params: | 
| 26958 | 131 | "(!!x y. PROP P x y) == (!!y x. PROP P x y)" .. | 
| 26570 | 132 | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 133 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 134 | subsection {* Meta-level conjunction *}
 | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 135 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 136 | lemma all_conjunction: | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 137 | "(!!x. PROP A x &&& PROP B x) == ((!!x. PROP A x) &&& (!!x. PROP B x))" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 138 | proof | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 139 | assume conj: "!!x. PROP A x &&& PROP B x" | 
| 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 140 | show "(!!x. PROP A x) &&& (!!x. PROP B x)" | 
| 19121 | 141 | proof - | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 142 | fix x | 
| 26958 | 143 | from conj show "PROP A x" by (rule conjunctionD1) | 
| 144 | from conj show "PROP B x" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 145 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 146 | next | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 147 | assume conj: "(!!x. PROP A x) &&& (!!x. PROP B x)" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 148 | fix x | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 149 | show "PROP A x &&& PROP B x" | 
| 19121 | 150 | proof - | 
| 26958 | 151 | show "PROP A x" by (rule conj [THEN conjunctionD1, rule_format]) | 
| 152 | show "PROP B x" by (rule conj [THEN conjunctionD2, rule_format]) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 153 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 154 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 155 | |
| 19121 | 156 | lemma imp_conjunction: | 
| 50603 | 157 | "(PROP A ==> PROP B &&& PROP C) == ((PROP A ==> PROP B) &&& (PROP A ==> PROP C))" | 
| 18836 | 158 | proof | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 159 | assume conj: "PROP A ==> PROP B &&& PROP C" | 
| 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 160 | show "(PROP A ==> PROP B) &&& (PROP A ==> PROP C)" | 
| 19121 | 161 | proof - | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 162 | assume "PROP A" | 
| 19121 | 163 | from conj [OF `PROP A`] show "PROP B" by (rule conjunctionD1) | 
| 164 | from conj [OF `PROP A`] show "PROP C" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 165 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 166 | next | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 167 | assume conj: "(PROP A ==> PROP B) &&& (PROP A ==> PROP C)" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 168 | assume "PROP A" | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 169 | show "PROP B &&& PROP C" | 
| 19121 | 170 | proof - | 
| 171 | from `PROP A` show "PROP B" by (rule conj [THEN conjunctionD1]) | |
| 172 | from `PROP A` show "PROP C" by (rule conj [THEN conjunctionD2]) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 173 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 174 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 175 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 176 | lemma conjunction_imp: | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 177 | "(PROP A &&& PROP B ==> PROP C) == (PROP A ==> PROP B ==> PROP C)" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 178 | proof | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 179 | assume r: "PROP A &&& PROP B ==> PROP C" | 
| 22933 | 180 | assume ab: "PROP A" "PROP B" | 
| 181 | show "PROP C" | |
| 182 | proof (rule r) | |
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 183 | from ab show "PROP A &&& PROP B" . | 
| 22933 | 184 | qed | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 185 | next | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 186 | assume r: "PROP A ==> PROP B ==> PROP C" | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 187 | assume conj: "PROP A &&& PROP B" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 188 | show "PROP C" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 189 | proof (rule r) | 
| 19121 | 190 | from conj show "PROP A" by (rule conjunctionD1) | 
| 191 | from conj show "PROP B" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 192 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 193 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 194 | |
| 48638 | 195 | end | 
| 196 |